diff --git a/lib/common/Pulse.Lib.Core.fsti b/lib/common/Pulse.Lib.Core.fsti index 37bb8b146..fe6b57a68 100644 --- a/lib/common/Pulse.Lib.Core.fsti +++ b/lib/common/Pulse.Lib.Core.fsti @@ -613,6 +613,16 @@ let match_rewrite_tac (_:unit) : T.Tac unit = T.mapply (`slprop_equiv_ext'); T.trefl () +let match_rename_tac (_:unit) : T.Tac unit = + let open T in + let b = T.nth_var (-1) in + T.norm []; // Needed, or otherwise the `tcut` in grewrite_eq can fail, unsure why. + try ( + T.grewrite_eq b; + T.trefl () + ) + with _ -> T.smt() + [@@deprecated "Use (rewrite .. as .. by ..) syntax instead!"] val rewrite_by (p:slprop) (q:slprop) (t:unit -> T.Tac unit) diff --git a/lib/pulse/lib/Pulse.Lib.AVLTree.fst b/lib/pulse/lib/Pulse.Lib.AVLTree.fst index 11db8d361..cb4372623 100644 --- a/lib/pulse/lib/Pulse.Lib.AVLTree.fst +++ b/lib/pulse/lib/Pulse.Lib.AVLTree.fst @@ -120,6 +120,7 @@ fn cases_of_is_tree #t (x:tree_t t) (ft:T.tree t) T.Leaf -> { unfold (is_tree x T.Leaf); fold (is_tree_cases None ft); + rewrite is_tree_cases None ft as is_tree_cases x ft; } T.Node data ltree rtree -> { unfold (is_tree x (T.Node data ltree rtree)); @@ -127,7 +128,8 @@ fn cases_of_is_tree #t (x:tree_t t) (ft:T.tree t) with n. assert p |-> n; with l'. rewrite is_tree lct l' as is_tree n.left l'; with r'. rewrite is_tree rct r' as is_tree n.right r'; - fold (is_tree_cases (Some p) (T.Node data ltree rtree)) + fold (is_tree_cases (Some p) ft); + rewrite (is_tree_cases (Some p) ft) as is_tree_cases x ft; } } } @@ -240,8 +242,6 @@ fn node_cons (#t:Type0) (v:t) (ltree:tree_t t) (rtree:tree_t t) (#l:(T.tree t)) ensures is_tree y (T.Node v l r) ** (pure (Some? y)) { let y = Box.alloc { data=v; left =ltree; right = rtree }; - rewrite each ltree as ({data = v; left = ltree; right = rtree}).left in (is_tree ltree l); - rewrite each rtree as ({data = v; left = ltree; right = rtree}).right in (is_tree rtree r); intro_is_tree_node (Some y) y; Some y } @@ -293,18 +293,12 @@ fn rec append_left (#t:Type0) (x:tree_t t) (v:t) (#ft:G.erased (T.tree t)) is_tree_case_some (Some vl) vl; - with _node _ltree _rtree._; - let node = !vl; let left_new = append_left node.left v; vl := {node with left = left_new}; - rewrite each left_new as ({ node with left = left_new }).left in (is_tree left_new ((T.append_left (reveal _ltree) v))); - - rewrite each node.right as ({ node with left = left_new }).right in (is_tree node.right _rtree); - intro_is_tree_node x vl; x @@ -344,18 +338,12 @@ fn rec append_right (#t:Type0) (x:tree_t t) (v:t) (#ft:G.erased (T.tree t)) Some np -> { is_tree_case_some (Some np) np; - with _node _ltree _rtree._; - let node = !np; let right_new = append_right node.right v; np := {node with right = right_new}; - rewrite each right_new as ({ node with right = right_new }).right in (is_tree right_new ((T.append_right (reveal _rtree) v))); - - rewrite each node.left as ({node with right = right_new}).left in (is_tree node.left _ltree); - intro_is_tree_node x np; x diff --git a/lib/pulse/lib/Pulse.Lib.Borrow.fst b/lib/pulse/lib/Pulse.Lib.Borrow.fst index 264eaef84..f74adf21c 100644 --- a/lib/pulse/lib/Pulse.Lib.Borrow.fst +++ b/lib/pulse/lib/Pulse.Lib.Borrow.fst @@ -41,7 +41,7 @@ noeq type blockchain_root = { type lifetime : Type0 = ref blockchain_root -let fpts_to #t (r: ref t) (x: t) = exists* p. pts_to r #p x +let fpts_to #t ([@@@mkey] r: ref t) (x: t) = exists* p. pts_to r #p x ghost fn dup_fpts_to u#t (t: Type u#t) r x () : duplicable_f (fpts_to #t r x) = { unfold fpts_to r x; @@ -189,7 +189,7 @@ type stored_shape = | Stored | StoredBoth : stored_shape -> stored_shape -> stored_shape -let rec bc_stored ([@@@mkey] x: blockchain_edge) ([@@@mkey] d: stored_shape) (y: slprop) : Tot slprop (decreases d) = +let rec bc_stored ([@@@mkey] x: blockchain_edge) (d: stored_shape) (y: slprop) : Tot slprop (decreases d) = match d with | Stored -> exists* z. slprop_ref_pts_to x.be_prop z ** later z ** trade (later z) y | StoredCheckedOut -> live x.be_ref ** y @@ -199,7 +199,7 @@ let rec bc_stored ([@@@mkey] x: blockchain_edge) ([@@@mkey] d: stored_shape) (y: bc_stored b db yb ** trade (ya ** yb) y -let rec rt_stored ([@@@mkey] x: ref blockchain_root) ([@@@mkey] n: unat) (b: slprop) : Tot slprop (decreases n) = +let rec rt_stored ([@@@mkey] x: ref blockchain_root) (n: unat) (b: slprop) : Tot slprop (decreases n) = match n with | Zero -> trade emp b | Succ n -> exists* r (b1 b2: slprop). fpts_to x r ** @@ -207,7 +207,7 @@ let rec rt_stored ([@@@mkey] x: ref blockchain_root) ([@@@mkey] n: unat) (b: slp rt_stored r.rt_next n b2 ** trade (b1 ** b2) b -let owns_end ([@@@mkey] x: ref blockchain_root) ([@@@mkey] n: unat) = +let owns_end ([@@@mkey] x: ref blockchain_root) (n: unat) = exists* y. root_idx x n y ** live y ghost fn elim_owns_end_zero x @@ -574,7 +574,7 @@ ghost fn share_borrow' #a (p q1 q2: slprop) dup (slprop_ref_pts_to eb.be_prop q2) (); fold valid_split l.be_prop ea.be_prop eb.be_prop; dup (bc_idx r.rt_tree is l) (); - dup (valid_split l.be_prop ra rb) (); + dup (valid_split l.be_prop ea.be_prop eb.be_prop) (); push_bc_idx r.rt_tree is false l; push_bc_idx r.rt_tree is true l; dup (root_idx' a n r) (); diff --git a/lib/pulse/lib/Pulse.Lib.CancellableInvariant.fsti b/lib/pulse/lib/Pulse.Lib.CancellableInvariant.fsti index 019207389..3be6a34d6 100644 --- a/lib/pulse/lib/Pulse.Lib.CancellableInvariant.fsti +++ b/lib/pulse/lib/Pulse.Lib.CancellableInvariant.fsti @@ -25,9 +25,9 @@ val cinv : Type0 instance val non_informative_cinv : NonInformative.non_informative cinv -val cinv_vp (c:cinv) (v:slprop) : slprop +val cinv_vp ([@@@mkey] c:cinv) (v:slprop) : slprop -val active (c:cinv) (p:perm) : slprop +val active ([@@@mkey] c:cinv) (p:perm) : slprop val active_timeless (c:cinv) (p:perm) : Lemma (timeless (active c p)) diff --git a/lib/pulse/lib/Pulse.Lib.ConditionVar.fst b/lib/pulse/lib/Pulse.Lib.ConditionVar.fst index b42031871..5108a75e9 100644 --- a/lib/pulse/lib/Pulse.Lib.ConditionVar.fst +++ b/lib/pulse/lib/Pulse.Lib.ConditionVar.fst @@ -167,7 +167,7 @@ ensures later_intro (cvar_inv b.core p); drop_ (Box.pts_to b.core.r #0.5R _) }; - drop_ _ + drop_ (inv _ _) } fn signal (c:cvar_t) (#p:slprop) diff --git a/lib/pulse/lib/Pulse.Lib.Deque.fst b/lib/pulse/lib/Pulse.Lib.Deque.fst index 7df190b38..8f7ebceef 100644 --- a/lib/pulse/lib/Pulse.Lib.Deque.fst +++ b/lib/pulse/lib/Pulse.Lib.Deque.fst @@ -21,15 +21,6 @@ type deque (t:Type0) = { tail: option (node_ptr t); } -(* Note: since within this module there is usually a *single* linked list -around, we mark the list predicated with no_mkeys so the matcher can be -more liberal. Crucially, this attribute is only set behind the interface, -and clients will just use the mkey in is_deque. - -This is a bit of a hack, the fact that F* allows the attributes to differ -between fst/fsti is probably wrong. Maybe we should have a typeclass? *) - -[@@no_mkeys] let rec is_deque_suffix (#t:Type0) ([@@@mkey] p:node_ptr t) @@ -82,7 +73,6 @@ fn fold_is_deque_suffix_cons -[@@no_mkeys] let is_deque #t ([@@@mkey] x:deque t) (l:list t) : Tot slprop (decreases l) = match l with @@ -120,7 +110,7 @@ fn push_front_empty (#t:Type) (l : deque t) (x : t) }; let node = Box.alloc vnode; - fold (is_deque_suffix node [x] None node); + fold (is_deque_suffix node [x] None node None); let l' = { head = Some node; @@ -287,10 +277,8 @@ let is_deque_suffix_factored : Tot slprop = exists* (v:node t). pts_to x v ** - pure ( - v.value == List.Tot.hd l /\ - v.dprev == prev - ) ** + pure (v.value == List.Tot.hd l) ** + pure (v.dprev == prev) ** is_deque_suffix_factored_next x l tail last v.dnext @@ -655,7 +643,7 @@ fn pop_front (#t:Type) (l : deque t) { let b = is_singleton l; if b { - pop_front_nil l; + pop_front_nil l #x; } else { pop_front_cons l; } diff --git a/lib/pulse/lib/Pulse.Lib.Forall.fst b/lib/pulse/lib/Pulse.Lib.Forall.fst index 5a37257bf..7343bd87c 100644 --- a/lib/pulse/lib/Pulse.Lib.Forall.fst +++ b/lib/pulse/lib/Pulse.Lib.Forall.fst @@ -47,7 +47,7 @@ fn elim_forall u#a { unfold (forall* x. p x); unfold uquant (F.on_dom a (fun x -> p x)); - with v. assert token v; unfold token v; + with v. unfold token v; extract_q v (F.on_domain a (fun x -> p x)) () x; } diff --git a/lib/pulse/lib/Pulse.Lib.GhostPCMReference.fst b/lib/pulse/lib/Pulse.Lib.GhostPCMReference.fst index 5824a85d5..8d97fcbdf 100644 --- a/lib/pulse/lib/Pulse.Lib.GhostPCMReference.fst +++ b/lib/pulse/lib/Pulse.Lib.GhostPCMReference.fst @@ -63,7 +63,7 @@ fn read u#a returns v:(v:a { compatible p x v /\ p.refine v }) ensures pts_to r (f v) { - with inst. assert small_token u#a inst; let inst = inst; + with inst. unfold small_token u#a inst; let inst = inst; fold small_token inst; U.downgrade_val (C.ghost_read #(U.raise_t a) #(raise p) r (hide (U.raise_val (reveal x))) (raise_refine p x f)); } @@ -92,7 +92,7 @@ fn write u#a requires pts_to r x ensures pts_to r y { - with inst. assert small_token u#a inst; let inst = inst; + with inst. unfold small_token u#a inst; let inst = inst; fold small_token inst; C.ghost_write #(U.raise_t a) #(raise p) r (hide (U.raise_val (reveal x))) (hide (U.raise_val (reveal y))) (raise_upd f) } @@ -107,18 +107,12 @@ fn share u#a requires pts_to r (v0 `op pcm` v1) ensures pts_to r v0 ** pts_to r v1 { - with inst. assert small_token u#a inst; let inst = inst; - fold small_token u#a inst; + with inst. unfold small_token u#a inst; let inst = inst; + fold small_token inst; + fold small_token inst; C.ghost_share #_ #(PR.raise pcm) r (U.raise_val v0) (U.raise_val v1) } -[@@allow_ambiguous] -ghost fn drop_amb (p: slprop) - requires p -{ - drop_ p -} - [@@allow_ambiguous] ghost fn gather u#a @@ -131,11 +125,9 @@ fn gather u#a returns squash (composable pcm v0 v1) ensures pts_to r (op pcm v0 v1) { - with inst. assert C.ghost_pcm_pts_to #_ #(raise #a #inst pcm) r (U.raise_val v1); - with inst'. assert C.ghost_pcm_pts_to #_ #(raise #a #inst' pcm) r (U.raise_val v1); - drop_amb (small_token u#a inst'); - let inst = inst; - C.ghost_gather #_ #(PR.raise pcm) r (U.raise_val v0) (U.raise_val v1) + with inst. assert C.ghost_pcm_pts_to #_ #(raise #a #inst pcm) r (U.raise_val #a #inst v1); + drop_ (small_token inst); + C.ghost_gather #_ #(PR.raise #a #inst pcm) r (U.raise_val #a #inst v0) (U.raise_val #a #inst v1) } ghost fn pts_to_not_null u#a (#a:Type u#a) diff --git a/lib/pulse/lib/Pulse.Lib.HashTable.fst b/lib/pulse/lib/Pulse.Lib.HashTable.fst index 52f273282..7cc326023 100644 --- a/lib/pulse/lib/Pulse.Lib.HashTable.fst +++ b/lib/pulse/lib/Pulse.Lib.HashTable.fst @@ -556,7 +556,6 @@ fn delete if (voff = ht.sz) { cont := false; - rewrite each vcont as false; // also relying on #110 } else { @@ -584,7 +583,6 @@ fn delete { cont := false; assert (pure (pht.repr == (PHT.delete pht k).repr)); - rewrite each vcont as false; // also relying on #110 } Zombie -> { diff --git a/lib/pulse/lib/Pulse.Lib.HashTable.fsti b/lib/pulse/lib/Pulse.Lib.HashTable.fsti index 10ceb2df8..fa5616a4e 100644 --- a/lib/pulse/lib/Pulse.Lib.HashTable.fsti +++ b/lib/pulse/lib/Pulse.Lib.HashTable.fsti @@ -44,13 +44,13 @@ let related #kt #vt (ht:ht_t kt vt) (pht:pht_t kt vt) : prop = SZ.v ht.sz == pht.repr.sz /\ pht.repr.hashf == lift_hash_fun ht.hashf -let models #kt #vt (ht:ht_t kt vt) (pht:pht_t kt vt) : slprop = +let models #kt #vt ([@@@mkey]ht: ht_t kt vt) (pht:pht_t kt vt) : slprop = V.pts_to ht.contents pht.repr.seq ** pure (related ht pht /\ V.is_full_vec ht.contents /\ SZ.fits (2 `op_Multiply` SZ.v ht.sz)) -val models_timeless #kt #vt (ht:ht_t kt vt) (pht:pht_t kt vt) +val models_timeless #kt #vt ([@@@mkey] ht:ht_t kt vt) (pht:pht_t kt vt) : Lemma (timeless (models ht pht)) [SMTPat (timeless (models ht pht))] diff --git a/lib/pulse/lib/Pulse.Lib.InsertionSort.fst b/lib/pulse/lib/Pulse.Lib.InsertionSort.fst index 16fb5e267..fa9757419 100644 --- a/lib/pulse/lib/Pulse.Lib.InsertionSort.fst +++ b/lib/pulse/lib/Pulse.Lib.InsertionSort.fst @@ -134,6 +134,7 @@ let step_outer_invariant #pop-options +#push-options "--z3rlimit_factor 2" fn insertion_sort u#a (#t:Type u#a) {| total_order t |} @@ -202,6 +203,7 @@ ensures exists* s'. (a |-> s') ** with s'. assert (a |-> s'); assert pure (Seq.slice s' 0 (Seq.length s') `Seq.equal` s'); } +#pop-options instance total_order_int : total_order int = { compare = FStar.Order.compare_int; diff --git a/lib/pulse/lib/Pulse.Lib.LinkedList.fst b/lib/pulse/lib/Pulse.Lib.LinkedList.fst index 7bb5eea1c..cc6d26c5a 100644 --- a/lib/pulse/lib/Pulse.Lib.LinkedList.fst +++ b/lib/pulse/lib/Pulse.Lib.LinkedList.fst @@ -290,6 +290,9 @@ fn move_next (#t:Type) (x:llist t) let node = !np; intro_yields_cons np; rewrite each (Some np) as x; + with tl. assert is_list node.tail tl; + rewrite trade (is_list node.tail tl) (is_list x (node.head :: tl)) + as trade (is_list node.tail tl) (is_list x 'l); node.tail } @@ -408,13 +411,15 @@ fn non_empty_list (#t:Type0) (x:llist t) intro_is_list_cons x v #n #tl; } +let is_cons #t ([@@@mkey] x: list t) hd tl = pure (x == hd :: tl) + fn move_next_forall (#t:Type) (x:llist t) requires is_list x 'l ** pure (Some? x) returns y:llist t ensures exists* hd tl. is_list y tl ** (forall* tl'. is_list y tl' @==> is_list x (hd::tl')) ** - pure ('l == hd::tl) + is_cons 'l hd tl { let np = Some?.v x; is_list_cases_some x np; @@ -422,6 +427,8 @@ fn move_next_forall (#t:Type) (x:llist t) intro (forall* tl'. is_list node.tail tl' @==> is_list x (node.head::tl')) #(np |-> node) fn _ tl' { intro_is_list_cons x np; }; + with tl. assert is_list node.tail tl; + fold is_cons 'l node.head tl; node.tail } @@ -458,14 +465,14 @@ fn append_iter (#t:Type) (x y:llist t) with ll pfx sfx. _; some_iff_cons ll; let next = move_next_forall Pulse.Lib.Reference.(!cur); - with hd tl. _; + with hd tl. unfold is_cons sfx hd tl; (* this is the key induction step *) FA.trans_compose (is_list next) (is_list ll) (is_list x) - (fun tl -> hd :: tl) + (fun tl -> reveal hd :: tl) (fun tl -> pfx @ tl); - rewrite (forall* tl. is_list next tl @==> is_list x (pfx@(hd::tl))) - as (forall* tl. is_list next tl @==> is_list x ((pfx@[hd])@tl)); + rewrite (forall* tl. is_list next tl @==> is_list x (pfx@(reveal hd::tl))) + as (forall* tl. is_list next tl @==> is_list x ((pfx@[reveal hd])@tl)); Pulse.Lib.Reference.(cur := next); }; with ll pfx sfx. _; @@ -526,7 +533,7 @@ fn split (#t:Type0) (x:llist t) (n:U32.t) (#xl:erased (list t)) { with i ll pfx sfx. _; let next = move_next_forall Pulse.Lib.Reference.(!cur); - with hd tl. _; + with hd tl. unfold is_cons sfx hd tl; (* this is the key induction step *) FA.trans_compose (is_list next) (is_list ll) (is_list x) diff --git a/lib/pulse/lib/Pulse.Lib.MonotonicGhostRef.fst b/lib/pulse/lib/Pulse.Lib.MonotonicGhostRef.fst index a785b3df2..67257e136 100644 --- a/lib/pulse/lib/Pulse.Lib.MonotonicGhostRef.fst +++ b/lib/pulse/lib/Pulse.Lib.MonotonicGhostRef.fst @@ -131,5 +131,5 @@ fn update (#t:Type) (#p:preorder t) (r:mref p) (#u:t) (v:t) unfold pts_to; with f h. assert (GR.pts_to r (f, h)); GR.write r _ _ (FP.mk_frame_preserving_upd p h v); - fold pts_to; + fold pts_to r #1.0R v; } diff --git a/lib/pulse/lib/Pulse.Lib.Mutex.fst b/lib/pulse/lib/Pulse.Lib.Mutex.fst index 41278e713..6c2e25b8d 100644 --- a/lib/pulse/lib/Pulse.Lib.Mutex.fst +++ b/lib/pulse/lib/Pulse.Lib.Mutex.fst @@ -94,7 +94,7 @@ fn unlock (#a:Type0) (#v:a -> slprop) (#p:perm) (m:mutex a) (mg:mutex_guard a) unfold (mg `belongs_to` m); with x. rewrite (pts_to mg x) as (R.pts_to (B.box_to_ref m.r) x); B.to_box_pts_to m.r; - fold lock_inv; + fold lock_inv m.r v; release m.l; fold (mutex_live m #p v) } diff --git a/lib/pulse/lib/Pulse.Lib.Mutex.fsti b/lib/pulse/lib/Pulse.Lib.Mutex.fsti index 30f49f817..2600d3d6f 100644 --- a/lib/pulse/lib/Pulse.Lib.Mutex.fsti +++ b/lib/pulse/lib/Pulse.Lib.Mutex.fsti @@ -31,7 +31,7 @@ val mutex_guard (a:Type0) : Type0 val mutex_live (#a:Type0) - (m:mutex a) + ([@@@mkey] m:mutex a) (#[T.exact (`1.0R)] p:perm) (v:a -> slprop) : slprop diff --git a/lib/pulse/lib/Pulse.Lib.OnRange.fst b/lib/pulse/lib/Pulse.Lib.OnRange.fst index 66cab197f..40327b29a 100644 --- a/lib/pulse/lib/Pulse.Lib.OnRange.fst +++ b/lib/pulse/lib/Pulse.Lib.OnRange.fst @@ -481,6 +481,7 @@ fn rec on_range_unzip (p q:nat -> slprop) (i j:nat) rewrite (on_range (fun k -> p k ** q k) i j) as pure False; unreachable (); } else if (j = i) { + rewrite each j as i; on_range_empty_elim (fun k -> p k ** q k) i; on_range_empty p i; on_range_empty q i; diff --git a/lib/pulse/lib/Pulse.Lib.PCMReference.fst b/lib/pulse/lib/Pulse.Lib.PCMReference.fst index 95ede7f31..a8955a43c 100644 --- a/lib/pulse/lib/Pulse.Lib.PCMReference.fst +++ b/lib/pulse/lib/Pulse.Lib.PCMReference.fst @@ -70,6 +70,8 @@ fn read u#a (#a:Type u#a) (#p:pcm a) (r:pcm_ref p) (x:erased a) ensures pcm_pts_to r (f v) { let inst = pts_to_small r _; + drop_ (small_token u#a _); + fold small_token u#a inst; U.downgrade_val (C.read #(U.raise_t a) #(raise p) r (hide (U.raise_val (reveal x))) (raise_refine p x f)); } @@ -79,6 +81,8 @@ fn write u#a (#a:Type u#a) (#p:pcm a) (r:pcm_ref p) (x y:erased a) ensures pcm_pts_to r y { let inst = pts_to_small r _; + drop_ (small_token u#a _); + fold small_token u#a inst; C.write #(U.raise_t a) #(raise p) r (hide (U.raise_val (reveal x))) (hide (U.raise_val (reveal y))) (raise_upd f) } @@ -90,17 +94,12 @@ ghost fn share u#a (#a:Type u#a) (#pcm:pcm a) (r:pcm_ref pcm) ensures pcm_pts_to r v1 { let inst = pts_to_small r _; + drop_ (small_token u#a _); + fold small_token inst; fold small_token inst; C.share #(U.raise_t a) #(raise pcm) r (U.raise_val v0) (U.raise_val v1); } -[@@allow_ambiguous] -ghost fn drop_amb (p: slprop) - requires p -{ - drop_ p -} - [@@allow_ambiguous] ghost fn gather u#a (#a:Type u#a) (#pcm:pcm a) (r:pcm_ref pcm) (v0:a) (v1:a) requires pcm_pts_to r v0 @@ -110,6 +109,7 @@ ghost fn gather u#a (#a:Type u#a) (#pcm:pcm a) (r:pcm_ref pcm) (v0:a) (v1:a) { let inst = pts_to_small r v0; with inst'. assert C.pcm_pts_to #_ #(raise #a #inst' pcm) r (U.raise_val #a #inst' v1); - drop_amb (small_token u#a inst'); + rewrite each inst' as inst; + drop_ (small_token u#a inst'); C.gather #(U.raise_t #inst a) #(raise #a #inst pcm) r (U.raise_val #a #inst v0) (U.raise_val #a #inst v1); } \ No newline at end of file diff --git a/lib/pulse/lib/Pulse.Lib.PCMReference.fsti b/lib/pulse/lib/Pulse.Lib.PCMReference.fsti index 8011a365a..65e56ebe7 100644 --- a/lib/pulse/lib/Pulse.Lib.PCMReference.fsti +++ b/lib/pulse/lib/Pulse.Lib.PCMReference.fsti @@ -67,7 +67,6 @@ ghost fn share u#a (#a:Type u#a) (#pcm:pcm a) (r:pcm_ref pcm) ensures pcm_pts_to r v0 ensures pcm_pts_to r v1 -[@@allow_ambiguous] ghost fn gather u#a (#a:Type u#a) (#pcm:pcm a) (r:pcm_ref pcm) (v0:a) (v1:a) requires pcm_pts_to r v0 requires pcm_pts_to r v1 diff --git a/lib/pulse/lib/Pulse.Lib.Slice.Util.fst b/lib/pulse/lib/Pulse.Lib.Slice.Util.fst index c4d4f8444..6dcb66041 100644 --- a/lib/pulse/lib/Pulse.Lib.Slice.Util.fst +++ b/lib/pulse/lib/Pulse.Lib.Slice.Util.fst @@ -32,7 +32,7 @@ fn append_split (#t: Type) (s: S.slice t) (#p: perm) (i: SZ.t) requires pts_to s #p (v1 `Seq.append` v2) returns res: (slice t & slice t) ensures - (let (s1, s2) = res in + (let s1, s2 = fst res, snd res in pts_to s1 #p v1 ** pts_to s2 #p v2 ** S.is_split s s1 s2) diff --git a/lib/pulse/lib/Pulse.Lib.Slice.fst b/lib/pulse/lib/Pulse.Lib.Slice.fst index d39a08f72..c7a35a125 100644 --- a/lib/pulse/lib/Pulse.Lib.Slice.fst +++ b/lib/pulse/lib/Pulse.Lib.Slice.fst @@ -222,7 +222,7 @@ fn split (#t: Type) (s: slice t) (#p: perm) (i: SZ.t) requires pts_to s #p v returns res : (slice t & slice t) ensures - (let (s1, s2) = res in + (let s1, s2 = fst res, snd res in pts_to s1 #p (Seq.slice v 0 (SZ.v i)) ** pts_to s2 #p (Seq.slice v (SZ.v i) (Seq.length v)) ** is_split s s1 s2) diff --git a/lib/pulse/lib/Pulse.Lib.Slice.fsti b/lib/pulse/lib/Pulse.Lib.Slice.fsti index 8e806be2c..ddf1c09fe 100644 --- a/lib/pulse/lib/Pulse.Lib.Slice.fsti +++ b/lib/pulse/lib/Pulse.Lib.Slice.fsti @@ -157,7 +157,7 @@ fn split (#t: Type) (s: slice t) (#p: perm) (i: SZ.t) requires pts_to s #p v returns res : (slice t & slice t) ensures - (let (s1, s2) = res in + (let s1, s2 = fst res, snd res in pts_to s1 #p (Seq.slice v 0 (SZ.v i)) ** pts_to s2 #p (Seq.slice v (SZ.v i) (Seq.length v)) ** is_split s s1 s2) diff --git a/lib/pulse/lib/Pulse.Lib.Sort.Merge.Slice.fst b/lib/pulse/lib/Pulse.Lib.Sort.Merge.Slice.fst index e754230f0..ea971f847 100644 --- a/lib/pulse/lib/Pulse.Lib.Sort.Merge.Slice.fst +++ b/lib/pulse/lib/Pulse.Lib.Sort.Merge.Slice.fst @@ -255,6 +255,7 @@ requires SM.seq_list_match_append_intro_trade vmatch c accu (c1' `Seq.append` c2') (l1' `List.Tot.append` l2'); List.Tot.append_length accu (List.Tot.append l1' l2'); Trade.trans _ _ (SM.seq_list_match c1 l1_0 vmatch ** SM.seq_list_match c2 l2_0 vmatch); + with va. assert pts_to a va; rewrite each Seq.Base.append c (Seq.Base.append c1' c2') as va; !pres } diff --git a/lib/pulse/lib/Pulse.Lib.SpinLock.fst b/lib/pulse/lib/Pulse.Lib.SpinLock.fst index fe5ea2247..96578ae15 100644 --- a/lib/pulse/lib/Pulse.Lib.SpinLock.fst +++ b/lib/pulse/lib/Pulse.Lib.SpinLock.fst @@ -148,7 +148,7 @@ fn release (#v:slprop) (#p:perm) (l:lock) unfold (lock_inv l.r l.gr v); unfold (lock_inv_aux l.r l.gr v); GR.pts_to_injective_eq l.gr; - GR.gather l.gr; + GR.gather l.gr #_ #1ul; with i. assert (pts_to l.gr i); rewrite (if (i = 0ul) then v else emp) as emp; write_atomic_box l.r 0ul; @@ -202,11 +202,9 @@ fn free (#v:slprop) (l:lock) unfold (lock_inv l.r l.gr v); unfold (lock_inv_aux l.r l.gr v); B.free l.r; - GR.gather l.gr; + GR.gather l.gr #_ #1ul; + with v. assert l.gr |-> v; rewrite each v as 1ul; // awkward GR.free l.gr; - with i. assert (if (i = 0ul) then v else emp); // awkward - rewrite (if (i = 0ul) then v else emp) as emp; - () } diff --git a/lib/pulse/lib/Pulse.Lib.Task.fst b/lib/pulse/lib/Pulse.Lib.Task.fst index 215deacff..2f142ecca 100644 --- a/lib/pulse/lib/Pulse.Lib.Task.fst +++ b/lib/pulse/lib/Pulse.Lib.Task.fst @@ -346,7 +346,7 @@ fn rec extract_state_pred take_one_h11 t ts'; intro (state_pred t.pre t.post t.h @==> all_state_pred ts) - #(task_thunk_typing t' ** all_state_pred ts') fn _ + #(task_thunk_typing t ** all_state_pred ts') fn _ { add_one_state_pred t ts'; }; @@ -415,9 +415,8 @@ ghost fn shift_up (x: slprop_ref) (y: slprop) { intro (shift (up x ** later_credit 1) y) #(slprop_ref_pts_to x y) fn _ { unfold up x; - slprop_ref_gather _; + slprop_ref_gather _ #_ #y; later_elim _; - equiv_comm _ _; equiv_elim _ _; drop_ (slprop_ref_pts_to _ _); }; @@ -1273,6 +1272,7 @@ fn try_await_pool (* We have permission on the queues + all_tasks_done. Obtain pool_done temporarilly to redeem. *) + with v. assert AR.pts_to_full p.g_runnable v; AR.drop_anchor p.g_runnable; AR.share p.g_runnable; fold (pool_done p); @@ -1280,7 +1280,7 @@ fn try_await_pool (*!*) assert q; unfold (pool_done p); AR.gather p.g_runnable; - AR.lift_anchor p.g_runnable _; + AR.lift_anchor p.g_runnable v; fold (ite true q (pledge is (pool_done p) q)); fold (lock_inv p.runnable p.g_runnable); @@ -1342,13 +1342,14 @@ fn rec teardown_pool let b = check_if_all_done runnable; unfold ite; if b { + with v. assert AR.pts_to_full p.g_runnable v; AR.drop_anchor p.g_runnable; AR.share p.g_runnable; fold (pool_done p); (* Drop the other ghost half + anchor. *) drop_ (AR.pts_to p.g_runnable #0.5R runnable); - drop_ (AR.anchored p.g_runnable _); + drop_ (AR.anchored p.g_runnable v); (* TODO: actually teardown. *) drop_ (pts_to p.runnable runnable); diff --git a/lib/pulse/lib/pledge/Pulse.Lib.Shift.fst b/lib/pulse/lib/pledge/Pulse.Lib.Shift.fst index f3779dace..cd16759b4 100644 --- a/lib/pulse/lib/pledge/Pulse.Lib.Shift.fst +++ b/lib/pulse/lib/pledge/Pulse.Lib.Shift.fst @@ -204,9 +204,9 @@ ghost fn shift_dup #is p q : duplicable_f (shift #is p q) = { unfold (shift #is p q); + let d = extract_duplicator _; dup_extra_duplicable _; dup_shift_elim_exists _ _; - let d = extract_duplicator _; Pulse.Class.Duplicable.dup _ #d (); fold (shift #is p q); fold (shift #is p q) diff --git a/share/pulse/examples/AuxPredicate.fst b/share/pulse/examples/AuxPredicate.fst index df27ba0a3..c1229f231 100644 --- a/share/pulse/examples/AuxPredicate.fst +++ b/share/pulse/examples/AuxPredicate.fst @@ -98,10 +98,8 @@ fn invar_introduces_ghost_alt (r:R.ref int) r := 0; while (let vr = !r; (vr = 0)) - invariant b. - exists* v. - pts_to r v ** - pure ( (v==0 \/ v == 1) /\ b == (v = 0) ) + invariant live r + invariant pure (!r == 0 \/ !r == 1) { r := 1; } diff --git a/share/pulse/examples/CustomSyntax.fst b/share/pulse/examples/CustomSyntax.fst index 3a2908288..296a5dccc 100644 --- a/share/pulse/examples/CustomSyntax.fst +++ b/share/pulse/examples/CustomSyntax.fst @@ -305,11 +305,9 @@ fn sum (r:ref nat) (n:nat) with (zero <> n); while (let m = !i; (m <> n)) - invariant b . exists* m s. ( - pts_to i m ** - pts_to sum s ** - pure (s == sum_spec m /\ - b == (m <> n))) + invariant live i + invariant live sum + invariant pure (!sum == sum_spec (!i)) { let m = !i; let s = !sum; @@ -339,10 +337,9 @@ fn sum2 (r:ref nat) (n:nat) let mut i = zero; let mut sum = zero; while (let m = !i; (m <> n)) - invariant b . exists* m s. - pts_to i m ** - pts_to sum s ** - pure (s == sum_spec m /\ b == (m <> n)) + invariant live i + invariant live sum + invariant pure (!sum == sum_spec (!i)) { let m = !i; let s = !sum; diff --git a/share/pulse/examples/Example.RefineCase.fst b/share/pulse/examples/Example.RefineCase.fst index b48778166..312d0be46 100644 --- a/share/pulse/examples/Example.RefineCase.fst +++ b/share/pulse/examples/Example.RefineCase.fst @@ -27,7 +27,7 @@ type t_rep = | AR of int | BR of bool -let t_perm (x:t) (v:t_rep) : slprop = +let t_perm ([@@@mkey]x:t) (v:t_rep) : slprop = match x with | A r -> ( match v with diff --git a/share/pulse/examples/PulseCorePaper.S2.Lock.fst b/share/pulse/examples/PulseCorePaper.S2.Lock.fst index 16c4edf02..75627db89 100644 --- a/share/pulse/examples/PulseCorePaper.S2.Lock.fst +++ b/share/pulse/examples/PulseCorePaper.S2.Lock.fst @@ -65,7 +65,8 @@ ensures protects l p with_invariants l.i { later_elim _; - drop_ (maybe _ _); + with v. assert l.r |-> v; + drop_ (maybe (v = 0ul) _); Pulse.Lib.Primitives.write_atomic_box l.r 0ul; later_intro (lock_inv l.r p); } diff --git a/share/pulse/examples/Quicksort.Base.fst b/share/pulse/examples/Quicksort.Base.fst index 7049dd8dc..5c1634027 100644 --- a/share/pulse/examples/Quicksort.Base.fst +++ b/share/pulse/examples/Quicksort.Base.fst @@ -240,20 +240,16 @@ fn partition (a: A.array int) (lo: nat) (hi:(hi:nat{lo < hi})) let mut j = lo; let mut k = lo; while (let vk = !k; (vk < hi - 1)) - invariant b . exists* s (vi vj vk : nat). ( + invariant exists* s. ( A.pts_to_range a lo hi s ** - R.pts_to i vi ** - R.pts_to j vj ** - R.pts_to k vk ** + live i ** live j ** live k ** pure ( - b == (vk < hi - 1) /\ - //eq2_prop #bool b (vk < hi - 1) /\ - lo <= vk /\ vk <= hi - 1 /\ - lo <= vi /\ vi <= vj /\ vj <= vk /\ + lo <= !k /\ !k <= hi - 1 /\ + lo <= !i /\ !i <= !j /\ !j <= !k /\ lb <= pivot /\ pivot <= rb /\ Seq.length s = hi - lo /\ Seq.index s (hi - 1 - lo) = pivot - /\ partition_pred s lo vk (vi, vj, pivot) + /\ partition_pred s lo (!k) (!i, !j, pivot) /\ permutation s0 s /\ between_bounds s lb rb )) diff --git a/share/pulse/examples/by-example/ParallelIncrement.fst b/share/pulse/examples/by-example/ParallelIncrement.fst index 8e12bd296..950acfd2d 100644 --- a/share/pulse/examples/by-example/ParallelIncrement.fst +++ b/share/pulse/examples/by-example/ParallelIncrement.fst @@ -55,18 +55,14 @@ requires L.lock_alive l #p (exists* v. pts_to x #0.5R v ** pred v) ** R.pts_to x ensures L.lock_alive l #p (exists* v. pts_to x #0.5R v ** pred v) ** R.pts_to x #0.5R ('i + 1) ** qpred ('i + 1) { let vx = !x; - rewrite (qpred 'i) as (qpred vx); L.acquire l; + with v. assert R.pts_to x #0.5R 'i ** R.pts_to x #0.5R v ** pred v; R.gather x; - with p v. rewrite (R.pts_to x #p v) as (R.pts_to x v); - x := (vx + 1); + rewrite each v as 'i; + x := (!x + 1); R.share x; - with p _v. rewrite (R.pts_to x #p _v) as (R.pts_to x #0.5R _v); - with _v. rewrite (pred _v) as (pred vx); f vx; L.release l; - with p _v. rewrite (R.pts_to x #p _v) as (R.pts_to x #0.5R _v); - rewrite (qpred (vx + 1)) as (qpred ('i + 1)); } @@ -86,7 +82,6 @@ ensures L.lock_alive l #p (exists* v. pts_to x v ** pred v) ** qpred ('i + 1) { L.acquire l; let vx = !x; - with _v. rewrite (pred _v) as (pred vx); x := vx + 1; f vx 'i; L.release l; @@ -198,7 +193,7 @@ ensures inv l (pts_to_refine x pred) ** qpred ('i + 1) with v. _; atomic_increment x; f v 'i; - fold pts_to_refine; + fold pts_to_refine x pred; later_intro (pts_to_refine x pred); } } @@ -232,7 +227,7 @@ ensures inv l (pts_to_refine x pred) ** qpred ('i + 1) FA.elim #_ #(fun vq -> (pred v ** qpred vq ** pts_to x (v + 1)) @==> (pred (v + 1) ** qpred (vq + 1) ** pts_to x (v + 1))) 'i; I.elim _ _; - fold pts_to_refine; + fold pts_to_refine x pred; later_intro (pts_to_refine x pred); } } @@ -324,11 +319,14 @@ ensures inv l invp ** qpred ('i + 1) let mut continue = true; fold (cond true (qpred 'i) (qpred ('i + 1))); while (!continue) - invariant b. + invariant + exists* b. inv l invp ** pts_to continue b ** cond b (qpred 'i) (qpred ('i + 1)) { + rewrite each (!continue) as true; // FIXME: rewrites_to goes the wrong direction? + elim_cond_true _ _ _; let v = read (); later_credit_buy 1; let next = @@ -340,13 +338,12 @@ ensures inv l invp ** qpred ('i + 1) { later_elim _; elim_inv (); - unfold cond; + with vv. assert pure (vv == !x); let b = cas x v (v + 1); if b { unfold cond; - with vv. assert (pred vv); - f vv _; + f vv 'i; intro_inv (); fold (cond false (qpred 'i) (qpred ('i + 1))); later_intro invp; @@ -363,6 +360,7 @@ ensures inv l invp ** qpred ('i + 1) }; continue := next }; + rewrite each (!continue) as false; // FIXME: rewrites_to goes the wrong direction? unfold cond; } diff --git a/share/pulse/examples/by-example/PulseTutorial.Array.fst b/share/pulse/examples/by-example/PulseTutorial.Array.fst index 41ca7df57..e16131d40 100644 --- a/share/pulse/examples/by-example/PulseTutorial.Array.fst +++ b/share/pulse/examples/by-example/PulseTutorial.Array.fst @@ -63,7 +63,6 @@ module A = Pulse.Lib.Array module R = Pulse.Lib.Reference open FStar.SizeT -#push-options "--debug prover,pulse.checker" fn incr (y:ref int) requires R.pts_to y 'v ensures R.pts_to y ('v + 2) @@ -101,11 +100,12 @@ fn compare } ) invariant - exists* (vi:_{SZ.v vi <= SZ.v l}). ( + exists* (vi:SZ.t{SZ.v vi <= SZ.v l}). ( R.pts_to i vi ** A.pts_to a1 #p1 's1 ** A.pts_to a2 #p2 's2 ** pure ( + SZ.v vi <= SZ.v l /\ (forall (i:nat). i < SZ.v vi ==> Seq.index 's1 i == Seq.index 's2 i) ) ) diff --git a/share/pulse/examples/by-example/PulseTutorial.Conditionals.fst b/share/pulse/examples/by-example/PulseTutorial.Conditionals.fst index eb37c0635..89817d831 100644 --- a/share/pulse/examples/by-example/PulseTutorial.Conditionals.fst +++ b/share/pulse/examples/by-example/PulseTutorial.Conditionals.fst @@ -96,7 +96,7 @@ ensures pts_to x #p 'vx ** pts_to y #q 'vy let nullable_ref (a: Type0) = option (ref a) let pts_to_or_null #a - (x:nullable_ref a) + ([@@@mkey] x:nullable_ref a) (#[default_arg (`1.0R)] p:perm) //implicit argument with a default (v:option a) : slprop @@ -166,7 +166,7 @@ fn intro_pts_to_or_null_some #a #p (r:nullable_ref a) (x:ref a) requires pts_to x #p 'v ** pure (r == Some x) ensures pts_to_or_null r #p (Some 'v) { - fold (pts_to_or_null (Some x) #p (Some 'v)); + fold (pts_to_or_null (Some x) #p (Some (reveal 'v))); rewrite each (Some x) as r; } //end pts_to_or_null_helpers$ @@ -188,6 +188,7 @@ ensures pts_to_or_null r #p 'v None -> { unfold pts_to_or_null None 'v; fold pts_to_or_null None 'v; + rewrite each None #(ref a) as r; None } } diff --git a/share/pulse/examples/by-example/PulseTutorial.Ghost.fst b/share/pulse/examples/by-example/PulseTutorial.Ghost.fst index 0d29c0458..597ebaf41 100644 --- a/share/pulse/examples/by-example/PulseTutorial.Ghost.fst +++ b/share/pulse/examples/by-example/PulseTutorial.Ghost.fst @@ -170,12 +170,14 @@ decreases l match l { [] -> { unfold (all_at_most [] n); - fold (all_at_most [] n); + fold (all_at_most [] m); + rewrite all_at_most [] m as all_at_most l m; } hd :: tl -> { unfold (all_at_most (hd::tl) n); weaken_at_most tl n m; fold (all_at_most (hd::tl) m); + rewrite all_at_most (hd::tl) m as all_at_most l m; } } } diff --git a/share/pulse/examples/by-example/PulseTutorial.LinkedList.fst b/share/pulse/examples/by-example/PulseTutorial.LinkedList.fst index 406bec6e1..60697e83d 100644 --- a/share/pulse/examples/by-example/PulseTutorial.LinkedList.fst +++ b/share/pulse/examples/by-example/PulseTutorial.LinkedList.fst @@ -35,7 +35,7 @@ and llist (t:Type0) = option (node_ptr t) //end llist$ //is_list$ -let rec is_list #t (x:llist t) (l:list t) +let rec is_list #t ([@@@mkey] x:llist t) (l:list t) : Tot slprop (decreases l) = match l with | [] -> pure (x == None) @@ -149,6 +149,7 @@ ensures is_list x l ** pure (l == []) rewrite each x as (None #(ref (node t))); unfold (is_list_cases None l); fold (is_list x []); + rewrite is_list x [] as is_list x l; } //end is_list_case_none$ @@ -178,11 +179,11 @@ returns n:nat ensures is_list x 'l ** pure (n == List.Tot.length 'l) { match x { - None -> { + norewrite None -> { is_list_case_none x; 0 } - Some vl -> { + norewrite Some vl -> { is_list_case_some x vl; let node = !vl; let n = length node.tail; @@ -202,11 +203,11 @@ returns n:nat ensures is_list x 'l ** pure (n == k + List.Tot.length 'l) { match x { - None -> { + norewrite None -> { is_list_case_none x; k } - Some vl -> { + norewrite Some vl -> { is_list_case_some x vl; with _node _tl. _; let n = !vl; @@ -222,21 +223,6 @@ ensures is_list x 'l ** pure (n == k + List.Tot.length 'l) module I = Pulse.Lib.Trade.Util open I -//tail_for_cons$ -ghost -fn tail_for_cons (#t:Type) (v:node_ptr t) (#n:node t) (tl:erased (list t)) -requires - pts_to v n -ensures - (is_list n.tail tl @==> is_list (Some v) (n.head::tl)) -{ - intro (is_list n.tail tl @==> is_list (Some v) (n.head::tl)) #(pts_to v n) fn _ - { - intro_is_list_cons (Some v) v - }; -} -//end tail_for_cons$ - //tail$ fn tail (#t:Type) (x:llist t) requires is_list x 'l ** pure (Some? x) @@ -248,9 +234,11 @@ ensures exists* tl. { let np = Some?.v x; is_list_case_some x np; - with node tl. _; let nd = !np; - tail_for_cons np tl; + with tl. assert is_list nd.tail tl; + intro (is_list nd.tail tl @==> is_list x 'l) #(np |-> nd) fn _ { + intro_is_list_cons x np; + }; nd.tail } //end tail$ @@ -316,13 +304,16 @@ ensures is_list x ('l1 @ 'l2) //end append$ //tail_alt$ +let is_cons #t ([@@@mkey] x: list t) (hd: t) (tl: list t) : slprop = + pure (x == hd :: tl) + fn tail_alt (#t:Type) (x:llist t) requires is_list x 'l ** pure (Some? x) returns y:llist t ensures exists* hd tl. is_list y tl ** (forall* tl'. is_list y tl' @==> is_list x (hd::tl')) ** - pure ('l == hd::tl) + is_cons 'l hd tl { let np = Some?.v x; is_list_case_some x np; @@ -332,6 +323,7 @@ ensures exists* hd tl. #(pts_to np node) fn _ tl' { intro_is_list_cons x np; }; + fold is_cons 'l node.head _tl; node.tail } //end tail_alt$ @@ -425,7 +417,7 @@ ensures is_list x ('l1 @ 'l2) with ll pfx sfx. _; let l = !cur; let next = tail_alt l; - with hd tl. _; + with hd tl. unfold is_cons sfx hd tl; (* this is the key induction step *) FA.trans_compose (is_list next) (is_list l) (is_list x) diff --git a/share/pulse/examples/by-example/PulseTutorial.MonotonicCounterShareableFreeable.fst b/share/pulse/examples/by-example/PulseTutorial.MonotonicCounterShareableFreeable.fst index accaad696..8fb789f8c 100644 --- a/share/pulse/examples/by-example/PulseTutorial.MonotonicCounterShareableFreeable.fst +++ b/share/pulse/examples/by-example/PulseTutorial.MonotonicCounterShareableFreeable.fst @@ -130,5 +130,6 @@ ensures emp share c; par (do_something c #_ #_) (do_something c #_ #_); gather c; + rewrite each (1.0R /. 2.0R +. 1.0R /. 2.0R) as 1.0R; destroy c } \ No newline at end of file diff --git a/share/pulse/examples/by-example/PulseTutorial.MonotonicRef.fst b/share/pulse/examples/by-example/PulseTutorial.MonotonicRef.fst index 4d05c614f..4e5e591b1 100644 --- a/share/pulse/examples/by-example/PulseTutorial.MonotonicRef.fst +++ b/share/pulse/examples/by-example/PulseTutorial.MonotonicRef.fst @@ -11,7 +11,7 @@ let mref (#t:Type) (p:preorder t) = GR.gref (FP.fp_pcm p) let pts_to (#t:Type) (#p:preorder t) - (r:mref p) + ([@@@mkey] r:mref p) (#f:perm) (v:t) = exists* h. GR.pts_to r (Some f, h) ** pure (Cons? h /\ PulseCore.Preorder.curval h == v) @@ -85,5 +85,5 @@ ensures pts_to r #1.0R v unfold pts_to u#a; with f h. assert (GR.pts_to r (f, h)); GR.write r _ _ (FP.mk_frame_preserving_upd p h v); - fold pts_to u#a; + fold pts_to r v; } \ No newline at end of file diff --git a/share/pulse/examples/by-example/PulseTutorial.PCMParallelIncrement.fst b/share/pulse/examples/by-example/PulseTutorial.PCMParallelIncrement.fst index 99cfcba75..f75ebe77e 100644 --- a/share/pulse/examples/by-example/PulseTutorial.PCMParallelIncrement.fst +++ b/share/pulse/examples/by-example/PulseTutorial.PCMParallelIncrement.fst @@ -477,6 +477,7 @@ decreases remaining drop_ (can_give gs remaining); later_credit_buy 1; has_given_zero #_ #capacity r ci; + rewrite has_given gs 0 as has_given gs remaining; } else { @@ -490,6 +491,7 @@ decreases remaining drop_ (inv _ _); CI.gather ci; gather_has_given gs; + rewrite has_given gs (1 + (remaining - 1)) as has_given gs remaining; } } diff --git a/share/pulse/examples/by-example/PulseTutorial.ParallelIncrement.fst b/share/pulse/examples/by-example/PulseTutorial.ParallelIncrement.fst index 8aa09915c..0589e4020 100644 --- a/share/pulse/examples/by-example/PulseTutorial.ParallelIncrement.fst +++ b/share/pulse/examples/by-example/PulseTutorial.ParallelIncrement.fst @@ -124,7 +124,7 @@ ensures L.lock_alive lock #p (lock_inv x i left right) ** GR.pts_to left #0.5R ( GR.write left ('vl + 1); GR.share left; fold (contributions left right i (v + 1)); - fold lock_inv; + fold lock_inv x i left right; L.release lock } //end incr_left$ @@ -333,12 +333,13 @@ ensures inv (C.iname_of c) (C.cinv_vp c (exists* v. pts_to x v ** refine v)) ** let mut continue = true; fold (cond true (aspec 'i) (aspec ('i + 1))); while (!continue) - invariant b. + invariant exists* b. inv (C.iname_of c) (C.cinv_vp c (exists* v. pts_to x v ** refine v)) ** pts_to continue b ** C.active c p ** cond b (aspec 'i) (aspec ('i + 1)) { + rewrite each (!continue) as true; // FIXME: rewrites_to goes in the wrong direction later_credit_buy 1; let v = read (); later_credit_buy 1; @@ -353,12 +354,12 @@ ensures inv (C.iname_of c) (C.cinv_vp c (exists* v. pts_to x v ** refine v)) ** later_elim _; C.unpack_cinv_vp c; unfold cond; + with vv. assert x |-> vv; let b = cas x v (v + 1); if b { unfold cond; - with vv. assert (refine vv); - f vv _; + f vv 'i; C.pack_cinv_vp #(exists* v. pts_to x v ** refine v) c; fold (cond false (aspec 'i) (aspec ('i + 1))); later_intro (C.cinv_vp c (exists* v. pts_to x v ** refine v)); @@ -376,6 +377,7 @@ ensures inv (C.iname_of c) (C.cinv_vp c (exists* v. pts_to x v ** refine v)) ** continue := next }; //end incr_atomic_body_loop$ + rewrite each (!continue) as false; // FIXME: rewrites_to goes in the wrong direction unfold cond; } //end incr_atomic_body$ @@ -562,14 +564,13 @@ fn incr_pcm_t (r:ref int) (ghost_r:GPR.gref pcm) (l:L.lock) (t1:bool) (#n:int) t1_perm ghost_r (add_one n) t1 { L.acquire l; - unfold lock_inv_pcm; - unfold lock_inv_ghost; + unfold lock_inv_pcm; with n'. _; + unfold lock_inv_ghost; with n1 n2. _; let v = !r; r := v + 1; if t1 { rewrite (t1_perm ghost_r n t1) as (GPR.pts_to ghost_r (half n, None)); - with n1 n2. assert (GPR.pts_to ghost_r (half n1, half n2)); GPR.gather ghost_r (half n, None) (half n1, half n2); rewrite (GPR.pts_to ghost_r ((half n, None) `op pcm` (half n1, half n2))) as (GPR.pts_to ghost_r (full n1, half n2)); @@ -580,14 +581,13 @@ fn incr_pcm_t (r:ref int) (ghost_r:GPR.gref pcm) (l:L.lock) (t1:bool) (#n:int) rewrite (GPR.pts_to ghost_r (full (add_one n1), half n2)) as (GPR.pts_to ghost_r ((half (add_one n1), half n2) `op pcm` (half (add_one n1), None))); GPR.share ghost_r (half (add_one n1), half n2) (half (add_one n1), None); - fold lock_inv_ghost; - fold lock_inv_pcm; + fold lock_inv_ghost ghost_r (!r); + fold lock_inv_pcm r; L.release l; fold (t1_perm ghost_r (add_one n) true); } else { rewrite (t1_perm ghost_r n t1) as (GPR.pts_to ghost_r (None, half n)); - with n1 n2. assert (GPR.pts_to ghost_r (half n1, half n2)); GPR.gather ghost_r (None, half n) (half n1, half n2); rewrite (GPR.pts_to ghost_r ((None, half n2) `op pcm` (half n1, half n2))) as (GPR.pts_to ghost_r (half n1, full n2)); @@ -598,7 +598,7 @@ fn incr_pcm_t (r:ref int) (ghost_r:GPR.gref pcm) (l:L.lock) (t1:bool) (#n:int) rewrite (GPR.pts_to ghost_r (half n1, full (add_one n2))) as (GPR.pts_to ghost_r ((half n1, half (add_one n2)) `op pcm` (None, half (add_one n2)))); GPR.share ghost_r (half n1, half (add_one n2)) (None, half (add_one n2)); - fold lock_inv_ghost; + fold lock_inv_ghost ghost_r (!r); fold lock_inv_pcm; L.release l; fold (t1_perm ghost_r (add_one n) false) @@ -672,7 +672,6 @@ fn incr_pcm_t_abstract (r:ref int) (l:L.lock) L.acquire l; let v = !r; r := v + 1; - with _v. rewrite (ghost_inv _v) as (ghost_inv v); ghost_steps v; L.release #(exists* v. pts_to r v ** ghost_inv v) l } diff --git a/share/pulse/examples/by-example/PulseTutorial.SpinLock.fst b/share/pulse/examples/by-example/PulseTutorial.SpinLock.fst index fbae2c996..a71bee46f 100644 --- a/share/pulse/examples/by-example/PulseTutorial.SpinLock.fst +++ b/share/pulse/examples/by-example/PulseTutorial.SpinLock.fst @@ -88,15 +88,15 @@ ensures lock_alive l p ** p opens [l.i] { later_elim _; unfold lock_inv; + with vv. assert l.r |-> vv; let b = cas_box l.r 0ul 1ul; if b { elim_cond_true _ _ _; - with _b. rewrite (maybe _b p) as p; + rewrite each vv as 0ul; fold (maybe false p); rewrite (maybe false p) as (maybe (1ul = 0ul) p); fold (lock_inv l.r p); - fold (maybe true p); later_intro (lock_inv l.r p); true } diff --git a/share/pulse/examples/by-example/PulseTutorial.UserDefinedPredicates.fst b/share/pulse/examples/by-example/PulseTutorial.UserDefinedPredicates.fst index 50244fa3f..07299189a 100644 --- a/share/pulse/examples/by-example/PulseTutorial.UserDefinedPredicates.fst +++ b/share/pulse/examples/by-example/PulseTutorial.UserDefinedPredicates.fst @@ -20,8 +20,8 @@ open Pulse.Lib.Pervasives open FStar.Mul //pts_to_diag$ let pts_to_diag - (#a: Type0) - (r:ref (a & a)) + (#a:Type0) + ([@@@mkey] r:ref (a & a)) (v:a) : slprop = pts_to r (v, v) diff --git a/share/pulse/examples/dice/cbor/CBOR.Pulse.Extern.fsti b/share/pulse/examples/dice/cbor/CBOR.Pulse.Extern.fsti index 8d881aae1..1c2766271 100644 --- a/share/pulse/examples/dice/cbor/CBOR.Pulse.Extern.fsti +++ b/share/pulse/examples/dice/cbor/CBOR.Pulse.Extern.fsti @@ -68,20 +68,20 @@ let sndp (#a1 #a2: Type) (x: (a1 & a2)) : Tot a2 = snd x val raw_data_item_match (p: perm) - (c: cbor) + ([@@@mkey]c: cbor) (v: Cbor.raw_data_item) : Tot slprop let raw_data_item_array_match (p: perm) - (c: Seq.seq cbor) + ([@@@mkey]c: Seq.seq cbor) (v: list Cbor.raw_data_item) : Tot slprop = SM.seq_list_match c v (raw_data_item_match p) let raw_data_item_map_entry_match (p: perm) - (c1: cbor_map_entry) + ([@@@mkey] c1: cbor_map_entry) (v1: (Cbor.raw_data_item & Cbor.raw_data_item)) : Tot slprop = raw_data_item_match p (cbor_map_entry_key c1) (fstp v1) ** @@ -89,7 +89,7 @@ let raw_data_item_map_entry_match let raw_data_item_map_match (p: perm) - (c: Seq.seq cbor_map_entry) + ([@@@mkey] c: Seq.seq cbor_map_entry) (v: list (Cbor.raw_data_item & Cbor.raw_data_item)) : Tot slprop (decreases v) diff --git a/share/pulse/examples/dice/cbor/CBOR.Pulse.fst b/share/pulse/examples/dice/cbor/CBOR.Pulse.fst index 5d6446fcc..208bc8b94 100644 --- a/share/pulse/examples/dice/cbor/CBOR.Pulse.fst +++ b/share/pulse/examples/dice/cbor/CBOR.Pulse.fst @@ -242,7 +242,7 @@ ensures let res = !pres; (res = 0s && not done) ) - invariant cont . exists* i1 i2 done res l1 l2 . + invariant exists* i1 i2 done res l1 l2 . pts_to pi1 i1 ** pts_to pi2 i2 ** pts_to pdone done ** pts_to pres res ** cbor_array_iterator_match p1 i1 l1 ** cbor_array_iterator_match p2 i2 l2 ** @@ -251,9 +251,7 @@ ensures pure ( List.Tot.length l1 == List.Tot.length l2 /\ Cbor.cbor_compare v1 v2 == (if res = 0s then Cbor.cbor_compare_array l1 l2 else I16.v res) /\ - (res == 0s ==> done == Nil? l1) /\ - cont == (Cons? l1 && res = 0s) - ) + (res == 0s ==> done == Nil? l1)) { let x1 = cbor_array_iterator_next pi1; with v1' . assert (raw_data_item_match p1 x1 v1'); @@ -328,9 +326,8 @@ ensures pure ( List.Tot.length l1 == List.Tot.length l2 /\ (Cbor.cbor_compare v1 v2 == (if res = 0s then Cbor.cbor_compare_map l1 l2 else I16.v res)) /\ - (res == 0s ==> done == Nil? l1) /\ - cont == (Cons? l1 && res = 0s) - ) + (res == 0s ==> done == Nil? l1)) ** + pure (cont == (Cons? l1 && res = 0s)) { let x1 = cbor_map_iterator_next pi1; with v1' . assert (raw_data_item_map_entry_match p1 x1 v1'); @@ -575,16 +572,13 @@ ensures assert (pts_to pres gres ** cbor_map_get_invariant pmap vkey vmap map gres i l); // FIXME: WHY WHY WHY? not (done || Found? res) ) - invariant cont . exists* (done: bool) (res: cbor_map_get_t) (i: cbor_map_iterator_t) (l: list (Cbor.raw_data_item & Cbor.raw_data_item)) . + invariant exists* (done: bool) (res: cbor_map_get_t) (i: cbor_map_iterator_t) (l: list (Cbor.raw_data_item & Cbor.raw_data_item)) . raw_data_item_match pkey key vkey ** pts_to pdone done ** pts_to pi i ** pts_to pres res ** cbor_map_get_invariant pmap vkey vmap map res i l ** - pure ( - done == Nil? l /\ - cont == not (done || Found? res) - ) + pure (done == Nil? l) { with gres gi l . assert (pts_to pres gres ** cbor_map_get_invariant pmap vkey vmap map gres gi l); rewrite each gres as NotFound; diff --git a/share/pulse/examples/dice/cbor/CDDL.Pulse.fst b/share/pulse/examples/dice/cbor/CDDL.Pulse.fst index 179b60c74..b761ca1b8 100644 --- a/share/pulse/examples/dice/cbor/CDDL.Pulse.fst +++ b/share/pulse/examples/dice/cbor/CDDL.Pulse.fst @@ -305,6 +305,7 @@ fn impl_array_group3_item } else { let c = cbor_array_iterator_next pi #p #l #gi; // FIXME: WHY WHY WHY those explicit arguments? with gc i' l' . assert ( + pi |-> i' ** raw_data_item_match p c gc ** cbor_array_iterator_match p i' l' ** ((raw_data_item_match p c gc ** cbor_array_iterator_match p i' l') @==> cbor_array_iterator_match p gi l) @@ -734,6 +735,7 @@ ensures let res = cbor_map_get key map; if (Found? res) { let fres = Found?._0 res; + assert rewrites_to res (Found fres); manurewrite (cbor_map_get_post pmap vkey vmap map res) (cbor_map_get_post_found pmap vkey vmap map fres); unfold (cbor_map_get_post_found pmap vkey vmap map fres); let test = ft fres; @@ -746,6 +748,7 @@ ensures NotFound } } else { + assert rewrites_to res NotFound; rewrite (cbor_map_get_post pmap vkey vmap map res) as (cbor_map_get_post_not_found pmap vkey vmap map); unfold (cbor_map_get_post_not_found pmap vkey vmap map); fold (cbor_map_get_with_typ_post t pmap vkey vmap map NotFound); diff --git a/share/pulse/examples/dice/dpe/DPE.fst b/share/pulse/examples/dice/dpe/DPE.fst index 208140952..ea32a3acc 100644 --- a/share/pulse/examples/dice/dpe/DPE.fst +++ b/share/pulse/examples/dice/dpe/DPE.fst @@ -71,7 +71,7 @@ fn drop_mutex_live (#a:Type0) (m:mutex a) (#p:perm) (v:a -> slprop) requires mutex_live m #p v ensures emp { - drop_ _; + drop_ (mutex_live m #p v); } [@@ Rust_const_fn] @@ -161,6 +161,35 @@ fn share_ (r:gref) GR.share r v0 v1; } +// +// A frame preserving update in the trace PCM, +// given a valid transition +// +noextract +let mk_frame_preserving_upd + (t:hist trace_preorder) + (s:g_session_state { valid_transition t s }) + : FStar.PCM.frame_preserving_upd trace_pcm (Some 1.0R, t) (Some 1.0R, next_trace t s) = + fun v -> + assert (trace_pcm.refine v); + assert (FStar.PCM.compatible trace_pcm (Some 1.0R, t) v); + let v_new = Some 1.0R, next_trace t s in + assert (trace_pcm.refine v_new); + FStar.PCM.compatible_refl trace_pcm v_new; + assert (FStar.PCM.compatible trace_pcm (Some 1.0R, next_trace t s) v_new); + let x = Some 1.0R, t in + let y = Some 1.0R, next_trace t s in + let p = trace_pcm in + let open FStar.PCM in + introduce + forall (frame:_{composable p x frame}). + composable p y frame /\ + (op p x frame == v ==> op p y frame == v_new) + with ( + assert (composable p x frame); + assert (fst frame == None) + ); + v_new noextract let full (t0:trace) = Some #perm 1.0R, t0 @@ -248,6 +277,8 @@ fn frame_session_perm_at_sid rewrite (session_perm r pht0 sid) as (session_state_perm r pht0 sid16); unfold session_state_perm; + with s. unfold pht_contains pht0 sid16 s; + fold pht_contains pht1 sid16 s; fold (session_state_perm r pht1 sid16); rewrite (session_state_perm r pht1 sid16) as (session_perm r pht1 sid) @@ -331,6 +362,7 @@ fn __open_session (s:st) emp_to_start_valid (); upd_sid_pts_to trace_ref ctr #emp_trace #emp_trace G_SessionStart; rewrite emp as (session_state_related SessionStart G_SessionStart); + fold pht_contains pht1 ctr SessionStart; fold (session_state_perm trace_ref pht1 ctr); rewrite (session_state_perm trace_ref pht1 ctr) as (session_perm trace_ref pht1 (U16.v ctr)); @@ -343,6 +375,7 @@ fn __open_session (s:st) tbl1 as (fst ret).st_tbl; fold (dpe_inv trace_ref (Some (fst ret))); fold (open_session_client_perm (Some ctr)); + rewrite each Some ctr as snd ret; ret } else { let s = { st_ctr = ctr; st_tbl = tbl1 }; @@ -416,7 +449,7 @@ fn open_session () mg := Some s; M.unlock (snd r) mg; - drop_mutex_live _ (dpe_inv (fst r)); + drop_mutex_live (snd r) _; sid_opt } @@ -505,13 +538,14 @@ fn replace_session { let r = Global.read_gvar gst; unfold (gvar_p r); + rewrite each fst r as trace_ref; let mg = M.lock (snd r); let sopt = M.replace mg None; match sopt { None -> { unfold (dpe_inv trace_ref None); - unfold sid_pts_to; + unfold sid_pts_to trace_ref; gather_ trace_ref all_sids_unused (singleton sid 0.5R t); unreachable () } @@ -528,37 +562,23 @@ fn replace_session on_range_get (U16.v sid) #(session_perm trace_ref pht0) #0 #(U16.v ctr); rewrite (session_perm trace_ref pht0 (U16.v sid)) as (session_state_perm trace_ref pht0 sid); - unfold session_state_perm; - gather_sid_pts_to sid; - with t1. assert (GR.pts_to trace_ref (singleton sid 1.0R t1)); - assert (pure (t1 == t)); - let ret = HT.lookup tbl sid; - let tbl = fst ret; - let idx = snd ret; - rewrite each - fst ret as tbl, - snd ret as idx; + unfold session_state_perm trace_ref pht0 sid; + with st' t1. assert pht_contains pht0 sid st' ** session_state_related st' (current_state t1); + unfold pht_contains pht0 sid st'; + gather_sid_pts_to sid #t #t1; + rewrite each t1 as t; + let tbl, idx = HT.lookup tbl sid; with pht. assert (models tbl pht); match idx { Some idx -> { - let ret = HT.replace #_ #_ #pht tbl idx sid sst (); - let tbl = fst ret; - let st = snd ret; - rewrite each - fst ret as tbl, - snd ret as st; - assert (session_state_related sst gsst); - with sst' gsst'. assert ( - session_state_related sst' gsst' ** - session_state_related sst gsst - ); - rewrite (session_state_related sst' gsst') - as (session_state_related st (current_state t1)); + let tbl, st = HT.replace #_ #_ #pht0 tbl idx sid sst (); + rewrite each st' as st; with pht. assert (models tbl pht); - upd_singleton sid #t1 gsst; - share_sid_pts_to sid; + upd_singleton sid #t gsst; + share_sid_pts_to sid #(next_trace t gsst); rewrite (session_state_related sst gsst) as - (session_state_related sst (current_state (next_trace t1 gsst))); + (session_state_related sst (current_state (next_trace t gsst))); + fold pht_contains pht sid sst; fold (session_state_perm trace_ref pht sid); rewrite (session_state_perm trace_ref pht sid) as (session_perm trace_ref pht (U16.v sid)); @@ -649,6 +669,19 @@ fn intro_session_state_tag_related (s:session_state) (gs:g_session_state) } } +ghost fn unrelated_session_state #x #y () + requires session_state_related x y + requires pure (match x, y with + | SessionStart, G_SessionStart + | InUse, G_InUse _ + | SessionClosed, G_SessionClosed _ + | SessionError, G_SessionError _ + | Available .., G_Available .. -> False + | _ -> True) + ensures pure False +{ + rewrite session_state_related x y as pure False; +} #push-options "--fuel 2 --ifuel 2 --split_queries no" @@ -669,7 +702,7 @@ fn initialize_context (sid:sid_t) match s { SessionStart -> { - rewrite (session_state_related s (current_state t)) as emp; + rewrite (session_state_related SessionStart (current_state t)) as emp; let context = init_engine_ctxt uds; let s = Available { context }; rewrite (context_perm context (Engine_context_repr uds_bytes)) as @@ -680,23 +713,19 @@ fn initialize_context (sid:sid_t) fold (initialize_context_client_perm sid uds_bytes) } InUse -> { - rewrite (session_state_related s (current_state t)) as - (pure False); + unrelated_session_state (); unreachable () } SessionClosed -> { - rewrite (session_state_related s (current_state t)) as - (pure False); + unrelated_session_state (); unreachable () } SessionError -> { - rewrite (session_state_related s (current_state t)) as - (pure False); + unrelated_session_state (); unreachable () } Available _ -> { - rewrite (session_state_related s (current_state t)) as - (pure False); + unrelated_session_state (); unreachable () } } @@ -1121,12 +1150,12 @@ fn derive_child (sid:sid_t) Available hc -> { match hc.context { L1_context _ -> { - rewrite (session_state_related s (current_state t)) as + rewrite (session_state_related (Available hc) (current_state t)) as (pure False); unreachable () } Engine_context _ -> { - let repr = rewrite_session_state_related_available hc.context s t; + let repr = rewrite_session_state_related_available hc.context (Available hc) t; intro_context_and_repr_tag_related hc.context repr; let ret = derive_child_from_context hc.context record #rrepr #repr (); @@ -1146,7 +1175,7 @@ fn derive_child (sid:sid_t) } None -> { let s = SessionError; - rewrite (maybe_context_perm repr rrepr ret) as emp; + rewrite (maybe_context_perm repr rrepr None) as emp; rewrite emp as (session_state_related s (G_SessionError (current_state t1))); let s = replace_session sid t1 s (G_SessionError (current_state t1)); intro_session_state_tag_related s (current_state t1); @@ -1157,7 +1186,7 @@ fn derive_child (sid:sid_t) } } L0_context _ -> { - let repr = rewrite_session_state_related_available hc.context s t; + let repr = rewrite_session_state_related_available hc.context (Available hc) t; intro_context_and_repr_tag_related hc.context repr; let ret = derive_child_from_context hc.context record #rrepr #repr (); @@ -1177,7 +1206,7 @@ fn derive_child (sid:sid_t) } None -> { let s = SessionError; - rewrite (maybe_context_perm repr rrepr ret) as emp; + rewrite (maybe_context_perm repr rrepr None) as emp; rewrite emp as (session_state_related s (G_SessionError (current_state t1))); let s = replace_session sid t1 s (G_SessionError (current_state t1)); intro_session_state_tag_related s (current_state t1); @@ -1191,23 +1220,19 @@ fn derive_child (sid:sid_t) } } SessionStart -> { - rewrite (session_state_related s (current_state t)) as - (pure False); + unrelated_session_state (); unreachable () } InUse -> { - rewrite (session_state_related s (current_state t)) as - (pure False); + unrelated_session_state (); unreachable () } SessionClosed -> { - rewrite (session_state_related s (current_state t)) as - (pure False); + unrelated_session_state (); unreachable () } SessionError -> { - rewrite (session_state_related s (current_state t)) as - (pure False); + unrelated_session_state (); unreachable () } } @@ -1222,7 +1247,7 @@ fn destroy_session_state (s:session_state) (t:G.erased trace) intro_session_state_tag_related s (current_state t); match s { Available hc -> { - rewrite_session_state_related_available hc.context s t; + rewrite_session_state_related_available hc.context (Available hc) t; destroy_ctxt hc.context } SessionStart -> { @@ -1288,7 +1313,7 @@ ensures with t1. assert (sid_pts_to trace_ref sid t1); match s { - Available hc -> { + norewrite Available hc -> { match hc.context { L1_context c -> { let c_crt_len = c.aliasKeyCRT_len; @@ -1344,23 +1369,19 @@ ensures } } SessionStart -> { - rewrite (session_state_related s (current_state t)) as - (pure False); + unrelated_session_state (); unreachable () } InUse -> { - rewrite (session_state_related s (current_state t)) as - (pure False); + unrelated_session_state (); unreachable () } SessionClosed -> { - rewrite (session_state_related s (current_state t)) as - (pure False); + unrelated_session_state (); unreachable () } SessionError -> { - rewrite (session_state_related s (current_state t)) as - (pure False); + unrelated_session_state (); unreachable () } } @@ -1391,7 +1412,7 @@ ensures with t1. assert (sid_pts_to trace_ref sid t1); match s { - Available hc -> { + norewrite Available hc -> { match hc.context { L1_context c -> { let r = rewrite_session_state_related_available (L1_context c) s t; @@ -1428,23 +1449,19 @@ ensures } } SessionStart -> { - rewrite (session_state_related s (current_state t)) as - (pure False); + unrelated_session_state (); unreachable () } InUse -> { - rewrite (session_state_related s (current_state t)) as - (pure False); + unrelated_session_state (); unreachable () } SessionClosed -> { - rewrite (session_state_related s (current_state t)) as - (pure False); + unrelated_session_state (); unreachable () } SessionError -> { - rewrite (session_state_related s (current_state t)) as - (pure False); + unrelated_session_state (); unreachable () } diff --git a/share/pulse/examples/dice/dpe/DPE.fsti b/share/pulse/examples/dice/dpe/DPE.fsti index 71c247d95..59b599fb3 100644 --- a/share/pulse/examples/dice/dpe/DPE.fsti +++ b/share/pulse/examples/dice/dpe/DPE.fsti @@ -221,36 +221,6 @@ let next_trace (t:trace) (s:g_session_state { valid_transition t s }) : trace = | [] -> [s]::t | l -> (s::l)::t -// -// A frame preserving update in the trace PCM, -// given a valid transition -// -noextract -let mk_frame_preserving_upd - (t:hist trace_preorder) - (s:g_session_state { valid_transition t s }) - : FStar.PCM.frame_preserving_upd trace_pcm (Some 1.0R, t) (Some 1.0R, next_trace t s) = - fun v -> - assert (trace_pcm.refine v); - assert (FStar.PCM.compatible trace_pcm (Some 1.0R, t) v); - let v_new = Some 1.0R, next_trace t s in - assert (trace_pcm.refine v_new); - FStar.PCM.compatible_refl trace_pcm v_new; - assert (FStar.PCM.compatible trace_pcm (Some 1.0R, next_trace t s) v_new); - let x = Some 1.0R, t in - let y = Some 1.0R, next_trace t s in - let p = trace_pcm in - let open FStar.PCM in - introduce - forall (frame:_{composable p x frame}). - composable p y frame /\ - (op p x frame == v ==> op p y frame == v_new) - with ( - assert (composable p x frame); - assert (fst frame == None) - ); - v_new - // // A snapshot of the trace PCM is the trace with no permission // @@ -293,7 +263,7 @@ let singleton (sid:sid_t) (p:perm) (t:trace) : GTot pcm_t = // to capture the session history for a sid // noextract -let sid_pts_to (r:gref) (sid:sid_t) (t:trace) : slprop = +let sid_pts_to ([@@@mkey] r:gref) (sid:sid_t) (t:trace) : slprop = GR.pts_to r (singleton sid 0.5R t) noextract @@ -303,7 +273,7 @@ type pht_t = PHT.pht_t sid_t session_state // Towards the global state invariant // -let session_state_related (s:session_state) (gs:g_session_state) : slprop = +let session_state_related ([@@@mkey] s:session_state) (gs:g_session_state) : slprop = match s, gs with | SessionStart, G_SessionStart | InUse, G_InUse _ @@ -317,9 +287,11 @@ let session_state_related (s:session_state) (gs:g_session_state) : slprop = // // Invariant for sessions that have been started // -let session_state_perm (r:gref) (pht:pht_t) (sid:sid_t) : slprop = +let pht_contains ([@@@mkey] pht:pht_t) ([@@@mkey] sid:sid_t) (s: session_state) = + pure (PHT.lookup pht sid == Some s) +let session_state_perm ([@@@mkey] r:gref) (pht:pht_t) (sid:sid_t) : slprop = exists* (s:session_state) (t:trace). - pure (PHT.lookup pht sid == Some s) ** + pht_contains pht sid s ** sid_pts_to r sid t ** session_state_related s (current_state t) @@ -328,7 +300,7 @@ let session_state_perm (r:gref) (pht:pht_t) (sid:sid_t) : slprop = // if we parameterized it over a typeclass, we can directly use u32 keys // and this should go away // -let session_perm (r:gref) (pht:pht_t) (sid:nat) = +let session_perm ([@@@mkey] r:gref) (pht:pht_t) (sid:nat) = if UInt.fits sid 16 then session_state_perm r pht (U16.uint_to_t sid) else emp @@ -346,7 +318,7 @@ let sids_above_unused (s:sid_t) : GTot pcm_t = map_literal (fun sid -> // // Main invariant // -let dpe_inv (r:gref) (s:option st) : slprop = +let dpe_inv ([@@@mkey] r:gref) (s:option st) : slprop = match s with // // Global state is not initialized, @@ -375,6 +347,7 @@ val trace_ref : gref // The DPE API // +[@@no_mkeys] let open_session_client_perm (s:option sid_t) : slprop = match s with | None -> emp @@ -390,6 +363,7 @@ noextract let trace_valid_for_initialize_context (t:trace) : prop = current_state t == G_SessionStart +[@@no_mkeys] let initialize_context_client_perm (sid:sid_t) (uds:Seq.seq U8.t) = exists* t. sid_pts_to trace_ref sid t ** pure (current_state t == G_Available (Engine_context_repr uds)) @@ -419,6 +393,7 @@ let derive_child_post_trace (r:repr_t) (t:trace) = | L0_repr r, G_Available (L1_context_repr lrepr) -> lrepr.repr == r | _ -> False +[@@no_mkeys] let derive_child_client_perm (sid:sid_t) (t0:trace) (repr:repr_t) (res:bool) : slprop = match res with @@ -448,6 +423,7 @@ let trace_valid_for_close (t:trace) : prop = | G_InUse _ -> False | _ -> True +[@@no_mkeys] let session_closed_client_perm (sid:sid_t) (t0:trace) = exists* t1. sid_pts_to trace_ref sid t1 ** pure (current_state t1 == G_SessionClosed (G_InUse (current_state t0))) @@ -466,6 +442,7 @@ let trace_valid_for_certify_key (t:trace) : prop = | G_Available (L1_context_repr _) -> True | _ -> False +[@@no_mkeys] let certify_key_client_perm (sid:sid_t) (t0:trace) : slprop = exists* t1. sid_pts_to trace_ref sid t1 ** pure (current_state t1 == current_state t0) @@ -493,6 +470,7 @@ let trace_valid_for_sign (t:trace) : prop = | G_Available (L1_context_repr _) -> True | _ -> False +[@@no_mkeys] let sign_client_perm (sid:sid_t) (t0:trace) : slprop = exists* t1. sid_pts_to trace_ref sid t1 ** pure (current_state t1 == current_state t0) diff --git a/share/pulse/examples/dice/dpe/DPETypes.fsti b/share/pulse/examples/dice/dpe/DPETypes.fsti index 5b645721b..fa4288d17 100644 --- a/share/pulse/examples/dice/dpe/DPETypes.fsti +++ b/share/pulse/examples/dice/dpe/DPETypes.fsti @@ -163,7 +163,7 @@ let mk_l0_context_repr_t : GTot l0_context_repr_t = {uds; cdi; repr} -let l0_context_perm (c:l0_context_t) (r:l0_context_repr_t): slprop +let l0_context_perm ([@@@mkey] c:l0_context_t) (r:l0_context_repr_t): slprop = pts_to c.cdi r.cdi ** pure (V.is_full_vec c.cdi) @@ -252,7 +252,7 @@ let mk_l1_context_repr_t { cdi; deviceID_label_len; aliasKey_label_len; deviceIDCSR_ingredients; aliasKeyCRT_ingredients; deviceID_pub; aliasKey_pub; aliasKey_priv; deviceIDCSR_len; deviceIDCSR; aliasKeyCRT_len; aliasKeyCRT; repr } -let l1_context_perm (c:l1_context_t) (r:l1_context_repr_t) +let l1_context_perm ([@@@mkey] c:l1_context_t) (r:l1_context_repr_t) : slprop = pts_to c.deviceID_pub r.deviceID_pub ** pts_to c.aliasKey_pub r.aliasKey_pub ** @@ -287,7 +287,7 @@ let mk_context_repr_t_l0 (r:erased l0_context_repr_t) let mk_context_repr_t_l1 (r:erased l1_context_repr_t) : erased context_repr_t = L1_context_repr r -let context_perm (context:context_t) (repr:context_repr_t): slprop = +let context_perm ([@@@mkey] context:context_t) (repr:context_repr_t): slprop = match context, repr with | Engine_context c, Engine_context_repr uds -> engine_context_perm c uds | L0_context c, L0_context_repr r -> l0_context_perm c r @@ -408,7 +408,7 @@ type repr_t = | Engine_repr : r:engine_record_repr -> repr_t | L0_repr : r:l0_record_repr_t -> repr_t -let record_perm (record:record_t) (p:perm) (repr:repr_t) : slprop = +let record_perm ([@@@mkey] record:record_t) (p:perm) (repr:repr_t) : slprop = match record, repr with | Engine_record r, Engine_repr r0 -> engine_record_perm r p r0 | L0_record r, L0_repr r0 -> l0_record_perm r p r0 diff --git a/share/pulse/examples/parallel/Async.fst b/share/pulse/examples/parallel/Async.fst index 344881e9c..9b7059c0d 100644 --- a/share/pulse/examples/parallel/Async.fst +++ b/share/pulse/examples/parallel/Async.fst @@ -100,11 +100,12 @@ fn __await let r = fst h; let th = snd h; unfold async_joinable; + rewrite each fst h as r; + rewrite each snd h as th; assert (joinable th); join th; (* join the thread *) assert (done th); - rewrite (done th) as (done (snd h)); - redeem_pledge emp_inames (done (snd h)) (box_solves_post r post); + redeem_pledge emp_inames (done th) (box_solves_post r post); assert (box_solves_post r post); unfold box_solves_post; with vv. assert (Box.pts_to r (Some vv)); diff --git a/share/pulse/examples/parallel/ParallelFor.fst b/share/pulse/examples/parallel/ParallelFor.fst index 348e5dab9..6ada98edb 100644 --- a/share/pulse/examples/parallel/ParallelFor.fst +++ b/share/pulse/examples/parallel/ParallelFor.fst @@ -207,6 +207,7 @@ fn rec redeem_range ensures f ** on_range p 0 n { if (n = 0) { + rewrite each n as 0; on_range_empty_elim (fun i -> pledge emp_inames f (p i)) 0; on_range_empty p 0; } else { diff --git a/share/pulse/examples/parallel/Promises.Examples3.fst b/share/pulse/examples/parallel/Promises.Examples3.fst index 73f4422ab..83d0b02cc 100644 --- a/share/pulse/examples/parallel/Promises.Examples3.fst +++ b/share/pulse/examples/parallel/Promises.Examples3.fst @@ -75,6 +75,7 @@ fn proof assert (pts_to done #0.5R v_done ** pts_to done #0.5R true ** pts_to res #0.5R v_res + ** pts_to claimed #0.5R false ** pts_to claimed #0.5R v_claimed ** (if not v_claimed then R.pts_to res #0.5R v_res else emp) ** pure (v_claimed ==> v_done) diff --git a/src/checker/Pulse.Checker.Admit.fst b/src/checker/Pulse.Checker.Admit.fst index 1e624e43d..ff332d273 100644 --- a/src/checker/Pulse.Checker.Admit.fst +++ b/src/checker/Pulse.Checker.Admit.fst @@ -22,7 +22,6 @@ open Pulse.Syntax open Pulse.Typing open Pulse.Checker.Pure open Pulse.Checker.Base -open Pulse.Checker.Prover module RU = Pulse.RuntimeUtils module P = Pulse.Syntax.Printer diff --git a/src/checker/Pulse.Checker.AssertWithBinders.fst b/src/checker/Pulse.Checker.AssertWithBinders.fst index afa84306b..2543ffbe7 100644 --- a/src/checker/Pulse.Checker.AssertWithBinders.fst +++ b/src/checker/Pulse.Checker.AssertWithBinders.fst @@ -31,6 +31,7 @@ module PC = Pulse.Checker.Pure module P = Pulse.Syntax.Printer module PS = Pulse.Checker.Prover.Substs module Prover = Pulse.Checker.Prover +open Pulse.Checker.Prover.Normalize module Env = Pulse.Typing.Env open Pulse.Show module RU = Pulse.RuntimeUtils @@ -39,81 +40,21 @@ let is_host_term (t:R.term) = not (R.Tv_Unknown? (R.inspect_ln t)) let debug_log = Pulse.Typing.debug_log "with_binders" -let option_must (f:option 'a) (msg:string) : T.Tac 'a = - match f with - | Some x -> x - | None -> T.fail msg - -let rec refl_abs_binders (t:R.term) (acc:list binder) : T.Tac (list binder) = - let open R in - match inspect_ln t with - | Tv_Abs b body -> - let {sort; ppname} = R.inspect_binder b in - refl_abs_binders body - ((mk_binder_ppname sort (mk_ppname ppname (RU.range_of_term t)))::acc) - | _ -> L.rev acc - -let infer_binder_types (g:env) (bs:list binder) (v:slprop) - : T.Tac (list binder) = - match bs with - | [] -> [] - | _ -> - let v_rng = Pulse.RuntimeUtils.range_of_term v in - let g = push_context g "infer_binder_types" v_rng in - let as_binder (b:binder) : R.binder = - let open R in - let bv : binder_view = - { sort = b.binder_ty; - ppname = b.binder_ppname.name; - qual = Q_Explicit; - attrs = [] } in - pack_binder bv - in - let abstraction = - L.fold_right - (fun b (tv:term) -> - let b = as_binder b in - R.pack_ln (R.Tv_Abs b tv)) - bs - v - in - let inst_abstraction, _ = PC.instantiate_term_implicits g (wr abstraction v_rng) None true in - refl_abs_binders inst_abstraction [] - -let rec open_binders (g:env) (bs:list binder) (uvs:env { disjoint uvs g }) (v:term) (body:st_term) - : T.Tac (uvs:env { disjoint uvs g } & term & st_term) = +let rec open_binders_with_uvars (g:env) (bs:list binder) (v:term) (body:st_term) (us: list term) + : T.Tac (uvs: list term & term & st_term) = match bs with - | [] -> (| uvs, v, body |) - | b::bs -> - // these binders are only lax checked so far - let _ = PC.check_universe (push_env g uvs) b.binder_ty in - let x = fresh (push_env g uvs) in - let ss = [ RT.DT 0 (tm_var {nm_index=x;nm_ppname=b.binder_ppname}) ] in + | [] -> (| List.rev us, v, body |) + | {binder_ty}::bs -> + let binder_ty, _ = PC.tc_type_phase1 g binder_ty in + let u, _ = PC.tc_term_phase1_with_type g tm_unknown binder_ty in + let ss = [ RT.DT 0 u ] in let bs = L.mapi (fun i b -> assume (i >= 0); subst_binder b (shift_subst_n i ss)) bs in let v = subst_term v (shift_subst_n (L.length bs) ss) in let body = subst_st_term body (shift_subst_n (L.length bs) ss) in - open_binders g bs (push_binding uvs x b.binder_ppname b.binder_ty) v body - -let closing (bs:list (ppname & var & typ)) : subst = - L.fold_right (fun (_, x, _) (n, ss) -> - n+1, - (RT.ND x n)::ss - ) bs (0, []) |> snd - -let rec close_binders (bs:list (ppname & var & typ)) - : Tot (list binder) (decreases L.length bs) = - match bs with - | [] -> [] - | (name, x, t)::bs -> - let bss = L.mapi (fun n (n1, x1, t1) -> - assume (n >= 0); - n1, x1, subst_term t1 [RT.ND x n]) bs in - let b = mk_binder_ppname t name in - assume (L.length bss == L.length bs); - b::(close_binders bss) + open_binders_with_uvars g bs v body (u::us) let unfold_all (g : env) (names : list string) (t:term) : T.Tac term @@ -220,22 +161,10 @@ let visit_and_rewrite (p: (R.term & R.term)) (t:term) : T.Tac term = let visit_and_rewrite_conjuncts (p: (R.term & R.term)) (tms:list term) : T.Tac (list term) = T.map (visit_and_rewrite p) tms -(* is_source: was this rewrite written in the source by the user? *) -let visit_and_rewrite_conjuncts_all (is_source:bool) (g:env) (p: list (R.term & R.term)) (goal:term) : T.Tac (term & term) = +let visit_and_rewrite_conjuncts_all (g:env) (p: list (R.term & R.term)) (goal:term) : T.Tac term = let tms = slprop_as_list goal in let tms' = T.fold_left (fun tms p -> visit_and_rewrite_conjuncts p tms) tms p in - assume (L.length tms' == L.length tms); - let lhs, rhs = - T.fold_left2 - (fun (lhs, rhs) t t' -> - if eq_tm t t' then lhs, rhs - else (t::lhs, t'::rhs)) - ([], []) - tms tms' - in - if is_source && Nil? lhs then - warn_nop g goal; - list_as_slprop lhs, list_as_slprop rhs + list_as_slprop tms' let disjoint (dom:list var) (cod:Set.set var) = L.for_all (fun d -> not (Set.mem d cod)) dom @@ -264,15 +193,15 @@ let rec as_subst (p : list (term & term)) -let rewrite_all (is_source:bool) (g:env) (p: list (term & term)) (t:term) pre elaborated tac_opt : T.Tac (term & term) = +let rewrite_all (is_source:bool) (g:env) (p: list (term & term)) (t:term) pre elaborated tac_opt : T.Tac (term & list (term & term)) = (* We only use the rewrites_to substitution if there is no tactic attached to the rewrite. Otherwise, tactics may become brittle as the goal is changed unexpectedly by other things in the context. See tests/Match.fst. *) let use_rwr = None? tac_opt in - let norm (t:term) : T.Tac term = dfst <| Pulse.Checker.Prover.normalize_slprop g t use_rwr in + let norm (t:term) : T.Tac term = dfst <| normalize_slprop g t use_rwr in let t = let t, _ = Pulse.Checker.Pure.instantiate_term_implicits g t None true in - let t = dfst <| Pulse.Checker.Prover.normalize_slprop g t use_rwr in + let t = dfst <| normalize_slprop g t use_rwr in t in let maybe_purify t = if elaborated then t else purify_term g {ctxt_now=pre;ctxt_old=None} t in @@ -291,21 +220,78 @@ let rewrite_all (is_source:bool) (g:env) (p: list (term & term)) (t:term) pre el let t' = subst_term t s in if is_source && eq_tm t t' then warn_nop g t; - t, t' + t', p | _ -> - let lhs, rhs = visit_and_rewrite_conjuncts_all is_source g p t in - debug_log g (fun _ -> Printf.sprintf "Rewrote %s to %s" (P.term_to_string lhs) (P.term_to_string rhs)); - lhs, rhs + let rhs = visit_and_rewrite_conjuncts_all g p t in + if is_source && eq_tm t rhs then + warn_nop g t; + debug_log g (fun _ -> Printf.sprintf "Rewrote %s to %s" (P.term_to_string t) (P.term_to_string rhs)); + rhs, p + +open Pulse.PP +let check_equiv_with_tac (g:env) (rng:Range.range) (lhs rhs ty:term) (tac_tm:term) +: T.Tac (option T.issues) += let g_env = elab_env g in + match Pulse.Typing.Util.universe_of_now g_env ty with + | None, issues -> + fail_doc_with_subissues g (Some rng) issues [ + text "rewrite: could not determine the universe of"; + pp ty; + ] + | Some u, _ -> + let goal = mk_squash u0 (RT.eq2 u ty lhs rhs) in + let goal_typing + : my_erased (T.typing_token (elab_env g) goal (RT.E_Total, R.pack_ln (R.Tv_Type u0))) + = magic() + in + let goal_typing_tok : squash (T.typing_token (elab_env g) goal (RT.E_Total, R.pack_ln (R.Tv_Type u0))) = + match goal_typing with | E x -> () + in + let res, issues = T.call_subtac_tm g_env tac_tm u0 goal in + if None? res then Some issues else None + +let check_equiv_maybe_tac (g:env) (rng:Range.range) (lhs rhs ty:term) (tac_opt:option term) +: T.Tac (option T.issues) += match tac_opt with + | None -> + let res, issues = + Pulse.Typing.Util.check_equiv_now (elab_env g) lhs rhs in + if None? res then Some issues else None + | Some tac_tm -> + check_equiv_with_tac g rng lhs rhs ty tac_tm + +let check_pair (g:env) rng (lhs rhs:term) (tac_opt:option term) : T.Tac unit = + let (| _, ty, _ |) = PC.core_compute_term_type g lhs in + let (| _, _ |) = PC.core_check_term_at_type g rhs ty in + let issues = check_equiv_maybe_tac g rng lhs rhs ty tac_opt in + match issues with + | Some issues -> + fail_doc_with_subissues g (Some rng) issues [ + text "rename: could not prove equality of"; + pp lhs; + pp rhs; + ] + | _ -> + () + +let rec check_pairs (g:env) rng (ps: list (term & term)) (tac_opt:option term) : T.Tac unit = + match ps with + | [] -> () + | (lhs,rhs)::ps -> check_pair g rng lhs rhs tac_opt; check_pairs g rng ps tac_opt let check_renaming (g:env) (pre:term) + (pre_typing:tot_typing g pre tm_slprop) + (post_hint:post_hint_opt g) + (res_ppname:ppname) (st:st_term { match st.term with | Tm_ProofHintWithBinders { hint_type = RENAME _ } -> true | _ -> false }) -: T.Tac st_term + (check:check_t) +: T.Tac (checker_result_t g pre post_hint) = let Tm_ProofHintWithBinders ht = st.term in let { hint_type=RENAME { pairs; goal; tac_opt; elaborated }; binders=bs; t=body } = ht in match bs, goal with @@ -320,6 +306,7 @@ let check_renaming // ... let body = {st with term = Tm_ProofHintWithBinders { ht with binders = [] }; source = Sealed.seal false; } in + check g pre pre_typing post_hint res_ppname { st with term = Tm_ProofHintWithBinders { hint_type=ASSERT { p = goal; elaborated = true }; binders=bs; t=body }; source = Sealed.seal false; @@ -327,30 +314,66 @@ let check_renaming | [], None -> // if there is no goal, take the goal to be the full current pre - let lhs, rhs = rewrite_all (T.unseal st.source) g pairs pre pre elaborated tac_opt in - let t = { st with term = Tm_Rewrite { t1 = lhs; t2 = rhs; tac_opt; elaborated = true }; - source = Sealed.seal false; } in - { st with - term = Tm_Bind { binder = as_binder tm_unit; head = t; body }; - source = Sealed.seal false; - } + let rhs, pairs = rewrite_all (T.unseal st.source) g pairs pre pre elaborated tac_opt in + check_pairs g st.range pairs tac_opt; + let h2: slprop_equiv g rhs pre = RU.magic () in + let h1: tot_typing g rhs tm_slprop = RU.magic () in + let (| x, g', ty, ctxt', k |) = check g rhs h1 post_hint res_ppname body in + (| x, g', ty, ctxt', k_elab_equiv k h2 (VE_Refl _ _) |) | [], Some goal -> ( let goal, _ = PC.instantiate_term_implicits g goal None false in - let lhs, rhs = rewrite_all (T.unseal st.source) g pairs goal pre elaborated tac_opt in - let t = { st with term = Tm_Rewrite { t1 = lhs; t2 = rhs; tac_opt; elaborated = true }; + let rhs, _ = rewrite_all (T.unseal st.source) g pairs goal pre elaborated tac_opt in + let t = { st with term = Tm_Rewrite { t1 = goal; t2 = rhs; tac_opt; elaborated = true }; source = Sealed.seal false; } in + check g pre pre_typing post_hint res_ppname { st with term = Tm_Bind { binder = as_binder tm_unit; head = t; body }; source = Sealed.seal false; } ) #restart-solver #push-options "--z3rlimit_factor 2 --fuel 0 --ifuel 1" +let rec peel_binders k (ex: slprop) pre r + (g:env) frame (bs: list binder) (t:term) (t_typ: tot_typing g t tm_slprop) : + T.Tac (g':env {env_extends g' g} & t': slprop & xs: list (universe & typ & nvar) & + continuation_elaborator + g (frame `tm_star` t) + g' (frame `tm_star` t')) = + match bs with + | [] -> (|g, t, [], k_elab_unit _ _|) + | with_b::bs -> + match inspect_term t with + | Tm_ExistsSL u b body -> + let x = with_b.binder_ppname, fresh g in + let ty = mk_erased u b.binder_ty in + let g' = push_binding g (snd x) (fst x) ty in + let t' = open_term' body (mk_reveal u b.binder_ty (term_of_nvar x)) 0 in + let t'_typ : tot_typing g' t' tm_slprop = RU.magic () in + let (|g'', t'', bs', k'|) = peel_binders k ex pre r g' frame bs t' t'_typ in + (| g'', t'', (u,b.binder_ty,x)::bs', k_elab_trans (Pulse.Checker.Prover.elim_exists g frame u b body x g') k' |) + | _ -> + fail_doc g (Some r) [ + text <| (Printf.sprintf "Expected an existential quantifier with at least %d binders; but only found %s with %d binders" + k (show ex) (k - List.Tot.length bs)); + text "The context was:" ^^ + indent (pp <| canon_slprop_print pre) + ] +let open_st_term_with_reveals (t: st_term) (xs: list (universe & typ & nvar)) : Dv st_term = + // FIXME: this probably does something very wrong when the variables depend on each other... + let rec mk_subst xs (i: nat) : subst = + match xs with + | [] -> [] + | (u,t,x)::xs -> RT.DT i (mk_reveal u t (term_of_nvar x)) :: mk_subst xs (i+1) in + subst_st_term t (mk_subst (List.Tot.rev xs) 0) let check_wild (g:env) (pre:term) + (pre_typing:tot_typing g pre tm_slprop) + (post_hint:post_hint_opt g) + (res_ppname:ppname) (st:st_term { head_wild st }) -: T.Tac st_term + (check:check_t) +: T.Tac (checker_result_t g pre post_hint) = let Tm_ProofHintWithBinders ht = st.term in let open Pulse.PP in let { binders=bs; t=body } = ht in @@ -374,41 +397,30 @@ let check_wild | [ex] -> let k = List.Tot.length bs in - let rec peel_binders (n:nat) (t:term) : T.Tac st_term = - if n = 0 - then ( - let ex_body = t in - { st with term = Tm_ProofHintWithBinders { ht with hint_type = ASSERT { p = ex_body; elaborated = true } }} - ) - else ( - match inspect_term t with - | Tm_ExistsSL u b body -> peel_binders (n-1) body - | _ -> - fail_doc g (Some st.range) [ - text <| (Printf.sprintf "Expected an existential quantifier with at least %d binders; but only found %s with %d binders" - k (show ex) (k - n)); - text "The context was:" ^^ - indent (pp <| canon_slprop_print pre) - ] - ) - in - peel_binders k ex + let frame = list_as_slprop rest in + let ex_typ : tot_typing g ex tm_slprop = RU.magic () in + let (|g', ex', bs, k|) = peel_binders k ex pre st.range g frame bs ex ex_typ in + let body = open_st_term_with_reveals body bs in + let pre_typ : tot_typing g' (tm_star frame ex') tm_slprop = RU.magic () in + let (| x'', g'', t'', ctxt'', k' |) = + check g' (frame `tm_star` ex') pre_typ post_hint res_ppname body in + assume pre == (frame `tm_star` ex); + (| x'', g'', t'', ctxt'', k_elab_trans k k' |) #pop-options + // // v is a partially applied slprop with type t // add uvars for the remaining arguments // -let rec add_rem_uvs (g:env) (t:typ) (uvs:env { Env.disjoint g uvs }) (v:slprop) - : T.Tac (uvs:env { Env.disjoint g uvs } & slprop) = +let rec add_rem_uvs (g:env) (t:typ) (v:term) + : T.Tac slprop = match is_arrow t with - | None -> (| uvs, v |) + | None -> v | Some (b, qopt, c) -> - let x = fresh (push_env g uvs) in - let ppname = ppname_for_uvar b.binder_ppname in - let ct = open_comp_nv c (ppname, x) in - let uvs = Env.push_binding uvs x ppname b.binder_ty in - let v = tm_pureapp v qopt (tm_var {nm_index = x; nm_ppname = ppname}) in - add_rem_uvs g (comp_res ct) uvs v + let u, _ = PC.tc_term_phase1_with_type g tm_unknown b.binder_ty in + let ct = open_comp' c u 0 in + let v = tm_pureapp v qopt u in + add_rem_uvs g (comp_res ct) v #push-options "--fuel 0 --ifuel 0" let check @@ -428,8 +440,7 @@ let check allow_invert hint_type; match hint_type with | WILD -> - let st = check_wild g pre st in - check g pre pre_typing post_hint res_ppname st + check_wild g pre pre_typing post_hint res_ppname st check | SHOW_PROOF_STATE r -> let open FStar.Pprint in @@ -441,8 +452,7 @@ let check fail_doc_env true g (Some r) msg | RENAME {} -> - let st = check_renaming g pre st in - check g pre pre_typing post_hint res_ppname st + check_renaming g pre pre_typing post_hint res_ppname st check | REWRITE { t1; t2; tac_opt; elaborated } -> ( match bs with @@ -459,89 +469,56 @@ let check | ASSERT { p = v; elaborated } -> FStar.Tactics.BreakVC.break_vc(); // Some stabilization - let bs = infer_binder_types g bs v in - let (| uvs, v_opened, body_opened |) = open_binders g bs (mk_env (fstar_env g)) v body in - let v, body = v_opened, body_opened in + let (| uvs, v, body |) = open_binders_with_uvars g bs v body [] in let ctxt = Pulse.Checker.ImpureSpec.({ctxt_now = pre; ctxt_old = None}) in - let (| v, d |) = + let v = if elaborated then - PC.check_slprop (push_env g uvs) v + fst (PC.tc_term_phase1_with_type g v tm_slprop) else - ImpureSpec.purify_and_check_spec (push_env g uvs) ctxt v in - let (| g1, nts, _, pre', k_frame |) = Prover.prove false pre_typing uvs d in - // - // No need to check effect labels for the uvs solution here, - // since we are checking the substituted body anyway, - // if some of them are ghost when they shouldn't be, - // it will get caught - // + ImpureSpec.purify_spec g ctxt v in + let (| g1, pre', k_frame |) = Prover.prove st.range g pre v false in + let v = RU.deep_compress_safe v in + let body = body in // TODO compress + let h: tot_typing g1 v tm_slprop = PC.core_check_term _ _ _ _ in + let h: tot_typing g1 (tm_star v pre') tm_slprop = RU.magic () in // TODO: propagate through prover let (| x, x_ty, pre'', g2, k |) = - check g1 (tm_star (PS.nt_subst_term v nts) pre') - (RU.magic ()) - post_hint res_ppname (PS.nt_subst_st_term body nts) in + check g1 (tm_star v pre') h post_hint res_ppname body in (| x, x_ty, pre'', g2, k_elab_trans k_frame k |) | UNFOLD { p=v } | FOLD { p=v } -> - - let (| uvs, v_opened, body_opened |) = - let bs = infer_binder_types g bs v in - open_binders g bs (mk_env (fstar_env g)) v body in + let (| uvs, v, body |) = open_binders_with_uvars g bs v body [] in check_unfoldable g v; - let v_opened = purify_term g { ctxt_now = pre; ctxt_old = None } v_opened in - let v_opened, t_rem = PC.instantiate_term_implicits (push_env g uvs) v_opened None false in - - let uvs, v_opened = - let (| uvs_rem, v_opened |) = - add_rem_uvs (push_env g uvs) t_rem (mk_env (fstar_env g)) v_opened in - push_env uvs uvs_rem, v_opened in + let v = purify_term g { ctxt_now = pre; ctxt_old = None } v in + let v = + let v, t, _ = PC.tc_term_phase1 g v in + add_rem_uvs g t v in - let lhs, rhs, tac = + let lhs, rhs = match hint_type with | UNFOLD _ -> - let rhs, head_sym = unfold_defs (push_env g uvs) None v_opened in - v_opened, rhs, Pulse.Reflection.Util.slprop_equiv_unfold_tm head_sym + let rhs, head_sym = unfold_defs g None v in + v, rhs | FOLD { names=ns } -> - let lhs, head_sym = unfold_defs (push_env g uvs) ns v_opened in - lhs, v_opened, Pulse.Reflection.Util.slprop_equiv_fold_tm head_sym + let lhs, head_sym = unfold_defs g ns v in + lhs, v in - let uvs_bs = uvs |> bindings_with_ppname |> L.rev in - let uvs_closing = uvs_bs |> closing in - let lhs = subst_term lhs uvs_closing in - let rhs = subst_term rhs uvs_closing in - let body = subst_st_term body_opened uvs_closing in - let bs = close_binders uvs_bs in - (* Since this rewrite is easy enough to show by unification, we always - mark them with the slprop_equiv_norm tactic. *) - FStar.Tactics.BreakVC.break_vc (); - if RU.debug_at_level (fstar_env g) "fold" then begin - (* If we're running interactively, print out the context - and environment. *) - let open FStar.Pprint in - let open Pulse.PP in - let msg = [ - text "Elaborated fold/unfold to rewrite"; - prefix 2 1 (text "lhs:") (pp lhs); - prefix 2 1 (text "rhs:") (pp rhs); - ] in - info_doc_env g (Some st.range) msg - end; - let rw = mk_term (Tm_Rewrite { t1 = lhs; t2 = rhs; tac_opt = Some tac; elaborated = true }) st.range in - let rw = { rw with effect_tag = as_effect_hint STT_Ghost } in - - let st = mk_term (Tm_Bind { binder = as_binder (wr (`unit) st.range); head = rw; body }) st.range in - let st = { st with effect_tag = body.effect_tag } in - - let st = - match bs with - | [] -> st - | _ -> - let t = mk_term (Tm_ProofHintWithBinders { hint_type = ASSERT { p = lhs; elaborated = true }; binders = bs; t = st }) st.range in - { t with effect_tag = st.effect_tag } - in - check g pre pre_typing post_hint res_ppname st -#pop-options \ No newline at end of file + let (| g', pre_remaining, k |) = Prover.prove st.range g pre lhs false in + + let norm t = T.norm_term_env (elab_env g') [primops; iota] (RU.deep_compress_safe t) in + // Clean up beta-redexes introduced by unification + let rhs' = norm rhs in + let v' = norm v in + + let _: tot_typing g v' tm_slprop = PC.check_slprop_with_core g v' in + + let h1: tot_typing g' (tm_star pre_remaining rhs') tm_slprop = RU.magic () in + let h2: slprop_equiv g' (tm_star pre_remaining rhs') (tm_star lhs pre_remaining) = RU.magic () in + + let (| x, g'', ty, ctxt', k' |) = + check g' (tm_star pre_remaining rhs') h1 post_hint res_ppname body in + (| x, g'', ty, ctxt', k_elab_trans k (k_elab_equiv k' h2 (VE_Refl _ _)) |) diff --git a/src/checker/Pulse.Checker.Base.fst b/src/checker/Pulse.Checker.Base.fst index 438475b5c..1202de927 100644 --- a/src/checker/Pulse.Checker.Base.fst +++ b/src/checker/Pulse.Checker.Base.fst @@ -410,8 +410,8 @@ let continuation_elaborator_with_bind' (#g:env) (ctxt:term) assert (open_term (comp_post c1) x == comp_pre c2); // we closed e2 with x assume (~ (x `Set.mem` freevars_st e2_closed)); - if x `Set.mem` freevars (comp_post c2) - then fail g' None "Impossible: freevar clash when constructing continuation elaborator for bind, please file a bug-report" + if x `Set.mem` freevars (RU.deep_compress_safe (comp_post c2)) + then fail g' None ("Impossible: freevar clash when constructing continuation elaborator for bind, please file a bug-report" ^ show (comp_post c2)) else ( let t_typing, post_typing = RU.record_stats "bind_res_and_post_typing" fun _ -> @@ -617,17 +617,16 @@ let return_in_ctxt (g:env) (y:var) (y_ppname:ppname) (u:universe) (ty:term) (ctx let match_comp_res_with_post_hint (#g:env) (#t:st_term) (#c:comp_st) (d:st_typing g t c) (post_hint:post_hint_opt g) - : T.Tac (t':st_term & - c':comp_st & - st_typing g t' c') = + : T.Tac (c':comp_st { comp_pre c' == comp_pre c } & + st_typing g t c') = match post_hint with - | NoHint -> (| t, c, d |) + | NoHint -> (| c, d |) | TypeHint ret_ty | PostHint { ret_ty } -> let cres = comp_res c in if eq_tm cres ret_ty - then (| t, c, d |) + then (| c, d |) else match Pulse.Typing.Util.check_equiv_now (elab_env g) cres ret_ty with | None, issues -> let open Pulse.PP in @@ -647,7 +646,7 @@ let match_comp_res_with_post_hint (#g:env) (#t:st_term) (#c:comp_st) ST_SLPropEquiv _ c c' _ cpre_typing cres_typing cpost_typing d_equiv (VE_Refl _ _) (VE_Refl _ _) in - (| t, c', Pulse.Typing.Combinators.t_equiv d d_stequiv |) + (| c', Pulse.Typing.Combinators.t_equiv d d_stequiv |) #pop-options #pop-options @@ -1211,6 +1210,8 @@ let infer_post #g #ctxt (r:checker_result_t g ctxt NoHint) post; x; post_typing_src; post_typing=RU.magic() } in + let post = RU.beta_lax (elab_env g) post in // clean up spurious dependencies on variables + let post = RU.deep_compress post in let close_post = close_post x dom_g g1 (bindings_with_ppname g1) post in Pulse.Checker.Util.debug g "pulse.infer_post" (fun _ -> Printf.sprintf "Original postcondition: %s |= %s\nInferred postcondition: %s |= %s\n" diff --git a/src/checker/Pulse.Checker.Base.fsti b/src/checker/Pulse.Checker.Base.fsti index 8a483a005..b3cd0c5c1 100644 --- a/src/checker/Pulse.Checker.Base.fsti +++ b/src/checker/Pulse.Checker.Base.fsti @@ -195,9 +195,8 @@ type check_t = val match_comp_res_with_post_hint (#g:env) (#t:st_term) (#c:comp_st) (d:st_typing g t c) (post_hint:post_hint_opt g) - : T.Tac (t':st_term & - c':comp_st & - st_typing g t' c') + : T.Tac (c':comp_st { comp_pre c' == comp_pre c } & + st_typing g t c') val apply_checker_result_k (#g:env) (#ctxt:slprop) (#post_hint:post_hint_for_env g) (r:checker_result_t g ctxt (PostHint post_hint)) diff --git a/src/checker/Pulse.Checker.Bind.fst b/src/checker/Pulse.Checker.Bind.fst index 639419e9c..8421abc0d 100644 --- a/src/checker/Pulse.Checker.Bind.fst +++ b/src/checker/Pulse.Checker.Bind.fst @@ -21,7 +21,7 @@ open Pulse.Typing open Pulse.Typing.Combinators open Pulse.Checker.Base open Pulse.Checker.Pure -open Pulse.Checker.Prover +open Pulse.Checker.Prover.Util open Pulse.Show open Pulse.Checker.Util @@ -107,6 +107,8 @@ let check_binder_typ // let x : t = x0 in // rename x0 x; ... // to leverage the pure case + let ty = Pulse.RuntimeUtils.deep_compress_safe ty in + let t = Pulse.RuntimeUtils.deep_compress_safe t in if not (eq_tm ty t) then let open Pulse.PP in fail_doc g (Some e1.range) [ diff --git a/src/checker/Pulse.Checker.Exists.fst b/src/checker/Pulse.Checker.Exists.fst index 0fdc68325..e85d67dd8 100644 --- a/src/checker/Pulse.Checker.Exists.fst +++ b/src/checker/Pulse.Checker.Exists.fst @@ -89,7 +89,8 @@ let check_elim_exists if eq_univ u u' then let x = fresh g in let d = T_ElimExists g u ty p x ty_typing t_typing in - prove_post_hint (try_frame_pre false pre_typing (match_comp_res_with_post_hint d post_hint) res_ppname) post_hint t_rng + let (|_,d|) = match_comp_res_with_post_hint d post_hint in + prove_post_hint (try_frame_pre false pre_typing (|_,_,d|) res_ppname) post_hint t_rng else fail g (Some t_rng) (Printf.sprintf "check_elim_exists: universe checking failed, computed %s, expected %s" (P.univ_to_string u') (P.univ_to_string u)) @@ -131,7 +132,8 @@ let check_intro_exists check_term g witness T.E_Ghost b.binder_ty in let d = T_IntroExists g u b p witness ty_typing t_typing witness_typing in let (| c, d |) : (c:_ & st_typing g _ c) = (| _, d |) in - prove_post_hint (try_frame_pre false pre_typing (match_comp_res_with_post_hint d post_hint) res_ppname) + let (| c, d |) = match_comp_res_with_post_hint d post_hint in + prove_post_hint (try_frame_pre false pre_typing (|_,_,d|) res_ppname) post_hint (Pulse.RuntimeUtils.range_of_term t) #pop-options diff --git a/src/checker/Pulse.Checker.If.fst b/src/checker/Pulse.Checker.If.fst index 085acc0e1..0283f1673 100644 --- a/src/checker/Pulse.Checker.If.fst +++ b/src/checker/Pulse.Checker.If.fst @@ -21,7 +21,6 @@ open Pulse.Typing open Pulse.Typing.Combinators open Pulse.Checker.Pure open Pulse.Checker.Base -open Pulse.Checker.Prover module T = FStar.Tactics.V2 module Metatheory = Pulse.Typing.Metatheory @@ -77,7 +76,7 @@ let check let t = mk_term (Tm_ProofHintWithBinders { binders = []; - hint_type = RENAME { pairs = [(b, eq_v)]; goal=None; tac_opt=None; elaborated=true }; + hint_type = RENAME { pairs = [(b, eq_v)]; goal=None; tac_opt = Some Pulse.Reflection.Util.match_rename_tac_tm; elaborated=true }; t = br; }) br.range in @@ -119,10 +118,12 @@ let check apply_checker_result_k r ppname in let br_name = if is_then then "then" else "else" in - if hyp `Set.mem` freevars_st br - then fail g (Some br.range) - (Printf.sprintf "check_if: branch hypothesis is in freevars of checked %s branch" br_name) - else (| br, c, d |) + // if hyp `Set.mem` freevars_st br + // then fail g (Some br.range) + // (Printf.sprintf "check_if: branch hypothesis is in freevars of checked %s branch" br_name) + // else + assume not (hyp `Set.mem` freevars_st br); + (| br, c, d |) in let (| e1, c1, e1_typing |) = extract then_ true in let (| e2, c2, e2_typing |) = extract else_ false in diff --git a/src/checker/Pulse.Checker.ImpureSpec.fst b/src/checker/Pulse.Checker.ImpureSpec.fst index 61071777a..a95f81d10 100644 --- a/src/checker/Pulse.Checker.ImpureSpec.fst +++ b/src/checker/Pulse.Checker.ImpureSpec.fst @@ -18,11 +18,11 @@ module Pulse.Checker.ImpureSpec module R = FStar.Reflection.V2 module T = FStar.Tactics.V2 module RU = Pulse.RuntimeUtils -module PS = Pulse.Checker.Prover.Substs open FStar.List.Tot open Pulse.Syntax.Base open Pulse.Syntax.Pure open Pulse.Checker.Prover.RewritesTo +open Pulse.Checker.Prover.Normalize open Pulse.Checker.Pure open Pulse.Typing.Env open Pulse.Checker.Base @@ -77,7 +77,7 @@ let prove_this (g: env) (goal: slprop) (ctxt: list slprop) : T.Tac (option (list | Tm_Star a b -> Some [a; b] | _ -> - let rec try_match (ctxt: list slprop) : bool = + let rec try_match (ctxt: list slprop) : Dv bool = match ctxt with | [] -> false | c::ctxt -> @@ -108,9 +108,9 @@ let rec prove_loop (g: env) (goals: list slprop) (ctxt: list slprop) : T.Tac (li | None -> goals let prove (g: env) (goal: slprop) (ctxt: slprop) (r: range) : T.Tac unit = - let (| goal, _ |) = Pulse.Checker.Prover.normalize_slprop g goal true in + let (| goal, _ |) = normalize_slprop g goal true in let goal = slprop_as_list goal in - let (| ctxt, _ |) = Pulse.Checker.Prover.normalize_slprop g ctxt true in + let (| ctxt, _ |) = normalize_slprop g ctxt true in let ctxt = slprop_as_list ctxt in match prove_loop g goal ctxt with | [] -> () @@ -134,7 +134,7 @@ let symb_eval_stateful_app (g: env) (ctxt: slprop) (t: term) : T.Tac R.term = let x_ppn = mk_ppname_no_range "result" in let g' = push_binding g x (mk_ppname_no_range "result") ty in let post = open_term_nv post (x_ppn, x) in - let (| post, _ |) = Pulse.Checker.Prover.normalize_slprop g' post true in + let (| post, _ |) = normalize_slprop g' post true in match get_rewrites_to_from_post g x post with | None -> let head, _ = T.collect_app_ln t in @@ -263,7 +263,7 @@ let rec run_elim_core (g: env) (ctxt: list slprop) : T.Tac (env & list nvar & li g', xs, c::ctxt' let run_elim (g: env) (ctxt: slprop) : T.Tac (env & list nvar & slprop) = - let (| ctxt, _ |) = Pulse.Checker.Prover.normalize_slprop g ctxt true in + let (| ctxt, _ |) = normalize_slprop g ctxt true in let g', xs, ctxt = run_elim_core g (slprop_as_list ctxt) in g', xs, list_as_slprop ctxt diff --git a/src/checker/Pulse.Checker.IntroPure.fst b/src/checker/Pulse.Checker.IntroPure.fst index 8a47f38b9..2337efd15 100644 --- a/src/checker/Pulse.Checker.IntroPure.fst +++ b/src/checker/Pulse.Checker.IntroPure.fst @@ -59,4 +59,5 @@ let check let (| p, p_typing |) = check_prop g p in let pv = check_prop_validity g p p_typing in let st_typing = T_IntroPure _ _ p_typing pv in - prove_post_hint (try_frame_pre false pre_typing (match_comp_res_with_post_hint st_typing post_hint) res_ppname) post_hint t.range + let (| c,d |) = match_comp_res_with_post_hint st_typing post_hint in + prove_post_hint (try_frame_pre false pre_typing (|_,c,d|) res_ppname) post_hint t.range diff --git a/src/checker/Pulse.Checker.Match.fst b/src/checker/Pulse.Checker.Match.fst index 4a6813795..5b7d89126 100644 --- a/src/checker/Pulse.Checker.Match.fst +++ b/src/checker/Pulse.Checker.Match.fst @@ -289,7 +289,7 @@ let check_branch binders = []; hint_type = RENAME { pairs = [(sc, elab_p_tm)]; goal = None; - tac_opt = Some Pulse.Reflection.Util.match_rewrite_tac_tm; + tac_opt = Some Pulse.Reflection.Util.match_rename_tac_tm; elaborated = true }; t = e; }) e.range diff --git a/src/checker/Pulse.Checker.Par.fst b/src/checker/Pulse.Checker.Par.fst index 19ecf13e0..db45c806c 100644 --- a/src/checker/Pulse.Checker.Par.fst +++ b/src/checker/Pulse.Checker.Par.fst @@ -61,5 +61,6 @@ let check let x = fresh g in assume (comp_u cL == comp_u cR); let d = T_Par _ _ _ _ _ x cL_typing cR_typing eL_typing eR_typing in - prove_post_hint (try_frame_pre false pre_typing (match_comp_res_with_post_hint d post_hint) res_ppname) post_hint t.range + let (|_,d|) = match_comp_res_with_post_hint d post_hint in + prove_post_hint (try_frame_pre false pre_typing (|_,_,d|) res_ppname) post_hint t.range #pop-options diff --git a/src/checker/Pulse.Checker.Prover.Base.fst b/src/checker/Pulse.Checker.Prover.Base.fst deleted file mode 100644 index a95348211..000000000 --- a/src/checker/Pulse.Checker.Prover.Base.fst +++ /dev/null @@ -1,188 +0,0 @@ -(* - Copyright 2023 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module Pulse.Checker.Prover.Base - -open Pulse.Syntax -open Pulse.Typing -open Pulse.Checker.Base -open Pulse.Typing.Combinators -module RU = Pulse.RuntimeUtils -module T = FStar.Tactics.V2 - -let rec list_as_slprop' (vp:slprop) (fvps:list slprop) - : Tot slprop (decreases fvps) = - match fvps with - | [] -> vp - | hd::tl -> list_as_slprop' (tm_star vp hd) tl - -let rec canon_right_aux (g:env) (vps:list slprop) (f:slprop -> T.Tac bool) - : T.Tac (vps' : list slprop & - fvps : list slprop & - slprop_equiv g (list_as_slprop vps) (list_as_slprop' (list_as_slprop vps') fvps)) = - - match vps with - | [] -> (| [], [], VE_Refl _ _ |) - | hd::rest -> - if f hd - then begin - let (| vps', fvps, _ |) = canon_right_aux g rest f in - let v_eq = RU.magic #(slprop_equiv _ _ _) () in - // let v_eq - // : slprop_equiv g (list_as_slprop vps) - // (list_as_slprop (hd :: (vps' @ fvps))) - // = list_as_slprop_ctx g [hd] _ rest (vps' @ fvps) (VE_Refl _ _) v_eq - // in - // let v_eq - // : slprop_equiv g (list_as_slprop vps) - // (list_as_slprop ((vps'@[hd]) @ fvps)) - // = VE_Trans _ _ _ _ v_eq (VE_Sym _ _ _ (slprop_equiv_swap_equiv _ _ _ hd _ (VE_Refl _ _))) - // in - // let v_eq - // : slprop_equiv g (list_as_slprop vps) - // (list_as_slprop (vps'@(hd::fvps))) - // = VE_Trans _ _ _ _ v_eq (VE_Sym _ _ _ (list_as_slprop_assoc _ _ _ _)) in - (| vps', hd :: fvps, v_eq |) - end - else begin - let (| vps', pures, _ |) = canon_right_aux g rest f in - let v_eq = RU.magic #(slprop_equiv _ _ _ ) () in //list_as_slprop_ctx g [hd] _ _ _ (VE_Refl _ _) v_eq in - (| hd::vps', pures, v_eq |) - end - -module VP = Pulse.Checker.SLPropEquiv - -let canon_right (#g:env) (#ctxt:term) (#frame:term) - (ctxt_frame_typing:tot_typing g (tm_star ctxt frame) tm_slprop) - (f:slprop -> T.Tac bool) - : T.Tac (ctxt':term & - tot_typing g (tm_star ctxt' frame) tm_slprop & - continuation_elaborator g (tm_star ctxt frame) g (tm_star ctxt' frame)) - = let (| vps', pures, veq |) = canon_right_aux g (slprop_as_list ctxt) f in - let veq : slprop_equiv g ctxt (list_as_slprop' (list_as_slprop vps') pures) - = RU.magic () in - let veq : slprop_equiv g (tm_star ctxt frame) (tm_star (list_as_slprop' (list_as_slprop vps') pures) frame) - = VE_Ctxt _ _ _ _ _ veq (VE_Refl _ _) in - (| _, VP.slprop_equiv_typing_fwd ctxt_frame_typing veq, k_elab_equiv (k_elab_unit _ _) (VE_Refl _ _) veq |) - -#push-options "--fuel 0 --ifuel 0" -let elim_one (#g:env) - (ctxt:term) (frame:slprop) (p:slprop) - (ctxt_frame_p_typing:tot_typing g (tm_star (tm_star ctxt frame) p) tm_slprop) - (nx:ppname) (e1:st_term) (c1:comp { stateful_comp c1 /\ comp_pre c1 == p }) - (e1_typing:st_typing g e1 c1) - (uvs:env { disjoint uvs g }) - : T.Tac (g':env { env_extends g' g /\ disjoint uvs g' } & - ctxt':term & - tot_typing g' (tm_star ctxt' frame) tm_slprop & - continuation_elaborator g (tm_star (tm_star ctxt frame) p) g' (tm_star ctxt' frame)) = - - let ctxt_frame_typing = star_typing_inversion_l ctxt_frame_p_typing in - let x = fresh (push_env g uvs) in - let k = - continuation_elaborator_with_bind (tm_star ctxt frame) e1_typing ctxt_frame_p_typing (nx, x) in - let g' = push_binding g x nx (comp_res c1) in - let ctxt' = tm_star (open_term_nv (comp_post c1) (nx, x)) ctxt in - let veq - : slprop_equiv g' (tm_star (open_term_nv (comp_post c1) (nx, x)) (tm_star ctxt frame)) - (tm_star ctxt' frame) = VE_Assoc _ _ _ _ in - let k - : continuation_elaborator - g (tm_star (tm_star ctxt frame) p) - g' (tm_star ctxt' frame) = - k_elab_equiv - #g #g' - #(tm_star (tm_star ctxt frame) p) - #(tm_star (tm_star ctxt frame) p) - #(tm_star (open_term_nv (comp_post c1) (nx, x)) (tm_star ctxt frame)) - #(tm_star ctxt' frame) - k (VE_Refl g (tm_star (tm_star ctxt frame) p)) veq in - - let ctxt'_frame_typing : tot_typing g' (tm_star ctxt' frame) tm_slprop = RU.magic () in - env_extends_push g x ppname_default (comp_res c1); - (| g', ctxt', ctxt'_frame_typing, k |) - -let rec elim_all (#g:env) - (f:slprop -> T.Tac bool) - (mk:mk_t) - (#ctxt:term) (#frame:term) (ctxt_frame_typing:tot_typing g (tm_star ctxt frame) tm_slprop) - (uvs:env { disjoint uvs g }) - : T.Tac (bool & - (g':env { env_extends g' g /\ disjoint uvs g' } & - ctxt':term & - tot_typing g' (tm_star ctxt' frame) tm_slprop & - continuation_elaborator g (tm_star ctxt frame) g' (tm_star ctxt' frame))) - = match inspect_term ctxt with - | Tm_Star ctxt' p -> - let p_typing = - star_typing_inversion_r #_ #ctxt' #p (star_typing_inversion_l ctxt_frame_typing) in - if f p - then match mk #_ #p p_typing with - | Some (| nx, e1, c1, e1_typing |) -> - let (| g', _, ctxt_typing', k |) = - elim_one ctxt' frame p (RU.magic ()) nx e1 c1 e1_typing uvs in - let k - : continuation_elaborator g (tm_star (tm_star ctxt' frame) p) - g' (tm_star _ frame) = k in - let k - : continuation_elaborator g (tm_star (tm_star ctxt' p) frame) - g' (tm_star _ frame) = - k_elab_equiv k - (RU.magic ()) (VE_Refl _ _) in - let _, (| g'', ctxt'', ctxt_typing'', k' |) = - elim_all #g' f mk ctxt_typing' uvs in - true, (| g'', ctxt'', ctxt_typing'', k_elab_trans k k' |) - | None -> - false, (| g, ctxt, ctxt_frame_typing, k_elab_unit _ _ |) - else begin - false, (| g, ctxt, ctxt_frame_typing, k_elab_unit _ _ |) - end - | _ -> - false, (| g, ctxt, ctxt_frame_typing, k_elab_unit _ _ |) - -let add_elims_aux (#g:env) (#ctxt:term) (#frame:term) - (f:slprop -> T.Tac bool) - (mk:mk_t) - (ctxt_frame_typing:tot_typing g (tm_star ctxt frame) tm_slprop) - (uvs:env { disjoint uvs g }) - : T.Tac (bool & - (g':env { env_extends g' g /\ disjoint uvs g' } & - ctxt':term & - tot_typing g' (tm_star ctxt' frame) tm_slprop & - continuation_elaborator g (tm_star ctxt frame) g' (tm_star ctxt' frame))) - = let (| ctxt', ctxt'_typing, k |) = canon_right ctxt_frame_typing f in - let progress, (| g', ctxt'', ctxt''_typing, k' |) = - elim_all f mk ctxt'_typing uvs in - progress, (| g', ctxt'', ctxt''_typing, k_elab_trans k k' |) - -let rec add_elims (#g:env) (#ctxt:term) (#frame:term) - (f:slprop -> T.Tac bool) - (mk:mk_t) - (ctxt_typing:tot_typing g (tm_star ctxt frame) tm_slprop) - (uvs:env { disjoint uvs g }) - : T.Tac (g':env { env_extends g' g /\ disjoint uvs g' } & - ctxt':term & - tot_typing g' (tm_star ctxt' frame) tm_slprop & - continuation_elaborator g (tm_star ctxt frame) g' (tm_star ctxt' frame)) - = let progress, res = add_elims_aux f mk ctxt_typing uvs in - if not progress - then res - else ( - let (| g', ctxt', ctxt'_typing, k |) = res in - let (| g'', ctxt'', ctxt''_typing, k' |) = add_elims f mk ctxt'_typing uvs in - (| g'', ctxt'', ctxt''_typing, k_elab_trans k k' |) - ) -#pop-options \ No newline at end of file diff --git a/src/checker/Pulse.Checker.Prover.Base.fsti b/src/checker/Pulse.Checker.Prover.Base.fsti deleted file mode 100644 index 5bbb7d142..000000000 --- a/src/checker/Pulse.Checker.Prover.Base.fsti +++ /dev/null @@ -1,152 +0,0 @@ -(* - Copyright 2023 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module Pulse.Checker.Prover.Base - -open Pulse.Syntax -open Pulse.Typing -open Pulse.Typing.Combinators - -open Pulse.Checker.Base - -module T = FStar.Tactics.V2 -module PS = Pulse.Checker.Prover.Substs - -let slprop_typing (g:env) (t:term) = tot_typing g t tm_slprop - -// -// Scaffolding for adding elims -// -// Given a function f : slprop -> T.Tac bool that decides whether a slprop -// should be elim-ed, -// and an mk function to create the elim term, comp, and typing, -// add_elims will create a continuation_elaborator -// - -type mk_t = - #g:env -> - #v:slprop -> - tot_typing g v tm_slprop -> - T.Tac (option (x:ppname & - t:st_term & - c:comp { stateful_comp c /\ comp_pre c == v } & - st_typing g t c)) - -val add_elims (#g:env) (#ctxt:term) (#frame:term) - (f:slprop -> T.Tac bool) - (mk:mk_t) - (ctxt_typing:tot_typing g (tm_star ctxt frame) tm_slprop) - (uvs:env { disjoint uvs g }) - : T.Tac (g':env { env_extends g' g /\ disjoint uvs g' } & - ctxt':term & - tot_typing g' (tm_star ctxt' frame) tm_slprop & - continuation_elaborator g (tm_star ctxt frame) g' (tm_star ctxt' frame)) - -// -// Prover state -// - -noeq type preamble = { - g0 : env; - - ctxt : slprop; - frame : slprop; - ctxt_frame_typing : slprop_typing g0 (tm_star ctxt frame); - - goals : slprop; -} - -let op_Array_Access (ss:PS.ss_t) (t:term) = - PS.ss_term t ss - -let op_Star = tm_star - -noeq -type prover_state (preamble:preamble) = { - pg : g:env { g `env_extends` preamble.g0 }; - - remaining_ctxt : list slprop; - remaining_ctxt_frame_typing : slprop_typing pg (list_as_slprop remaining_ctxt * preamble.frame); - - uvs : uvs:env { disjoint uvs pg }; - ss : PS.ss_t; - rwr_ss : PS.ss_t; - - // - // these are the typed ss - // sometimes a step in the prover (e.g., intro exists) may compute these - // in such cases, the prover doesn't need to compute again - // so we cache them here - // - nts : option (nts:PS.nt_substs & effect_labels:list T.tot_or_ghost { - PS.well_typed_nt_substs pg uvs nts effect_labels /\ - PS.is_permutation nts ss - }); - - solved : slprop; - unsolved : list slprop; - - (* Note: we substitute solved, as it can contain some of the uvars in - uvs, but not remaining_ctxt or preamble.frame which do not mention - them (see their typing hypotheses above, in pg, which is disjoint - from uvs. Substituting would not be wrong, just useless. *) - k : continuation_elaborator preamble.g0 (preamble.ctxt * preamble.frame) - pg ((list_as_slprop remaining_ctxt * preamble.frame) * ss.(solved)); - - // // GM: Why isn't the rhs substituted here? - goals_inv : slprop_equiv (push_env pg uvs) - preamble.goals - (list_as_slprop unsolved * solved); - solved_inv : squash (freevars ss.(solved) `Set.subset` dom pg); - - progress : bool; (* used by prover and rules to check for progress and restart if so *) - allow_ambiguous : bool; - (* ^ true if we can skip the ambiguity check, used for functions like gather which - are inherently ambiguous, but safely so. *) -} - -let is_terminal (#preamble:_) (st:prover_state preamble) = - st.unsolved == [] - -irreducible -let extend_post_hint_opt_g (g:env) (post_hint:post_hint_opt g) (g1:env { g1 `env_extends` g }) - : p:post_hint_opt g1 { p == post_hint } = - match post_hint with - | PostHint post_hint -> - assert (g `env_extends` post_hint.g); - assert (g1 `env_extends` g); - assert (g1 `env_extends` post_hint.g); - PostHint post_hint - | p -> p - - -let ss_extends (ss1 ss2:PS.ss_t) = - Set.subset (PS.dom ss2) (PS.dom ss1) /\ - (forall (x:var). PS.contains ss2 x ==> PS.sel ss1 x == PS.sel ss2 x) - -let pst_extends (#preamble1 #preamble2:_) - (pst1 : prover_state preamble1) - (pst2 : prover_state preamble2) -= - pst1.pg `env_extends` pst2.pg /\ - pst1.uvs `env_extends` pst2.uvs /\ - pst1.ss `ss_extends` pst2.ss - -type prover_t = - #preamble:_ -> - pst1:prover_state preamble -> - T.Tac (pst2:prover_state preamble { pst2 `pst_extends` pst1 /\ - is_terminal pst2 }) diff --git a/src/checker/Pulse.Checker.Prover.ElimExists.fst b/src/checker/Pulse.Checker.Prover.ElimExists.fst deleted file mode 100644 index baf20b829..000000000 --- a/src/checker/Pulse.Checker.Prover.ElimExists.fst +++ /dev/null @@ -1,121 +0,0 @@ -(* - Copyright 2023 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module Pulse.Checker.Prover.ElimExists -#set-options "--z3rlimit 10" //base rlimit x 2 -open Pulse.Syntax -open Pulse.Typing -open Pulse.Typing.Combinators - -module T = FStar.Tactics.V2 -module RU = Pulse.RuntimeUtils - -open Pulse.Checker.SLPropEquiv - -open Pulse.Checker.Prover.Base - -let should_elim_exists (v:slprop) : T.Tac bool = - match inspect_term v with - | Tm_ExistsSL _ _ _ -> true - | _ -> false - -let mk (#g:env) (#v:slprop) (v_typing:tot_typing g v tm_slprop) - : T.Tac (option (x:ppname & - t:st_term & - c:comp { stateful_comp c /\ comp_pre c == v } & - st_typing g t c)) = - - match inspect_term v with - | Tm_ExistsSL u { binder_ppname=nm; binder_ty = t } p -> - let x = fresh g in - let c = Pulse.Typing.comp_elim_exists u t p (nm, x) in - let tm_typing : st_typing g _ c = - T_ElimExists g (comp_u c) t p x (RU.magic()) (RU.magic()) - in - Some (| nm, _, c, tm_typing |) - | _ -> None - -let elim_exists_frame (#g:env) (#ctxt #frame:slprop) - (ctxt_frame_typing:tot_typing g (ctxt * frame) tm_slprop) - (uvs:env { disjoint uvs g }) - : T.Tac (g':env { env_extends g' g /\ disjoint uvs g' } & - ctxt':term & - tot_typing g' (ctxt' * frame) tm_slprop & - continuation_elaborator g (ctxt * frame) g' (ctxt' * frame)) = - add_elims should_elim_exists mk ctxt_frame_typing uvs - -let elim_exists (#g:env) (#ctxt:term) - (ctxt_typing:tot_typing g ctxt tm_slprop) - : T.Tac (g':env { env_extends g' g } & - ctxt':term & - tot_typing g' ctxt' tm_slprop & - continuation_elaborator g ctxt g' ctxt') = - - let ctxt_emp_typing : tot_typing g (tm_star ctxt tm_emp) tm_slprop = RU.magic () in - let (| g', ctxt', ctxt'_emp_typing, k |) = - elim_exists_frame ctxt_emp_typing (mk_env (fstar_env g)) in - let k = k_elab_equiv k (VE_Trans _ _ _ _ (VE_Comm _ _ _) (VE_Unit _ _)) - (VE_Trans _ _ _ _ (VE_Comm _ _ _) (VE_Unit _ _)) in - (| g', ctxt', star_typing_inversion_l ctxt'_emp_typing, k |) - -let elim_exists_pst (#preamble:_) (pst:prover_state preamble) - : T.Tac (list (list Pprint.document) & pst':prover_state preamble { pst' `pst_extends` pst} ) = - - (* Hacking progress checking: we eliminate all exists, so if - there's any in the ctxt then we will make progress. *) - let prog = List.Tot.existsb (fun t -> Tm_ExistsSL? (inspect_term t)) pst.remaining_ctxt in - - // if not prog then [], pst else - let (| g', remaining_ctxt', ty, k |) = - elim_exists_frame - #pst.pg - #(list_as_slprop pst.remaining_ctxt) - #(preamble.frame * pst.ss.(pst.solved)) - (RU.magic ()) - pst.uvs in - - let k - : continuation_elaborator - pst.pg (list_as_slprop pst.remaining_ctxt * (preamble.frame * pst.ss.(pst.solved))) - g' (remaining_ctxt' * (preamble.frame * pst.ss.(pst.solved))) = k in - - // some *s - let k - : continuation_elaborator - pst.pg ((list_as_slprop pst.remaining_ctxt * preamble.frame) * pst.ss.(pst.solved)) - g' ((remaining_ctxt' * preamble.frame) * pst.ss.(pst.solved)) = - - k_elab_equiv k (RU.magic ()) (RU.magic ()) in - - let k_new - : continuation_elaborator - preamble.g0 (preamble.ctxt * preamble.frame) - g' ((remaining_ctxt' * preamble.frame) * pst.ss.(pst.solved)) = - k_elab_trans pst.k k in - - assume (list_as_slprop (slprop_as_list remaining_ctxt') == remaining_ctxt'); - - [], - { pst with - progress = prog; - pg = g'; - remaining_ctxt = slprop_as_list remaining_ctxt'; - remaining_ctxt_frame_typing = RU.magic (); - nts = None; - k = k_new; - goals_inv = RU.magic (); // weakening of pst.goals_inv - solved_inv = () - } diff --git a/src/checker/Pulse.Checker.Prover.ElimExists.fsti b/src/checker/Pulse.Checker.Prover.ElimExists.fsti deleted file mode 100644 index b1d803260..000000000 --- a/src/checker/Pulse.Checker.Prover.ElimExists.fsti +++ /dev/null @@ -1,35 +0,0 @@ -(* - Copyright 2023 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module Pulse.Checker.Prover.ElimExists - -open Pulse.Syntax -open Pulse.Typing - -open Pulse.Checker.Base -open Pulse.Checker.Prover.Base - -module T = FStar.Tactics.V2 - -val elim_exists (#g:env) (#ctxt:term) (ctxt_typing:tot_typing g ctxt tm_slprop) - : T.Tac (g':env { env_extends g' g } & - ctxt':term & - tot_typing g' ctxt' tm_slprop & - continuation_elaborator g ctxt g' ctxt') - -val elim_exists_pst (#preamble:_) (pst:prover_state preamble) - : T.Tac (list (list Pprint.document) & pst':prover_state preamble { pst' `pst_extends` pst }) - ///\ pst'.unsolved == pst.unsolved }) diff --git a/src/checker/Pulse.Checker.Prover.ElimPure.fst b/src/checker/Pulse.Checker.Prover.ElimPure.fst deleted file mode 100644 index 25ff68628..000000000 --- a/src/checker/Pulse.Checker.Prover.ElimPure.fst +++ /dev/null @@ -1,178 +0,0 @@ -(* - Copyright 2023 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module Pulse.Checker.Prover.ElimPure -#set-options "--z3rlimit 10" //base rlimit x 2 -module RT = FStar.Reflection.Typing -module R = FStar.Reflection.V2 -module T = FStar.Tactics.V2 -open FStar.List.Tot -open Pulse.Syntax -open Pulse.Checker.Pure -open Pulse.Checker.SLPropEquiv - -open Pulse.Typing -open Pulse.Typing.Combinators -module Metatheory = Pulse.Typing.Metatheory -open Pulse.Reflection.Util -open Pulse.Checker.Prover.Base -open Pulse.Checker.Prover.Util -module PS = Pulse.Checker.Prover.Substs -module RU = Pulse.RuntimeUtils - -let elim_pure_head = - let elim_pure_explicit_lid = mk_pulse_lib_core_lid "elim_pure_explicit" in - tm_fvar (as_fv elim_pure_explicit_lid) - -let elim_pure_head_ty = - // let open Pulse.Steel.Wrapper in - // let open Steel.Effect.Common in - let squash_p = mk_squash u0 (RT.bound_var 0) in - let pure_p = mk_pure (RT.bound_var 0) in - let post = - mk_abs squash_p R.Q_Explicit (R.pack_ln (R.Tv_FVar (R.pack_fv emp_lid))) - in - let cod = mk_stt_ghost_comp u0 squash_p tm_emp_inames pure_p post in - mk_arrow - (R.pack_ln (R.Tv_FVar (R.pack_fv R.prop_qn)), R.Q_Explicit) - cod - // Following crashes in extraction - // `(p:prop -> stt_ghost (squash p) emp_inames - // (pure p) - // (fun _ -> emp)) - -let wr t = wr t Range.range_0 - -let elim_pure_head_typing (g:env) - : tot_typing g elim_pure_head (wr elim_pure_head_ty) - = admit() - -let mk_elim_pure (p:term) - : st_term - = let t = Tm_ST { t = T.mk_app elim_pure_head [p, T.Q_Explicit]; args = [] } in - wtag (Some STT_Ghost) t - - -let elim_pure_comp (p:term) = - let st : st_comp = { - u=u_zero; - res=wr (mk_squash u0 p); - pre=tm_pure (wr p); - post=tm_emp - } in - C_STGhost tm_emp_inames st - -let elim_pure_typing (g:env) (p:term) - (p_prop:tot_typing g (wr p) (wr RT.tm_prop)) - : st_typing g (mk_elim_pure (wr p)) (elim_pure_comp p) - = admit() - -let is_elim_pure (vp:term) : T.Tac bool = - match inspect_term vp with - | Tm_Pure _ -> true - | _ -> false - -let mk (#g:env) (#v:slprop) (v_typing:tot_typing g v tm_slprop) - : T.Tac (option (x:ppname & - t:st_term & - c:comp { stateful_comp c /\ comp_pre c == v } & - st_typing g t c)) = - match inspect_term v with - | Tm_Pure pp -> - let p_typing = - Metatheory.pure_typing_inversion #g #(wr pp) v_typing in - debug_prover g (fun _ -> Printf.sprintf "elim_pure: %s\n" (T.term_to_string pp)); - Some (| ppname_default, - mk_elim_pure (wr pp), - elim_pure_comp pp, - elim_pure_typing g pp p_typing |) - | _ -> None - -let elim_pure_frame (#g:env) (#ctxt:term) (#frame:term) - (ctxt_frame_typing:tot_typing g (ctxt * frame) tm_slprop) - (uvs:env { disjoint uvs g }) - : T.Tac (g':env { env_extends g' g /\ disjoint uvs g' } & - ctxt':term & - tot_typing g' (ctxt' * frame) tm_slprop & - continuation_elaborator g (ctxt * frame) g' (ctxt' * frame)) = - add_elims is_elim_pure mk ctxt_frame_typing uvs - -let elim_pure (#g:env) (#ctxt:term) (ctxt_typing:tot_typing g ctxt tm_slprop) - : T.Tac (g':env { env_extends g' g } & - ctxt':term & - tot_typing g' ctxt' tm_slprop & - continuation_elaborator g ctxt g' ctxt') = - let ctxt_emp_typing : tot_typing g (tm_star ctxt tm_emp) tm_slprop = RU.magic () in - let (| g', ctxt', ctxt'_emp_typing, k |) = - elim_pure_frame ctxt_emp_typing (mk_env (fstar_env g)) in - let k = k_elab_equiv k (VE_Trans _ _ _ _ (VE_Comm _ _ _) (VE_Unit _ _)) - (VE_Trans _ _ _ _ (VE_Comm _ _ _) (VE_Unit _ _)) in - (| g', ctxt', star_typing_inversion_l ctxt'_emp_typing, k |) - -// a lot of this is copy-pasted from elim_exists_pst, -// can add a common function to Prover.Common -let elim_pure_pst (#preamble:_) (pst:prover_state preamble) - : T.Tac (pst':prover_state preamble { pst' `pst_extends` pst /\ - pst'.unsolved == pst.unsolved }) = - - (* Hacking progress checking: we eliminate all exists, so if - there's any in the ctxt then we will make progress. *) - let prog = T.existsb is_elim_pure pst.remaining_ctxt in - if not prog then pst else - - let (| g', remaining_ctxt', ty, k |) = - elim_pure_frame - #pst.pg - #(list_as_slprop pst.remaining_ctxt) - #(preamble.frame * pst.ss.(pst.solved)) - (RU.magic ()) - pst.uvs in - - let rwr_ss : PS.ss_t = // extends pst.ss with new rewrites_to_p substitutions - RewritesTo.get_new_substs_from_env pst.pg g' pst.rwr_ss in - - let k - : continuation_elaborator - pst.pg (list_as_slprop pst.remaining_ctxt * (preamble.frame * pst.ss.(pst.solved))) - g' (remaining_ctxt' * (preamble.frame * pst.ss.(pst.solved))) = k in - - // some *s - let k - : continuation_elaborator - pst.pg ((list_as_slprop pst.remaining_ctxt * preamble.frame) * pst.ss.(pst.solved)) - g' ((remaining_ctxt' * preamble.frame) * pst.ss.(pst.solved)) = - - k_elab_equiv k (RU.magic ()) (RU.magic ()) in - - let k_new - : continuation_elaborator - preamble.g0 (preamble.ctxt * preamble.frame) - g' ((remaining_ctxt' * preamble.frame) * pst.ss.(pst.solved)) = - k_elab_trans pst.k k in - - assume (list_as_slprop (slprop_as_list remaining_ctxt') == remaining_ctxt'); - - { pst with - progress=prog; - pg = g'; - remaining_ctxt = slprop_as_list remaining_ctxt'; - remaining_ctxt_frame_typing = RU.magic (); - rwr_ss; - nts = None; - k = k_new; - goals_inv = RU.magic (); // weakening of pst.goals_inv - solved_inv = (); - } diff --git a/src/checker/Pulse.Checker.Prover.ElimPure.fsti b/src/checker/Pulse.Checker.Prover.ElimPure.fsti deleted file mode 100644 index f8577b3b0..000000000 --- a/src/checker/Pulse.Checker.Prover.ElimPure.fsti +++ /dev/null @@ -1,36 +0,0 @@ -(* - Copyright 2023 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module Pulse.Checker.Prover.ElimPure - -open Pulse.Syntax -open Pulse.Typing - -open Pulse.Checker.Base -open Pulse.Checker.Prover.Base - -module T = FStar.Tactics.V2 - - -val elim_pure (#g:env) (#ctxt:term) (ctxt_typing:tot_typing g ctxt tm_slprop) - : T.Tac (g':env { env_extends g' g } & - ctxt':term & - tot_typing g' ctxt' tm_slprop & - continuation_elaborator g ctxt g' ctxt') - -val elim_pure_pst (#preamble:_) (pst:prover_state preamble) - : T.Tac (pst':prover_state preamble { pst' `pst_extends` pst /\ - pst'.unsolved == pst.unsolved }) \ No newline at end of file diff --git a/src/checker/Pulse.Checker.Prover.ElimWithPure.fst b/src/checker/Pulse.Checker.Prover.ElimWithPure.fst deleted file mode 100644 index 49a2dacba..000000000 --- a/src/checker/Pulse.Checker.Prover.ElimWithPure.fst +++ /dev/null @@ -1,139 +0,0 @@ -(* - Copyright 2025 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module Pulse.Checker.Prover.ElimWithPure - -module T = FStar.Tactics.V2 -open Pulse.Syntax - -open Pulse.Typing -open Pulse.Checker.Base -open Pulse.Checker.Prover.Base -open Pulse.Checker.Prover.Util -module PS = Pulse.Checker.Prover.Substs -module RU = Pulse.RuntimeUtils - -let elim_with_pure_head = - let elim_pure_explicit_lid = ["Pulse"; "Lib"; "WithPure"; "elim_with_pure"] in - tm_fvar (as_fv elim_pure_explicit_lid) - -let wr t = wr t Range.range_0 - -let mk_elim_with_pure (pred: term) (p:term) - : st_term - = let t = - Tm_ST { - t = T.mk_app - (tm_pureapp elim_with_pure_head None pred) - [tm_pureabs ppname_default.name (mk_squash u0 pred) None p Range.range_0, T.Q_Explicit]; - args = []; - } - in - wtag (Some STT_Ghost) t - - -let elim_with_pure_comp (p:term) (n:ppname) (v:term) = - let st : st_comp = { - u=u_zero; - res=wr (mk_squash u0 p); - pre=tm_with_pure (wr p) n (wr v); - post=open_term' v unit_const 0; - } in - C_STGhost tm_emp_inames st - -let is_elim_with_pure (vp:term) : T.Tac bool = - match inspect_term vp with - | Tm_WithPure .. -> true - | _ -> false - -#push-options "--fuel 0 --ifuel 0" -let mk (#g:env) (#v:slprop) (v_typing:tot_typing g v tm_slprop) - : T.Tac (option (x:ppname & - t:st_term & - c:comp { stateful_comp c /\ comp_pre c == v } & - st_typing g t c)) = - match inspect_term v with - | Tm_WithPure pred n pp -> - debug_prover g (fun _ -> Printf.sprintf "elim_with_pure: %s\n" (T.term_to_string v)); - let t = mk_elim_with_pure (wr pred) (wr pp) in - let c = elim_with_pure_comp pred n pp in - let typ : st_typing g t c = RU.magic () in - Some (| ppname_default, t, c, typ |) - | _ -> None - -let elim_with_pure_frame (#g:env) (#ctxt:term) (#frame:term) - (ctxt_frame_typing:tot_typing g (ctxt * frame) tm_slprop) - (uvs:env { disjoint uvs g }) - : T.Tac (g':env { env_extends g' g /\ disjoint uvs g' } & - ctxt':term & - tot_typing g' (ctxt' * frame) tm_slprop & - continuation_elaborator g (ctxt * frame) g' (ctxt' * frame)) = - add_elims is_elim_with_pure mk ctxt_frame_typing uvs - -// a lot of this is copy-pasted from elim_exists_pst, -// can add a common function to Prover.Common -#push-options "--z3rlimit_factor 4 --ifuel 1" -let elim_with_pure_pst (#preamble:_) (pst:prover_state preamble) - : T.Tac (pst':prover_state preamble { pst' `pst_extends` pst /\ - pst'.unsolved == pst.unsolved }) = - - let prog = T.existsb is_elim_with_pure pst.remaining_ctxt in - if not prog then pst else - let (| g', remaining_ctxt', ty, k |) = - elim_with_pure_frame - #pst.pg - #(list_as_slprop pst.remaining_ctxt) - #(preamble.frame * pst.ss.(pst.solved)) - (RU.magic ()) - pst.uvs in - - let rwr_ss : PS.ss_t = // extends pst.ss with new rewrites_to_p substitutions - RewritesTo.get_new_substs_from_env pst.pg g' pst.rwr_ss in - - let k - : continuation_elaborator - pst.pg (list_as_slprop pst.remaining_ctxt * (preamble.frame * pst.ss.(pst.solved))) - g' (remaining_ctxt' * (preamble.frame * pst.ss.(pst.solved))) = k in - - // some *s - let k - : continuation_elaborator - pst.pg ((list_as_slprop pst.remaining_ctxt * preamble.frame) * pst.ss.(pst.solved)) - g' ((remaining_ctxt' * preamble.frame) * pst.ss.(pst.solved)) = - - k_elab_equiv k (RU.magic ()) (RU.magic ()) in - - let k_new - : continuation_elaborator - preamble.g0 (preamble.ctxt * preamble.frame) - g' ((remaining_ctxt' * preamble.frame) * pst.ss.(pst.solved)) = - k_elab_trans pst.k k in - - assume (list_as_slprop (slprop_as_list remaining_ctxt') == remaining_ctxt'); - - { pst with - progress=true; - pg = g'; - remaining_ctxt = slprop_as_list remaining_ctxt'; - remaining_ctxt_frame_typing = RU.magic (); - rwr_ss; - nts = None; - k = k_new; - goals_inv = RU.magic (); // weakening of pst.goals_inv - solved_inv = (); - } -#pop-options -#pop-options \ No newline at end of file diff --git a/src/checker/Pulse.Checker.Prover.ElimWithPure.fsti b/src/checker/Pulse.Checker.Prover.ElimWithPure.fsti deleted file mode 100644 index c95a2af8e..000000000 --- a/src/checker/Pulse.Checker.Prover.ElimWithPure.fsti +++ /dev/null @@ -1,24 +0,0 @@ -(* - Copyright 2025 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module Pulse.Checker.Prover.ElimWithPure - -open Pulse.Checker.Prover.Base -module T = FStar.Tactics.V2 - -val elim_with_pure_pst (#preamble:_) (pst:prover_state preamble) - : T.Tac (pst':prover_state preamble { pst' `pst_extends` pst /\ - pst'.unsolved == pst.unsolved }) \ No newline at end of file diff --git a/src/checker/Pulse.Checker.Prover.Explode.fst b/src/checker/Pulse.Checker.Prover.Explode.fst deleted file mode 100644 index 1290634cf..000000000 --- a/src/checker/Pulse.Checker.Prover.Explode.fst +++ /dev/null @@ -1,98 +0,0 @@ -(* - Copyright 2023 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module Pulse.Checker.Prover.Explode - -module T = FStar.Tactics - -open Pulse.Syntax -open Pulse.Typing -open Pulse.PP - -open Pulse.Checker.Base -open Pulse.Checker.Prover.Base -open Pulse.Typing.Combinators - -let has_structure (q:slprop) : bool = - match inspect_term q with - | Tm_Star _ _ -> true - | _ -> false - -val slprop_as_list_roundtrip (g:env) (v:slprop) - : slprop_equiv g v (list_as_slprop (slprop_as_list v)) -let slprop_as_list_roundtrip g v = admit() - -let __explode1 - (g:env) - (q:slprop) -: T.Tac (option (qs: list slprop & slprop_equiv g q (list_as_slprop qs))) -= - // info_doc pst.pg None [ - // text "Trying to explode" ^/^ pp q_ss; - // text "pst.ss = " ^/^ pp pst.ss; - // text "pst.uvs = " ^/^ pp pst.uvs; - // text "q= " ^/^ pp q; - // text "q_ss= " ^/^ pp q_ss; - // ]; - if has_structure q - then ( - // info_doc pst.pg None [ - // text "Exploding" ^/^ pp q_ss; - // text "into" ^/^ pp (slprop_as_list q_ss); - // ]; - Some (| slprop_as_list q, slprop_as_list_roundtrip _ _ |) - ) else None - -let explode1 - (#preamble:_) - (pst:prover_state preamble) - (q:slprop) -: T.Tac (option (qs: list slprop & slprop_equiv (push_env pst.pg pst.uvs) pst.ss.(q) (list_as_slprop qs))) -= - let q_ss = pst.ss.(q) in - __explode1 (push_env pst.pg pst.uvs) q_ss - -let rec explode_aux - (#preamble:_) - (pst:prover_state preamble) - (prog : bool) - (acc : list slprop) (todo : list slprop) : T.Tac (list slprop & bool) -= - match todo with - | [] -> acc, prog - | q::qs -> - let acc', prog' = - match explode1 pst q with - | Some (|qs, _|) -> qs, true - | None -> [q], false - in - explode_aux pst (prog || prog') (acc @ acc') qs - -let explode - (#preamble:_) (pst:prover_state preamble) -: T.Tac (list (list Pprint.document) & pst':prover_state preamble {pst' `pst_extends` pst}) -= - let remaining_ctxt, p1 = explode_aux pst false [] pst.remaining_ctxt in - let unsolved', p2 = explode_aux pst false [] pst.unsolved in - [], - { pst with - unsolved = unsolved'; - goals_inv = magic(); - remaining_ctxt = remaining_ctxt; - remaining_ctxt_frame_typing = magic(); - k = (admit(); pst.k); - progress = p1 || p2; - } diff --git a/src/checker/Pulse.Checker.Prover.Explode.fsti b/src/checker/Pulse.Checker.Prover.Explode.fsti deleted file mode 100644 index a0b48a745..000000000 --- a/src/checker/Pulse.Checker.Prover.Explode.fsti +++ /dev/null @@ -1,26 +0,0 @@ -(* - Copyright 2023 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module Pulse.Checker.Prover.Explode - -open Pulse.Checker.Prover.Base -module T = FStar.Tactics - -(* Explodes any stars in the unsolved field of the pst into -its individual components. *) -val explode - (#preamble:_) (pst:prover_state preamble) -: T.Tac (list (list Pprint.document) & pst':prover_state preamble {pst' `pst_extends` pst}) diff --git a/src/checker/Pulse.Checker.Prover.IntroExists.fst b/src/checker/Pulse.Checker.Prover.IntroExists.fst deleted file mode 100644 index 109ab4d5a..000000000 --- a/src/checker/Pulse.Checker.Prover.IntroExists.fst +++ /dev/null @@ -1,358 +0,0 @@ -(* - Copyright 2023 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module Pulse.Checker.Prover.IntroExists -#set-options "--z3rlimit 20" //base rlimit x 4 -open Pulse.Syntax -open Pulse.Typing -open Pulse.Typing.Combinators -open Pulse.Typing.Metatheory -open Pulse.Checker.SLPropEquiv -open Pulse.Checker.Prover.Base -open Pulse.Checker.Base - -module RT = FStar.Reflection.Typing -module RU = Pulse.RuntimeUtils -module T = FStar.Tactics.V2 -module PS = Pulse.Checker.Prover.Substs -module Metatheory = Pulse.Typing.Metatheory - -let coerce_eq (#a #b:Type) (x:a) (_:squash (a == b)) : b = x - -// there will be some side conditions related to the typings -let k_intro_exists (#g:env) (#u:universe) (#b:binder) (#p:slprop) - (ex_typing:tot_typing g (tm_exists_sl u b p) tm_slprop) - (#e:term) - (e_typing:tot_typing g e b.binder_ty) - (#frame:slprop) - (frame_typing:tot_typing g frame tm_slprop) - : T.Tac (continuation_elaborator g (frame * subst_term p [ RT.DT 0 e ]) - g (frame * tm_exists_sl u b p)) = - - let t = wtag (Some STT_Ghost) (Tm_IntroExists { p = tm_exists_sl u b p; - witnesses = [e] }) in - - let c = comp_intro_exists u b p e in - - let t_typing = T_IntroExists g u b p e (RU.magic ()) ex_typing (lift_typing_to_ghost_typing e_typing) in - - assert (comp_pre c == subst_term p [ RT.DT 0 e ]); - assert (comp_post c == tm_exists_sl u b p); - - let x = fresh g in - assume (open_term (comp_post c) x == comp_post c); - - let ppname = mk_ppname_no_range "_pintroe" in - let k - : continuation_elaborator - g (frame * subst_term p [ RT.DT 0 e ]) - (push_binding g x ppname_default (comp_res c)) (tm_exists_sl u b p * frame) = - continuation_elaborator_with_bind frame t_typing (RU.magic ()) (ppname, x) in - - let k - : continuation_elaborator - g (frame * subst_term p [ RT.DT 0 e ]) - (push_binding g x ppname_default (comp_res c)) (frame * tm_exists_sl u b p) = - k_elab_equiv k (VE_Refl _ _) (VE_Comm _ _ _) in - - fun post_hint r -> - let (| t1, c1, d1 |) = r in - let d1 : st_typing g t1 c1 = d1 in - let empty_env = mk_env (fstar_env g) in - assert (equal g (push_env g empty_env)); - assert (equal (push_env (push_binding g x ppname_default (comp_res c)) empty_env) - (push_binding g x ppname_default (comp_res c))); - let d1 : st_typing (push_binding g x ppname_default (comp_res c)) t1 c1 = - Pulse.Typing.Metatheory.st_typing_weakening - g - empty_env - t1 c1 d1 - (push_binding g x ppname_default (comp_res c)) in - - k post_hint (| t1, c1, d1 |) - -#push-options "--z3rlimit_factor 8 --ifuel 2 --fuel 1" - -let intro_exists (#preamble:_) (pst:prover_state preamble) - (u:universe) (b:binder) (body:slprop) - (unsolved':list slprop) - (_:squash (pst.unsolved == (tm_exists_sl u b body)::unsolved')) - (prover:prover_t) - : T.Tac (pst':prover_state preamble { pst' `pst_extends` pst /\ - is_terminal pst' }) = - - let x = fresh (push_env pst.pg pst.uvs) in - let ppname = ppname_for_uvar b.binder_ppname in - let px = ppname, x in - let preamble_sub = { - g0 = pst.pg; - ctxt = list_as_slprop pst.remaining_ctxt; - frame = preamble.frame * pst.ss.(pst.solved); - ctxt_frame_typing = RU.magic (); - goals = open_term_nv body px * (list_as_slprop unsolved'); - } in - let k_sub: - continuation_elaborator - preamble_sub.g0 (preamble_sub.ctxt * preamble_sub.frame) - pst.pg ((list_as_slprop (slprop_as_list preamble_sub.ctxt) * preamble_sub.frame) * pst.ss.(tm_emp)) = - let k = k_elab_unit preamble_sub.g0 (preamble_sub.ctxt * preamble_sub.frame) in - let k = k_elab_equiv k - (VE_Refl _ _) - (RU.magic () <: - slprop_equiv - preamble_sub.g0 - (preamble_sub.ctxt * preamble_sub.frame) - ((list_as_slprop (slprop_as_list preamble_sub.ctxt) * preamble_sub.frame) * pst.ss.(tm_emp))) in - coerce_eq k () - in - assume (pst.ss.(tm_emp) == tm_emp); - let pst_sub : prover_state preamble_sub = { - pg = pst.pg; - remaining_ctxt = slprop_as_list preamble_sub.ctxt; - remaining_ctxt_frame_typing = RU.magic (); - uvs = push_binding pst.uvs x ppname b.binder_ty; - ss = pst.ss; - rwr_ss = pst.rwr_ss; - nts = None; - solved = tm_emp; - unsolved = (slprop_as_list (open_term_nv body px)) @ unsolved'; - k = k_sub; - goals_inv = RU.magic (); - solved_inv = (); - progress = false; - allow_ambiguous = pst.allow_ambiguous; - } in - let pst_sub = prover pst_sub in - let pst_sub_goals_inv - : slprop_equiv (push_env pst_sub.pg pst_sub.uvs) - preamble_sub.goals - (list_as_slprop [] * pst_sub.solved) = pst_sub.goals_inv in - let (| nt, effect_labels |) - : nts:PS.nt_substs & - effect_labels:list T.tot_or_ghost { - PS.well_typed_nt_substs pst_sub.pg pst_sub.uvs nts effect_labels /\ - PS.is_permutation nts pst_sub.ss - } = - match pst_sub.nts with - | Some nts -> nts - | None -> - let r = PS.ss_to_nt_substs pst_sub.pg pst_sub.uvs pst_sub.ss in - match r with - | Inr msg -> - fail_doc pst_sub.pg None [ - Pulse.PP.text "Resulted substitution after intro exists protocol is not well-typed."; - Pulse.PP.text msg; - ] - | Inl nt -> nt in - assert (PS.well_typed_nt_substs pst_sub.pg pst_sub.uvs nt effect_labels); - let pst_sub_goals_inv - : slprop_equiv pst_sub.pg - pst_sub.ss.(preamble_sub.goals) - pst_sub.ss.(list_as_slprop [] * pst_sub.solved) = - PS.slprop_equiv_nt_substs_derived pst_sub.pg pst_sub.uvs pst_sub_goals_inv nt effect_labels in - assume (pst_sub.ss.(list_as_slprop [] * pst_sub.solved) == - tm_emp * pst_sub.ss.(pst_sub.solved)); - let pst_sub_goals_inv - : slprop_equiv pst_sub.pg - pst_sub.ss.(preamble_sub.goals) - (tm_emp * pst_sub.ss.(pst_sub.solved)) = coerce_eq pst_sub_goals_inv () in - let pst_sub_goals_inv - : slprop_equiv pst_sub.pg - pst_sub.ss.(preamble_sub.goals) - pst_sub.ss.(pst_sub.solved) = VE_Trans _ _ _ _ pst_sub_goals_inv (VE_Unit _ _) in - let k_sub - : continuation_elaborator - preamble_sub.g0 (preamble_sub.ctxt * preamble_sub.frame) - pst_sub.pg ((list_as_slprop pst_sub.remaining_ctxt * preamble_sub.frame) * pst_sub.ss.(pst_sub.solved)) = - pst_sub.k in - // replacing pst_sub.ss.(pst_sub.solved) with - // pst_sub.ss.(preamble_sub.goals) using the equiv relation - let k_sub - : continuation_elaborator - preamble_sub.g0 (preamble_sub.ctxt * preamble_sub.frame) - pst_sub.pg ((list_as_slprop pst_sub.remaining_ctxt * preamble_sub.frame) * pst_sub.ss.(preamble_sub.goals)) = - k_elab_equiv k_sub (VE_Refl _ _) (RU.magic ()) in - // substitute preamble_sub.goals - let k_sub - : continuation_elaborator - preamble_sub.g0 (preamble_sub.ctxt * preamble_sub.frame) - pst_sub.pg ((list_as_slprop pst_sub.remaining_ctxt * preamble_sub.frame) * - pst_sub.ss.(open_term_nv body px * (list_as_slprop unsolved'))) = - coerce_eq k_sub () in - assume (pst_sub.ss.(open_term_nv body px * (list_as_slprop unsolved')) == - pst_sub.ss.(open_term_nv body px) * pst_sub.ss.(list_as_slprop unsolved')); - let witness = pst_sub.ss.(null_var x) in - assume (pst_sub.ss.(open_term_nv body px) == subst_term (pst_sub.ss.(body)) [RT.DT 0 witness]); - // rewrite - let k_sub - : continuation_elaborator - preamble_sub.g0 (preamble_sub.ctxt * preamble_sub.frame) - pst_sub.pg ((list_as_slprop pst_sub.remaining_ctxt * preamble_sub.frame) * - (subst_term (pst_sub.ss.(body)) [RT.DT 0 witness] * pst_sub.ss.(list_as_slprop unsolved'))) = - coerce_eq k_sub () in - // some * rearrangement - let k_sub - : continuation_elaborator - preamble_sub.g0 (preamble_sub.ctxt * preamble_sub.frame) - pst_sub.pg ((list_as_slprop pst_sub.remaining_ctxt * - preamble_sub.frame * - pst_sub.ss.(list_as_slprop unsolved')) * - (subst_term (pst_sub.ss.(body)) [RT.DT 0 witness])) = - k_elab_equiv k_sub (VE_Refl _ _) (RU.magic ()) in - - let k_intro_exists - : continuation_elaborator - pst_sub.pg ((list_as_slprop pst_sub.remaining_ctxt * - preamble_sub.frame * - pst_sub.ss.(list_as_slprop unsolved')) * - (subst_term (pst_sub.ss.(body)) [RT.DT 0 witness])) - pst_sub.pg ( _ * - (tm_exists_sl u (PS.nt_subst_binder b nt) pst_sub.ss.(body))) = - k_intro_exists - #pst_sub.pg - #u - #(PS.nt_subst_binder b nt) - #pst_sub.ss.(body) - (RU.magic ()) // typing of tm_exists_sl with pst_sub.ss applied - #witness - (RU.magic ()) // witness typing - #_ - (RU.magic ()) // frame typing - in - assume (tm_exists_sl u (PS.nt_subst_binder b nt) pst_sub.ss.(body) == - pst_sub.ss.(tm_exists_sl u b body)); - // pst_sub.ss extends pst.ss, and pst.ss already solved all of pst.solved - assume (pst.ss.(pst.solved) == pst_sub.ss.(pst.solved)); - // also substitute preamble_sub.frame - let k_intro_exists - : continuation_elaborator - pst_sub.pg ((list_as_slprop pst_sub.remaining_ctxt * - preamble_sub.frame * - pst_sub.ss.(list_as_slprop unsolved')) * - (subst_term (pst_sub.ss.(body)) [RT.DT 0 witness])) - pst_sub.pg ((list_as_slprop pst_sub.remaining_ctxt * - (preamble.frame * pst_sub.ss.(pst.solved)) * - pst_sub.ss.(list_as_slprop unsolved')) * - (pst_sub.ss.(tm_exists_sl u b body))) = coerce_eq k_intro_exists () in - - // rejig some *s in the continuation context - let k_intro_exists - : continuation_elaborator - pst_sub.pg ((list_as_slprop pst_sub.remaining_ctxt * - preamble_sub.frame * - pst_sub.ss.(list_as_slprop unsolved')) * - (subst_term (pst_sub.ss.(body)) [RT.DT 0 witness])) - pst_sub.pg ((list_as_slprop pst_sub.remaining_ctxt * preamble.frame) * - (pst_sub.ss.(pst.solved) * - pst_sub.ss.(tm_exists_sl u b body) * - pst_sub.ss.(list_as_slprop unsolved'))) = - k_elab_equiv k_intro_exists (VE_Refl _ _) (RU.magic ()) in - assume (pst_sub.ss.(pst.solved) * - pst_sub.ss.(tm_exists_sl u b body) * - pst_sub.ss.(list_as_slprop unsolved') == - pst_sub.ss.(pst.solved * tm_exists_sl u b body * list_as_slprop unsolved')); - let k_intro_exists - : continuation_elaborator - pst_sub.pg ((list_as_slprop pst_sub.remaining_ctxt * - preamble_sub.frame * - pst_sub.ss.(list_as_slprop unsolved')) * - (subst_term (pst_sub.ss.(body)) [RT.DT 0 witness])) - pst_sub.pg ((list_as_slprop pst_sub.remaining_ctxt * preamble.frame) * - (pst_sub.ss.(pst.solved * tm_exists_sl u b body * list_as_slprop unsolved'))) = - coerce_eq k_intro_exists () in - let k_sub - : continuation_elaborator - preamble_sub.g0 (preamble_sub.ctxt * preamble_sub.frame) - pst_sub.pg ((list_as_slprop pst_sub.remaining_ctxt * preamble.frame) * - (pst_sub.ss.(pst.solved * tm_exists_sl u b body * list_as_slprop unsolved'))) = - k_elab_trans k_sub k_intro_exists in - // pst.unsolved == tm_exists_sl u b body::unsolved' - let k_sub - : continuation_elaborator - preamble_sub.g0 (preamble_sub.ctxt * preamble_sub.frame) - pst_sub.pg ((list_as_slprop pst_sub.remaining_ctxt * preamble.frame) * - (pst_sub.ss.(pst.solved * list_as_slprop pst.unsolved))) = - k_elab_equiv k_sub (VE_Refl _ _) (RU.magic ()) in - - let k_sub - : continuation_elaborator - pst.pg (list_as_slprop pst.remaining_ctxt * (preamble.frame * pst.ss.(pst.solved))) - pst_sub.pg ((list_as_slprop pst_sub.remaining_ctxt * preamble.frame) * - (pst_sub.ss.(pst.solved * list_as_slprop pst.unsolved))) = - coerce_eq k_sub () in - - // rejig *s in the elab ctxt - let k_sub - : continuation_elaborator - pst.pg ((list_as_slprop pst.remaining_ctxt * preamble.frame) * pst.ss.(pst.solved)) - pst_sub.pg ((list_as_slprop pst_sub.remaining_ctxt * preamble.frame) * - (pst_sub.ss.(pst.solved * list_as_slprop pst.unsolved))) = - k_elab_equiv k_sub (RU.magic ()) (VE_Refl _ _) in - - let k - : continuation_elaborator - preamble.g0 (preamble.ctxt * preamble.frame) - pst_sub.pg ((list_as_slprop pst_sub.remaining_ctxt * preamble.frame) * - (pst_sub.ss.(pst.solved * list_as_slprop pst.unsolved))) = - k_elab_trans pst.k k_sub in - - let goals_inv - : slprop_equiv (push_env pst.pg pst.uvs) - preamble.goals - (list_as_slprop pst.unsolved * pst.solved) = - pst.goals_inv in - // weakening - let goals_inv - : slprop_equiv (push_env pst_sub.pg pst_sub.uvs) - preamble.goals - (pst.solved * list_as_slprop pst.unsolved) = - let d = Metatheory.veq_weakening pst.pg pst.uvs goals_inv pst_sub.pg in - let d = Metatheory.veq_weakening_end pst_sub.pg pst.uvs d pst_sub.uvs in - VE_Trans _ _ _ _ d (VE_Comm _ _ _) in - - let goals_inv - : slprop_equiv pst_sub.pg - (pst_sub.ss.(preamble.goals)) - (pst_sub.ss.(pst.solved * list_as_slprop pst.unsolved)) = - PS.slprop_equiv_nt_substs_derived pst_sub.pg pst_sub.uvs goals_inv nt effect_labels in - - // rewrite k using goals_inv - let k - : continuation_elaborator - preamble.g0 (preamble.ctxt * preamble.frame) - pst_sub.pg ((list_as_slprop pst_sub.remaining_ctxt * preamble.frame) * - (pst_sub.ss.(preamble.goals))) = - k_elab_equiv k (VE_Refl _ _) (RU.magic ()) in - - let pst' : prover_state preamble = { - pg = pst_sub.pg; - remaining_ctxt = pst_sub.remaining_ctxt; - remaining_ctxt_frame_typing = RU.magic (); - uvs = pst_sub.uvs; - ss = pst_sub.ss; - rwr_ss = pst_sub.rwr_ss; - nts = Some (| nt, effect_labels |); - solved = preamble.goals; - unsolved = []; - k; - goals_inv = RU.magic (); - solved_inv = RU.magic (); - progress = false; - allow_ambiguous = pst_sub.allow_ambiguous; - } in - - pst' -#pop-options diff --git a/src/checker/Pulse.Checker.Prover.IntroPure.fst b/src/checker/Pulse.Checker.Prover.IntroPure.fst deleted file mode 100644 index 7ae753b7f..000000000 --- a/src/checker/Pulse.Checker.Prover.IntroPure.fst +++ /dev/null @@ -1,362 +0,0 @@ -(* - Copyright 2023 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module Pulse.Checker.Prover.IntroPure - -open Pulse.Syntax -open Pulse.Typing -open Pulse.Typing.Combinators -open Pulse.Typing.Metatheory -open Pulse.Checker.Pure -open Pulse.Checker.SLPropEquiv -open Pulse.Checker.Prover.Base -open Pulse.Checker.Base -open Pulse.Checker.Prover.Util -open Pulse.PP -module RU = Pulse.RuntimeUtils -module T = FStar.Tactics.V2 -module P = Pulse.Syntax.Printer -module PS = Pulse.Checker.Prover.Substs - -let coerce_eq (#a #b:Type) (x:a) (_:squash (a == b)) : y:b{y == x} = x -#push-options "--z3rlimit_factor 4" -let k_intro_pure (g:env) (p:term) - (d:tot_typing g p tm_prop) - (token:prop_validity g p) (frame:slprop) - - : T.Tac (continuation_elaborator g frame g (frame * tm_pure p)) = - - let t = wtag (Some STT_Ghost) (Tm_IntroPure {p}) in - let c = comp_intro_pure p in - let d : st_typing g t c = T_IntroPure g p d token in - - let x = fresh g in - - // p is well-typed in g, so it does not have x free - assume (open_term p x == p); - - let ppname = mk_ppname_no_range "_pintrop" in - let k : continuation_elaborator - g (frame * tm_emp) - (push_binding g x ppname_default tm_unit) (tm_pure p * frame) = - continuation_elaborator_with_bind frame d (RU.magic ()) (ppname, x) in - - let k : continuation_elaborator - g frame - (push_binding g x ppname_default tm_unit) (frame * tm_pure p) = - k_elab_equiv k (RU.magic ()) (RU.magic ()) in - - fun post_hint r -> - let (| t1, c1, d1 |) = r in - let d1 : st_typing g t1 c1 = d1 in - let empty_env = mk_env (fstar_env g) in - assert (equal g (push_env g empty_env)); - assert (equal (push_env (push_binding g x ppname_default tm_unit) empty_env) - (push_binding g x ppname_default tm_unit)); - let d1 : st_typing (push_binding g x ppname_default tm_unit) t1 c1 = - st_typing_weakening - g - empty_env - t1 c1 d1 - (push_binding g x ppname_default tm_unit) in - - k post_hint (| t1, c1, d1 |) -#pop-options -module R = FStar.Reflection.V2 - -// let is_eq2 (t:R.term) : option (R.term & R.term) = -// let head, args = R.collect_app_ln t in -// match R.inspect_ln head, args with -// | R.Tv_FVar fv, [_; (a1, _); (a2, _)] -// | R.Tv_UInst fv _, [_; (a1, _); (a2, _)] -> -// let l = R.inspect_fv fv in -// if l = ["Pulse"; "Steel"; "Wrapper"; "eq2_prop"] || -// l = ["Prims"; "eq2"] -// then Some (a1, a2) -// else None -// | _ -> None - -let pure_uv_heuristic_t = - uvs:env -> t:term -> - T.Tac (option (uv:var { uv `Set.mem` freevars t } & term)) - -let is_eq2_uvar -: pure_uv_heuristic_t -= fun (uvs:env) (t:term) -> - match is_eq2 t with - | None -> None - | Some (_, l, r) -> - match is_var l with - | Some nm -> - if Set.mem nm.nm_index (dom uvs) - then Some (| nm.nm_index, r |) - else None - | None -> - match is_var r with - | Some nm -> - if Set.mem nm.nm_index (dom uvs) - then Some (| nm.nm_index, l |) - else None - | _ -> None - -module RF = FStar.Reflection.V2.Formula -let is_host_var (x:R.term) = - match R.inspect_ln x with - | R.Tv_Var nv -> - let nv_view = R.inspect_namedv nv in - Some {nm_index=nv_view.uniq; - nm_ppname=mk_ppname (nv_view.ppname) (R.range_of_term x)} - | _ -> None - -let is_uvar_implication -: pure_uv_heuristic_t -= fun (uvs:env) (t:term) -> - debug uvs (fun _ -> Printf.sprintf "is_uvar_implication??: %s\n" (P.term_to_string t)); - match inspect_term t with - | Tm_FStar _ -> ( - let tt = t in - let f = RF.term_as_formula' tt in - match f with - | RF.Implies t0 t1 -> ( - debug uvs (fun _ -> Printf.sprintf "is_uvar_implication, LHS=: %s\n" (T.term_to_string t0)); - match R.inspect_ln t0 with - | R.Tv_Unknown -> None - | _ -> ( - let t0 = wr t0 FStar.Range.range_0 in - match is_eq2 t0 with - | None -> None - | Some (ty, lhs, rhs) -> - if eq_tm ty (wr (`bool) FStar.Range.range_0) - then ( - let try_negated maybe_var other_side - : T.Tac (option (uv:var { uv `Set.mem` freevars t } & term)) - = match is_var lhs with - | None -> None - | Some nm -> - if Set.mem nm.nm_index (dom uvs) - then ( - match inspect_term rhs with - | Tm_FStar _ -> - let rhs = wr (`(not (`#(rhs)))) FStar.Range.range_0 in - assume (nm.nm_index `Set.mem` freevars t); - Some (| nm.nm_index, rhs |) - | _ -> None - ) - else None - in - match try_negated lhs rhs with - | None -> try_negated rhs lhs - | x -> x - ) - else None - ) - ) - | _ -> None - ) - | _ -> None - -let pure_uvar_heursitics : pure_uv_heuristic_t - = let h = [is_eq2_uvar; is_uvar_implication] in - fun (uvs:env) (t:term) -> - let rec loop (h:list pure_uv_heuristic_t) - : T.Tac (option (uv:var { uv `Set.mem` freevars t } & term)) = - match h with - | [] -> None - | h::hs -> - match h uvs t with - | None -> loop hs - | Some (| uv, e |) -> Some (| uv, e |) - in loop h - -let rec try_collect_substs (uvs:env) (t:term) - : T.Tac (ss:PS.ss_t { PS.dom ss `Set.subset` freevars t }) = - assume (PS.dom PS.empty == Set.empty); - match inspect_term t with - | Tm_FStar _ -> - let rt = t in - let f = RF.term_as_formula' rt in - begin - match f with - | RF.And rt0 rt1 -> - let ss0 = try_collect_substs uvs (wr rt0 FStar.Range.range_0) in - let ss1 = try_collect_substs uvs (wr rt1 FStar.Range.range_0) in - if PS.check_disjoint ss0 ss1 - then let r = PS.push_ss ss0 ss1 in - assume (PS.dom r `Set.subset` freevars t); - r - else PS.empty - | _ -> - match pure_uvar_heursitics uvs t with - | Some (| uv, e |) -> - assume (~ (uv `Set.mem` (PS.dom PS.empty))); - let r = PS.push PS.empty uv e in - assume (PS.dom r `Set.subset` freevars t); - r - | None -> PS.empty - end - - | _ -> PS.empty - -#push-options "--z3rlimit_factor 4" -let intro_pure (#preamble:_) (pst:prover_state preamble) - (t:term) - (unsolved':list slprop) - (_:squash (pst.unsolved == (tm_pure t)::unsolved')) - : T.Tac (option (pst':prover_state preamble { pst' `pst_extends` pst })) = - - let t_ss = pst.ss.(t) in - - debug_prover pst.pg (fun _ -> - Printf.sprintf "Intro pure trying to typecheck prop: %s with uvs: %s\n" - (P.term_to_string t_ss) - (env_to_string pst.uvs)); - - - let ss' = try_collect_substs pst.uvs t_ss in - assume (PS.dom pst.ss `Set.disjoint` PS.dom ss'); - let ss_new = PS.push_ss pst.ss ss' in - assume (ss_new `ss_extends` pst.ss); - - let t_ss = ss_new.(t) in - let _ = - let fvs = freevars t_ss in - if check_disjoint pst.uvs fvs - then () - else - fail_doc pst.pg (Some (Pulse.RuntimeUtils.range_of_term t_ss)) - [prefix 2 1 (text "Could not resolve all free variables in the proposition:") - (pp t_ss);] - in - let d = core_check_tot_term pst.pg t_ss tm_prop in - let d_valid = T.with_policy T.ForceSMT (fun () -> check_prop_validity pst.pg t_ss d) in - - // let (| ss_new, t_ss, d, d_valid |) : ss_new:PS.ss_t { ss_new `ss_extends` pst.ss } & - // t_ss:term { t_ss == ss_new.(t) } & - // tot_typing pst.pg t_ss tm_prop & - // prop_validity pst.pg t_ss = - // match is_eq2_uvar pst.pg pst.uvs t_ss with - // | Some (| uv, e |) -> - // // uv is a freevar in t_ss, - // // which is obtained by applying pst.ss to t - // // so uv can't possibly in the domain of pst.ss - // // or it could be a check? - // assume (~ (uv `Set.mem` (PS.dom pst.ss))); - // assume (~ (PS.contains PS.empty uv)); - // let ss_uv = PS.push PS.empty uv e in - // let t_ss_new = ss_uv.(t_ss) in - // assume (Set.disjoint (PS.dom ss_uv) (PS.dom pst.ss)); - // let ss_new = PS.push_ss pst.ss ss_uv in - // assume (t_ss_new == ss_new.(t)); - // // we know this is refl, can we avoid this call? - // let token = check_prop_validity pst.pg t_ss_new (RU.magic ()) in - // (| ss_new, - // t_ss_new, - // RU.magic (), - // token |) - // | None -> - // // - // // we need to check that t is closed in pst.pg - // // this is one way - // // - // let d = core_check_term_with_expected_type pst.pg t_ss tm_prop in - // (| pst.ss, t_ss, E d, check_prop_validity pst.pg t_ss (E d) |) in - - let x = fresh (push_env pst.pg pst.uvs) in - - let solved_new = (tm_pure t) * pst.solved in - let unsolved_new = unsolved' in - - let k : continuation_elaborator - preamble.g0 (preamble.ctxt * preamble.frame) - pst.pg ((list_as_slprop pst.remaining_ctxt * preamble.frame) * ss_new.(solved_new)) = - let frame = (list_as_slprop pst.remaining_ctxt * preamble.frame) * ss_new.(pst.solved) in - let k_pure - : continuation_elaborator - pst.pg frame - pst.pg (frame * (tm_pure t_ss)) = - k_intro_pure _ _ d d_valid frame in - // some *s - let veq - : slprop_equiv pst.pg - (((list_as_slprop pst.remaining_ctxt * preamble.frame) * ss_new.(pst.solved)) * tm_pure t_ss) - ((list_as_slprop pst.remaining_ctxt * preamble.frame) * (tm_pure t_ss * ss_new.(pst.solved))) = - RU.magic () in - - // need lemmas in Prover.Subst - assume (tm_pure ss_new.(t) * ss_new.(pst.solved) == - ss_new.(tm_pure t * pst.solved)); - - let k_pure : continuation_elaborator - pst.pg ((list_as_slprop pst.remaining_ctxt * preamble.frame) * ss_new.(pst.solved)) - pst.pg ((list_as_slprop pst.remaining_ctxt * preamble.frame) * ss_new.(solved_new)) = - - k_elab_equiv k_pure (VE_Refl _ _) veq in - - - let k_pst : continuation_elaborator - preamble.g0 (preamble.ctxt * preamble.frame) - pst.pg ((list_as_slprop pst.remaining_ctxt * preamble.frame) * pst.ss.(pst.solved)) = pst.k in - - assume (pst.ss.(pst.solved) == ss_new.(pst.solved)); - let k_pst : continuation_elaborator - preamble.g0 (preamble.ctxt * preamble.frame) - pst.pg ((list_as_slprop pst.remaining_ctxt * preamble.frame) * ss_new.(pst.solved)) = coerce_eq k_pst () in - - k_elab_trans k_pst k_pure in - - let goals_inv - : slprop_equiv (push_env pst.pg pst.uvs) - preamble.goals - (list_as_slprop ((tm_pure t)::unsolved_new) * pst.solved) = pst.goals_inv in - - let veq : slprop_equiv (push_env pst.pg pst.uvs) - (list_as_slprop ((tm_pure t)::unsolved_new)) - (list_as_slprop unsolved_new * tm_pure t) = RU.magic () in - - let veq : slprop_equiv (push_env pst.pg pst.uvs) - (list_as_slprop ((tm_pure t)::unsolved_new) * pst.solved) - ((list_as_slprop unsolved_new * tm_pure t) * pst.solved) = - VE_Ctxt _ _ _ _ _ veq (VE_Refl _ _) in - - let goals_inv - : slprop_equiv (push_env pst.pg pst.uvs) - preamble.goals - ((list_as_slprop unsolved_new * tm_pure t) * pst.solved) = - VE_Trans _ _ _ _ goals_inv veq in - - let veq : slprop_equiv (push_env pst.pg pst.uvs) - ((list_as_slprop unsolved_new * tm_pure t) * pst.solved) - (list_as_slprop unsolved_new * (tm_pure t * pst.solved)) = - VE_Sym _ _ _ (VE_Assoc _ _ _ _) in - - let goals_inv - : slprop_equiv (push_env pst.pg pst.uvs) - preamble.goals - (list_as_slprop unsolved_new * solved_new) = - VE_Trans _ _ _ _ goals_inv veq in - - assume (freevars ss_new.(solved_new) `Set.subset` dom pst.pg); - - let pst_new : prover_state preamble = { pst with ss = ss_new; - nts = None; - solved = solved_new; - unsolved = unsolved_new; - k; - goals_inv; - solved_inv = () } in - - Some pst_new -#pop-options \ No newline at end of file diff --git a/src/checker/Pulse.Checker.Prover.IntroPure.fsti b/src/checker/Pulse.Checker.Prover.IntroPure.fsti deleted file mode 100644 index 37c2ce2c7..000000000 --- a/src/checker/Pulse.Checker.Prover.IntroPure.fsti +++ /dev/null @@ -1,29 +0,0 @@ -(* - Copyright 2025 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module Pulse.Checker.Prover.IntroPure - -module T = FStar.Tactics - -open Pulse.Syntax -open Pulse.Typing -open Pulse.Checker.Prover.Base - -val intro_pure (#preamble:_) (pst:prover_state preamble) - (t:term) - (unsolved':list slprop) - (_:squash (pst.unsolved == (tm_pure t)::unsolved')) - : T.Tac (option (pst':prover_state preamble { pst' `pst_extends` pst })) diff --git a/src/checker/Pulse.Checker.Prover.IntroWithPure.fst b/src/checker/Pulse.Checker.Prover.IntroWithPure.fst deleted file mode 100644 index f20a436f1..000000000 --- a/src/checker/Pulse.Checker.Prover.IntroWithPure.fst +++ /dev/null @@ -1,225 +0,0 @@ -(* - Copyright 2025 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module Pulse.Checker.Prover.IntroWithPure - -open Pulse.Syntax -open Pulse.Typing -open Pulse.Typing.Combinators -open Pulse.Typing.Metatheory -open Pulse.Checker.Pure -open Pulse.Checker.SLPropEquiv -open Pulse.Checker.Prover.Base -open Pulse.Checker.Base -open Pulse.Checker.Prover.Util -open Pulse.PP -module RU = Pulse.RuntimeUtils -module T = FStar.Tactics.V2 -module P = Pulse.Syntax.Printer -module PS = Pulse.Checker.Prover.Substs - -let coerce_eq (#a #b:Type) (x:a) (_:squash (a == b)) : y:b{y == x} = x - -let open_with_unit_const (v: slprop) : slprop = - open_term' v unit_const 0 - -let intro_with_pure_head = - let intro_pure_explicit_lid = ["Pulse"; "Lib"; "WithPure"; "intro_with_pure"] in - tm_fvar (as_fv intro_pure_explicit_lid) - -let wr t = wr t Range.range_0 - -let mk_intro_with_pure (pred: term) (p:term) - : st_term - = let t = - Tm_ST { - t = T.mk_app (tm_pureapp (tm_pureapp intro_with_pure_head None pred) None - (tm_pureabs ppname_default.name (mk_squash u0 pred) None p Range.range_0)) - [unit_const, T.Q_Explicit]; - args = [] - } - in - wtag (Some STT_Ghost) t - -let intro_with_pure_comp (p:term) (n:ppname) (v:term) = - let st : st_comp = { - u=u_zero; - res=tm_unit; - pre=open_with_unit_const v; - post=tm_with_pure (wr p) n (wr v); - } in - C_STGhost tm_emp_inames st - -let k_intro_with_pure (g:env) (p: term) (n: ppname) (v: slprop) - (d:tot_typing g p tm_prop) - (token:prop_validity g p) (frame:slprop) - - : T.Tac (continuation_elaborator g (frame * open_with_unit_const v) g (frame * tm_pure p)) = - - let t = mk_intro_with_pure p v in - let c = intro_with_pure_comp p n v in - let d : st_typing g t c = RU.magic () in - - let x = fresh g in - - let ppname = n in - let k : continuation_elaborator - g (frame * open_with_unit_const v) - (push_binding g x ppname_default tm_unit) (tm_with_pure p n v * frame) = - admit (); - continuation_elaborator_with_bind frame d (RU.magic ()) (ppname, x) in - - let k : continuation_elaborator - g frame - (push_binding g x ppname_default tm_unit) (frame * tm_pure p) = - k_elab_equiv k (RU.magic ()) (RU.magic ()) in - - fun post_hint r -> - let (| t1, c1, d1 |) = r in - let d1 : st_typing g t1 c1 = d1 in - let empty_env = mk_env (fstar_env g) in - assert (equal g (push_env g empty_env)); - assert (equal (push_env (push_binding g x ppname_default tm_unit) empty_env) - (push_binding g x ppname_default tm_unit)); - let d1 : st_typing (push_binding g x ppname_default tm_unit) t1 c1 = - st_typing_weakening - g - empty_env - t1 c1 d1 - (push_binding g x ppname_default tm_unit) in - - k post_hint (| t1, c1, d1 |) - -module R = FStar.Reflection.V2 - -module RF = FStar.Reflection.V2.Formula - -let intro_with_pure (#preamble:_) (pst:prover_state preamble) - (p: term) (n: ppname) (v: term) - (unsolved':list slprop) - (_:squash (pst.unsolved == (tm_with_pure p n v)::unsolved')) - (prover:prover_t) - : T.Tac (pst':prover_state preamble { pst' `pst_extends` pst /\ - is_terminal pst' }) = - let p_ss = pst.ss.(p) in - - debug_prover pst.pg (fun _ -> - Printf.sprintf "Intro with pure trying to typecheck prop: %s with uvs: %s\n" - (P.term_to_string p_ss) - (env_to_string pst.uvs)); - - let _ = - let fvs = freevars p_ss in - if check_disjoint pst.uvs fvs - then () - else - fail_doc pst.pg (Some (Pulse.RuntimeUtils.range_of_term p_ss)) - [prefix 2 1 (text "Could not resolve all free variables in the proposition:") - (pp p_ss);] - in - - let d = core_check_tot_term pst.pg p_ss tm_prop in - let d_valid = T.with_policy T.ForceSMT (fun () -> check_prop_validity pst.pg p_ss d) in - - admit (); - - let preamble_sub = { - g0 = pst.pg; - ctxt = list_as_slprop pst.remaining_ctxt; - frame = preamble.frame * pst.ss.(pst.solved); - ctxt_frame_typing = RU.magic (); - goals = open_with_unit_const v * (list_as_slprop unsolved'); - } in - - let k_sub: - continuation_elaborator - preamble_sub.g0 (preamble_sub.ctxt * preamble_sub.frame) - pst.pg ((list_as_slprop (slprop_as_list preamble_sub.ctxt) * preamble_sub.frame) * pst.ss.(tm_emp)) = - coerce_eq (k_elab_unit preamble_sub.g0 (preamble_sub.ctxt * preamble_sub.frame)) (RU.magic ()) - in - assume (pst.ss.(tm_emp) == tm_emp); - let pst_sub : prover_state preamble_sub = { - pg = pst.pg; - remaining_ctxt = slprop_as_list preamble_sub.ctxt; - remaining_ctxt_frame_typing = RU.magic (); - uvs = pst.uvs; - ss = pst.ss; - rwr_ss = pst.rwr_ss; - nts = None; - solved = tm_emp; - unsolved = (slprop_as_list (open_with_unit_const v)) @ unsolved'; - k = k_sub; - goals_inv = RU.magic (); - solved_inv = (); - progress = false; - allow_ambiguous = pst.allow_ambiguous; - } in - let pst_sub = prover pst_sub in - let k_sub - : continuation_elaborator - preamble_sub.g0 (preamble_sub.ctxt * preamble_sub.frame) - pst_sub.pg ((list_as_slprop pst_sub.remaining_ctxt * preamble_sub.frame) * pst_sub.ss.(pst_sub.solved)) = - pst_sub.k in - - let k_intro_with_pure - : continuation_elaborator - pst_sub.pg ((list_as_slprop pst_sub.remaining_ctxt * - preamble_sub.frame * - pst_sub.ss.(list_as_slprop unsolved')) * - (pst_sub.ss.(open_with_unit_const v))) - pst_sub.pg ( _ * - (tm_with_pure pst_sub.ss.(p) n pst_sub.ss.(v))) = - k_intro_with_pure - pst_sub.pg - p n v - d - d_valid - (list_as_slprop pst_sub.remaining_ctxt * preamble_sub.frame * pst_sub.ss.(list_as_slprop unsolved')) - in - let k_sub - : continuation_elaborator - preamble_sub.g0 (preamble_sub.ctxt * preamble_sub.frame) - pst_sub.pg ((list_as_slprop pst_sub.remaining_ctxt * preamble.frame) * - (pst_sub.ss.(pst.solved * tm_with_pure p n v * list_as_slprop unsolved'))) = - k_elab_trans k_sub k_intro_with_pure in - - let k - : continuation_elaborator - preamble.g0 (preamble.ctxt * preamble.frame) - pst_sub.pg ((list_as_slprop pst_sub.remaining_ctxt * preamble.frame) * - (pst_sub.ss.(pst.solved * list_as_slprop pst.unsolved))) = - k_elab_trans pst.k k_sub in - - let pst' : prover_state preamble = { - pg = pst_sub.pg; - remaining_ctxt = pst_sub.remaining_ctxt; - remaining_ctxt_frame_typing = RU.magic (); - uvs = pst_sub.uvs; - ss = pst_sub.ss; - rwr_ss = pst_sub.rwr_ss; - nts = None; - solved = preamble.goals; - unsolved = []; - k; - goals_inv = RU.magic (); - solved_inv = RU.magic (); - progress = false; - allow_ambiguous = pst_sub.allow_ambiguous; - } in - - assume pst' `pst_extends` pst; - - pst' \ No newline at end of file diff --git a/src/checker/Pulse.Checker.Prover.IntroWithPure.fsti b/src/checker/Pulse.Checker.Prover.IntroWithPure.fsti deleted file mode 100644 index 309cdf709..000000000 --- a/src/checker/Pulse.Checker.Prover.IntroWithPure.fsti +++ /dev/null @@ -1,30 +0,0 @@ -(* - Copyright 2023 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module Pulse.Checker.Prover.IntroWithPure - -module T = FStar.Tactics - -open Pulse.Syntax -open Pulse.Checker.Prover.Base - -val intro_with_pure (#preamble:_) (pst:prover_state preamble) - p n v - (unsolved':list slprop) - (_:squash (pst.unsolved == tm_with_pure p n v::unsolved')) - (prover:prover_t) - : T.Tac (pst':prover_state preamble { pst' `pst_extends` pst /\ - is_terminal pst' }) diff --git a/src/checker/Pulse.Checker.Prover.Match.Base.fst b/src/checker/Pulse.Checker.Prover.Match.Base.fst deleted file mode 100644 index 29e0f0d94..000000000 --- a/src/checker/Pulse.Checker.Prover.Match.Base.fst +++ /dev/null @@ -1,268 +0,0 @@ -(* - Copyright 2023 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module Pulse.Checker.Prover.Match.Base - -open FStar.Tactics.V2 -open FStar.List.Tot -open FStar.Ghost - -open Pulse.Syntax -open Pulse.Typing -open Pulse.Typing.Combinators -open Pulse.Checker.Base -open Pulse.Checker.Prover.Base -open Pulse.PP - -module RU = Pulse.RuntimeUtils -module T = FStar.Tactics.V2 -module PS = Pulse.Checker.Prover.Substs - -let slprop_list_equiv_refl (g:env) (vs:list slprop) : slprop_list_equiv g vs vs = admit() -let slprop_list_equiv_append (g:env) (vs1 vs2 vs3 vs4 : list slprop) - (d1 : slprop_list_equiv g vs1 vs2) - (d2 : slprop_list_equiv g vs3 vs4) - : slprop_list_equiv g (vs1 @ vs3) (vs2 @ vs4) - = admit() -let slprop_list_equiv_trans (g:env) (vs1 vs2 vs3 : list slprop) - (d1 : slprop_list_equiv g vs1 vs2) - (d2 : slprop_list_equiv g vs2 vs3) - : slprop_list_equiv g vs1 vs3 - = admit() -let slprop_list_equiv_assoc_l (g:env) (vs1 vs2 vs3 : list slprop) - : slprop_list_equiv g (vs1 @ (vs2 @ vs3)) ((vs1 @ vs2) @ vs3) - = admit() -let slprop_list_equiv_assoc_r (g:env) (vs1 vs2 vs3 : list slprop) - : slprop_list_equiv g ((vs1 @ vs2) @ vs3) (vs1 @ (vs2 @ vs3)) - = admit() -let slprop_list_equiv_elim_append (g:env) (vs1 vs2 : list slprop) - : slprop_equiv g (list_as_slprop (vs1 @ vs2)) (list_as_slprop vs1 * list_as_slprop vs2) - = admit() -let slprop_list_equiv_elim_append' (g:env) (vs1 vs2 : list slprop) - : slprop_equiv g (list_as_slprop (vs1 @ vs2)) (list_as_slprop vs2 * list_as_slprop vs1) - = admit() - -(* local aliases *) -let (>>>) #g #t0 #t1 #t2 = VE_Trans g t0 t1 t2 -let (>>*) #g #t0 #t1 #t2 = slprop_list_equiv_trans g t0 t1 t2 -let cong_r #g #t0 #t1 #t2 (d : slprop_equiv g t1 t2) : slprop_equiv g (t0 * t1) (t0 * t2) = - VE_Ctxt _ _ _ _ _ (VE_Refl _ _) d -let cong_l #g #t0 #t1 #t2 (d : slprop_equiv g t1 t2) : slprop_equiv g (t1 * t0) (t2 * t0) = - VE_Ctxt _ _ _ _ _ d (VE_Refl _ _) -let ve_refl_pf (#g:env) (p q : slprop) (s : squash (p == q)) : slprop_equiv g p q = VE_Refl g p - -(* MORE LEMMAS *) - -(* This is not trivial since list_as_slprop distinguishes the singleton case... -maybe it should not? *) -let list_as_slprop_cons_equiv - (g : env) - (p : slprop) - (ps : list slprop) - : slprop_equiv g (list_as_slprop (p::ps)) (p * list_as_slprop ps) -= match ps with - | [] -> - let pf : slprop_equiv g p (p * tm_emp) = - VE_Sym _ _ _ (VE_Unit _ _) >>> VE_Comm _ _ _ - in - pf - | _ -> VE_Refl _ _ - -let slprop_list_equiv_cons g p q ps qs d1 d2 = - list_as_slprop_cons_equiv g p ps >>> - VE_Ctxt _ _ _ _ _ d1 d2 >>> - VE_Sym _ _ _ (list_as_slprop_cons_equiv g q qs) - -let star_flip g p q r = - VE_Assoc _ _ _ _ >>> - VE_Ctxt _ _ _ _ _ (VE_Comm _ _ _) (VE_Refl _ _) >>> - VE_Sym _ _ _ (VE_Assoc _ _ _ _) - -let slprop_list_equiv_flip g p q ps = - VE_Ctxt _ _ _ _ _ (VE_Refl _ _) (list_as_slprop_cons_equiv g q ps) >>> - star_flip _ _ _ _ >>> - VE_Ctxt _ _ _ _ _ (VE_Refl _ _) (VE_Sym _ _ _ <| list_as_slprop_cons_equiv g p ps) - -let slprop_list_equiv_push_append g p ps qs = admit() - -let slprop_list_equiv_append_r g ps qs rs pf = slprop_list_equiv_append _ _ _ _ _ (VE_Refl _ _) pf - -let subst_append ss ps qs = - List.Tot.Properties.map_append - (fun t -> PS.ss_term t ss) - ps qs - -let subst_append_equiv g ss ps qs = - subst_append ss ps qs; - VE_Refl _ _ - -let subst_list_as_slprop = admit() - -(* Is this somewhere else? *) -let subst_star ss p q = admit() -(* Compose two match pass results. *) -let compose_mpr (g:env) (ss : PS.ss_t) (ctxt0 unsolved0 ctxt1 unsolved1 : list slprop) - (mpr1 : match_pass_result g ss ctxt0 unsolved0) - (mpr2 : match_pass_result g mpr1.ss' mpr1.ctxt1 mpr1.unsolved1) - : match_pass_result g ss ctxt0 unsolved0 -= { ss' = mpr2.ss'; - - ctxt_matched = mpr1.ctxt_matched @ mpr2.ctxt_matched; - ctxt1 = mpr2.ctxt1; - ctxt_ok = - (* quite brittle to write these proofs. *) - mpr1.ctxt_ok >>* - (slprop_list_equiv_append _ _ _ _ _ (slprop_list_equiv_refl _ _) mpr2.ctxt_ok) >>* - (slprop_list_equiv_assoc_l _ _ _ _) - ; - - unsolved_matched = mpr1.unsolved_matched @ mpr2.unsolved_matched; - unsolved1 = mpr2.unsolved1; - unsolved_ok = ( - subst_append ss mpr1.unsolved_matched mpr2.unsolved_matched; (* lemma call *) - (* proof: *) - mpr1.unsolved_ok >>* - (slprop_list_equiv_append _ _ _ _ _ (slprop_list_equiv_refl _ _) mpr2.unsolved_ok) >>* - (slprop_list_equiv_assoc_l _ _ _ _) - ); - - match_ok = ( - assume (mpr1.ss' $$ mpr1.unsolved_matched == mpr2.ss' $$ mpr1.unsolved_matched); - (* ^ should hold since mpr1.ss' fully defines the variables in unsolved_matched *) - slprop_list_equiv_append _ _ _ _ _ mpr1.match_ok mpr2.match_ok >>* - // VE_Sym _ _ _ (subst_append_equiv _ _ _ _) >>* - slprop_list_equiv_append _ _ _ _ _ - (ve_refl_pf #g - (list_as_slprop <| (mpr1.ss' $$ mpr1.unsolved_matched)) - (list_as_slprop <| (mpr2.ss' $$ mpr1.unsolved_matched)) - ()) (* relies on assumption above *) - (VE_Refl _ _) - >>* VE_Sym _ _ _ (subst_append_equiv _ _ _ _) - ); - - msgs = mpr1.msgs @ mpr2.msgs; - } - -let apply_mpr_vequiv_proof - (#preamble:_) (pst:prover_state preamble) - (mpr : mpr_t pst) - : slprop_equiv (push_env pst.pg pst.uvs) - ((list_as_slprop pst.remaining_ctxt * preamble.frame) * (mpr.ss'.(pst.solved))) - ((list_as_slprop mpr.ctxt1 * preamble.frame) * (mpr.ss'.(list_as_slprop mpr.unsolved_matched * pst.solved))) -= - let pf0 : slprop_equiv (push_env pst.pg pst.uvs) - (list_as_slprop pst.remaining_ctxt) - (list_as_slprop mpr.ctxt_matched * list_as_slprop mpr.ctxt1) - = mpr.ctxt_ok >>> - slprop_list_equiv_elim_append _ _ _ - in - let pf1 : slprop_equiv (push_env pst.pg pst.uvs) - (list_as_slprop mpr.ctxt_matched * list_as_slprop mpr.ctxt1) - (list_as_slprop mpr.ctxt1 * (mpr.ss'.(list_as_slprop mpr.unsolved_matched))) - = subst_list_as_slprop mpr.ss' mpr.unsolved_matched; - VE_Comm _ _ _ >>> - VE_Ctxt _ _ _ _ _ (VE_Refl _ _) mpr.match_ok - in - let pf2 : slprop_equiv (push_env pst.pg pst.uvs) - ((list_as_slprop pst.remaining_ctxt * preamble.frame) * (mpr.ss'.(pst.solved))) - ((list_as_slprop mpr.ctxt1 * (mpr.ss'.(list_as_slprop mpr.unsolved_matched) * preamble.frame)) * (mpr.ss'.(pst.solved))) - = cong_l (cong_l (pf0 >>> pf1)) >>> - cong_l (VE_Sym _ _ _ (VE_Assoc _ _ _ _)) - in - let eq : squash ((mpr.ss'.(list_as_slprop mpr.unsolved_matched) * mpr.ss'.(pst.solved)) - == - mpr.ss'.(list_as_slprop mpr.unsolved_matched * pst.solved)) - = subst_star mpr.ss' (list_as_slprop mpr.unsolved_matched) pst.solved; - () - in - let pf3 : slprop_equiv (push_env pst.pg pst.uvs) - ((list_as_slprop mpr.ctxt1 * (mpr.ss'.(list_as_slprop mpr.unsolved_matched) * preamble.frame)) * (mpr.ss'.(pst.solved))) - ((list_as_slprop mpr.ctxt1 * preamble.frame) * (mpr.ss'.(list_as_slprop mpr.unsolved_matched * pst.solved))) - = - FStar.Pure.BreakVC.break_vc (); - // ^ this line seems to make the proof much faster - // and more reliable, and work without splitting! - VE_Sym _ _ _ (VE_Assoc _ _ _ _) >>> - cong_r ( - cong_l (VE_Comm _ _ _) >>> - VE_Sym _ _ _ (VE_Assoc _ _ _ _) >>> - cong_r (ve_refl_pf ((mpr.ss'.(list_as_slprop mpr.unsolved_matched)) * mpr.ss'.(pst.solved)) - (mpr.ss'.(list_as_slprop mpr.unsolved_matched * pst.solved)) - eq) - ) >>> - VE_Assoc _ _ _ _ - in - pf2 >>> pf3 - -(* Apply a match pass result to the prover state. *) -let apply_mpr - (#preamble:_) (pst:prover_state preamble) - (mpr : mpr_t pst) - : T.Tac (pst':prover_state preamble { pst' `pst_extends` pst }) -= { pst with - pg = pst.pg; - uvs = pst.uvs; - ss = mpr.ss'; - - nts = None; (* Clear substitution cache *) - - remaining_ctxt = mpr.ctxt1; - remaining_ctxt_frame_typing = RU.magic (); (* inversion? otherwise prove *) - - unsolved = mpr.unsolved1; - solved = list_as_slprop mpr.unsolved_matched * pst.solved; - solved_inv = RU.magic (); (* freevars *) - - (* boring proof *) - goals_inv = - pst.goals_inv >>> - (VE_Ctxt _ _ _ _ _ - (VE_Trans _ _ _ _ mpr.unsolved_ok (slprop_list_equiv_elim_append' _ _ _)) - (VE_Refl _ _)) >>> - (VE_Sym _ _ _ <| VE_Assoc _ _ _ _) - ; - - k = ( - assume (pst.pg == push_env pst.pg pst.uvs); // FIXME!!! Is the environment on the continuation_elaborator right?? - let k0 : continuation_elaborator - preamble.g0 (preamble.ctxt * preamble.frame) - pst.pg ((list_as_slprop pst.remaining_ctxt * preamble.frame) * pst.ss.(pst.solved)) - = pst.k - in - let k1 : continuation_elaborator - pst.pg ((list_as_slprop pst.remaining_ctxt * preamble.frame) * pst.ss.(pst.solved)) - pst.pg ((list_as_slprop pst.remaining_ctxt * preamble.frame) * mpr.ss'.(pst.solved)) - = (admit(); k_elab_unit _ _) - in - let k2 : continuation_elaborator - pst.pg ((list_as_slprop pst.remaining_ctxt * preamble.frame) * mpr.ss'.(pst.solved)) - pst.pg ((list_as_slprop mpr.ctxt1 * preamble.frame) * mpr.ss'.(list_as_slprop mpr.unsolved_matched * pst.solved)) - = k_elab_equiv (k_elab_unit pst.pg ((list_as_slprop pst.remaining_ctxt * preamble.frame) * mpr.ss'.(pst.solved))) - (VE_Refl _ _) - (apply_mpr_vequiv_proof pst mpr) - in - let k3 : continuation_elaborator - preamble.g0 (preamble.ctxt * preamble.frame) - pst.pg ((list_as_slprop mpr.ctxt1 * preamble.frame) * mpr.ss'.(list_as_slprop mpr.unsolved_matched * pst.solved)) - = k0 `k_elab_trans` k1 `k_elab_trans` k2 - in - k3 - ); - - (* Set progress=true if we matched something. *) - progress = pst.progress || Cons? mpr.ctxt_matched; - } - diff --git a/src/checker/Pulse.Checker.Prover.Match.Base.fsti b/src/checker/Pulse.Checker.Prover.Match.Base.fsti deleted file mode 100644 index 687851ac7..000000000 --- a/src/checker/Pulse.Checker.Prover.Match.Base.fsti +++ /dev/null @@ -1,173 +0,0 @@ -(* - Copyright 2023 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module Pulse.Checker.Prover.Match.Base - -open Pulse.Syntax -open Pulse.Typing -open Pulse.Checker.Prover.Base -open Pulse.VC -open FStar.List.Tot - -module T = FStar.Tactics.V2 -module PS = Pulse.Checker.Prover.Substs - -(* Move? *) -let slprop_list_equiv (g:env) (vs1 vs2 : list slprop) : Type = - slprop_equiv g (list_as_slprop vs1) (list_as_slprop vs2) - -let ( $$ ) (ss : PS.ss_t) (ts : list slprop) : list slprop = - ts |> List.Tot.map (fun t -> PS.ss_term t ss) - -(* All of these should be easy. *) -val slprop_list_equiv_refl (g:env) (vs:list slprop) : slprop_list_equiv g vs vs -val slprop_list_equiv_append (g:env) (vs1 vs2 vs3 vs4 : list slprop) - (d1 : slprop_list_equiv g vs1 vs2) - (d2 : slprop_list_equiv g vs3 vs4) - : slprop_list_equiv g (vs1 @ vs3) (vs2 @ vs4) -val slprop_list_equiv_trans (g:env) (vs1 vs2 vs3 : list slprop) - (d1 : slprop_list_equiv g vs1 vs2) - (d2 : slprop_list_equiv g vs2 vs3) - : slprop_list_equiv g vs1 vs3 -val slprop_list_equiv_assoc_l (g:env) (vs1 vs2 vs3 : list slprop) - : slprop_list_equiv g (vs1 @ (vs2 @ vs3)) ((vs1 @ vs2) @ vs3) -val slprop_list_equiv_assoc_r (g:env) (vs1 vs2 vs3 : list slprop) - : slprop_list_equiv g ((vs1 @ vs2) @ vs3) (vs1 @ (vs2 @ vs3)) - -val slprop_list_equiv_elim_append (g:env) (vs1 vs2 : list slprop) - : slprop_equiv g (list_as_slprop (vs1 @ vs2)) (list_as_slprop vs1 * list_as_slprop vs2) -val slprop_list_equiv_elim_append' (g:env) (vs1 vs2 : list slprop) - : slprop_equiv g (list_as_slprop (vs1 @ vs2)) (list_as_slprop vs2 * list_as_slprop vs1) -val slprop_list_equiv_cons (g:env) (p q : slprop) (ps qs : list slprop) - (d1 : slprop_equiv g p q) - (d2 : slprop_list_equiv g ps qs) - : slprop_list_equiv g (p::ps) (q::qs) - -(* The result of a matching pass. Essentially equating part of the context -and goals ("unsolved") and proving it. Essentially: - - ctxt0 = ctxt_matched + ctxt1 - unsolved0 = unsolved_matched + unsolved1 - ctxt_matched = unsolved_matched - - The ss is the substitution that defines some uvars in the goal/unsolved set. - We do not need it for the ctxt slprops, those do not mention them. -*) -noeq -type match_pass_result (g:env) (ss:PS.ss_t) (ctxt0 unsolved0 : list slprop) = { - ss' : (ss' : PS.ss_t {ss' `ss_extends` ss}); - - ctxt_matched : list slprop; - ctxt1 : list slprop; - ctxt_ok : slprop_list_equiv g ctxt0 (ctxt_matched @ ctxt1); - - unsolved_matched : list slprop; - unsolved1 : list slprop; - unsolved_ok : slprop_list_equiv g unsolved0 (unsolved_matched @ unsolved1); - - match_ok : slprop_list_equiv g ctxt_matched (ss' $$ unsolved_matched); - - (* Some information for the user explaining why unsolved1 couldn't be - fully matched. E.g. ambiguity messages. *) - msgs : list (list Pprint.document); -} - -(* A zero for the match pass result, no progress at all. -Copilot wrote this almost correctly in one shot! *) -let mpr_zero (#g #ss #ctxt0 #unsolved0 :_) : match_pass_result g ss ctxt0 unsolved0 = - { ss' = ss; - - ctxt_matched = []; - ctxt1 = ctxt0; - ctxt_ok = slprop_list_equiv_refl _ _; - - unsolved_matched = []; - unsolved1 = unsolved0; - unsolved_ok = slprop_list_equiv_refl _ _; - - match_ok = slprop_list_equiv_refl _ _; - - msgs = []; - } - -(* FIXME: probably do not have to be in this interface, and can -be moved to a separate module anyway. *) -val star_flip (g:env) (p q r : slprop) : slprop_equiv g (p * (q * r)) (q * (p * r)) -val slprop_list_equiv_flip (g:env) (p q : slprop) (ps : list slprop) - : slprop_list_equiv g (p :: q :: ps) (q :: p :: ps) -val slprop_list_equiv_push_append (g:env) (p : slprop) (ps qs : list slprop) - : slprop_list_equiv g (p :: ps @ qs) (ps @ (p :: qs)) -val slprop_list_equiv_append_r (g:env) (ps qs rs : list slprop) - (pf : slprop_list_equiv g qs rs) - : slprop_list_equiv g (ps @ qs) (ps @ rs) -val subst_append (ss : PS.ss_t) (ps qs : list slprop) - : Lemma (ensures ss $$ ps @ qs == (ss $$ ps) @ (ss $$ qs)) -val subst_append_equiv (g:env) (ss : PS.ss_t) (ps qs : list slprop) - : slprop_list_equiv g (ss $$ ps @ qs) ((ss $$ ps) @ (ss $$ qs)) -val subst_list_as_slprop (ss : PS.ss_t) (ps : list slprop) - : Lemma (ensures list_as_slprop (ss $$ ps) == ss.(list_as_slprop ps)) -val subst_star (ss : PS.ss_t) (p q : slprop) - : Lemma (ensures ss.(p * q) == ss.(p) * ss.(q)) - -(* Compose two match pass results. *) - -val compose_mpr (g:env) (ss : PS.ss_t) (ctxt0 unsolved0 ctxt1 unsolved1 : list slprop) - (mpr1 : match_pass_result g ss ctxt0 unsolved0) - (mpr2 : match_pass_result g mpr1.ss' mpr1.ctxt1 mpr1.unsolved1) - : match_pass_result g ss ctxt0 unsolved0 - -(* Just an alias: match_pass_result for a given prover state *) -type mpr_t (#preamble:_) (pst:prover_state preamble) = - match_pass_result (push_env pst.pg pst.uvs) pst.ss pst.remaining_ctxt pst.unsolved - -(* Apply a match pass result to the prover state. *) -val apply_mpr - (#preamble:_) (pst:prover_state preamble) - (mpr : mpr_t pst) - : T.Tac (pst':prover_state preamble { pst' `pst_extends` pst }) - -(******************************************************************) -(******************************************************************) -(******************************************************************) - -(* The result of a matcher. *) - -noeq -type match_res_t - (#preamble:_) (pst : prover_state preamble) - (p q : slprop) : Type = - | NoMatch : reason:string -> match_res_t pst p q - | Matched : - ss_ext : PS.ss_t -> - vc : vc_t -> - with_vc vc (slprop_equiv pst.pg p ss_ext.(q)) -> - match_res_t pst p q - (* Note: I would prefer to just write - Matched of guarded (ss:PS.ss_t & slprop_equiv pst.pg p ss.(q)) - but that seems to incur many more inference failures. *) - -(* A matcher can also raise this to signal a graceful failure, it's just -converted to NoMatch. *) -exception ENoMatch of string - -(* The type of a 1-to-1 matcher. The pst is "read-only". If there's -no match, it returns NoMatch or raises ENoMatch. Other exceptions are -not caught. These matchers never try to discharge VCs, but instead return -guarded results (see match_res_t). *) -type matcher_t = - (#preamble:_) -> (pst : prover_state preamble) -> - (p : slprop) -> (q : slprop) -> - T.Tac (match_res_t pst p q) diff --git a/src/checker/Pulse.Checker.Prover.Match.Comb.fst b/src/checker/Pulse.Checker.Prover.Match.Comb.fst deleted file mode 100644 index 33379971b..000000000 --- a/src/checker/Pulse.Checker.Prover.Match.Comb.fst +++ /dev/null @@ -1,337 +0,0 @@ -(* - Copyright 2023 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module Pulse.Checker.Prover.Match.Comb - -open FStar.Tactics.V2 -open FStar.List.Tot -open FStar.Ghost - -open Pulse.Typing.Combinators -open Pulse.Syntax -open Pulse.Typing -open Pulse.PP -open Pulse.Show -open Pulse.VC - -module RU = Pulse.RuntimeUtils -module PS = Pulse.Checker.Prover.Substs - - -(* Ambig (q, p1, p2): q (in goals) can be matched by p1 or p2 (in ctx). -This is internal to this module, which is in charge of calling the matchers -1 by 1 and detecting ambiguity. *) -exception Ambig of (slprop & slprop & slprop) - -(* local aliases *) -let (>>>) #g #t0 #t1 #t2 = VE_Trans g t0 t1 t2 -let (>>*) #g #t0 #t1 #t2 = slprop_list_equiv_trans g t0 t1 t2 -let cong_r #g #t0 #t1 #t2 (d : slprop_equiv g t1 t2) : slprop_equiv g (t0 * t1) (t0 * t2) = - VE_Ctxt _ _ _ _ _ (VE_Refl _ _) d -let cong_l #g #t0 #t1 #t2 (d : slprop_equiv g t1 t2) : slprop_equiv g (t1 * t0) (t2 * t0) = - VE_Ctxt _ _ _ _ _ d (VE_Refl _ _) -let ve_refl_pf (#g:env) (p q : slprop) (s : squash (p == q)) : slprop_equiv g p q = VE_Refl g p - -let wrap_matcher - (label : string) - (matcher : matcher_t) - (#preamble:_) (pst : prover_state preamble) - (p q : slprop) - : T.Tac (match_res_t pst p q) -= - if RU.debug_at_level (fstar_env pst.pg) "prover.match" then - info_doc pst.pg (Some <| range_of_env pst.pg) [ - text "Trying to match"; - prefix 2 1 (doc_of_string "p =") (pp p); - prefix 2 1 (doc_of_string "q =") (pp pst.ss.(q)); - text <| "with matcher: " ^ label; - ]; - let res = try matcher pst p q with | ENoMatch s -> NoMatch s | e -> raise e in - if RU.debug_at_level (fstar_env pst.pg) "prover.match" then begin - info_doc pst.pg (Some <| range_of_env pst.pg) [ - text "Result:" ^/^ - (match res with - | NoMatch s -> text "No match: " ^/^ text s - | Matched _ _ _ -> text "Matched (with some guard)") - ] - end; - res - -(* The type of a match of q from something in a context ps. *) noeq -type match_from_context_t - (#preamble:_) (pst : prover_state preamble) - (q : slprop) - (ps : list slprop) -= { - p : slprop; - rest : list slprop; - rest_ok : slprop_list_equiv pst.pg ps (p::rest); - ss' : PS.ss_t; - ss_extends : squash (ss' `ss_extends` pst.ss); - proof : guarded (slprop_equiv pst.pg p ss'.(q)); -} - -(* If we matched q from ps, we can match it from any extension of ps. *) -let frame_match_from_context - (#preamble:_) (pst : prover_state preamble) - (q : slprop) - (ps : list slprop) - (mm : match_from_context_t pst q ps) - (frame : list slprop) - : match_from_context_t pst q (frame @ ps) -= { - p = mm.p; - rest = frame @ mm.rest; - rest_ok = slprop_list_equiv_append _ _ _ _ _ (VE_Refl _ _) mm.rest_ok >>> - VE_Sym _ _ _ (slprop_list_equiv_push_append _ _ _ _); - ss' = mm.ss'; - ss_extends = mm.ss_extends; - proof = mm.proof; -} - -#push-options "--z3rlimit_factor 4" -(* The result of a match pass, trying to match all of ctxt to all of unsolved. *) -(* Returns all successful matches of q in ps. *) -let rec get_all_matches_aux - (label : string) - (matcher : matcher_t) - (#preamble:_) (pst : prover_state preamble) - (q : slprop) - (ps0 ps : list slprop) - : T.Tac (list (match_from_context_t pst q (List.rev ps0 @ ps))) -= - match ps with - | [] -> [] - | p::ps' -> - let thisone : list (match_from_context_t pst q (List.rev ps0 @ ps)) = - match wrap_matcher label matcher pst p pst.ss.(q) with - | NoMatch s -> [] - | Matched ss_ext vc ff -> - assume (Set.disjoint (PS.dom pst.ss) (PS.dom ss_ext)); - let ss' = PS.push_ss pst.ss ss_ext in - assume (ss'.(q) == ss_ext.(pst.ss.(q))); (* this should be true since it's just composing substs? *) - - if RU.debug_at_level (fstar_env pst.pg) "prover.match" then - info_doc pst.pg (Some <| range_of_env pst.pg) [ - text ("Matched with " ^ label); - prefix 2 1 (doc_of_string "p =") (pp p); - prefix 2 1 (doc_of_string "q =") (pp pst.ss.(q)); - prefix 2 1 (doc_of_string "ss' =") (pp ss'); - prefix 2 1 (doc_of_string "ss'.(q) =") (pp ss'.(q)); - ]; - - let ff : with_vc vc (slprop_equiv pst.pg p ss_ext.(pst.ss.(q))) = ff in - let ff : with_vc vc (slprop_equiv pst.pg p ss'.(q)) = coerce_eq () ff in - (* ^ FIXME: why is the coerce needed? *) - - let mm : match_from_context_t pst q ps = { - p = p; - rest = ps'; - rest_ok = slprop_list_equiv_refl _ _; - ss' = ss'; - ss_extends = (); - proof = Guarded vc ff; - } - in - let mm = frame_match_from_context pst q ps mm (List.rev ps0) in - [mm] - in - let rest : list (match_from_context_t pst q (List.rev ps0 @ ps)) = - assume (List.rev (p::ps0) @ ps' == List.rev ps0 @ (p::ps')); // boring - get_all_matches_aux label matcher pst q (p::ps0) ps' - in - thisone @ rest -#pop-options - -let get_all_matches - (label : string) - (matcher : matcher_t) - (#preamble:_) (pst : prover_state preamble) - (q : slprop) - (ps : list slprop) - : T.Tac (list (match_from_context_t pst q ps)) -= get_all_matches_aux label matcher pst q [] ps - -let rec coallesce_equal_matches - (#preamble:_) (#pst : prover_state preamble) - (#q : slprop) - (#ps : list slprop) - (ms : list (match_from_context_t pst q ps)) - : list (match_from_context_t pst q ps) -= - match ms with - | [] -> [] - | m1::ms' -> - let ms' = coallesce_equal_matches ms' in - match ms' with - | [] -> [m1] - | m2::ms'' -> - if FStar.Reflection.TermEq.term_eq m1.p m2.p then - m1::ms'' - else - m1::m2::ms'' - -(* Returns the slprop that was matched, the remaining context, the - needed substitution and a **guarded** proof of the match. This basically - turns the list returned by get_all_matches into a single match, failing - if any ambiguity is detected. - - In the error case (Inl) it can add some messages for the user, for example - explaning why ambiguity was detected. -*) -let match_f_1n - (label : string) - (matcher : matcher_t) - (#preamble:_) (pst : prover_state preamble) - (q : slprop) - (ps : list slprop) - : T.Tac (either (list document) (match_from_context_t pst q ps)) -= let ms = get_all_matches label matcher pst q ps in - let ms = coallesce_equal_matches ms in - match ms with - | [] -> Inl [] (* no matches *) - | [m] -> - if RU.debug_at_level (fstar_env pst.pg) "prover.match" then - info_doc pst.pg (Some <| range_of_env pst.pg) [ - text ("Successful unambiguous match pass with " ^ label); - prefix 2 1 (doc_of_string "q =") (pp q); - prefix 2 1 (doc_of_string "m.p =") (pp m.p); - prefix 2 1 (doc_of_string "m.ss' =") (pp m.ss'); - prefix 2 1 (doc_of_string "m.ss'.(q) =") (pp m.ss'.(q)); - ]; - Inr m (* yay! *) - | m1::m2::_ -> - if pst.allow_ambiguous - then ( - if RU.debug_at_level (fstar_env pst.pg) "prover.match" then - info_doc pst.pg (Some <| range_of_env pst.pg) [ - text ("Successful AMBIGUOUS Match pass with " ^ label); - prefix 2 1 (doc_of_string "q =") (pp q); - prefix 2 1 (doc_of_string "m.p =") (pp m1.p); - prefix 2 1 (doc_of_string "m.ss' =") (pp m1.ss'); - prefix 2 1 (doc_of_string "m.ss'.(q) =") (pp m1.ss'.(q)); - ]; - Inr m1 - ) else - Inl [ - text "Ambiguous match for resource:" ^^ - indent (pp q); - text "It can be matched by both:" ^^ - indent (pp m1.p) ^^ hardline ^^ - text "and:" ^^ - indent (pp m2.p) ^^ hardline ^^ - text "in the context."; - ] - -let weaken_slprop_equiv_env - (#g1 : env) - (#g2 : env{g2 `env_extends` g1}) - (#p #q : slprop) - (pf : slprop_equiv g1 p q) : slprop_equiv g2 p q = magic() - -(* Tries to match all of ctxt to all of unsolved, and returns the mpr *) -let rec match_f_nn - (label : string) - (matcher_f : matcher_t) - (#preamble:_) (pst : prover_state preamble) - (ctxt:list slprop) (unsolved:list slprop) - : T.Tac (match_pass_result (push_env pst.pg pst.uvs) pst.ss ctxt unsolved) -= match unsolved with - | [] -> - (* Done *) - mpr_zero - | q::qs -> - (* Try to match the tail. We get back some possibly reduced ctxt. *) - let mpr = match_f_nn label matcher_f pst ctxt qs in - (* Now, try to match q, in the reduced context. *) - let pst' = { pst with - ss = mpr.ss'; - nts = None; - k = coerce_eq (magic()) pst.k; - solved_inv = magic(); - } in - if RU.debug_at_level (fstar_env pst.pg) "prover.match" then - T.print ("Trying to match goal " ^ show q ^ " from context (" ^ label ^ ")"); - - match match_f_1n label matcher_f pst' q mpr.ctxt1 with - | Inl docs -> - (* If that fails we're done, just adding q as unsolved. *) - let mpr' = { mpr with - unsolved_matched = mpr.unsolved_matched; - unsolved1 = q::mpr.unsolved1; - unsolved_ok = - slprop_list_equiv_cons _ q q _ _ (VE_Refl _ _) mpr.unsolved_ok >>> - slprop_list_equiv_push_append _ _ _ _; - - msgs = docs :: mpr.msgs; - } - in mpr' - - | Inr { p; rest; rest_ok; ss'; ss_extends; proof } -> - (* conditionally matched, try to discharge *) - - match unguard proof with - | Inl iss -> - fail_doc_with_subissues pst.pg (Some <| range_of_env pst.pg) iss [ - text "Failed to discharge match guard for goal:" ^^ - indent (pp q) ^^ hardline ^^ - text "with resource from context:" ^^ - indent (pp p); - ] - - | Inr pq -> - (* Got another match, extend the mpr. FIXME: We could accumulate all guards - instead of discharging here. *) - let unsolved_ok : slprop_list_equiv (push_env pst.pg pst.uvs) qs (mpr.unsolved_matched @ mpr.unsolved1) = - mpr.unsolved_ok - in - let ctxt_ok () : slprop_list_equiv (push_env pst.pg pst.uvs) ctxt ((p :: mpr.ctxt_matched) @ rest) = - mpr.ctxt_ok >>> - weaken_slprop_equiv_env (slprop_list_equiv_append_r _ _ _ _ rest_ok) >>> - VE_Sym _ _ _ (slprop_list_equiv_push_append _ p mpr.ctxt_matched rest) - in - let unsolved_ok : slprop_list_equiv (push_env pst.pg pst.uvs) (q::qs) ((q :: mpr.unsolved_matched) @ mpr.unsolved1) = - slprop_list_equiv_cons _ q q _ _ (VE_Refl _ _) unsolved_ok - in - let mpr1 : match_pass_result (push_env pst.pg pst.uvs) pst.ss ctxt unsolved = { - ss' = ss'; - - ctxt_matched = p :: mpr.ctxt_matched; - ctxt1 = rest; - ctxt_ok = ctxt_ok (); - - unsolved_matched = q :: mpr.unsolved_matched; - unsolved1 = coerce_eq () mpr.unsolved1; - unsolved_ok = unsolved_ok; - - match_ok = ( - slprop_list_equiv_cons _ _ _ _ _ (weaken_slprop_equiv_env pq) mpr.match_ok >>* - ve_refl_pf _ _ (admit()) - ); - - msgs = mpr.msgs; - } - in - mpr1 - -let match_with - (label : string) - (matcher : matcher_t) - (#preamble:_) (pst:prover_state preamble) - : T.Tac (list (list document) & pst':prover_state preamble { pst' `pst_extends` pst }) -= - let mpr : mpr_t pst = match_f_nn label matcher pst pst.remaining_ctxt pst.unsolved in - mpr.msgs, apply_mpr pst mpr diff --git a/src/checker/Pulse.Checker.Prover.Match.Comb.fsti b/src/checker/Pulse.Checker.Prover.Match.Comb.fsti deleted file mode 100644 index 0d1f2167c..000000000 --- a/src/checker/Pulse.Checker.Prover.Match.Comb.fsti +++ /dev/null @@ -1,37 +0,0 @@ -(* - Copyright 2023 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module Pulse.Checker.Prover.Match.Comb - -module T = FStar.Tactics - -open FStar.Pprint - -open Pulse.Syntax -open Pulse.Typing - -open Pulse.Checker.Base -open Pulse.Checker.Prover.Base -open Pulse.Checker.Prover.Match.Base - -(* Combinators for matching passes (currently a single one) *) - -(* Do a pass over all unsolved goals to see if any can be matched with the given matcher. *) -val match_with - (label : string) - (matcher : matcher_t) - (#preamble:_) (pst:prover_state preamble) - : T.Tac (list (list document) & pst':prover_state preamble { pst' `pst_extends` pst }) diff --git a/src/checker/Pulse.Checker.Prover.Match.MKeys.fst b/src/checker/Pulse.Checker.Prover.Match.MKeys.fst index 45c78f738..1a4d3f723 100644 --- a/src/checker/Pulse.Checker.Prover.Match.MKeys.fst +++ b/src/checker/Pulse.Checker.Prover.Match.MKeys.fst @@ -79,14 +79,22 @@ let rec zip3 (l1:list 'a) (l2:list 'b) (l3:list 'c) : T.Tac (list ('a & 'b & 'c) | _, _, _ -> T.fail "zip3: length mismatch" -let same_head (t0 t1:term) - : T.Tac bool - = match T.hua t0, T.hua t1 with - | Some (h0, us0, args0), Some (h1, us1, args1) -> - T.inspect_fv h0 = T.inspect_fv h1 && - L.length args0 = L.length args1 - | _ -> - true // conservative +let get_mkeys (g:env) (p: slprop) : T.Tac (option (list term)) = + match T.hua p with + | None -> None + | Some (hfv, _, args) -> + if Pulse.Reflection.Util.fv_has_attr_string "Pulse.Lib.Core.no_mkeys" hfv then Some [] else + match type_of_fv g hfv with | None -> None | Some t -> + let bs, _ = R.collect_arr_ln_bs t in + if L.length bs <> L.length args then None else + let bs_args = T.zip bs args in + let mkeys = T.filter_map (fun (b, (a, _)) -> + if binder_is_mkey b then + Some a + else + None + ) bs_args in + if Nil? mkeys then None else Some mkeys exception GFalse of string exception GTrue of string @@ -184,17 +192,3 @@ let rec eligible_for_smt_equality (g:env) (t0 t1 : term) // ]; true | e -> raise e - -let same_head_strict (t0 t1:term) - : T.Tac bool - = match T.hua t0, T.hua t1 with - | Some (h0, us0, args0), Some (h1, us1, args1) -> - T.inspect_fv h0 = T.inspect_fv h1 && - L.length args0 = L.length args1 - | _ -> - false - -let mkey_mismatch (g:env) (t0 t1 : term) : T.Tac bool = - if same_head_strict t0 t1 - then not (eligible_for_smt_equality g t0 t1) - else false diff --git a/src/checker/Pulse.Checker.Prover.Match.MKeys.fsti b/src/checker/Pulse.Checker.Prover.Match.MKeys.fsti index 90d3bacc3..0f32f55c9 100644 --- a/src/checker/Pulse.Checker.Prover.Match.MKeys.fsti +++ b/src/checker/Pulse.Checker.Prover.Match.MKeys.fsti @@ -18,11 +18,8 @@ module Pulse.Checker.Prover.Match.MKeys open FStar.Tactics.V2 open Pulse.Syntax -open Pulse.Typing +open Pulse.Typing.Env -(* remove!*) -val same_head (t0 t1 : term): Tac bool +val get_mkeys (g:env) (p: slprop) : Tac (option (list term)) val eligible_for_smt_equality (g:env) (t0 t1 : term) : Tac bool - -val mkey_mismatch (g:env) (t0 t1 : term) : Tac bool diff --git a/src/checker/Pulse.Checker.Prover.Match.Matchers.fst b/src/checker/Pulse.Checker.Prover.Match.Matchers.fst deleted file mode 100644 index 22468d903..000000000 --- a/src/checker/Pulse.Checker.Prover.Match.Matchers.fst +++ /dev/null @@ -1,205 +0,0 @@ -(* - Copyright 2023 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module Pulse.Checker.Prover.Match.Matchers - -open FStar.Tactics.V2 -open FStar.List.Tot -open FStar.Ghost - -open Pulse.Syntax -open Pulse.Typing -open Pulse.Typing.Combinators -open Pulse.Checker.Base -open Pulse.Checker.Prover.Match.Base -open Pulse.PP -open Pulse.Show -open Pulse.Checker.Prover.Base - -module T = FStar.Tactics.V2 -module R = FStar.Reflection.V2 -module L = FStar.List.Tot -module TermEq = FStar.Reflection.TermEq -module RT = FStar.Reflection.Typing -module MKeys = Pulse.Checker.Prover.Match.MKeys - -module RU = Pulse.RuntimeUtils -module PS = Pulse.Checker.Prover.Substs -module PTU = Pulse.Typing.Util - -(* local aliases *) -let (>>>) #g #t0 #t1 #t2 = VE_Trans g t0 t1 t2 -let (>>*) #g #t0 #t1 #t2 = slprop_list_equiv_trans g t0 t1 t2 -let cong_r #g #t0 #t1 #t2 (d : slprop_equiv g t1 t2) : slprop_equiv g (t0 * t1) (t0 * t2) = - VE_Ctxt _ _ _ _ _ (VE_Refl _ _) d -let cong_l #g #t0 #t1 #t2 (d : slprop_equiv g t1 t2) : slprop_equiv g (t1 * t0) (t2 * t0) = - VE_Ctxt _ _ _ _ _ d (VE_Refl _ _) -let ve_refl_pf (#g:env) (p q : slprop) (s : squash (p == q)) : slprop_equiv g p q = VE_Refl g p - -// -// Call into the F* unifier to solve for uvs by unifying p and q -// -let try_solve_uvars (g:env) (uvs:env { disjoint uvs g }) (p q:term) - : T.Tac (ss:PS.ss_t { PS.dom ss `Set.subset` freevars q }) = - - let uvs = uvs - |> bindings_with_ppname - |> L.rev - |> L.map (fun (({name}, x, t):(ppname & _ & _)) -> - let nv_view = { - R.uniq = x; - R.sort = t; - R.ppname = name; - } in - let nv = R.pack_namedv nv_view in - nv, t - ) in - - let l, issues = - RU.with_context (get_context g) (fun _ -> - T.with_policy T.Force (fun () -> - T.try_unify (elab_env g) uvs p q)) - in - - T.log_issues issues; - - // build ss - let ss = PS.empty in - assume (PS.dom ss `Set.subset` freevars q); - match l with - | None -> ss - | Some l -> - let q_names = freevars q in - L.fold_left (fun (ss:(ss:PS.ss_t { PS.dom ss `Set.subset` freevars q })) (x, t) -> - let nv_view = R.inspect_namedv x in - if Set.mem nv_view.uniq q_names && - not (Set.mem nv_view.uniq (PS.dom ss)) - then begin - let ss_new = PS.push ss nv_view.uniq t in - assert (nv_view.uniq `Set.mem` freevars q); - assert (PS.dom ss `Set.subset` freevars q); - assume (PS.dom ss_new `Set.subset` freevars q); - ss_new - end - else ss - ) ss l - -let rec unascribe (g:env) (t:term) : (t':term & erased (RT.equiv (elab_env g) t t')) = - match R.inspect_ln t with - | R.Tv_AscribedT t' _ _ _ - | R.Tv_AscribedC t' _ _ _ -> - let tok : erased (RT.equiv (elab_env g) t t') = magic() in // Rel_ascribed? - let (| t'', tok' |) = unascribe g t' in - (| t'', hide <| RT.Rel_trans _ _ _ _ _ tok tok' |) - | _ -> (| t, hide <| RT.Rel_refl _ _ _ |) - -(* Syntactic equality check, but we also unascribe at the top level. Returns -an equality proof. *) -let eq_tm_unascribe (g:env) (p q:term) : option (erased (RT.equiv (elab_env g) p q)) = - let (| up, ptok |) = unascribe g p in - let (| uq, qtok |) = unascribe g q in - if eq_tm up uq - then ( - let t1 : erased (RT.equiv (elab_env g) p q) = - let trans #t0 #t1 #t2 - (d1 : (RT.equiv (elab_env g) t0 t1)) - (d2 : (RT.equiv (elab_env g) t1 t2)) - : (RT.equiv (elab_env g) t0 t2) - = RT.Rel_trans (elab_env g) _ _ _ RT.R_Eq d1 d2 - in - ptok `trans` RT.Rel_refl _ _ _ `trans` RT.Rel_sym _ _ _ qtok - in - Some t1 - ) - else None - -(**************** The actual matchers *) - -open Pulse.VC - -(* The syntactic matcher *) -let match_syntactic_11 - (#preamble:_) (pst : prover_state preamble) - (p q : slprop) - : T.Tac (match_res_t pst p q) -= (* term_eq gives us provable equality between p and q, so we can use VE_Refl. *) - if TermEq.term_eq p q - then Matched PS.empty Trivial (fun () -> VE_Refl _ _) - else NoMatch "not term_eq" - -(* Fast unification / strict matcher *) -let match_fastunif_11 - (#preamble:_) (pst : prover_state preamble) - (p q : slprop) - : T.Tac (match_res_t pst p q) -= match PTU.check_equiv_now_nosmt (elab_env pst.pg) p q with - | Some tok, _ -> - Matched PS.empty Trivial (fun () -> VE_Ext _ _ _ (RT.Rel_eq_token _ _ _ ())) - | None, _ -> - NoMatch "no unif" - -let match_fastunif_inst_11 - (#preamble:_) (pst : prover_state preamble) - (p q : slprop) - : T.Tac (match_res_t pst p q) -= let g = pst.pg in - let q0 = q in - - (* If the heads of p and q differ, skip. This makes sure we don't unfold - stuff implicitly. *) - if not <| MKeys.same_head p q then - raise (ENoMatch "head mismatch"); - - (* Try to instantiate q's uvars by matching it to p. We do not trust - this call so we then typecheck the result (and normalize it too). *) - let ss' = try_solve_uvars pst.pg pst.uvs p q in - let q_subst = ss'.(q) in - - match PTU.check_equiv_now_nosmt (elab_env pst.pg) p q_subst with - | None, issues -> - if RU.debug_at_level (fstar_env g) "ggg" then - info_doc_with_subissues g (Some <| range_of_env g) issues [ - text "match_fastunif_inst_11: check_equiv failed, no unif"; - ]; - raise (ENoMatch "no unif") - - | Some token, _ -> - Matched ss' Trivial (fun _ -> VE_Ext _ _ _ (RT.Rel_eq_token _ _ _ ())) - -let match_full_11 - (#preamble:_) (pst : prover_state preamble) - (p q : slprop) - : T.Tac (match_res_t pst p q) -= let g = pst.pg in - let q0 = q in - - (* Similar to fastunif_inst. Instead of unifiying without SMT, - we check if the mkeys match. If so, we return a guarded result. *) - - if not <| MKeys.same_head p q then - raise (ENoMatch "head mismatch"); - - let ss' = try_solve_uvars pst.pg pst.uvs p q in - let q_subst = ss'.(q) in - - (* Check now, after normalizing etc, that we are allowed to try an SMT query - to match them. This is the part that looks at the binder attributes for strictness. - If this check doesn't pass, skip. *) - if not (MKeys.eligible_for_smt_equality g p q_subst) then - raise (ENoMatch "not eligible for SMT"); - - (* Return a guarded result *) - Matched ss' (Equiv g p q_subst) (fun equiv -> VE_Ext _ _ _ equiv) diff --git a/src/checker/Pulse.Checker.Prover.Match.Matchers.fsti b/src/checker/Pulse.Checker.Prover.Match.Matchers.fsti deleted file mode 100644 index 88953fccf..000000000 --- a/src/checker/Pulse.Checker.Prover.Match.Matchers.fsti +++ /dev/null @@ -1,26 +0,0 @@ -(* - Copyright 2023 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module Pulse.Checker.Prover.Match.Matchers - -open Pulse.Checker.Prover.Match.Base - -(* The actual matchers, in 1-to-1 form. Usually called with Match.Comb.match_with *) - -val match_syntactic_11 : matcher_t -val match_fastunif_11 : matcher_t -val match_fastunif_inst_11 : matcher_t -val match_full_11 : matcher_t diff --git a/src/checker/Pulse.Checker.Prover.Match.fst b/src/checker/Pulse.Checker.Prover.Match.fst deleted file mode 100644 index 2bac89ca4..000000000 --- a/src/checker/Pulse.Checker.Prover.Match.fst +++ /dev/null @@ -1,25 +0,0 @@ -(* - Copyright 2023 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module Pulse.Checker.Prover.Match - -open Pulse.Checker.Prover.Match.Comb -open Pulse.Checker.Prover.Match.Matchers - -let match_syntactic = match_with "SYNTACTIC" match_syntactic_11 -let match_fastunif = match_with "FASTUNIF" match_fastunif_11 -let match_fastunif_i = match_with "FASTUNIF_INST" match_fastunif_inst_11 -let match_full = match_with "FULL" match_full_11 diff --git a/src/checker/Pulse.Checker.Prover.Match.fsti b/src/checker/Pulse.Checker.Prover.Match.fsti deleted file mode 100644 index da59d9931..000000000 --- a/src/checker/Pulse.Checker.Prover.Match.fsti +++ /dev/null @@ -1,36 +0,0 @@ -(* - Copyright 2023 Microsoft Research - - Licensed under the Apache License, Version 2.0 (the "License"); - you may not use this file except in compliance with the License. - You may obtain a copy of the License at - - http://www.apache.org/licenses/LICENSE-2.0 - - Unless required by applicable law or agreed to in writing, software - distributed under the License is distributed on an "AS IS" BASIS, - WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. - See the License for the specific language governing permissions and - limitations under the License. -*) - -module Pulse.Checker.Prover.Match - -module T = FStar.Tactics -open FStar.Pprint - -open Pulse.Checker.Prover.Base - -(* Full matching passes. *) - -val match_syntactic (#preamble:preamble) (pst:prover_state preamble) -: T.Tac (list (list document) & pst':prover_state preamble { pst' `pst_extends` pst }) - -val match_fastunif (#preamble:preamble) (pst:prover_state preamble) -: T.Tac (list (list document) & pst':prover_state preamble { pst' `pst_extends` pst }) - -val match_fastunif_i (#preamble:preamble) (pst:prover_state preamble) -: T.Tac (list (list document) & pst':prover_state preamble { pst' `pst_extends` pst }) - -val match_full (#preamble:preamble) (pst:prover_state preamble) -: T.Tac (list (list document) & pst':prover_state preamble { pst' `pst_extends` pst }) diff --git a/src/checker/Pulse.Checker.Prover.Normalize.fst b/src/checker/Pulse.Checker.Prover.Normalize.fst new file mode 100644 index 000000000..ed08e5109 --- /dev/null +++ b/src/checker/Pulse.Checker.Prover.Normalize.fst @@ -0,0 +1,77 @@ +(* + Copyright 2023 Microsoft Research + + Licensed under the Apache License, Version 2.0 (the "License"); + you may not use this file except in compliance with the License. + You may obtain a copy of the License at + + http://www.apache.org/licenses/LICENSE-2.0 + + Unless required by applicable law or agreed to in writing, software + distributed under the License is distributed on an "AS IS" BASIS, + WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. + See the License for the specific language governing permissions and + limitations under the License. +*) + +module Pulse.Checker.Prover.Normalize + +module T = FStar.Tactics.V2 + +open Pulse.Syntax +open Pulse.Typing +module PCP = Pulse.Checker.Pure +module RU = Pulse.RuntimeUtils +module PS = Pulse.Checker.Prover.Substs +open Pulse.Checker.Base + +let __normalize_slprop + (g:env) + (v:slprop) + : T.Tac (v':slprop & slprop_equiv g v v') += + (* Keep things reduced *) + let steps = [unascribe; primops; iota] in + + (* NOTE: whatever we unfold or reduce here will also apply for pure + slprops and under lambdas, so be conservative. Adding fst/snd here + to reduce into the projectors caused some breakages in pure slprops + mentioning them (L.map fst l == ...) as it introduced hash-consed + lambdas. *) + + (* Unfold anything marked with the "pulse_unfold" attribute. *) + let steps = steps @ [delta_attr ["Pulse.Lib.Core.pulse_unfold"; "Pulse.Lib.Core.pulse_eager_unfold"]] in + (* Unfold anything marked with F*'s "unfold" qualifier . *) + let steps = steps @ [delta_qualifier ["unfold"]] in + (* Unfold recursive definitions too, but only the ones that match the filters above. *) + let steps = steps @ [zeta] in + + let v' = PCP.norm_well_typed_term (elab_env g) steps v in + let v' = Pulse.Simplify.simplify v' in (* NOTE: the simplify stage is unverified *) + let v_equiv_v' = VE_Ext _ _ _ (RU.magic ()) in + (| v', v_equiv_v' |) + +let normalize_slprop + (g:env) + (v:slprop) + (use_rewrites_to : bool) + : T.Tac (v':slprop & slprop_equiv g v v') += + if use_rewrites_to then + let rwr = Pulse.Checker.Prover.RewritesTo.get_subst_from_env g in + let v' = PS.ss_term v rwr in + let eq_v_v' : slprop_equiv g v v' = VE_Ext _ _ _ (RU.magic ()) in + let (| v'', eq_v'_v'' |) = __normalize_slprop g v' in + (| v'', VE_Trans _ _ _ _ eq_v_v' eq_v'_v'' |) + else + __normalize_slprop g v + +let normalize_slprop_welltyped + (g:env) + (v:slprop) + (v_typing:tot_typing g v tm_slprop) + : T.Tac (v':slprop & slprop_equiv g v v' & tot_typing g v' tm_slprop) += + let (| v', v_equiv_v' |) = normalize_slprop g v true in + // FIXME: prove (or add axiom) that equiv preserves typing + (| v', v_equiv_v', E (magic()) |) \ No newline at end of file diff --git a/src/checker/Pulse.Checker.Prover.IntroExists.fsti b/src/checker/Pulse.Checker.Prover.Normalize.fsti similarity index 59% rename from src/checker/Pulse.Checker.Prover.IntroExists.fsti rename to src/checker/Pulse.Checker.Prover.Normalize.fsti index d55ebe307..539d6a372 100644 --- a/src/checker/Pulse.Checker.Prover.IntroExists.fsti +++ b/src/checker/Pulse.Checker.Prover.Normalize.fsti @@ -14,18 +14,26 @@ limitations under the License. *) -module Pulse.Checker.Prover.IntroExists +module Pulse.Checker.Prover.Normalize -module T = FStar.Tactics +module T = FStar.Tactics.V2 open Pulse.Syntax open Pulse.Typing -open Pulse.Checker.Prover.Base - -val intro_exists (#preamble:_) (pst:prover_state preamble) - (u:universe) (b:binder) (body:slprop) - (unsolved':list slprop) - (_:squash (pst.unsolved == (tm_exists_sl u b body)::unsolved')) - (prover:prover_t) - : T.Tac (pst':prover_state preamble { pst' `pst_extends` pst /\ - is_terminal pst' }) \ No newline at end of file + +val __normalize_slprop + (g:env) + (v:slprop) + : T.Tac (v':slprop & slprop_equiv g v v') + +val normalize_slprop + (g:env) + (v:slprop) + (use_rewrites_to : bool) + : T.Tac (v':slprop & slprop_equiv g v v') + +val normalize_slprop_welltyped + (g:env) + (v:slprop) + (v_typing:tot_typing g v tm_slprop) + : T.Tac (v':slprop & slprop_equiv g v v' & tot_typing g v' tm_slprop) diff --git a/src/checker/Pulse.Checker.Prover.fst b/src/checker/Pulse.Checker.Prover.fst index c4b55f540..786803c29 100644 --- a/src/checker/Pulse.Checker.Prover.fst +++ b/src/checker/Pulse.Checker.Prover.fst @@ -1,5 +1,5 @@ (* - Copyright 2023 Microsoft Research + Copyright 2025 Microsoft Research Licensed under the Apache License, Version 2.0 (the "License"); you may not use this file except in compliance with the License. @@ -15,693 +15,691 @@ *) module Pulse.Checker.Prover - -open FStar.List.Tot - -open Pulse -open Pulse.Syntax -open Pulse.Typing -open Pulse.Typing.Combinators open Pulse.Checker.Base -open Pulse.PP +open Pulse.Checker.Pure +open Pulse.Typing.Combinators +open Pulse.Typing open Pulse.Show - +open Pulse.Syntax.Base +open Pulse.Syntax.Pure +open Pulse.Syntax.Naming +open Pulse.Checker.Prover.Util +open Pulse.Checker.Prover.Normalize +open FStar.List { (@) } module RU = Pulse.RuntimeUtils -module L = FStar.List.Tot module T = FStar.Tactics.V2 -module P = Pulse.Syntax.Printer -module Metatheory = Pulse.Typing.Metatheory -module PS = Pulse.Checker.Prover.Substs -module RT = FStar.Reflection.Typing - -module ElimExists = Pulse.Checker.Prover.ElimExists -module ElimPure = Pulse.Checker.Prover.ElimPure -module ElimWithPure = Pulse.Checker.Prover.ElimWithPure -module Match = Pulse.Checker.Prover.Match -module IntroExists = Pulse.Checker.Prover.IntroExists -module IntroPure = Pulse.Checker.Prover.IntroPure -module IntroWithPure = Pulse.Checker.Prover.IntroWithPure -module Explode = Pulse.Checker.Prover.Explode -module PCP = Pulse.Checker.Pure - -let coerce_eq (#a #b:Type) (x:a) (_:squash (a == b)) : y:b{y == x} = x - -(* Checks if `p` is equivalent to emp, using the core checker. *) -let check_equiv_emp' (g:env) (p:slprop) : T.Tac (option (slprop_equiv g p tm_emp)) = - match check_equiv_emp g p with - | Some t -> Some t - | None -> - match Pulse.Typing.Util.check_equiv_now_nosmt (elab_env g) p tm_emp with - | Some tok, _ -> - Some (VE_Ext _ _ _ (RT.Rel_eq_token _ _ _ ())) - | None, _ -> None - -let __normalize_slprop - (g:env) - (v:slprop) - : T.Tac (v':slprop & slprop_equiv g v v') -= - (* Keep things reduced *) - let steps = [unascribe; primops; iota] in - - (* NOTE: whatever we unfold or reduce here will also apply for pure - slprops and under lambdas, so be conservative. Adding fst/snd here - to reduce into the projectors caused some breakages in pure slprops - mentioning them (L.map fst l == ...) as it introduced hash-consed - lambdas. *) - - (* Unfold anything marked with the "pulse_unfold" attribute. *) - let steps = steps @ [delta_attr ["Pulse.Lib.Core.pulse_unfold"; "Pulse.Lib.Core.pulse_eager_unfold"]] in - (* Unfold anything marked with F*'s "unfold" qualifier . *) - let steps = steps @ [delta_qualifier ["unfold"]] in - (* Unfold recursive definitions too, but only the ones that match the filters above. *) - let steps = steps @ [zeta] in - - let v' = PCP.norm_well_typed_term (elab_env g) steps v in - let v' = Simplify.simplify v' in (* NOTE: the simplify stage is unverified *) - let v_equiv_v' = VE_Ext _ _ _ (RU.magic ()) in - (| v', v_equiv_v' |) - -let normalize_slprop - (g:env) - (v:slprop) - (use_rewrites_to : bool) - : T.Tac (v':slprop & slprop_equiv g v v') -= - if use_rewrites_to then - let rwr = Pulse.Checker.Prover.RewritesTo.get_subst_from_env g in - let v' = rwr.(v) in - let eq_v_v' : slprop_equiv g v v' = VE_Ext _ _ _ (RU.magic ()) in - let (| v'', eq_v'_v'' |) = __normalize_slprop g v' in - (| v'', VE_Trans _ _ _ _ eq_v_v' eq_v'_v'' |) - else - __normalize_slprop g v - -let normalize_slprop_welltyped - (g:env) - (v:slprop) - (v_typing:tot_typing g v tm_slprop) - : T.Tac (v':slprop & slprop_equiv g v v' & tot_typing g v' tm_slprop) -= - let (| v', v_equiv_v' |) = normalize_slprop g v true in - // FIXME: prove (or add axiom) that equiv preserves typing - (| v', v_equiv_v', E (magic()) |) - -(* normalizes ctx and unsolved *) -let normalize_slprop_context - (#preamble:_) - (pst:prover_state preamble) - : T.Tac (pst':prover_state preamble { pst' `pst_extends` pst }) = - - let norm1 (v : slprop) : T.Tac slprop = - dfst (__normalize_slprop pst.pg pst.ss.(pst.rwr_ss.(v))) - in - - let ctxt = pst.remaining_ctxt in - let ctxt' = ctxt |> Tactics.Util.map norm1 in - - let unsolved = pst.unsolved in - let unsolved' = unsolved |> Tactics.Util.map norm1 in - - if RU.debug_at_level (fstar_env pst.pg) "ggg" then - info_doc pst.pg None [ - text "PROVER Normalized context"; - pp ctxt; - pp ctxt'; - ]; - - { pst with - remaining_ctxt = ctxt'; - remaining_ctxt_frame_typing = RU.magic (); - k = k_elab_equiv pst.k (VE_Refl _ _) (RU.magic ()); - - unsolved = unsolved'; - goals_inv = RU.magic (); - } - -#push-options "--z3rlimit_factor 4" -let rec __intro_any_exists (n:nat) - (#preamble:_) - (pst:prover_state preamble) - (prover:prover_t) - : T.Tac (pst':prover_state preamble { pst' `pst_extends` pst }) -= - if n = 0 then pst else ( - match pst.unsolved with - | [] -> pst - | hd::unsolved' -> - // info_doc pst.pg None [ - // text "Trying to introduce existential quantifier:" ^/^ - // pp hd; - // ]; - assume (hd == pst.ss.(hd)); - let hd = pst.ss.(hd) in - match inspect_term hd with - | Tm_ExistsSL u b body -> - // info_doc pst.pg (Some (RU.range_of_term hd)) [ - // text "Introducing existential quantifier:" ^/^ - // pp hd; - // ]; - IntroExists.intro_exists pst u b body unsolved' () prover - | Tm_WithPure p n v -> - IntroWithPure.intro_with_pure pst p n v unsolved' () prover - | _ -> - let pst = { - pst with - unsolved = unsolved'@[hd]; - goals_inv = magic(); - } in - __intro_any_exists (n-1) pst prover - ) -#pop-options - -let unsolved_equiv_pst (#preamble:_) (pst:prover_state preamble) (unsolved':list slprop) - (d:slprop_equiv (push_env pst.pg pst.uvs) (list_as_slprop pst.unsolved) (list_as_slprop unsolved')) - : prover_state preamble = - { pst with unsolved = unsolved'; goals_inv = RU.magic () } - -let rec collect_exists (g:env) (l:list slprop) - : exs:list slprop & - rest:list slprop & - slprop_equiv g (list_as_slprop l) (list_as_slprop (exs @ rest)) = - - match l with - | [] -> (| [], [], VE_Refl _ _ |) - | hd::tl -> - let (| exs, rest, _ |) = collect_exists g tl in - match inspect_term hd with - | Tm_ExistsSL _ _ _ -> - (| hd::exs, rest, RU.magic #(slprop_equiv _ _ _) () |) - | _ -> (| exs, hd::rest, RU.magic #(slprop_equiv _ _ _) () |) - -let rec collect_pures (g:env) (l:list slprop) - : pures:list slprop & - rest:list slprop & - slprop_equiv g (list_as_slprop l) (list_as_slprop (rest @ pures)) = - - match l with - | [] -> (| [], [], VE_Refl _ _ |) - | hd::tl -> - let (| pures, rest, _ |) = collect_pures g tl in - match inspect_term hd with - | Tm_Pure _ -> (| hd::pures, rest, RU.magic #(slprop_equiv _ _ _) () |) - | _ -> (| pures, hd::rest, RU.magic #(slprop_equiv _ _ _) () |) -#push-options "--fuel 0" -#push-options "--z3rlimit_factor 4" -let rec prove_pures #preamble (pst:prover_state preamble) - : T.Tac (pst':prover_state preamble { pst' `pst_extends` pst /\ - is_terminal pst' }) = - - match pst.unsolved with - | [] -> pst - | p::unsolved' -> - match inspect_term p with - | Tm_Pure p -> - let pst_opt = IntroPure.intro_pure pst p unsolved' () in - (match pst_opt with - | None -> - fail_doc pst.pg None [ - text "Cannot prove pure proposition" ^/^ - pp p - ] - | Some pst1 -> - let pst2 = prove_pures pst1 in - assert (pst1 `pst_extends` pst); - assert (pst2 `pst_extends` pst1); - assert (pst2 `pst_extends` pst); - pst2) - | _ -> - fail pst.pg None - (Printf.sprintf "Impossible! prover.prove_pures: %s is not a pure, please file a bug-report" - (P.term_to_string (L.hd pst.unsolved))) -#pop-options - -let intro_any_exists - (#preamble:_) - (pst:prover_state preamble) - (prover:prover_t) - : T.Tac (pst':prover_state preamble { pst' `pst_extends` pst }) = - __intro_any_exists (List.length pst.unsolved) pst prover - -noeq -type prover_iteration_res_t (#preamble:_) (pst0:prover_state preamble) = - | Stepped : lbl:string -> pst':prover_state preamble { pst' `pst_extends` pst0 } -> prover_iteration_res_t pst0 - | NoProgress of list (list document) (* possible hints *) - -(* a "subtyping" in the pst for the type above. *) -let res_advance (#preamble:_) - (#pst0 : prover_state preamble) - (#pst1 : prover_state preamble{ pst1 `pst_extends` pst0 }) - (ir : prover_iteration_res_t pst1) - : prover_iteration_res_t pst0 = - match ir with - | Stepped lbl pst1' -> Stepped lbl pst1' - | NoProgress ms -> NoProgress ms - -(* Move this type to a common module to make sure all passes adhere. *) -let prover_pass_t : Type = - (#preamble:_) -> - (pst0:prover_state preamble) -> - T.Tac (list (list document) & pst:prover_state preamble{ pst `pst_extends` pst0 }) - -(* A helper to avoid F* issue #3339. *) -noeq -type prover_pass = - | P : string -> prover_pass_t -> prover_pass - -let adddocs docs - (#preamble:_) (#pst : prover_state preamble) - (r : prover_iteration_res_t pst) -: prover_iteration_res_t pst -= - match r with - | NoProgress ms -> NoProgress (ms@docs) - | Stepped lbl pst -> Stepped lbl pst - -(* Going over the passes, stopping as soon as one makes progress. *) -let rec prover_iteration_loop - (#preamble:_) - (pst0:prover_state preamble) - (passes : list prover_pass) - : T.Tac (prover_iteration_res_t pst0) -= - match passes with - | [] -> NoProgress [] - | (P name pass)::passes' -> - let (docs, pst) = pass pst0 in - if pst.progress then ( - debug_prover pst.pg (fun _ -> - Printf.sprintf "prover: %s: made progress, remaining_ctxt after pass = %s\n" - name (show pst.remaining_ctxt)); - Stepped name pst - ) else ( - debug_prover pst.pg (fun _ -> - Printf.sprintf "prover: %s: no progress\n" name); - (* TODO: start from pst0? *) - // res_advance <| prover_iteration_loop pst passes' - adddocs docs (prover_iteration_loop pst0 passes') - ) - -let prover_pass_collect_exists (#preamble:_) (pst0:prover_state preamble) - : T.Tac (list (list document) & pst:prover_state preamble{ pst `pst_extends` pst0 }) -= - let (| exs, rest, d |) = collect_exists (push_env pst0.pg pst0.uvs) pst0.unsolved in - [], unsolved_equiv_pst pst0 (exs@rest) d - -(* One prover iteration is applying these passes until one succeeds. -If so, we return a "Stepped" with the new pst (and the prover starts -from the beginning again). If none succeeds, we return "NoProgress". *) -let prover_iteration - (#preamble:_) - (pst0:prover_state preamble) - : T.Tac (prover_iteration_res_t pst0) -= let pst = pst0 in - let pst = { pst with progress = false } in - - res_advance <| prover_iteration_loop pst [ - // P "elim_pure_pst" ElimPure.elim_pure_pst; - // ^ This is done explicitly below, but check again since some proofs - // seem to lack this step. - P "elim_exists" ElimExists.elim_exists_pst; - P "collect_exists" prover_pass_collect_exists; - P "explode" Explode.explode; - P "match_syntactic" Match.match_syntactic; - P "match_fastunif" Match.match_fastunif; - P "match_fastunif_i" Match.match_fastunif_i; - P "match_full" Match.match_full; - ] - -#push-options "--z3rlimit_factor 40 --ifuel 2" +module R = FStar.Reflection.V2 + +type head_id = + | FVarHead of R.name + | VarHead of nat + | MatchHead + +noeq type slprop_view = + | Pure : term -> slprop_view + | WithPure : term -> ppname -> body: term -> slprop_view // body is opened with () + | Exists : u: universe -> b: binder -> body: term -> slprop_view + | Atom : head: head_id -> mkeys: option (list term) -> t: slprop -> slprop_view + | Unknown : slprop -> slprop_view + +let elab_slprop (p: slprop_view) : slprop = + match p with + | Pure p -> tm_pure p + | Exists u b body -> tm_exists_sl u b body + | Atom _ _ t -> t + | WithPure p n b -> tm_with_pure p n b + | Unknown p -> p + +let rec elab_slprops (ps: list slprop_view) : slprop = + match ps with + | [] -> tm_emp + | [p] -> elab_slprop p + | p::ps -> elab_slprop p `tm_star` elab_slprops ps + +let slprop_eqv (p q: slprop) : prop = + forall g. squash (slprop_equiv g p q) + +let slprop_eqv_intro #p #q (h: (g:env -> slprop_equiv g p q)) : squash (slprop_eqv p q) = admit () +let slprop_eqv_refl (p: slprop) : squash (slprop_eqv p p) = slprop_eqv_intro fun g -> VE_Refl g p +let slprop_eqv_trans (p q r: slprop) : Lemma (requires slprop_eqv p q /\ slprop_eqv q r) (ensures slprop_eqv p r) = admit () +let slprop_eqv_star p1 q1 p2 q2 : Lemma (requires slprop_eqv p1 p2 /\ slprop_eqv q1 q2) (ensures slprop_eqv (tm_star p1 q1) (tm_star p2 q2)) = admit () +let elab_slprops_append ps qs : squash (elab_slprops (ps@qs) `slprop_eqv` (elab_slprops ps `tm_star` elab_slprops qs)) = admit () + +#push-options "--fuel 1 --ifuel 1 --z3rlimit_factor 4" #restart-solver -let rec prover - (#preamble:_) - (pst0:prover_state preamble) - : T.Tac (pst':prover_state preamble { pst' `pst_extends` pst0 /\ - is_terminal pst' }) -= (* Beta/iota/primops normalization in the context and goals. - FIXME: do this incrementally instead of on every entry to the - prover. *) - // if Nil? pst0.unsolved then pst0 else - let pst = normalize_slprop_context pst0 in - let pst = { pst with progress = false } in - (* Always eagerly eliminate pure, even if the goals are empty, - so we don't complain about a possible "leak". I think it'd be nicer - to use a typeclass for safely-droppable resources. *) - let _, pst = ElimExists.elim_exists_pst pst in - let pst = ElimWithPure.elim_with_pure_pst pst in - let pst = ElimPure.elim_pure_pst pst in - debug_prover pst.pg (fun _ -> - Printf.sprintf "At the prover top-level with\n remaining_ctxt: %s\n unsolved: %s\n allow_ambiguous: %s\n ss: %s\n rwr_ss: %s\n env: %s\n" - // (RU.stack_dump ()) - (show (list_as_slprop pst.remaining_ctxt)) - (show (list_as_slprop pst.unsolved)) - (show pst.allow_ambiguous) - (show pst.ss) - (show pst.rwr_ss) - (Pprint.render (env_to_doc pst.pg))); - - if pst.progress then begin - (* We made progress, and need to renormalize, so start over. *) - prover pst - end else match pst.unsolved with - | [] -> - (* We happen to be called on a fully-solved pst, do nothing. *) - pst - +let rec inspect_slprop (g: env) (p: slprop) : T.Tac (v:list slprop_view { elab_slprops v `slprop_eqv` p }) = + slprop_eqv_refl p; + match inspect_term p with + | Tm_Pure p -> [Pure p] + | Tm_WithPure a n b -> + assume tm_with_pure a n (open_term' b unit_const 0) `slprop_eqv` p; + [WithPure a n (open_term' b unit_const 0)] + | Tm_ExistsSL u b body -> [Exists u b body] + | Tm_Emp -> [] + | Tm_Star a b -> + let a' = inspect_slprop g a in + let b' = inspect_slprop g b in + elab_slprops_append a' b'; assert elab_slprops (a' @ b') `slprop_eqv` (elab_slprops a' `tm_star` elab_slprops b'); + slprop_eqv_star (elab_slprops a') (elab_slprops b') a b; assert (elab_slprops a' `tm_star` elab_slprops b') `slprop_eqv` tm_star a b; + slprop_eqv_trans (elab_slprops (a' @ b')) (elab_slprops a' `tm_star` elab_slprops b') (tm_star a b); assert elab_slprops (a' @ b') `slprop_eqv` tm_star a b; + a' @ b' | _ -> - match prover_iteration pst with - | Stepped name pst' -> - (* We made progress, so we start over. *) - debug_prover pst.pg (fun _ -> - Printf.sprintf "prover: made progress with pass '%s', remaining_ctxt after iteration = %s\n" - name - (show (list_as_slprop pst.remaining_ctxt))); - prover pst' - | NoProgress msgs -> - let pst = intro_any_exists pst prover in - if pst.progress then prover pst else let () = () in - - match pst.unsolved with - | [] -> pst - | hd::unsolved' -> - (* Push all pures to the back. We try them last. *) - (* TODO: can't we just call prove_pures and be done? *) - let (| pures, rest, d |) = collect_pures (push_env pst.pg pst.uvs) pst.unsolved in - let pst = unsolved_equiv_pst pst (rest@pures) d in - match pst.unsolved with - | [] -> pst - | q::tl -> - match inspect_term q with - | Tm_Pure _ -> prove_pures pst - | _ -> - (* We have a first unsolved goal that is not a pure, we fail right - now, reporting all non-pure goals. *) - let non_pures = T.filter (fun t -> not (Tm_Pure? (inspect_term t))) pst.unsolved in - let non_pures = T.map (fun q -> pst.ss.(q)) non_pures in - let q_norm : slprop = pst.ss.(q) in - match check_equiv_emp' pst.pg q_norm with // MOVE BEFORE, filter all emps before trying fastunif - | Some tok -> - (* It's emp, so just prove it here. *) - let pst' = { pst with - unsolved = unsolved'; - goals_inv = magic(); - } in - prover pst' - | None -> - let msg = [ - text ( - if List.length non_pures > 1 - then "Cannot prove any of:" - else "Cannot prove:" - ) ^^ - indent (pp <| canon_slprop_list_print non_pures); - text "In the context:" ^^ - indent (pp <| canon_slprop_list_print pst.remaining_ctxt) - ] @ (if Pulse.Config.debug_flag "initial_solver_state" then [ - text "The prover was started with goal:" ^^ - indent (pp preamble.goals); - text "and initial context:" ^^ - indent (pp preamble.ctxt); - ] else []) - in - let pass_hints = - if Cons? (List.flatten msgs) - then [ text "Some hints:" ] @ List.flatten msgs - else [] - in - // GM: I feel I should use (Some q.range) instead of None, but that makes - // several error locations worse. - fail_doc pst.pg None (msg @ pass_hints) -#pop-options -#pop-options -let rec get_q_at_hd (g:env) (l:list slprop) (q:slprop { L.existsb (fun v -> eq_tm v q) l }) - : l':list slprop & - slprop_equiv g (list_as_slprop l) (q * list_as_slprop l') = - - match l with - | hd::tl -> - if eq_tm hd q then (| tl, RU.magic #(slprop_equiv _ _ _) () |) - else let (| tl', _ |) = get_q_at_hd g tl q in - (| hd::tl', RU.magic #(slprop_equiv _ _ _) () |) - -// When we elaborate a term like `foo : x:ref int -> #y:erased int -> #_:squash (y < 1) -> ...`, -// the implicit squashed argument remain unresolved uvars after running the prover. -// This function instantiates them with (). -let prove_squash_uvars #preamble (g: env) (pst: prover_state preamble) : - T.Tac (pst':prover_state preamble { pst_extends pst' pst /\ (is_terminal pst ==> is_terminal pst') }) = - let rec check bs pst : T.Tac (pst':prover_state preamble { pst_extends pst' pst /\ (is_terminal pst ==> is_terminal pst') }) = - match bs with - | (x,t)::bs -> - if not (PS.contains pst.ss x) then - match is_squash t with - | Some t' -> - let ss = PS.push pst.ss x unit_const in - assume (ss `ss_extends` pst.ss); - let pst = { pst with - ss; - nts = None; - solved_inv = RU.magic (); - k = k_elab_equiv pst.k (VE_Refl _ _) (RU.magic ()); - } in - // NOTE: we don't need to check prop validity here, since we'll recheck the term anyhow - check bs pst - | None -> check bs pst - else - check bs pst - | [] -> pst in - check (bindings pst.uvs) pst - -#push-options "--z3rlimit_factor 16 --ifuel 2 --fuel 1 --split_queries no" -#restart-solver -let prove - (allow_ambiguous : bool) - (#g:env) (#ctxt:slprop) (ctxt_typing:slprop_typing g ctxt) - (uvs:env { disjoint g uvs }) - (#goals:slprop) (goals_typing:slprop_typing (push_env g uvs) goals) - - : T.Tac (g1 : env { g1 `env_extends` g /\ disjoint g1 uvs } & - nts : PS.nt_substs & - effect_labels:list T.tot_or_ghost { PS.well_typed_nt_substs g1 uvs nts effect_labels } & - remaining_ctxt : slprop & - continuation_elaborator g ctxt g1 ((PS.nt_subst_term goals nts) * remaining_ctxt)) = - - debug_prover g (fun _ -> - Printf.sprintf "\nEnter top-level prove with ctxt: %s\ngoals: %s\n" - (P.term_to_string ctxt) (P.term_to_string goals)); - - let ctxt_l = slprop_as_list ctxt in - - let ctxt_frame_typing : slprop_typing g (ctxt * tm_emp) = RU.magic () in - let preamble = { - g0 = g; - ctxt; - frame = tm_emp; - ctxt_frame_typing; - goals; - } in - assume (list_as_slprop (slprop_as_list ctxt) == ctxt); - assume ((PS.empty).(tm_emp) == tm_emp); - let rwr_ss = Pulse.Checker.Prover.RewritesTo.get_subst_from_env g in - let pst0 : prover_state preamble = { - pg = g; - remaining_ctxt = slprop_as_list ctxt; - remaining_ctxt_frame_typing = ctxt_frame_typing; - uvs = uvs; - ss = PS.empty; - rwr_ss; - nts = None; - solved = tm_emp; - unsolved = slprop_as_list goals; - k = k_elab_equiv_continuation (k_elab_unit g (preamble.ctxt * preamble.frame)) (RU.magic ()); - goals_inv = RU.magic (); - solved_inv = (); - progress = false; - allow_ambiguous = allow_ambiguous; - } in - - let pst = RU.record_stats "Pulse.prover" fun _ -> prover pst0 in - - let pst = prove_squash_uvars g pst in - - let (| nts, effect_labels |) - : nts:PS.nt_substs & - effect_labels:list T.tot_or_ghost { - PS.well_typed_nt_substs pst.pg pst.uvs nts effect_labels /\ - PS.is_permutation nts pst.ss - } = - match pst.nts with - | Some nts -> nts + match T.hua p with + | Some (h, _, _) -> + let h = R.inspect_fv h in + let mkeys = Pulse.Checker.Prover.Match.MKeys.get_mkeys g p in + [Atom (FVarHead h) mkeys p] | None -> - // warn_doc pst.pg None [ - // text <| Printf.sprintf "About to translate prover state to nts (nts is None)"; - // prefix 2 1 (text "pst.pg =") (pp pst.pg); - // prefix 2 1 (text "pst.uvs =") (pp pst.uvs); - // prefix 2 1 (text "pst.ss =") (pp pst.ss); - // prefix 2 1 (text "pst.remaining_ctxt =") (pp pst.remaining_ctxt); - // prefix 2 1 (text "pst.unsolved =") (pp pst.unsolved); - // ]; - let r = PS.ss_to_nt_substs pst.pg pst.uvs pst.ss in - match r with - | Inr msg -> - fail_doc pst.pg None [ - text <| Printf.sprintf "Prover error: ill-typed substitutions (%s)" msg; - prefix 2 1 (text "pst.pg =") (pp pst.pg); - prefix 2 1 (text "pst.uvs =") (pp pst.uvs); - prefix 2 1 (text "pst.ss =") (pp pst.ss); - prefix 2 1 (text "pst.remaining_ctxt =") (pp pst.remaining_ctxt); - prefix 2 1 (text "pst.unsolved =") (pp pst.unsolved); - ] - | Inl nts -> nts - in - let nts_uvs, nts_uvs_effect_labels = - PS.well_typed_nt_substs_prefix pst.pg pst.uvs nts effect_labels uvs in - let k () - : continuation_elaborator - g (ctxt * tm_emp) - pst.pg ((list_as_slprop pst.remaining_ctxt * tm_emp) * (PS.nt_subst_term pst.solved nts)) = pst.k in - // admit () - let goals_inv - : slprop_equiv (push_env pst.pg pst.uvs) goals (list_as_slprop [] * pst.solved) = pst.goals_inv in - let goals_inv - : slprop_equiv pst.pg (PS.nt_subst_term goals nts) (PS.nt_subst_term (list_as_slprop [] * pst.solved) nts) = - PS.slprop_equiv_nt_substs_derived pst.pg pst.uvs goals_inv nts effect_labels in - - // goals is well-typed in initial g + uvs - // so any of the remaining uvs in pst.uvs should not be in goals - // so we can drop their substitutions from the tail of nts - assume (PS.nt_subst_term goals nts == PS.nt_subst_term goals nts_uvs); - - (| pst.pg, nts_uvs, nts_uvs_effect_labels, - list_as_slprop pst.remaining_ctxt, - k_elab_equiv (k ()) (RU.magic ()) (RU.magic ()) |) -#pop-options - -let canon_post (c:comp_st) : comp_st = - let canon_st_comp_post (c:st_comp) : st_comp = - match inspect_term c.post with - | Tm_FStar _ -> c - | post_v -> { c with post=pack_term_view_wr post_v (RU.range_of_term c.post) } - in - match c with - | C_ST s -> C_ST (canon_st_comp_post s) - | C_STAtomic i obs s -> C_STAtomic i obs (canon_st_comp_post s) - | C_STGhost i s -> C_STGhost i (canon_st_comp_post s) - -irreducible -let typing_canon #g #t (#c:comp_st) (d:st_typing g t c) : st_typing g t (canon_post c) = - assume false; - d - -#push-options "--z3rlimit_factor 10 --fuel 0 --ifuel 2 --split_queries no" -#restart-solver -let try_frame_pre_uvs - (allow_ambiguous : bool) - (#g:env) (#ctxt:slprop) (ctxt_typing:tot_typing g ctxt tm_slprop) - (uvs:env { disjoint g uvs }) - (d:(t:st_term & c:comp_st & st_typing (push_env g uvs) t c)) - (res_ppname:ppname) - - : T.Tac (checker_result_t g ctxt NoHint) = - - let (| t, c, d |) = d in - - let g = push_context g "try_frame_pre" t.range in - - let (| g1, nts, effect_labels, remaining_ctxt, k_frame |) = - RU.record_stats "Pulse.prove" fun _ -> - prove allow_ambiguous #g #_ ctxt_typing uvs #(comp_pre c) (RU.magic ()) in - // assert (nts == []); - - let d : st_typing (push_env g1 uvs) t c = - Metatheory.st_typing_weakening g uvs t c d g1 in - assert (comp_pre (PS.nt_subst_comp c nts) == PS.nt_subst_term (comp_pre c) nts); - let t = PS.nt_subst_st_term t nts in - let c = PS.nt_subst_comp c nts in - - let d : st_typing g1 t c = - let r = PS.st_typing_nt_substs_derived g1 uvs d nts effect_labels in - match r with - | Inr (x, x_t) -> - let open Pulse.PP in - fail_doc g1 (Some t.range) [ - text "Prover error"; - text "For term" ^/^ pp t ^/^ text "implicit solution" ^/^ pp x_t ^/^ text "has ghost effect."; - ] - | Inl d -> d in - - (* shouldn't need this once term becomes a view; currently we sometimes end up with a computation - type whose postcondition is Tm_FStar (`(p1 ** p2))) rather than a Tm_Star p1 p2; - canon_post normalizes that *) - let c = canon_post c in - let d = typing_canon d in - - let k_frame : continuation_elaborator g ctxt g1 (comp_pre c * remaining_ctxt) = coerce_eq k_frame () in - - let x = fresh g1 in - let ty = comp_res c in - let g2 = push_binding g1 x res_ppname ty in - assert (g2 `env_extends` g1); - let ctxt' = (open_term_nv (comp_post c) (res_ppname, x) * remaining_ctxt) in - - let d : st_typing g1 t c = Metatheory.st_typing_weakening_standard d g1 in - - let k - : continuation_elaborator g1 (remaining_ctxt * comp_pre c) - g2 ctxt' = - continuation_elaborator_with_bind remaining_ctxt d (RU.magic #(tot_typing _ _ _) ()) (res_ppname, x) in - let k - : continuation_elaborator g1 (comp_pre c * remaining_ctxt) - g2 ctxt' = - k_elab_equiv k (VE_Comm _ _ _) (VE_Refl _ _) in - - let k = k_elab_trans k_frame k in - - let comp_res_typing_in_g1, _, f = - Metatheory.st_comp_typing_inversion_cofinite - (fst <| Metatheory.comp_typing_inversion (Metatheory.st_typing_correctness d)) in - - let d_ty - : universe_of g2 ty (comp_u c) = - Metatheory.tot_typing_weakening_single comp_res_typing_in_g1 x (comp_res c) in - - assume (~ (x `Set.mem` freevars (comp_post c))); - let d_post - : slprop_typing g2 (open_term_nv (comp_post c) (res_ppname, x)) = - f x in - - // the RU.magic is for the ctxt' typing - // see d_post for post typing - // then the remaining_ctxt typing should come from the prover state - // TODO: add it there - // and then ctxt' is just their `*` - - let t : (u:universe & t:typ & universe_of g2 t u) = (| comp_u c, ty, d_ty |) in - let ctxt' : (ctxt':slprop & tot_typing g2 ctxt' tm_slprop) = (| ctxt', RU.magic #(tot_typing _ _ _) () |) in - let k : continuation_elaborator g ctxt g2 (dfst ctxt') = k in - assert_spinoff (lookup g2 x == Some ty); - assert_spinoff (checker_result_inv g NoHint x g2 t ctxt'); - (| x, g2, t, ctxt', k |) + let hd, _ = R.collect_app_ln p in + match R.inspect_ln hd with + | R.Tv_Var x -> + let x = R.inspect_namedv x in + [Atom (VarHead x.uniq) None p] + | _ -> + match R.inspect_ln p with + | R.Tv_Match sc _ _ -> + let sc = RU.deep_compress_safe sc in + if RU.no_uvars_in_term sc then + [Atom MatchHead (Some [sc]) p] + else + [Unknown p] + | _ -> + [Unknown p] #pop-options -let try_frame_pre - (allow_ambiguous : bool) - (#g:env) (#ctxt:slprop) (ctxt_typing:tot_typing g ctxt tm_slprop) - (d:(t:st_term & c:comp_st & st_typing g t c)) - (res_ppname:ppname) +let cont_elab g ps g' ps' = + frame: list slprop_view -> continuation_elaborator g (elab_slprops (frame @ ps)) g' (elab_slprops (frame @ ps')) + +let cont_elab_refl g ps ps' (h: slprop_equiv g (elab_slprops ps) (elab_slprops ps')) : cont_elab g ps g ps' = + fun frame -> k_elab_equiv (k_elab_unit g (elab_slprops (frame @ ps))) (VE_Refl _ _) (RU.magic ()) + +let cont_elab_trans #g1 (#g2: env { g2 `env_extends` g1 }) (#g3: env { g3 `env_extends` g2 }) #ps1 #ps2 #ps2' #ps3 + (k1: cont_elab g1 ps1 g2 ps2) + (k2: cont_elab g2 ps2' g3 ps3) + (h: slprop_equiv g2 (elab_slprops ps2) (elab_slprops ps2')) : + cont_elab g1 ps1 g3 ps3 = + fun frame -> k_elab_trans (k1 frame) (k_elab_equiv (k2 frame) (RU.magic ()) (VE_Refl _ _)) + +let cont_elab_equiv #g1 #ps1 #ps1' #g2 #ps2 #ps2' + (k: cont_elab g1 ps1 g2 ps2) + (h1: slprop_equiv g1 (elab_slprops ps1) (elab_slprops ps1')) + (h2: slprop_equiv g2 (elab_slprops ps2) (elab_slprops ps2')) : + cont_elab g1 ps1' g2 ps2' = + fun frame -> k_elab_equiv (k frame) (RU.magic ()) (RU.magic ()) + +let cont_elab_frame #g #ps #g' #ps' (k: cont_elab g ps g' ps') frame : + cont_elab g (frame @ ps) g' (frame @ ps') = + fun frame' -> k_elab_equiv (k (frame' @ frame)) (RU.magic()) (RU.magic()) + +let prover_result (g: env) (ctxt goals: list slprop_view) = + g':env { env_extends g' g } & + ctxt': list slprop_view & + goals': list slprop_view & + solved: list slprop_view & + (g'': env { env_extends g' g } -> T.Tac + (cont_elab g ctxt g' (solved @ ctxt') & + (cont_elab g'' (solved @ goals') g'' goals))) + +let prover_result_join #g #ctxt #goals #g1 #ctxt1 #goals1 + (r: prover_result g ctxt goals { let (| g1', ctxt1', goals1', _, _ |) = r in g1' == g1 /\ ctxt1' == ctxt1 /\ goals1' == goals1 }) + (r1: prover_result g1 ctxt1 goals1) : + prover_result g ctxt goals = + let (| g1, ctxt1, goals1, solved1, k1 |) = r in + let (| g2, ctxt2, goals2, solved2, k2 |) = r1 in + (| g2, ctxt2, goals2, solved1 @ solved2, fun g3 -> + let before1, after1 = k1 g3 in + let before2, after2 = k2 g3 in + (fun frame -> + let h1: slprop_equiv g1 (elab_slprops ((frame @ solved1) @ ctxt1)) (elab_slprops (frame @ solved1 @ ctxt1)) = RU.magic () in + let h2: slprop_equiv g2 (elab_slprops ((frame @ solved1) @ solved2 @ ctxt2)) (elab_slprops (frame @ (solved1 @ solved2) @ ctxt2)) = RU.magic () in + k_elab_trans (before1 frame) (k_elab_equiv (before2 (frame @ solved1)) h1 h2)), + (fun frame -> + let h1: slprop_equiv g3 (elab_slprops ((frame @ solved1) @ solved2 @ goals2)) (elab_slprops (frame @ (solved1 @ solved2) @ goals2)) = RU.magic () in + let h2: slprop_equiv g3 (elab_slprops ((frame @ solved1) @ goals1)) (elab_slprops (frame @ solved1 @ goals1)) = RU.magic () in + k_elab_trans (k_elab_equiv (after2 (frame @ solved1)) h1 h2) (after1 frame)) + <: T.Tac _ |) + +let prove_first (g: env) (ctxt goals: list slprop_view) + (k: (goal: slprop_view -> T.Tac (option (prover_result g ctxt [goal])))) : + T.Tac (option (prover_result g ctxt goals)) = + let goals0 = goals in + let rec go (goals_left_rev: list slprop_view) (goals: list slprop_view { List.rev goals_left_rev @ goals == goals0 }) : + T.Tac (option (prover_result g ctxt goals0)) = + match goals with + | [] -> None + | goal::goals -> + match k goal with + | Some (| g', ctxt', goals', solved, res |) -> + Some (| g', ctxt', List.rev goals_left_rev @ goals' @ goals, solved, fun g'' -> + let before, after = res g'' in + before, + (fun frame -> + let h1 : slprop_equiv g'' + (elab_slprops ((frame @ List.Tot.Base.rev goals_left_rev @ goals) @ solved @ goals')) + (elab_slprops (frame @ solved @ List.Tot.Base.rev goals_left_rev @ goals' @ goals)) = RU.magic () in + let h2 : slprop_equiv g'' + (elab_slprops ((frame @ List.Tot.Base.rev goals_left_rev @ goals) @ [goal])) + (elab_slprops (frame @ goals0)) = RU.magic () in + k_elab_equiv (after (frame @ List.rev goals_left_rev @ goals)) h1 h2) |) + | None -> + assert List.rev goals_left_rev @ (goal::goals) == goals0; + assume List.rev (goal::goals_left_rev) @ goals == goals0; + go (goal::goals_left_rev) goals in + go [] goals + +let deep_compress_st_comp (c:st_comp) : st_comp = + { u = c.u; res = RU.deep_compress_safe c.res; pre = RU.deep_compress_safe c.pre; post = RU.deep_compress_safe c.post } + +let deep_compress_comp (c:comp {stateful_comp c}) : comp = + with_st_comp c (deep_compress_st_comp (st_comp_of_comp c)) + +let continuation_elaborator_with_bind_nondep (#g:env) (ctxt:term) + (#c1:comp{stateful_comp c1}) + (#e1:st_term) + (e1_typing:st_typing g e1 c1) + (ctxt_pre1_typing:tot_typing g (tm_star ctxt (comp_pre c1)) tm_slprop) + : T.Tac (continuation_elaborator + g + (tm_star ctxt (comp_pre c1)) + g + (tm_star (comp_post c1) ctxt)) = + let x = fresh g in + admit (); + continuation_elaborator_with_bind (RU.deep_compress_safe ctxt) #(deep_compress_comp c1) e1_typing ctxt_pre1_typing (ppname_default, x) + +let tot_typing_tm_unit (g: env) : tot_typing g tm_unit (tm_type u0) = RU.magic () + +let intro_pure (g: env) (frame: slprop) (p: term) : + continuation_elaborator g frame g (frame `tm_star` tm_pure p) = + fun post t -> + let g = push_context g "check_intro_pure" (RU.range_of_term p) in + let p_typing: tot_typing g p tm_prop = RU.magic() in // implied by t2_typing + let pv = check_prop_validity g p p_typing in + let frame_typ : tot_typing g frame tm_slprop = RU.magic () in // implied by t2_typing + let h: tot_typing g (tm_star frame (comp_pre (comp_intro_pure p))) tm_slprop = RU.magic () in + debug_prover g (fun _ -> Printf.sprintf "intro_pure p=%s\nframe=%s\n" (show p) (show frame)); + k_elab_equiv (continuation_elaborator_with_bind_nondep frame (T_IntroPure g p p_typing pv) h) (RU.magic ()) (RU.magic ()) + post t + +let is_uvar (t:term) : bool = + match R.inspect_ln t with + | R.Tv_Uvar .. -> true + | _ -> false + +// solve `pure (t == ?u)` via unification +let is_eq_unif (g: env) (p: term) : Dv (option (term & term)) = + match is_eq2 (RU.deep_compress_safe p) with + | Some (_, lhs, rhs) -> + if is_uvar lhs || is_uvar rhs then + Some (lhs, rhs) + else + None + | None -> None +let pure_eq_unif (g: env) (p: term) skip_eq_uvar : Dv bool = + match is_eq_unif g p with + | Some (lhs, rhs) -> + if skip_eq_uvar then + true + else ( + ignore (RU.teq_nosmt_force (elab_env g) lhs rhs); + false + ) + | None -> false + +// skip_eq_uvar to support (assert foo. with pure (foo == ....)) +let prove_pure (g: env) (ctxt: list slprop_view) (skip_eq_uvar: bool) (goal: slprop_view) : + T.Tac (option (prover_result g ctxt [goal])) = + match goal with + | Pure p -> + if pure_eq_unif g p skip_eq_uvar then None else + + Some (| g, ctxt, [], [], fun g'' -> + cont_elab_refl g ctxt ([] @ ctxt) (VE_Refl _ _), + (fun frame -> + let h1: slprop_equiv g'' (elab_slprops frame) (elab_slprops (frame @ [] @ [])) = RU.magic () in + let h2: slprop_equiv g'' (tm_star (elab_slprops frame) (tm_pure p)) (elab_slprops (frame @ [goal])) = RU.magic () in + k_elab_equiv (intro_pure g'' (elab_slprops frame) p) h1 h2) + <: T.Tac _ |) + | _ -> None + +// let foo = u2 + +let intro_with_pure (g: env) (frame: slprop) (p: term) (n: ppname) (v: term) : + continuation_elaborator g (frame `tm_star` v) g (frame `tm_star` tm_with_pure p n v) = + fun post t -> + let g = push_context g "check_intro_with_pure" (RU.range_of_term p) in + let p_typing: tot_typing g p tm_prop = RU.magic() in // implied by t2_typing + let pv = check_prop_validity g p p_typing in + let frame_typ : tot_typing g frame tm_slprop = RU.magic () in // implied by t2_typing + let ty = mk_squash u0 p in + let st = wtag (Some STT_Ghost) (Tm_ST { t = tm_unknown; args = [] }) in + let c = C_STGhost tm_emp_inames { u=u0; res=tm_unit; pre=v; post=tm_with_pure p n v } in + let typing: st_typing g st c = RU.magic () in + let h: tot_typing g (tm_star frame (comp_pre c)) tm_slprop = RU.magic () in + debug_prover g (fun _ -> Printf.sprintf "intro_pure p=%s\nframe=%s\n" (show p) (show frame)); + k_elab_equiv (continuation_elaborator_with_bind_nondep frame typing h) (RU.magic ()) (RU.magic ()) + post t + +let prove_with_pure (g: env) (ctxt: list slprop_view) skip_eq_uvar (goal: slprop_view) : + T.Tac (option (prover_result g ctxt [goal])) = + match goal with + | WithPure p n v -> + if pure_eq_unif g p skip_eq_uvar then None else + + Some (| g, ctxt, [Unknown v], [], fun g'' -> + cont_elab_refl g ctxt ([] @ ctxt) (VE_Refl _ _), + (fun frame -> + let h1: slprop_equiv g'' (tm_star (elab_slprops frame) v) (elab_slprops (frame @ [Unknown v] @ [])) = RU.magic () in + let h2: slprop_equiv g'' (tm_star (elab_slprops frame) (tm_with_pure p n v)) + (elab_slprops (frame @ [goal])) = RU.magic () in + k_elab_equiv (intro_with_pure g'' (elab_slprops frame) p n v) h1 h2) + <: T.Tac _ |) + | _ -> None + +let intro_exists (g: env) (frame: slprop) (u: universe) (b: binder) (body: slprop) (e: term) : + continuation_elaborator g (frame `tm_star` open_term' body e 0) g (frame `tm_star` tm_exists_sl u b body) = + fun post t -> + let g = push_context g "check_intro_exists" (RU.range_of_term body) in + let frame_typ : tot_typing g frame tm_slprop = RU.magic () in // implied by t2_typing + let binder_ty_typ : tot_typing g b.binder_ty (tm_type u) = RU.magic() in // implied by t2_typing + let tm_ex_typ : tot_typing g (tm_exists_sl u b body) tm_slprop = RU.magic() in // implied by t2_typing + let e_typ = core_check_term' g e T.E_Ghost b.binder_ty (fun _ -> let open Pulse.PP in + [text "Cannot find witness for" ^/^ pp (tm_exists_sl u b body)]) in + let h1: tot_typing g (tm_star frame (comp_pre (comp_intro_exists u b body e))) tm_slprop = RU.magic () in + let h2: slprop_equiv g (tm_star frame (comp_pre (comp_intro_exists u b body e))) (tm_star frame (open_term' body e 0)) = RU.magic () in + let h3: slprop_equiv g (tm_star (comp_post (comp_intro_exists u b body e)) frame) (tm_star frame (tm_exists_sl u b body)) = RU.magic () in + debug_prover g (fun _ -> Printf.sprintf "intro_exists %s\nframe=%s\n" (show (tm_exists_sl u b body)) (show frame)); + k_elab_equiv (continuation_elaborator_with_bind_nondep frame (T_IntroExists g u b body e binder_ty_typ tm_ex_typ e_typ) h1) h2 h3 + post t + +let mk_uvar (g: env) (ty: term) : T.Tac term = + // TODO + fst (tc_term_phase1_with_type g tm_unknown ty) + +let prove_exists (g: env) (ctxt: list slprop_view) (goal: slprop_view) : + T.Tac (option (prover_result g ctxt [goal])) = + match goal with + | Exists u b body -> + let e = mk_uvar g b.binder_ty in // unnecessarily restrictive environment for uvar + Some (| g, ctxt, [Unknown (open_term' body e 0)], [], fun g'' -> + cont_elab_refl g ctxt ([] @ ctxt) (VE_Refl _ _), + (fun frame -> + let h1: slprop_equiv g'' (tm_star (elab_slprops frame) (open_term' body e 0)) (elab_slprops (frame @ [] @ [Unknown (open_term' body e 0)])) = RU.magic () in + let h2: slprop_equiv g'' (tm_star (elab_slprops frame) (tm_exists_sl u b body)) (elab_slprops (frame @ [goal])) = RU.magic () in + k_elab_equiv (intro_exists g'' (elab_slprops frame) u b body e) h1 h2) + <: T.Tac _ |) + | _ -> None + +let unpack_and_norm_goal (g: env) (ctxt: list slprop_view) (goal: slprop_view) : + T.Tac (option (prover_result g ctxt [goal])) = + match goal with + | Unknown goal -> + let (| goal', goal_eq_goal' |) = normalize_slprop g goal false in + let goal'' = inspect_slprop g goal' in + (match goal'' with + | [Unknown _] -> None + | _ -> Some (| g, ctxt, goal'', [], fun g' -> + let h: slprop_equiv g' (elab_slprops ([] @ goal'')) (elab_slprops [Unknown goal]) = RU.magic () in + cont_elab_refl _ _ _ (VE_Refl _ _), cont_elab_refl _ _ _ h + <: T.Tac _ |)) + | _ -> None + +let prover_result_samegoals g ctxt goals = + r:prover_result g ctxt goals { let (| _, _, goals', _, _ |) = r in goals' == goals } + +let prover_result_nogoals g ctxt = + prover_result_samegoals g ctxt [] + +let elim_first' (g: env) (ctxt0 goals: list slprop_view) + (k: (ctxt: slprop_view -> T.Tac (option (prover_result_nogoals g [ctxt])))) : + T.Tac (option (prover_result_samegoals g ctxt0 goals)) = + let rec go (ctxt_left_rev: list slprop_view) (ctxt: list slprop_view { List.rev ctxt_left_rev @ ctxt == ctxt0 }) : + T.Tac (option (prover_result_samegoals g ctxt0 goals)) = + match ctxt with + | [] -> None + | c::ctxt -> + match k c with + | Some (| g', ctxt', goals', solved, res |) -> + assert goals' == []; + Some (| g', List.rev ctxt_left_rev @ ctxt' @ ctxt, goals, solved, fun g'' -> + let before, after = res g'' in + let h1: slprop_equiv g (elab_slprops ((List.Tot.Base.rev ctxt_left_rev @ ctxt) @ [c])) (elab_slprops ctxt0) = RU.magic () in + let h2: slprop_equiv g' (elab_slprops ((List.Tot.Base.rev ctxt_left_rev @ ctxt) @ solved @ ctxt')) + (elab_slprops (solved @ List.Tot.Base.rev ctxt_left_rev @ ctxt' @ ctxt)) = RU.magic () in + let h3: slprop_equiv g'' (elab_slprops (goals @ solved @ goals')) (elab_slprops (solved @ goals)) = RU.magic () in + let h4: slprop_equiv g'' (elab_slprops (goals @ [])) (elab_slprops goals) = RU.magic () in + cont_elab_equiv (cont_elab_frame before (List.rev ctxt_left_rev @ ctxt)) h1 h2, + cont_elab_equiv (cont_elab_frame after goals) h3 h4 |) + | None -> + assert List.rev ctxt_left_rev @ (c::ctxt) == ctxt0; + assume List.rev (c::ctxt_left_rev) @ ctxt == ctxt0; + go (c::ctxt_left_rev) ctxt in + go [] ctxt0 + +let elim_first (g: env) (ctxt0 goals: list slprop_view) + (k: (ctxt: slprop_view -> T.Tac (option (prover_result_nogoals g [ctxt])))) : + T.Tac (option (prover_result g ctxt0 goals)) = + match elim_first' g ctxt0 goals k with + | Some r -> Some r + | None -> None + +let unpack_and_norm_ctxt (g: env) (ctxt: slprop_view) : + T.Tac (option (prover_result_nogoals g [ctxt])) = + match ctxt with + | Unknown ctxt -> + let (| ctxt', ctxt_eq_ctxt' |) = normalize_slprop g ctxt false in + let ctxt'' = inspect_slprop g ctxt' in + (match ctxt'' with + | [Unknown _] -> None + | _ -> Some (| g, ctxt'', [], [], fun g' -> + let h: slprop_equiv g ctxt (elab_slprops ([] @ ctxt'')) = RU.magic () in + cont_elab_refl _ _ _ h, cont_elab_refl _ _ _ (VE_Refl _ _) + <: T.Tac _ |)) + | _ -> None + +let elim_pure (g: env) (frame: slprop) (p: term) (x: nvar { ~(Set.mem (snd x) (dom g)) }) + (g': env { g' == push_binding g (snd x) (fst x) (mk_squash u0 p) }) : + continuation_elaborator g (frame `tm_star` tm_pure p) g' frame = fun post t -> + let ty = mk_squash u0 p in + let st = wtag (Some STT_Ghost) (Tm_ST { t = tm_unknown; args = [] }) in + let c = C_STGhost tm_emp_inames { u=u0; res=ty; pre=tm_pure p; post=tm_emp } in + let typing: st_typing g st c = RU.magic () in + let h: tot_typing g (tm_star frame (comp_pre c)) tm_slprop = RU.magic () in + let h2: slprop_equiv g' (tm_star (open_term_nv (comp_post c) x) frame) frame = RU.magic () in + let k: continuation_elaborator g (tm_star frame (tm_pure p)) g' (tm_star tm_emp frame) = + continuation_elaborator_with_bind frame typing h x in + k_elab_equiv k (VE_Refl _ _) h2 post t + +let elim_pure_step (g: env) (ctxt: slprop_view) : + T.Tac (option (prover_result_nogoals g [ctxt])) = + match ctxt with + | Pure p -> + let x = ppname_default, fresh g in + let ty = mk_squash u0 p in + let g' = push_binding g (snd x) (fst x) ty in + Some (| g', [], [], [], fun g'' -> + (fun frame -> + let h1: slprop_equiv g (tm_star (elab_slprops frame) (tm_pure p)) (elab_slprops (frame @ [ctxt])) = RU.magic () in + let h2: slprop_equiv g' (elab_slprops frame) (elab_slprops (frame @ [] @ [])) = RU.magic () in + k_elab_equiv (elim_pure g (elab_slprops frame) p x g') h1 h2), + cont_elab_refl _ _ _ (VE_Refl _ _) + <: T.Tac _ |) + | _ -> None + +let elim_with_pure (g: env) (frame: slprop) (p: term) (x: nvar { ~(Set.mem (snd x) (dom g)) }) (v: term) + (g': env { g' == push_binding g (snd x) (fst x) (mk_squash u0 p) }) : + continuation_elaborator g (frame `tm_star` tm_with_pure p (fst x) v) g' + (frame `tm_star` v) = fun post t -> + let ty = mk_squash u0 p in + let st = wtag (Some STT_Ghost) (Tm_ST { t = tm_unknown; args = [] }) in + let c = C_STGhost tm_emp_inames { u=u0; res=ty; pre=tm_with_pure p (fst x) v; post=v } in + assume open_term v (snd x) == v; // no loose bvars + let typing: st_typing g st c = RU.magic () in + let h: tot_typing g (tm_star frame (comp_pre c)) tm_slprop = RU.magic () in + let h2: slprop_equiv g' (tm_star (open_term_nv (comp_post c) x) frame) (tm_star frame v) = RU.magic () in + let k: continuation_elaborator g (tm_star frame (tm_with_pure p (fst x) v)) g' (tm_star v frame) = + continuation_elaborator_with_bind frame typing h x in + k_elab_equiv k (VE_Refl _ _) h2 post t + +let elim_with_pure_step (g: env) (ctxt: slprop_view) : + T.Tac (option (prover_result_nogoals g [ctxt])) = + match ctxt with + | WithPure p n v -> + let x = n, fresh g in + let ty = mk_squash u0 p in + let g' = push_binding g (snd x) (fst x) ty in + Some (| g', [Unknown v], [], [], fun g'' -> + (fun frame -> + let h1: slprop_equiv g (tm_star (elab_slprops frame) (tm_with_pure p (fst x) v)) (elab_slprops (frame @ [ctxt])) = RU.magic () in + let h2: slprop_equiv g' (tm_star (elab_slprops frame) v) (elab_slprops (frame @ [Unknown v] @ [])) = RU.magic () in + k_elab_equiv (elim_with_pure g (elab_slprops frame) p x v g') h1 h2), + cont_elab_refl _ _ _ (VE_Refl _ _) + <: T.Tac _ |) + | _ -> None + +// TODO: don't always erase +let elim_exists (g: env) (frame: slprop) u b body (x: nvar { ~(Set.mem (snd x) (dom g)) }) + (g': env { g' == push_binding g (snd x) (fst x) (mk_erased u b.binder_ty) }) : + continuation_elaborator g (frame `tm_star` tm_exists_sl u b body) + g' (frame `tm_star` open_term' body (mk_reveal u b.binder_ty (term_of_nvar x)) 0) = fun post t -> + let c = comp_elim_exists u b.binder_ty body x in + let h1: tot_typing g b.binder_ty (tm_type u) = RU.magic () in + let h2: tot_typing g (tm_exists_sl u (as_binder b.binder_ty) body) tm_slprop = RU.magic () in + let typing: st_typing g _ c = T_ElimExists g u b.binder_ty body (snd x) h1 h2 in + let h: tot_typing g (tm_star frame (comp_pre c)) tm_slprop = RU.magic () in + let c_post_x = open_term' body (mk_reveal u b.binder_ty (term_of_nvar x)) 0 in + assume open_term (comp_post c) (snd x) == c_post_x; + let h2: slprop_equiv g' (tm_star c_post_x frame) (tm_star frame c_post_x) = RU.magic () in + let k: continuation_elaborator g (tm_star frame (tm_exists_sl u b body)) g' (tm_star c_post_x frame) = + continuation_elaborator_with_bind frame typing h x in + k_elab_equiv k (VE_Refl _ _) h2 post t + +let elim_exists_step (g: env) (ctxt: slprop_view) : + T.Tac (option (prover_result_nogoals g [ctxt])) = + match ctxt with + | Exists u b body -> + let n = "_" ^ T.unseal b.binder_ppname.name ^ string_of_int (List.length (bindings g)) in + let n: ppname = { name = T.seal n; range = RU.range_of_term body } in + let x = n, fresh g in + let ty = mk_erased u b.binder_ty in + let g' = push_binding g (snd x) (fst x) ty in + let result = open_term' body (mk_reveal u b.binder_ty (term_of_nvar x)) 0 in + Some (| g', [Unknown result], [], [], fun g'' -> + (fun frame -> + let h1: slprop_equiv g (tm_star (elab_slprops frame) (tm_exists_sl u b body)) (elab_slprops (frame @ [ctxt])) = RU.magic () in + let h2: slprop_equiv g' (tm_star (elab_slprops frame) result) (elab_slprops (frame @ [] @ [Unknown result])) = RU.magic () in + k_elab_equiv (elim_exists g (elab_slprops frame) u b body x g') h1 h2), + cont_elab_refl _ _ _ (VE_Refl _ _) + <: T.Tac _ |) + | _ -> None + +exception AbortUFTransaction of bool +let with_uf_transaction (k: unit -> T.Tac bool) : T.Tac bool = + let open FStar.Tactics.V2 in + try + T.raise <| AbortUFTransaction <| k () + with + | AbortUFTransaction res -> res + | ex -> T.raise ex + +let forallb (k: 'a -> T.Tac bool) (xs: list 'a) = + not (T.existsb (fun x -> not (k x)) xs) - : T.Tac (checker_result_t g ctxt NoHint) = +open Pulse.PP - let uvs = mk_env (fstar_env g) in - assert (equal g (push_env g uvs)); - try_frame_pre_uvs allow_ambiguous ctxt_typing uvs d res_ppname +module RT = FStar.Reflection.Typing +let check_slprop_equiv_ext r (g:env) (p q:slprop) +: T.Tac (slprop_equiv g p q) += + let p = RU.deep_compress_safe p in + let q = RU.deep_compress_safe q in + let res, issues = Pulse.Typing.Util.check_equiv_now (elab_env g) p q in + match res with + | None -> + fail_doc_with_subissues g (Some r) issues [ + text "rewrite: could not prove equality of"; + pp p; + pp q; + ] + | Some token -> + VE_Ext g p q (RT.Rel_eq_token _ _ _ ()) + +let teq_nosmt_force_args (g: R.env) (x y: term) (fail_fast: bool) : Dv bool = + let rec go (xs ys: list R.argv) : Dv bool = + match xs, ys with + | [], [] -> true + | (x,_)::xs, (y,_)::ys -> + if RU.teq_nosmt_force g x y then + go xs ys + else ( + if not fail_fast then ignore (go xs ys); + false + ) + | _ -> false in + let xh, xa = R.collect_app_ln x in + let yh, ya = R.collect_app_ln y in + go ((xh, R.Q_Explicit) :: xa) ((yh, R.Q_Explicit) :: ya) + +module TermEq = FStar.Reflection.TermEq + +let is_unamb (cands: list (int & slprop_view)) : bool = + match cands with + | [] | [_] -> true + | (_, x)::cands -> List.Tot.for_all (fun (_, y) -> TermEq.term_eq (elab_slprop x) (elab_slprop y)) cands + +// this matches atoms when they're the only unifiable pair +// necessary for various gather like functions where you don't specify all arguments +// needed for pcm_pts_to gather, where we need to specify the order of the ambiguous arguments +// needed for mask_mext, when we need to disambiguate mkey-ambiguous resources +let prove_atom_unamb (g: env) (ctxt: list slprop_view) (goal: slprop_view) : + T.Tac (option (prover_result g ctxt [goal])) = + match goal with + | Atom hd _ goal -> + let matches_mkeys (ctxt: slprop_view) : T.Tac bool = + match ctxt with + | Atom hd' _ ctxt -> + if hd <> hd' then false else + with_uf_transaction (fun _ -> + teq_nosmt_force_args (elab_env g) ctxt goal true + ) + | _ -> false in + let ictxt = List.Tot.mapi (fun i ctxt -> i, ctxt) ctxt in + let cands = T.filter (fun (i, ctxt) -> matches_mkeys ctxt) ictxt in + if Nil? cands then None else + if not (is_unamb cands) then None else + let (i, cand) :: _ = cands in + debug_prover g (fun _ -> Printf.sprintf "prove_atom_unamb: commiting to unify %s and %s\n" (show (elab_slprop cand)) (show goal)); + ignore (teq_nosmt_force_args (elab_env g) (elab_slprop cand) goal false); + let rest_ctxt = List.Tot.filter (fun (j, _) -> j <> i) ictxt |> List.Tot.map snd in + Some (| g, rest_ctxt, [], [cand], fun g' -> + let h2: slprop_equiv g' (elab_slprop cand) goal = check_slprop_equiv_ext (RU.range_of_term goal) _ _ _ in + let h1: slprop_equiv g (elab_slprops ctxt) (elab_slprops ([cand] @ rest_ctxt)) = RU.magic () in + let h2: slprop_equiv g' (elab_slprops ([cand] @ [])) goal = h2 in + cont_elab_refl _ _ _ h1, + cont_elab_refl _ _ _ h2 + <: T.Tac _ |) + | _ -> None + +let prove_atom (g: env) (ctxt: list slprop_view) (allow_amb: bool) (goal: slprop_view) : + T.Tac (option (prover_result g ctxt [goal])) = + match goal with + | Atom hd mkeys goal -> + let matches_mkeys (ctxt: slprop_view) : T.Tac bool = + match ctxt with + | Atom hd' mkeys' ctxt -> + if hd <> hd' then false else + with_uf_transaction (fun _ -> + match mkeys, mkeys' with + | Some mkeys, Some mkeys' -> + T.zip mkeys mkeys' |> forallb (fun (a, b) -> RU.teq_nosmt_force (elab_env g) a b) + | _, _ -> + teq_nosmt_force_args (elab_env g) ctxt goal true + ) + | _ -> false in + let ictxt = List.Tot.mapi (fun i ctxt -> i, ctxt) ctxt in + let cands = T.filter (fun (i, ctxt) -> matches_mkeys ctxt) ictxt in + if Nil? cands then None else + if (if allow_amb then false else not (is_unamb cands)) then None else + let (i, cand)::_ = cands in + debug_prover g (fun _ -> Printf.sprintf "commiting to unify %s and %s\n" (show (elab_slprop cand)) (show goal)); + ignore (teq_nosmt_force_args (elab_env g) (elab_slprop cand) goal false); + let rest_ctxt = List.Tot.filter (fun (j, _) -> j <> i) ictxt |> List.Tot.map snd in + Some (| g, rest_ctxt, [], [cand], fun g' -> + let h2: slprop_equiv g' (elab_slprop cand) goal = check_slprop_equiv_ext (RU.range_of_term goal) _ _ _ in + let h1: slprop_equiv g (elab_slprops ctxt) (elab_slprops ([cand] @ rest_ctxt)) = RU.magic () in + let h2: slprop_equiv g' (elab_slprops ([cand] @ [])) goal = h2 in + cont_elab_refl _ _ _ h1, + cont_elab_refl _ _ _ h2 + <: T.Tac _ + |) + | _ -> None + +let rec first_some #a (ks: list (unit -> T.Tac (option a))) : T.Tac (option a) = + match ks with + | [] -> None + | k::ks -> + match k () with + | Some res -> Some res + | None -> first_some ks + +let show_slprops ps = + String.concat "\n" (T.map (fun p -> show (elab_slprop p)) ps) + +let rec try_prove_core (g: env) (ctxt goals: list slprop_view) allow_amb : T.Tac (prover_result g ctxt goals) = + let noop () : prover_result g ctxt goals = + (| g, ctxt, goals, [], fun g'' -> + cont_elab_refl g _ _ (VE_Refl _ _), + cont_elab_refl g'' ([] @ goals) goals (VE_Refl _ _) + <: T.Tac _ |) in + match goals with + | [] -> noop () + | goals -> + debug_prover g (fun _ -> Printf.sprintf "proving\n%s\n from\n%s\n" (show_slprops goals) (show_slprops ctxt)); + let step : option (prover_result g ctxt goals) = + first_some [ + (fun _ -> elim_first g ctxt goals (unpack_and_norm_ctxt g)); + (fun _ -> elim_first g ctxt goals (elim_pure_step g)); + (fun _ -> elim_first g ctxt goals (elim_with_pure_step g)); + (fun _ -> elim_first g ctxt goals (elim_exists_step g)); + (fun _ -> prove_first g ctxt goals (prove_pure g ctxt true)); + (fun _ -> prove_first g ctxt goals (prove_with_pure g ctxt true)); + (fun _ -> prove_first g ctxt goals (prove_exists g ctxt)); + (fun _ -> prove_first g ctxt goals (unpack_and_norm_goal g ctxt)); + (fun _ -> prove_first g ctxt goals (prove_atom_unamb g ctxt)); + (fun _ -> prove_first g ctxt goals (prove_atom g ctxt allow_amb)); + (fun _ -> prove_first g ctxt goals (prove_pure g ctxt false)); + (fun _ -> prove_first g ctxt goals (prove_with_pure g ctxt false)); + ] in + match step with + | Some step -> + let (| g1, ctxt1, goals1, solved1, k1 |) = step in + let step2 = try_prove_core g1 ctxt1 goals1 allow_amb in + prover_result_join step step2 + | None -> noop () + +let try_prove (g: env) (ctxt goals: slprop) allow_amb : T.Tac (prover_result g [Unknown ctxt] [Unknown goals]) = + let ctxt = RU.deep_compress_safe ctxt in + let goals = RU.deep_compress_safe goals in + let ss = Pulse.Checker.Prover.RewritesTo.get_subst_from_env g in + let ctxt' = Pulse.Checker.Prover.Substs.ss_term ctxt ss in + let goals' = Pulse.Checker.Prover.Substs.ss_term goals ss in + let (| g1, ctxt1, goals1, solved1, k1 |) = try_prove_core g [Unknown ctxt'] [Unknown goals'] allow_amb in + (| g1, ctxt1, goals1, solved1, fun g2 -> + let before, after = k1 g2 in + let h1: slprop_equiv g ctxt' ctxt = RU.magic () in + let h2: slprop_equiv g2 goals' goals = RU.magic () in + cont_elab_equiv before h1 (VE_Refl _ _), + cont_elab_equiv after (VE_Refl _ _) h2 |) + +let pp_slprops ts = ts + |> List.Tot.map elab_slprop + |> sort_terms + |> T.map pp + |> separate hardline + +let prove rng (g: env) (ctxt goals: slprop) allow_amb : + T.Tac (g':env { env_extends g' g } & + ctxt': slprop & + continuation_elaborator g ctxt g' (goals `tm_star` ctxt')) = + let (| g', ctxt', goals', solved', k |) = try_prove g ctxt goals allow_amb in + if Cons? goals' then + T.fail_doc_at ([ + text (if List.length goals' > 1 then "Cannot prove any of:" else "Cannot prove:") ^^ + indent (pp_slprops goals'); + text "In the context:" ^^ indent (pp_slprops ctxt'); + ] @ (if Pulse.Config.debug_flag "initial_solver_state" then [ + text "The prover was started with goal:" ^^ indent (pp goals); + text "and initial context:" ^^ indent (pp ctxt); + ] else [])) + (Some rng) + else + let before, after = k g' in + let h: slprop_equiv g' (elab_slprops (solved' @ ctxt')) (elab_slprops (ctxt' @ solved' @ goals')) = RU.magic () in + let k = cont_elab_trans before (cont_elab_frame after ctxt') h [] in + let h: slprop_equiv g' (elab_slprops (ctxt' @ [Unknown goals])) (tm_star goals (elab_slprops ctxt')) = RU.magic () in + (| g', RU.deep_compress_safe (elab_slprops ctxt'), k_elab_equiv k (VE_Refl _ _) h |) -#push-options "--z3rlimit_factor 4" -#restart-solver -let prove_post_hint (#g:env) (#ctxt:slprop) - (r:checker_result_t g ctxt NoHint) - (post_hint:post_hint_opt g) - (rng:range) - +let prove_post_hint (#g:env) (#ctxt:slprop) (r:checker_result_t g ctxt NoHint) (post_hint:post_hint_opt g) (rng:range) : T.Tac (checker_result_t g ctxt post_hint) = let g = push_context g "prove_post_hint" rng in @@ -716,57 +714,83 @@ let prove_post_hint (#g:env) (#ctxt:slprop) let post_hint_opened = open_term_nv post_hint.post (ppname, x) in // TODO: subtyping - if not (eq_tm ty post_hint.ret_ty) + if not (eq_tm (RU.deep_compress_safe ty) (RU.deep_compress_safe post_hint.ret_ty)) then fail_doc g (Some rng) [ - text "The return type" ^^ - indent (pp ty) ^/^ - text "does not match the expected" ^^ - indent (pp post_hint.ret_ty) + text "The return type" ^^ indent (pp ty) ^/^ + text "does not match the expected" ^^ indent (pp post_hint.ret_ty) ] else if eq_tm post_hint_opened ctxt' then (| x, g2, (| u_ty, ty, ty_typing |), (| ctxt', ctxt'_typing |), k |) else - let (| g3, nts, _, remaining_ctxt, k_post |) = - prove false #g2 #ctxt' ctxt'_typing (mk_env (fstar_env g2)) #post_hint_opened (RU.magic ()) in - - assert (nts == []); - let k_post - : continuation_elaborator g2 ctxt' g3 (post_hint_opened * remaining_ctxt) = - coerce_eq k_post () in - - match check_equiv_emp' g3 remaining_ctxt with - | None -> + let (| g3, remaining_ctxt, k_post |) = + prove rng g2 ctxt' post_hint_opened false in + + let tv = inspect_term remaining_ctxt in + if not (Tm_Emp? tv) then fail_doc g (Some rng) [ prefix 2 1 (text "Leftover resources:") - (align (P.term_to_doc remaining_ctxt)); - (let tv = inspect_term remaining_ctxt in - if Tm_Star? tv + (align (pp remaining_ctxt)); + (if Tm_Star? tv then text "Did you forget to free these resources?" else text "Did you forget to free this resource?"); ] - | Some d -> - let k_post - : continuation_elaborator g2 ctxt' g3 post_hint_opened = - k_elab_equiv k_post (VE_Refl _ _) (RU.magic #(slprop_equiv _ _ _) ()) in - // - // for the typing of ty in g3, - // we have typing of ty in g2 above, and g3 `env_extends` g2 - // - // - // for the typing of post_hint_opened, - // again post_hint is well-typed in g, and g3 `env_extends` g - // - (| x, g3, (| u_ty, ty, RU.magic #(tot_typing _ _ _) () |), - (| post_hint_opened, RU.magic #(tot_typing _ _ _) () |), k_elab_trans k k_post |) + else + let h3: slprop_equiv g3 (tm_star post_hint_opened remaining_ctxt) post_hint_opened = RU.magic () in + // for the typing of ty in g3, we have typing of ty in g2 above, and g3 `env_extends` g2 + let h1: universe_of g3 ty u_ty = RU.magic () in + // for the typing of post_hint_opened, again post_hint is well-typed in g, and g3 `env_extends` g + let h2: tot_typing g3 post_hint_opened tm_slprop = RU.magic () in + (| x, g3, (| u_ty, ty, h1 |), (| post_hint_opened, h2 |), + k_elab_trans k (k_elab_equiv k_post (VE_Refl _ _) h3) |) + +let try_frame_pre (allow_ambiguous : bool) (#g:env) + (#ctxt:slprop) (ctxt_typing:tot_typing g ctxt tm_slprop) + (d:(t:st_term & c:comp_st & st_typing g t c)) + (res_ppname:ppname) : + T.Tac (checker_result_t g ctxt NoHint) = + let (| t, c, d |) = d in + let (| g', ctxt', k |) = prove t.range g ctxt (comp_pre c) allow_ambiguous in + let d: st_typing g' t c = RU.magic () in // weakening from g to g' + let h1: tot_typing g' ctxt' tm_slprop = RU.magic() in // weakening from to g' + checker_result_for_st_typing (k _ (| t, add_frame c ctxt', T_Frame _ _ _ ctxt' h1 d |)) res_ppname + +let rec try_elim_core (g: env) (ctxt: list slprop_view) : + T.Tac (prover_result_nogoals g ctxt) = + let noop () : prover_result g ctxt [] = + (| g, ctxt, [], [], fun g'' -> + cont_elab_refl g _ _ (VE_Refl _ _), + cont_elab_refl g'' [] [] (VE_Refl _ _) + <: T.Tac _ |) in + debug_prover g (fun _ -> Printf.sprintf "eliminating\n%s\n" (show_slprops ctxt)); + let step : option (prover_result_nogoals g ctxt) = + first_some [ + (fun _ -> elim_first' g ctxt [] (unpack_and_norm_ctxt g)); + (fun _ -> elim_first' g ctxt [] (elim_pure_step g)); + (fun _ -> elim_first' g ctxt [] (elim_with_pure_step g)); + (fun _ -> elim_first' g ctxt [] (elim_exists_step g)); + ] in + match step with + | Some step -> + let (| g1, ctxt1, goals1, solved1, k1 |) = step in + let step2 = try_elim_core g1 ctxt1 in + prover_result_join step step2 + | None -> noop () let elim_exists_and_pure (#g:env) (#ctxt:slprop) - (ctxt_typing:tot_typing g ctxt tm_slprop) - : T.Tac (g':env { env_extends g' g } & - ctxt':term & - tot_typing g' ctxt' tm_slprop & - continuation_elaborator g ctxt g' ctxt') = - - let (| g', _nts, _labels, ctxt', k |) = prove false #g #ctxt ctxt_typing (mk_env (fstar_env g)) emp_typing in - (| g', ctxt', E (RU.magic ()), k_elab_equiv k (VE_Refl _ _) (VE_Unit _ _) |) -#pop-options \ No newline at end of file + (ctxt_typing:tot_typing g ctxt tm_slprop) + : T.Tac (g':env { env_extends g' g } & + ctxt':term & + tot_typing g' ctxt' tm_slprop & + continuation_elaborator g ctxt g' ctxt') = + let ss = Pulse.Checker.Prover.RewritesTo.get_subst_from_env g in + let ctxt' = Pulse.Checker.Prover.Substs.ss_term ctxt ss in + let (| g', ctxt'', goals'', solved, k |) = try_elim_core g [Unknown ctxt'] in + let h: tot_typing g' (elab_slprops ctxt'') tm_slprop = RU.magic () in // TODO thread through prover + let h1: slprop_equiv g (elab_slprops ([] @ [Unknown ctxt'])) ctxt = (RU.magic() <: slprop_equiv g ctxt' ctxt) in + let h2: slprop_equiv g' (elab_slprops (ctxt'' @ solved @ goals'')) (elab_slprops ([] @ solved @ ctxt'')) = RU.magic () in + let h3: slprop_equiv g' (elab_slprops (ctxt'' @ [])) (elab_slprops ctxt'') = RU.magic () in + let before, after = k g' in + (| g', elab_slprops ctxt'', h, + k_elab_trans (k_elab_equiv (before []) h1 (VE_Refl _ _)) + (k_elab_equiv (after ctxt'') h2 h3) |) diff --git a/src/checker/Pulse.Checker.Prover.fsti b/src/checker/Pulse.Checker.Prover.fsti index 9f6dfc475..4c6f68319 100644 --- a/src/checker/Pulse.Checker.Prover.fsti +++ b/src/checker/Pulse.Checker.Prover.fsti @@ -1,5 +1,5 @@ (* - Copyright 2023 Microsoft Research + Copyright 2025 Microsoft Research Licensed under the Apache License, Version 2.0 (the "License"); you may not use this file except in compliance with the License. @@ -15,67 +15,34 @@ *) module Pulse.Checker.Prover - -module T = FStar.Tactics.V2 - -open Pulse.Syntax -open Pulse.Typing open Pulse.Checker.Base +open Pulse.Typing +open Pulse.Syntax.Base +open Pulse.Syntax.Pure +open Pulse.Syntax.Naming +module T = FStar.Tactics.V2 -module PS = Pulse.Checker.Prover.Substs - -include Pulse.Checker.Prover.Base -include Pulse.Checker.Prover.Util - -val normalize_slprop - (g:env) - (v:slprop) - (use_rewrites_to : bool) - : T.Tac (v':slprop & slprop_equiv g v v') - -val normalize_slprop_welltyped - (g:env) - (v:slprop) - (v_typing:tot_typing g v tm_slprop) - : T.Tac (v':slprop & slprop_equiv g v v' & tot_typing g v' tm_slprop) - -val prove - (allow_ambiguous : bool) - (#g:env) (#ctxt:slprop) (ctxt_typing:tot_typing g ctxt tm_slprop) - (uvs:env { disjoint g uvs }) - (#goals:slprop) (goals_typing:tot_typing (push_env g uvs) goals tm_slprop) - - : T.Tac (g1 : env { g1 `env_extends` g /\ disjoint g1 uvs } & - nts : PS.nt_substs & - effect_labels:list T.tot_or_ghost { PS.well_typed_nt_substs g1 uvs nts effect_labels } & - remaining_ctxt : slprop & - continuation_elaborator g ctxt g1 ((PS.nt_subst_term goals nts) * remaining_ctxt)) +val elim_exists (g: env) (frame: slprop) u b body (x: nvar { ~(Set.mem (snd x) (dom g)) }) + (g': env { g' == push_binding g (snd x) (fst x) (mk_erased u b.binder_ty) }) : + continuation_elaborator g (frame `tm_star` tm_exists_sl u b body) + g' (frame `tm_star` open_term' body (mk_reveal u b.binder_ty (term_of_nvar x)) 0) -val try_frame_pre_uvs - (allow_ambiguous : bool) - (#g:env) (#ctxt:slprop) (ctxt_typing:tot_typing g ctxt tm_slprop) - (uvs:env { disjoint g uvs }) - (d:(t:st_term & c:comp_st & st_typing (push_env g uvs) t c)) - (res_ppname:ppname) - : T.Tac (checker_result_t g ctxt NoHint) +val prove (rng: range) (g: env) (ctxt goals: slprop) (allow_amb: bool) : + T.Tac (g':env { env_extends g' g } & ctxt': slprop & + continuation_elaborator g ctxt g' (goals `tm_star` ctxt')) -val try_frame_pre - (allow_ambiguous : bool) - (#g:env) (#ctxt:slprop) (ctxt_typing:tot_typing g ctxt tm_slprop) - (d:(t:st_term & c:comp_st & st_typing g t c)) - (res_ppname:ppname) - : T.Tac (checker_result_t g ctxt NoHint) +val prove_post_hint (#g:env) (#ctxt:slprop) (r:checker_result_t g ctxt NoHint) (post_hint:post_hint_opt g) (rng:range) : + T.Tac (checker_result_t g ctxt post_hint) -val prove_post_hint (#g:env) (#ctxt:slprop) - (r:checker_result_t g ctxt NoHint) - (post_hint:post_hint_opt g) - (rng:range) - - : T.Tac (checker_result_t g ctxt post_hint) +val try_frame_pre (allow_ambiguous : bool) (#g:env) + (#ctxt:slprop) (ctxt_typing:tot_typing g ctxt tm_slprop) + (d:(t:st_term & c:comp_st & st_typing g t c)) + (res_ppname:ppname) : + T.Tac (checker_result_t g ctxt NoHint) val elim_exists_and_pure (#g:env) (#ctxt:slprop) - (ctxt_typing:tot_typing g ctxt tm_slprop) - : T.Tac (g':env { env_extends g' g } & - ctxt':term & - tot_typing g' ctxt' tm_slprop & - continuation_elaborator g ctxt g' ctxt') \ No newline at end of file + (ctxt_typing:tot_typing g ctxt tm_slprop) + : T.Tac (g':env { env_extends g' g } & + ctxt':term & + tot_typing g' ctxt' tm_slprop & + continuation_elaborator g ctxt g' ctxt') \ No newline at end of file diff --git a/src/checker/Pulse.Checker.Pure.fst b/src/checker/Pulse.Checker.Pure.fst index 7023523b3..f38fe100d 100644 --- a/src/checker/Pulse.Checker.Pure.fst +++ b/src/checker/Pulse.Checker.Pure.fst @@ -418,6 +418,7 @@ let tc_term_phase1 g (t:term) : T.Tac (term & term & T.tot_or_ghost) = let fg = elab_env g in let t = RU.deep_transform_to_unary_applications t in let instantiate_imps = true in + debug g (fun _ -> "tc_term_phase1: " ^ show t); let res, issues = catch_all fun _ -> RU.with_context (RU.extend_context "tc_term_phase1" (Some (range_of_term t)) (get_context g)) fun _ -> RU.tc_term_phase1 fg t instantiate_imps in @@ -465,7 +466,7 @@ let core_compute_term_type (g:env) (t:term) in RU.record_stats "Pulse.core_compute_term_type" aux -let core_check_term g e eff t +let core_check_term' g e eff t extra_msg = let aux () : T.Tac (typing g e eff t) = let fg = elab_env g in let topt, issues = @@ -475,11 +476,15 @@ let core_check_term g e eff t fg e eff t) in match topt with | None -> - fail_doc_with_subissues g (Some <| RU.range_of_term e) issues (ill_typed_term e (Some t) None) + fail_doc_with_subissues g (Some <| RU.range_of_term e) issues (extra_msg () @ ill_typed_term e (Some t) None) | Some tok -> E (RT.T_Token _ _ _ (FStar.Squash.return_squash tok)) in RU.record_stats "Pulse.core_check_term" aux + +let core_check_term g e eff t += core_check_term' g e eff t fun _ -> [] + let core_check_term_at_type g e t = let aux () : T.Tac (eff:T.tot_or_ghost & typing g e eff t) = let fg = elab_env g in diff --git a/src/checker/Pulse.Checker.Pure.fsti b/src/checker/Pulse.Checker.Pure.fsti index bd779fc68..ac591fe31 100644 --- a/src/checker/Pulse.Checker.Pure.fsti +++ b/src/checker/Pulse.Checker.Pure.fsti @@ -65,6 +65,10 @@ val core_compute_term_type (g:env) (t:term) ty:term & typing g t eff ty) +val core_check_term' (g:env) (e:term) (eff:T.tot_or_ghost) (t:term) + (extra_msg: unit -> T.Tac (list Pprint.document)) + : T.Tac (typing g e eff t) + val core_check_term (g:env) (e:term) (eff:T.tot_or_ghost) (t:term) : T.Tac (typing g e eff t) diff --git a/src/checker/Pulse.Checker.Return.fst b/src/checker/Pulse.Checker.Return.fst index 68efab5d4..c74cf1722 100644 --- a/src/checker/Pulse.Checker.Return.fst +++ b/src/checker/Pulse.Checker.Return.fst @@ -137,15 +137,14 @@ let check_core assume (open_term (close_term post_opened x) x == post_opened); let post = close_term post_opened x in let d = T_Return g c use_eq u ty t post x uty d post_typing in - let dd = (match_comp_res_with_post_hint d post_hint) in + let (|c',d'|) = match_comp_res_with_post_hint d post_hint in Pulse.Checker.Util.debug g "pulse.return" (fun _ -> - let (| _, c, _ |) = dd in Printf.sprintf "Return comp is: %s" - (Pulse.Syntax.Printer.comp_to_string c)); + (Pulse.Syntax.Printer.comp_to_string c')); prove_post_hint #g - (try_frame_pre false #g ctxt_typing dd res_ppname) + (try_frame_pre false #g ctxt_typing (|_,c',d'|) res_ppname) post_hint - (Pulse.RuntimeUtils.range_of_term t) + st.range #pop-options let check diff --git a/src/checker/Pulse.Checker.Rewrite.fst b/src/checker/Pulse.Checker.Rewrite.fst index e2752fd41..3e9b04964 100644 --- a/src/checker/Pulse.Checker.Rewrite.fst +++ b/src/checker/Pulse.Checker.Rewrite.fst @@ -141,4 +141,5 @@ let check "Pulse.Checker.Rewrite.check_slprop_equiv_tac" in let d = T_Rewrite _ p q p_typing equiv_p_q in - prove_post_hint (try_frame_pre false pre_typing (match_comp_res_with_post_hint d post_hint) res_ppname) post_hint t.range + let (| c,d |) = match_comp_res_with_post_hint d post_hint in + prove_post_hint (try_frame_pre false pre_typing (| _,c,d |) res_ppname) post_hint t.range diff --git a/src/checker/Pulse.Checker.ST.fst b/src/checker/Pulse.Checker.ST.fst index 2b22a0e64..6c13f6d39 100644 --- a/src/checker/Pulse.Checker.ST.fst +++ b/src/checker/Pulse.Checker.ST.fst @@ -26,66 +26,33 @@ module RT = FStar.Reflection.Typing module R = FStar.Reflection.V2 module RU = Pulse.RuntimeUtils module P = Pulse.Syntax.Printer -module Prover = Pulse.Checker.Prover +open Pulse.Checker.Prover open Pulse.Show let should_allow_ambiguous (t:term) : T.Tac bool = Pulse.Reflection.Util.head_has_attr_string "Pulse.Lib.Core.allow_ambiguous" t -// HACK: Instantiates implicits at the end with uvs -// If we let instantiate_term_implicits_uvs instantiate these implicits, -// then it will instantiate squash arguments with (). -// The () will usually only type-check after running the prover, -// however the prover requires a type-correct term. -// Therefore we manually instantiate these implicits... -let instantiate_term_implicits_uvs (g:env) (t:term) - : T.Tac (uvs:env { disjoint g uvs } & term & term) = - let (| uvs, t, ty |) = Pulse.Checker.Pure.instantiate_term_implicits_uvs g t false in - let rec add_implicits (uvs: env { disjoint g uvs }) (t ty:term) - : T.Tac (uvs:env { disjoint g uvs } & term & term) = - match inspect_ln_unascribe ty with - | R.Tv_Arrow b c -> - (match R.inspect_comp c with - | C_Total ty -> - let b = R.inspect_binder b in - if Q_Implicit? b.qual then - let x = fresh uvs in - let px = ppname_default, x in - let uvs = push_binding uvs x (fst px) b.sort in - assume disjoint g uvs; - let ty = open_term_nv ty px in - add_implicits uvs (pack_ln (R.Tv_App t (term_of_nvar px, b.qual))) ty - else - (| uvs, t, ty |) - | _ -> (| uvs, t, ty |)) - | _ -> (| uvs, t, ty |) in - add_implicits uvs t ty - open Pulse.PP #push-options "--fuel 0 --ifuel 1 --z3rlimit_factor 2" let check - (g0:env) + (g:env) (ctxt:slprop) - (ctxt_typing:tot_typing g0 ctxt tm_slprop) - (post_hint:post_hint_opt g0) + (ctxt_typing:tot_typing g ctxt tm_slprop) + (post_hint:post_hint_opt g) (res_ppname:ppname) (t:st_term { Tm_ST? t.term }) - : T.Tac (checker_result_t g0 ctxt post_hint) = + : T.Tac (checker_result_t g ctxt post_hint) = - let g = push_context "st" t.range g0 in + let g = push_context "st" t.range g in let post_hint: post_hint_opt g = post_hint in let range = t.range in let Tm_ST { t=e; args } = t.term in if Cons? args then - fail_doc g0 (Some t.range) [ + fail_doc g (Some t.range) [ text "Internal error: trailing combinator arguments not lifted in Tm_ST"; pp t ]; - let (| uvs, e, _ |) = instantiate_term_implicits_uvs g e in - let g' : env = push_env g uvs in - assert (g' `env_extends` g); - let post_hint: post_hint_opt g' = post_hint in - let (| e, eff, ty, typing |) = Pulse.Checker.Pure.compute_term_type g' e in + let e, ty, eff = tc_term_phase1 g e in match Pulse.Readback.readback_comp ty with | None -> fail g (Some range) (Printf.sprintf "readback of %s failed" (show ty)) | Some (C_Tot _) -> @@ -100,36 +67,45 @@ let check text "is a total term"; text "Maybe it is not fully applied?"] - | Some c -> ( + | Some c0 -> ( let allow_ambiguous = should_allow_ambiguous e in + let (| g', ctxt', k |) = prove (RU.range_of_term e) g ctxt (comp_pre c0) allow_ambiguous in + + // remove spurious beta-redexes + let e = e |> RU.beta_lax (elab_env g) |> RU.deep_compress in + let ty = ty |> RU.beta_lax (elab_env g) |> RU.deep_compress in + assume elab_comp c0 == ty; + let Some c = Pulse.Readback.readback_comp ty in + + let (| eff, typing |) = core_check_term_at_type g' e ty in let t = { t with term = Tm_ST { t=e; args=[] } } in - let d : ( st_typing g' t c ) = - if eff = T.E_Total - then T_ST g' e c typing - else ( - match c with - | C_ST _ | C_STAtomic .. -> - let open Pulse.PP in - fail_doc g0 (Some range) - [text "Application of a stateful or atomic computation cannot have a ghost effect"; - pp t; - text "has computation type"; - pp c] - | C_STGhost .. -> - let d_non_info : non_informative g' c = - let token = is_non_informative g' c in - match token with - | None -> - fail g (Some range) - (Printf.sprintf "Unexpected informative result for %s" (P.comp_to_string c)) - | Some token -> - E <| RT.Non_informative_token _ _ (FStar.Squash.return_squash token) - in - T_STGhost g' e c typing d_non_info - ) - in - let (| st', c', st_typing' |) = match_comp_res_with_post_hint d post_hint in - let framed = - RU.record_stats "try_frame_pre_uvs" fun _ -> Prover.try_frame_pre_uvs allow_ambiguous ctxt_typing uvs (| st', c', st_typing' |) res_ppname in - RU.record_stats "prove_post_hint" fun _ -> Prover.prove_post_hint framed post_hint range + let d : st_typing g' t c = + if eff = T.E_Total + then T_ST g' e c typing + else ( + match c with + | C_ST _ | C_STAtomic .. -> + let open Pulse.PP in + fail_doc g (Some range) + [text "Application of a stateful or atomic computation cannot have a ghost effect"; + pp t; + text "has computation type"; + pp c] + | C_STGhost .. -> + let d_non_info : non_informative g' c = + let token = is_non_informative g' c in + match token with + | None -> + fail g' (Some range) + (Printf.sprintf "Unexpected informative result for %s" (P.comp_to_string c)) + | Some token -> + E <| RT.Non_informative_token _ _ (FStar.Squash.return_squash token) + in + T_STGhost g' e c typing d_non_info + ) + in + let (| c, d |) = match_comp_res_with_post_hint d post_hint in + let h: tot_typing g' ctxt' tm_slprop = RU.magic () in // TODO: thread through prover + let framed = checker_result_for_st_typing (k _ (| t, add_frame c ctxt', T_Frame _ _ _ ctxt' h d |)) res_ppname in + RU.record_stats "prove_post_hint" fun _ -> prove_post_hint framed post_hint range ) \ No newline at end of file diff --git a/src/checker/Pulse.Checker.Unreachable.fst b/src/checker/Pulse.Checker.Unreachable.fst index 0a98e86a9..09c578abb 100644 --- a/src/checker/Pulse.Checker.Unreachable.fst +++ b/src/checker/Pulse.Checker.Unreachable.fst @@ -22,7 +22,6 @@ open Pulse.Syntax open Pulse.Typing open Pulse.Checker.Pure open Pulse.Checker.Base -open Pulse.Checker.Prover let check diff --git a/src/checker/Pulse.Checker.Unreachable.fsti b/src/checker/Pulse.Checker.Unreachable.fsti index d67485353..41c7312b7 100644 --- a/src/checker/Pulse.Checker.Unreachable.fsti +++ b/src/checker/Pulse.Checker.Unreachable.fsti @@ -22,7 +22,6 @@ open Pulse.Syntax open Pulse.Typing open Pulse.Checker.Pure open Pulse.Checker.Base -open Pulse.Checker.Prover val check diff --git a/src/checker/Pulse.Checker.While.fst b/src/checker/Pulse.Checker.While.fst index af9892c9a..482cb6cf0 100644 --- a/src/checker/Pulse.Checker.While.fst +++ b/src/checker/Pulse.Checker.While.fst @@ -116,7 +116,8 @@ let check if eq_comp body_comp (comp_while_body nm inv) then let d = T_While g inv cond body inv_typing cond_typing body_typing in - prove_post_hint (try_frame_pre false pre_typing (match_comp_res_with_post_hint d post_hint) res_ppname) post_hint t.range + let (| c,d |) = match_comp_res_with_post_hint d post_hint in + prove_post_hint (try_frame_pre false pre_typing (|_,c,d|) res_ppname) post_hint t.range else fail g None (Printf.sprintf "Could not prove the inferred type of the while body matches the annotation\n\ Inferred type = %s\n\ @@ -173,8 +174,8 @@ let check_nuwhile { ctxt_now = pre; ctxt_old = Some pre } inv in - let (| g1, nts, labs, remaining, k |) = Pulse.Checker.Prover.prove false pre_typing (empty_env g) inv_typing in - let inv = tm_star (Pulse.Checker.Prover.Substs.nt_subst_term inv nts) remaining in + let (| g1, remaining, k |) = Pulse.Checker.Prover.prove t.range g pre inv false in + let inv = tm_star (RU.deep_compress_safe inv) remaining in let inv_typing : tot_typing g1 inv tm_slprop = RU.magic () in let (| post_cond, r_cond |) : (ph:post_hint_for_env g1 & checker_result_t g1 inv (PostHint ph)) = let r_cond = check (push_context "check_while_condition" cond.range g1) inv inv_typing (TypeHint tm_bool) ppname_default cond in diff --git a/src/checker/Pulse.Checker.WithInv.fst b/src/checker/Pulse.Checker.WithInv.fst index a519215f2..e73b26d28 100644 --- a/src/checker/Pulse.Checker.WithInv.fst +++ b/src/checker/Pulse.Checker.WithInv.fst @@ -20,7 +20,7 @@ open Pulse.Syntax open Pulse.Typing open Pulse.Checker.Pure open Pulse.Checker.Base -open Pulse.Checker.Prover +open Pulse.Checker.Prover.Normalize open Pulse.Checker.Comp open Pulse.Show open FStar.Pprint @@ -232,8 +232,8 @@ let withinv_post (#g:env) (#p:term) (#i:term) (#post:term) : T.Tac (option (post':term & tot_typing g post' tm_slprop)) = - let (| p, _, p_typing |) = Prover.normalize_slprop_welltyped g p p_typing in - let (| post, _, post_typing |) = Prover.normalize_slprop_welltyped g post post_typing in + let (| p, _, p_typing |) = normalize_slprop_welltyped g p p_typing in + let (| post, _, post_typing |) = normalize_slprop_welltyped g post post_typing in __withinv_post #g #p #i #post p_typing i_typing post_typing #push-options "--fuel 0 --ifuel 0 --query_stats" @@ -497,7 +497,7 @@ let norm_and_check (t:st_term{Tm_WithInv? t.term}) (check:check_t) : T.Tac (checker_result_t g pre post_hint) -= let (| pre', pre_equiv, pre'_typing |) = Prover.normalize_slprop_welltyped g pre pre_typing in += let (| pre', pre_equiv, pre'_typing |) = normalize_slprop_welltyped g pre pre_typing in let r = check0 g pre' pre'_typing post_hint res_ppname t check in checker_result_t_equiv_ctxt _ _ _ _ (VE_Sym _ _ _ pre_equiv) r diff --git a/src/checker/Pulse.Checker.WithLocal.fst b/src/checker/Pulse.Checker.WithLocal.fst index 8b34a657d..55ed64074 100644 --- a/src/checker/Pulse.Checker.WithLocal.fst +++ b/src/checker/Pulse.Checker.WithLocal.fst @@ -100,9 +100,7 @@ let check else let x = fresh g in let px = binder.binder_ppname, x in - if x `Set.mem` freevars_st body - then fail g (Some body.range) (Printf.sprintf "withlocal: %s is free in body" (T.unseal binder.binder_ppname.name)) - else + assume not (x `Set.mem` freevars_st body); let x_tm = term_of_nvar px in let g_extended = push_binding g x binder.binder_ppname (mk_ref init_t) in let body_pre = comp_withlocal_body_pre pre init_t x_tm init in @@ -110,10 +108,7 @@ let check // elaborating this post here, // so that later we can check the computed post to be equal to this one let post : post_hint_for_env g = post in - if x `Set.mem` freevars post.post - then fail g None "Impossible! check_withlocal: unexpected name clash in with_local,\ - please file a bug-report" - else + assume not (x `Set.mem` freevars post.post); let open Pulse.Typing.Combinators in let body_post : post_hint_for_env g_extended = extend_post_hint_for_local g post init_t x in let r = check g_extended body_pre body_pre_typing (PostHint body_post) binder.binder_ppname (open_st_term_nv body px) in diff --git a/src/checker/Pulse.JoinComp.fst b/src/checker/Pulse.JoinComp.fst index a71084f67..137ba1db4 100644 --- a/src/checker/Pulse.JoinComp.fst +++ b/src/checker/Pulse.JoinComp.fst @@ -22,6 +22,7 @@ open Pulse.Typing.Combinators open Pulse.Checker.Pure open Pulse.Checker.Base open Pulse.Checker.Prover +open Pulse.Checker.Prover.Normalize open FStar.List.Tot open Pulse.Show module T = FStar.Tactics.V2 @@ -221,7 +222,7 @@ let join_post #g #hyp #b (T.term_to_string p1.post) (T.term_to_string p2.post) ); - if not (T.term_eq p1.ret_ty p2.ret_ty) + if not (T.term_eq (RU.deep_compress_safe p1.ret_ty) (RU.deep_compress_safe p2.ret_ty)) then ( fail_doc g (Some (T.range_of_term p1.ret_ty)) Pulse.PP.( @@ -230,19 +231,12 @@ let join_post #g #hyp #b ) ); let x = fresh g in + let g' = push_binding_def g x p1.ret_ty in let p1_post = open_term_nv p1.post (ppname_default, x) in - let (| p1_post, _ |) = - Pulse.Checker.Prover.normalize_slprop - (push_binding_def g x p1.ret_ty) - p1_post true - in + let (| p1_post, _ |) = normalize_slprop g' p1_post true in let p2_post = open_term_nv p2.post (ppname_default, x) in - let (| p2_post, _ |) = - Pulse.Checker.Prover.normalize_slprop - (push_binding_def g x p1.ret_ty) - p2_post true - in - let joined_post = join_slprop g b [] [] p1_post p2_post in + let (| p2_post, _ |) = normalize_slprop g' p2_post true in + let joined_post = join_slprop g' b [] [] p1_post p2_post in let joined_post = close_term joined_post x in Pulse.Checker.Util.debug g "pulse.join_comp" (fun _ -> Printf.sprintf "Inferred joint postcondition:\n%s\n" @@ -251,7 +245,6 @@ let join_post #g #hyp #b assume (fresh_wrt x g (freevars joined_post)); let (| u, ty_typing |) = Pulse.Checker.Pure.check_universe g p1.ret_ty in let joined_post' = open_term_nv joined_post (ppname_default, x) in - let g' = push_binding g x ppname_default p1.ret_ty in let post_typing_src = Pulse.Checker.Pure.check_slprop_with_core g' joined_post' in let (| eff, eff_ty |) = join_effect_annot g p1.effect_annot p2.effect_annot in let res : post_hint_for_env g = diff --git a/src/checker/Pulse.Main.fst b/src/checker/Pulse.Main.fst index b0680026e..53063daef 100644 --- a/src/checker/Pulse.Main.fst +++ b/src/checker/Pulse.Main.fst @@ -72,7 +72,7 @@ let check_fndefn let (| body, c, t_typing |) = Pulse.Checker.Abs.check_abs g body Pulse.Checker.check in - Pulse.Checker.Prover.debug_prover g + Pulse.Checker.Prover.Util.debug_prover g (fun _ -> Printf.sprintf "\ncheck call returned in main with:\n%s\nat type %s\n" (P.st_term_to_string body) (P.comp_to_string c)); diff --git a/src/checker/Pulse.Reflection.Util.fst b/src/checker/Pulse.Reflection.Util.fst index de7afc13d..f8d5b03b1 100644 --- a/src/checker/Pulse.Reflection.Util.fst +++ b/src/checker/Pulse.Reflection.Util.fst @@ -73,6 +73,7 @@ let slprop_equiv_unfold_tm (s:string) = let t = R.pack_ln (R.Tv_App head (s, R.Q_Explicit)) in t let match_rewrite_tac_tm = R.pack_ln (R.Tv_FVar (R.pack_fv (mk_pulse_lib_core_lid "match_rewrite_tac"))) +let match_rename_tac_tm = R.pack_ln (R.Tv_FVar (R.pack_fv (mk_pulse_lib_core_lid "match_rename_tac"))) (* The "plicities" *) let ex t : R.argv = (t, R.Q_Explicit) diff --git a/src/checker/Pulse.RuntimeUtils.fsti b/src/checker/Pulse.RuntimeUtils.fsti index 6fceb8936..af9e8a474 100644 --- a/src/checker/Pulse.RuntimeUtils.fsti +++ b/src/checker/Pulse.RuntimeUtils.fsti @@ -47,6 +47,8 @@ module R = FStar.Reflection.V2 val push_univ_vars (g: env) (us: list R.univ_name) : g':env { g' == g } val debug_subst (s:list R.subst_elt) (t:T.term) (r1 r2:T.term): y:T.term{ y == r1 } val deep_compress (t:T.term) : r:T.term { t == r } +val deep_compress_safe (t:T.term) : r:T.term { t == r } // does not bail on uvars +val no_uvars_in_term (t:T.term) : Dv bool (* ***WARNING*** *) (* This function is natively implemented against the F* compiler libraries to transform terms to use unary applications. @@ -67,9 +69,10 @@ val float_one : FStar.Float.float val lax_check_term_with_unknown_universes (g:env) (t:T.term) : Dv (option T.term) module T = FStar.Tactics.V2 val tc_term_phase1 (g:env) (t:T.term) (instantiate_imps:bool) : Dv (option (T.term & T.term & T.tot_or_ghost) & T.issues) -val teq_nosmt_force (g:env) (ty1 ty2:T.term) : bool +val teq_nosmt_force (g:env) (ty1 ty2:T.term) : Dv bool val whnf_lax (g:env) (t:T.term) : T.term val hnf_lax (g:env) (t:T.term) : T.term +val beta_lax (g:env) (t:T.term) : T.term module RT = FStar.Reflection.Typing val norm_well_typed_term (#g:T.env) diff --git a/src/checker/Pulse.Syntax.Printer.fst b/src/checker/Pulse.Syntax.Printer.fst index b21c452da..780ed695a 100644 --- a/src/checker/Pulse.Syntax.Printer.fst +++ b/src/checker/Pulse.Syntax.Printer.fst @@ -377,10 +377,11 @@ let rec st_term_to_string' (level:string) (t:st_term) (st_term_to_string' level body2) (term_to_string post2) - | Tm_Rewrite { t1; t2 } -> - sprintf "rewrite %s as %s" + | Tm_Rewrite { t1; t2; tac_opt } -> + sprintf "rewrite %s as %s (with %s)" (term_to_string t1) (term_to_string t2) + (match tac_opt with | None -> "no tactic" | Some tac -> term_to_string tac) | Tm_WithLocal { binder; initializer; body } -> sprintf "let mut %s = %s;\n%s%s" @@ -426,12 +427,13 @@ let rec st_term_to_string' (level:string) (t:st_term) | ASSERT { p } -> "assert", term_to_string p | UNFOLD { names; p } -> sprintf "unfold%s" (names_to_string names), term_to_string p | FOLD { names; p } -> sprintf "fold%s" (names_to_string names), term_to_string p - | RENAME { pairs; goal } -> - sprintf "rewrite each %s" + | RENAME { pairs; goal; tac_opt } -> + sprintf "rewrite each %s (with %s)" (String.concat ", " (T.map (fun (x, y) -> sprintf "%s as %s" (term_to_string x) (term_to_string y)) - pairs)), + pairs)) + (match tac_opt with | None -> "no tactic" | Some tac -> term_to_string tac), (match goal with | None -> "" | Some t -> sprintf " in %s" (term_to_string t)) diff --git a/src/checker/Pulse.Typing.Combinators.fst b/src/checker/Pulse.Typing.Combinators.fst index fe1b5c265..32dd830d8 100644 --- a/src/checker/Pulse.Typing.Combinators.fst +++ b/src/checker/Pulse.Typing.Combinators.fst @@ -467,7 +467,7 @@ let bind_res_and_post_typing g c2 x post_hint let (| u, res_typing |) = check_universe g s2.res in if not (eq_univ u s2.u) then fail g None "Unexpected universe for result type" - else if x `Set.mem` freevars s2.post + else if x `Set.mem` freevars (RU.deep_compress_safe s2.post) then fail g None (Printf.sprintf "Bound variable %d escapes scope in postcondition %s" x (P.term_to_string s2.post)) else ( let y = x in //fresh g in diff --git a/src/ml/Pulse_RuntimeUtils.ml b/src/ml/Pulse_RuntimeUtils.ml index 37c458659..e8d0dd4d8 100644 --- a/src/ml/Pulse_RuntimeUtils.ml +++ b/src/ml/Pulse_RuntimeUtils.ml @@ -132,6 +132,8 @@ let deep_transform_to_unary_applications (t:S.term) = t let deep_compress (t:S.term) = FStarC_Syntax_Compress.deep_compress_uvars t +let deep_compress_safe (t:S.term) = FStarC_Syntax_Compress.deep_compress true true t +let no_uvars_in_term = FStarC_Tactics_V2_Basic.no_uvars_in_term let map_seal (t:'a) (f:'a -> 'b) : 'b = f t let float_one = FStarC_Util.float_of_string "1.0" module TcEnv = FStarC_TypeChecker_Env @@ -202,6 +204,9 @@ let whnf_lax (g:TcEnv.env) (t:S.term) : S.term = let hnf_lax (g:TcEnv.env) (t:S.term) : S.term = FStarC_TypeChecker_Normalize.normalize [TcEnv.Unascribe; TcEnv.Primops; TcEnv.HNF; TcEnv.UnfoldUntil S.delta_constant; TcEnv.Beta] g t +let beta_lax (g:TcEnv.env) (t:S.term) : S.term = + FStarC_TypeChecker_Normalize.normalize [TcEnv.Unascribe; TcEnv.Beta] g t + let norm_well_typed_term (g:TcEnv.env) (t:S.term) diff --git a/test/ArrayTests.fst b/test/ArrayTests.fst index e62aeb56f..dad516012 100644 --- a/test/ArrayTests.fst +++ b/test/ArrayTests.fst @@ -44,9 +44,9 @@ fn compare (#t:eqtype) (#p1 #p2:perm) (l:US.t) (#s1 #s2:elseq t l) (a1 a2:A.larr R.pts_to i vi ** A.pts_to a1 #p1 s1 ** A.pts_to a2 #p2 s2 ** + pure (b == (US.(vi <^ l) && Seq.index s1 (US.v vi) = Seq.index s2 (US.v vi))) ** pure ( US.v vi <= US.v l /\ - (b == (US.(vi <^ l) && Seq.index s1 (US.v vi) = Seq.index s2 (US.v vi))) /\ (forall (i:nat). i < US.v vi ==> Seq.index s1 i == Seq.index s2 i) ) ) @@ -73,11 +73,10 @@ fn fill_array (#t:Type0) (l:US.t) (a:(a:A.array t{ US.v l == A.length a })) (v:t { let mut i = 0sz; while (let vi = !i; US.(vi <^ l)) - invariant b. exists* (s:Seq.seq t) (vi:US.t). ( + invariant exists* (s:Seq.seq t) (vi:US.t). ( A.pts_to a s ** R.pts_to i vi ** - pure ((b == US.(vi <^ l)) /\ - US.v vi <= US.v l /\ + pure (US.v vi <= US.v l /\ Seq.length s == A.length a /\ (forall (i:nat). i < US.v vi ==> Seq.index s i == v)) ) diff --git a/test/LoopInvariants.fst.output.expected b/test/LoopInvariants.fst.output.expected index 5382588e5..b4287eb7c 100644 --- a/test/LoopInvariants.fst.output.expected +++ b/test/LoopInvariants.fst.output.expected @@ -4,20 +4,18 @@ * Info at LoopInvariants.fst(77,9-77,10): - Expected failure: - - Failed to discharge match guard for goal: - Pulse.Lib.Reference.pts_to s 2 - with resource from context: - Pulse.Lib.Reference.pts_to s 3 + - rewrite: could not prove equality of + - Pulse.Lib.Reference.pts_to s 3 + - Pulse.Lib.Reference.pts_to s 2 - Assertion failed - The SMT solver could not prove the query. Use --query_stats for more details. * Info at LoopInvariants.fst(88,9-88,10): - Expected failure: - - Failed to discharge match guard for goal: - Pulse.Lib.Reference.pts_to s 2 - with resource from context: - Pulse.Lib.Reference.pts_to s 3 + - rewrite: could not prove equality of + - Pulse.Lib.Reference.pts_to s 3 + - Pulse.Lib.Reference.pts_to s 2 - Assertion failed - The SMT solver could not prove the query. Use --query_stats for more details. diff --git a/test/Test.Recursion.fst.output.expected b/test/Test.Recursion.fst.output.expected index 67a033405..5365dd8c2 100644 --- a/test/Test.Recursion.fst.output.expected +++ b/test/Test.Recursion.fst.output.expected @@ -1,6 +1,11 @@ * Info at Test.Recursion.fst(94,2-94,20): - Expected failure: - Ill-typed term: test_ghost_loop () + - Expected a term of type + Pulse.Lib.Core.stt_ghost Prims.unit + Pulse.Lib.Core.emp_inames + Pulse.Lib.Core.emp + (fun _ -> Pulse.Lib.Core.emp) - Could not prove termination - The SMT solver could not prove the query. Use --query_stats for more details. @@ -8,6 +13,11 @@ * Info at Test.Recursion.fst(141,4-141,22): - Expected failure: - Ill-typed term: test5' (z - 1) (y - 1) + - Expected a term of type + Pulse.Lib.Core.stt_ghost Prims.unit + Pulse.Lib.Core.emp_inames + Pulse.Lib.Core.emp + (fun _ -> Pulse.Lib.Core.emp) - Could not prove termination - The SMT solver could not prove the query. Use --query_stats for more details. diff --git a/test/Test.ReflikeClass.fst b/test/Test.ReflikeClass.fst index f37310a9d..bd2a92d3b 100644 --- a/test/Test.ReflikeClass.fst +++ b/test/Test.ReflikeClass.fst @@ -7,12 +7,14 @@ module Box = Pulse.Lib.Box [@@fundeps [0]; pulse_unfold] class reflike (vt:Type) (rt:Type) = { - ( |-> ) : rt -> vt -> slprop; + ( |-> ) : [@@@mkey]rt -> vt -> slprop; alloc : v:vt -> stt rt emp (fun r -> r |-> v); (!) : r:rt -> #v0:erased vt -> stt vt (r |-> v0) (fun v -> (r |-> v0) ** rewrites_to v (reveal v0)); (:=) : r:rt -> v:vt -> #v0:erased vt -> stt unit (r |-> v0) (fun _ -> r |-> v); } +(* // Disabling test because the typeclass feature mixes up Mkreflike?.op_Bar_Subtraction_Greater and op_Bar_Subtraction_Greater + #push-options "--warn_error -288" fn weakened_ref_alloc u#a (#a:Type u#a) {| small_type u#a |} (x: a) returns r: Pulse.Lib.Reference.ref a @@ -64,3 +66,4 @@ fn test2 (r : Box.box int) () } +*) \ No newline at end of file diff --git a/test/bug-reports/Bug.DesugaringError.fst.output.expected b/test/bug-reports/Bug.DesugaringError.fst.output.expected index f52bc0f47..c95b375aa 100644 --- a/test/bug-reports/Bug.DesugaringError.fst.output.expected +++ b/test/bug-reports/Bug.DesugaringError.fst.output.expected @@ -14,5 +14,6 @@ * Info at Bug.DesugaringError.fst(46,4-46,8): - Expected failure: + - Ill-typed term: Bug.DesugaringError.f true - Expected expression of type Prims.nat got expression true of type Prims.bool diff --git a/test/bug-reports/Bug.SpinLock.fst.output.expected b/test/bug-reports/Bug.SpinLock.fst.output.expected index 031929361..6be5754da 100644 --- a/test/bug-reports/Bug.SpinLock.fst.output.expected +++ b/test/bug-reports/Bug.SpinLock.fst.output.expected @@ -2,7 +2,7 @@ - Expected failure: - Tactic failed - Cannot prove: - Pulse.Lib.Reference.pts_to __ ?__ + Pulse.Lib.Reference.pts_to __ (*?u85*)_ - In the context: Pulse.Lib.SpinLock.lock_alive my_lock (exists* (v: Prims.int). Pulse.Lib.Reference.pts_to r v) diff --git a/test/bug-reports/Bug107.fst b/test/bug-reports/Bug107.fst index d30b5abc3..5319e0fe5 100644 --- a/test/bug-reports/Bug107.fst +++ b/test/bug-reports/Bug107.fst @@ -36,8 +36,6 @@ fn test2 () } -[@@expect_failure] // should work - fn test3 () requires foo 1 2 ensures emp @@ -47,8 +45,6 @@ fn test3 () } -[@@expect_failure] // should work - fn test4 () requires foo 1 2 ensures emp diff --git a/test/bug-reports/Bug107.fst.output.expected b/test/bug-reports/Bug107.fst.output.expected deleted file mode 100644 index 9e2936ac6..000000000 --- a/test/bug-reports/Bug107.fst.output.expected +++ /dev/null @@ -1,10 +0,0 @@ -* Info at Bug107.fst(45,9-45,16): - - Expected failure: - - Tactic failed - - Could not resolve implicit uu___#87 of type Prims.int in term Bug107.foo 1 _ - -* Info at Bug107.fst(56,9-56,16): - - Expected failure: - - Tactic failed - - Could not resolve implicit uu___#87 of type Prims.int in term Bug107.foo _ 2 - diff --git a/test/bug-reports/Bug110.fst b/test/bug-reports/Bug110.fst index 837175099..cf06cfb88 100644 --- a/test/bug-reports/Bug110.fst +++ b/test/bug-reports/Bug110.fst @@ -3,7 +3,7 @@ module Bug110 #lang-pulse open Pulse -assume val foo : int -> slprop +[@@no_mkeys] assume val foo : int -> slprop (* OK *) fn test1 (i j : int) @@ -14,6 +14,7 @@ fn test1 (i j : int) } (* Should fail, we cannot prove i == j *) +[@@expect_failure [19]] fn test2 (i j : int) requires foo i ** pure (foo i == foo j) ensures foo j diff --git a/test/bug-reports/Bug137.fst b/test/bug-reports/Bug137.fst index 85b3803d2..7b2e542fe 100644 --- a/test/bug-reports/Bug137.fst +++ b/test/bug-reports/Bug137.fst @@ -19,25 +19,12 @@ module Bug137 open Pulse.Lib.Pervasives +let tag p : slprop = p fn test_elim_pure (x:option bool) -requires exists* q. q ** pure (Some? x) -ensures exists* q. q ** pure (Some? x) +requires exists* q. tag q ** pure (Some? x) +ensures exists* q. tag q ** pure (Some? x) { let v = Some?.v x; () } - - -// This previously had an exists for p,q in the pre and post, -// but then that's very ambiguous. I think this captures the idea -// of the test anyway. - -fn test_elim_pure2 (x:option bool) (p q : slprop) -requires p ** (exists* r. r ** pure (Some? x)) ** q -ensures p ** q ** (exists* r. r ** pure (Some? x)) -{ - let v = Some?.v x; - () -} - diff --git a/test/bug-reports/Bug174.fst.output.expected b/test/bug-reports/Bug174.fst.output.expected index dc1bf4c5b..1995e8fad 100644 --- a/test/bug-reports/Bug174.fst.output.expected +++ b/test/bug-reports/Bug174.fst.output.expected @@ -1,9 +1,11 @@ * Info at Bug174.fst(15,1-17,10): - Expected failure: - - Ill-typed term: !r + - rewrite: could not prove equality of + - Pulse.Lib.Reference.pts_to r v + - Pulse.Lib.Reference.pts_to r v - Assertion failed - The SMT solver could not prove the query. Use --query_stats for more details. - - Also see: Bug174.fst(16,30-16,42) - - Other related locations: Bug174.fst(16,39-16,40) + - Also see: Bug174.fst(11,11-16,42) + - Other related locations: Bug174.fst(16,10-16,40) diff --git a/test/bug-reports/Bug206.fst.output.expected b/test/bug-reports/Bug206.fst.output.expected index 9bd7455aa..269dbfcd4 100644 --- a/test/bug-reports/Bug206.fst.output.expected +++ b/test/bug-reports/Bug206.fst.output.expected @@ -1,6 +1,10 @@ * Info at Bug206.fst(18,2-18,6): - Expected failure: - Ill-typed term: Bug206.test x + - Expected a term of type + Pulse.Lib.Core.stt Prims.unit + Pulse.Lib.Core.emp + (fun _ -> Pulse.Lib.Core.emp) - Assertion failed - The SMT solver could not prove the query. Use --query_stats for more details. diff --git a/test/bug-reports/Bug267.fst b/test/bug-reports/Bug267.fst index 217cb9871..9fcb62c16 100644 --- a/test/bug-reports/Bug267.fst +++ b/test/bug-reports/Bug267.fst @@ -45,14 +45,13 @@ fn multiply_by_repeated_addition (x y:nat) let c = !ctr; (c < x) ) - invariant b. + invariant exists* (c a : int). pts_to ctr c ** Pulse.Lib.Reference.pts_to acc a ** pure (//0 <= c /\ c <= x /\ - a == (c * y) /\ - b == (c < x)) + a == (c * y)) { let a = !acc; acc := a + y; @@ -62,8 +61,8 @@ fn multiply_by_repeated_addition (x y:nat) !acc } -[@@expect_failure [12]] +[@@expect_failure [228]] fn foo (n: nat) requires pure (n >= 42) ensures emp { - assert (n >= 1); + assert (n >= 1); //misses a pure } diff --git a/test/bug-reports/Bug267.fst.output.expected b/test/bug-reports/Bug267.fst.output.expected index b9d0544ca..ec37f5006 100644 --- a/test/bug-reports/Bug267.fst.output.expected +++ b/test/bug-reports/Bug267.fst.output.expected @@ -3,7 +3,7 @@ - Tactic failed - Expected a total term, but this term has Ghost effect -* Info at Bug267.fst(62,4-62,5): +* Info at Bug267.fst(61,5-61,8): - Expected failure: - Could not prove equality between computed type Prims.int and expected type Prims.nat @@ -11,9 +11,13 @@ - The SMT solver could not prove the query. Use --query_stats for more details. - Also see: Prims.fst(161,28-183,79) - - See also Bug267.fst(62,4-62,8) + - See also Bug267.fst(61,4-61,8) -* Info at Bug267.fst(67,13-67,14): +* Info at Bug267.fst(65,53-66,17): - Expected failure: - - Expected type Pulse.Lib.Core.slprop but n >= 1 has type Prims.bool + - Tactic failed + - Cannot prove: + n >= 1 + - In the context: + diff --git a/test/bug-reports/Bug274.fst.output.expected b/test/bug-reports/Bug274.fst.output.expected index bebbc3aa9..a3a6f4983 100644 --- a/test/bug-reports/Bug274.fst.output.expected +++ b/test/bug-reports/Bug274.fst.output.expected @@ -1,111 +1,31 @@ -* Info at Bug274.fst(29,4-29,15): +* Info at Bug274.fst(29,4-29,9): - Expected failure: - Tactic failed - Cannot prove any of: - Pulse.Lib.Trade.trade _ _ ** Pulse.Lib.Trade.trade _ _ + Pulse.Lib.Trade.trade (*?u99*)_ (*?u100*)_ + Pulse.Lib.Trade.trade (*?u100*)_ (*?u101*)_ - In the context: - Pulse.Lib.Trade.trade p q ** Pulse.Lib.Trade.trade q r - - Some hints: - - Ambiguous match for resource: - Pulse.Lib.Trade.trade _ _ - - It can be matched by both: Pulse.Lib.Trade.trade p q - and: Pulse.Lib.Trade.trade q r - in the context. - - Ambiguous match for resource: - Pulse.Lib.Trade.trade _ _ - - It can be matched by both: - Pulse.Lib.Trade.trade p q - and: - Pulse.Lib.Trade.trade q r - in the context. - - Ambiguous match for resource: - Pulse.Lib.Trade.trade _ _ - - It can be matched by both: - Pulse.Lib.Trade.trade p q - and: - Pulse.Lib.Trade.trade q r - in the context. - - Ambiguous match for resource: - Pulse.Lib.Trade.trade _ _ - - It can be matched by both: - Pulse.Lib.Trade.trade p q - and: - Pulse.Lib.Trade.trade q r - in the context. -* Info at Bug274.fst(41,4-41,15): +* Info at Bug274.fst(41,4-41,9): - Expected failure: - Tactic failed - Cannot prove any of: - Pulse.Lib.Trade.trade _ _ ** Pulse.Lib.Trade.trade _ _ + Pulse.Lib.Trade.trade (*?u99*)_ (*?u100*)_ + Pulse.Lib.Trade.trade (*?u100*)_ (*?u101*)_ - In the context: - Pulse.Lib.Trade.trade p q ** Pulse.Lib.Trade.trade q r - - Some hints: - - Ambiguous match for resource: - Pulse.Lib.Trade.trade _ _ - - It can be matched by both: - Pulse.Lib.Trade.trade q r - and: Pulse.Lib.Trade.trade p q - in the context. - - Ambiguous match for resource: - Pulse.Lib.Trade.trade _ _ - - It can be matched by both: Pulse.Lib.Trade.trade q r - and: - Pulse.Lib.Trade.trade p q - in the context. - - Ambiguous match for resource: - Pulse.Lib.Trade.trade _ _ - - It can be matched by both: - Pulse.Lib.Trade.trade q r - and: - Pulse.Lib.Trade.trade p q - in the context. - - Ambiguous match for resource: - Pulse.Lib.Trade.trade _ _ - - It can be matched by both: - Pulse.Lib.Trade.trade q r - and: - Pulse.Lib.Trade.trade p q - in the context. -* Info at Bug274.fst(73,4-73,12): +* Info at Bug274.fst(73,4-73,8): - Expected failure: - Tactic failed - Cannot prove any of: - _ ** Pulse.Lib.Trade.trade _ _ + Pulse.Lib.Trade.trade (*?u119*)_ (*?u120*)_ + (*?u119*)_ - In the context: - p ** Pulse.Lib.Trade.trade p q ** Pulse.Lib.Trade.trade r r - - Some hints: - - Ambiguous match for resource: - Pulse.Lib.Trade.trade _ _ - - It can be matched by both: - Pulse.Lib.Trade.trade r r - and: - Pulse.Lib.Trade.trade p q - in the context. - - Ambiguous match for resource: - _ - - It can be matched by both: - Pulse.Lib.Trade.trade r r - and: - Pulse.Lib.Trade.trade p q - in the context. - - Ambiguous match for resource: - Pulse.Lib.Trade.trade _ _ - - It can be matched by both: - Pulse.Lib.Trade.trade r r - and: + p Pulse.Lib.Trade.trade p q - in the context. - - Ambiguous match for resource: - _ - - It can be matched by both: Pulse.Lib.Trade.trade r r - and: - p - in the context. diff --git a/test/bug-reports/Bug278.fst.output.expected b/test/bug-reports/Bug278.fst.output.expected index 2dbc7d305..0e15d8bf6 100644 --- a/test/bug-reports/Bug278.fst.output.expected +++ b/test/bug-reports/Bug278.fst.output.expected @@ -1,4 +1,5 @@ -* Info at Bug278.fst(36,11-36,16): +* Info at Bug278.fst(35,5-37,5): - Expected failure: - - Elaborated term has unresolved univ uvars: Bug278.spawn (Bug278.pth n) + - Tactic failed + - Internal error: unexpected unresolved (universe) uvar in deep_compress: 21 diff --git a/test/bug-reports/Bug29.fst.output.expected b/test/bug-reports/Bug29.fst.output.expected index e7f1e68db..95077164a 100644 --- a/test/bug-reports/Bug29.fst.output.expected +++ b/test/bug-reports/Bug29.fst.output.expected @@ -4,5 +4,5 @@ - Cannot prove: Pulse.Lib.Reference.pts_to r v - In the context: - emp + diff --git a/test/bug-reports/Bug45.fst.output.expected b/test/bug-reports/Bug45.fst.output.expected index b3594c056..e487a3b57 100644 --- a/test/bug-reports/Bug45.fst.output.expected +++ b/test/bug-reports/Bug45.fst.output.expected @@ -1,4 +1,5 @@ -* Info at Bug45.fst(64,13-64,18): +* Info at Bug45.fst(61,10-66,5): - Expected failure: - - Elaborated term has unresolved univ uvars: Bug45.spawn (Bug45.pth n) + - Tactic failed + - Internal error: unexpected unresolved (universe) uvar in deep_compress: 75 diff --git a/test/bug-reports/BugIfFalse.fst b/test/bug-reports/BugIfFalse.fst new file mode 100644 index 000000000..5f3cfacea --- /dev/null +++ b/test/bug-reports/BugIfFalse.fst @@ -0,0 +1,17 @@ +module BugIfFalse +#lang-pulse +open Pulse.Lib.Pervasives + +fn text () +requires emp +ensures emp +{ + if (false) + { + () + } + else + { + () + } +} \ No newline at end of file diff --git a/test/bug-reports/BugUnificationUnderBinder.fst b/test/bug-reports/BugUnificationUnderBinder.fst index 59248e3f8..6d985b233 100644 --- a/test/bug-reports/BugUnificationUnderBinder.fst +++ b/test/bug-reports/BugUnificationUnderBinder.fst @@ -27,7 +27,7 @@ requires exists* x. p x 5 ensures emp { with zz. assert (exists* x. p x zz); - drop_ _; + drop_ (p _ zz); } @@ -37,7 +37,7 @@ requires forall* x. p x 5 ensures emp { with zz. assert (forall* x. p x zz); - drop_ _; + drop_ (forall* x. p x zz); } diff --git a/test/bug-reports/BugWhileInv.fst b/test/bug-reports/BugWhileInv.fst index bba7cf9a3..78c9760a7 100644 --- a/test/bug-reports/BugWhileInv.fst +++ b/test/bug-reports/BugWhileInv.fst @@ -19,6 +19,7 @@ module BugWhileInv open Pulse.Lib.Pervasives module R = Pulse.Lib.Reference +[@@no_mkeys] let workaround (b:bool) (v:nat) : slprop = pure (not b ==> v == 0) @@ -70,9 +71,6 @@ ensures R.pts_to x 0 true; } ) - invariant b. - exists* v. - R.pts_to x v ** - pure ((b == false) ==> (v == 0)) + invariant live x { () }; } diff --git a/test/bug-reports/ExistsErasedAndPureEqualities.fst b/test/bug-reports/ExistsErasedAndPureEqualities.fst index 1881c60cb..588b45d39 100644 --- a/test/bug-reports/ExistsErasedAndPureEqualities.fst +++ b/test/bug-reports/ExistsErasedAndPureEqualities.fst @@ -1,4 +1,4 @@ -(* + (* Copyright 2023 Microsoft Research Licensed under the Apache License, Version 2.0 (the "License"); @@ -22,9 +22,6 @@ module R = Pulse.Lib.Reference assume val some_pred (x:R.ref int) (v:int) : slprop -//Intro exists* with an erased variable fails -[@@expect_failure] - fn test1 (x:R.ref int) (#v:Ghost.erased int) requires some_pred x v ensures some_pred x v @@ -37,11 +34,6 @@ fn test1 (x:R.ref int) (#v:Ghost.erased int) } -//Intro exists* with an erased variable in an equality bound on the left, -//fails weirdly with an SMT failure, where it tries to prove earsed int == int -//and hide v == v -//Intro exists* with an erased variable fails -[@@expect_failure] fn test2 (x:R.ref int) (#v:Ghost.erased int) requires some_pred x v diff --git a/test/bug-reports/ExistsErasedAndPureEqualities.fst.output.expected b/test/bug-reports/ExistsErasedAndPureEqualities.fst.output.expected index f46c01e1b..84e8c01e9 100644 --- a/test/bug-reports/ExistsErasedAndPureEqualities.fst.output.expected +++ b/test/bug-reports/ExistsErasedAndPureEqualities.fst.output.expected @@ -1,29 +1,17 @@ -* Info at ExistsErasedAndPureEqualities.fst(32,4-33,21): - - Expected failure: - - Tactic failed - - Could not resolve all free variables in the proposition: v == v_ - - See also ExistsErasedAndPureEqualities.fst(32,4-34,12) - -* Info at ExistsErasedAndPureEqualities.fst(50,4-51,21): - - Expected failure: - - Tactic failed - - Could not resolve all free variables in the proposition: v == v_ - - See also ExistsErasedAndPureEqualities.fst(50,4-52,12) - -* Info at ExistsErasedAndPureEqualities.fst(68,4-68,11): +* Info at ExistsErasedAndPureEqualities.fst(60,4-60,11): - Admitting continuation. - Current context: ExistsErasedAndPureEqualities.some_pred x v - In typing environment: - __#586 : Prims.squash (v_ == v), - v_#585 : FStar.Ghost.erased Prims.int, - __#451 : Prims.squash (v == v), + __#545 : Prims.squash (_v_4 == v), + _v_4#544 : FStar.Ghost.erased Prims.int, + __#410 : Prims.squash (v == v), v#267 : FStar.Ghost.erased Prims.int, x#262 : Pulse.Lib.Reference.ref Prims.int -* Info at ExistsErasedAndPureEqualities.fst(79,4-80,21): +* Info at ExistsErasedAndPureEqualities.fst(71,4-73,12): - Expected failure: - - Tactic failed - - Could not resolve all free variables in the proposition: v == v_ - - See also ExistsErasedAndPureEqualities.fst(79,4-81,12) + - Ill-typed term: pure (v == (*?u106*)_) + - Expected a term of type slprop + - Fstarcompiler.FStarC_Tactics_Common.TacticFailure(_) diff --git a/test/bug-reports/IntroGhost.fst.output.expected b/test/bug-reports/IntroGhost.fst.output.expected index a816f5bb3..6280218c3 100644 --- a/test/bug-reports/IntroGhost.fst.output.expected +++ b/test/bug-reports/IntroGhost.fst.output.expected @@ -2,9 +2,9 @@ - Expected failure: - Tactic failed - Cannot prove: - Pulse.Lib.Reference.pts_to r __ + Pulse.Lib.Reference.pts_to r (*?u258*)_ - In the context: - IntroGhost.my_inv b r + IntroGhost.my_inv _b2 r * Info at IntroGhost.fst(71,2-71,4): - Expected failure: @@ -12,16 +12,13 @@ - Cannot prove: Pulse.Lib.Reference.pts_to r 0 - In the context: - IntroGhost.my_inv b r + IntroGhost.my_inv _b3 r * Info at IntroGhost.fst(89,46-92,24): - Expected failure: - - Failed to discharge match guard for goal: + - Tactic failed + - Cannot prove: IntroGhost.my_inv true r - with resource from context: - IntroGhost.my_inv b r - - Assertion failed - - The SMT solver could not prove the query. Use --query_stats for more - details. - - Also see: IntroGhost.fst(92,17-92,21) + - In the context: + IntroGhost.my_inv _b3 r diff --git a/test/bug-reports/OrderDepHigherOrder.fst b/test/bug-reports/OrderDepHigherOrder.fst index 8528da206..580645dc2 100644 --- a/test/bug-reports/OrderDepHigherOrder.fst +++ b/test/bug-reports/OrderDepHigherOrder.fst @@ -34,23 +34,11 @@ fn test2 (m:mutex int) -[@@allow_ambiguous] - -fn flip (#p #q : slprop) (_:unit) - requires p ** q - ensures q ** p -{ - () -} - - - fn test3 (m:mutex int) requires mutex_live m p ensures mutex_live m p { let r = lock m; - flip (); unlock m r; () } diff --git a/test/bug-reports/UnificationVariableEscapes.fst b/test/bug-reports/UnificationVariableEscapes.fst index cbd8777a2..0f95b3b25 100644 --- a/test/bug-reports/UnificationVariableEscapes.fst +++ b/test/bug-reports/UnificationVariableEscapes.fst @@ -33,11 +33,10 @@ fn fill_array (#t:Type0) (a:A.array t) (l:(l:US.t { US.v l == A.length a })) (v: { let mut i = 0sz; while (let vi = !i; US.(vi <^ l)) - invariant b. exists* (s:Seq.seq t) (vi:US.t). ( + invariant exists* (s:Seq.seq t) (vi:US.t). ( A.pts_to a s ** R.pts_to i vi ** - pure ((b == US.(vi <^ l)) /\ - US.v vi <= US.v l /\ + pure (US.v vi <= US.v l /\ Seq.length s == A.length a /\ (forall (i:nat). i < US.v vi ==> Seq.index s i == v)) ) diff --git a/test/nolib/AssertWildcard.fst b/test/nolib/AssertWildcard.fst index 54a5bb88e..b5e04f9c3 100644 --- a/test/nolib/AssertWildcard.fst +++ b/test/nolib/AssertWildcard.fst @@ -6,8 +6,6 @@ open Pulse.Nolib assume val foo (x y z w : int) : slprop -(* Would be nice if this worked. *) -[@@expect_failure] fn test () (#x:int) requires foo x 'y 'z 'w ensures foo x 'y 'z 'w diff --git a/test/nolib/Bug32.fst b/test/nolib/Bug32.fst new file mode 100644 index 000000000..453b003fc --- /dev/null +++ b/test/nolib/Bug32.fst @@ -0,0 +1,10 @@ +module Bug32 +open Pulse.Nolib +#lang-pulse + +[@@expect_failure [189]] +ghost fn test () +{ + rewrite emp as (if true then emp else 1); + admit(); +} \ No newline at end of file diff --git a/test/nolib/Bug357.fst b/test/nolib/Bug357.fst index ccd3ef97c..9c592ae1d 100644 --- a/test/nolib/Bug357.fst +++ b/test/nolib/Bug357.fst @@ -15,7 +15,6 @@ fn test1 (x : t2) match x { norewrite C y z -> { let foo = z; - rewrite each x as C y z; (); } } @@ -27,7 +26,7 @@ fn test11 (x : t2) { match x { y -> { - (); + rewrite each y as x; } } } @@ -38,7 +37,7 @@ fn test2 (x : t2) { match x { C y z -> { - (); + rewrite each C y z as x; } } } @@ -56,5 +55,5 @@ fn test4 (x : t2) ensures foo x { let C y z = x; - (); + rewrite each C y z as x; } diff --git a/test/nolib/ErrCantFindWitness.fst b/test/nolib/ErrCantFindWitness.fst new file mode 100644 index 000000000..180dc5adc --- /dev/null +++ b/test/nolib/ErrCantFindWitness.fst @@ -0,0 +1,8 @@ +module ErrCantFindWitness +#lang-pulse +open Pulse.Nolib + +[@@expect_failure] +fn foo () + ensures exists* (x: nat). emp +{} \ No newline at end of file diff --git a/test/nolib/ErrCantFindWitness.fst.output.expected b/test/nolib/ErrCantFindWitness.fst.output.expected new file mode 100644 index 000000000..0a240e7cc --- /dev/null +++ b/test/nolib/ErrCantFindWitness.fst.output.expected @@ -0,0 +1,7 @@ +* Info at ErrCantFindWitness.fst(8,1-8,1): + - Expected failure: + - Cannot find witness for (exists* (x:Prims.nat). emp) + - Ill-typed term: (*?u56*)_ + - Expected a term of type Prims.nat + - Fstarcompiler.FStarC_Tactics_Common.TacticFailure(_) + diff --git a/test/nolib/InameReduce.fst b/test/nolib/InameReduce.fst index d74c2a518..2a118ca7e 100644 --- a/test/nolib/InameReduce.fst +++ b/test/nolib/InameReduce.fst @@ -3,6 +3,7 @@ module InameReduce #lang-pulse open Pulse.Nolib +[@@no_mkeys] assume val foo : inames -> slprop #push-options "--no_smt" diff --git a/test/nolib/Match.fst b/test/nolib/Match.fst index 70a5f03e8..81c1c0730 100644 --- a/test/nolib/Match.fst +++ b/test/nolib/Match.fst @@ -51,9 +51,9 @@ fn test (x : t1) (y z : t2) () } B -> { - let D = y; - let D = z; - (); + norewrite let D = y; + norewrite let D = z; + rewrite each B as x; } } } diff --git a/test/nolib/RewriteNopWarning.fst b/test/nolib/RewriteNopWarning.fst index 2089929a9..2d6fd7f9e 100644 --- a/test/nolib/RewriteNopWarning.fst +++ b/test/nolib/RewriteNopWarning.fst @@ -7,6 +7,7 @@ assume val foo : int -> slprop (* The rewrite is a nop. *) fn test1 (y : int) + requires pure (y == 1) preserves foo 1 { rewrite each y as 1; @@ -14,6 +15,7 @@ fn test1 (y : int) (* Same *) fn test2 (y : int) (z : int) + requires pure (z == 42) preserves foo 1 ** foo y { rewrite each z as 42; diff --git a/test/nolib/RewriteNopWarning.fst.output.expected b/test/nolib/RewriteNopWarning.fst.output.expected index 2eff747cb..6fe4a7d30 100644 --- a/test/nolib/RewriteNopWarning.fst.output.expected +++ b/test/nolib/RewriteNopWarning.fst.output.expected @@ -1,9 +1,9 @@ -* Warning at RewriteNopWarning.fst(11,1-12,21): +* Warning at RewriteNopWarning.fst(12,1-13,21): - No rewrites performed. - Rewriting in: RewriteNopWarning.foo 1 -* Warning at RewriteNopWarning.fst(18,1-19,22): +* Warning at RewriteNopWarning.fst(20,1-21,22): - No rewrites performed. - Rewriting in: RewriteNopWarning.foo y ** RewriteNopWarning.foo 1 diff --git a/test/nolib/UnfoldArgs.fst b/test/nolib/UnfoldArgs.fst index 373e61adc..5d91bf3b4 100644 --- a/test/nolib/UnfoldArgs.fst +++ b/test/nolib/UnfoldArgs.fst @@ -42,5 +42,6 @@ fn test4 (y : _) ensures trade emp (rel (g y)) { rewrite each rel2 (g y) as rel (g (f (g y))); // ideally automated? + rewrite each g (f (g y)) as g y; (); }