From 370cbc7aa45280f836a6396a45ac69291b1a9989 Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Fri, 9 Dec 2022 10:05:50 +0000 Subject: [PATCH 01/18] syntax changes to remove universes from Total and GTotal comps --- src/syntax/FStar.Syntax.Free.fst | 8 +-- src/syntax/FStar.Syntax.Hash.fst | 17 +++--- src/syntax/FStar.Syntax.InstFV.fst | 4 +- src/syntax/FStar.Syntax.Print.fst | 14 ++--- src/syntax/FStar.Syntax.Resugar.fst | 30 ++-------- src/syntax/FStar.Syntax.Subst.fst | 14 ++--- src/syntax/FStar.Syntax.Syntax.fst | 6 +- src/syntax/FStar.Syntax.Syntax.fsti | 6 +- src/syntax/FStar.Syntax.Util.fst | 93 +++++++++++++++-------------- 9 files changed, 77 insertions(+), 115 deletions(-) diff --git a/src/syntax/FStar.Syntax.Free.fst b/src/syntax/FStar.Syntax.Free.fst index af4f5c5b7e3..4f37ddb0118 100644 --- a/src/syntax/FStar.Syntax.Free.fst +++ b/src/syntax/FStar.Syntax.Free.fst @@ -216,14 +216,10 @@ and free_names_and_uvars_comp c use_cache = else n, new_fv_set () | _ -> let n = match c.n with - | GTotal (t, None) - | Total (t, None) -> + | GTotal t + | Total t -> free_names_and_uvars t use_cache - | GTotal (t, Some u) - | Total (t, Some u) -> - union (free_univs u) (free_names_and_uvars t use_cache) - | Comp ct -> //collect from the decreases clause let decreases_vars = diff --git a/src/syntax/FStar.Syntax.Hash.fst b/src/syntax/FStar.Syntax.Hash.fst index f97c2b61b09..003c3fd7954 100644 --- a/src/syntax/FStar.Syntax.Hash.fst +++ b/src/syntax/FStar.Syntax.Hash.fst @@ -132,16 +132,14 @@ and hash_term' (t:term) and hash_comp' (c:comp) : mm H.hash_code = match c.n with - | Total (t, ou) -> + | Total t -> mix_list_lit [of_int 811; - hash_term t; - hash_option hash_universe ou] - | GTotal (t, ou) -> + hash_term t] + | GTotal t -> mix_list_lit [of_int 821; - hash_term t; - hash_option hash_universe ou] + hash_term t] | Comp ct -> mix_list_lit [of_int 823; @@ -413,10 +411,9 @@ let rec equal_term (t1 t2:term) and equal_comp c1 c2 = if physical_equality c1 c2 then true else match c1.n, c2.n with - | Total (t1, u1), Total (t2, u2) - | GTotal (t1, u1), GTotal (t2, u2) -> - equal_term t1 t2 && - equal_opt equal_universe u1 u2 + | Total t1, Total t2 + | GTotal t1, GTotal t2 -> + equal_term t1 t2 | Comp ct1, Comp ct2 -> Ident.lid_equals ct1.effect_name ct2.effect_name && equal_list equal_universe ct1.comp_univs ct2.comp_univs && diff --git a/src/syntax/FStar.Syntax.InstFV.fst b/src/syntax/FStar.Syntax.InstFV.fst index a7f8bcc832f..a1af732de5a 100644 --- a/src/syntax/FStar.Syntax.InstFV.fst +++ b/src/syntax/FStar.Syntax.InstFV.fst @@ -106,8 +106,8 @@ and inst_binders s bs = bs |> List.map (inst_binder s) and inst_args s args = args |> List.map (fun (a, imp) -> inst s a, imp) and inst_comp s c = match c.n with - | Total (t, uopt) -> S.mk_Total' (inst s t) uopt - | GTotal (t, uopt) -> S.mk_GTotal' (inst s t) uopt + | Total t -> S.mk_Total (inst s t) + | GTotal t -> S.mk_GTotal (inst s t) | Comp ct -> let ct = {ct with result_typ=inst s ct.result_typ; effect_args=inst_args s ct.effect_args; flags=ct.flags |> List.map (function diff --git a/src/syntax/FStar.Syntax.Print.fst b/src/syntax/FStar.Syntax.Print.fst index aba3e5b75ce..bb9e357077e 100644 --- a/src/syntax/FStar.Syntax.Print.fst +++ b/src/syntax/FStar.Syntax.Print.fst @@ -529,21 +529,15 @@ and comp_to_string c = else Errors.with_ctx "While ugly-printing a computation" (fun () -> match c.n with - | Total (t, uopt) -> + | Total t -> begin match (compress t).n with | Tm_type _ when not (Options.print_implicits() || Options.print_universes()) -> term_to_string t - | _ -> - match uopt with - | Some u when Options.print_universes() -> U.format2 "Tot<%s> %s" (univ_to_string u) (term_to_string t) - | _ -> U.format1 "Tot %s" (term_to_string t) + | _ -> U.format1 "Tot %s" (term_to_string t) end - | GTotal (t, uopt) -> + | GTotal t -> begin match (compress t).n with | Tm_type _ when not (Options.print_implicits() || Options.print_universes()) -> term_to_string t - | _ -> - match uopt with - | Some u when Options.print_universes() -> U.format2 "GTot<%s> %s" (univ_to_string u) (term_to_string t) - | _ -> U.format1 "GTot %s" (term_to_string t) + | _ -> U.format1 "GTot %s" (term_to_string t) end | Comp c -> let basic = diff --git a/src/syntax/FStar.Syntax.Resugar.fst b/src/syntax/FStar.Syntax.Resugar.fst index 4cb547b21e4..b8dbf0f6f72 100644 --- a/src/syntax/FStar.Syntax.Resugar.fst +++ b/src/syntax/FStar.Syntax.Resugar.fst @@ -958,37 +958,15 @@ and resugar_comp' (env: DsEnv.env) (c:S.comp) : A.term = A.mk_term a c.pos A.Un in match (c.n) with - | Total (typ, u) -> + | Total typ -> let t = resugar_term' env typ in if not (Options.print_implicits()) then t - else ( - begin match u with - | None -> - mk (A.Construct(C.effect_Tot_lid, [(t, A.Nothing)])) - | Some u -> - // add the universe as the first argument - if (Options.print_universes()) then - let u = resugar_universe u c.pos in - mk (A.Construct(C.effect_Tot_lid, [(u, A.UnivApp);(t, A.Nothing)])) - else - mk (A.Construct(C.effect_Tot_lid, [(t, A.Nothing)])) - end - ) + else mk (A.Construct(C.effect_Tot_lid, [(t, A.Nothing)])) - | GTotal (typ, u) -> + | GTotal typ -> let t = resugar_term' env typ in - begin match u with - | None -> - mk (A.Construct(C.effect_GTot_lid, [(t, A.Nothing)])) - | Some u -> - // add the universe as the first argument - if (Options.print_universes()) then - let u = resugar_universe u c.pos in - mk (A.Construct(C.effect_GTot_lid, [(u, A.UnivApp);(t, A.Nothing)])) - else - mk (A.Construct(C.effect_GTot_lid, [(t, A.Nothing)])) - end + mk (A.Construct(C.effect_GTot_lid, [(t, A.Nothing)])) | Comp c -> let result = (resugar_term' env c.result_typ, A.Nothing) in diff --git a/src/syntax/FStar.Syntax.Subst.fst b/src/syntax/FStar.Syntax.Subst.fst index dc89facadd1..b259542f48b 100644 --- a/src/syntax/FStar.Syntax.Subst.fst +++ b/src/syntax/FStar.Syntax.Subst.fst @@ -263,8 +263,8 @@ let subst_comp' s t = | [[]], NoUseRange -> t | _ -> match t.n with - | Total (t, uopt) -> mk_Total' (subst' s t) (Option.map (subst_univ (fst s)) uopt) - | GTotal (t, uopt) -> mk_GTotal' (subst' s t) (Option.map (subst_univ (fst s)) uopt) + | Total t -> mk_Total (subst' s t) + | GTotal t -> mk_GTotal (subst' s t) | Comp ct -> mk_Comp(subst_comp_typ' s ct) let subst_ascription' s (asc:ascription) = @@ -969,13 +969,11 @@ and deep_compress_cflags allow_uvars flags = and deep_compress_comp allow_uvars (c:comp) : comp = let mk x = S.mk x c.pos in match c.n with - | Total (t, uopt) -> - let uopt = map_opt uopt (deep_compress_univ allow_uvars) in - mk (Total (deep_compress allow_uvars t, uopt)) + | Total t -> + mk (Total (deep_compress allow_uvars t)) - | GTotal (t, uopt) -> - let uopt = map_opt uopt (deep_compress_univ allow_uvars) in - mk (GTotal (deep_compress allow_uvars t, uopt)) + | GTotal t -> + mk (GTotal (deep_compress allow_uvars t)) | Comp ct -> let ct = { diff --git a/src/syntax/FStar.Syntax.Syntax.fst b/src/syntax/FStar.Syntax.Syntax.fst index 392c71d16ee..a5485768061 100644 --- a/src/syntax/FStar.Syntax.Syntax.fst +++ b/src/syntax/FStar.Syntax.Syntax.fst @@ -132,10 +132,8 @@ let extend_app_n t args' r = match t.n with | _ -> mk_Tm_app t args' r let extend_app t arg r = extend_app_n t [arg] r let mk_Tm_delayed lr pos : term = mk (Tm_delayed lr) pos -let mk_Total' t u: comp = mk (Total(t, u)) t.pos -let mk_GTotal' t u: comp = mk (GTotal(t, u)) t.pos -let mk_Total t = mk_Total' t None -let mk_GTotal t = mk_GTotal' t None +let mk_Total t = mk (Total t) t.pos +let mk_GTotal t : comp = mk (GTotal t) t.pos let mk_Comp (ct:comp_typ) : comp = mk (Comp ct) ct.result_typ.pos let mk_lb (x, univs, eff, t, e, attrs, pos) = { lbname=x; diff --git a/src/syntax/FStar.Syntax.Syntax.fsti b/src/syntax/FStar.Syntax.Syntax.fsti index 75cf9ddcff3..d63290a0825 100644 --- a/src/syntax/FStar.Syntax.Syntax.fsti +++ b/src/syntax/FStar.Syntax.Syntax.fsti @@ -193,8 +193,8 @@ and comp_typ = { flags:list cflag } and comp' = - | Total of typ * option universe - | GTotal of typ * option universe + | Total of typ + | GTotal of typ | Comp of comp_typ and term = syntax term' and typ = term (* sometimes we use typ to emphasize that a term is a type *) @@ -625,8 +625,6 @@ val extend_app_n: term -> args -> range -> term val mk_Tm_delayed: (term * subst_ts) -> Range.range -> term val mk_Total: typ -> comp val mk_GTotal: typ -> comp -val mk_Total': typ -> option universe -> comp -val mk_GTotal': typ -> option universe -> comp val mk_Tac : typ -> comp val mk_Comp: comp_typ -> comp val bv_to_tm: bv -> term diff --git a/src/syntax/FStar.Syntax.Util.fst b/src/syntax/FStar.Syntax.Util.fst index 3804a5899ba..f3b8ac03e67 100644 --- a/src/syntax/FStar.Syntax.Util.fst +++ b/src/syntax/FStar.Syntax.Util.fst @@ -241,32 +241,38 @@ let comp_flags c = match c.n with | GTotal _ -> [SOMETRIVIAL] | Comp ct -> ct.flags -(* Duplicate of the function below not failing when universe is not provided *) -let comp_to_comp_typ_nouniv (c:comp) : comp_typ = - match c.n with - | Comp c -> c - | Total (t, u_opt) - | GTotal(t, u_opt) -> - {comp_univs=dflt [] (map_opt u_opt (fun x -> [x])); - effect_name=comp_effect_name c; - result_typ=t; - effect_args=[]; - flags=comp_flags c} - -let comp_set_flags (c:comp) f = - {c with n=Comp ({comp_to_comp_typ_nouniv c with flags=f})} - -let comp_to_comp_typ (c:comp) : comp_typ = - match c.n with - | Comp c -> c - | Total (t, Some u) - | GTotal(t, Some u) -> - {comp_univs=[u]; - effect_name=comp_effect_name c; - result_typ=t; - effect_args=[]; - flags=comp_flags c} - | _ -> failwith "Assertion failed: Computation type without universe" +let comp_eff_name_res_and_args (c:comp) : lident & typ & args = + match c.n with + | Total t -> PC.effect_Tot_lid, t, [] + | GTotal t -> PC.effect_GTot_lid, t, [] + | Comp c -> c.effect_name, c.result_typ, c.effect_args + +// (* Duplicate of the function below not failing when universe is not provided *) +// let comp_to_comp_typ_nouniv (c:comp) : comp_typ = +// match c.n with +// | Comp c -> c +// | Total (t, u_opt) +// | GTotal(t, u_opt) -> +// {comp_univs=dflt [] (map_opt u_opt (fun x -> [x])); +// effect_name=comp_effect_name c; +// result_typ=t; +// effect_args=[]; +// flags=comp_flags c} + +// let comp_set_flags (c:comp) f = +// {c with n=Comp ({comp_to_comp_typ_nouniv c with flags=f})} + +// let comp_to_comp_typ (c:comp) : comp_typ = +// match c.n with +// | Comp c -> c +// | Total (t, Some u) +// | GTotal(t, Some u) -> +// {comp_univs=[u]; +// effect_name=comp_effect_name c; +// result_typ=t; +// effect_args=[]; +// flags=comp_flags c} +// | _ -> failwith "Assertion failed: Computation type without universe" (* @@ -288,7 +294,7 @@ let effect_indices_from_repr (repr:term) (is_layered:bool) (r:Range.range) (err: | Tm_app (_, _::is) -> is |> List.map fst | _ -> err () else match repr.n with - | Tm_arrow (_, c) -> c |> comp_to_comp_typ |> (fun ct -> ct.effect_args |> List.map fst) + | Tm_arrow (_, c) -> c |> comp_eff_name_res_and_args |> (fun (_, _, args) -> args |> List.map fst) | _ -> err () let destruct_comp c : (universe * typ * typ) = @@ -421,8 +427,8 @@ let is_ml_comp c = match c.n with | _ -> false let comp_result c = match c.n with - | Total (t, _) - | GTotal (t, _) -> t + | Total t + | GTotal t -> t | Comp ct -> ct.result_typ let set_result_typ c t = match c.n with @@ -744,11 +750,9 @@ and eq_univs_list (us:universes) (vs:universes) : bool = and eq_comp (c1 c2:comp) : eq_result = match c1.n, c2.n with - | Total (t1, u1opt), Total (t2, u2opt) - | GTotal (t1, u1opt), GTotal (t2, u2opt) -> - eq_tm t1 t2 //Rel ignores the universe annotations when comparing these comps for equality. - //So, ignore them here too - //But, this should be sorted out: Can we always populate the universe? + | Total t1, Total t2 + | GTotal t1, GTotal t2 -> + eq_tm t1 t2 | Comp ct1, Comp ct2 -> eq_and (equal_if (eq_univs_list ct1.comp_univs ct2.comp_univs)) (fun _ -> @@ -997,7 +1001,7 @@ let flat_arrow bs c = match (Subst.compress t).n with | Tm_arrow(bs, c) -> begin match c.n with - | Total (tres, _) -> + | Total tres -> begin match (Subst.compress tres).n with | Tm_arrow(bs', c') -> mk (Tm_arrow(bs@bs', c')) t.pos | _ -> t @@ -1010,7 +1014,7 @@ let rec canon_arrow t = match (compress t).n with | Tm_arrow (bs, c) -> let cn = match c.n with - | Total (t, u) -> Total (canon_arrow t, u) + | Total t -> Total (canon_arrow t) | _ -> c.n in let c = { c with n = cn } in @@ -1081,9 +1085,8 @@ let let_rec_arity (lb:letbinding) : int * option (list bool) = match k.n with | Tm_arrow(bs, c) -> let bs, c = Subst.open_comp bs c in - let ct = comp_to_comp_typ c in (match - ct.flags |> U.find_opt (function DECREASES _ -> true | _ -> false) + c |> comp_flags |> U.find_opt (function DECREASES _ -> true | _ -> false) with | Some (DECREASES d) -> bs, Some d @@ -1639,7 +1642,7 @@ let destruct_typ_as_formula f : option connective = if not (is_tot_or_gtot_comp c) then None else - let q = (comp_to_comp_typ_nouniv c).result_typ in + let q = comp_result c in if is_free_in b.binder_bv q then ( let pats, q = patterns q in @@ -1928,11 +1931,11 @@ and binder_eq_dbg (dbg : bool) b1 b2 = (check "binder attrs" (eqlist (term_eq_dbg dbg) b1.binder_attrs b2.binder_attrs)) and comp_eq_dbg (dbg : bool) c1 c2 = - let c1 = comp_to_comp_typ_nouniv c1 in - let c2 = comp_to_comp_typ_nouniv c2 in - (check "comp eff" (lid_equals c1.effect_name c2.effect_name)) && + let eff1, res1, args1 = comp_eff_name_res_and_args c1 in + let eff2, res2, args2 = comp_eff_name_res_and_args c2 in + (check "comp eff" (lid_equals eff1 eff2)) && //(check "comp univs" (c1.comp_univs = c2.comp_univs)) && - (check "comp result typ" (term_eq_dbg dbg c1.result_typ c2.result_typ)) && + (check "comp result typ" (term_eq_dbg dbg res1 res2)) && (* (check "comp args" (eqlist arg_eq_dbg dbg c1.effect_args c2.effect_args)) && *) true //eq_flags c1.flags c2.flags and branch_eq_dbg (dbg : bool) (p1,w1,t1) (p2,w2,t2) = @@ -2144,8 +2147,8 @@ and unbound_variables_ascription asc = and unbound_variables_comp c = match c.n with - | GTotal (t, _) - | Total (t, _) -> + | GTotal t + | Total t -> unbound_variables t | Comp ct -> From af53d7780a1293e84920fe422bd9eb282782792c Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Fri, 9 Dec 2022 10:06:15 +0000 Subject: [PATCH 02/18] tc changes, mostly relying on Env.comp_to_comp_typ --- src/typechecker/FStar.TypeChecker.Core.fst | 20 ++++----- src/typechecker/FStar.TypeChecker.DMFF.fst | 16 +++---- src/typechecker/FStar.TypeChecker.Env.fst | 23 ++++++---- src/typechecker/FStar.TypeChecker.Env.fsti | 1 + src/typechecker/FStar.TypeChecker.Err.fst | 4 +- src/typechecker/FStar.TypeChecker.NBE.fst | 8 ++-- .../FStar.TypeChecker.Normalize.fst | 34 ++++++--------- src/typechecker/FStar.TypeChecker.Rel.fst | 42 ++++++++----------- .../FStar.TypeChecker.TcEffect.fst | 16 +++---- src/typechecker/FStar.TypeChecker.TcTerm.fst | 20 ++++----- src/typechecker/FStar.TypeChecker.Util.fst | 41 +++++++++--------- 11 files changed, 107 insertions(+), 118 deletions(-) diff --git a/src/typechecker/FStar.TypeChecker.Core.fst b/src/typechecker/FStar.TypeChecker.Core.fst index 05c137983e5..c2a1244795d 100644 --- a/src/typechecker/FStar.TypeChecker.Core.fst +++ b/src/typechecker/FStar.TypeChecker.Core.fst @@ -1110,19 +1110,19 @@ and check_relation_comp (g:env) rel (c0 c1:comp) if U.eq_comp c0 c1 = U.Equal then return () else ( - let ct_eq ct0 ct1 = - check_relation g EQUALITY ct0.result_typ ct1.result_typ ;! - check_relation_args g EQUALITY ct0.effect_args ct1.effect_args + let ct_eq res0 args0 res1 args1 = + check_relation g EQUALITY res0 res1 ;! + check_relation_args g EQUALITY args0 args1 in - let ct0 = U.comp_to_comp_typ_nouniv c0 in - let ct1 = U.comp_to_comp_typ_nouniv c1 in - if I.lid_equals ct0.effect_name ct1.effect_name - then ct_eq ct0 ct1 + let eff0, res0, args0 = U.comp_eff_name_res_and_args c0 in + let eff1, res1, args1 = U.comp_eff_name_res_and_args c1 in + if I.lid_equals eff0 eff1 + then ct_eq res0 args0 res1 args1 else ( let ct0 = Env.unfold_effect_abbrev g.tcenv c0 in let ct1 = Env.unfold_effect_abbrev g.tcenv c1 in if I.lid_equals ct0.effect_name ct1.effect_name - then ct_eq ct0 ct1 + then ct_eq ct0.result_typ ct0.effect_args ct1.result_typ ct1.effect_args else fail (BU.format2 "Subcomp failed: Unequal computation types %s and %s" (Ident.string_of_lid ct0.effect_name) (Ident.string_of_lid ct1.effect_name)) @@ -1528,8 +1528,8 @@ and check_binders (g_initial:env) (xs:binders) and check_comp (g:env) (c:comp) : result universe = match c.n with - | Total(t, _) - | GTotal(t, _) -> + | Total t + | GTotal t -> let! _, t = check "(G)Tot comp result" g (U.comp_result c) in is_type g t | Comp ct -> diff --git a/src/typechecker/FStar.TypeChecker.DMFF.fst b/src/typechecker/FStar.TypeChecker.DMFF.fst index 9bf72bf3296..72f4984a746 100644 --- a/src/typechecker/FStar.TypeChecker.DMFF.fst +++ b/src/typechecker/FStar.TypeChecker.DMFF.fst @@ -83,7 +83,7 @@ let gen_wps_for_free | Tm_arrow (bs, comp) -> // TODO: dubious, assert no nested arrows let rest = match comp.n with - | Total (t, _) -> t + | Total t -> t | _ -> raise_error (Error_UnexpectedDM4FType, BU.format1 "wp_a contains non-Tot arrow: %s" (Print.comp_to_string comp)) comp.pos @@ -336,8 +336,8 @@ let gen_wps_for_free | Tm_type _ -> (* BU.print2 "type0, x=%s, y=%s\n" (Print.term_to_string x) (Print.term_to_string y); *) rel x y - | Tm_arrow ([ binder ], { n = GTotal (b, _) }) - | Tm_arrow ([ binder ], { n = Total (b, _) }) -> + | Tm_arrow ([ binder ], { n = GTotal b }) + | Tm_arrow ([ binder ], { n = Total b }) -> let a = binder.binder_bv.sort in if is_monotonic a || is_monotonic b //this is an important special case; most monads have zero-order results then let a1 = S.gen_bv "a1" None a in @@ -386,8 +386,8 @@ let gen_wps_for_free | rel0 :: rels -> rel0, rels in List.fold_left U.mk_conj rel0 rels - | Tm_arrow (binders, { n = GTotal (b, _) }) - | Tm_arrow (binders, { n = Total (b, _) }) -> + | Tm_arrow (binders, { n = GTotal b }) + | Tm_arrow (binders, { n = Total b }) -> let bvs = List.mapi (fun i ({binder_bv=bv;binder_qual=q}) -> S.gen_bv ("a" ^ string_of_int i) None bv.sort) binders in let args = List.map (fun ai -> S.as_arg (S.bv_to_name ai)) bvs in let body = mk_stronger b (U.mk_app x args) (U.mk_app y args) in @@ -483,7 +483,7 @@ type nm = | N of typ | M of typ type nm_ = nm let nm_of_comp c = match c.n with - | Total (t, _) -> + | Total t -> N t | Comp c when c.flags |> BU.for_some (function CPS -> true | _ -> false) -> //lid_equals c.effect_name PC.monadic_lid -> @@ -540,7 +540,7 @@ and star_type' env t = (* Catch the GTotal case early; it seems relatively innocuous to allow * GTotal to appear. TODO fix this as a clean, single pattern-matching. *) begin match t.n with - | Tm_arrow (_, { n = GTotal (hn, _) }) -> + | Tm_arrow (_, { n = GTotal hn }) -> mk (Tm_arrow (binders, mk_GTotal (star_type' env hn))) | _ -> match is_monadic_arrow t.n with @@ -1024,7 +1024,7 @@ and infer (env: env) (e: term): nm * term * term = let is_arrow t = match (SS.compress t).n with | Tm_arrow _ -> true | _ -> false in // TODO: replace with BU.arrow_formals_comp let rec flatten t = match (SS.compress t).n with - | Tm_arrow (binders, { n = Total (t, _) }) when is_arrow t -> + | Tm_arrow (binders, { n = Total t }) when is_arrow t -> let binders', comp = flatten t in binders @ binders', comp | Tm_arrow (binders, comp) -> diff --git a/src/typechecker/FStar.TypeChecker.Env.fst b/src/typechecker/FStar.TypeChecker.Env.fst index c73d834c8e0..a9af8c38e7a 100644 --- a/src/typechecker/FStar.TypeChecker.Env.fst +++ b/src/typechecker/FStar.TypeChecker.Env.fst @@ -1084,15 +1084,20 @@ let wp_sig_aux decls m = let wp_signature env m = wp_sig_aux env.effects.decls m let comp_to_comp_typ (env:env) c = - let c = match c.n with - | Total (t, None) -> - let u = env.universe_of env t in - S.mk_Total' t (Some u) - | GTotal (t, None) -> - let u = env.universe_of env t in - S.mk_GTotal' t (Some u) - | _ -> c in - U.comp_to_comp_typ c + match c.n with + | Comp ct -> ct + | _ -> + let effect_name, result_typ = + match c.n with + | Total t -> Const.effect_Tot_lid, t + | GTotal t -> Const.effect_GTot_lid, t in + {comp_univs = [env.universe_of env result_typ]; + effect_name; + result_typ; + effect_args = []; + flags = U.comp_flags c} + +let comp_set_flags env c f = {c with n=Comp ({comp_to_comp_typ env c with flags=f})} let rec unfold_effect_abbrev env comp = let c = comp_to_comp_typ env comp in diff --git a/src/typechecker/FStar.TypeChecker.Env.fsti b/src/typechecker/FStar.TypeChecker.Env.fsti index d8d14f9bb49..f43e3ad2fc6 100644 --- a/src/typechecker/FStar.TypeChecker.Env.fsti +++ b/src/typechecker/FStar.TypeChecker.Env.fsti @@ -403,6 +403,7 @@ val get_top_level_effect : env -> lident -> option lident val is_layered_effect : env -> lident -> bool val wp_signature : env -> lident -> (bv * term) val comp_to_comp_typ : env -> comp -> comp_typ +val comp_set_flags : env -> comp -> list S.cflag -> comp val unfold_effect_abbrev : env -> comp -> comp_typ val effect_repr : env -> comp -> universe -> option term val reify_comp : env -> comp -> universe -> term diff --git a/src/typechecker/FStar.TypeChecker.Err.fst b/src/typechecker/FStar.TypeChecker.Err.fst index fcdf8c08a63..b0d988bd0c6 100644 --- a/src/typechecker/FStar.TypeChecker.Err.fst +++ b/src/typechecker/FStar.TypeChecker.Err.fst @@ -243,8 +243,8 @@ let disjunctive_pattern_vars v1 v2 = (vars v1) (vars v2))) let name_and_result c = match c.n with - | Total (t, _) -> "Tot", t - | GTotal (t, _) -> "GTot", t + | Total t -> "Tot", t + | GTotal t -> "GTot", t | Comp ct -> Print.lid_to_string ct.effect_name, ct.result_typ let computed_computation_type_does_not_match_annotation env e c c' = diff --git a/src/typechecker/FStar.TypeChecker.NBE.fst b/src/typechecker/FStar.TypeChecker.NBE.fst index f503ca9fc6b..b8297f023d0 100644 --- a/src/typechecker/FStar.TypeChecker.NBE.fst +++ b/src/typechecker/FStar.TypeChecker.NBE.fst @@ -710,8 +710,8 @@ let rec translate (cfg:config) (bs:list t) (e:term) : t = and translate_comp cfg bs (c:S.comp) : comp = match c.n with - | S.Total (typ, u) -> Tot (translate cfg bs typ, fmap_opt (translate_univ cfg bs) u) - | S.GTotal (typ, u) -> GTot (translate cfg bs typ, fmap_opt (translate_univ cfg bs) u) + | S.Total typ -> Tot (translate cfg bs typ, None) + | S.GTotal typ -> GTot (translate cfg bs typ, None) | S.Comp ctyp -> Comp (translate_comp_typ cfg bs ctyp) (* uncurried application *) @@ -974,8 +974,8 @@ and translate_constant (c : sconst) : constant = and readback_comp cfg (c: comp) : S.comp = let c' = match c with - | Tot (typ, u) -> S.Total (readback cfg typ, u) - | GTot (typ, u) -> S.GTotal (readback cfg typ, u) + | Tot (typ, _u) -> S.Total (readback cfg typ) + | GTot (typ, _u) -> S.GTotal (readback cfg typ) | Comp ctyp -> S.Comp (readback_comp_typ cfg ctyp) in S.mk c' Range.dummyRange diff --git a/src/typechecker/FStar.TypeChecker.Normalize.fst b/src/typechecker/FStar.TypeChecker.Normalize.fst index 5212057b3c7..840edaa89b7 100644 --- a/src/typechecker/FStar.TypeChecker.Normalize.fst +++ b/src/typechecker/FStar.TypeChecker.Normalize.fst @@ -554,13 +554,11 @@ and close_comp cfg env c = | [] when not <| cfg.steps.compress_uvars -> c | _ -> match c.n with - | Total (t, uopt) -> - mk_Total' (inline_closure_env cfg env [] t) - (Option.map (norm_universe cfg env) uopt) + | Total t -> + mk_Total (inline_closure_env cfg env [] t) - | GTotal (t, uopt) -> - mk_GTotal' (inline_closure_env cfg env [] t) - (Option.map (norm_universe cfg env) uopt) + | GTotal t -> + mk_GTotal (inline_closure_env cfg env [] t) | Comp c -> let rt = inline_closure_env cfg env [] c.result_typ in @@ -816,8 +814,8 @@ let should_reify cfg stack = let rec maybe_weakly_reduced tm : bool = let aux_comp c = match c.n with - | GTotal (t, _) - | Total (t, _) -> + | GTotal t + | Total t -> maybe_weakly_reduced t | Comp ct -> @@ -2139,21 +2137,13 @@ and norm_comp : cfg -> env -> comp -> comp = (Print.comp_to_string comp) (BU.string_of_int (List.length env))); match comp.n with - | Total (t, uopt) -> + | Total t -> let t = norm cfg env [] t in - let uopt = match uopt with - | Some u -> Some <| norm_universe cfg env u - | None -> None - in - { mk_Total' t uopt with pos = comp.pos } + { mk_Total t with pos = comp.pos } - | GTotal (t, uopt) -> + | GTotal t -> let t = norm cfg env [] t in - let uopt = match uopt with - | Some u -> Some <| norm_universe cfg env u - | None -> None - in - { mk_GTotal' t uopt with pos = comp.pos } + { mk_GTotal t with pos = comp.pos } | Comp ct -> //if for extraction then erase effect args to unit @@ -3063,8 +3053,8 @@ let maybe_promote_t env non_informative_only t = let ghost_to_pure_aux env non_informative_only c = match c.n with | Total _ -> c - | GTotal (t, uopt) -> - if maybe_promote_t env non_informative_only t then {c with n = Total (t, uopt)} else c + | GTotal t -> + if maybe_promote_t env non_informative_only t then {c with n = Total t} else c | Comp ct -> let l = Env.norm_eff_name env ct.effect_name in if U.is_ghost_effect l diff --git a/src/typechecker/FStar.TypeChecker.Rel.fst b/src/typechecker/FStar.TypeChecker.Rel.fst index 3ecdaa0c6f0..f7ed0de8401 100644 --- a/src/typechecker/FStar.TypeChecker.Rel.fst +++ b/src/typechecker/FStar.TypeChecker.Rel.fst @@ -1079,7 +1079,7 @@ let restrict_ctx env (tgt:ctx_uvar) (bs:binders) (src:ctx_uvar) wl : worklist = if List.length bs = 0 then aux (U.ctx_uvar_typ src) (fun src' -> src') //no abstraction over bs else begin aux - (let t = U.ctx_uvar_typ src in t |> env.universe_of env |> Some |> S.mk_Total' t |> U.arrow bs) //bs -> Tot t_t + (let t = U.ctx_uvar_typ src in t |> S.mk_Total |> U.arrow bs) //bs -> Tot t_t (fun src' -> S.mk_Tm_app //?u bs src' (bs |> S.binders_to_names |> List.map S.as_arg) @@ -1893,7 +1893,7 @@ let apply_substitutive_indexed_subcomp (env:Env.env) // apply the substitutions to subcomp_c, // and get the precondition from the PURE wp - let subcomp_ct = subcomp_c |> SS.subst_comp subst |> U.comp_to_comp_typ in + let subcomp_ct = subcomp_c |> SS.subst_comp subst |> Env.comp_to_comp_typ env in let fml = let u, wp = List.hd subcomp_ct.comp_univs, fst (List.hd subcomp_ct.effect_args) in @@ -1975,7 +1975,7 @@ let apply_ad_hoc_indexed_subcomp (env:Env.env) ps@[p], wl ) ([], wl) f_sort_is (ct1.effect_args |> List.map fst) in - let subcomp_ct = subcomp_c |> SS.subst_comp substs |> U.comp_to_comp_typ in + let subcomp_ct = subcomp_c |> SS.subst_comp substs |> Env.comp_to_comp_typ env in let g_sub_probs, wl = let g_sort_is = @@ -2442,22 +2442,16 @@ and imitate_arrow (orig:prob) (env:Env.env) (wl:worklist) let bs_lhs_args = List.map (fun ({binder_bv=x;binder_qual=i}) -> S.bv_to_name x, i) bs_lhs in let (Flex (_, u_lhs, _)) = lhs in let imitate_comp bs bs_terms c wl = - let imitate_tot_or_gtot t uopt f wl = - let k, univ = - match uopt with - | None -> - U.type_u() - | Some univ -> - S.mk (Tm_type univ) t.pos, univ - in + let imitate_tot_or_gtot t f wl = + let k, _ = U.type_u () in let _, u, wl = copy_uvar u_lhs (bs_lhs@bs) k wl in - f u (Some univ), wl + f u, wl in match c.n with - | Total (t, uopt) -> - imitate_tot_or_gtot t uopt S.mk_Total' wl - | GTotal (t, uopt) -> - imitate_tot_or_gtot t uopt S.mk_GTotal' wl + | Total t -> + imitate_tot_or_gtot t S.mk_Total wl + | GTotal t -> + imitate_tot_or_gtot t S.mk_GTotal wl | Comp ct -> let out_args, wl = List.fold_right @@ -2852,7 +2846,7 @@ and solve_t_flex_rigid_eq env (orig:prob) wl let _, lhs', wl = let b = S.null_binder t_last_arg in copy_uvar u_lhs (bs_lhs@[b]) - (t_res_lhs |> env.universe_of env |> Some |> S.mk_Total' t_res_lhs |> U.arrow [b]) wl + (t_res_lhs |> S.mk_Total |> U.arrow [b]) wl in let _, lhs'_last_arg, wl = copy_uvar u_lhs bs_lhs t_last_arg wl in lhs', lhs'_last_arg, wl @@ -4228,7 +4222,7 @@ and solve_c (env:Env.env) (problem:problem comp) (wl:worklist) : solution = let lift_c1 (edge:edge) : comp_typ * guard_t = c1 |> S.mk_Comp |> edge.mlift.mlift_wp env - |> (fun (c, g) -> U.comp_to_comp_typ c, g) in + |> (fun (c, g) -> Env.comp_to_comp_typ env c, g) in let c1, g_lift, stronger_t_opt, kind, num_eff_params, is_polymonadic = match Env.exists_polymonadic_subcomp env c1.effect_name c2.effect_name with @@ -4366,7 +4360,7 @@ and solve_c (env:Env.env) (problem:problem comp) (wl:worklist) : solution = then raise_error (Errors.Fatal_UnexpectedEffect, BU.format2 "Lift between wp-effects (%s~>%s) should not have returned a non-trivial guard" (Ident.string_of_lid c1.effect_name) (Ident.string_of_lid c2.effect_name)) r - else U.comp_to_comp_typ c) + else Env.comp_to_comp_typ env c) in if should_fail_since_repr_subcomp_not_allowed wl.repr_subcomp_allowed @@ -4455,20 +4449,20 @@ and solve_c (env:Env.env) (problem:problem comp) (wl:worklist) : solution = else N.ghost_to_pure2 env (c1, c2) in match c1.n, c2.n with - | GTotal (t1, _), Total (t2, _) when (Env.non_informative env t2) -> + | GTotal t1, Total t2 when (Env.non_informative env t2) -> solve_t env (problem_using_guard orig t1 problem.relation t2 None "result type") wl | GTotal _, Total _ -> giveup env (Thunk.mkv "incompatible monad ordering: GTot //rigid-rigid 1 + | Total t1, Total t2 + | GTotal t1, GTotal t2 -> //rigid-rigid 1 solve_t env (problem_using_guard orig t1 problem.relation t2 None "result type") wl - | Total (t1, _), GTotal (t2, _) when problem.relation = SUB -> + | Total t1, GTotal t2 when problem.relation = SUB -> solve_t env (problem_using_guard orig t1 problem.relation t2 None "result type") wl - | Total (t1, _), GTotal (t2, _) -> + | Total t1, GTotal t2 -> giveup env (Thunk.mkv "GTot =/= Tot") orig | GTotal _, Comp _ diff --git a/src/typechecker/FStar.TypeChecker.TcEffect.fst b/src/typechecker/FStar.TypeChecker.TcEffect.fst index a54d7f6e2cb..c43969a264d 100644 --- a/src/typechecker/FStar.TypeChecker.TcEffect.fst +++ b/src/typechecker/FStar.TypeChecker.TcEffect.fst @@ -434,7 +434,7 @@ let validate_indexed_effect_bind_shape (env:env) n_repr_ts (U_name u_b) (b_b.binder_bv |> S.bv_to_name) in - S.gen_bv "g" None (U.arrow [x_a] (S.mk_Total' repr (Some (new_u_univ ())))) |> S.mk_binder, + S.gen_bv "g" None (U.arrow [x_a] (S.mk_Total repr)) |> S.mk_binder, g in // return repr type p_repr ?us @@ -1223,7 +1223,7 @@ Errors.with_ctx (BU.format1 "While checking layered effect definition `%s`" (str let signature_ts = let us, t, _ = signature in (us, t) in TcUtil.layered_effect_indices_as_binders env r ed.mname signature_ts u (a.binder_bv |> S.bv_to_name) in let bs = a::rest_bs in - let k = U.arrow bs (U.type_u () |> (fun (t, u) -> S.mk_Total' t (Some (new_u_univ ())))) in //note the universe of Tot need not be u + let k = U.arrow bs (U.type_u () |> (fun (t, u) -> S.mk_Total t)) in //note the universe of Tot need not be u let g = Rel.teq env ty k in Rel.force_trivial_guard env g; (repr_us, repr_t, SS.close_univ_vars us (k |> N.remove_uvar_solutions env)) @@ -1277,7 +1277,7 @@ Errors.with_ctx (BU.format1 "While checking layered effect definition `%s`" (str | _ -> not_an_arrow_error "return" 2 ty r in let bs = a::x_a::rest_bs in let repr, g = fresh_repr r (Env.push_binders env bs) u_a (a.binder_bv |> S.bv_to_name) in - let k = U.arrow bs (S.mk_Total' repr (Some u_a)) in + let k = U.arrow bs (S.mk_Total repr) in let g_eq = Rel.teq env ty k in Rel.force_trivial_guard env (Env.conj_guard g g_eq); @@ -1630,7 +1630,7 @@ Errors.with_ctx (BU.format1 "While checking layered effect definition `%s`" (str let act_typ = match (SS.compress act.action_typ).n with | Tm_arrow (bs, c) -> - let ct = U.comp_to_comp_typ c in + let ct = Env.comp_to_comp_typ env c in if lid_equals ct.effect_name ed.mname then let repr_ts = let us, t, _ = repr in (us, t) in @@ -1639,7 +1639,7 @@ Errors.with_ctx (BU.format1 "While checking layered effect definition `%s`" (str repr (S.as_arg ct.result_typ::ct.effect_args) r in - let c = S.mk_Total' repr (Some (new_u_univ ())) in + let c = S.mk_Total repr in U.arrow bs c else act.action_typ | _ -> act.action_typ in @@ -1664,7 +1664,7 @@ Errors.with_ctx (BU.format1 "While checking layered effect definition `%s`" (str (string_of_lid ed.mname) (string_of_lid act.action_name) in let a_tm, _, g_tm = TcUtil.new_implicit_var reason r env t in let repr, g = fresh_repr r env u a_tm in - U.arrow bs (S.mk_Total' repr (Env.new_u_univ () |> Some)), Env.conj_guard g g_tm + U.arrow bs (S.mk_Total repr), Env.conj_guard g g_tm | _ -> raise_error (Errors.Fatal_ActionMustHaveFunctionType, BU.format3 "Unexpected non-function type for action %s:%s (%s)" (string_of_lid ed.mname) (string_of_lid act.action_name) (Print.term_to_string act_typ)) r in @@ -1823,7 +1823,7 @@ Errors.with_ctx (BU.format1 "While checking effect definition `%s`" (string_of_l //generalize the universes in bs //bs are closed with us and closed let us, bs = - let tmp_t = U.arrow bs (S.mk_Total' S.t_unit (U_zero |> Some)) in //create a temporary bs -> Tot unit + let tmp_t = U.arrow bs (S.mk_Total S.t_unit) in //create a temporary bs -> Tot unit let us, tmp_t = Gen.generalize_universes env0 tmp_t in us, tmp_t |> U.arrow_formals |> fst |> SS.close_binders in @@ -2103,7 +2103,7 @@ Errors.with_ctx (BU.format1 "While checking effect definition `%s`" (string_of_l let act_typ = match (SS.compress act.action_typ).n with | Tm_arrow (bs, c) -> - let c = U.comp_to_comp_typ c in + let c = Env.comp_to_comp_typ env c in if lid_equals c.effect_name ed.mname then U.arrow bs (S.mk_Total (mk_repr' c.result_typ (fst (List.hd c.effect_args)))) else act.action_typ diff --git a/src/typechecker/FStar.TypeChecker.TcTerm.fst b/src/typechecker/FStar.TypeChecker.TcTerm.fst index 4db65ceb901..4659eab278f 100644 --- a/src/typechecker/FStar.TypeChecker.TcTerm.fst +++ b/src/typechecker/FStar.TypeChecker.TcTerm.fst @@ -866,7 +866,7 @@ and tc_maybe_toplevel_term env (e:term) : term (* type-checked let repr = Env.effect_repr env0 (expected_ct |> S.mk_Comp) u_c |> must in // e <: Tot repr - let e = S.mk (Tm_ascribed (e, (Inr (S.mk_Total' repr (Some u_c)), None, use_eq), None)) e.pos in + let e = S.mk (Tm_ascribed (e, (Inr (S.mk_Total repr), None, use_eq), None)) e.pos in if Env.debug env0 <| Options.Extreme then BU.print1 "Typechecking ascribed reflect, inner ascribed term: %s\n" @@ -1477,7 +1477,7 @@ and tc_match (env : Env.env) (top : term) : term * lcomp * guard_t = if erasable then (* promote cres to ghost *) let e = U.exp_true_bool in - let c = mk_GTotal' U.t_bool (Some U_zero) in + let c = mk_GTotal U.t_bool in TcUtil.bind e.pos env (Some e) (TcComm.lcomp_of_comp c) (None, cres) else cres in @@ -1837,15 +1837,15 @@ and tc_comp env c : comp (* checked version * guard_t = (* logical guard for the well-formedness of c *) let c0 = c in match c.n with - | Total (t, _) -> + | Total t -> let k, u = U.type_u () in let t, _, g = tc_check_tot_or_gtot_term env t k "" in - mk_Total' t (Some u), u, g + mk_Total t, u, g - | GTotal (t, _) -> + | GTotal t -> let k, u = U.type_u () in let t, _, g = tc_check_tot_or_gtot_term env t k "" in - mk_GTotal' t (Some u), u, g + mk_GTotal t, u, g | Comp c -> let head = S.fvar c.effect_name delta_constant None in @@ -3791,7 +3791,7 @@ and check_top_level_let env e = * * Note that for top-level lets, this cres is not used anyway *) - let cres = S.mk_Total' S.t_unit (Some S.U_zero) in + let cres = S.mk_Total S.t_unit in (*close*)let lb = U.close_univs_and_mk_letbinding None lb.lbname univ_vars (U.comp_result c1) (U.comp_effect_name c1) e1 lb.lbattrs lb.lbpos in mk (Tm_let((false, [lb]), e2)) @@ -4649,9 +4649,9 @@ let rec typeof_tot_or_gtot_term_fastpath (env:env) (t:term) (must_tot:bool) : op | Tm_abs(bs, body, Some ({residual_effect=eff; residual_typ=tbody})) -> //AR: maybe keep residual univ too? let mk_comp = if Ident.lid_equals eff Const.effect_Tot_lid - then Some S.mk_Total' + then Some S.mk_Total else if Ident.lid_equals eff Const.effect_GTot_lid - then Some S.mk_GTotal' + then Some S.mk_GTotal else None in bind_opt mk_comp (fun f -> @@ -4664,7 +4664,7 @@ let rec typeof_tot_or_gtot_term_fastpath (env:env) (t:term) (must_tot:bool) : op bind_opt tbody (fun tbody -> let bs, tbody = SS.open_term bs tbody in let u = universe_of (Env.push_binders env bs) tbody in - Some (U.arrow bs (f tbody (Some u))))) + Some (U.arrow bs (f tbody)))) | Tm_abs _ -> None diff --git a/src/typechecker/FStar.TypeChecker.Util.fst b/src/typechecker/FStar.TypeChecker.Util.fst index 9c7f102d029..740af919c1b 100644 --- a/src/typechecker/FStar.TypeChecker.Util.fst +++ b/src/typechecker/FStar.TypeChecker.Util.fst @@ -207,7 +207,7 @@ let extract_let_rec_annotation env {lbname=lbname; lbunivs=univ_vars; lbtyp=t; l let bs, c = un_arrow tarr in (match get_decreases c with | Some (pfx, DECREASES d, sfx) -> - let c = U.comp_set_flags c (pfx @ sfx) in + let c = Env.comp_set_flags env c (pfx @ sfx) in U.arrow bs c, tarr, true | _ -> tarr, tarr, true) @@ -222,9 +222,9 @@ let extract_let_rec_annotation env {lbname=lbname; lbunivs=univ_vars; lbtyp=t; l let s = U.rename_binders bs bs' in SS.subst_decreasing_order s d in - let c = U.comp_set_flags c flags in + let c = Env.comp_set_flags env c flags in let tarr = U.arrow bs c in - let c' = U.comp_set_flags c' (DECREASES d'::flags') in + let c' = Env.comp_set_flags env c' (DECREASES d'::flags') in let tannot = U.arrow bs' c' in tarr, tannot, true in @@ -460,8 +460,7 @@ let extract_let_rec_annotation env {lbname=lbname; lbunivs=univ_vars; lbtyp=t; l let comp_univ_opt c = match c.n with - | Total (_, uopt) - | GTotal (_, uopt) -> uopt + | Total _ | GTotal _ -> None | Comp c -> match c.comp_univs with | [] -> None @@ -491,7 +490,7 @@ let effect_args_from_repr (repr:term) (is_layered:bool) (r:Range.range) : list t | Tm_app (_, _::is) -> is |> List.map fst | _ -> err () else match repr.n with - | Tm_arrow (_, c) -> c |> U.comp_to_comp_typ |> (fun ct -> ct.effect_args |> List.map fst) + | Tm_arrow (_, c) -> c |> U.comp_eff_name_res_and_args |> (fun (_, _, args) -> args |> List.map fst) | _ -> err () @@ -506,7 +505,7 @@ let mk_wp_return env (ed:S.eff_decl) (u_a:universe) (a:typ) (e:term) (r:Range.ra if not <| Env.lid_exists env C.effect_GTot_lid //we're still in prims, not yet having fully defined the primitive effects then mk_Total a else if U.is_unit a - then S.mk_Total' a (Some U_zero) + then S.mk_Total a else let wp = if env.lax && Options.ml_ish() //NS: Disabling this optimization temporarily @@ -615,7 +614,7 @@ let is_pure_or_ghost_effect env l = let lax_mk_tot_or_comp_l mname u_result result flags = if Ident.lid_equals mname C.effect_Tot_lid - then S.mk_Total' result (Some u_result) + then S.mk_Total result else mk_comp_l mname u_result result S.tun flags let is_function t = match (compress t).n with @@ -1004,11 +1003,11 @@ let mk_indexed_return env (ed:S.eff_decl) (u_a:universe) (a:typ) (e:term) (r:Ran "%s.return %s does not have proper shape (reason:%s)" (Ident.string_of_lid ed.mname) (Print.term_to_string return_t) s) in - let a_b, x_b, rest_bs, return_ct = + let a_b, x_b, rest_bs, return_typ = match (SS.compress return_t).n with | Tm_arrow (bs, c) when List.length bs >= 2 -> let ((a_b::x_b::bs, c)) = SS.open_comp bs c in - a_b, x_b, bs, U.comp_to_comp_typ c + a_b, x_b, bs, U.comp_result c | _ -> raise_error (return_t_shape_error "Either not an arrow or not enough binders") r in let rest_bs_uvars, g_uvars = @@ -1027,7 +1026,7 @@ let mk_indexed_return env (ed:S.eff_decl) (u_a:universe) (a:typ) (e:term) (r:Ran (a_b::x_b::rest_bs) (a::e::rest_bs_uvars) in let is = - effect_args_from_repr (SS.compress return_ct.result_typ) (U.is_layered ed) r + effect_args_from_repr (SS.compress return_typ) (U.is_layered ed) r |> List.map (SS.subst subst) in let c = mk_Comp ({ @@ -1093,7 +1092,7 @@ let mk_indexed_bind env substitutive_indexed_bind_substs env m_ed n_ed p_ed bind_t_bs binder_kinds ct1 b ct2 r1 num_effect_params has_range_binders in - let bind_ct = bind_c |> SS.subst_comp subst |> U.comp_to_comp_typ in + let bind_ct = bind_c |> SS.subst_comp subst |> Env.comp_to_comp_typ env in //compute the formula `bind_c.wp (fun _ -> True)` and add it to the final guard let fml = @@ -1176,7 +1175,7 @@ let mk_bind env * so it's fine to return g_return as is *) let m, c1, c2, g_lift = lift_comps env c1 c2 b true in - let ct1, ct2 = U.comp_to_comp_typ c1, U.comp_to_comp_typ c2 in + let ct1, ct2 = Env.comp_to_comp_typ env c1, Env.comp_to_comp_typ env c2 in let c, g_bind = if Env.is_layered_effect env m @@ -1708,10 +1707,10 @@ let assume_result_eq_pure_term_in_m env (m_opt:option lident) (e:term) (lc:lcomp let retc, g_retc = return_value env m (Some u_t) (U.comp_result c) e in let g_c = Env.conj_guard g_c g_retc in if not (U.is_pure_comp c) //it started in GTot, so it should end up in Ghost - then let retc = U.comp_to_comp_typ retc in + then let retc = Env.comp_to_comp_typ env retc in let retc = {retc with effect_name=C.effect_GHOST_lid; flags=flags} in S.mk_Comp retc, g_c - else U.comp_set_flags retc flags, g_c + else Env.comp_set_flags env retc flags, g_c else //AR: augment c's post-condition with a M.return let c = Env.unfold_effect_abbrev env c in let t = c.result_typ in @@ -1720,11 +1719,11 @@ let assume_result_eq_pure_term_in_m env (m_opt:option lident) (e:term) (lc:lcomp let xexp = S.bv_to_name x in let env_x = Env.push_bv env x in let ret, g_ret = return_value env_x m (Some u_t) t xexp in - let ret = TcComm.lcomp_of_comp <| U.comp_set_flags ret [PARTIAL_RETURN] in + let ret = TcComm.lcomp_of_comp <| Env.comp_set_flags env ret [PARTIAL_RETURN] in let eq = U.mk_eq2 u_t t xexp e in let eq_ret = weaken_precondition env_x ret (NonTrivial eq) in let bind_c, g_bind = TcComm.lcomp_comp (bind e.pos env None (TcComm.lcomp_of_comp c) (Some x, eq_ret)) in - U.comp_set_flags bind_c flags, Env.conj_guards [g_c; g_ret; g_bind] + Env.comp_set_flags env bind_c flags, Env.conj_guards [g_c; g_ret; g_bind] in if should_not_inline_lc lc @@ -2144,7 +2143,7 @@ let bind_cases env0 (res_t:typ) lift_comps_sep_guards env cthen celse None false in let md = Env.get_effect_decl env m in md, - cthen |> U.comp_to_comp_typ, celse |> U.comp_to_comp_typ, + cthen |> Env.comp_to_comp_typ env, celse |> Env.comp_to_comp_typ env, g_lift_then, g_lift_else in //function to apply the if-then-else combinator @@ -3467,7 +3466,7 @@ let lift_tf_layered_effect (tgt:lident) (lift_ts:tscheme) (kind:S.indexed_effect let r = Env.get_range env in - let ct = U.comp_to_comp_typ c in + let ct = Env.comp_to_comp_typ env c in check_non_informative_type_for_lift env ct.effect_name tgt ct.result_typ r; @@ -3484,7 +3483,7 @@ let lift_tf_layered_effect (tgt:lident) (lift_ts:tscheme) (kind:S.indexed_effect then ad_hoc_indexed_lift_substs env bs ct (lift_name ()) r else substitutive_indexed_lift_substs env bs ct (lift_name ()) r in - let lift_ct = lift_c |> SS.subst_comp substs |> U.comp_to_comp_typ in + let lift_ct = lift_c |> SS.subst_comp substs |> Env.comp_to_comp_typ env in let is = effect_args_from_repr lift_ct.result_typ (Env.is_layered_effect env tgt) r in @@ -3561,7 +3560,7 @@ let get_mlift_for_subeff env (sub:S.sub_eff) : Env.mlift = else let mk_mlift_wp ts env c = - let ct = U.comp_to_comp_typ c in + let ct = Env.comp_to_comp_typ env c in check_non_informative_type_for_lift env ct.effect_name sub.target ct.result_typ env.range; let _, lift_t = inst_tscheme_with ts ct.comp_univs in let wp = List.hd ct.effect_args in From 6688d489a4832361070f3622d7783bbe98368f4d Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Fri, 9 Dec 2022 10:06:41 +0000 Subject: [PATCH 03/18] ulib reflection changes, also make decreases as part of C_Eff --- ulib/FStar.Reflection.Compare.fst | 17 ++++----- ulib/FStar.Reflection.Data.fsti | 5 +-- ulib/FStar.Reflection.Derived.Lemmas.fst | 10 +++--- ulib/FStar.Reflection.Derived.fst | 6 ++-- ulib/FStar.Tactics.Derived.fst | 2 +- ulib/FStar.Tactics.Print.fst | 6 ++-- ulib/FStar.Tactics.SyntaxHelpers.fst | 12 +++---- ulib/FStar.Tactics.Typeclasses.fst | 2 +- ulib/FStar.Tactics.Visit.fst | 15 ++++---- .../FStar.InteractiveHelpers.Base.fst | 36 +++++++++---------- .../FStar.InteractiveHelpers.Effectful.fst | 6 ++-- .../FStar.InteractiveHelpers.ExploreTerm.fst | 14 ++++---- ulib/experimental/Steel.Effect.Common.fsti | 15 ++++---- 13 files changed, 69 insertions(+), 77 deletions(-) diff --git a/ulib/FStar.Reflection.Compare.fst b/ulib/FStar.Reflection.Compare.fst index 96a34a0329b..1fa8f734a29 100644 --- a/ulib/FStar.Reflection.Compare.fst +++ b/ulib/FStar.Reflection.Compare.fst @@ -174,12 +174,9 @@ and compare_comp (c1 c2 : comp) : Tot order (decreases c1) = let cv1 = inspect_comp c1 in let cv2 = inspect_comp c2 in match cv1, cv2 with - | C_Total t1 u1 md1, C_Total t2 u2 md2 + | C_Total t1, C_Total t2 - | C_GTotal t1 u1 md1, C_GTotal t2 u2 md2 -> - lex (compare_term t1 t2) - (fun _ -> lex (compare_universe u1 u2) - (fun _ -> compare_term_list md1 md2)) + | C_GTotal t1, C_GTotal t2 -> compare_term t1 t2 | C_Lemma p1 q1 s1, C_Lemma p2 q2 s2 -> lex (compare_term p1 p2) @@ -188,14 +185,14 @@ and compare_comp (c1 c2 : comp) : Tot order (decreases c1) = (fun () -> compare_term s1 s2) ) - | C_Eff us1 eff1 res1 args1, - C_Eff us2 eff2 res2 args2 -> + | C_Eff us1 eff1 res1 args1 _decrs1, + C_Eff us2 eff2 res2 args2 _decrs2 -> (* This could be more complex, not sure it is worth it *) lex (compare_universes us1 us2) (fun _ -> lex (compare_name eff1 eff2) (fun _ -> compare_term res1 res2)) - | C_Total _ _ _, _ -> Lt | _, C_Total _ _ _ -> Gt - | C_GTotal _ _ _, _ -> Lt | _, C_GTotal _ _ _ -> Gt + | C_Total _, _ -> Lt | _, C_Total _ -> Gt + | C_GTotal _, _ -> Lt | _, C_GTotal _ -> Gt | C_Lemma _ _ _, _ -> Lt | _, C_Lemma _ _ _ -> Gt - | C_Eff _ _ _ _, _ -> Lt | _, C_Eff _ _ _ _ -> Gt + | C_Eff _ _ _ _ _, _ -> Lt | _, C_Eff _ _ _ _ _ -> Gt diff --git a/ulib/FStar.Reflection.Data.fsti b/ulib/FStar.Reflection.Data.fsti index 26974934b43..7daeb7a04a5 100644 --- a/ulib/FStar.Reflection.Data.fsti +++ b/ulib/FStar.Reflection.Data.fsti @@ -95,13 +95,14 @@ let notAscription (tv:term_view) : bool = // Very basic for now noeq type comp_view = - | C_Total : ret:typ -> u:universe -> decr:(list term) -> comp_view - | C_GTotal : ret:typ -> u:universe -> decr:(list term) -> comp_view + | C_Total : ret:typ -> comp_view + | C_GTotal : ret:typ -> comp_view | C_Lemma : term -> term -> term -> comp_view // pre, post, patterns | C_Eff : us:universes -> eff_name:name -> result:term -> eff_args:(list argv) -> + decrs:list term -> comp_view (* Constructor for an inductive type. See explanation in diff --git a/ulib/FStar.Reflection.Derived.Lemmas.fst b/ulib/FStar.Reflection.Derived.Lemmas.fst index eef795630bb..84e36711e74 100644 --- a/ulib/FStar.Reflection.Derived.Lemmas.fst +++ b/ulib/FStar.Reflection.Derived.Lemmas.fst @@ -105,7 +105,7 @@ let rec collect_arr_order' (bds: binders) (tt: term) (c: comp) bds' <<: tt /\ c' << tt)) (decreases c) = match inspect_comp c with - | C_Total ret _ _ -> + | C_Total ret -> ( match inspect_ln_unascribe ret with | Tv_Arrow b c -> collect_arr_order' (b::bds) tt c | _ -> ()) @@ -115,17 +115,17 @@ val collect_arr_ln_bs_order : (t:term) -> Lemma (ensures forall bds c. (bds, c) == collect_arr_ln_bs t ==> (c << t /\ bds <<: t) - \/ (c == pack_comp (C_Total t u_unk []) /\ bds == []) + \/ (c == pack_comp (C_Total t) /\ bds == []) ) let collect_arr_ln_bs_order t = match inspect_ln_unascribe t with | Tv_Arrow b c -> collect_arr_order' [b] t c; Classical.forall_intro_2 (rev_memP #binder); - inspect_pack_comp_inv (C_Total t u_unk []) - | _ -> inspect_pack_comp_inv (C_Total t u_unk []) + inspect_pack_comp_inv (C_Total t) + | _ -> inspect_pack_comp_inv (C_Total t) val collect_arr_ln_bs_ref : (t:term) -> list (bd:binder{bd << t}) - * (c:comp{ c == pack_comp (C_Total t u_unk []) \/ c << t}) + * (c:comp{ c == pack_comp (C_Total t) \/ c << t}) let collect_arr_ln_bs_ref t = let bds, c = collect_arr_ln_bs t in collect_arr_ln_bs_order t; diff --git a/ulib/FStar.Reflection.Derived.fst b/ulib/FStar.Reflection.Derived.fst index d48489582d9..ba553bf69dd 100644 --- a/ulib/FStar.Reflection.Derived.fst +++ b/ulib/FStar.Reflection.Derived.fst @@ -91,11 +91,11 @@ let u_unk : universe = pack_universe Uv_Unk let rec mk_tot_arr_ln (bs: list binder) (cod : term) : Tot term (decreases bs) = match bs with | [] -> cod - | (b::bs) -> pack_ln (Tv_Arrow b (pack_comp (C_Total (mk_tot_arr_ln bs cod) u_unk []))) + | (b::bs) -> pack_ln (Tv_Arrow b (pack_comp (C_Total (mk_tot_arr_ln bs cod)))) let rec collect_arr' (bs : list binder) (c : comp) : Tot (list binder * comp) (decreases c) = begin match inspect_comp c with - | C_Total t _ _ -> + | C_Total t -> begin match inspect_ln_unascribe t with | Tv_Arrow b c -> collect_arr' (b::bs) c @@ -107,7 +107,7 @@ let rec collect_arr' (bs : list binder) (c : comp) : Tot (list binder * comp) (d val collect_arr_ln_bs : typ -> list binder * comp let collect_arr_ln_bs t = - let (bs, c) = collect_arr' [] (pack_comp (C_Total t u_unk [])) in + let (bs, c) = collect_arr' [] (pack_comp (C_Total t)) in (List.Tot.Base.rev bs, c) val collect_arr_ln : typ -> list typ * comp diff --git a/ulib/FStar.Tactics.Derived.fst b/ulib/FStar.Tactics.Derived.fst index 28bef2c9c1a..6867a08cd32 100644 --- a/ulib/FStar.Tactics.Derived.fst +++ b/ulib/FStar.Tactics.Derived.fst @@ -701,7 +701,7 @@ let rec apply_squash_or_lem d t = | _ -> fail "mapply: can't apply (1)" end - | C_Total rt _ _ -> + | C_Total rt -> begin match unsquash_term rt with (* If the function returns a squash, just apply it, since our goals are squashed *) | Some rt -> diff --git a/ulib/FStar.Tactics.Print.fst b/ulib/FStar.Tactics.Print.fst index 0b1909e784f..79d4249feb4 100644 --- a/ulib/FStar.Tactics.Print.fst +++ b/ulib/FStar.Tactics.Print.fst @@ -91,10 +91,10 @@ and branch_to_ast_string (b:branch) : Tac string = and comp_to_ast_string (c:comp) : Tac string = match inspect_comp c with - | C_Total t u _ -> "Tot" ^ universe_to_ast_string u ^ term_to_ast_string t - | C_GTotal t u _ -> "GTot" ^ universe_to_ast_string u ^ term_to_ast_string t + | C_Total t -> "Tot " ^ term_to_ast_string t + | C_GTotal t -> "GTot " ^ term_to_ast_string t | C_Lemma pre post _ -> "Lemma " ^ term_to_ast_string pre ^ " " ^ term_to_ast_string post - | C_Eff us eff res _ -> + | C_Eff us eff res _ _ -> "Effect" ^ "<" ^ universes_to_ast_string us ^ "> " ^ paren (implode_qn eff ^ ", " ^ term_to_ast_string res) and const_to_ast_string (c:vconst) : Tac string = diff --git a/ulib/FStar.Tactics.SyntaxHelpers.fst b/ulib/FStar.Tactics.SyntaxHelpers.fst index 7119897e224..00d05ce2942 100644 --- a/ulib/FStar.Tactics.SyntaxHelpers.fst +++ b/ulib/FStar.Tactics.SyntaxHelpers.fst @@ -10,7 +10,7 @@ open FStar.Tactics.Types private let rec collect_arr' (bs : list binder) (c : comp) : Tac (list binder * comp) = begin match inspect_comp c with - | C_Total t _ _ -> + | C_Total t -> begin match inspect t with | Tv_Arrow b c -> collect_arr' (b::bs) c @@ -22,12 +22,12 @@ let rec collect_arr' (bs : list binder) (c : comp) : Tac (list binder * comp) = val collect_arr_bs : typ -> Tac (list binder * comp) let collect_arr_bs t = - let (bs, c) = collect_arr' [] (pack_comp (C_Total t u_unk [])) in + let (bs, c) = collect_arr' [] (pack_comp (C_Total t)) in (List.Tot.Base.rev bs, c) val collect_arr : typ -> Tac (list typ * comp) let collect_arr t = - let (bs, c) = collect_arr' [] (pack_comp (C_Total t u_unk [])) in + let (bs, c) = collect_arr' [] (pack_comp (C_Total t)) in let ts = List.Tot.Base.map type_of_binder bs in (List.Tot.Base.rev ts, c) @@ -52,19 +52,19 @@ let rec mk_arr (bs: list binder) (cod : comp) : Tac term = | [] -> fail "mk_arr, empty binders" | [b] -> pack (Tv_Arrow b cod) | (b::bs) -> - pack (Tv_Arrow b (pack_comp (C_Total (mk_arr bs cod) u_unk []))) + pack (Tv_Arrow b (pack_comp (C_Total (mk_arr bs cod)))) let rec mk_arr_curried (bs: list binder) (cod : comp) : Tac term = match bs with | [] -> fail "mk_arr, empty binders" | [b] -> pack_curried (Tv_Arrow b cod) - | (b::bs) -> pack_curried (Tv_Arrow b (pack_comp (C_Total (mk_arr_curried bs cod) u_unk []))) + | (b::bs) -> pack_curried (Tv_Arrow b (pack_comp (C_Total (mk_arr_curried bs cod)))) let rec mk_tot_arr (bs: list binder) (cod : term) : Tac term = match bs with | [] -> cod | (b::bs) -> - pack (Tv_Arrow b (pack_comp (C_Total (mk_tot_arr bs cod) u_unk []))) + pack (Tv_Arrow b (pack_comp (C_Total (mk_tot_arr bs cod)))) let lookup_lb_view (lbs:list letbinding) (nm:name) : Tac lb_view = let o = FStar.List.Tot.Base.find diff --git a/ulib/FStar.Tactics.Typeclasses.fst b/ulib/FStar.Tactics.Typeclasses.fst index 91f358643c3..ffac34ab991 100644 --- a/ulib/FStar.Tactics.Typeclasses.fst +++ b/ulib/FStar.Tactics.Typeclasses.fst @@ -133,7 +133,7 @@ let mk_class (nm:string) : Tac decls = let bs, cod = collect_arr_bs ty in let r = inspect_comp cod in guard (C_Total? r); - let C_Total cod _ _ = r in (* must be total *) + let C_Total cod = r in (* must be total *) (* print ("n_univs = " ^ string_of_int (List.Tot.Base.length us)); *) diff --git a/ulib/FStar.Tactics.Visit.fst b/ulib/FStar.Tactics.Visit.fst index 34a34661eb1..87676599407 100644 --- a/ulib/FStar.Tactics.Visit.fst +++ b/ulib/FStar.Tactics.Visit.fst @@ -118,15 +118,13 @@ and visit_comp (ff : term -> Tac term) (c : comp) : Tac comp = let cv = inspect_comp c in let cv' = match cv with - | C_Total ret uopt decr -> + | C_Total ret -> let ret = visit_tm ff ret in - let decr = map (visit_tm ff) decr in - C_Total ret uopt decr + C_Total ret - | C_GTotal ret uopt decr -> + | C_GTotal ret -> let ret = visit_tm ff ret in - let decr = map (visit_tm ff) decr in - C_GTotal ret uopt decr + C_GTotal ret | C_Lemma pre post pats -> let pre = visit_tm ff pre in @@ -134,9 +132,10 @@ and visit_comp (ff : term -> Tac term) (c : comp) : Tac comp = let pats = visit_tm ff pats in C_Lemma pre post pats - | C_Eff us eff res args -> + | C_Eff us eff res args decrs -> let res = visit_tm ff res in let args = map (fun (a, q) -> (visit_tm ff a, q)) args in - C_Eff us eff res args + let decrs = map (visit_tm ff) decrs in + C_Eff us eff res args decrs in pack_comp cv' diff --git a/ulib/experimental/FStar.InteractiveHelpers.Base.fst b/ulib/experimental/FStar.InteractiveHelpers.Base.fst index cebf74a026c..6c2a655b4ef 100644 --- a/ulib/experimental/FStar.InteractiveHelpers.Base.fst +++ b/ulib/experimental/FStar.InteractiveHelpers.Base.fst @@ -111,13 +111,13 @@ let print_binders_info (full : bool) (e:env) : Tac unit = let acomp_to_string (c:comp) : Tac string = match inspect_comp c with - | C_Total ret _ decr -> + | C_Total ret -> "C_Total (" ^ term_to_string ret ^ ")" - | C_GTotal ret _ decr -> + | C_GTotal ret -> "C_GTotal (" ^ term_to_string ret ^ ")" | C_Lemma pre post patterns -> "C_Lemma (" ^ term_to_string pre ^ ") (" ^ term_to_string post ^ ")" - | C_Eff us eff_name result eff_args -> + | C_Eff us eff_name result eff_args _ -> let eff_arg_to_string (a : term) : Tac string = " (" ^ term_to_string a ^ ")" in @@ -398,23 +398,22 @@ let norm_apply_subst_in_comp e c subst = | Q_Meta t -> Q_Meta (subst t) in match inspect_comp c with - | C_Total ret uopt decr -> + | C_Total ret -> let ret = subst ret in - let decr = map subst decr in - pack_comp (C_Total ret uopt decr) - | C_GTotal ret uopt decr -> + pack_comp (C_Total ret) + | C_GTotal ret -> let ret = subst ret in - let decr = map subst decr in - pack_comp (C_GTotal ret uopt decr) + pack_comp (C_GTotal ret) | C_Lemma pre post patterns -> let pre = subst pre in let post = subst post in let patterns = subst patterns in pack_comp (C_Lemma pre post patterns) - | C_Eff us eff_name result eff_args -> + | C_Eff us eff_name result eff_args decrs -> let result = subst result in let eff_args = map (fun (x, a) -> (subst x, subst_in_aqualv a)) eff_args in - pack_comp (C_Eff us eff_name result eff_args) + let decrs = map subst decrs in + pack_comp (C_Eff us eff_name result eff_args decrs) /// As substitution with normalization is very expensive, we implemented another /// technique which works by exploring terms. This is super fast, but the terms @@ -531,23 +530,22 @@ and deep_apply_subst_in_comp e c subst = | Q_Meta t -> Q_Meta (subst t) in match inspect_comp c with - | C_Total ret uopt decr -> + | C_Total ret -> let ret = subst ret in - let decr = map subst decr in - pack_comp (C_Total ret uopt decr) - | C_GTotal ret uopt decr -> + pack_comp (C_Total ret) + | C_GTotal ret -> let ret = subst ret in - let decr = map subst decr in - pack_comp (C_GTotal ret uopt decr) + pack_comp (C_GTotal ret) | C_Lemma pre post patterns -> let pre = subst pre in let post = subst post in let patterns = subst patterns in pack_comp (C_Lemma pre post patterns) - | C_Eff us eff_name result eff_args -> + | C_Eff us eff_name result eff_args decrs -> let result = subst result in let eff_args = map (fun (x, a) -> (subst x, subst_in_aqualv a)) eff_args in - pack_comp (C_Eff us eff_name result eff_args) + let decrs = map subst decrs in + pack_comp (C_Eff us eff_name result eff_args decrs) and deep_apply_subst_in_pattern e pat subst = match pat with diff --git a/ulib/experimental/FStar.InteractiveHelpers.Effectful.fst b/ulib/experimental/FStar.InteractiveHelpers.Effectful.fst index b02094222d7..2c3771def61 100644 --- a/ulib/experimental/FStar.InteractiveHelpers.Effectful.fst +++ b/ulib/experimental/FStar.InteractiveHelpers.Effectful.fst @@ -113,10 +113,10 @@ val comp_view_to_effect_info : dbg:bool -> comp_view -> Tac (option effect_info) let comp_view_to_effect_info dbg cv = match cv with - | C_Total ret_ty _ decr -> + | C_Total ret_ty -> let ret_type_info = get_type_info_from_type ret_ty in Some (mk_effect_info E_Total ret_type_info None None) - | C_GTotal ret_ty _ decr -> + | C_GTotal ret_ty -> let ret_type_info = get_type_info_from_type ret_ty in Some (mk_effect_info E_Total ret_type_info None None) | C_Lemma pre post patterns -> @@ -124,7 +124,7 @@ let comp_view_to_effect_info dbg cv = let pre = prettify_term dbg pre in let post = prettify_term dbg post in Some (mk_effect_info E_Lemma unit_type_info (Some pre) (Some post)) - | C_Eff univs eff_name ret_ty eff_args -> + | C_Eff univs eff_name ret_ty eff_args _ -> print_dbg dbg ("comp_view_to_effect_info: C_Eff " ^ flatten_name eff_name); let ret_type_info = get_type_info_from_type ret_ty in let etype = effect_name_to_type eff_name in diff --git a/ulib/experimental/FStar.InteractiveHelpers.ExploreTerm.fst b/ulib/experimental/FStar.InteractiveHelpers.ExploreTerm.fst index c2b90ca3bcd..10855fe22d7 100644 --- a/ulib/experimental/FStar.InteractiveHelpers.ExploreTerm.fst +++ b/ulib/experimental/FStar.InteractiveHelpers.ExploreTerm.fst @@ -29,10 +29,10 @@ val comp_qualifier (c : comp) : Tac string #push-options "--ifuel 1" let comp_qualifier (c : comp) : Tac string = match inspect_comp c with - | C_Total _ _ _ -> "C_Total" - | C_GTotal _ _ _ -> "C_GTotal" + | C_Total _ -> "C_Total" + | C_GTotal _ -> "C_GTotal" | C_Lemma _ _ _ -> "C_Lemma" - | C_Eff _ _ _ _ -> "C_Eff" + | C_Eff _ _ _ _ _ -> "C_Eff" #pop-options /// Effect information: we list the current supported effects @@ -118,14 +118,14 @@ let get_type_info (e:env) (t:term) : Tac (option type_info) = val get_total_or_gtotal_ret_type : comp -> Tot (option typ) let get_total_or_gtotal_ret_type c = match inspect_comp c with - | C_Total ret_ty _ _ | C_GTotal ret_ty _ _ -> Some ret_ty + | C_Total ret_ty | C_GTotal ret_ty -> Some ret_ty | _ -> None val get_comp_ret_type : comp -> Tot typ let get_comp_ret_type c = match inspect_comp c with - | C_Total ret_ty _ _ | C_GTotal ret_ty _ _ - | C_Eff _ _ ret_ty _ -> ret_ty + | C_Total ret_ty | C_GTotal ret_ty + | C_Eff _ _ ret_ty _ _ -> ret_ty | C_Lemma _ _ _ -> (`unit) val is_total_or_gtotal : comp -> Tot bool @@ -371,7 +371,7 @@ let flush_typ_or_comp dbg e tyc = in try begin match tyc with | TC_Typ ty pl n -> - let c = pack_comp (C_Total ty u_unk []) in + let c = pack_comp (C_Total ty) in flush_comp pl n c | TC_Comp c pl n -> flush_comp pl n c end diff --git a/ulib/experimental/Steel.Effect.Common.fsti b/ulib/experimental/Steel.Effect.Common.fsti index 76235b0b55b..c799fe0663b 100644 --- a/ulib/experimental/Steel.Effect.Common.fsti +++ b/ulib/experimental/Steel.Effect.Common.fsti @@ -707,21 +707,18 @@ and visit_br (ff : term -> Tac unit) (b:branch) : Tac unit = and visit_comp (ff : term -> Tac unit) (c : comp) : Tac unit = let cv = inspect_comp c in match cv with - | C_Total ret _ decr -> - visit_tm ff ret; - iter (visit_tm ff) decr - | C_GTotal ret _ decr -> - visit_tm ff ret; - iter (visit_tm ff) decr + | C_Total ret -> visit_tm ff ret + | C_GTotal ret -> visit_tm ff ret | C_Lemma pre post pats -> visit_tm ff pre; visit_tm ff post; visit_tm ff pats - | C_Eff us eff res args -> + | C_Eff us eff res args decrs -> visit_tm ff res; - iter (fun (a, q) -> visit_tm ff a) args + iter (fun (a, q) -> visit_tm ff a) args; + iter (visit_tm ff) decrs /// Decides whether a top-level name [nm] syntactically /// appears in the term [t]. @@ -958,7 +955,7 @@ let rec new_args_for_smt_attrs (env:env) (l:list argv) (ty:typ) : Tac (list argv in begin match inspect_comp comp with - | C_Total ty2 _ _ -> + | C_Total ty2 -> let tl_argv, tl_terms = new_args_for_smt_attrs env tl ty2 in new_hd::tl_argv, (if needs_smt then arg::tl_terms else tl_terms) | _ -> fail "computation type not supported in definition of slprops" From 8f2148a1e37bc14d0d015d134ea0ecd0b84d5a97 Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Fri, 9 Dec 2022 10:07:10 +0000 Subject: [PATCH 04/18] src reflection changes, remove univs from Total GTotal --- src/reflection/FStar.Reflection.Basic.fst | 54 +++++-------------- src/reflection/FStar.Reflection.Data.fsti | 6 +-- .../FStar.Reflection.Embeddings.fst | 34 +++++------- .../FStar.Reflection.NBEEmbeddings.fst | 34 +++++------- 4 files changed, 45 insertions(+), 83 deletions(-) diff --git a/src/reflection/FStar.Reflection.Basic.fst b/src/reflection/FStar.Reflection.Basic.fst index ad6ee11acdd..e544d9d273b 100644 --- a/src/reflection/FStar.Reflection.Basic.fst +++ b/src/reflection/FStar.Reflection.Basic.fst @@ -298,8 +298,8 @@ let inspect_comp (c : comp) : comp_view = | _ -> failwith "Impossible!" in match c.n with - | Total (t, uopt) -> C_Total (t, BU.dflt U_unknown uopt, []) - | GTotal (t, uopt) -> C_GTotal (t, BU.dflt U_unknown uopt, []) + | Total t -> C_Total t + | GTotal t -> C_GTotal t | Comp ct -> begin let uopt = if List.length ct.comp_univs = 0 @@ -311,18 +311,13 @@ let inspect_comp (c : comp) : comp_view = C_Lemma (pre, post, pats) | _ -> failwith "inspect_comp: Lemma does not have enough arguments?" - else if Ident.lid_equals ct.effect_name PC.effect_Tot_lid then - let md = get_dec ct.flags in - C_Total (ct.result_typ, uopt, md) - else if Ident.lid_equals ct.effect_name PC.effect_GTot_lid then - let md = get_dec ct.flags in - C_GTotal (ct.result_typ, uopt, md) else let inspect_arg (a, q) = (a, inspect_aqual q) in C_Eff (ct.comp_univs, Ident.path_of_lid ct.effect_name, ct.result_typ, - List.map inspect_arg ct.effect_args) + List.map inspect_arg ct.effect_args, + get_dec ct.flags) end let pack_comp (cv : comp_view) : comp = @@ -335,26 +330,8 @@ let pack_comp (cv : comp_view) : comp = then None else Some u in match cv with - | C_Total (t, u, []) -> mk_Total' t (urefl_to_univ_opt u) - | C_Total (t, u, l) -> - let ct = { comp_univs=urefl_to_univs u - ; effect_name=PC.effect_Tot_lid - ; result_typ = t - ; effect_args = [] - ; flags = [DECREASES (Decreases_lex l)] } - in - S.mk_Comp ct - - | C_GTotal (t, u, []) -> mk_GTotal' t (urefl_to_univ_opt u) - | C_GTotal (t, u, l) -> - let ct = { comp_univs=urefl_to_univs u - ; effect_name=PC.effect_GTot_lid - ; result_typ = t - ; effect_args = [] - ; flags = [DECREASES (Decreases_lex l)] } - in - S.mk_Comp ct - + | C_Total t -> mk_Total t + | C_GTotal t -> mk_GTotal t | C_Lemma (pre, post, pats) -> let ct = { comp_univs = [] ; effect_name = PC.effect_Lemma_lid @@ -363,13 +340,13 @@ let pack_comp (cv : comp_view) : comp = ; flags = [] } in S.mk_Comp ct - | C_Eff (us, ef, res, args) -> + | C_Eff (us, ef, res, args, decrs) -> let pack_arg (a, q) = (a, pack_aqual q) in let ct = { comp_univs = us ; effect_name = Ident.lid_of_path ef Range.dummyRange ; result_typ = res ; effect_args = List.map pack_arg args - ; flags = [] } in + ; flags = [DECREASES (Decreases_lex decrs)] } in S.mk_Comp ct let pack_const (c:vconst) : sconst = @@ -875,22 +852,19 @@ and bv_eq (bv1 : bv) (bv2 : bv) : bool = and comp_eq (c1 : comp) (c2 : comp) : bool = match inspect_comp c1, inspect_comp c2 with - | C_Total (t1, u1, dec1), C_Total (t2, u2, dec2) -> - term_eq t1 t2 && univ_eq u1 u2 - && eqlist term_eq dec1 dec2 - - | C_GTotal (t1, u1, dec1), C_GTotal (t2, u2, dec2) -> - term_eq t1 t2 && univ_eq u1 u2 - && eqlist term_eq dec1 dec2 + | C_Total t1, C_Total t2 + | C_GTotal t1, C_GTotal t2 -> + term_eq t1 t2 | C_Lemma (pre1, post1, pats1), C_Lemma (pre2, post2, pats2) -> term_eq pre1 pre2 && term_eq post1 post2 && term_eq pats1 pats2 - | C_Eff (us1, name1, t1, args1), C_Eff (us2, name2, t2, args2) -> + | C_Eff (us1, name1, t1, args1, decrs1), C_Eff (us2, name2, t2, args2, decrs2) -> univs_eq us1 us2 && name1 = name2 && term_eq t1 t2 && - eqlist arg_eq args1 args2 + eqlist arg_eq args1 args2 && + eqlist term_eq decrs1 decrs2 | _ -> false diff --git a/src/reflection/FStar.Reflection.Data.fsti b/src/reflection/FStar.Reflection.Data.fsti index 469356f73d8..a50784c2bc0 100644 --- a/src/reflection/FStar.Reflection.Data.fsti +++ b/src/reflection/FStar.Reflection.Data.fsti @@ -126,10 +126,10 @@ type bv_view = { type binder_view = bv * (aqualv * list term) type comp_view = - | C_Total of typ * universe * list term //decreases clause - | C_GTotal of typ * universe * list term //idem + | C_Total of typ + | C_GTotal of typ | C_Lemma of term * term * term - | C_Eff of universes * name * term * list argv + | C_Eff of universes * name * term * list argv * list term // list term is the decreases clause type ctor = name * typ diff --git a/src/reflection/FStar.Reflection.Embeddings.fst b/src/reflection/FStar.Reflection.Embeddings.fst index 57f1eb82eae..188425267c1 100644 --- a/src/reflection/FStar.Reflection.Embeddings.fst +++ b/src/reflection/FStar.Reflection.Embeddings.fst @@ -631,28 +631,25 @@ let e_bv_view = let e_comp_view = let embed_comp_view (rng:Range.range) (cv : comp_view) : term = match cv with - | C_Total (t, u, md) -> - S.mk_Tm_app ref_C_Total.t [S.as_arg (embed e_term rng t); - S.as_arg (embed e_universe rng u); - S.as_arg (embed (e_list e_term) rng md)] + | C_Total t -> + S.mk_Tm_app ref_C_Total.t [S.as_arg (embed e_term rng t)] rng - | C_GTotal (t, u, md) -> - S.mk_Tm_app ref_C_GTotal.t [S.as_arg (embed e_term rng t); - S.as_arg (embed e_universe rng u); - S.as_arg (embed (e_list e_term) rng md)] + | C_GTotal t -> + S.mk_Tm_app ref_C_GTotal.t [S.as_arg (embed e_term rng t)] rng | C_Lemma (pre, post, pats) -> S.mk_Tm_app ref_C_Lemma.t [S.as_arg (embed e_term rng pre); S.as_arg (embed e_term rng post); S.as_arg (embed e_term rng pats)] rng - | C_Eff (us, eff, res, args) -> + | C_Eff (us, eff, res, args, decrs) -> S.mk_Tm_app ref_C_Eff.t [ S.as_arg (embed (e_list e_universe) rng us) ; S.as_arg (embed e_string_list rng eff) ; S.as_arg (embed e_term rng res) - ; S.as_arg (embed (e_list e_argv) rng args)] rng + ; S.as_arg (embed (e_list e_argv) rng args) + ; S.as_arg (embed (e_list e_term) rng decrs)] rng in @@ -660,19 +657,15 @@ let e_comp_view = let t = U.unascribe t in let hd, args = U.head_and_args t in match (U.un_uinst hd).n, args with - | Tm_fvar fv, [(t, _); (u, _); (md, _)] + | Tm_fvar fv, [(t, _)] when S.fv_eq_lid fv ref_C_Total.lid -> BU.bind_opt (unembed' w e_term t) (fun t -> - BU.bind_opt (unembed' w e_universe u) (fun u -> - BU.bind_opt (unembed' w (e_list e_term) md) (fun md -> - Some <| C_Total (t, u, md)))) + Some <| C_Total t) - | Tm_fvar fv, [(t, _); (u, _); (md, _)] + | Tm_fvar fv, [(t, _)] when S.fv_eq_lid fv ref_C_GTotal.lid -> BU.bind_opt (unembed' w e_term t) (fun t -> - BU.bind_opt (unembed' w e_universe u) (fun u -> - BU.bind_opt (unembed' w (e_list e_term) md) (fun md -> - Some <| C_GTotal (t, u, md)))) + Some <| C_GTotal t) | Tm_fvar fv, [(pre, _); (post, _); (pats, _)] when S.fv_eq_lid fv ref_C_Lemma.lid -> BU.bind_opt (unembed' w e_term pre) (fun pre -> @@ -680,13 +673,14 @@ let e_comp_view = BU.bind_opt (unembed' w e_term pats) (fun pats -> Some <| C_Lemma (pre, post, pats)))) - | Tm_fvar fv, [(us, _); (eff, _); (res, _); (args, _)] + | Tm_fvar fv, [(us, _); (eff, _); (res, _); (args, _); (decrs, _)] when S.fv_eq_lid fv ref_C_Eff.lid -> BU.bind_opt (unembed' w (e_list e_universe) us) (fun us -> BU.bind_opt (unembed' w e_string_list eff) (fun eff -> BU.bind_opt (unembed' w e_term res) (fun res-> BU.bind_opt (unembed' w (e_list e_argv) args) (fun args -> - Some <| C_Eff ([], eff, res, args))))) + BU.bind_opt (unembed' w (e_list e_term) decrs) (fun decrs -> + Some <| C_Eff (us, eff, res, args, decrs)))))) | _ -> if w then diff --git a/src/reflection/FStar.Reflection.NBEEmbeddings.fst b/src/reflection/FStar.Reflection.NBEEmbeddings.fst index 5d941a22f5e..33aeb601106 100644 --- a/src/reflection/FStar.Reflection.NBEEmbeddings.fst +++ b/src/reflection/FStar.Reflection.NBEEmbeddings.fst @@ -605,43 +605,36 @@ let e_bv_view = let e_comp_view = let embed_comp_view cb (cv : comp_view) : t = match cv with - | C_Total (t, u, md) -> + | C_Total t -> mkConstruct ref_C_Total.fv [] [ - as_arg (embed e_term cb t); - as_arg (embed e_universe cb u); - as_arg (embed (e_list e_term) cb md)] + as_arg (embed e_term cb t)] - | C_GTotal (t, u, md) -> + | C_GTotal t -> mkConstruct ref_C_GTotal.fv [] [ - as_arg (embed e_term cb t); - as_arg (embed e_universe cb u); - as_arg (embed (e_list e_term) cb md)] + as_arg (embed e_term cb t)] | C_Lemma (pre, post, pats) -> mkConstruct ref_C_Lemma.fv [] [as_arg (embed e_term cb pre); as_arg (embed e_term cb post); as_arg (embed e_term cb pats)] - | C_Eff (us, eff, res, args) -> + | C_Eff (us, eff, res, args, decrs) -> mkConstruct ref_C_Eff.fv [] [ as_arg (embed (e_list e_universe) cb us) ; as_arg (embed e_string_list cb eff) ; as_arg (embed e_term cb res) - ; as_arg (embed (e_list e_argv) cb args)] + ; as_arg (embed (e_list e_argv) cb args) + ; as_arg (embed (e_list e_term) cb decrs)] in let unembed_comp_view cb (t : t) : option comp_view = match t.nbe_t with - | Construct (fv, _, [(md, _); (u, _); (t, _)]) + | Construct (fv, _, [(t, _)]) when S.fv_eq_lid fv ref_C_Total.lid -> BU.bind_opt (unembed e_term cb t) (fun t -> - BU.bind_opt (unembed e_universe cb u) (fun u -> - BU.bind_opt (unembed (e_list e_term) cb md) (fun md -> - Some <| C_Total (t, u, md)))) + Some <| C_Total t) - | Construct (fv, _, [(md, _); (u, _); (t, _)]) + | Construct (fv, _, [(t, _)]) when S.fv_eq_lid fv ref_C_GTotal.lid -> BU.bind_opt (unembed e_term cb t) (fun t -> - BU.bind_opt (unembed e_universe cb u) (fun u -> - BU.bind_opt (unembed (e_list e_term) cb md) (fun md -> - Some <| C_GTotal (t, u, md)))) + Some <| C_GTotal t) | Construct (fv, _, [(post, _); (pre, _); (pats, _)]) when S.fv_eq_lid fv ref_C_Lemma.lid -> BU.bind_opt (unembed e_term cb pre) (fun pre -> @@ -649,13 +642,14 @@ let e_comp_view = BU.bind_opt (unembed e_term cb pats) (fun pats -> Some <| C_Lemma (pre, post, pats)))) - | Construct (fv, _, [(args, _); (res, _); (eff, _); (us, _)]) + | Construct (fv, _, [(decrs, _); (args, _); (res, _); (eff, _); (us, _)]) when S.fv_eq_lid fv ref_C_Eff.lid -> BU.bind_opt (unembed (e_list e_universe) cb us) (fun us -> BU.bind_opt (unembed e_string_list cb eff) (fun eff -> BU.bind_opt (unembed e_term cb res) (fun res-> BU.bind_opt (unembed (e_list e_argv) cb args) (fun args -> - Some <| C_Eff (us, eff, res, args))))) + BU.bind_opt (unembed (e_list e_term) cb decrs) (fun decrs -> + Some <| C_Eff (us, eff, res, args, decrs)))))) | _ -> Err.log_issue Range.dummyRange (Err.Warning_NotEmbedded, (BU.format1 "Not an embedded comp_view: %s" (t_to_string t))); From 2064baba36e067c2536f327dd834309a4a2e0b37 Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Fri, 9 Dec 2022 10:07:25 +0000 Subject: [PATCH 05/18] tests wip --- tests/bug-reports/Bug1902.fst | 15 ++++++----- tests/micro-benchmarks/BinderAttributes.fst | 2 +- .../MultipleAttributesBinder.fst | 2 +- .../RecordFieldAttributes.fst | 2 +- tests/tactics/InspectEffComp.fst | 4 +-- tests/tactics/ReflectionMisc.fst | 26 +++++++++---------- tests/tactics/Splice.fst | 7 ++++- 7 files changed, 33 insertions(+), 25 deletions(-) diff --git a/tests/bug-reports/Bug1902.fst b/tests/bug-reports/Bug1902.fst index 52e19add358..4106d0f0f08 100644 --- a/tests/bug-reports/Bug1902.fst +++ b/tests/bug-reports/Bug1902.fst @@ -42,11 +42,14 @@ let decr_at_every_level = false let rec mk_tot_arr_decr (bs: list binder) (cod : term) decr : Tac term = match bs with | [] -> cod - | (b::bs) -> pack (Tv_Arrow b (pack_comp (C_Total (mk_tot_arr_decr bs cod decr) (pack_universe Uv_Zero) ( - if decr_at_every_level || FStar.List.Tot.length bs = 0 - then [decr] - else [] - )))) + | (b::bs) -> pack (Tv_Arrow b + (pack_comp (C_Eff [pack_universe Uv_Zero] + ["Prims"; "Tot"] + (mk_tot_arr_decr bs cod decr) + [] + (if decr_at_every_level || FStar.List.Tot.length bs = 0 + then [decr] + else [])))) let craft_f' use_f_type: Tac decls = let name = pack_fv (cur_module () @ ["f'" ^ ( @@ -85,4 +88,4 @@ let craft_f' use_f_type: Tac decls = // however, crafting f' with a crafted type // result in F* failing because of termination issues -%splice[] (craft_f' false) +// %splice[] (craft_f' false) diff --git a/tests/micro-benchmarks/BinderAttributes.fst b/tests/micro-benchmarks/BinderAttributes.fst index b339a80beef..5b4a9aca280 100644 --- a/tests/micro-benchmarks/BinderAttributes.fst +++ b/tests/micro-benchmarks/BinderAttributes.fst @@ -70,7 +70,7 @@ let rec binders_from_arrow (ty : T.term) : T.Tac binders = | T.Tv_Arrow b comp -> begin let ba = binder_from_term b in match T.inspect_comp comp with - | T.C_Total ty2 _ _ -> ba :: binders_from_arrow ty2 + | T.C_Total ty2 -> ba :: binders_from_arrow ty2 | _ -> T.fail "Unsupported computation type" end | T.Tv_FVar fv -> [] //last part diff --git a/tests/micro-benchmarks/MultipleAttributesBinder.fst b/tests/micro-benchmarks/MultipleAttributesBinder.fst index 8ac857c403c..dc643d2fa7e 100644 --- a/tests/micro-benchmarks/MultipleAttributesBinder.fst +++ b/tests/micro-benchmarks/MultipleAttributesBinder.fst @@ -34,7 +34,7 @@ let rec binders_from_arrow (ty : T.term) : T.Tac binders = | T.Tv_Arrow b comp -> begin let ba = binder_from_term b in match T.inspect_comp comp with - | T.C_Total ty2 _ _ -> ba :: binders_from_arrow ty2 + | T.C_Total ty2 -> ba :: binders_from_arrow ty2 | _ -> T.fail "Unsupported computation type" end | T.Tv_FVar fv -> [] //last part diff --git a/tests/micro-benchmarks/RecordFieldAttributes.fst b/tests/micro-benchmarks/RecordFieldAttributes.fst index 1129c060443..65fa94a4ffe 100644 --- a/tests/micro-benchmarks/RecordFieldAttributes.fst +++ b/tests/micro-benchmarks/RecordFieldAttributes.fst @@ -38,7 +38,7 @@ let rec unpack_fields (qname : list string) (ty : T.term) : T.Tac (list (string | T.Tv_Arrow binder comp -> begin let f = unpack_field binder in match T.inspect_comp comp with - | T.C_Total ty2 _ _ -> f :: unpack_fields qname ty2 + | T.C_Total ty2 -> f :: unpack_fields qname ty2 | _ -> T.fail "Unsupported computation type" end | T.Tv_FVar fv -> begin diff --git a/tests/tactics/InspectEffComp.fst b/tests/tactics/InspectEffComp.fst index a6ef982672e..c6d911c63fb 100644 --- a/tests/tactics/InspectEffComp.fst +++ b/tests/tactics/InspectEffComp.fst @@ -11,9 +11,9 @@ let test () : Type0 = | Tv_Arrow bv c -> let c' = begin match inspect_comp c with - | C_Eff us eff res args -> + | C_Eff us eff res args decrs -> let args' = [(`(as_pure_wp (fun p -> p 17)), Q_Explicit)] in - pack_comp (C_Eff us eff res args') + pack_comp (C_Eff us eff res args' decrs) | _ -> fail "no" end in diff --git a/tests/tactics/ReflectionMisc.fst b/tests/tactics/ReflectionMisc.fst index 31d38e8c2d5..b43c38ea60a 100644 --- a/tests/tactics/ReflectionMisc.fst +++ b/tests/tactics/ReflectionMisc.fst @@ -66,18 +66,18 @@ let _ = assert True // Taking the universe from Tot // And applying Succ to it // -let get_u () : Tac term = - let t = `(int -> Tot u#0 int) in - match inspect t with - | Tv_Arrow _ c -> - (match inspect_comp c with - | C_Total _ u _ -> pack_ln (Tv_Type (pack_universe (Uv_Succ u))) - | _ -> fail "2") - | _ -> fail "3" +// let get_u () : Tac term = +// let t = `(int -> Tot u#0 int) in +// match inspect t with +// | Tv_Arrow _ c -> +// (match inspect_comp c with +// | C_Total _ u _ -> pack_ln (Tv_Type (pack_universe (Uv_Succ u))) +// | _ -> fail "2") +// | _ -> fail "3" -let test0 (#[exact (get_u ())]t:Type u#2) () : Type u#2 = t -let test1 () = test0 () +// let test0 (#[exact (get_u ())]t:Type u#2) () : Type u#2 = t +// let test1 () = test0 () -let test2 (#[exact (get_u ())]t:Type u#1) () : Type u#1 = t -[@@ expect_failure [228]] -let test3 () = test2 () +// let test2 (#[exact (get_u ())]t:Type u#1) () : Type u#1 = t +// [@@ expect_failure [228]] +// let test3 () = test2 () diff --git a/tests/tactics/Splice.fst b/tests/tactics/Splice.fst index aaf8d6aad1c..495224c7409 100644 --- a/tests/tactics/Splice.fst +++ b/tests/tactics/Splice.fst @@ -53,7 +53,12 @@ let _ = assert (x == 42) Sg_Let true [pack_lb ({ lb_fv = recursive_fv ; lb_us = [] - ; lb_typ = mk_arr [n] (pack_comp (C_Total (`Type0) (pack_universe (Uv_Succ (pack_universe Uv_Zero))) [`(10 - (`#(binder_to_term n)))])) + ; lb_typ = mk_arr [n] + (pack_comp (C_Eff [pack_universe (Uv_Succ (pack_universe Uv_Zero))] + ["Prims"; "Tot"] + (`Type0) + [] + [`(10 - (`#(binder_to_term n)))])) ; lb_def = `(fun n -> if n>=10 then int else int * (`#recursive) (n + 1)) })] ) From 940bf05484345210809daaed4147b5c7ee25932c Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Fri, 9 Dec 2022 10:07:32 +0000 Subject: [PATCH 06/18] snap --- examples/tactics/Printers.fst | 2 +- src/ocaml-output/FStar_Reflection_Basic.ml | 113 +++----- src/ocaml-output/FStar_Reflection_Data.ml | 26 +- .../FStar_Reflection_Embeddings.ml | 132 ++++------ .../FStar_Reflection_NBEEmbeddings.ml | 130 ++++----- src/ocaml-output/FStar_Syntax_Free.ml | 15 +- src/ocaml-output/FStar_Syntax_Hash.ml | 30 +-- src/ocaml-output/FStar_Syntax_InstFV.ml | 8 +- src/ocaml-output/FStar_Syntax_Print.ml | 28 +- src/ocaml-output/FStar_Syntax_Resugar.ml | 54 +--- src/ocaml-output/FStar_Syntax_Subst.ml | 28 +- src/ocaml-output/FStar_Syntax_Syntax.ml | 20 +- src/ocaml-output/FStar_Syntax_Util.ml | 171 ++++-------- src/ocaml-output/FStar_Tactics_Basic.ml | 66 +++-- src/ocaml-output/FStar_Tests_Util.ml | 4 +- src/ocaml-output/FStar_TypeChecker_Core.ml | 176 +++++++------ src/ocaml-output/FStar_TypeChecker_DMFF.ml | 248 +++++++++--------- src/ocaml-output/FStar_TypeChecker_Env.ml | 59 ++++- src/ocaml-output/FStar_TypeChecker_Err.ml | 4 +- src/ocaml-output/FStar_TypeChecker_NBE.ml | 20 +- .../FStar_TypeChecker_Normalize.ml | 44 +--- src/ocaml-output/FStar_TypeChecker_Rel.ml | 137 ++++------ .../FStar_TypeChecker_TcEffect.ml | 71 ++--- src/ocaml-output/FStar_TypeChecker_TcTerm.ml | 59 ++--- src/ocaml-output/FStar_TypeChecker_Util.ml | 78 +++--- src/tactics/FStar.Tactics.Basic.fst | 12 +- src/tests/FStar.Tests.Util.fst | 2 +- 27 files changed, 707 insertions(+), 1030 deletions(-) diff --git a/examples/tactics/Printers.fst b/examples/tactics/Printers.fst index 9aef524bca9..889236052ff 100644 --- a/examples/tactics/Printers.fst +++ b/examples/tactics/Printers.fst @@ -48,7 +48,7 @@ let mk_print_bv (self : name) (f : term) (bv : bv) : Tac term = let mk_printer_type (t : term) : Tac term = let b = fresh_binder_named "arg" t in let str = pack (Tv_FVar (pack_fv string_lid)) in - let c = pack_comp (C_Total str u_unk []) in + let c = pack_comp (C_Total str) in pack (Tv_Arrow b c) diff --git a/src/ocaml-output/FStar_Reflection_Basic.ml b/src/ocaml-output/FStar_Reflection_Basic.ml index be7b4941982..e24ef550d65 100644 --- a/src/ocaml-output/FStar_Reflection_Basic.ml +++ b/src/ocaml-output/FStar_Reflection_Basic.ml @@ -321,14 +321,8 @@ let (inspect_comp : []) | uu___1 -> failwith "Impossible!" in match c.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.Total (t, uopt) -> - FStar_Reflection_Data.C_Total - (t, (FStar_Compiler_Util.dflt FStar_Syntax_Syntax.U_unknown uopt), - []) - | FStar_Syntax_Syntax.GTotal (t, uopt) -> - FStar_Reflection_Data.C_GTotal - (t, (FStar_Compiler_Util.dflt FStar_Syntax_Syntax.U_unknown uopt), - []) + | FStar_Syntax_Syntax.Total t -> FStar_Reflection_Data.C_Total t + | FStar_Syntax_Syntax.GTotal t -> FStar_Reflection_Data.C_GTotal t | FStar_Syntax_Syntax.Comp ct -> let uopt = if @@ -349,37 +343,19 @@ let (inspect_comp : | uu___1 -> failwith "inspect_comp: Lemma does not have enough arguments?") else - (let uu___2 = - FStar_Ident.lid_equals ct.FStar_Syntax_Syntax.effect_name - FStar_Parser_Const.effect_Tot_lid in - if uu___2 - then - let md = get_dec ct.FStar_Syntax_Syntax.flags in - FStar_Reflection_Data.C_Total - ((ct.FStar_Syntax_Syntax.result_typ), uopt, md) - else - (let uu___4 = - FStar_Ident.lid_equals ct.FStar_Syntax_Syntax.effect_name - FStar_Parser_Const.effect_GTot_lid in - if uu___4 - then - let md = get_dec ct.FStar_Syntax_Syntax.flags in - FStar_Reflection_Data.C_GTotal - ((ct.FStar_Syntax_Syntax.result_typ), uopt, md) - else - (let inspect_arg uu___6 = - match uu___6 with - | (a, q) -> let uu___7 = inspect_aqual q in (a, uu___7) in - let uu___6 = - let uu___7 = - FStar_Ident.path_of_lid - ct.FStar_Syntax_Syntax.effect_name in - let uu___8 = - FStar_Compiler_List.map inspect_arg - ct.FStar_Syntax_Syntax.effect_args in - ((ct.FStar_Syntax_Syntax.comp_univs), uu___7, - (ct.FStar_Syntax_Syntax.result_typ), uu___8) in - FStar_Reflection_Data.C_Eff uu___6))) + (let inspect_arg uu___2 = + match uu___2 with + | (a, q) -> let uu___3 = inspect_aqual q in (a, uu___3) in + let uu___2 = + let uu___3 = + FStar_Ident.path_of_lid ct.FStar_Syntax_Syntax.effect_name in + let uu___4 = + FStar_Compiler_List.map inspect_arg + ct.FStar_Syntax_Syntax.effect_args in + let uu___5 = get_dec ct.FStar_Syntax_Syntax.flags in + ((ct.FStar_Syntax_Syntax.comp_univs), uu___3, + (ct.FStar_Syntax_Syntax.result_typ), uu___4, uu___5) in + FStar_Reflection_Data.C_Eff uu___2) let (pack_comp : FStar_Reflection_Data.comp_view -> FStar_Syntax_Syntax.comp) = fun cv -> @@ -390,36 +366,8 @@ let (pack_comp : FStar_Reflection_Data.comp_view -> FStar_Syntax_Syntax.comp) then FStar_Pervasives_Native.None else FStar_Pervasives_Native.Some u in match cv with - | FStar_Reflection_Data.C_Total (t, u, []) -> - FStar_Syntax_Syntax.mk_Total' t (urefl_to_univ_opt u) - | FStar_Reflection_Data.C_Total (t, u, l) -> - let ct = - { - FStar_Syntax_Syntax.comp_univs = (urefl_to_univs u); - FStar_Syntax_Syntax.effect_name = - FStar_Parser_Const.effect_Tot_lid; - FStar_Syntax_Syntax.result_typ = t; - FStar_Syntax_Syntax.effect_args = []; - FStar_Syntax_Syntax.flags = - [FStar_Syntax_Syntax.DECREASES - (FStar_Syntax_Syntax.Decreases_lex l)] - } in - FStar_Syntax_Syntax.mk_Comp ct - | FStar_Reflection_Data.C_GTotal (t, u, []) -> - FStar_Syntax_Syntax.mk_GTotal' t (urefl_to_univ_opt u) - | FStar_Reflection_Data.C_GTotal (t, u, l) -> - let ct = - { - FStar_Syntax_Syntax.comp_univs = (urefl_to_univs u); - FStar_Syntax_Syntax.effect_name = - FStar_Parser_Const.effect_GTot_lid; - FStar_Syntax_Syntax.result_typ = t; - FStar_Syntax_Syntax.effect_args = []; - FStar_Syntax_Syntax.flags = - [FStar_Syntax_Syntax.DECREASES - (FStar_Syntax_Syntax.Decreases_lex l)] - } in - FStar_Syntax_Syntax.mk_Comp ct + | FStar_Reflection_Data.C_Total t -> FStar_Syntax_Syntax.mk_Total t + | FStar_Reflection_Data.C_GTotal t -> FStar_Syntax_Syntax.mk_GTotal t | FStar_Reflection_Data.C_Lemma (pre, post, pats) -> let ct = let uu___ = @@ -439,7 +387,7 @@ let (pack_comp : FStar_Reflection_Data.comp_view -> FStar_Syntax_Syntax.comp) FStar_Syntax_Syntax.flags = [] } in FStar_Syntax_Syntax.mk_Comp ct - | FStar_Reflection_Data.C_Eff (us, ef, res, args) -> + | FStar_Reflection_Data.C_Eff (us, ef, res, args, decrs) -> let pack_arg uu___ = match uu___ with | (a, q) -> let uu___1 = pack_aqual q in (a, uu___1) in @@ -452,7 +400,9 @@ let (pack_comp : FStar_Reflection_Data.comp_view -> FStar_Syntax_Syntax.comp) FStar_Syntax_Syntax.effect_name = uu___; FStar_Syntax_Syntax.result_typ = res; FStar_Syntax_Syntax.effect_args = uu___1; - FStar_Syntax_Syntax.flags = [] + FStar_Syntax_Syntax.flags = + [FStar_Syntax_Syntax.DECREASES + (FStar_Syntax_Syntax.Decreases_lex decrs)] } in FStar_Syntax_Syntax.mk_Comp ct let (pack_const : FStar_Reflection_Data.vconst -> FStar_Syntax_Syntax.sconst) @@ -1239,22 +1189,19 @@ and (comp_eq : let uu___1 = inspect_comp c1 in let uu___2 = inspect_comp c2 in (uu___1, uu___2) in match uu___ with - | (FStar_Reflection_Data.C_Total (t1, u1, dec1), - FStar_Reflection_Data.C_Total (t2, u2, dec2)) -> - ((term_eq t1 t2) && (univ_eq u1 u2)) && - ((eqlist ()) term_eq dec1 dec2) - | (FStar_Reflection_Data.C_GTotal (t1, u1, dec1), - FStar_Reflection_Data.C_GTotal (t2, u2, dec2)) -> - ((term_eq t1 t2) && (univ_eq u1 u2)) && - ((eqlist ()) term_eq dec1 dec2) + | (FStar_Reflection_Data.C_Total t1, FStar_Reflection_Data.C_Total t2) + -> term_eq t1 t2 + | (FStar_Reflection_Data.C_GTotal t1, FStar_Reflection_Data.C_GTotal + t2) -> term_eq t1 t2 | (FStar_Reflection_Data.C_Lemma (pre1, post1, pats1), FStar_Reflection_Data.C_Lemma (pre2, post2, pats2)) -> ((term_eq pre1 pre2) && (term_eq post1 post2)) && (term_eq pats1 pats2) - | (FStar_Reflection_Data.C_Eff (us1, name1, t1, args1), - FStar_Reflection_Data.C_Eff (us2, name2, t2, args2)) -> - (((univs_eq us1 us2) && (name1 = name2)) && (term_eq t1 t2)) && - ((eqlist ()) arg_eq args1 args2) + | (FStar_Reflection_Data.C_Eff (us1, name1, t1, args1, decrs1), + FStar_Reflection_Data.C_Eff (us2, name2, t2, args2, decrs2)) -> + ((((univs_eq us1 us2) && (name1 = name2)) && (term_eq t1 t2)) && + ((eqlist ()) arg_eq args1 args2)) + && ((eqlist ()) term_eq decrs1 decrs2) | uu___1 -> false and (match_ret_asc_eq : FStar_Syntax_Syntax.match_returns_ascription -> diff --git a/src/ocaml-output/FStar_Reflection_Data.ml b/src/ocaml-output/FStar_Reflection_Data.ml index e88f2183f7d..25957025e66 100644 --- a/src/ocaml-output/FStar_Reflection_Data.ml +++ b/src/ocaml-output/FStar_Reflection_Data.ml @@ -353,28 +353,21 @@ let (__proj__Mkbv_view__item__bv_sort : bv_view -> typ) = type binder_view = (FStar_Syntax_Syntax.bv * (aqualv * FStar_Syntax_Syntax.term Prims.list)) type comp_view = - | C_Total of (typ * FStar_Syntax_Syntax.universe * FStar_Syntax_Syntax.term - Prims.list) - | C_GTotal of (typ * FStar_Syntax_Syntax.universe * - FStar_Syntax_Syntax.term Prims.list) + | C_Total of typ + | C_GTotal of typ | C_Lemma of (FStar_Syntax_Syntax.term * FStar_Syntax_Syntax.term * FStar_Syntax_Syntax.term) - | C_Eff of (universes * name * FStar_Syntax_Syntax.term * argv Prims.list) + | C_Eff of (universes * name * FStar_Syntax_Syntax.term * argv Prims.list * + FStar_Syntax_Syntax.term Prims.list) let (uu___is_C_Total : comp_view -> Prims.bool) = fun projectee -> match projectee with | C_Total _0 -> true | uu___ -> false -let (__proj__C_Total__item___0 : - comp_view -> - (typ * FStar_Syntax_Syntax.universe * FStar_Syntax_Syntax.term - Prims.list)) - = fun projectee -> match projectee with | C_Total _0 -> _0 +let (__proj__C_Total__item___0 : comp_view -> typ) = + fun projectee -> match projectee with | C_Total _0 -> _0 let (uu___is_C_GTotal : comp_view -> Prims.bool) = fun projectee -> match projectee with | C_GTotal _0 -> true | uu___ -> false -let (__proj__C_GTotal__item___0 : - comp_view -> - (typ * FStar_Syntax_Syntax.universe * FStar_Syntax_Syntax.term - Prims.list)) - = fun projectee -> match projectee with | C_GTotal _0 -> _0 +let (__proj__C_GTotal__item___0 : comp_view -> typ) = + fun projectee -> match projectee with | C_GTotal _0 -> _0 let (uu___is_C_Lemma : comp_view -> Prims.bool) = fun projectee -> match projectee with | C_Lemma _0 -> true | uu___ -> false let (__proj__C_Lemma__item___0 : @@ -386,7 +379,8 @@ let (uu___is_C_Eff : comp_view -> Prims.bool) = fun projectee -> match projectee with | C_Eff _0 -> true | uu___ -> false let (__proj__C_Eff__item___0 : comp_view -> - (universes * name * FStar_Syntax_Syntax.term * argv Prims.list)) + (universes * name * FStar_Syntax_Syntax.term * argv Prims.list * + FStar_Syntax_Syntax.term Prims.list)) = fun projectee -> match projectee with | C_Eff _0 -> _0 type ctor = (name * typ) type lb_view = diff --git a/src/ocaml-output/FStar_Reflection_Embeddings.ml b/src/ocaml-output/FStar_Reflection_Embeddings.ml index cd2957e6ce3..03695ff56da 100644 --- a/src/ocaml-output/FStar_Reflection_Embeddings.ml +++ b/src/ocaml-output/FStar_Reflection_Embeddings.ml @@ -1509,45 +1509,21 @@ let (e_comp_view : FStar_Reflection_Data.comp_view FStar_Syntax_Embeddings.embedding) = let embed_comp_view rng cv = match cv with - | FStar_Reflection_Data.C_Total (t, u, md) -> + | FStar_Reflection_Data.C_Total t -> let uu___ = let uu___1 = let uu___2 = embed e_term rng t in FStar_Syntax_Syntax.as_arg uu___2 in - let uu___2 = - let uu___3 = - let uu___4 = embed e_universe rng u in - FStar_Syntax_Syntax.as_arg uu___4 in - let uu___4 = - let uu___5 = - let uu___6 = - let uu___7 = FStar_Syntax_Embeddings.e_list e_term in - embed uu___7 rng md in - FStar_Syntax_Syntax.as_arg uu___6 in - [uu___5] in - uu___3 :: uu___4 in - uu___1 :: uu___2 in + [uu___1] in FStar_Syntax_Syntax.mk_Tm_app FStar_Reflection_Constants.ref_C_Total.FStar_Reflection_Constants.t uu___ rng - | FStar_Reflection_Data.C_GTotal (t, u, md) -> + | FStar_Reflection_Data.C_GTotal t -> let uu___ = let uu___1 = let uu___2 = embed e_term rng t in FStar_Syntax_Syntax.as_arg uu___2 in - let uu___2 = - let uu___3 = - let uu___4 = embed e_universe rng u in - FStar_Syntax_Syntax.as_arg uu___4 in - let uu___4 = - let uu___5 = - let uu___6 = - let uu___7 = FStar_Syntax_Embeddings.e_list e_term in - embed uu___7 rng md in - FStar_Syntax_Syntax.as_arg uu___6 in - [uu___5] in - uu___3 :: uu___4 in - uu___1 :: uu___2 in + [uu___1] in FStar_Syntax_Syntax.mk_Tm_app FStar_Reflection_Constants.ref_C_GTotal.FStar_Reflection_Constants.t uu___ rng @@ -1570,7 +1546,7 @@ let (e_comp_view : FStar_Syntax_Syntax.mk_Tm_app FStar_Reflection_Constants.ref_C_Lemma.FStar_Reflection_Constants.t uu___ rng - | FStar_Reflection_Data.C_Eff (us, eff, res, args) -> + | FStar_Reflection_Data.C_Eff (us, eff, res, args, decrs) -> let uu___ = let uu___1 = let uu___2 = @@ -1592,7 +1568,14 @@ let (e_comp_view : let uu___9 = FStar_Syntax_Embeddings.e_list e_argv in embed uu___9 rng args in FStar_Syntax_Syntax.as_arg uu___8 in - [uu___7] in + let uu___8 = + let uu___9 = + let uu___10 = + let uu___11 = FStar_Syntax_Embeddings.e_list e_term in + embed uu___11 rng decrs in + FStar_Syntax_Syntax.as_arg uu___10 in + [uu___9] in + uu___7 :: uu___8 in uu___5 :: uu___6 in uu___3 :: uu___4 in uu___1 :: uu___2 in @@ -1610,46 +1593,26 @@ let (e_comp_view : uu___3.FStar_Syntax_Syntax.n in (uu___2, args) in (match uu___1 with - | (FStar_Syntax_Syntax.Tm_fvar fv, - (t2, uu___2)::(u, uu___3)::(md, uu___4)::[]) when + | (FStar_Syntax_Syntax.Tm_fvar fv, (t2, uu___2)::[]) when FStar_Syntax_Syntax.fv_eq_lid fv FStar_Reflection_Constants.ref_C_Total.FStar_Reflection_Constants.lid -> - let uu___5 = unembed' w e_term t2 in - FStar_Compiler_Util.bind_opt uu___5 + let uu___3 = unembed' w e_term t2 in + FStar_Compiler_Util.bind_opt uu___3 (fun t3 -> - let uu___6 = unembed' w e_universe u in - FStar_Compiler_Util.bind_opt uu___6 - (fun u1 -> - let uu___7 = - let uu___8 = FStar_Syntax_Embeddings.e_list e_term in - unembed' w uu___8 md in - FStar_Compiler_Util.bind_opt uu___7 - (fun md1 -> - FStar_Compiler_Effect.op_Less_Bar - (fun uu___8 -> - FStar_Pervasives_Native.Some uu___8) - (FStar_Reflection_Data.C_Total (t3, u1, md1))))) - | (FStar_Syntax_Syntax.Tm_fvar fv, - (t2, uu___2)::(u, uu___3)::(md, uu___4)::[]) when + FStar_Compiler_Effect.op_Less_Bar + (fun uu___4 -> FStar_Pervasives_Native.Some uu___4) + (FStar_Reflection_Data.C_Total t3)) + | (FStar_Syntax_Syntax.Tm_fvar fv, (t2, uu___2)::[]) when FStar_Syntax_Syntax.fv_eq_lid fv FStar_Reflection_Constants.ref_C_GTotal.FStar_Reflection_Constants.lid -> - let uu___5 = unembed' w e_term t2 in - FStar_Compiler_Util.bind_opt uu___5 + let uu___3 = unembed' w e_term t2 in + FStar_Compiler_Util.bind_opt uu___3 (fun t3 -> - let uu___6 = unembed' w e_universe u in - FStar_Compiler_Util.bind_opt uu___6 - (fun u1 -> - let uu___7 = - let uu___8 = FStar_Syntax_Embeddings.e_list e_term in - unembed' w uu___8 md in - FStar_Compiler_Util.bind_opt uu___7 - (fun md1 -> - FStar_Compiler_Effect.op_Less_Bar - (fun uu___8 -> - FStar_Pervasives_Native.Some uu___8) - (FStar_Reflection_Data.C_GTotal (t3, u1, md1))))) + FStar_Compiler_Effect.op_Less_Bar + (fun uu___4 -> FStar_Pervasives_Native.Some uu___4) + (FStar_Reflection_Data.C_GTotal t3)) | (FStar_Syntax_Syntax.Tm_fvar fv, (pre, uu___2)::(post, uu___3)::(pats, uu___4)::[]) when FStar_Syntax_Syntax.fv_eq_lid fv @@ -1670,34 +1633,41 @@ let (e_comp_view : (FStar_Reflection_Data.C_Lemma (pre1, post1, pats1))))) | (FStar_Syntax_Syntax.Tm_fvar fv, - (us, uu___2)::(eff, uu___3)::(res, uu___4)::(args1, uu___5)::[]) - when + (us, uu___2)::(eff, uu___3)::(res, uu___4)::(args1, uu___5):: + (decrs, uu___6)::[]) when FStar_Syntax_Syntax.fv_eq_lid fv FStar_Reflection_Constants.ref_C_Eff.FStar_Reflection_Constants.lid -> - let uu___6 = - let uu___7 = FStar_Syntax_Embeddings.e_list e_universe in - unembed' w uu___7 us in - FStar_Compiler_Util.bind_opt uu___6 + let uu___7 = + let uu___8 = FStar_Syntax_Embeddings.e_list e_universe in + unembed' w uu___8 us in + FStar_Compiler_Util.bind_opt uu___7 (fun us1 -> - let uu___7 = + let uu___8 = unembed' w FStar_Syntax_Embeddings.e_string_list eff in - FStar_Compiler_Util.bind_opt uu___7 + FStar_Compiler_Util.bind_opt uu___8 (fun eff1 -> - let uu___8 = unembed' w e_term res in - FStar_Compiler_Util.bind_opt uu___8 + let uu___9 = unembed' w e_term res in + FStar_Compiler_Util.bind_opt uu___9 (fun res1 -> - let uu___9 = - let uu___10 = + let uu___10 = + let uu___11 = FStar_Syntax_Embeddings.e_list e_argv in - unembed' w uu___10 args1 in - FStar_Compiler_Util.bind_opt uu___9 + unembed' w uu___11 args1 in + FStar_Compiler_Util.bind_opt uu___10 (fun args2 -> - FStar_Compiler_Effect.op_Less_Bar - (fun uu___10 -> - FStar_Pervasives_Native.Some uu___10) - (FStar_Reflection_Data.C_Eff - ([], eff1, res1, args2)))))) + let uu___11 = + let uu___12 = + FStar_Syntax_Embeddings.e_list e_term in + unembed' w uu___12 decrs in + FStar_Compiler_Util.bind_opt uu___11 + (fun decrs1 -> + FStar_Compiler_Effect.op_Less_Bar + (fun uu___12 -> + FStar_Pervasives_Native.Some + uu___12) + (FStar_Reflection_Data.C_Eff + (us1, eff1, res1, args2, decrs1))))))) | uu___2 -> (if w then diff --git a/src/ocaml-output/FStar_Reflection_NBEEmbeddings.ml b/src/ocaml-output/FStar_Reflection_NBEEmbeddings.ml index e817f618bdb..29a02076816 100644 --- a/src/ocaml-output/FStar_Reflection_NBEEmbeddings.ml +++ b/src/ocaml-output/FStar_Reflection_NBEEmbeddings.ml @@ -1514,45 +1514,21 @@ let (e_comp_view : FStar_Reflection_Data.comp_view FStar_TypeChecker_NBETerm.embedding) = let embed_comp_view cb cv = match cv with - | FStar_Reflection_Data.C_Total (t, u, md) -> + | FStar_Reflection_Data.C_Total t -> let uu___ = let uu___1 = let uu___2 = FStar_TypeChecker_NBETerm.embed e_term cb t in FStar_TypeChecker_NBETerm.as_arg uu___2 in - let uu___2 = - let uu___3 = - let uu___4 = FStar_TypeChecker_NBETerm.embed e_universe cb u in - FStar_TypeChecker_NBETerm.as_arg uu___4 in - let uu___4 = - let uu___5 = - let uu___6 = - let uu___7 = FStar_TypeChecker_NBETerm.e_list e_term in - FStar_TypeChecker_NBETerm.embed uu___7 cb md in - FStar_TypeChecker_NBETerm.as_arg uu___6 in - [uu___5] in - uu___3 :: uu___4 in - uu___1 :: uu___2 in + [uu___1] in mkConstruct FStar_Reflection_Constants.ref_C_Total.FStar_Reflection_Constants.fv [] uu___ - | FStar_Reflection_Data.C_GTotal (t, u, md) -> + | FStar_Reflection_Data.C_GTotal t -> let uu___ = let uu___1 = let uu___2 = FStar_TypeChecker_NBETerm.embed e_term cb t in FStar_TypeChecker_NBETerm.as_arg uu___2 in - let uu___2 = - let uu___3 = - let uu___4 = FStar_TypeChecker_NBETerm.embed e_universe cb u in - FStar_TypeChecker_NBETerm.as_arg uu___4 in - let uu___4 = - let uu___5 = - let uu___6 = - let uu___7 = FStar_TypeChecker_NBETerm.e_list e_term in - FStar_TypeChecker_NBETerm.embed uu___7 cb md in - FStar_TypeChecker_NBETerm.as_arg uu___6 in - [uu___5] in - uu___3 :: uu___4 in - uu___1 :: uu___2 in + [uu___1] in mkConstruct FStar_Reflection_Constants.ref_C_GTotal.FStar_Reflection_Constants.fv [] uu___ @@ -1575,7 +1551,7 @@ let (e_comp_view : mkConstruct FStar_Reflection_Constants.ref_C_Lemma.FStar_Reflection_Constants.fv [] uu___ - | FStar_Reflection_Data.C_Eff (us, eff, res, args) -> + | FStar_Reflection_Data.C_Eff (us, eff, res, args, decrs) -> let uu___ = let uu___1 = let uu___2 = @@ -1598,7 +1574,14 @@ let (e_comp_view : let uu___9 = FStar_TypeChecker_NBETerm.e_list e_argv in FStar_TypeChecker_NBETerm.embed uu___9 cb args in FStar_TypeChecker_NBETerm.as_arg uu___8 in - [uu___7] in + let uu___8 = + let uu___9 = + let uu___10 = + let uu___11 = FStar_TypeChecker_NBETerm.e_list e_term in + FStar_TypeChecker_NBETerm.embed uu___11 cb decrs in + FStar_TypeChecker_NBETerm.as_arg uu___10 in + [uu___9] in + uu___7 :: uu___8 in uu___5 :: uu___6 in uu___3 :: uu___4 in uu___1 :: uu___2 in @@ -1607,44 +1590,26 @@ let (e_comp_view : [] uu___ in let unembed_comp_view cb t = match t.FStar_TypeChecker_NBETerm.nbe_t with - | FStar_TypeChecker_NBETerm.Construct - (fv, uu___, (md, uu___1)::(u, uu___2)::(t1, uu___3)::[]) when + | FStar_TypeChecker_NBETerm.Construct (fv, uu___, (t1, uu___1)::[]) when FStar_Syntax_Syntax.fv_eq_lid fv FStar_Reflection_Constants.ref_C_Total.FStar_Reflection_Constants.lid -> - let uu___4 = FStar_TypeChecker_NBETerm.unembed e_term cb t1 in - FStar_Compiler_Util.bind_opt uu___4 + let uu___2 = FStar_TypeChecker_NBETerm.unembed e_term cb t1 in + FStar_Compiler_Util.bind_opt uu___2 (fun t2 -> - let uu___5 = FStar_TypeChecker_NBETerm.unembed e_universe cb u in - FStar_Compiler_Util.bind_opt uu___5 - (fun u1 -> - let uu___6 = - let uu___7 = FStar_TypeChecker_NBETerm.e_list e_term in - FStar_TypeChecker_NBETerm.unembed uu___7 cb md in - FStar_Compiler_Util.bind_opt uu___6 - (fun md1 -> - FStar_Compiler_Effect.op_Less_Bar - (fun uu___7 -> FStar_Pervasives_Native.Some uu___7) - (FStar_Reflection_Data.C_Total (t2, u1, md1))))) - | FStar_TypeChecker_NBETerm.Construct - (fv, uu___, (md, uu___1)::(u, uu___2)::(t1, uu___3)::[]) when + FStar_Compiler_Effect.op_Less_Bar + (fun uu___3 -> FStar_Pervasives_Native.Some uu___3) + (FStar_Reflection_Data.C_Total t2)) + | FStar_TypeChecker_NBETerm.Construct (fv, uu___, (t1, uu___1)::[]) when FStar_Syntax_Syntax.fv_eq_lid fv FStar_Reflection_Constants.ref_C_GTotal.FStar_Reflection_Constants.lid -> - let uu___4 = FStar_TypeChecker_NBETerm.unembed e_term cb t1 in - FStar_Compiler_Util.bind_opt uu___4 + let uu___2 = FStar_TypeChecker_NBETerm.unembed e_term cb t1 in + FStar_Compiler_Util.bind_opt uu___2 (fun t2 -> - let uu___5 = FStar_TypeChecker_NBETerm.unembed e_universe cb u in - FStar_Compiler_Util.bind_opt uu___5 - (fun u1 -> - let uu___6 = - let uu___7 = FStar_TypeChecker_NBETerm.e_list e_term in - FStar_TypeChecker_NBETerm.unembed uu___7 cb md in - FStar_Compiler_Util.bind_opt uu___6 - (fun md1 -> - FStar_Compiler_Effect.op_Less_Bar - (fun uu___7 -> FStar_Pervasives_Native.Some uu___7) - (FStar_Reflection_Data.C_GTotal (t2, u1, md1))))) + FStar_Compiler_Effect.op_Less_Bar + (fun uu___3 -> FStar_Pervasives_Native.Some uu___3) + (FStar_Reflection_Data.C_GTotal t2)) | FStar_TypeChecker_NBETerm.Construct (fv, uu___, (post, uu___1)::(pre, uu___2)::(pats, uu___3)::[]) when FStar_Syntax_Syntax.fv_eq_lid fv @@ -1665,35 +1630,44 @@ let (e_comp_view : (FStar_Reflection_Data.C_Lemma (pre1, post1, pats1))))) | FStar_TypeChecker_NBETerm.Construct (fv, uu___, - (args, uu___1)::(res, uu___2)::(eff, uu___3)::(us, uu___4)::[]) + (decrs, uu___1)::(args, uu___2)::(res, uu___3)::(eff, uu___4):: + (us, uu___5)::[]) when FStar_Syntax_Syntax.fv_eq_lid fv FStar_Reflection_Constants.ref_C_Eff.FStar_Reflection_Constants.lid -> - let uu___5 = - let uu___6 = FStar_TypeChecker_NBETerm.e_list e_universe in - FStar_TypeChecker_NBETerm.unembed uu___6 cb us in - FStar_Compiler_Util.bind_opt uu___5 + let uu___6 = + let uu___7 = FStar_TypeChecker_NBETerm.e_list e_universe in + FStar_TypeChecker_NBETerm.unembed uu___7 cb us in + FStar_Compiler_Util.bind_opt uu___6 (fun us1 -> - let uu___6 = + let uu___7 = FStar_TypeChecker_NBETerm.unembed FStar_TypeChecker_NBETerm.e_string_list cb eff in - FStar_Compiler_Util.bind_opt uu___6 + FStar_Compiler_Util.bind_opt uu___7 (fun eff1 -> - let uu___7 = + let uu___8 = FStar_TypeChecker_NBETerm.unembed e_term cb res in - FStar_Compiler_Util.bind_opt uu___7 + FStar_Compiler_Util.bind_opt uu___8 (fun res1 -> - let uu___8 = - let uu___9 = FStar_TypeChecker_NBETerm.e_list e_argv in - FStar_TypeChecker_NBETerm.unembed uu___9 cb args in - FStar_Compiler_Util.bind_opt uu___8 + let uu___9 = + let uu___10 = + FStar_TypeChecker_NBETerm.e_list e_argv in + FStar_TypeChecker_NBETerm.unembed uu___10 cb args in + FStar_Compiler_Util.bind_opt uu___9 (fun args1 -> - FStar_Compiler_Effect.op_Less_Bar - (fun uu___9 -> - FStar_Pervasives_Native.Some uu___9) - (FStar_Reflection_Data.C_Eff - (us1, eff1, res1, args1)))))) + let uu___10 = + let uu___11 = + FStar_TypeChecker_NBETerm.e_list e_term in + FStar_TypeChecker_NBETerm.unembed uu___11 cb + decrs in + FStar_Compiler_Util.bind_opt uu___10 + (fun decrs1 -> + FStar_Compiler_Effect.op_Less_Bar + (fun uu___11 -> + FStar_Pervasives_Native.Some uu___11) + (FStar_Reflection_Data.C_Eff + (us1, eff1, res1, args1, decrs1))))))) | uu___ -> ((let uu___2 = let uu___3 = diff --git a/src/ocaml-output/FStar_Syntax_Free.ml b/src/ocaml-output/FStar_Syntax_Free.ml index 18062e2ce44..347032798e3 100644 --- a/src/ocaml-output/FStar_Syntax_Free.ml +++ b/src/ocaml-output/FStar_Syntax_Free.ml @@ -384,20 +384,9 @@ and (free_names_and_uvars_comp : | uu___1 -> let n = match c.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.GTotal (t, FStar_Pervasives_Native.None) -> + | FStar_Syntax_Syntax.GTotal t -> free_names_and_uvars t use_cache - | FStar_Syntax_Syntax.Total (t, FStar_Pervasives_Native.None) -> - free_names_and_uvars t use_cache - | FStar_Syntax_Syntax.GTotal (t, FStar_Pervasives_Native.Some u) - -> - let uu___2 = free_univs u in - let uu___3 = free_names_and_uvars t use_cache in - union uu___2 uu___3 - | FStar_Syntax_Syntax.Total (t, FStar_Pervasives_Native.Some u) - -> - let uu___2 = free_univs u in - let uu___3 = free_names_and_uvars t use_cache in - union uu___2 uu___3 + | FStar_Syntax_Syntax.Total t -> free_names_and_uvars t use_cache | FStar_Syntax_Syntax.Comp ct -> let decreases_vars = let uu___2 = diff --git a/src/ocaml-output/FStar_Syntax_Hash.ml b/src/ocaml-output/FStar_Syntax_Hash.ml index 6d5d314eb2c..57dd69fd395 100644 --- a/src/ocaml-output/FStar_Syntax_Hash.ml +++ b/src/ocaml-output/FStar_Syntax_Hash.ml @@ -211,25 +211,17 @@ and (hash_term' : FStar_Syntax_Syntax.term -> FStar_Hash.hash_code mm) = and (hash_comp' : FStar_Syntax_Syntax.comp -> FStar_Hash.hash_code mm) = fun c -> match c.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.Total (t, ou) -> + | FStar_Syntax_Syntax.Total t -> let uu___ = let uu___1 = of_int (Prims.of_int (811)) in - let uu___2 = - let uu___3 = hash_term t in - let uu___4 = - let uu___5 = hash_option hash_universe ou in [uu___5] in - uu___3 :: uu___4 in - uu___1 :: uu___2 in + let uu___2 = let uu___3 = hash_term t in [uu___3] in uu___1 :: + uu___2 in mix_list_lit uu___ - | FStar_Syntax_Syntax.GTotal (t, ou) -> + | FStar_Syntax_Syntax.GTotal t -> let uu___ = let uu___1 = of_int (Prims.of_int (821)) in - let uu___2 = - let uu___3 = hash_term t in - let uu___4 = - let uu___5 = hash_option hash_universe ou in [uu___5] in - uu___3 :: uu___4 in - uu___1 :: uu___2 in + let uu___2 = let uu___3 = hash_term t in [uu___3] in uu___1 :: + uu___2 in mix_list_lit uu___ | FStar_Syntax_Syntax.Comp ct -> let uu___ = @@ -776,12 +768,10 @@ and (equal_comp : then true else (match ((c1.FStar_Syntax_Syntax.n), (c2.FStar_Syntax_Syntax.n)) with - | (FStar_Syntax_Syntax.Total (t1, u1), FStar_Syntax_Syntax.Total - (t2, u2)) -> - (equal_term t1 t2) && (equal_opt equal_universe u1 u2) - | (FStar_Syntax_Syntax.GTotal (t1, u1), FStar_Syntax_Syntax.GTotal - (t2, u2)) -> - (equal_term t1 t2) && (equal_opt equal_universe u1 u2) + | (FStar_Syntax_Syntax.Total t1, FStar_Syntax_Syntax.Total t2) -> + equal_term t1 t2 + | (FStar_Syntax_Syntax.GTotal t1, FStar_Syntax_Syntax.GTotal t2) -> + equal_term t1 t2 | (FStar_Syntax_Syntax.Comp ct1, FStar_Syntax_Syntax.Comp ct2) -> ((((FStar_Ident.lid_equals ct1.FStar_Syntax_Syntax.effect_name ct2.FStar_Syntax_Syntax.effect_name) diff --git a/src/ocaml-output/FStar_Syntax_InstFV.ml b/src/ocaml-output/FStar_Syntax_InstFV.ml index 09b1773fbb7..e005b735aa2 100644 --- a/src/ocaml-output/FStar_Syntax_InstFV.ml +++ b/src/ocaml-output/FStar_Syntax_InstFV.ml @@ -213,10 +213,10 @@ and (inst_comp : fun s -> fun c -> match c.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.Total (t, uopt) -> - let uu___ = inst s t in FStar_Syntax_Syntax.mk_Total' uu___ uopt - | FStar_Syntax_Syntax.GTotal (t, uopt) -> - let uu___ = inst s t in FStar_Syntax_Syntax.mk_GTotal' uu___ uopt + | FStar_Syntax_Syntax.Total t -> + let uu___ = inst s t in FStar_Syntax_Syntax.mk_Total uu___ + | FStar_Syntax_Syntax.GTotal t -> + let uu___ = inst s t in FStar_Syntax_Syntax.mk_GTotal uu___ | FStar_Syntax_Syntax.Comp ct -> let ct1 = let uu___ = inst s ct.FStar_Syntax_Syntax.result_typ in diff --git a/src/ocaml-output/FStar_Syntax_Print.ml b/src/ocaml-output/FStar_Syntax_Print.ml index 61b4d1b9d24..256a8c7f4e1 100644 --- a/src/ocaml-output/FStar_Syntax_Print.ml +++ b/src/ocaml-output/FStar_Syntax_Print.ml @@ -1040,7 +1040,7 @@ and (comp_to_string : FStar_Syntax_Syntax.comp -> Prims.string) = FStar_Errors.with_ctx "While ugly-printing a computation" (fun uu___2 -> match c.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.Total (t, uopt) -> + | FStar_Syntax_Syntax.Total t -> let uu___3 = let uu___4 = FStar_Syntax_Subst.compress t in uu___4.FStar_Syntax_Syntax.n in @@ -1051,17 +1051,9 @@ and (comp_to_string : FStar_Syntax_Syntax.comp -> Prims.string) = (FStar_Options.print_universes ()) in Prims.op_Negation uu___5 -> term_to_string t | uu___4 -> - (match uopt with - | FStar_Pervasives_Native.Some u when - FStar_Options.print_universes () -> - let uu___5 = univ_to_string u in - let uu___6 = term_to_string t in - FStar_Compiler_Util.format2 "Tot<%s> %s" uu___5 - uu___6 - | uu___5 -> - let uu___6 = term_to_string t in - FStar_Compiler_Util.format1 "Tot %s" uu___6)) - | FStar_Syntax_Syntax.GTotal (t, uopt) -> + let uu___5 = term_to_string t in + FStar_Compiler_Util.format1 "Tot %s" uu___5) + | FStar_Syntax_Syntax.GTotal t -> let uu___3 = let uu___4 = FStar_Syntax_Subst.compress t in uu___4.FStar_Syntax_Syntax.n in @@ -1072,16 +1064,8 @@ and (comp_to_string : FStar_Syntax_Syntax.comp -> Prims.string) = (FStar_Options.print_universes ()) in Prims.op_Negation uu___5 -> term_to_string t | uu___4 -> - (match uopt with - | FStar_Pervasives_Native.Some u when - FStar_Options.print_universes () -> - let uu___5 = univ_to_string u in - let uu___6 = term_to_string t in - FStar_Compiler_Util.format2 "GTot<%s> %s" uu___5 - uu___6 - | uu___5 -> - let uu___6 = term_to_string t in - FStar_Compiler_Util.format1 "GTot %s" uu___6)) + let uu___5 = term_to_string t in + FStar_Compiler_Util.format1 "GTot %s" uu___5) | FStar_Syntax_Syntax.Comp c1 -> let basic = let uu___3 = FStar_Options.print_effect_args () in diff --git a/src/ocaml-output/FStar_Syntax_Resugar.ml b/src/ocaml-output/FStar_Syntax_Resugar.ml index f79b5bae5bf..3efe05f2d9d 100644 --- a/src/ocaml-output/FStar_Syntax_Resugar.ml +++ b/src/ocaml-output/FStar_Syntax_Resugar.ml @@ -1705,7 +1705,7 @@ and (resugar_comp' : FStar_Parser_AST.mk_term a c.FStar_Syntax_Syntax.pos FStar_Parser_AST.Un in match c.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.Total (typ, u) -> + | FStar_Syntax_Syntax.Total typ -> let t = resugar_term' env typ in let uu___ = let uu___1 = FStar_Options.print_implicits () in @@ -1713,50 +1713,16 @@ and (resugar_comp' : if uu___ then t else - (match u with - | FStar_Pervasives_Native.None -> - mk - (FStar_Parser_AST.Construct - (FStar_Parser_Const.effect_Tot_lid, - [(t, FStar_Parser_AST.Nothing)])) - | FStar_Pervasives_Native.Some u1 -> - let uu___2 = FStar_Options.print_universes () in - if uu___2 - then - let u2 = resugar_universe u1 c.FStar_Syntax_Syntax.pos in - mk - (FStar_Parser_AST.Construct - (FStar_Parser_Const.effect_Tot_lid, - [(u2, FStar_Parser_AST.UnivApp); - (t, FStar_Parser_AST.Nothing)])) - else - mk - (FStar_Parser_AST.Construct - (FStar_Parser_Const.effect_Tot_lid, - [(t, FStar_Parser_AST.Nothing)]))) - | FStar_Syntax_Syntax.GTotal (typ, u) -> + mk + (FStar_Parser_AST.Construct + (FStar_Parser_Const.effect_Tot_lid, + [(t, FStar_Parser_AST.Nothing)])) + | FStar_Syntax_Syntax.GTotal typ -> let t = resugar_term' env typ in - (match u with - | FStar_Pervasives_Native.None -> - mk - (FStar_Parser_AST.Construct - (FStar_Parser_Const.effect_GTot_lid, - [(t, FStar_Parser_AST.Nothing)])) - | FStar_Pervasives_Native.Some u1 -> - let uu___ = FStar_Options.print_universes () in - if uu___ - then - let u2 = resugar_universe u1 c.FStar_Syntax_Syntax.pos in - mk - (FStar_Parser_AST.Construct - (FStar_Parser_Const.effect_GTot_lid, - [(u2, FStar_Parser_AST.UnivApp); - (t, FStar_Parser_AST.Nothing)])) - else - mk - (FStar_Parser_AST.Construct - (FStar_Parser_Const.effect_GTot_lid, - [(t, FStar_Parser_AST.Nothing)]))) + mk + (FStar_Parser_AST.Construct + (FStar_Parser_Const.effect_GTot_lid, + [(t, FStar_Parser_AST.Nothing)])) | FStar_Syntax_Syntax.Comp c1 -> let result = let uu___ = resugar_term' env c1.FStar_Syntax_Syntax.result_typ in diff --git a/src/ocaml-output/FStar_Syntax_Subst.ml b/src/ocaml-output/FStar_Syntax_Subst.ml index 2ad4932d382..a6270981207 100644 --- a/src/ocaml-output/FStar_Syntax_Subst.ml +++ b/src/ocaml-output/FStar_Syntax_Subst.ml @@ -462,18 +462,12 @@ let (subst_comp' : | ([]::[], FStar_Syntax_Syntax.NoUseRange) -> t | uu___ -> (match t.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.Total (t1, uopt) -> + | FStar_Syntax_Syntax.Total t1 -> let uu___1 = subst' s t1 in - let uu___2 = - FStar_Compiler_Option.map - (subst_univ (FStar_Pervasives_Native.fst s)) uopt in - FStar_Syntax_Syntax.mk_Total' uu___1 uu___2 - | FStar_Syntax_Syntax.GTotal (t1, uopt) -> + FStar_Syntax_Syntax.mk_Total uu___1 + | FStar_Syntax_Syntax.GTotal t1 -> let uu___1 = subst' s t1 in - let uu___2 = - FStar_Compiler_Option.map - (subst_univ (FStar_Pervasives_Native.fst s)) uopt in - FStar_Syntax_Syntax.mk_GTotal' uu___1 uu___2 + FStar_Syntax_Syntax.mk_GTotal uu___1 | FStar_Syntax_Syntax.Comp ct -> let uu___1 = subst_comp_typ' s ct in FStar_Syntax_Syntax.mk_Comp uu___1) @@ -2029,20 +2023,14 @@ and (deep_compress_comp : fun c -> let mk x = FStar_Syntax_Syntax.mk x c.FStar_Syntax_Syntax.pos in match c.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.Total (t, uopt) -> - let uopt1 = - FStar_Compiler_Util.map_opt uopt (deep_compress_univ allow_uvars) in + | FStar_Syntax_Syntax.Total t -> let uu___ = - let uu___1 = - let uu___2 = deep_compress allow_uvars t in (uu___2, uopt1) in + let uu___1 = deep_compress allow_uvars t in FStar_Syntax_Syntax.Total uu___1 in mk uu___ - | FStar_Syntax_Syntax.GTotal (t, uopt) -> - let uopt1 = - FStar_Compiler_Util.map_opt uopt (deep_compress_univ allow_uvars) in + | FStar_Syntax_Syntax.GTotal t -> let uu___ = - let uu___1 = - let uu___2 = deep_compress allow_uvars t in (uu___2, uopt1) in + let uu___1 = deep_compress allow_uvars t in FStar_Syntax_Syntax.GTotal uu___1 in mk uu___ | FStar_Syntax_Syntax.Comp ct -> diff --git a/src/ocaml-output/FStar_Syntax_Syntax.ml b/src/ocaml-output/FStar_Syntax_Syntax.ml index 7b5ae8d4f4d..34128fdcd37 100644 --- a/src/ocaml-output/FStar_Syntax_Syntax.ml +++ b/src/ocaml-output/FStar_Syntax_Syntax.ml @@ -263,8 +263,8 @@ and comp_typ = (term' syntax * arg_qualifier FStar_Pervasives_Native.option) Prims.list ; flags: cflag Prims.list } and comp' = - | Total of (term' syntax * universe FStar_Pervasives_Native.option) - | GTotal of (term' syntax * universe FStar_Pervasives_Native.option) + | Total of term' syntax + | GTotal of term' syntax | Comp of comp_typ and binder = { @@ -645,13 +645,11 @@ let (__proj__Mkcomp_typ__item__flags : comp_typ -> cflag Prims.list) = | { comp_univs; effect_name; result_typ; effect_args; flags;_} -> flags let (uu___is_Total : comp' -> Prims.bool) = fun projectee -> match projectee with | Total _0 -> true | uu___ -> false -let (__proj__Total__item___0 : - comp' -> (term' syntax * universe FStar_Pervasives_Native.option)) = +let (__proj__Total__item___0 : comp' -> term' syntax) = fun projectee -> match projectee with | Total _0 -> _0 let (uu___is_GTotal : comp' -> Prims.bool) = fun projectee -> match projectee with | GTotal _0 -> true | uu___ -> false -let (__proj__GTotal__item___0 : - comp' -> (term' syntax * universe FStar_Pervasives_Native.option)) = +let (__proj__GTotal__item___0 : comp' -> term' syntax) = fun projectee -> match projectee with | GTotal _0 -> _0 let (uu___is_Comp : comp' -> Prims.bool) = fun projectee -> match projectee with | Comp _0 -> true | uu___ -> false @@ -1792,14 +1790,8 @@ let (extend_app : term -> arg -> FStar_Compiler_Range.range -> term) = fun t -> fun arg1 -> fun r -> extend_app_n t [arg1] r let (mk_Tm_delayed : (term * subst_ts) -> FStar_Compiler_Range.range -> term) = fun lr -> fun pos -> mk (Tm_delayed lr) pos -let (mk_Total' : typ -> universe FStar_Pervasives_Native.option -> comp) = - fun t -> fun u -> mk (Total (t, u)) t.pos -let (mk_GTotal' : typ -> universe FStar_Pervasives_Native.option -> comp) = - fun t -> fun u -> mk (GTotal (t, u)) t.pos -let (mk_Total : typ -> comp) = - fun t -> mk_Total' t FStar_Pervasives_Native.None -let (mk_GTotal : typ -> comp) = - fun t -> mk_GTotal' t FStar_Pervasives_Native.None +let (mk_Total : typ -> comp) = fun t -> mk (Total t) t.pos +let (mk_GTotal : typ -> comp) = fun t -> mk (GTotal t) t.pos let (mk_Comp : comp_typ -> comp) = fun ct -> mk (Comp ct) (ct.result_typ).pos let (mk_lb : (lbname * univ_name Prims.list * FStar_Ident.lident * typ * term * diff --git a/src/ocaml-output/FStar_Syntax_Util.ml b/src/ocaml-output/FStar_Syntax_Util.ml index 6c213a2597c..a092ceae4a6 100644 --- a/src/ocaml-output/FStar_Syntax_Util.ml +++ b/src/ocaml-output/FStar_Syntax_Util.ml @@ -398,83 +398,20 @@ let (comp_flags : | FStar_Syntax_Syntax.Total uu___ -> [FStar_Syntax_Syntax.TOTAL] | FStar_Syntax_Syntax.GTotal uu___ -> [FStar_Syntax_Syntax.SOMETRIVIAL] | FStar_Syntax_Syntax.Comp ct -> ct.FStar_Syntax_Syntax.flags -let (comp_to_comp_typ_nouniv : - FStar_Syntax_Syntax.comp -> FStar_Syntax_Syntax.comp_typ) = - fun c -> - match c.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.Comp c1 -> c1 - | FStar_Syntax_Syntax.Total (t, u_opt) -> - let uu___ = - let uu___1 = FStar_Compiler_Util.map_opt u_opt (fun x -> [x]) in - FStar_Compiler_Util.dflt [] uu___1 in - { - FStar_Syntax_Syntax.comp_univs = uu___; - FStar_Syntax_Syntax.effect_name = (comp_effect_name c); - FStar_Syntax_Syntax.result_typ = t; - FStar_Syntax_Syntax.effect_args = []; - FStar_Syntax_Syntax.flags = (comp_flags c) - } - | FStar_Syntax_Syntax.GTotal (t, u_opt) -> - let uu___ = - let uu___1 = FStar_Compiler_Util.map_opt u_opt (fun x -> [x]) in - FStar_Compiler_Util.dflt [] uu___1 in - { - FStar_Syntax_Syntax.comp_univs = uu___; - FStar_Syntax_Syntax.effect_name = (comp_effect_name c); - FStar_Syntax_Syntax.result_typ = t; - FStar_Syntax_Syntax.effect_args = []; - FStar_Syntax_Syntax.flags = (comp_flags c) - } -let (comp_set_flags : +let (comp_eff_name_res_and_args : FStar_Syntax_Syntax.comp -> - FStar_Syntax_Syntax.cflag Prims.list -> - FStar_Syntax_Syntax.comp' FStar_Syntax_Syntax.syntax) + (FStar_Ident.lident * FStar_Syntax_Syntax.typ * FStar_Syntax_Syntax.args)) = - fun c -> - fun f -> - let uu___ = - let uu___1 = - let uu___2 = comp_to_comp_typ_nouniv c in - { - FStar_Syntax_Syntax.comp_univs = - (uu___2.FStar_Syntax_Syntax.comp_univs); - FStar_Syntax_Syntax.effect_name = - (uu___2.FStar_Syntax_Syntax.effect_name); - FStar_Syntax_Syntax.result_typ = - (uu___2.FStar_Syntax_Syntax.result_typ); - FStar_Syntax_Syntax.effect_args = - (uu___2.FStar_Syntax_Syntax.effect_args); - FStar_Syntax_Syntax.flags = f - } in - FStar_Syntax_Syntax.Comp uu___1 in - { - FStar_Syntax_Syntax.n = uu___; - FStar_Syntax_Syntax.pos = (c.FStar_Syntax_Syntax.pos); - FStar_Syntax_Syntax.vars = (c.FStar_Syntax_Syntax.vars); - FStar_Syntax_Syntax.hash_code = (c.FStar_Syntax_Syntax.hash_code) - } -let (comp_to_comp_typ : - FStar_Syntax_Syntax.comp -> FStar_Syntax_Syntax.comp_typ) = fun c -> match c.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.Comp c1 -> c1 - | FStar_Syntax_Syntax.Total (t, FStar_Pervasives_Native.Some u) -> - { - FStar_Syntax_Syntax.comp_univs = [u]; - FStar_Syntax_Syntax.effect_name = (comp_effect_name c); - FStar_Syntax_Syntax.result_typ = t; - FStar_Syntax_Syntax.effect_args = []; - FStar_Syntax_Syntax.flags = (comp_flags c) - } - | FStar_Syntax_Syntax.GTotal (t, FStar_Pervasives_Native.Some u) -> - { - FStar_Syntax_Syntax.comp_univs = [u]; - FStar_Syntax_Syntax.effect_name = (comp_effect_name c); - FStar_Syntax_Syntax.result_typ = t; - FStar_Syntax_Syntax.effect_args = []; - FStar_Syntax_Syntax.flags = (comp_flags c) - } - | uu___ -> failwith "Assertion failed: Computation type without universe" + | FStar_Syntax_Syntax.Total t -> + (FStar_Parser_Const.effect_Tot_lid, t, []) + | FStar_Syntax_Syntax.GTotal t -> + (FStar_Parser_Const.effect_GTot_lid, t, []) + | FStar_Syntax_Syntax.Comp c1 -> + ((c1.FStar_Syntax_Syntax.effect_name), + (c1.FStar_Syntax_Syntax.result_typ), + (c1.FStar_Syntax_Syntax.effect_args)) let (effect_indices_from_repr : FStar_Syntax_Syntax.term -> Prims.bool -> @@ -500,12 +437,15 @@ let (effect_indices_from_repr : (match repr1.FStar_Syntax_Syntax.n with | FStar_Syntax_Syntax.Tm_arrow (uu___1, c) -> let uu___2 = - FStar_Compiler_Effect.op_Bar_Greater c comp_to_comp_typ in + FStar_Compiler_Effect.op_Bar_Greater c + comp_eff_name_res_and_args in FStar_Compiler_Effect.op_Bar_Greater uu___2 - (fun ct -> - FStar_Compiler_Effect.op_Bar_Greater - ct.FStar_Syntax_Syntax.effect_args - (FStar_Compiler_List.map FStar_Pervasives_Native.fst)) + (fun uu___3 -> + match uu___3 with + | (uu___4, uu___5, args) -> + FStar_Compiler_Effect.op_Bar_Greater args + (FStar_Compiler_List.map + FStar_Pervasives_Native.fst)) | uu___1 -> err1 ()) let (destruct_comp : FStar_Syntax_Syntax.comp_typ -> @@ -735,8 +675,8 @@ let (comp_result : = fun c -> match c.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.Total (t, uu___) -> t - | FStar_Syntax_Syntax.GTotal (t, uu___) -> t + | FStar_Syntax_Syntax.Total t -> t + | FStar_Syntax_Syntax.GTotal t -> t | FStar_Syntax_Syntax.Comp ct -> ct.FStar_Syntax_Syntax.result_typ let (set_result_typ : FStar_Syntax_Syntax.comp' FStar_Syntax_Syntax.syntax -> @@ -1253,10 +1193,10 @@ and (eq_comp : fun c1 -> fun c2 -> match ((c1.FStar_Syntax_Syntax.n), (c2.FStar_Syntax_Syntax.n)) with - | (FStar_Syntax_Syntax.Total (t1, u1opt), FStar_Syntax_Syntax.Total - (t2, u2opt)) -> eq_tm t1 t2 - | (FStar_Syntax_Syntax.GTotal (t1, u1opt), FStar_Syntax_Syntax.GTotal - (t2, u2opt)) -> eq_tm t1 t2 + | (FStar_Syntax_Syntax.Total t1, FStar_Syntax_Syntax.Total t2) -> + eq_tm t1 t2 + | (FStar_Syntax_Syntax.GTotal t1, FStar_Syntax_Syntax.GTotal t2) -> + eq_tm t1 t2 | (FStar_Syntax_Syntax.Comp ct1, FStar_Syntax_Syntax.Comp ct2) -> let uu___ = let uu___1 = @@ -1735,21 +1675,21 @@ let (flat_arrow : match uu___ with | FStar_Syntax_Syntax.Tm_arrow (bs1, c1) -> (match c1.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.Total (tres, uu___1) -> - let uu___2 = - let uu___3 = FStar_Syntax_Subst.compress tres in - uu___3.FStar_Syntax_Syntax.n in - (match uu___2 with + | FStar_Syntax_Syntax.Total tres -> + let uu___1 = + let uu___2 = FStar_Syntax_Subst.compress tres in + uu___2.FStar_Syntax_Syntax.n in + (match uu___1 with | FStar_Syntax_Syntax.Tm_arrow (bs', c') -> FStar_Syntax_Syntax.mk (FStar_Syntax_Syntax.Tm_arrow ((FStar_Compiler_List.op_At bs1 bs'), c')) t.FStar_Syntax_Syntax.pos - | uu___3 -> t) + | uu___2 -> t) | uu___1 -> t) | uu___1 -> t let rec (canon_arrow : - FStar_Syntax_Syntax.term' FStar_Syntax_Syntax.syntax -> + FStar_Syntax_Syntax.term -> FStar_Syntax_Syntax.term' FStar_Syntax_Syntax.syntax) = fun t -> @@ -1760,9 +1700,8 @@ let rec (canon_arrow : | FStar_Syntax_Syntax.Tm_arrow (bs, c) -> let cn = match c.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.Total (t1, u) -> - let uu___1 = let uu___2 = canon_arrow t1 in (uu___2, u) in - FStar_Syntax_Syntax.Total uu___1 + | FStar_Syntax_Syntax.Total t1 -> + let uu___1 = canon_arrow t1 in FStar_Syntax_Syntax.Total uu___1 | uu___1 -> c.FStar_Syntax_Syntax.n in let c1 = { @@ -1884,15 +1823,15 @@ let (let_rec_arity : let uu___ = FStar_Syntax_Subst.open_comp bs c in (match uu___ with | (bs1, c1) -> - let ct = comp_to_comp_typ c1 in let uu___1 = - FStar_Compiler_Effect.op_Bar_Greater - ct.FStar_Syntax_Syntax.flags + let uu___2 = + FStar_Compiler_Effect.op_Bar_Greater c1 comp_flags in + FStar_Compiler_Effect.op_Bar_Greater uu___2 (FStar_Compiler_Util.find_opt - (fun uu___2 -> - match uu___2 with - | FStar_Syntax_Syntax.DECREASES uu___3 -> true - | uu___3 -> false)) in + (fun uu___3 -> + match uu___3 with + | FStar_Syntax_Syntax.DECREASES uu___4 -> true + | uu___4 -> false)) in (match uu___1 with | FStar_Pervasives_Native.Some (FStar_Syntax_Syntax.DECREASES d) -> (bs1, (FStar_Pervasives_Native.Some d)) @@ -3338,9 +3277,7 @@ let (destruct_typ_as_formula : if uu___2 then FStar_Pervasives_Native.None else - (let q = - let uu___4 = comp_to_comp_typ_nouniv c in - uu___4.FStar_Syntax_Syntax.result_typ in + (let q = comp_result c in let uu___4 = is_free_in b.FStar_Syntax_Syntax.binder_bv q in if uu___4 then @@ -3902,17 +3839,17 @@ and (comp_eq_dbg : fun dbg -> fun c1 -> fun c2 -> - let c11 = comp_to_comp_typ_nouniv c1 in - let c21 = comp_to_comp_typ_nouniv c2 in - ((let uu___ = - FStar_Ident.lid_equals c11.FStar_Syntax_Syntax.effect_name - c21.FStar_Syntax_Syntax.effect_name in - check "comp eff" uu___) && - (let uu___ = - term_eq_dbg dbg c11.FStar_Syntax_Syntax.result_typ - c21.FStar_Syntax_Syntax.result_typ in - check "comp result typ" uu___)) - && true + let uu___ = comp_eff_name_res_and_args c1 in + match uu___ with + | (eff1, res1, args1) -> + let uu___1 = comp_eff_name_res_and_args c2 in + (match uu___1 with + | (eff2, res2, args2) -> + ((let uu___2 = FStar_Ident.lid_equals eff1 eff2 in + check "comp eff" uu___2) && + (let uu___2 = term_eq_dbg dbg res1 res2 in + check "comp result typ" uu___2)) + && true) and (branch_eq_dbg : Prims.bool -> (FStar_Syntax_Syntax.pat' FStar_Syntax_Syntax.withinfo_t * @@ -4272,8 +4209,8 @@ and (unbound_variables_comp : FStar_Syntax_Syntax.comp -> FStar_Syntax_Syntax.bv Prims.list) = fun c -> match c.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.GTotal (t, uu___) -> unbound_variables t - | FStar_Syntax_Syntax.Total (t, uu___) -> unbound_variables t + | FStar_Syntax_Syntax.GTotal t -> unbound_variables t + | FStar_Syntax_Syntax.Total t -> unbound_variables t | FStar_Syntax_Syntax.Comp ct -> let uu___ = unbound_variables ct.FStar_Syntax_Syntax.result_typ in let uu___1 = diff --git a/src/ocaml-output/FStar_Tactics_Basic.ml b/src/ocaml-output/FStar_Tactics_Basic.ml index ffd2f1583b2..79ea1bc15f7 100644 --- a/src/ocaml-output/FStar_Tactics_Basic.ml +++ b/src/ocaml-output/FStar_Tactics_Basic.ml @@ -3015,40 +3015,38 @@ let (lemma_or_sq : FStar_Pervasives_Native.option) = fun c -> - let ct = FStar_Syntax_Util.comp_to_comp_typ_nouniv c in - let uu___ = - FStar_Ident.lid_equals ct.FStar_Syntax_Syntax.effect_name - FStar_Parser_Const.effect_Lemma_lid in - if uu___ - then - let uu___1 = - match ct.FStar_Syntax_Syntax.effect_args with - | pre::post::uu___2 -> - ((FStar_Pervasives_Native.fst pre), - (FStar_Pervasives_Native.fst post)) - | uu___2 -> failwith "apply_lemma: impossible: not a lemma" in - match uu___1 with - | (pre, post) -> - let post1 = - let uu___2 = - let uu___3 = - FStar_Syntax_Syntax.as_arg FStar_Syntax_Util.exp_unit in - [uu___3] in - FStar_Syntax_Util.mk_app post uu___2 in - FStar_Pervasives_Native.Some (pre, post1) - else - (let uu___2 = - (FStar_Syntax_Util.is_pure_effect ct.FStar_Syntax_Syntax.effect_name) - || - (FStar_Syntax_Util.is_ghost_effect - ct.FStar_Syntax_Syntax.effect_name) in - if uu___2 - then - let uu___3 = - FStar_Syntax_Util.un_squash ct.FStar_Syntax_Syntax.result_typ in - FStar_Compiler_Util.map_opt uu___3 - (fun post -> (FStar_Syntax_Util.t_true, post)) - else FStar_Pervasives_Native.None) + let uu___ = FStar_Syntax_Util.comp_eff_name_res_and_args c in + match uu___ with + | (eff_name, res, args) -> + let uu___1 = + FStar_Ident.lid_equals eff_name FStar_Parser_Const.effect_Lemma_lid in + if uu___1 + then + let uu___2 = + match args with + | pre::post::uu___3 -> + ((FStar_Pervasives_Native.fst pre), + (FStar_Pervasives_Native.fst post)) + | uu___3 -> failwith "apply_lemma: impossible: not a lemma" in + (match uu___2 with + | (pre, post) -> + let post1 = + let uu___3 = + let uu___4 = + FStar_Syntax_Syntax.as_arg FStar_Syntax_Util.exp_unit in + [uu___4] in + FStar_Syntax_Util.mk_app post uu___3 in + FStar_Pervasives_Native.Some (pre, post1)) + else + (let uu___3 = + (FStar_Syntax_Util.is_pure_effect eff_name) || + (FStar_Syntax_Util.is_ghost_effect eff_name) in + if uu___3 + then + let uu___4 = FStar_Syntax_Util.un_squash res in + FStar_Compiler_Util.map_opt uu___4 + (fun post -> (FStar_Syntax_Util.t_true, post)) + else FStar_Pervasives_Native.None) let rec fold_left : 'a 'b . ('a -> 'b -> 'b FStar_Tactics_Monad.tac) -> diff --git a/src/ocaml-output/FStar_Tests_Util.ml b/src/ocaml-output/FStar_Tests_Util.ml index 99922fa3a94..8e6a84c9e73 100644 --- a/src/ocaml-output/FStar_Tests_Util.ml +++ b/src/ocaml-output/FStar_Tests_Util.ml @@ -74,8 +74,8 @@ let rec (term_eq' : uu___2 = FStar_Syntax_Util.Equal)) xs ys) in let comp_eq c d = match ((c.FStar_Syntax_Syntax.n), (d.FStar_Syntax_Syntax.n)) with - | (FStar_Syntax_Syntax.Total (t, uu___), FStar_Syntax_Syntax.Total - (s, uu___1)) -> term_eq' t s + | (FStar_Syntax_Syntax.Total t, FStar_Syntax_Syntax.Total s) -> + term_eq' t s | (FStar_Syntax_Syntax.Comp ct1, FStar_Syntax_Syntax.Comp ct2) -> ((FStar_Ident.lid_equals ct1.FStar_Syntax_Syntax.effect_name ct2.FStar_Syntax_Syntax.effect_name) diff --git a/src/ocaml-output/FStar_TypeChecker_Core.ml b/src/ocaml-output/FStar_TypeChecker_Core.ml index 6048661083b..9c42bb28c8a 100644 --- a/src/ocaml-output/FStar_TypeChecker_Core.ml +++ b/src/ocaml-output/FStar_TypeChecker_Core.ml @@ -2226,46 +2226,50 @@ and (check_relation_comp : if uu___2 then return () else - (let ct_eq ct0 ct1 = - let uu___4 = - check_relation g EQUALITY - ct0.FStar_Syntax_Syntax.result_typ - ct1.FStar_Syntax_Syntax.result_typ in + (let ct_eq res0 args0 res1 args1 = + let uu___4 = check_relation g EQUALITY res0 res1 in op_let_Bang uu___4 (fun uu___5 -> - check_relation_args g EQUALITY - ct0.FStar_Syntax_Syntax.effect_args - ct1.FStar_Syntax_Syntax.effect_args) in - let ct0 = FStar_Syntax_Util.comp_to_comp_typ_nouniv c0 in - let ct1 = FStar_Syntax_Util.comp_to_comp_typ_nouniv c1 in - let uu___4 = - FStar_Ident.lid_equals ct0.FStar_Syntax_Syntax.effect_name - ct1.FStar_Syntax_Syntax.effect_name in - if uu___4 - then ct_eq ct0 ct1 - else - (let ct01 = - FStar_TypeChecker_Env.unfold_effect_abbrev g.tcenv c0 in - let ct11 = - FStar_TypeChecker_Env.unfold_effect_abbrev g.tcenv c1 in - let uu___6 = - FStar_Ident.lid_equals - ct01.FStar_Syntax_Syntax.effect_name - ct11.FStar_Syntax_Syntax.effect_name in - if uu___6 - then ct_eq ct01 ct11 - else - (let uu___8 = - let uu___9 = - FStar_Ident.string_of_lid - ct01.FStar_Syntax_Syntax.effect_name in - let uu___10 = - FStar_Ident.string_of_lid - ct11.FStar_Syntax_Syntax.effect_name in - FStar_Compiler_Util.format2 - "Subcomp failed: Unequal computation types %s and %s" - uu___9 uu___10 in - fail uu___8))) + check_relation_args g EQUALITY args0 args1) in + let uu___4 = FStar_Syntax_Util.comp_eff_name_res_and_args c0 in + match uu___4 with + | (eff0, res0, args0) -> + let uu___5 = + FStar_Syntax_Util.comp_eff_name_res_and_args c1 in + (match uu___5 with + | (eff1, res1, args1) -> + let uu___6 = FStar_Ident.lid_equals eff0 eff1 in + if uu___6 + then ct_eq res0 args0 res1 args1 + else + (let ct0 = + FStar_TypeChecker_Env.unfold_effect_abbrev + g.tcenv c0 in + let ct1 = + FStar_TypeChecker_Env.unfold_effect_abbrev + g.tcenv c1 in + let uu___8 = + FStar_Ident.lid_equals + ct0.FStar_Syntax_Syntax.effect_name + ct1.FStar_Syntax_Syntax.effect_name in + if uu___8 + then + ct_eq ct0.FStar_Syntax_Syntax.result_typ + ct0.FStar_Syntax_Syntax.effect_args + ct1.FStar_Syntax_Syntax.result_typ + ct1.FStar_Syntax_Syntax.effect_args + else + (let uu___10 = + let uu___11 = + FStar_Ident.string_of_lid + ct0.FStar_Syntax_Syntax.effect_name in + let uu___12 = + FStar_Ident.string_of_lid + ct1.FStar_Syntax_Syntax.effect_name in + FStar_Compiler_Util.format2 + "Subcomp failed: Unequal computation types %s and %s" + uu___11 uu___12 in + fail uu___10)))) | (uu___1, FStar_Pervasives_Native.None) -> let uu___2 = let uu___3 = FStar_Syntax_Util.eq_comp c0 c1 in @@ -2273,46 +2277,50 @@ and (check_relation_comp : if uu___2 then return () else - (let ct_eq ct0 ct1 = - let uu___4 = - check_relation g EQUALITY - ct0.FStar_Syntax_Syntax.result_typ - ct1.FStar_Syntax_Syntax.result_typ in + (let ct_eq res0 args0 res1 args1 = + let uu___4 = check_relation g EQUALITY res0 res1 in op_let_Bang uu___4 (fun uu___5 -> - check_relation_args g EQUALITY - ct0.FStar_Syntax_Syntax.effect_args - ct1.FStar_Syntax_Syntax.effect_args) in - let ct0 = FStar_Syntax_Util.comp_to_comp_typ_nouniv c0 in - let ct1 = FStar_Syntax_Util.comp_to_comp_typ_nouniv c1 in - let uu___4 = - FStar_Ident.lid_equals ct0.FStar_Syntax_Syntax.effect_name - ct1.FStar_Syntax_Syntax.effect_name in - if uu___4 - then ct_eq ct0 ct1 - else - (let ct01 = - FStar_TypeChecker_Env.unfold_effect_abbrev g.tcenv c0 in - let ct11 = - FStar_TypeChecker_Env.unfold_effect_abbrev g.tcenv c1 in - let uu___6 = - FStar_Ident.lid_equals - ct01.FStar_Syntax_Syntax.effect_name - ct11.FStar_Syntax_Syntax.effect_name in - if uu___6 - then ct_eq ct01 ct11 - else - (let uu___8 = - let uu___9 = - FStar_Ident.string_of_lid - ct01.FStar_Syntax_Syntax.effect_name in - let uu___10 = - FStar_Ident.string_of_lid - ct11.FStar_Syntax_Syntax.effect_name in - FStar_Compiler_Util.format2 - "Subcomp failed: Unequal computation types %s and %s" - uu___9 uu___10 in - fail uu___8))) + check_relation_args g EQUALITY args0 args1) in + let uu___4 = FStar_Syntax_Util.comp_eff_name_res_and_args c0 in + match uu___4 with + | (eff0, res0, args0) -> + let uu___5 = + FStar_Syntax_Util.comp_eff_name_res_and_args c1 in + (match uu___5 with + | (eff1, res1, args1) -> + let uu___6 = FStar_Ident.lid_equals eff0 eff1 in + if uu___6 + then ct_eq res0 args0 res1 args1 + else + (let ct0 = + FStar_TypeChecker_Env.unfold_effect_abbrev + g.tcenv c0 in + let ct1 = + FStar_TypeChecker_Env.unfold_effect_abbrev + g.tcenv c1 in + let uu___8 = + FStar_Ident.lid_equals + ct0.FStar_Syntax_Syntax.effect_name + ct1.FStar_Syntax_Syntax.effect_name in + if uu___8 + then + ct_eq ct0.FStar_Syntax_Syntax.result_typ + ct0.FStar_Syntax_Syntax.effect_args + ct1.FStar_Syntax_Syntax.result_typ + ct1.FStar_Syntax_Syntax.effect_args + else + (let uu___10 = + let uu___11 = + FStar_Ident.string_of_lid + ct0.FStar_Syntax_Syntax.effect_name in + let uu___12 = + FStar_Ident.string_of_lid + ct1.FStar_Syntax_Syntax.effect_name in + FStar_Compiler_Util.format2 + "Subcomp failed: Unequal computation types %s and %s" + uu___11 uu___12 in + fail uu___10)))) | (FStar_Pervasives_Native.Some (E_TOTAL, t0), FStar_Pervasives_Native.Some (uu___1, t1)) -> check_relation g rel t0 t1 @@ -3360,16 +3368,16 @@ and (check_comp : fun g -> fun c -> match c.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.Total (t, uu___) -> - let uu___1 = + | FStar_Syntax_Syntax.Total t -> + let uu___ = check "(G)Tot comp result" g (FStar_Syntax_Util.comp_result c) in - op_let_Bang uu___1 - (fun uu___2 -> match uu___2 with | (uu___3, t1) -> is_type g t1) - | FStar_Syntax_Syntax.GTotal (t, uu___) -> - let uu___1 = + op_let_Bang uu___ + (fun uu___1 -> match uu___1 with | (uu___2, t1) -> is_type g t1) + | FStar_Syntax_Syntax.GTotal t -> + let uu___ = check "(G)Tot comp result" g (FStar_Syntax_Util.comp_result c) in - op_let_Bang uu___1 - (fun uu___2 -> match uu___2 with | (uu___3, t1) -> is_type g t1) + op_let_Bang uu___ + (fun uu___1 -> match uu___1 with | (uu___2, t1) -> is_type g t1) | FStar_Syntax_Syntax.Comp ct -> if (FStar_Compiler_List.length ct.FStar_Syntax_Syntax.comp_univs) <> diff --git a/src/ocaml-output/FStar_TypeChecker_DMFF.ml b/src/ocaml-output/FStar_TypeChecker_DMFF.ml index 9e592bc2eb4..ff891043e8f 100644 --- a/src/ocaml-output/FStar_TypeChecker_DMFF.ml +++ b/src/ocaml-output/FStar_TypeChecker_DMFF.ml @@ -64,7 +64,7 @@ let (gen_wps_for_free : | FStar_Syntax_Syntax.Tm_arrow (bs, comp) -> let rest = match comp.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.Total (t2, uu___2) -> t2 + | FStar_Syntax_Syntax.Total t2 -> t2 | uu___2 -> let uu___3 = let uu___4 = @@ -745,20 +745,28 @@ let (gen_wps_for_free : (binder::[], { FStar_Syntax_Syntax.n = FStar_Syntax_Syntax.GTotal - (b, uu___4); - FStar_Syntax_Syntax.pos = uu___5; - FStar_Syntax_Syntax.vars = uu___6; - FStar_Syntax_Syntax.hash_code = uu___7;_}) + b; + FStar_Syntax_Syntax.pos = uu___4; + FStar_Syntax_Syntax.vars = uu___5; + FStar_Syntax_Syntax.hash_code = uu___6;_}) -> let a2 = (binder.FStar_Syntax_Syntax.binder_bv).FStar_Syntax_Syntax.sort in - let uu___8 = (is_monotonic a2) || (is_monotonic b) in - if uu___8 + let uu___7 = (is_monotonic a2) || (is_monotonic b) in + if uu___7 then let a11 = FStar_Syntax_Syntax.gen_bv "a1" FStar_Pervasives_Native.None a2 in let body = + let uu___8 = + let uu___9 = + let uu___10 = + let uu___11 = + FStar_Syntax_Syntax.bv_to_name a11 in + FStar_Syntax_Syntax.as_arg uu___11 in + [uu___10] in + FStar_Syntax_Util.mk_app x uu___9 in let uu___9 = let uu___10 = let uu___11 = @@ -766,16 +774,8 @@ let (gen_wps_for_free : FStar_Syntax_Syntax.bv_to_name a11 in FStar_Syntax_Syntax.as_arg uu___12 in [uu___11] in - FStar_Syntax_Util.mk_app x uu___10 in - let uu___10 = - let uu___11 = - let uu___12 = - let uu___13 = - FStar_Syntax_Syntax.bv_to_name a11 in - FStar_Syntax_Syntax.as_arg uu___13 in - [uu___12] in - FStar_Syntax_Util.mk_app y uu___11 in - mk_rel1 b uu___9 uu___10 in + FStar_Syntax_Util.mk_app y uu___10 in + mk_rel1 b uu___8 uu___9 in mk_forall a11 body else (let a11 = @@ -785,51 +785,59 @@ let (gen_wps_for_free : FStar_Syntax_Syntax.gen_bv "a2" FStar_Pervasives_Native.None a2 in let body = - let uu___10 = - let uu___11 = + let uu___9 = + let uu___10 = FStar_Syntax_Syntax.bv_to_name a11 in - let uu___12 = + let uu___11 = FStar_Syntax_Syntax.bv_to_name a21 in - mk_rel1 a2 uu___11 uu___12 in - let uu___11 = + mk_rel1 a2 uu___10 uu___11 in + let uu___10 = + let uu___11 = + let uu___12 = + let uu___13 = + let uu___14 = + FStar_Syntax_Syntax.bv_to_name a11 in + FStar_Syntax_Syntax.as_arg uu___14 in + [uu___13] in + FStar_Syntax_Util.mk_app x uu___12 in let uu___12 = let uu___13 = let uu___14 = let uu___15 = - FStar_Syntax_Syntax.bv_to_name a11 in + FStar_Syntax_Syntax.bv_to_name a21 in FStar_Syntax_Syntax.as_arg uu___15 in [uu___14] in - FStar_Syntax_Util.mk_app x uu___13 in - let uu___13 = - let uu___14 = - let uu___15 = - let uu___16 = - FStar_Syntax_Syntax.bv_to_name a21 in - FStar_Syntax_Syntax.as_arg uu___16 in - [uu___15] in - FStar_Syntax_Util.mk_app y uu___14 in - mk_rel1 b uu___12 uu___13 in - FStar_Syntax_Util.mk_imp uu___10 uu___11 in - let uu___10 = mk_forall a21 body in - mk_forall a11 uu___10) + FStar_Syntax_Util.mk_app y uu___13 in + mk_rel1 b uu___11 uu___12 in + FStar_Syntax_Util.mk_imp uu___9 uu___10 in + let uu___9 = mk_forall a21 body in + mk_forall a11 uu___9) | FStar_Syntax_Syntax.Tm_arrow (binder::[], { FStar_Syntax_Syntax.n = FStar_Syntax_Syntax.Total - (b, uu___4); - FStar_Syntax_Syntax.pos = uu___5; - FStar_Syntax_Syntax.vars = uu___6; - FStar_Syntax_Syntax.hash_code = uu___7;_}) + b; + FStar_Syntax_Syntax.pos = uu___4; + FStar_Syntax_Syntax.vars = uu___5; + FStar_Syntax_Syntax.hash_code = uu___6;_}) -> let a2 = (binder.FStar_Syntax_Syntax.binder_bv).FStar_Syntax_Syntax.sort in - let uu___8 = (is_monotonic a2) || (is_monotonic b) in - if uu___8 + let uu___7 = (is_monotonic a2) || (is_monotonic b) in + if uu___7 then let a11 = FStar_Syntax_Syntax.gen_bv "a1" FStar_Pervasives_Native.None a2 in let body = + let uu___8 = + let uu___9 = + let uu___10 = + let uu___11 = + FStar_Syntax_Syntax.bv_to_name a11 in + FStar_Syntax_Syntax.as_arg uu___11 in + [uu___10] in + FStar_Syntax_Util.mk_app x uu___9 in let uu___9 = let uu___10 = let uu___11 = @@ -837,16 +845,8 @@ let (gen_wps_for_free : FStar_Syntax_Syntax.bv_to_name a11 in FStar_Syntax_Syntax.as_arg uu___12 in [uu___11] in - FStar_Syntax_Util.mk_app x uu___10 in - let uu___10 = - let uu___11 = - let uu___12 = - let uu___13 = - FStar_Syntax_Syntax.bv_to_name a11 in - FStar_Syntax_Syntax.as_arg uu___13 in - [uu___12] in - FStar_Syntax_Util.mk_app y uu___11 in - mk_rel1 b uu___9 uu___10 in + FStar_Syntax_Util.mk_app y uu___10 in + mk_rel1 b uu___8 uu___9 in mk_forall a11 body else (let a11 = @@ -856,33 +856,33 @@ let (gen_wps_for_free : FStar_Syntax_Syntax.gen_bv "a2" FStar_Pervasives_Native.None a2 in let body = - let uu___10 = - let uu___11 = + let uu___9 = + let uu___10 = FStar_Syntax_Syntax.bv_to_name a11 in - let uu___12 = + let uu___11 = FStar_Syntax_Syntax.bv_to_name a21 in - mk_rel1 a2 uu___11 uu___12 in - let uu___11 = + mk_rel1 a2 uu___10 uu___11 in + let uu___10 = + let uu___11 = + let uu___12 = + let uu___13 = + let uu___14 = + FStar_Syntax_Syntax.bv_to_name a11 in + FStar_Syntax_Syntax.as_arg uu___14 in + [uu___13] in + FStar_Syntax_Util.mk_app x uu___12 in let uu___12 = let uu___13 = let uu___14 = let uu___15 = - FStar_Syntax_Syntax.bv_to_name a11 in + FStar_Syntax_Syntax.bv_to_name a21 in FStar_Syntax_Syntax.as_arg uu___15 in [uu___14] in - FStar_Syntax_Util.mk_app x uu___13 in - let uu___13 = - let uu___14 = - let uu___15 = - let uu___16 = - FStar_Syntax_Syntax.bv_to_name a21 in - FStar_Syntax_Syntax.as_arg uu___16 in - [uu___15] in - FStar_Syntax_Util.mk_app y uu___14 in - mk_rel1 b uu___12 uu___13 in - FStar_Syntax_Util.mk_imp uu___10 uu___11 in - let uu___10 = mk_forall a21 body in - mk_forall a11 uu___10) + FStar_Syntax_Util.mk_app y uu___13 in + mk_rel1 b uu___11 uu___12 in + FStar_Syntax_Util.mk_imp uu___9 uu___10 in + let uu___9 = mk_forall a21 body in + mk_forall a11 uu___9) | FStar_Syntax_Syntax.Tm_arrow (binder::binders1, comp) -> let t2 = @@ -968,40 +968,40 @@ let (gen_wps_for_free : (binders1, { FStar_Syntax_Syntax.n = - FStar_Syntax_Syntax.GTotal (b, uu___4); - FStar_Syntax_Syntax.pos = uu___5; - FStar_Syntax_Syntax.vars = uu___6; - FStar_Syntax_Syntax.hash_code = uu___7;_}) + FStar_Syntax_Syntax.GTotal b; + FStar_Syntax_Syntax.pos = uu___4; + FStar_Syntax_Syntax.vars = uu___5; + FStar_Syntax_Syntax.hash_code = uu___6;_}) -> let bvs = FStar_Compiler_List.mapi (fun i -> - fun uu___8 -> - match uu___8 with + fun uu___7 -> + match uu___7 with | { FStar_Syntax_Syntax.binder_bv = bv; FStar_Syntax_Syntax.binder_qual = q; FStar_Syntax_Syntax.binder_attrs = - uu___9;_} + uu___8;_} -> - let uu___10 = - let uu___11 = + let uu___9 = + let uu___10 = FStar_Compiler_Util.string_of_int i in - Prims.op_Hat "a" uu___11 in - FStar_Syntax_Syntax.gen_bv uu___10 + Prims.op_Hat "a" uu___10 in + FStar_Syntax_Syntax.gen_bv uu___9 FStar_Pervasives_Native.None bv.FStar_Syntax_Syntax.sort) binders1 in let args = FStar_Compiler_List.map (fun ai -> - let uu___8 = + let uu___7 = FStar_Syntax_Syntax.bv_to_name ai in - FStar_Syntax_Syntax.as_arg uu___8) bvs in + FStar_Syntax_Syntax.as_arg uu___7) bvs in let body = - let uu___8 = FStar_Syntax_Util.mk_app x args in - let uu___9 = FStar_Syntax_Util.mk_app y args in - mk_stronger b uu___8 uu___9 in + let uu___7 = FStar_Syntax_Util.mk_app x args in + let uu___8 = FStar_Syntax_Util.mk_app y args in + mk_stronger b uu___7 uu___8 in FStar_Compiler_List.fold_right (fun bv -> fun body1 -> mk_forall bv body1) bvs body @@ -1009,40 +1009,40 @@ let (gen_wps_for_free : (binders1, { FStar_Syntax_Syntax.n = - FStar_Syntax_Syntax.Total (b, uu___4); - FStar_Syntax_Syntax.pos = uu___5; - FStar_Syntax_Syntax.vars = uu___6; - FStar_Syntax_Syntax.hash_code = uu___7;_}) + FStar_Syntax_Syntax.Total b; + FStar_Syntax_Syntax.pos = uu___4; + FStar_Syntax_Syntax.vars = uu___5; + FStar_Syntax_Syntax.hash_code = uu___6;_}) -> let bvs = FStar_Compiler_List.mapi (fun i -> - fun uu___8 -> - match uu___8 with + fun uu___7 -> + match uu___7 with | { FStar_Syntax_Syntax.binder_bv = bv; FStar_Syntax_Syntax.binder_qual = q; FStar_Syntax_Syntax.binder_attrs = - uu___9;_} + uu___8;_} -> - let uu___10 = - let uu___11 = + let uu___9 = + let uu___10 = FStar_Compiler_Util.string_of_int i in - Prims.op_Hat "a" uu___11 in - FStar_Syntax_Syntax.gen_bv uu___10 + Prims.op_Hat "a" uu___10 in + FStar_Syntax_Syntax.gen_bv uu___9 FStar_Pervasives_Native.None bv.FStar_Syntax_Syntax.sort) binders1 in let args = FStar_Compiler_List.map (fun ai -> - let uu___8 = + let uu___7 = FStar_Syntax_Syntax.bv_to_name ai in - FStar_Syntax_Syntax.as_arg uu___8) bvs in + FStar_Syntax_Syntax.as_arg uu___7) bvs in let body = - let uu___8 = FStar_Syntax_Util.mk_app x args in - let uu___9 = FStar_Syntax_Util.mk_app y args in - mk_stronger b uu___8 uu___9 in + let uu___7 = FStar_Syntax_Util.mk_app x args in + let uu___8 = FStar_Syntax_Util.mk_app y args in + mk_stronger b uu___7 uu___8 in FStar_Compiler_List.fold_right (fun bv -> fun body1 -> mk_forall bv body1) bvs body @@ -1311,7 +1311,7 @@ let (nm_of_comp : FStar_Syntax_Syntax.comp' FStar_Syntax_Syntax.syntax -> nm) = fun c -> match c.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.Total (t, uu___) -> N t + | FStar_Syntax_Syntax.Total t -> N t | FStar_Syntax_Syntax.Comp c1 when FStar_Compiler_Effect.op_Bar_Greater c1.FStar_Syntax_Syntax.flags (FStar_Compiler_Util.for_some @@ -1427,21 +1427,19 @@ and (star_type' : (match t1.FStar_Syntax_Syntax.n with | FStar_Syntax_Syntax.Tm_arrow (uu___1, - { - FStar_Syntax_Syntax.n = FStar_Syntax_Syntax.GTotal - (hn, uu___2); - FStar_Syntax_Syntax.pos = uu___3; - FStar_Syntax_Syntax.vars = uu___4; - FStar_Syntax_Syntax.hash_code = uu___5;_}) + { FStar_Syntax_Syntax.n = FStar_Syntax_Syntax.GTotal hn; + FStar_Syntax_Syntax.pos = uu___2; + FStar_Syntax_Syntax.vars = uu___3; + FStar_Syntax_Syntax.hash_code = uu___4;_}) -> - let uu___6 = - let uu___7 = - let uu___8 = - let uu___9 = star_type' env1 hn in - FStar_Syntax_Syntax.mk_GTotal uu___9 in - (binders1, uu___8) in - FStar_Syntax_Syntax.Tm_arrow uu___7 in - mk uu___6 + let uu___5 = + let uu___6 = + let uu___7 = + let uu___8 = star_type' env1 hn in + FStar_Syntax_Syntax.mk_GTotal uu___8 in + (binders1, uu___7) in + FStar_Syntax_Syntax.Tm_arrow uu___6 in + mk uu___5 | uu___1 -> let uu___2 = is_monadic_arrow t1.FStar_Syntax_Syntax.n in (match uu___2 with @@ -2530,15 +2528,13 @@ and (infer : match uu___2 with | FStar_Syntax_Syntax.Tm_arrow (binders, - { - FStar_Syntax_Syntax.n = FStar_Syntax_Syntax.Total - (t1, uu___3); - FStar_Syntax_Syntax.pos = uu___4; - FStar_Syntax_Syntax.vars = uu___5; - FStar_Syntax_Syntax.hash_code = uu___6;_}) + { FStar_Syntax_Syntax.n = FStar_Syntax_Syntax.Total t1; + FStar_Syntax_Syntax.pos = uu___3; + FStar_Syntax_Syntax.vars = uu___4; + FStar_Syntax_Syntax.hash_code = uu___5;_}) when is_arrow t1 -> - let uu___7 = flatten t1 in - (match uu___7 with + let uu___6 = flatten t1 in + (match uu___6 with | (binders', comp) -> ((FStar_Compiler_List.op_At binders binders'), comp)) diff --git a/src/ocaml-output/FStar_TypeChecker_Env.ml b/src/ocaml-output/FStar_TypeChecker_Env.ml index ac1961e7c63..5658f72f074 100644 --- a/src/ocaml-output/FStar_TypeChecker_Env.ml +++ b/src/ocaml-output/FStar_TypeChecker_Env.ml @@ -4349,16 +4349,55 @@ let (comp_to_comp_typ : env -> FStar_Syntax_Syntax.comp -> FStar_Syntax_Syntax.comp_typ) = fun env1 -> fun c -> - let c1 = - match c.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.Total (t, FStar_Pervasives_Native.None) -> - let u = env1.universe_of env1 t in - FStar_Syntax_Syntax.mk_Total' t (FStar_Pervasives_Native.Some u) - | FStar_Syntax_Syntax.GTotal (t, FStar_Pervasives_Native.None) -> - let u = env1.universe_of env1 t in - FStar_Syntax_Syntax.mk_GTotal' t (FStar_Pervasives_Native.Some u) - | uu___ -> c in - FStar_Syntax_Util.comp_to_comp_typ c1 + match c.FStar_Syntax_Syntax.n with + | FStar_Syntax_Syntax.Comp ct -> ct + | uu___ -> + let uu___1 = + match c.FStar_Syntax_Syntax.n with + | FStar_Syntax_Syntax.Total t -> + (FStar_Parser_Const.effect_Tot_lid, t) + | FStar_Syntax_Syntax.GTotal t -> + (FStar_Parser_Const.effect_GTot_lid, t) in + (match uu___1 with + | (effect_name, result_typ) -> + let uu___2 = + let uu___3 = env1.universe_of env1 result_typ in [uu___3] in + { + FStar_Syntax_Syntax.comp_univs = uu___2; + FStar_Syntax_Syntax.effect_name = effect_name; + FStar_Syntax_Syntax.result_typ = result_typ; + FStar_Syntax_Syntax.effect_args = []; + FStar_Syntax_Syntax.flags = (FStar_Syntax_Util.comp_flags c) + }) +let (comp_set_flags : + env -> + FStar_Syntax_Syntax.comp -> + FStar_Syntax_Syntax.cflag Prims.list -> FStar_Syntax_Syntax.comp) + = + fun env1 -> + fun c -> + fun f -> + let uu___ = + let uu___1 = + let uu___2 = comp_to_comp_typ env1 c in + { + FStar_Syntax_Syntax.comp_univs = + (uu___2.FStar_Syntax_Syntax.comp_univs); + FStar_Syntax_Syntax.effect_name = + (uu___2.FStar_Syntax_Syntax.effect_name); + FStar_Syntax_Syntax.result_typ = + (uu___2.FStar_Syntax_Syntax.result_typ); + FStar_Syntax_Syntax.effect_args = + (uu___2.FStar_Syntax_Syntax.effect_args); + FStar_Syntax_Syntax.flags = f + } in + FStar_Syntax_Syntax.Comp uu___1 in + { + FStar_Syntax_Syntax.n = uu___; + FStar_Syntax_Syntax.pos = (c.FStar_Syntax_Syntax.pos); + FStar_Syntax_Syntax.vars = (c.FStar_Syntax_Syntax.vars); + FStar_Syntax_Syntax.hash_code = (c.FStar_Syntax_Syntax.hash_code) + } let rec (unfold_effect_abbrev : env -> FStar_Syntax_Syntax.comp -> FStar_Syntax_Syntax.comp_typ) = fun env1 -> diff --git a/src/ocaml-output/FStar_TypeChecker_Err.ml b/src/ocaml-output/FStar_TypeChecker_Err.ml index b96871c291c..6398c097893 100644 --- a/src/ocaml-output/FStar_TypeChecker_Err.ml +++ b/src/ocaml-output/FStar_TypeChecker_Err.ml @@ -433,8 +433,8 @@ let (name_and_result : = fun c -> match c.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.Total (t, uu___) -> ("Tot", t) - | FStar_Syntax_Syntax.GTotal (t, uu___) -> ("GTot", t) + | FStar_Syntax_Syntax.Total t -> ("Tot", t) + | FStar_Syntax_Syntax.GTotal t -> ("GTot", t) | FStar_Syntax_Syntax.Comp ct -> let uu___ = FStar_Syntax_Print.lid_to_string ct.FStar_Syntax_Syntax.effect_name in diff --git a/src/ocaml-output/FStar_TypeChecker_NBE.ml b/src/ocaml-output/FStar_TypeChecker_NBE.ml index af3e258bd7b..6df3a5b96ba 100644 --- a/src/ocaml-output/FStar_TypeChecker_NBE.ml +++ b/src/ocaml-output/FStar_TypeChecker_NBE.ml @@ -1321,17 +1321,15 @@ and (translate_comp : fun bs -> fun c -> match c.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.Total (typ, u) -> + | FStar_Syntax_Syntax.Total typ -> let uu___ = let uu___1 = translate cfg bs typ in - let uu___2 = fmap_opt (translate_univ cfg bs) u in - (uu___1, uu___2) in + (uu___1, FStar_Pervasives_Native.None) in FStar_TypeChecker_NBETerm.Tot uu___ - | FStar_Syntax_Syntax.GTotal (typ, u) -> + | FStar_Syntax_Syntax.GTotal typ -> let uu___ = let uu___1 = translate cfg bs typ in - let uu___2 = fmap_opt (translate_univ cfg bs) u in - (uu___1, uu___2) in + (uu___1, FStar_Pervasives_Native.None) in FStar_TypeChecker_NBETerm.GTot uu___ | FStar_Syntax_Syntax.Comp ctyp -> let uu___ = translate_comp_typ cfg bs ctyp in @@ -2038,12 +2036,10 @@ and (readback_comp : fun c -> let c' = match c with - | FStar_TypeChecker_NBETerm.Tot (typ, u) -> - let uu___ = let uu___1 = readback cfg typ in (uu___1, u) in - FStar_Syntax_Syntax.Total uu___ - | FStar_TypeChecker_NBETerm.GTot (typ, u) -> - let uu___ = let uu___1 = readback cfg typ in (uu___1, u) in - FStar_Syntax_Syntax.GTotal uu___ + | FStar_TypeChecker_NBETerm.Tot (typ, _u) -> + let uu___ = readback cfg typ in FStar_Syntax_Syntax.Total uu___ + | FStar_TypeChecker_NBETerm.GTot (typ, _u) -> + let uu___ = readback cfg typ in FStar_Syntax_Syntax.GTotal uu___ | FStar_TypeChecker_NBETerm.Comp ctyp -> let uu___ = readback_comp_typ cfg ctyp in FStar_Syntax_Syntax.Comp uu___ in diff --git a/src/ocaml-output/FStar_TypeChecker_Normalize.ml b/src/ocaml-output/FStar_TypeChecker_Normalize.ml index 0d816b91642..b2ad925d247 100644 --- a/src/ocaml-output/FStar_TypeChecker_Normalize.ml +++ b/src/ocaml-output/FStar_TypeChecker_Normalize.ml @@ -1144,16 +1144,12 @@ and (close_comp : -> c | uu___ -> (match c.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.Total (t, uopt) -> + | FStar_Syntax_Syntax.Total t -> let uu___1 = inline_closure_env cfg env1 [] t in - let uu___2 = - FStar_Compiler_Option.map (norm_universe cfg env1) uopt in - FStar_Syntax_Syntax.mk_Total' uu___1 uu___2 - | FStar_Syntax_Syntax.GTotal (t, uopt) -> + FStar_Syntax_Syntax.mk_Total uu___1 + | FStar_Syntax_Syntax.GTotal t -> let uu___1 = inline_closure_env cfg env1 [] t in - let uu___2 = - FStar_Compiler_Option.map (norm_universe cfg env1) uopt in - FStar_Syntax_Syntax.mk_GTotal' uu___1 uu___2 + FStar_Syntax_Syntax.mk_GTotal uu___1 | FStar_Syntax_Syntax.Comp c1 -> let rt = inline_closure_env cfg env1 [] @@ -1848,8 +1844,8 @@ let rec (maybe_weakly_reduced : fun tm -> let aux_comp c = match c.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.GTotal (t, uu___) -> maybe_weakly_reduced t - | FStar_Syntax_Syntax.Total (t, uu___) -> maybe_weakly_reduced t + | FStar_Syntax_Syntax.GTotal t -> maybe_weakly_reduced t + | FStar_Syntax_Syntax.Total t -> maybe_weakly_reduced t | FStar_Syntax_Syntax.Comp ct -> (maybe_weakly_reduced ct.FStar_Syntax_Syntax.result_typ) || (FStar_Compiler_Util.for_some @@ -5289,17 +5285,9 @@ and (norm_comp : FStar_Compiler_Util.print2 ">>> %s\nNormComp with with %s env elements\n" uu___2 uu___3); (match comp.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.Total (t, uopt) -> + | FStar_Syntax_Syntax.Total t -> let t1 = norm cfg env1 [] t in - let uopt1 = - match uopt with - | FStar_Pervasives_Native.Some u -> - let uu___1 = norm_universe cfg env1 u in - FStar_Compiler_Effect.op_Less_Bar - (fun uu___2 -> FStar_Pervasives_Native.Some uu___2) - uu___1 - | FStar_Pervasives_Native.None -> FStar_Pervasives_Native.None in - let uu___1 = FStar_Syntax_Syntax.mk_Total' t1 uopt1 in + let uu___1 = FStar_Syntax_Syntax.mk_Total t1 in { FStar_Syntax_Syntax.n = (uu___1.FStar_Syntax_Syntax.n); FStar_Syntax_Syntax.pos = (comp.FStar_Syntax_Syntax.pos); @@ -5307,17 +5295,9 @@ and (norm_comp : FStar_Syntax_Syntax.hash_code = (uu___1.FStar_Syntax_Syntax.hash_code) } - | FStar_Syntax_Syntax.GTotal (t, uopt) -> + | FStar_Syntax_Syntax.GTotal t -> let t1 = norm cfg env1 [] t in - let uopt1 = - match uopt with - | FStar_Pervasives_Native.Some u -> - let uu___1 = norm_universe cfg env1 u in - FStar_Compiler_Effect.op_Less_Bar - (fun uu___2 -> FStar_Pervasives_Native.Some uu___2) - uu___1 - | FStar_Pervasives_Native.None -> FStar_Pervasives_Native.None in - let uu___1 = FStar_Syntax_Syntax.mk_GTotal' t1 uopt1 in + let uu___1 = FStar_Syntax_Syntax.mk_GTotal t1 in { FStar_Syntax_Syntax.n = (uu___1.FStar_Syntax_Syntax.n); FStar_Syntax_Syntax.pos = (comp.FStar_Syntax_Syntax.pos); @@ -7804,12 +7784,12 @@ let (ghost_to_pure_aux : fun c -> match c.FStar_Syntax_Syntax.n with | FStar_Syntax_Syntax.Total uu___ -> c - | FStar_Syntax_Syntax.GTotal (t, uopt) -> + | FStar_Syntax_Syntax.GTotal t -> let uu___ = maybe_promote_t env1 non_informative_only t in if uu___ then { - FStar_Syntax_Syntax.n = (FStar_Syntax_Syntax.Total (t, uopt)); + FStar_Syntax_Syntax.n = (FStar_Syntax_Syntax.Total t); FStar_Syntax_Syntax.pos = (c.FStar_Syntax_Syntax.pos); FStar_Syntax_Syntax.vars = (c.FStar_Syntax_Syntax.vars); FStar_Syntax_Syntax.hash_code = diff --git a/src/ocaml-output/FStar_TypeChecker_Rel.ml b/src/ocaml-output/FStar_TypeChecker_Rel.ml index 03574b2cdc7..3cf68640c99 100644 --- a/src/ocaml-output/FStar_TypeChecker_Rel.ml +++ b/src/ocaml-output/FStar_TypeChecker_Rel.ml @@ -2330,11 +2330,12 @@ let (gamma_until : (match uu___3 with | FStar_Pervasives_Native.None -> [] | FStar_Pervasives_Native.Some (uu___4, bx, rest) -> bx :: rest) -let (restrict_ctx : - FStar_TypeChecker_Env.env -> - FStar_Syntax_Syntax.ctx_uvar -> - FStar_Syntax_Syntax.binders -> - FStar_Syntax_Syntax.ctx_uvar -> worklist -> worklist) +let restrict_ctx : + 'uuuuu . + 'uuuuu -> + FStar_Syntax_Syntax.ctx_uvar -> + FStar_Syntax_Syntax.binders -> + FStar_Syntax_Syntax.ctx_uvar -> worklist -> worklist = fun env -> fun tgt -> @@ -2412,14 +2413,8 @@ let (restrict_ctx : (let uu___3 = let t = FStar_Syntax_Util.ctx_uvar_typ src in let uu___4 = - let uu___5 = - let uu___6 = - FStar_Compiler_Effect.op_Bar_Greater t - (env.FStar_TypeChecker_Env.universe_of env) in - FStar_Compiler_Effect.op_Bar_Greater uu___6 - (fun uu___7 -> FStar_Pervasives_Native.Some uu___7) in - FStar_Compiler_Effect.op_Bar_Greater uu___5 - (FStar_Syntax_Syntax.mk_Total' t) in + FStar_Compiler_Effect.op_Bar_Greater t + FStar_Syntax_Syntax.mk_Total in FStar_Compiler_Effect.op_Bar_Greater uu___4 (FStar_Syntax_Util.arrow bs1) in aux uu___3 @@ -2433,11 +2428,12 @@ let (restrict_ctx : FStar_Syntax_Syntax.as_arg) in FStar_Syntax_Syntax.mk_Tm_app src' uu___4 src.FStar_Syntax_Syntax.ctx_uvar_range)) -let (restrict_all_uvars : - FStar_TypeChecker_Env.env -> - FStar_Syntax_Syntax.ctx_uvar -> - FStar_Syntax_Syntax.binders -> - FStar_Syntax_Syntax.ctx_uvar Prims.list -> worklist -> worklist) +let restrict_all_uvars : + 'uuuuu . + 'uuuuu -> + FStar_Syntax_Syntax.ctx_uvar -> + FStar_Syntax_Syntax.binders -> + FStar_Syntax_Syntax.ctx_uvar Prims.list -> worklist -> worklist = fun env -> fun tgt -> @@ -4277,7 +4273,8 @@ let (apply_substitutive_indexed_subcomp : subst4) in FStar_Compiler_Effect.op_Bar_Greater uu___5 - FStar_Syntax_Util.comp_to_comp_typ in + (FStar_TypeChecker_Env.comp_to_comp_typ + env) in let fml = let uu___5 = let uu___6 = @@ -4479,7 +4476,8 @@ let (apply_ad_hoc_indexed_subcomp : (FStar_Syntax_Subst.subst_comp substs) in FStar_Compiler_Effect.op_Bar_Greater uu___3 - FStar_Syntax_Util.comp_to_comp_typ in + (FStar_TypeChecker_Env.comp_to_comp_typ + env) in let uu___3 = let g_sort_is = let uu___4 = @@ -5698,34 +5696,23 @@ and (imitate_arrow : match uu___ with | Flex (uu___1, u_lhs, uu___2) -> let imitate_comp bs bs_terms c wl1 = - let imitate_tot_or_gtot t uopt f wl2 = - let uu___3 = - match uopt with - | FStar_Pervasives_Native.None -> - FStar_Syntax_Util.type_u () - | FStar_Pervasives_Native.Some univ -> - let uu___4 = - FStar_Syntax_Syntax.mk - (FStar_Syntax_Syntax.Tm_type univ) - t.FStar_Syntax_Syntax.pos in - (uu___4, univ) in + let imitate_tot_or_gtot t f wl2 = + let uu___3 = FStar_Syntax_Util.type_u () in match uu___3 with - | (k, univ) -> - let uu___4 = + | (k, uu___4) -> + let uu___5 = copy_uvar u_lhs (FStar_Compiler_List.op_At bs_lhs bs) k wl2 in - (match uu___4 with - | (uu___5, u, wl3) -> - let uu___6 = - f u (FStar_Pervasives_Native.Some univ) in - (uu___6, wl3)) in + (match uu___5 with + | (uu___6, u, wl3) -> + let uu___7 = f u in (uu___7, wl3)) in match c.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.Total (t, uopt) -> - imitate_tot_or_gtot t uopt - FStar_Syntax_Syntax.mk_Total' wl1 - | FStar_Syntax_Syntax.GTotal (t, uopt) -> - imitate_tot_or_gtot t uopt - FStar_Syntax_Syntax.mk_GTotal' wl1 + | FStar_Syntax_Syntax.Total t -> + imitate_tot_or_gtot t + FStar_Syntax_Syntax.mk_Total wl1 + | FStar_Syntax_Syntax.GTotal t -> + imitate_tot_or_gtot t + FStar_Syntax_Syntax.mk_GTotal wl1 | FStar_Syntax_Syntax.Comp ct -> let uu___3 = let uu___4 = @@ -6472,21 +6459,9 @@ and (solve_t_flex_rigid_eq : t_last_arg in let uu___10 = let uu___11 = - let uu___12 = - let uu___13 = - FStar_Compiler_Effect.op_Bar_Greater - t_res_lhs - (env1.FStar_TypeChecker_Env.universe_of - env1) in - FStar_Compiler_Effect.op_Bar_Greater - uu___13 - (fun uu___14 -> - FStar_Pervasives_Native.Some - uu___14) in FStar_Compiler_Effect.op_Bar_Greater - uu___12 - (FStar_Syntax_Syntax.mk_Total' - t_res_lhs) in + t_res_lhs + FStar_Syntax_Syntax.mk_Total in FStar_Compiler_Effect.op_Bar_Greater uu___11 (FStar_Syntax_Util.arrow [b]) in @@ -12012,7 +11987,8 @@ and (solve_c : (fun uu___5 -> match uu___5 with | (c, g) -> - let uu___6 = FStar_Syntax_Util.comp_to_comp_typ c in + let uu___6 = + FStar_TypeChecker_Env.comp_to_comp_typ env c in (uu___6, g)) in let uu___4 = let uu___5 = @@ -12342,7 +12318,7 @@ and (solve_c : uu___7 uu___8 in (FStar_Errors.Fatal_UnexpectedEffect, uu___6) in FStar_Errors.raise_error uu___5 r - else FStar_Syntax_Util.comp_to_comp_typ c) in + else FStar_TypeChecker_Env.comp_to_comp_typ env c) in let uu___1 = should_fail_since_repr_subcomp_not_allowed wl.repr_subcomp_allowed c11.FStar_Syntax_Syntax.effect_name @@ -12629,48 +12605,47 @@ and (solve_c : (match ((c11.FStar_Syntax_Syntax.n), (c21.FStar_Syntax_Syntax.n)) with - | (FStar_Syntax_Syntax.GTotal (t1, uu___4), - FStar_Syntax_Syntax.Total (t2, uu___5)) when - FStar_TypeChecker_Env.non_informative env t2 -> - let uu___6 = + | (FStar_Syntax_Syntax.GTotal t1, FStar_Syntax_Syntax.Total + t2) when FStar_TypeChecker_Env.non_informative env t2 -> + let uu___4 = problem_using_guard orig t1 problem.FStar_TypeChecker_Common.relation t2 FStar_Pervasives_Native.None "result type" in - solve_t env uu___6 wl + solve_t env uu___4 wl | (FStar_Syntax_Syntax.GTotal uu___4, FStar_Syntax_Syntax.Total uu___5) -> let uu___6 = FStar_Thunk.mkv "incompatible monad ordering: GTot - let uu___6 = + | (FStar_Syntax_Syntax.Total t1, FStar_Syntax_Syntax.Total + t2) -> + let uu___4 = problem_using_guard orig t1 problem.FStar_TypeChecker_Common.relation t2 FStar_Pervasives_Native.None "result type" in - solve_t env uu___6 wl - | (FStar_Syntax_Syntax.GTotal (t1, uu___4), - FStar_Syntax_Syntax.GTotal (t2, uu___5)) -> - let uu___6 = + solve_t env uu___4 wl + | (FStar_Syntax_Syntax.GTotal t1, FStar_Syntax_Syntax.GTotal + t2) -> + let uu___4 = problem_using_guard orig t1 problem.FStar_TypeChecker_Common.relation t2 FStar_Pervasives_Native.None "result type" in - solve_t env uu___6 wl - | (FStar_Syntax_Syntax.Total (t1, uu___4), - FStar_Syntax_Syntax.GTotal (t2, uu___5)) when + solve_t env uu___4 wl + | (FStar_Syntax_Syntax.Total t1, FStar_Syntax_Syntax.GTotal + t2) when problem.FStar_TypeChecker_Common.relation = FStar_TypeChecker_Common.SUB -> - let uu___6 = + let uu___4 = problem_using_guard orig t1 problem.FStar_TypeChecker_Common.relation t2 FStar_Pervasives_Native.None "result type" in - solve_t env uu___6 wl - | (FStar_Syntax_Syntax.Total (t1, uu___4), - FStar_Syntax_Syntax.GTotal (t2, uu___5)) -> - let uu___6 = FStar_Thunk.mkv "GTot =/= Tot" in - giveup env uu___6 orig + solve_t env uu___4 wl + | (FStar_Syntax_Syntax.Total t1, FStar_Syntax_Syntax.GTotal + t2) -> + let uu___4 = FStar_Thunk.mkv "GTot =/= Tot" in + giveup env uu___4 orig | (FStar_Syntax_Syntax.GTotal uu___4, FStar_Syntax_Syntax.Comp uu___5) -> let uu___6 = diff --git a/src/ocaml-output/FStar_TypeChecker_TcEffect.ml b/src/ocaml-output/FStar_TypeChecker_TcEffect.ml index 4547516d027..4feb35c8990 100644 --- a/src/ocaml-output/FStar_TypeChecker_TcEffect.ml +++ b/src/ocaml-output/FStar_TypeChecker_TcEffect.ml @@ -1563,14 +1563,8 @@ let (validate_indexed_effect_bind_shape : let uu___6 = let uu___7 = let uu___8 = - let uu___9 = - let uu___10 = - FStar_TypeChecker_Env.new_u_univ - () in - FStar_Pervasives_Native.Some - uu___10 in - FStar_Syntax_Syntax.mk_Total' - repr uu___9 in + FStar_Syntax_Syntax.mk_Total + repr in FStar_Syntax_Util.arrow [x_a] uu___8 in FStar_Syntax_Syntax.gen_bv @@ -4065,14 +4059,8 @@ let (tc_layered_eff_decl : (fun uu___11 -> match uu___11 with | (t, u1) -> - let uu___12 = - let uu___13 = - FStar_TypeChecker_Env.new_u_univ - () in - FStar_Pervasives_Native.Some - uu___13 in - FStar_Syntax_Syntax.mk_Total' - t uu___12) in + FStar_Syntax_Syntax.mk_Total + t) in FStar_Syntax_Util.arrow bs uu___9 in let g = FStar_TypeChecker_Rel.teq env ty k in @@ -4217,10 +4205,8 @@ let (tc_layered_eff_decl : | (repr1, g) -> let k = let uu___11 = - FStar_Syntax_Syntax.mk_Total' - repr1 - (FStar_Pervasives_Native.Some - u_a) in + FStar_Syntax_Syntax.mk_Total + repr1 in FStar_Syntax_Util.arrow bs uu___11 in let g_eq = @@ -5270,8 +5256,8 @@ let (tc_layered_eff_decl : | FStar_Syntax_Syntax.Tm_arrow (bs, c) -> let ct = - FStar_Syntax_Util.comp_to_comp_typ - c in + FStar_TypeChecker_Env.comp_to_comp_typ + env1 c in let uu___17 = FStar_Ident.lid_equals ct.FStar_Syntax_Syntax.effect_name @@ -5301,14 +5287,8 @@ let (tc_layered_eff_decl : FStar_Syntax_Syntax.mk_Tm_app repr1 uu___18 r in let c1 = - let uu___18 = - let uu___19 = - FStar_TypeChecker_Env.new_u_univ - () in - FStar_Pervasives_Native.Some - uu___19 in - FStar_Syntax_Syntax.mk_Total' - repr2 uu___18 in + FStar_Syntax_Syntax.mk_Total + repr2 in FStar_Syntax_Util.arrow bs c1 else @@ -5576,22 +5556,8 @@ let (tc_layered_eff_decl : = let uu___29 = - let uu___30 - = - let uu___31 - = - FStar_TypeChecker_Env.new_u_univ - () in - FStar_Compiler_Effect.op_Bar_Greater - uu___31 - (fun - uu___32 - -> - FStar_Pervasives_Native.Some - uu___32) in - FStar_Syntax_Syntax.mk_Total' - repr1 - uu___30 in + FStar_Syntax_Syntax.mk_Total + repr1 in FStar_Syntax_Util.arrow bs1 uu___29 in @@ -6272,13 +6238,8 @@ let (tc_non_layered_eff_decl : let uu___8 = let tmp_t = let uu___9 = - let uu___10 = - FStar_Compiler_Effect.op_Bar_Greater - FStar_Syntax_Syntax.U_zero - (fun uu___11 -> - FStar_Pervasives_Native.Some uu___11) in - FStar_Syntax_Syntax.mk_Total' - FStar_Syntax_Syntax.t_unit uu___10 in + FStar_Syntax_Syntax.mk_Total + FStar_Syntax_Syntax.t_unit in FStar_Syntax_Util.arrow bs1 uu___9 in let uu___9 = FStar_TypeChecker_Generalize.generalize_universes @@ -7681,8 +7642,8 @@ let (tc_non_layered_eff_decl : | FStar_Syntax_Syntax.Tm_arrow (bs1, c) -> let c1 = - FStar_Syntax_Util.comp_to_comp_typ - c in + FStar_TypeChecker_Env.comp_to_comp_typ + env1 c in let uu___23 = FStar_Ident.lid_equals c1.FStar_Syntax_Syntax.effect_name diff --git a/src/ocaml-output/FStar_TypeChecker_TcTerm.ml b/src/ocaml-output/FStar_TypeChecker_TcTerm.ml index 59015dbe595..bde8b546602 100644 --- a/src/ocaml-output/FStar_TypeChecker_TcTerm.ml +++ b/src/ocaml-output/FStar_TypeChecker_TcTerm.ml @@ -2264,8 +2264,7 @@ and (tc_maybe_toplevel_term : let uu___12 = let uu___13 = let uu___14 = - FStar_Syntax_Syntax.mk_Total' repr - (FStar_Pervasives_Native.Some u_c) in + FStar_Syntax_Syntax.mk_Total repr in FStar_Pervasives.Inr uu___14 in (uu___13, FStar_Pervasives_Native.None, use_eq) in @@ -3769,10 +3768,8 @@ and (tc_match : then let e = FStar_Syntax_Util.exp_true_bool in let c = - FStar_Syntax_Syntax.mk_GTotal' - FStar_Syntax_Util.t_bool - (FStar_Pervasives_Native.Some - FStar_Syntax_Syntax.U_zero) in + FStar_Syntax_Syntax.mk_GTotal + FStar_Syntax_Util.t_bool in let uu___6 = FStar_TypeChecker_Common.lcomp_of_comp c in FStar_TypeChecker_Util.bind @@ -4691,28 +4688,24 @@ and (tc_comp : fun c -> let c0 = c in match c.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.Total (t, uu___) -> - let uu___1 = FStar_Syntax_Util.type_u () in - (match uu___1 with + | FStar_Syntax_Syntax.Total t -> + let uu___ = FStar_Syntax_Util.type_u () in + (match uu___ with | (k, u) -> - let uu___2 = tc_check_tot_or_gtot_term env t k "" in - (match uu___2 with - | (t1, uu___3, g) -> - let uu___4 = - FStar_Syntax_Syntax.mk_Total' t1 - (FStar_Pervasives_Native.Some u) in - (uu___4, u, g))) - | FStar_Syntax_Syntax.GTotal (t, uu___) -> - let uu___1 = FStar_Syntax_Util.type_u () in - (match uu___1 with + let uu___1 = tc_check_tot_or_gtot_term env t k "" in + (match uu___1 with + | (t1, uu___2, g) -> + let uu___3 = FStar_Syntax_Syntax.mk_Total t1 in + (uu___3, u, g))) + | FStar_Syntax_Syntax.GTotal t -> + let uu___ = FStar_Syntax_Util.type_u () in + (match uu___ with | (k, u) -> - let uu___2 = tc_check_tot_or_gtot_term env t k "" in - (match uu___2 with - | (t1, uu___3, g) -> - let uu___4 = - FStar_Syntax_Syntax.mk_GTotal' t1 - (FStar_Pervasives_Native.Some u) in - (uu___4, u, g))) + let uu___1 = tc_check_tot_or_gtot_term env t k "" in + (match uu___1 with + | (t1, uu___2, g) -> + let uu___3 = FStar_Syntax_Syntax.mk_GTotal t1 in + (uu___3, u, g))) | FStar_Syntax_Syntax.Comp c1 -> let head = FStar_Syntax_Syntax.fvar c1.FStar_Syntax_Syntax.effect_name @@ -10086,10 +10079,8 @@ and (check_top_level_let : "Let binding AFTER tcnorm: %s\n" uu___6 else ()); (let cres = - FStar_Syntax_Syntax.mk_Total' - FStar_Syntax_Syntax.t_unit - (FStar_Pervasives_Native.Some - FStar_Syntax_Syntax.U_zero) in + FStar_Syntax_Syntax.mk_Total + FStar_Syntax_Syntax.t_unit in let lb1 = FStar_Syntax_Util.close_univs_and_mk_letbinding FStar_Pervasives_Native.None @@ -12664,15 +12655,14 @@ let rec (typeof_tot_or_gtot_term_fastpath : let uu___1 = FStar_Ident.lid_equals eff FStar_Parser_Const.effect_Tot_lid in if uu___1 - then FStar_Pervasives_Native.Some FStar_Syntax_Syntax.mk_Total' + then FStar_Pervasives_Native.Some FStar_Syntax_Syntax.mk_Total else (let uu___3 = FStar_Ident.lid_equals eff FStar_Parser_Const.effect_GTot_lid in if uu___3 then - FStar_Pervasives_Native.Some - FStar_Syntax_Syntax.mk_GTotal' + FStar_Pervasives_Native.Some FStar_Syntax_Syntax.mk_GTotal else FStar_Pervasives_Native.None) in FStar_Compiler_Util.bind_opt mk_comp (fun f -> @@ -12700,8 +12690,7 @@ let rec (typeof_tot_or_gtot_term_fastpath : FStar_TypeChecker_Env.push_binders env bs1 in universe_of uu___2 tbody3 in let uu___2 = - let uu___3 = - f tbody3 (FStar_Pervasives_Native.Some u) in + let uu___3 = f tbody3 in FStar_Syntax_Util.arrow bs1 uu___3 in FStar_Pervasives_Native.Some uu___2)) | FStar_Syntax_Syntax.Tm_abs uu___ -> FStar_Pervasives_Native.None diff --git a/src/ocaml-output/FStar_TypeChecker_Util.ml b/src/ocaml-output/FStar_TypeChecker_Util.ml index f538d8cb143..a6662b69c72 100644 --- a/src/ocaml-output/FStar_TypeChecker_Util.ml +++ b/src/ocaml-output/FStar_TypeChecker_Util.ml @@ -203,8 +203,8 @@ let (extract_let_rec_annotation : (pfx, FStar_Syntax_Syntax.DECREASES d, sfx) -> let c1 = - FStar_Syntax_Util.comp_set_flags c - (FStar_Compiler_List.op_At pfx sfx) in + FStar_TypeChecker_Env.comp_set_flags env1 + c (FStar_Compiler_List.op_At pfx sfx) in let uu___8 = FStar_Syntax_Util.arrow bs c1 in (uu___8, tarr, true) | uu___8 -> (tarr, tarr, true))) @@ -232,12 +232,13 @@ let (extract_let_rec_annotation : FStar_Syntax_Subst.subst_decreasing_order s d in let c1 = - FStar_Syntax_Util.comp_set_flags c - flags in + FStar_TypeChecker_Env.comp_set_flags + env1 c flags in let tarr1 = FStar_Syntax_Util.arrow bs c1 in let c'1 = - FStar_Syntax_Util.comp_set_flags c' + FStar_TypeChecker_Env.comp_set_flags + env1 c' ((FStar_Syntax_Syntax.DECREASES d') :: flags') in let tannot = @@ -581,8 +582,8 @@ let (comp_univ_opt : = fun c -> match c.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.Total (uu___, uopt) -> uopt - | FStar_Syntax_Syntax.GTotal (uu___, uopt) -> uopt + | FStar_Syntax_Syntax.Total uu___ -> FStar_Pervasives_Native.None + | FStar_Syntax_Syntax.GTotal uu___ -> FStar_Pervasives_Native.None | FStar_Syntax_Syntax.Comp c1 -> (match c1.FStar_Syntax_Syntax.comp_univs with | [] -> FStar_Pervasives_Native.None @@ -664,12 +665,14 @@ let (effect_args_from_repr : | FStar_Syntax_Syntax.Tm_arrow (uu___1, c) -> let uu___2 = FStar_Compiler_Effect.op_Bar_Greater c - FStar_Syntax_Util.comp_to_comp_typ in + FStar_Syntax_Util.comp_eff_name_res_and_args in FStar_Compiler_Effect.op_Bar_Greater uu___2 - (fun ct -> - FStar_Compiler_Effect.op_Bar_Greater - ct.FStar_Syntax_Syntax.effect_args - (FStar_Compiler_List.map FStar_Pervasives_Native.fst)) + (fun uu___3 -> + match uu___3 with + | (uu___4, uu___5, args) -> + FStar_Compiler_Effect.op_Bar_Greater args + (FStar_Compiler_List.map + FStar_Pervasives_Native.fst)) | uu___1 -> err ()) let (mk_wp_return : FStar_TypeChecker_Env.env -> @@ -696,10 +699,7 @@ let (mk_wp_return : else (let uu___2 = FStar_Syntax_Util.is_unit a in if uu___2 - then - FStar_Syntax_Syntax.mk_Total' a - (FStar_Pervasives_Native.Some - FStar_Syntax_Syntax.U_zero) + then FStar_Syntax_Syntax.mk_Total a else (let wp = let uu___4 = @@ -979,9 +979,7 @@ let (lax_mk_tot_or_comp_l : let uu___ = FStar_Ident.lid_equals mname FStar_Parser_Const.effect_Tot_lid in if uu___ - then - FStar_Syntax_Syntax.mk_Total' result - (FStar_Pervasives_Native.Some u_result) + then FStar_Syntax_Syntax.mk_Total result else mk_comp_l mname u_result result FStar_Syntax_Syntax.tun flags let (is_function : FStar_Syntax_Syntax.term -> Prims.bool) = fun t -> @@ -1981,16 +1979,15 @@ let (mk_indexed_return : let uu___5 = FStar_Syntax_Subst.open_comp bs c in (match uu___5 with | (a_b::x_b::bs1, c1) -> - let uu___6 = - FStar_Syntax_Util.comp_to_comp_typ c1 in - (a_b, x_b, bs1, uu___6)) + (a_b, x_b, bs1, + (FStar_Syntax_Util.comp_result c1))) | uu___5 -> let uu___6 = return_t_shape_error "Either not an arrow or not enough binders" in FStar_Errors.raise_error uu___6 r in (match uu___3 with - | (a_b, x_b, rest_bs, return_ct) -> + | (a_b, x_b, rest_bs, return_typ) -> let uu___4 = FStar_TypeChecker_Env.uvars_for_binders env rest_bs [FStar_Syntax_Syntax.NT @@ -2027,8 +2024,7 @@ let (mk_indexed_return : let is = let uu___5 = let uu___6 = - FStar_Syntax_Subst.compress - return_ct.FStar_Syntax_Syntax.result_typ in + FStar_Syntax_Subst.compress return_typ in let uu___7 = FStar_Syntax_Util.is_layered ed in effect_args_from_repr uu___6 uu___7 r in FStar_Compiler_Effect.op_Bar_Greater uu___5 @@ -2246,7 +2242,8 @@ let (mk_indexed_bind : subst) in FStar_Compiler_Effect.op_Bar_Greater uu___8 - FStar_Syntax_Util.comp_to_comp_typ in + (FStar_TypeChecker_Env.comp_to_comp_typ + env) in let fml = let uu___8 = let uu___9 = @@ -2443,9 +2440,11 @@ let (mk_bind : | (m, c11, c21, g_lift) -> let uu___3 = let uu___4 = - FStar_Syntax_Util.comp_to_comp_typ c11 in + FStar_TypeChecker_Env.comp_to_comp_typ env + c11 in let uu___5 = - FStar_Syntax_Util.comp_to_comp_typ c21 in + FStar_TypeChecker_Env.comp_to_comp_typ env + c21 in (uu___4, uu___5) in (match uu___3 with | (ct11, ct21) -> @@ -3541,7 +3540,8 @@ let (assume_result_eq_pure_term_in_m : Prims.op_Negation uu___5 in if uu___4 then - let retc1 = FStar_Syntax_Util.comp_to_comp_typ retc in + let retc1 = + FStar_TypeChecker_Env.comp_to_comp_typ env retc in let retc2 = { FStar_Syntax_Syntax.comp_univs = @@ -3558,7 +3558,8 @@ let (assume_result_eq_pure_term_in_m : (uu___5, g_c1) else (let uu___6 = - FStar_Syntax_Util.comp_set_flags retc flags in + FStar_TypeChecker_Env.comp_set_flags env retc + flags in (uu___6, g_c1))) else (let c1 = FStar_TypeChecker_Env.unfold_effect_abbrev env c in @@ -3577,7 +3578,7 @@ let (assume_result_eq_pure_term_in_m : | (ret, g_ret) -> let ret1 = let uu___5 = - FStar_Syntax_Util.comp_set_flags ret + FStar_TypeChecker_Env.comp_set_flags env ret [FStar_Syntax_Syntax.PARTIAL_RETURN] in FStar_Compiler_Effect.op_Less_Bar FStar_TypeChecker_Common.lcomp_of_comp uu___5 in @@ -3596,7 +3597,8 @@ let (assume_result_eq_pure_term_in_m : (match uu___5 with | (bind_c, g_bind) -> let uu___6 = - FStar_Syntax_Util.comp_set_flags bind_c flags in + FStar_TypeChecker_Env.comp_set_flags env bind_c + flags in let uu___7 = FStar_TypeChecker_Env.conj_guards [g_c; g_ret; g_bind] in @@ -4557,11 +4559,13 @@ let (bind_cases : let uu___15 = FStar_Compiler_Effect.op_Bar_Greater cthen2 - FStar_Syntax_Util.comp_to_comp_typ in + (FStar_TypeChecker_Env.comp_to_comp_typ + env) in let uu___16 = FStar_Compiler_Effect.op_Bar_Greater celse1 - FStar_Syntax_Util.comp_to_comp_typ in + (FStar_TypeChecker_Env.comp_to_comp_typ + env) in (md1, uu___15, uu___16, g_lift_then, @@ -7598,7 +7602,7 @@ let (lift_tf_layered_effect : "Lifting indexed comp %s to %s {\n" uu___1 uu___2) else (); (let r = FStar_TypeChecker_Env.get_range env in - let ct = FStar_Syntax_Util.comp_to_comp_typ c in + let ct = FStar_TypeChecker_Env.comp_to_comp_typ env c in check_non_informative_type_for_lift env ct.FStar_Syntax_Syntax.effect_name tgt ct.FStar_Syntax_Syntax.result_typ r; @@ -7638,7 +7642,7 @@ let (lift_tf_layered_effect : FStar_Compiler_Effect.op_Bar_Greater lift_c (FStar_Syntax_Subst.subst_comp substs) in FStar_Compiler_Effect.op_Bar_Greater uu___6 - FStar_Syntax_Util.comp_to_comp_typ in + (FStar_TypeChecker_Env.comp_to_comp_typ env) in let is = let uu___6 = FStar_TypeChecker_Env.is_layered_effect env @@ -7859,7 +7863,7 @@ let (get_mlift_for_subeff : } else (let mk_mlift_wp ts env1 c = - let ct = FStar_Syntax_Util.comp_to_comp_typ c in + let ct = FStar_TypeChecker_Env.comp_to_comp_typ env1 c in check_non_informative_type_for_lift env1 ct.FStar_Syntax_Syntax.effect_name sub.FStar_Syntax_Syntax.target ct.FStar_Syntax_Syntax.result_typ diff --git a/src/tactics/FStar.Tactics.Basic.fst b/src/tactics/FStar.Tactics.Basic.fst index 93e6ba491de..f6c94c8c76a 100644 --- a/src/tactics/FStar.Tactics.Basic.fst +++ b/src/tactics/FStar.Tactics.Basic.fst @@ -990,18 +990,18 @@ let t_apply (uopt:bool) (only_match:bool) (tc_resolved_uvars:bool) (tm:term) : t // returns pre and post let lemma_or_sq (c : comp) : option (term * term) = - let ct = U.comp_to_comp_typ_nouniv c in - if lid_equals ct.effect_name PC.effect_Lemma_lid then - let pre, post = match ct.effect_args with + let eff_name, res, args = U.comp_eff_name_res_and_args c in + if lid_equals eff_name PC.effect_Lemma_lid then + let pre, post = match args with | pre::post::_ -> fst pre, fst post | _ -> failwith "apply_lemma: impossible: not a lemma" in // Lemma post is thunked let post = U.mk_app post [S.as_arg U.exp_unit] in Some (pre, post) - else if U.is_pure_effect ct.effect_name - || U.is_ghost_effect ct.effect_name then - map_opt (U.un_squash ct.result_typ) (fun post -> (U.t_true, post)) + else if U.is_pure_effect eff_name + || U.is_ghost_effect eff_name then + map_opt (U.un_squash res) (fun post -> (U.t_true, post)) else None diff --git a/src/tests/FStar.Tests.Util.fst b/src/tests/FStar.Tests.Util.fst index efe1e6ad94d..fd2e7f734c1 100644 --- a/src/tests/FStar.Tests.Util.fst +++ b/src/tests/FStar.Tests.Util.fst @@ -58,7 +58,7 @@ let rec term_eq' t1 t2 = && List.forall2 (fun (a, imp) (b, imp') -> term_eq' a b && U.eq_aqual imp imp'=U.Equal) xs ys in let comp_eq (c:S.comp) (d:S.comp) = match c.n, d.n with - | S.Total (t, _), S.Total (s, _) -> term_eq' t s + | S.Total t, S.Total s -> term_eq' t s | S.Comp ct1, S.Comp ct2 -> I.lid_equals ct1.effect_name ct2.effect_name && term_eq' ct1.result_typ ct2.result_typ From 62ba91d06a9448dc1dc61b8a98291ecc8d3836cf Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Fri, 9 Dec 2022 16:46:33 +0000 Subject: [PATCH 07/18] nits --- examples/param/Param.fst | 2 +- ulib/experimental/Steel.Effect.Common.fsti | 16 +++++++++++----- 2 files changed, 12 insertions(+), 6 deletions(-) diff --git a/examples/param/Param.fst b/examples/param/Param.fst index 97f83a47956..23477ac063e 100644 --- a/examples/param/Param.fst +++ b/examples/param/Param.fst @@ -53,7 +53,7 @@ let rec param' (bvmap : bvmap) (t:term) : Tac term = | Tv_Arrow b c -> let bv, (q, _attrs) = inspect_binder b in begin match inspect_comp c with - | C_Total t2 _ _ -> + | C_Total t2 -> let t1 = (inspect_bv bv).bv_sort in // bv:t1 -> t2 diff --git a/ulib/experimental/Steel.Effect.Common.fsti b/ulib/experimental/Steel.Effect.Common.fsti index c799fe0663b..bb0ad81174d 100644 --- a/ulib/experimental/Steel.Effect.Common.fsti +++ b/ulib/experimental/Steel.Effect.Common.fsti @@ -954,11 +954,17 @@ let rec new_args_for_smt_attrs (env:env) (l:list argv) (ty:typ) : Tac (list argv ) else (arg, aqualv) in begin - match inspect_comp comp with - | C_Total ty2 -> - let tl_argv, tl_terms = new_args_for_smt_attrs env tl ty2 in - new_hd::tl_argv, (if needs_smt then arg::tl_terms else tl_terms) - | _ -> fail "computation type not supported in definition of slprops" + let ty2 = + match inspect_comp comp with + | C_Total ty2 -> ty2 + | C_Eff _ eff_name ty2 _ _ -> + if eff_name = ["Prims"; "Tot"] + then ty2 + else fail "computation type not supported in definition of slprops" + | _ -> fail "computation type not supported in definition of slprops" in + + let tl_argv, tl_terms = new_args_for_smt_attrs env tl ty2 in + new_hd::tl_argv, (if needs_smt then arg::tl_terms else tl_terms) end | [], Tv_FVar fv -> [], [] | _ -> fail "should not happen. Is an slprop partially applied?" From 980df2c2beed6279aa2819597dc0e6abbb197601 Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Mon, 12 Dec 2022 22:26:55 +0530 Subject: [PATCH 08/18] wip --- .common.mk | 5 ++- diff_runlim.py | 94 ++++++++++++++++++++++++++++++++++++++++++++++++++ 2 files changed, 98 insertions(+), 1 deletion(-) create mode 100644 diff_runlim.py diff --git a/.common.mk b/.common.mk index 59d24416159..d03001b53dc 100644 --- a/.common.mk +++ b/.common.mk @@ -33,5 +33,8 @@ endif # Passing MON=1 will create .runlim files through the source tree with # information about the time and space taken by each F* invocation. ifneq ($(MON),) - PREF=runlim -p -o $@.runlim + ifneq ($(MONID),) + MONPREFIX=$(MONID). + endif + PREF=runlim -p -o $@.$(MONPREFIX)runlim endif diff --git a/diff_runlim.py b/diff_runlim.py new file mode 100644 index 00000000000..3a745ca10ee --- /dev/null +++ b/diff_runlim.py @@ -0,0 +1,94 @@ +# a function that takes as input a filename, 3 hashtables, and a line +# it removes the "[runlim] " prefix from the line +# it splits the rest of the line into a list of strings +# if the first element of the list is "real:", it adds the filename and the second string to the first hashtable +# if the first element of the list is "time:", it adds the filename and the second string to the second hashtable +# if the first element of the list is "space:", it adds the filename and the second string to the third hashtable +# it returns the three hashtables +def parse_line(filename, real_times, user_times, space, line): + line = line.replace("[runlim] ", "") + line = line.split() + if line[0] == "real:": + real_times[filename] = line[1] + elif line[0] == "time:": + user_times[filename] = line[1] + elif line[0] == "space:": + space[filename] = line[1] + return real_times, user_times, space + +# a function that takes as input a filename, an identifier, and 3 hashtables +# it opens the file and reads each line +# it strips .identifier".runlim" from the filename +# it calls parse_line on the on the resulting string, the 3 hashtables, and each line +# it returns the 3 hashtables +def parse_file(filename, identifier, real_times, user_times, space): + with open(filename) as f: + for line in f: + filename = filename.replace(".{}.runlim".format(identifier), "") + real_times, user_times, space = parse_line(filename, real_times, user_times, space, line) + return real_times, user_times, space + +# a function that takes as input an identifier +# it creates 3 empty hashtables +# it calls parse_file on all the files in the current directory with suffix .identifier".runlim" +# it returns the 3 hashtables +def parse_all(identifier): + real_times = {} + user_times = {} + space = {} + for filename in os.listdir("."): + if filename.endswith(".{}.runlim".format(identifier)): + real_times, user_times, space = parse_file(filename, identifier, real_times, user_times, space) + return real_times, user_times, space + +# a function that takes as input 2 hashtables +# it creates a third hashtable +# it iterates over the keys of the first hashtable +# it adds the key and the percentage difference of values of the first and second hashtables to the third hashtable +# it returns the third hashtable +def diff_hashtables(hashtable1, hashtable2): + diff = {} + for key in hashtable1: + diff[key] = (float(hashtable1[key]) - float(hashtable2[key])) / float(hashtable2[key]) * 100 + return diff + +# a function that takes as input 2 identifiers +# it calls parse_all on the first identifier +# it calls parse_all on the second identifier +# it calls diff_hashtables on the two resulting hashtables +# it returns the resulting hashtable +def diff(identifier1, identifier2): + real_times1, user_times1, space1 = parse_all(identifier1) + real_times2, user_times2, space2 = parse_all(identifier2) + real_times_diff = diff_hashtables(real_times1, real_times2) + user_times_diff = diff_hashtables(user_times1, user_times2) + space_diff = diff_hashtables(space1, space2) + return real_times_diff, user_times_diff, space_diff + +# a function that takes as input two identifiers +# it calls diff on the two identifiers +# it plots 3 scatter plots for the the resulting hashtables and saves them to files +def plot_diff(identifier1, identifier2): + real_times_diff, user_times_diff, space_diff = diff(identifier1, identifier2) + plt.scatter(real_times_diff.keys(), real_times_diff.values()) + plt.xlabel("filename") + plt.ylabel("real time difference (%)") + plt.savefig("real_times_diff_{}_{}.png".format(identifier1, identifier2)) + plt.clf() + plt.scatter(user_times_diff.keys(), user_times_diff.values()) + plt.xlabel("filename") + plt.ylabel("user time difference (%)") + plt.savefig("user_times_diff_{}_{}.png".format(identifier1, identifier2)) + plt.clf() + plt.scatter(space_diff.keys(), space_diff.values()) + plt.xlabel("filename") + plt.ylabel("space difference (%)") + plt.savefig("space_diff_{}_{}.png".format(identifier1, identifier2)) + plt.clf() + +# a function that reads two identifiers from the command line +# it calls plot_diff on the two identifiers +if __name__ == "__main__": + identifier1 = sys.argv[1] + identifier2 = sys.argv[2] + plot_diff(identifier1, identifier2) From 3e226816a260c864f9e2dbf52a552385a7f21398 Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Tue, 13 Dec 2022 15:40:17 +0530 Subject: [PATCH 09/18] a script to compare two runlim runs --- diff_runlim.py | 197 +++++++++++++++++++++++++++++++++---------------- 1 file changed, 133 insertions(+), 64 deletions(-) diff --git a/diff_runlim.py b/diff_runlim.py index 3a745ca10ee..5c84cae9ee9 100644 --- a/diff_runlim.py +++ b/diff_runlim.py @@ -1,94 +1,163 @@ -# a function that takes as input a filename, 3 hashtables, and a line +# a function that takes as input a filename, 2 hashtables, and a line # it removes the "[runlim] " prefix from the line # it splits the rest of the line into a list of strings -# if the first element of the list is "real:", it adds the filename and the second string to the first hashtable -# if the first element of the list is "time:", it adds the filename and the second string to the second hashtable -# if the first element of the list is "space:", it adds the filename and the second string to the third hashtable -# it returns the three hashtables -def parse_line(filename, real_times, user_times, space, line): +# if the first element of the list is "time:", it adds the filename and the second string to the first hashtable +# if the first element of the list is "space:", it adds the filename and the second string to the second hashtable +# # it returns the two hashtables +def parse_line(filename, time, space, line): line = line.replace("[runlim] ", "") line = line.split() - if line[0] == "real:": - real_times[filename] = line[1] - elif line[0] == "time:": - user_times[filename] = line[1] + if line[0] == "time:": + time[filename] = line[1] elif line[0] == "space:": space[filename] = line[1] - return real_times, user_times, space + return time, space -# a function that takes as input a filename, an identifier, and 3 hashtables + +# a function that takes as input a filename, an identifier, and 2 hashtables # it opens the file and reads each line # it strips .identifier".runlim" from the filename -# it calls parse_line on the on the resulting string, the 3 hashtables, and each line -# it returns the 3 hashtables -def parse_file(filename, identifier, real_times, user_times, space): +# it calls parse_line on the on the resulting string, the 2 hashtables, and each line +# it returns the 2 hashtables +def parse_file(filename, identifier, time, space): with open(filename) as f: for line in f: filename = filename.replace(".{}.runlim".format(identifier), "") - real_times, user_times, space = parse_line(filename, real_times, user_times, space, line) - return real_times, user_times, space + time, space = parse_line( + filename, time, space, line) + return time, space + + +# a function that takes as input a directory +# it lists all the files in the directory recursively +# it returns the list of files + + +def list_files(directory): + import os + files = [] + for filename in os.listdir(directory): + if os.path.isfile(os.path.join(directory, filename)): + files.append(os.path.join(directory, filename)) + else: + files = files + list_files(os.path.join(directory, filename)) + return files + # a function that takes as input an identifier -# it creates 3 empty hashtables -# it calls parse_file on all the files in the current directory with suffix .identifier".runlim" -# it returns the 3 hashtables +# it creates 2 empty hashtables +# it calls list_files on "." (the current directory) +# it calls parse_file on all the files in the list with suffix .identifier".runlim" +# it returns the 2 hashtables def parse_all(identifier): - real_times = {} - user_times = {} + time = {} space = {} - for filename in os.listdir("."): + files = list_files(".") + for filename in files: if filename.endswith(".{}.runlim".format(identifier)): - real_times, user_times, space = parse_file(filename, identifier, real_times, user_times, space) - return real_times, user_times, space + time, space = parse_file(filename, identifier, time, space) + return time, space # a function that takes as input 2 hashtables -# it creates a third hashtable +# it creates one array # it iterates over the keys of the first hashtable -# it adds the key and the percentage difference of values of the first and second hashtables to the third hashtable -# it returns the third hashtable +# it adds the tuple (key, value in the first hashtable, value in the second hashtable, percentage difference of values of the first and second hashtables) to the array +# it returns the array + + def diff_hashtables(hashtable1, hashtable2): - diff = {} + diff = [] for key in hashtable1: - diff[key] = (float(hashtable1[key]) - float(hashtable2[key])) / float(hashtable2[key]) * 100 + diff.append((key, hashtable1[key], hashtable2[key], (float( + hashtable1[key]) - float(hashtable2[key])) / float(hashtable2[key]) * 100)) return diff -# a function that takes as input 2 identifiers +# a function that takes as input an array of tuples +# it sorts the array by the second element of the tuples +# it returns the sorted array + + +def sort_array(array): + return sorted(array, key=lambda x: x[1]) + + +def generate_scatter_plot(sorted_lines, xlabel, ylabel, title): + import matplotlib.pyplot as plt + import numpy as np + # create an array of the query timing differences + x_axis = [] + y_axis = [] + for line in sorted_lines: + x_axis.append(float(line[1])) + y_axis.append(float(line[2])) + # plot the query timing differences + plt.scatter(x_axis, y_axis) + # label x-axis as xlabel + plt.xlabel(xlabel) + # label y-axis as ylabel + plt.ylabel(ylabel) + # add a linear regression line + # Fit linear regression via least squares with numpy.polyfit + # It returns an slope (b) and intercept (a) + # deg=1 means linear fit (i.e. polynomial of degree 1) + b, a = np.polyfit(x_axis, y_axis, deg=1) + print(title + " slope: " + str(b)) + #print (title + "intercept: " + str(a)) + # Create sequence of 100 numbers from 0 to 100 + # find the maximum of the x_axis + max_x = max(x_axis) + xseq = np.linspace(0, max_x, num=1000) + + # Plot regression line + plt.plot(xseq, a + b * xseq, color="k", lw=2.5) + + # add a title + plt.title(title + "; Linear regression slope = " + str(b)) + + plt.show() + +# a function that takes as input a hashtable +# it prints the hashtable +# it exits the program + + +def print_hashtable(hashtable): + import sys + for key in hashtable: + print(key + " " + hashtable[key]) + sys.exit(0) + +# a function that takes as input two identifiers # it calls parse_all on the first identifier # it calls parse_all on the second identifier -# it calls diff_hashtables on the two resulting hashtables -# it returns the resulting hashtable +# it calls diff_hashtables on the two hashtables +# it calls sort_array on the two arrays +# it calls generate_scatter_plot on the two arrays with xlabel as the second identifier and ylabel as the first identifier and title as "F* runlim" + + def diff(identifier1, identifier2): - real_times1, user_times1, space1 = parse_all(identifier1) - real_times2, user_times2, space2 = parse_all(identifier2) - real_times_diff = diff_hashtables(real_times1, real_times2) - user_times_diff = diff_hashtables(user_times1, user_times2) + time1, space1 = parse_all(identifier1) + time2, space2 = parse_all(identifier2) + time_diff = diff_hashtables(time1, time2) space_diff = diff_hashtables(space1, space2) - return real_times_diff, user_times_diff, space_diff + time_diff = sort_array(time_diff) + space_diff = sort_array(space_diff) + generate_scatter_plot( + time_diff, "ID " + identifier2, "ID " + identifier1, "F* runlim time") + generate_scatter_plot( + space_diff, "ID " + identifier2, "ID " + identifier1, "F* runlim space") -# a function that takes as input two identifiers +# main function that parses two identifiers from the command line # it calls diff on the two identifiers -# it plots 3 scatter plots for the the resulting hashtables and saves them to files -def plot_diff(identifier1, identifier2): - real_times_diff, user_times_diff, space_diff = diff(identifier1, identifier2) - plt.scatter(real_times_diff.keys(), real_times_diff.values()) - plt.xlabel("filename") - plt.ylabel("real time difference (%)") - plt.savefig("real_times_diff_{}_{}.png".format(identifier1, identifier2)) - plt.clf() - plt.scatter(user_times_diff.keys(), user_times_diff.values()) - plt.xlabel("filename") - plt.ylabel("user time difference (%)") - plt.savefig("user_times_diff_{}_{}.png".format(identifier1, identifier2)) - plt.clf() - plt.scatter(space_diff.keys(), space_diff.values()) - plt.xlabel("filename") - plt.ylabel("space difference (%)") - plt.savefig("space_diff_{}_{}.png".format(identifier1, identifier2)) - plt.clf() - -# a function that reads two identifiers from the command line -# it calls plot_diff on the two identifiers -if __name__ == "__main__": - identifier1 = sys.argv[1] - identifier2 = sys.argv[2] - plot_diff(identifier1, identifier2) + + +def main(): + import sys + if len(sys.argv) != 3: + print( + "Usage: python diff_runlim.py identifier1(the new run) identifier2(the old run)") + return + diff(sys.argv[1], sys.argv[2]) + + +main() From db4f32be9db4f9fb60e35c9b69362a71f14a48fd Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Tue, 13 Dec 2022 16:03:51 +0530 Subject: [PATCH 10/18] some comments --- diff_runlim.py => .scripts/runlim_diff.py | 32 +++++++++++++++++------ 1 file changed, 24 insertions(+), 8 deletions(-) rename diff_runlim.py => .scripts/runlim_diff.py (79%) diff --git a/diff_runlim.py b/.scripts/runlim_diff.py similarity index 79% rename from diff_runlim.py rename to .scripts/runlim_diff.py index 5c84cae9ee9..8ae3000f2ac 100644 --- a/diff_runlim.py +++ b/.scripts/runlim_diff.py @@ -1,9 +1,22 @@ +# This script compares two runlim runs and produces a scatter plot with linear regression +# +# It takes as input two identifiers, the identifiers for the new run and the old run resp. +# +# Suppose the requirement is to compare the runlim runs across two branches b1 and b2 +# First you would git checkout b1, and run F* with MON=1 and MONID=1 (or any other identifier) +# Then you would git checkout b2, and run F* with MON=1 and MONID=2 (or something else) +# And then you can invoke this script as `python3 runlim_diff.py 2 1`, +# and it would compare the run 2 with the baseline run 1 +# +# It can be used across local changes also, not just to compare two branches +# + # a function that takes as input a filename, 2 hashtables, and a line # it removes the "[runlim] " prefix from the line # it splits the rest of the line into a list of strings # if the first element of the list is "time:", it adds the filename and the second string to the first hashtable # if the first element of the list is "space:", it adds the filename and the second string to the second hashtable -# # it returns the two hashtables +# it returns the two hashtables def parse_line(filename, time, space, line): line = line.replace("[runlim] ", "") line = line.split() @@ -67,18 +80,21 @@ def parse_all(identifier): def diff_hashtables(hashtable1, hashtable2): diff = [] - for key in hashtable1: - diff.append((key, hashtable1[key], hashtable2[key], (float( - hashtable1[key]) - float(hashtable2[key])) / float(hashtable2[key]) * 100)) + for key in hashtable2: + if key in hashtable1: + diff.append((key, hashtable1[key], hashtable2[key], (float( + hashtable1[key]) - float(hashtable2[key])) / float(hashtable2[key]) * 100)) + else: + print(key + " is in the base run but not the new run, dropping") return diff # a function that takes as input an array of tuples -# it sorts the array by the second element of the tuples +# it sorts the array by the third element of the tuples # it returns the sorted array def sort_array(array): - return sorted(array, key=lambda x: x[1]) + return sorted(array, key=lambda x: x[2]) def generate_scatter_plot(sorted_lines, xlabel, ylabel, title): @@ -88,8 +104,8 @@ def generate_scatter_plot(sorted_lines, xlabel, ylabel, title): x_axis = [] y_axis = [] for line in sorted_lines: - x_axis.append(float(line[1])) - y_axis.append(float(line[2])) + x_axis.append(float(line[2])) + y_axis.append(float(line[1])) # plot the query timing differences plt.scatter(x_axis, y_axis) # label x-axis as xlabel From 6ddc24a6365750737a28bb57588f4241d4e2ee07 Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Tue, 13 Dec 2022 16:05:45 +0530 Subject: [PATCH 11/18] nit --- .scripts/runlim_diff.py | 1 + 1 file changed, 1 insertion(+) diff --git a/.scripts/runlim_diff.py b/.scripts/runlim_diff.py index 8ae3000f2ac..bbe6e910070 100644 --- a/.scripts/runlim_diff.py +++ b/.scripts/runlim_diff.py @@ -96,6 +96,7 @@ def diff_hashtables(hashtable1, hashtable2): def sort_array(array): return sorted(array, key=lambda x: x[2]) +# generate_scatter_plot is lift and shift from runlim_stats.py def generate_scatter_plot(sorted_lines, xlabel, ylabel, title): import matplotlib.pyplot as plt From 24e61051139d8034c843e4ab051a67306952ff3f Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Tue, 13 Dec 2022 16:35:06 +0530 Subject: [PATCH 12/18] avoid a comp_to_comp_typ when checking positivity --- src/ocaml-output/FStar_TypeChecker_Util.ml | 21 +++++++++++---------- src/typechecker/FStar.TypeChecker.Util.fst | 8 +++++--- 2 files changed, 16 insertions(+), 13 deletions(-) diff --git a/src/ocaml-output/FStar_TypeChecker_Util.ml b/src/ocaml-output/FStar_TypeChecker_Util.ml index a6662b69c72..f6d7678f634 100644 --- a/src/ocaml-output/FStar_TypeChecker_Util.ml +++ b/src/ocaml-output/FStar_TypeChecker_Util.ml @@ -8747,18 +8747,19 @@ and (ty_strictly_positive_in_type : (fun uu___5 -> "Checking strict positivity in Tm_arrow"); (let check_comp1 = - let c1 = - let uu___5 = - FStar_TypeChecker_Env.unfold_effect_abbrev env c in - FStar_Compiler_Effect.op_Bar_Greater uu___5 - FStar_Syntax_Syntax.mk_Comp in - (FStar_Syntax_Util.is_pure_or_ghost_comp c1) || + (FStar_Syntax_Util.is_pure_or_ghost_comp c) || (let uu___5 = - FStar_TypeChecker_Env.lookup_effect_quals env - (FStar_Syntax_Util.comp_effect_name c1) in + let uu___6 = + let uu___7 = + FStar_Compiler_Effect.op_Bar_Greater c + FStar_Syntax_Util.comp_effect_name in + FStar_Compiler_Effect.op_Bar_Greater uu___7 + (FStar_TypeChecker_Env.norm_eff_name env) in + FStar_Compiler_Effect.op_Bar_Greater uu___6 + (FStar_TypeChecker_Env.lookup_effect_quals env) in FStar_Compiler_Effect.op_Bar_Greater uu___5 - (FStar_Compiler_List.existsb - (fun q -> q = FStar_Syntax_Syntax.TotalEffect))) in + (FStar_Compiler_List.contains + FStar_Syntax_Syntax.TotalEffect)) in if Prims.op_Negation check_comp1 then (debug_positivity env diff --git a/src/typechecker/FStar.TypeChecker.Util.fst b/src/typechecker/FStar.TypeChecker.Util.fst index 740af919c1b..8035648094a 100644 --- a/src/typechecker/FStar.TypeChecker.Util.fst +++ b/src/typechecker/FStar.TypeChecker.Util.fst @@ -3953,9 +3953,11 @@ and ty_strictly_positive_in_type env (ty_lid:lident) (btype:term) (unfolded:unfo | Tm_arrow (sbs, c) -> //binder type is an arrow type debug_positivity env (fun () -> "Checking strict positivity in Tm_arrow"); let check_comp = - let c = Env.unfold_effect_abbrev env c |> mk_Comp in - U.is_pure_or_ghost_comp c || (Env.lookup_effect_quals env (U.comp_effect_name c) |> List.existsb (fun q -> q = S.TotalEffect)) - in + U.is_pure_or_ghost_comp c || + (c |> U.comp_effect_name + |> Env.norm_eff_name env + |> Env.lookup_effect_quals env + |> List.contains S.TotalEffect) in if not check_comp then let _ = debug_positivity env (fun () -> "Checking strict positivity , the arrow is impure, so return true") in true From ddaebcc6e6e59a2bde3e7c0c24384582066f48b7 Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Tue, 13 Dec 2022 20:47:29 +0530 Subject: [PATCH 13/18] snap --- src/ocaml-output/FStar_TypeChecker_Env.ml | 7 + src/ocaml-output/FStar_TypeChecker_NBE.ml | 6 + .../FStar_TypeChecker_Normalize.ml | 487 +++++++++++++----- 3 files changed, 358 insertions(+), 142 deletions(-) diff --git a/src/ocaml-output/FStar_TypeChecker_Env.ml b/src/ocaml-output/FStar_TypeChecker_Env.ml index 5658f72f074..cf9f152b17e 100644 --- a/src/ocaml-output/FStar_TypeChecker_Env.ml +++ b/src/ocaml-output/FStar_TypeChecker_Env.ml @@ -16,6 +16,7 @@ type step = | UnfoldFully of FStar_Ident.lid Prims.list | UnfoldAttr of FStar_Ident.lid Prims.list | UnfoldQual of Prims.string Prims.list + | UnfoldNamespace of Prims.string Prims.list | UnfoldTac | PureSubtermsWithinComputations | Simplify @@ -81,6 +82,11 @@ let (uu___is_UnfoldQual : step -> Prims.bool) = match projectee with | UnfoldQual _0 -> true | uu___ -> false let (__proj__UnfoldQual__item___0 : step -> Prims.string Prims.list) = fun projectee -> match projectee with | UnfoldQual _0 -> _0 +let (uu___is_UnfoldNamespace : step -> Prims.bool) = + fun projectee -> + match projectee with | UnfoldNamespace _0 -> true | uu___ -> false +let (__proj__UnfoldNamespace__item___0 : step -> Prims.string Prims.list) = + fun projectee -> match projectee with | UnfoldNamespace _0 -> _0 let (uu___is_UnfoldTac : step -> Prims.bool) = fun projectee -> match projectee with | UnfoldTac -> true | uu___ -> false let (uu___is_PureSubtermsWithinComputations : step -> Prims.bool) = @@ -164,6 +170,7 @@ let rec (eq_step : step -> step -> Prims.bool) = && (FStar_Compiler_List.forall2 FStar_Ident.lid_equals lids1 lids2) | (UnfoldQual strs1, UnfoldQual strs2) -> strs1 = strs2 + | (UnfoldNamespace strs1, UnfoldNamespace strs2) -> strs1 = strs2 | uu___ -> false type sig_binding = (FStar_Ident.lident Prims.list * FStar_Syntax_Syntax.sigelt) diff --git a/src/ocaml-output/FStar_TypeChecker_NBE.ml b/src/ocaml-output/FStar_TypeChecker_NBE.ml index 6df3a5b96ba..28f12f18f8a 100644 --- a/src/ocaml-output/FStar_TypeChecker_NBE.ml +++ b/src/ocaml-output/FStar_TypeChecker_NBE.ml @@ -180,6 +180,8 @@ let (zeta_false : config -> config) = (uu___.FStar_TypeChecker_Cfg.unfold_attr); FStar_TypeChecker_Cfg.unfold_qual = (uu___.FStar_TypeChecker_Cfg.unfold_qual); + FStar_TypeChecker_Cfg.unfold_namespace = + (uu___.FStar_TypeChecker_Cfg.unfold_namespace); FStar_TypeChecker_Cfg.unfold_tac = (uu___.FStar_TypeChecker_Cfg.unfold_tac); FStar_TypeChecker_Cfg.pure_subterms_within_computations = @@ -3196,6 +3198,8 @@ let (normalize : (uu___.FStar_TypeChecker_Cfg.unfold_attr); FStar_TypeChecker_Cfg.unfold_qual = (uu___.FStar_TypeChecker_Cfg.unfold_qual); + FStar_TypeChecker_Cfg.unfold_namespace = + (uu___.FStar_TypeChecker_Cfg.unfold_namespace); FStar_TypeChecker_Cfg.unfold_tac = (uu___.FStar_TypeChecker_Cfg.unfold_tac); FStar_TypeChecker_Cfg.pure_subterms_within_computations = @@ -3304,6 +3308,8 @@ let (normalize_for_unit_test : (uu___.FStar_TypeChecker_Cfg.unfold_attr); FStar_TypeChecker_Cfg.unfold_qual = (uu___.FStar_TypeChecker_Cfg.unfold_qual); + FStar_TypeChecker_Cfg.unfold_namespace = + (uu___.FStar_TypeChecker_Cfg.unfold_namespace); FStar_TypeChecker_Cfg.unfold_tac = (uu___.FStar_TypeChecker_Cfg.unfold_tac); FStar_TypeChecker_Cfg.pure_subterms_within_computations = diff --git a/src/ocaml-output/FStar_TypeChecker_Normalize.ml b/src/ocaml-output/FStar_TypeChecker_Normalize.ml index b2ad925d247..aa7b8f5deba 100644 --- a/src/ocaml-output/FStar_TypeChecker_Normalize.ml +++ b/src/ocaml-output/FStar_TypeChecker_Normalize.ml @@ -1552,6 +1552,8 @@ let (reduce_equality : (FStar_TypeChecker_Cfg.default_steps.FStar_TypeChecker_Cfg.unfold_attr); FStar_TypeChecker_Cfg.unfold_qual = (FStar_TypeChecker_Cfg.default_steps.FStar_TypeChecker_Cfg.unfold_qual); + FStar_TypeChecker_Cfg.unfold_namespace = + (FStar_TypeChecker_Cfg.default_steps.FStar_TypeChecker_Cfg.unfold_namespace); FStar_TypeChecker_Cfg.unfold_tac = (FStar_TypeChecker_Cfg.default_steps.FStar_TypeChecker_Cfg.unfold_tac); FStar_TypeChecker_Cfg.pure_subterms_within_computations = @@ -2003,7 +2005,8 @@ let (should_unfold : ((cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_only), ((cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_fully), ((cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_attr), - ((cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_qual)) + ((cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_qual), + ((cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_namespace)) with | (FStar_Pervasives_Native.Some (FStar_Pervasives.Inr @@ -2018,16 +2021,16 @@ let (should_unfold : FStar_Syntax_Syntax.sigopts = uu___5;_}, uu___6), uu___7), - uu___8, uu___9, uu___10, uu___11) when + uu___8, uu___9, uu___10, uu___11, uu___12) when FStar_Compiler_List.contains FStar_Syntax_Syntax.HasMaskedEffect qs -> (FStar_TypeChecker_Cfg.log_unfolding cfg - (fun uu___13 -> + (fun uu___14 -> FStar_Compiler_Util.print_string " >> HasMaskedEffect, not unfolding\n"); no) - | (uu___, uu___1, uu___2, uu___3, uu___4) when + | (uu___, uu___1, uu___2, uu___3, uu___4, uu___5) when (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_tac && (FStar_Compiler_Util.for_some @@ -2035,7 +2038,7 @@ let (should_unfold : FStar_Syntax_Util.tac_opaque_attr) attrs) -> (FStar_TypeChecker_Cfg.log_unfolding cfg - (fun uu___6 -> + (fun uu___7 -> FStar_Compiler_Util.print_string " >> tac_opaque, not unfolding\n"); no) @@ -2052,7 +2055,7 @@ let (should_unfold : FStar_Syntax_Syntax.sigopts = uu___5;_}, uu___6), uu___7), - uu___8, uu___9, uu___10, uu___11) when + uu___8, uu___9, uu___10, uu___11, uu___12) when (is_rec && (Prims.op_Negation (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.zeta)) @@ -2061,53 +2064,53 @@ let (should_unfold : (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.zeta_full) -> (FStar_TypeChecker_Cfg.log_unfolding cfg - (fun uu___13 -> + (fun uu___14 -> FStar_Compiler_Util.print_string " >> It's a recursive definition but we're not doing Zeta, not unfolding\n"); no) | (uu___, FStar_Pervasives_Native.Some uu___1, uu___2, - uu___3, uu___4) -> + uu___3, uu___4, uu___5) -> (FStar_TypeChecker_Cfg.log_unfolding cfg - (fun uu___6 -> - let uu___7 = FStar_Syntax_Print.fv_to_string fv in + (fun uu___7 -> + let uu___8 = FStar_Syntax_Print.fv_to_string fv in FStar_Compiler_Util.print1 "should_unfold: Reached a %s with selective unfolding\n" - uu___7); + uu___8); (let meets_some_criterion = - let uu___6 = - let uu___7 = + let uu___7 = + let uu___8 = if (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.for_extraction then - let uu___8 = - let uu___9 = + let uu___9 = + let uu___10 = FStar_TypeChecker_Env.lookup_definition_qninfo [FStar_TypeChecker_Env.Eager_unfolding_only; FStar_TypeChecker_Env.InliningDelta] (fv.FStar_Syntax_Syntax.fv_name).FStar_Syntax_Syntax.v qninfo in - FStar_Compiler_Option.isSome uu___9 in - FStar_Compiler_Effect.op_Less_Bar yesno uu___8 + FStar_Compiler_Option.isSome uu___10 in + FStar_Compiler_Effect.op_Less_Bar yesno uu___9 else no in - let uu___8 = - let uu___9 = + let uu___9 = + let uu___10 = match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_only with | FStar_Pervasives_Native.None -> no | FStar_Pervasives_Native.Some lids -> - let uu___10 = + let uu___11 = FStar_Compiler_Util.for_some (FStar_Syntax_Syntax.fv_eq_lid fv) lids in FStar_Compiler_Effect.op_Less_Bar yesno - uu___10 in - let uu___10 = - let uu___11 = + uu___11 in + let uu___11 = + let uu___12 = match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_attr with | FStar_Pervasives_Native.None -> no | FStar_Pervasives_Native.Some lids -> - let uu___12 = + let uu___13 = FStar_Compiler_Util.for_some (fun at -> FStar_Compiler_Util.for_some @@ -2115,86 +2118,106 @@ let (should_unfold : FStar_Syntax_Util.is_fvar lid at) lids) attrs in FStar_Compiler_Effect.op_Less_Bar yesno - uu___12 in - let uu___12 = - let uu___13 = + uu___13 in + let uu___13 = + let uu___14 = match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_fully with | FStar_Pervasives_Native.None -> no | FStar_Pervasives_Native.Some lids -> - let uu___14 = + let uu___15 = FStar_Compiler_Util.for_some (FStar_Syntax_Syntax.fv_eq_lid fv) lids in FStar_Compiler_Effect.op_Less_Bar - fullyno uu___14 in - let uu___14 = - let uu___15 = + fullyno uu___15 in + let uu___15 = + let uu___16 = match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_qual with | FStar_Pervasives_Native.None -> no | FStar_Pervasives_Native.Some qs -> - let uu___16 = + let uu___17 = FStar_Compiler_Util.for_some (fun q -> FStar_Compiler_Util.for_some (fun qual -> - let uu___17 = + let uu___18 = FStar_Syntax_Print.qual_to_string qual in - uu___17 = q) quals) qs in + uu___18 = q) quals) qs in FStar_Compiler_Effect.op_Less_Bar - yesno uu___16 in - [uu___15] in - uu___13 :: uu___14 in - uu___11 :: uu___12 in - uu___9 :: uu___10 in - uu___7 :: uu___8 in - comb_or uu___6 in + yesno uu___17 in + let uu___17 = + let uu___18 = + match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_namespace + with + | FStar_Pervasives_Native.None -> no + | FStar_Pervasives_Native.Some + namespaces -> + let uu___19 = + FStar_Compiler_Util.for_some + (fun ns -> + let uu___20 = + let uu___21 = + FStar_Syntax_Syntax.lid_of_fv + fv in + FStar_Ident.nsstr uu___21 in + FStar_Compiler_Util.starts_with + uu___20 ns) namespaces in + FStar_Compiler_Effect.op_Less_Bar + yesno uu___19 in + [uu___18] in + uu___16 :: uu___17 in + uu___14 :: uu___15 in + uu___12 :: uu___13 in + uu___10 :: uu___11 in + uu___8 :: uu___9 in + comb_or uu___7 in meets_some_criterion)) | (uu___, uu___1, FStar_Pervasives_Native.Some uu___2, - uu___3, uu___4) -> + uu___3, uu___4, uu___5) -> (FStar_TypeChecker_Cfg.log_unfolding cfg - (fun uu___6 -> - let uu___7 = FStar_Syntax_Print.fv_to_string fv in + (fun uu___7 -> + let uu___8 = FStar_Syntax_Print.fv_to_string fv in FStar_Compiler_Util.print1 "should_unfold: Reached a %s with selective unfolding\n" - uu___7); + uu___8); (let meets_some_criterion = - let uu___6 = - let uu___7 = + let uu___7 = + let uu___8 = if (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.for_extraction then - let uu___8 = - let uu___9 = + let uu___9 = + let uu___10 = FStar_TypeChecker_Env.lookup_definition_qninfo [FStar_TypeChecker_Env.Eager_unfolding_only; FStar_TypeChecker_Env.InliningDelta] (fv.FStar_Syntax_Syntax.fv_name).FStar_Syntax_Syntax.v qninfo in - FStar_Compiler_Option.isSome uu___9 in - FStar_Compiler_Effect.op_Less_Bar yesno uu___8 + FStar_Compiler_Option.isSome uu___10 in + FStar_Compiler_Effect.op_Less_Bar yesno uu___9 else no in - let uu___8 = - let uu___9 = + let uu___9 = + let uu___10 = match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_only with | FStar_Pervasives_Native.None -> no | FStar_Pervasives_Native.Some lids -> - let uu___10 = + let uu___11 = FStar_Compiler_Util.for_some (FStar_Syntax_Syntax.fv_eq_lid fv) lids in FStar_Compiler_Effect.op_Less_Bar yesno - uu___10 in - let uu___10 = - let uu___11 = + uu___11 in + let uu___11 = + let uu___12 = match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_attr with | FStar_Pervasives_Native.None -> no | FStar_Pervasives_Native.Some lids -> - let uu___12 = + let uu___13 = FStar_Compiler_Util.for_some (fun at -> FStar_Compiler_Util.for_some @@ -2202,86 +2225,106 @@ let (should_unfold : FStar_Syntax_Util.is_fvar lid at) lids) attrs in FStar_Compiler_Effect.op_Less_Bar yesno - uu___12 in - let uu___12 = - let uu___13 = + uu___13 in + let uu___13 = + let uu___14 = match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_fully with | FStar_Pervasives_Native.None -> no | FStar_Pervasives_Native.Some lids -> - let uu___14 = + let uu___15 = FStar_Compiler_Util.for_some (FStar_Syntax_Syntax.fv_eq_lid fv) lids in FStar_Compiler_Effect.op_Less_Bar - fullyno uu___14 in - let uu___14 = - let uu___15 = + fullyno uu___15 in + let uu___15 = + let uu___16 = match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_qual with | FStar_Pervasives_Native.None -> no | FStar_Pervasives_Native.Some qs -> - let uu___16 = + let uu___17 = FStar_Compiler_Util.for_some (fun q -> FStar_Compiler_Util.for_some (fun qual -> - let uu___17 = + let uu___18 = FStar_Syntax_Print.qual_to_string qual in - uu___17 = q) quals) qs in + uu___18 = q) quals) qs in FStar_Compiler_Effect.op_Less_Bar - yesno uu___16 in - [uu___15] in - uu___13 :: uu___14 in - uu___11 :: uu___12 in - uu___9 :: uu___10 in - uu___7 :: uu___8 in - comb_or uu___6 in + yesno uu___17 in + let uu___17 = + let uu___18 = + match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_namespace + with + | FStar_Pervasives_Native.None -> no + | FStar_Pervasives_Native.Some + namespaces -> + let uu___19 = + FStar_Compiler_Util.for_some + (fun ns -> + let uu___20 = + let uu___21 = + FStar_Syntax_Syntax.lid_of_fv + fv in + FStar_Ident.nsstr uu___21 in + FStar_Compiler_Util.starts_with + uu___20 ns) namespaces in + FStar_Compiler_Effect.op_Less_Bar + yesno uu___19 in + [uu___18] in + uu___16 :: uu___17 in + uu___14 :: uu___15 in + uu___12 :: uu___13 in + uu___10 :: uu___11 in + uu___8 :: uu___9 in + comb_or uu___7 in meets_some_criterion)) | (uu___, uu___1, uu___2, FStar_Pervasives_Native.Some - uu___3, uu___4) -> + uu___3, uu___4, uu___5) -> (FStar_TypeChecker_Cfg.log_unfolding cfg - (fun uu___6 -> - let uu___7 = FStar_Syntax_Print.fv_to_string fv in + (fun uu___7 -> + let uu___8 = FStar_Syntax_Print.fv_to_string fv in FStar_Compiler_Util.print1 "should_unfold: Reached a %s with selective unfolding\n" - uu___7); + uu___8); (let meets_some_criterion = - let uu___6 = - let uu___7 = + let uu___7 = + let uu___8 = if (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.for_extraction then - let uu___8 = - let uu___9 = + let uu___9 = + let uu___10 = FStar_TypeChecker_Env.lookup_definition_qninfo [FStar_TypeChecker_Env.Eager_unfolding_only; FStar_TypeChecker_Env.InliningDelta] (fv.FStar_Syntax_Syntax.fv_name).FStar_Syntax_Syntax.v qninfo in - FStar_Compiler_Option.isSome uu___9 in - FStar_Compiler_Effect.op_Less_Bar yesno uu___8 + FStar_Compiler_Option.isSome uu___10 in + FStar_Compiler_Effect.op_Less_Bar yesno uu___9 else no in - let uu___8 = - let uu___9 = + let uu___9 = + let uu___10 = match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_only with | FStar_Pervasives_Native.None -> no | FStar_Pervasives_Native.Some lids -> - let uu___10 = + let uu___11 = FStar_Compiler_Util.for_some (FStar_Syntax_Syntax.fv_eq_lid fv) lids in FStar_Compiler_Effect.op_Less_Bar yesno - uu___10 in - let uu___10 = - let uu___11 = + uu___11 in + let uu___11 = + let uu___12 = match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_attr with | FStar_Pervasives_Native.None -> no | FStar_Pervasives_Native.Some lids -> - let uu___12 = + let uu___13 = FStar_Compiler_Util.for_some (fun at -> FStar_Compiler_Util.for_some @@ -2289,86 +2332,106 @@ let (should_unfold : FStar_Syntax_Util.is_fvar lid at) lids) attrs in FStar_Compiler_Effect.op_Less_Bar yesno - uu___12 in - let uu___12 = - let uu___13 = + uu___13 in + let uu___13 = + let uu___14 = match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_fully with | FStar_Pervasives_Native.None -> no | FStar_Pervasives_Native.Some lids -> - let uu___14 = + let uu___15 = FStar_Compiler_Util.for_some (FStar_Syntax_Syntax.fv_eq_lid fv) lids in FStar_Compiler_Effect.op_Less_Bar - fullyno uu___14 in - let uu___14 = - let uu___15 = + fullyno uu___15 in + let uu___15 = + let uu___16 = match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_qual with | FStar_Pervasives_Native.None -> no | FStar_Pervasives_Native.Some qs -> - let uu___16 = + let uu___17 = FStar_Compiler_Util.for_some (fun q -> FStar_Compiler_Util.for_some (fun qual -> - let uu___17 = + let uu___18 = FStar_Syntax_Print.qual_to_string qual in - uu___17 = q) quals) qs in + uu___18 = q) quals) qs in FStar_Compiler_Effect.op_Less_Bar - yesno uu___16 in - [uu___15] in - uu___13 :: uu___14 in - uu___11 :: uu___12 in - uu___9 :: uu___10 in - uu___7 :: uu___8 in - comb_or uu___6 in + yesno uu___17 in + let uu___17 = + let uu___18 = + match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_namespace + with + | FStar_Pervasives_Native.None -> no + | FStar_Pervasives_Native.Some + namespaces -> + let uu___19 = + FStar_Compiler_Util.for_some + (fun ns -> + let uu___20 = + let uu___21 = + FStar_Syntax_Syntax.lid_of_fv + fv in + FStar_Ident.nsstr uu___21 in + FStar_Compiler_Util.starts_with + uu___20 ns) namespaces in + FStar_Compiler_Effect.op_Less_Bar + yesno uu___19 in + [uu___18] in + uu___16 :: uu___17 in + uu___14 :: uu___15 in + uu___12 :: uu___13 in + uu___10 :: uu___11 in + uu___8 :: uu___9 in + comb_or uu___7 in meets_some_criterion)) | (uu___, uu___1, uu___2, uu___3, - FStar_Pervasives_Native.Some uu___4) -> + FStar_Pervasives_Native.Some uu___4, uu___5) -> (FStar_TypeChecker_Cfg.log_unfolding cfg - (fun uu___6 -> - let uu___7 = FStar_Syntax_Print.fv_to_string fv in + (fun uu___7 -> + let uu___8 = FStar_Syntax_Print.fv_to_string fv in FStar_Compiler_Util.print1 "should_unfold: Reached a %s with selective unfolding\n" - uu___7); + uu___8); (let meets_some_criterion = - let uu___6 = - let uu___7 = + let uu___7 = + let uu___8 = if (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.for_extraction then - let uu___8 = - let uu___9 = + let uu___9 = + let uu___10 = FStar_TypeChecker_Env.lookup_definition_qninfo [FStar_TypeChecker_Env.Eager_unfolding_only; FStar_TypeChecker_Env.InliningDelta] (fv.FStar_Syntax_Syntax.fv_name).FStar_Syntax_Syntax.v qninfo in - FStar_Compiler_Option.isSome uu___9 in - FStar_Compiler_Effect.op_Less_Bar yesno uu___8 + FStar_Compiler_Option.isSome uu___10 in + FStar_Compiler_Effect.op_Less_Bar yesno uu___9 else no in - let uu___8 = - let uu___9 = + let uu___9 = + let uu___10 = match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_only with | FStar_Pervasives_Native.None -> no | FStar_Pervasives_Native.Some lids -> - let uu___10 = + let uu___11 = FStar_Compiler_Util.for_some (FStar_Syntax_Syntax.fv_eq_lid fv) lids in FStar_Compiler_Effect.op_Less_Bar yesno - uu___10 in - let uu___10 = - let uu___11 = + uu___11 in + let uu___11 = + let uu___12 = match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_attr with | FStar_Pervasives_Native.None -> no | FStar_Pervasives_Native.Some lids -> - let uu___12 = + let uu___13 = FStar_Compiler_Util.for_some (fun at -> FStar_Compiler_Util.for_some @@ -2376,42 +2439,169 @@ let (should_unfold : FStar_Syntax_Util.is_fvar lid at) lids) attrs in FStar_Compiler_Effect.op_Less_Bar yesno - uu___12 in - let uu___12 = - let uu___13 = + uu___13 in + let uu___13 = + let uu___14 = match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_fully with | FStar_Pervasives_Native.None -> no | FStar_Pervasives_Native.Some lids -> - let uu___14 = + let uu___15 = FStar_Compiler_Util.for_some (FStar_Syntax_Syntax.fv_eq_lid fv) lids in FStar_Compiler_Effect.op_Less_Bar - fullyno uu___14 in + fullyno uu___15 in + let uu___15 = + let uu___16 = + match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_qual + with + | FStar_Pervasives_Native.None -> no + | FStar_Pervasives_Native.Some qs -> + let uu___17 = + FStar_Compiler_Util.for_some + (fun q -> + FStar_Compiler_Util.for_some + (fun qual -> + let uu___18 = + FStar_Syntax_Print.qual_to_string + qual in + uu___18 = q) quals) qs in + FStar_Compiler_Effect.op_Less_Bar + yesno uu___17 in + let uu___17 = + let uu___18 = + match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_namespace + with + | FStar_Pervasives_Native.None -> no + | FStar_Pervasives_Native.Some + namespaces -> + let uu___19 = + FStar_Compiler_Util.for_some + (fun ns -> + let uu___20 = + let uu___21 = + FStar_Syntax_Syntax.lid_of_fv + fv in + FStar_Ident.nsstr uu___21 in + FStar_Compiler_Util.starts_with + uu___20 ns) namespaces in + FStar_Compiler_Effect.op_Less_Bar + yesno uu___19 in + [uu___18] in + uu___16 :: uu___17 in + uu___14 :: uu___15 in + uu___12 :: uu___13 in + uu___10 :: uu___11 in + uu___8 :: uu___9 in + comb_or uu___7 in + meets_some_criterion)) + | (uu___, uu___1, uu___2, uu___3, uu___4, + FStar_Pervasives_Native.Some uu___5) -> + (FStar_TypeChecker_Cfg.log_unfolding cfg + (fun uu___7 -> + let uu___8 = FStar_Syntax_Print.fv_to_string fv in + FStar_Compiler_Util.print1 + "should_unfold: Reached a %s with selective unfolding\n" + uu___8); + (let meets_some_criterion = + let uu___7 = + let uu___8 = + if + (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.for_extraction + then + let uu___9 = + let uu___10 = + FStar_TypeChecker_Env.lookup_definition_qninfo + [FStar_TypeChecker_Env.Eager_unfolding_only; + FStar_TypeChecker_Env.InliningDelta] + (fv.FStar_Syntax_Syntax.fv_name).FStar_Syntax_Syntax.v + qninfo in + FStar_Compiler_Option.isSome uu___10 in + FStar_Compiler_Effect.op_Less_Bar yesno uu___9 + else no in + let uu___9 = + let uu___10 = + match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_only + with + | FStar_Pervasives_Native.None -> no + | FStar_Pervasives_Native.Some lids -> + let uu___11 = + FStar_Compiler_Util.for_some + (FStar_Syntax_Syntax.fv_eq_lid fv) + lids in + FStar_Compiler_Effect.op_Less_Bar yesno + uu___11 in + let uu___11 = + let uu___12 = + match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_attr + with + | FStar_Pervasives_Native.None -> no + | FStar_Pervasives_Native.Some lids -> + let uu___13 = + FStar_Compiler_Util.for_some + (fun at -> + FStar_Compiler_Util.for_some + (fun lid -> + FStar_Syntax_Util.is_fvar + lid at) lids) attrs in + FStar_Compiler_Effect.op_Less_Bar yesno + uu___13 in + let uu___13 = let uu___14 = - let uu___15 = + match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_fully + with + | FStar_Pervasives_Native.None -> no + | FStar_Pervasives_Native.Some lids -> + let uu___15 = + FStar_Compiler_Util.for_some + (FStar_Syntax_Syntax.fv_eq_lid fv) + lids in + FStar_Compiler_Effect.op_Less_Bar + fullyno uu___15 in + let uu___15 = + let uu___16 = match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_qual with | FStar_Pervasives_Native.None -> no | FStar_Pervasives_Native.Some qs -> - let uu___16 = + let uu___17 = FStar_Compiler_Util.for_some (fun q -> FStar_Compiler_Util.for_some (fun qual -> - let uu___17 = + let uu___18 = FStar_Syntax_Print.qual_to_string qual in - uu___17 = q) quals) qs in + uu___18 = q) quals) qs in FStar_Compiler_Effect.op_Less_Bar - yesno uu___16 in - [uu___15] in - uu___13 :: uu___14 in - uu___11 :: uu___12 in - uu___9 :: uu___10 in - uu___7 :: uu___8 in - comb_or uu___6 in + yesno uu___17 in + let uu___17 = + let uu___18 = + match (cfg.FStar_TypeChecker_Cfg.steps).FStar_TypeChecker_Cfg.unfold_namespace + with + | FStar_Pervasives_Native.None -> no + | FStar_Pervasives_Native.Some + namespaces -> + let uu___19 = + FStar_Compiler_Util.for_some + (fun ns -> + let uu___20 = + let uu___21 = + FStar_Syntax_Syntax.lid_of_fv + fv in + FStar_Ident.nsstr uu___21 in + FStar_Compiler_Util.starts_with + uu___20 ns) namespaces in + FStar_Compiler_Effect.op_Less_Bar + yesno uu___19 in + [uu___18] in + uu___16 :: uu___17 in + uu___14 :: uu___15 in + uu___12 :: uu___13 in + uu___10 :: uu___11 in + uu___8 :: uu___9 in + comb_or uu___7 in meets_some_criterion)) | uu___ -> default_unfolding ()) in FStar_TypeChecker_Cfg.log_unfolding cfg @@ -2518,6 +2708,8 @@ let decide_unfolding : FStar_Pervasives_Native.None; FStar_TypeChecker_Cfg.unfold_qual = FStar_Pervasives_Native.None; + FStar_TypeChecker_Cfg.unfold_namespace = + FStar_Pervasives_Native.None; FStar_TypeChecker_Cfg.unfold_tac = (uu___.FStar_TypeChecker_Cfg.unfold_tac); FStar_TypeChecker_Cfg.pure_subterms_within_computations @@ -2855,6 +3047,8 @@ let rec (norm : (uu___3.FStar_TypeChecker_Cfg.unfold_attr); FStar_TypeChecker_Cfg.unfold_qual = (uu___3.FStar_TypeChecker_Cfg.unfold_qual); + FStar_TypeChecker_Cfg.unfold_namespace = + (uu___3.FStar_TypeChecker_Cfg.unfold_namespace); FStar_TypeChecker_Cfg.unfold_tac = (uu___3.FStar_TypeChecker_Cfg.unfold_tac); FStar_TypeChecker_Cfg.pure_subterms_within_computations @@ -3027,6 +3221,8 @@ let rec (norm : (uu___5.FStar_TypeChecker_Cfg.unfold_attr); FStar_TypeChecker_Cfg.unfold_qual = (uu___5.FStar_TypeChecker_Cfg.unfold_qual); + FStar_TypeChecker_Cfg.unfold_namespace = + (uu___5.FStar_TypeChecker_Cfg.unfold_namespace); FStar_TypeChecker_Cfg.unfold_tac = (uu___5.FStar_TypeChecker_Cfg.unfold_tac); FStar_TypeChecker_Cfg.pure_subterms_within_computations @@ -3654,6 +3850,8 @@ let rec (norm : (uu___2.FStar_TypeChecker_Cfg.unfold_attr); FStar_TypeChecker_Cfg.unfold_qual = (uu___2.FStar_TypeChecker_Cfg.unfold_qual); + FStar_TypeChecker_Cfg.unfold_namespace = + (uu___2.FStar_TypeChecker_Cfg.unfold_namespace); FStar_TypeChecker_Cfg.unfold_tac = (uu___2.FStar_TypeChecker_Cfg.unfold_tac); FStar_TypeChecker_Cfg.pure_subterms_within_computations @@ -7035,6 +7233,8 @@ and (rebuild : FStar_Pervasives_Native.None; FStar_TypeChecker_Cfg.unfold_qual = FStar_Pervasives_Native.None; + FStar_TypeChecker_Cfg.unfold_namespace = + FStar_Pervasives_Native.None; FStar_TypeChecker_Cfg.unfold_tac = false; FStar_TypeChecker_Cfg.pure_subterms_within_computations = @@ -7305,6 +7505,9 @@ and (rebuild : (uu___10.FStar_TypeChecker_Cfg.unfold_attr); FStar_TypeChecker_Cfg.unfold_qual = (uu___10.FStar_TypeChecker_Cfg.unfold_qual); + FStar_TypeChecker_Cfg.unfold_namespace + = + (uu___10.FStar_TypeChecker_Cfg.unfold_namespace); FStar_TypeChecker_Cfg.unfold_tac = (uu___10.FStar_TypeChecker_Cfg.unfold_tac); FStar_TypeChecker_Cfg.pure_subterms_within_computations From 649611de1c1b083475dcb49324e1b78bf5acfbdc Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Tue, 13 Dec 2022 21:27:45 +0530 Subject: [PATCH 14/18] nits --- src/ocaml-output/FStar_Syntax_Util.ml | 2 +- src/ocaml-output/FStar_TypeChecker_NBE.ml | 12 +++---- src/ocaml-output/FStar_TypeChecker_NBETerm.ml | 20 ++++-------- src/syntax/FStar.Syntax.Util.fst | 32 ++----------------- src/typechecker/FStar.TypeChecker.NBE.fst | 10 +++--- src/typechecker/FStar.TypeChecker.NBETerm.fst | 2 +- .../FStar.TypeChecker.NBETerm.fsti | 4 +-- 7 files changed, 22 insertions(+), 60 deletions(-) diff --git a/src/ocaml-output/FStar_Syntax_Util.ml b/src/ocaml-output/FStar_Syntax_Util.ml index a092ceae4a6..80e87722d5a 100644 --- a/src/ocaml-output/FStar_Syntax_Util.ml +++ b/src/ocaml-output/FStar_Syntax_Util.ml @@ -4209,8 +4209,8 @@ and (unbound_variables_comp : FStar_Syntax_Syntax.comp -> FStar_Syntax_Syntax.bv Prims.list) = fun c -> match c.FStar_Syntax_Syntax.n with - | FStar_Syntax_Syntax.GTotal t -> unbound_variables t | FStar_Syntax_Syntax.Total t -> unbound_variables t + | FStar_Syntax_Syntax.GTotal t -> unbound_variables t | FStar_Syntax_Syntax.Comp ct -> let uu___ = unbound_variables ct.FStar_Syntax_Syntax.result_typ in let uu___1 = diff --git a/src/ocaml-output/FStar_TypeChecker_NBE.ml b/src/ocaml-output/FStar_TypeChecker_NBE.ml index 28f12f18f8a..fe6d7091c67 100644 --- a/src/ocaml-output/FStar_TypeChecker_NBE.ml +++ b/src/ocaml-output/FStar_TypeChecker_NBE.ml @@ -1324,14 +1324,10 @@ and (translate_comp : fun c -> match c.FStar_Syntax_Syntax.n with | FStar_Syntax_Syntax.Total typ -> - let uu___ = - let uu___1 = translate cfg bs typ in - (uu___1, FStar_Pervasives_Native.None) in + let uu___ = translate cfg bs typ in FStar_TypeChecker_NBETerm.Tot uu___ | FStar_Syntax_Syntax.GTotal typ -> - let uu___ = - let uu___1 = translate cfg bs typ in - (uu___1, FStar_Pervasives_Native.None) in + let uu___ = translate cfg bs typ in FStar_TypeChecker_NBETerm.GTot uu___ | FStar_Syntax_Syntax.Comp ctyp -> let uu___ = translate_comp_typ cfg bs ctyp in @@ -2038,9 +2034,9 @@ and (readback_comp : fun c -> let c' = match c with - | FStar_TypeChecker_NBETerm.Tot (typ, _u) -> + | FStar_TypeChecker_NBETerm.Tot typ -> let uu___ = readback cfg typ in FStar_Syntax_Syntax.Total uu___ - | FStar_TypeChecker_NBETerm.GTot (typ, _u) -> + | FStar_TypeChecker_NBETerm.GTot typ -> let uu___ = readback cfg typ in FStar_Syntax_Syntax.GTotal uu___ | FStar_TypeChecker_NBETerm.Comp ctyp -> let uu___ = readback_comp_typ cfg ctyp in diff --git a/src/ocaml-output/FStar_TypeChecker_NBETerm.ml b/src/ocaml-output/FStar_TypeChecker_NBETerm.ml index 6d69316bd3b..fe40147c534 100644 --- a/src/ocaml-output/FStar_TypeChecker_NBETerm.ml +++ b/src/ocaml-output/FStar_TypeChecker_NBETerm.ml @@ -86,10 +86,8 @@ and t = { nbe_t: t' ; nbe_r: FStar_Compiler_Range.range } and comp = - | Tot of (t * FStar_Syntax_Syntax.universe FStar_Pervasives_Native.option) - - | GTot of (t * FStar_Syntax_Syntax.universe FStar_Pervasives_Native.option) - + | Tot of t + | GTot of t | Comp of comp_typ and comp_typ = { @@ -264,14 +262,12 @@ let (__proj__Mkt__item__nbe_r : t -> FStar_Compiler_Range.range) = fun projectee -> match projectee with | { nbe_t; nbe_r;_} -> nbe_r let (uu___is_Tot : comp -> Prims.bool) = fun projectee -> match projectee with | Tot _0 -> true | uu___ -> false -let (__proj__Tot__item___0 : - comp -> (t * FStar_Syntax_Syntax.universe FStar_Pervasives_Native.option)) - = fun projectee -> match projectee with | Tot _0 -> _0 +let (__proj__Tot__item___0 : comp -> t) = + fun projectee -> match projectee with | Tot _0 -> _0 let (uu___is_GTot : comp -> Prims.bool) = fun projectee -> match projectee with | GTot _0 -> true | uu___ -> false -let (__proj__GTot__item___0 : - comp -> (t * FStar_Syntax_Syntax.universe FStar_Pervasives_Native.option)) - = fun projectee -> match projectee with | GTot _0 -> _0 +let (__proj__GTot__item___0 : comp -> t) = + fun projectee -> match projectee with | GTot _0 -> _0 let (uu___is_Comp : comp -> Prims.bool) = fun projectee -> match projectee with | Comp _0 -> true | uu___ -> false let (__proj__Comp__item___0 : comp -> comp_typ) = @@ -800,9 +796,7 @@ let (make_arrow1 : t -> arg -> t) = fun t1 -> fun a -> FStar_Compiler_Effect.op_Less_Bar mk_t - (Arrow - (FStar_Pervasives.Inr - ([a], (Tot (t1, FStar_Pervasives_Native.None))))) + (Arrow (FStar_Pervasives.Inr ([a], (Tot t1)))) let lazy_embed : 'a . FStar_Syntax_Syntax.emb_typ -> 'a -> (unit -> t) -> t = fun et -> fun x -> diff --git a/src/syntax/FStar.Syntax.Util.fst b/src/syntax/FStar.Syntax.Util.fst index f3b8ac03e67..a8796ddb6d2 100644 --- a/src/syntax/FStar.Syntax.Util.fst +++ b/src/syntax/FStar.Syntax.Util.fst @@ -247,34 +247,6 @@ let comp_eff_name_res_and_args (c:comp) : lident & typ & args = | GTotal t -> PC.effect_GTot_lid, t, [] | Comp c -> c.effect_name, c.result_typ, c.effect_args -// (* Duplicate of the function below not failing when universe is not provided *) -// let comp_to_comp_typ_nouniv (c:comp) : comp_typ = -// match c.n with -// | Comp c -> c -// | Total (t, u_opt) -// | GTotal(t, u_opt) -> -// {comp_univs=dflt [] (map_opt u_opt (fun x -> [x])); -// effect_name=comp_effect_name c; -// result_typ=t; -// effect_args=[]; -// flags=comp_flags c} - -// let comp_set_flags (c:comp) f = -// {c with n=Comp ({comp_to_comp_typ_nouniv c with flags=f})} - -// let comp_to_comp_typ (c:comp) : comp_typ = -// match c.n with -// | Comp c -> c -// | Total (t, Some u) -// | GTotal(t, Some u) -> -// {comp_univs=[u]; -// effect_name=comp_effect_name c; -// result_typ=t; -// effect_args=[]; -// flags=comp_flags c} -// | _ -> failwith "Assertion failed: Computation type without universe" - - (* * For layered effects, given a (repr a is), return is * For wp effects, given a (unit -> M a wp), return wp @@ -2147,8 +2119,8 @@ and unbound_variables_ascription asc = and unbound_variables_comp c = match c.n with - | GTotal t - | Total t -> + | Total t + | GTotal t -> unbound_variables t | Comp ct -> diff --git a/src/typechecker/FStar.TypeChecker.NBE.fst b/src/typechecker/FStar.TypeChecker.NBE.fst index b8297f023d0..a96a05669d5 100644 --- a/src/typechecker/FStar.TypeChecker.NBE.fst +++ b/src/typechecker/FStar.TypeChecker.NBE.fst @@ -710,9 +710,9 @@ let rec translate (cfg:config) (bs:list t) (e:term) : t = and translate_comp cfg bs (c:S.comp) : comp = match c.n with - | S.Total typ -> Tot (translate cfg bs typ, None) - | S.GTotal typ -> GTot (translate cfg bs typ, None) - | S.Comp ctyp -> Comp (translate_comp_typ cfg bs ctyp) + | S.Total typ -> Tot (translate cfg bs typ) + | S.GTotal typ -> GTot (translate cfg bs typ) + | S.Comp ctyp -> Comp (translate_comp_typ cfg bs ctyp) (* uncurried application *) and iapp (cfg : config) (f:t) (args:args) : t = @@ -974,8 +974,8 @@ and translate_constant (c : sconst) : constant = and readback_comp cfg (c: comp) : S.comp = let c' = match c with - | Tot (typ, _u) -> S.Total (readback cfg typ) - | GTot (typ, _u) -> S.GTotal (readback cfg typ) + | Tot typ -> S.Total (readback cfg typ) + | GTot typ -> S.GTotal (readback cfg typ) | Comp ctyp -> S.Comp (readback_comp_typ cfg ctyp) in S.mk c' Range.dummyRange diff --git a/src/typechecker/FStar.TypeChecker.NBETerm.fst b/src/typechecker/FStar.TypeChecker.NBETerm.fst index 22f2717473a..5ac046f9768 100644 --- a/src/typechecker/FStar.TypeChecker.NBETerm.fst +++ b/src/typechecker/FStar.TypeChecker.NBETerm.fst @@ -244,7 +244,7 @@ let as_iarg (a:t) : arg = (a, S.as_aqual_implicit true) let as_arg (a:t) : arg = (a, None) // Non-dependent total arrow -let make_arrow1 t1 (a:arg) : t = mk_t <| Arrow (Inr ([a], Tot (t1, None))) +let make_arrow1 t1 (a:arg) : t = mk_t <| Arrow (Inr ([a], Tot t1)) let lazy_embed (et:emb_typ) (x:'a) (f:unit -> t) = if !Options.debug_embedding diff --git a/src/typechecker/FStar.TypeChecker.NBETerm.fsti b/src/typechecker/FStar.TypeChecker.NBETerm.fsti index 709f5063ceb..9437e79c351 100644 --- a/src/typechecker/FStar.TypeChecker.NBETerm.fsti +++ b/src/typechecker/FStar.TypeChecker.NBETerm.fsti @@ -166,8 +166,8 @@ and t = { } and comp = - | Tot of t * option universe - | GTot of t * option universe + | Tot of t + | GTot of t | Comp of comp_typ and comp_typ = { From ca232e5fd148e5c84bc273b9c787eac93c79b5fe Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Wed, 14 Dec 2022 12:12:58 +0000 Subject: [PATCH 15/18] nit --- src/ocaml-output/FStar_Reflection_Basic.ml | 10 +++++++--- src/ocaml-output/FStar_Tactics_Hooks.ml | 2 +- src/reflection/FStar.Reflection.Basic.fst | 6 +++++- src/tactics/FStar.Tactics.Hooks.fst | 2 +- 4 files changed, 14 insertions(+), 6 deletions(-) diff --git a/src/ocaml-output/FStar_Reflection_Basic.ml b/src/ocaml-output/FStar_Reflection_Basic.ml index e24ef550d65..eb4a4431175 100644 --- a/src/ocaml-output/FStar_Reflection_Basic.ml +++ b/src/ocaml-output/FStar_Reflection_Basic.ml @@ -391,6 +391,12 @@ let (pack_comp : FStar_Reflection_Data.comp_view -> FStar_Syntax_Syntax.comp) let pack_arg uu___ = match uu___ with | (a, q) -> let uu___1 = pack_aqual q in (a, uu___1) in + let flags = + if (FStar_Compiler_List.length decrs) = Prims.int_zero + then [] + else + [FStar_Syntax_Syntax.DECREASES + (FStar_Syntax_Syntax.Decreases_lex decrs)] in let ct = let uu___ = FStar_Ident.lid_of_path ef FStar_Compiler_Range.dummyRange in @@ -400,9 +406,7 @@ let (pack_comp : FStar_Reflection_Data.comp_view -> FStar_Syntax_Syntax.comp) FStar_Syntax_Syntax.effect_name = uu___; FStar_Syntax_Syntax.result_typ = res; FStar_Syntax_Syntax.effect_args = uu___1; - FStar_Syntax_Syntax.flags = - [FStar_Syntax_Syntax.DECREASES - (FStar_Syntax_Syntax.Decreases_lex decrs)] + FStar_Syntax_Syntax.flags = flags } in FStar_Syntax_Syntax.mk_Comp ct let (pack_const : FStar_Reflection_Data.vconst -> FStar_Syntax_Syntax.sconst) diff --git a/src/ocaml-output/FStar_Tactics_Hooks.ml b/src/ocaml-output/FStar_Tactics_Hooks.ml index ad69c9fbf3f..a0276a4476a 100644 --- a/src/ocaml-output/FStar_Tactics_Hooks.ml +++ b/src/ocaml-output/FStar_Tactics_Hooks.ml @@ -1705,7 +1705,7 @@ let (splice : FStar_Common.string_of_list FStar_Syntax_Print.sigelt_to_string sigelts in FStar_Compiler_Util.print1 - "splice: got decls = %s\n" uu___7 + "splice: got decls = {\n\n%s\n\n}\n" uu___7 else ()); (let sigelts1 = FStar_Compiler_Effect.op_Bar_Greater sigelts diff --git a/src/reflection/FStar.Reflection.Basic.fst b/src/reflection/FStar.Reflection.Basic.fst index e544d9d273b..de07cf9a1d0 100644 --- a/src/reflection/FStar.Reflection.Basic.fst +++ b/src/reflection/FStar.Reflection.Basic.fst @@ -342,11 +342,15 @@ let pack_comp (cv : comp_view) : comp = | C_Eff (us, ef, res, args, decrs) -> let pack_arg (a, q) = (a, pack_aqual q) in + let flags = + if List.length decrs = 0 + then [] + else [DECREASES (Decreases_lex decrs)] in let ct = { comp_univs = us ; effect_name = Ident.lid_of_path ef Range.dummyRange ; result_typ = res ; effect_args = List.map pack_arg args - ; flags = [DECREASES (Decreases_lex decrs)] } in + ; flags = flags } in S.mk_Comp ct let pack_const (c:vconst) : sconst = diff --git a/src/tactics/FStar.Tactics.Hooks.fst b/src/tactics/FStar.Tactics.Hooks.fst index 118004f3899..4abb5810757 100644 --- a/src/tactics/FStar.Tactics.Hooks.fst +++ b/src/tactics/FStar.Tactics.Hooks.fst @@ -828,7 +828,7 @@ let splice (env:Env.env) (rng:Range.range) (tau:term) : list sigelt = then Err.raise_error (Err.Fatal_OpenGoalsInSynthesis, "splice left open goals") typ.pos; if !tacdbg then - BU.print1 "splice: got decls = %s\n" + BU.print1 "splice: got decls = {\n\n%s\n\n}\n" (FStar.Common.string_of_list Print.sigelt_to_string sigelts); let sigelts = sigelts |> List.map (fun se -> From 5aa99c4f8a9483aa6a995740b119b834acb308e0 Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Thu, 15 Dec 2022 09:01:10 +0530 Subject: [PATCH 16/18] incrementing checked file version number --- src/fstar/FStar.CheckedFiles.fst | 2 +- src/ocaml-output/FStar_CheckedFiles.ml | 2 +- ulib/prims.fst | 2 +- 3 files changed, 3 insertions(+), 3 deletions(-) diff --git a/src/fstar/FStar.CheckedFiles.fst b/src/fstar/FStar.CheckedFiles.fst index 3bc62ef4c2b..82a38a8a492 100644 --- a/src/fstar/FStar.CheckedFiles.fst +++ b/src/fstar/FStar.CheckedFiles.fst @@ -40,7 +40,7 @@ module Dep = FStar.Parser.Dep * detect when loading the cache that the version number is same * It needs to be kept in sync with prims.fst *) -let cache_version_number = 46 +let cache_version_number = 47 (* * Abbreviation for what we store in the checked files (stages as described below) diff --git a/src/ocaml-output/FStar_CheckedFiles.ml b/src/ocaml-output/FStar_CheckedFiles.ml index 2174dde8c92..2a0475606f6 100644 --- a/src/ocaml-output/FStar_CheckedFiles.ml +++ b/src/ocaml-output/FStar_CheckedFiles.ml @@ -1,5 +1,5 @@ open Prims -let (cache_version_number : Prims.int) = (Prims.of_int (46)) +let (cache_version_number : Prims.int) = (Prims.of_int (47)) type tc_result = { checked_module: FStar_Syntax_Syntax.modul ; diff --git a/ulib/prims.fst b/ulib/prims.fst index 85a22a1dcea..ebdbf515479 100644 --- a/ulib/prims.fst +++ b/ulib/prims.fst @@ -730,4 +730,4 @@ let labeled (r: range) (msg: string) (b: Type) : Type = b (** THIS IS MEANT TO BE KEPT IN SYNC WITH FStar.CheckedFiles.fs Incrementing this forces all .checked files to be invalidated *) irreducible -let __cache_version_number__ = 46 +let __cache_version_number__ = 47 From b79e39db1c8fb60c57cdfe8f42363388bdca2031 Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Thu, 15 Dec 2022 09:01:15 +0530 Subject: [PATCH 17/18] CHANGES.md --- CHANGES.md | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/CHANGES.md b/CHANGES.md index a812e326ff3..8a8174dfe0e 100644 --- a/CHANGES.md +++ b/CHANGES.md @@ -14,6 +14,13 @@ Guidelines for the changelog: # Version 0.9.7.0 ## Tactics & Reflection + * PR https://github.com/FStarLang/FStar/pull/2785 changes the reflection syntax + for computation types, by removing universe field from the Total and GTotal + comps. It also moves the decreases clause to the general C_Eff case. + + This is a breaking change for the reflection clients, but the regressions should + only be syntactic. + * As a better fix for Bug2635, F* now has a memoizing core typechecker for total (including ghost terms that are total) terms. The unification solutions, including those computed in the tactics engine, are typechecked using this core typechecker. From b71a523bafbb3e1ed3403819e4719dca70522487 Mon Sep 17 00:00:00 2001 From: Aseem Rastogi Date: Thu, 15 Dec 2022 09:04:17 +0530 Subject: [PATCH 18/18] deprecating runlim_stats.py --- .scripts/{runlim_stats.py => runlim_diff_old.py} | 10 ++++++++++ 1 file changed, 10 insertions(+) rename .scripts/{runlim_stats.py => runlim_diff_old.py} (93%) diff --git a/.scripts/runlim_stats.py b/.scripts/runlim_diff_old.py similarity index 93% rename from .scripts/runlim_stats.py rename to .scripts/runlim_diff_old.py index 34672a23d2c..9b5d616b738 100755 --- a/.scripts/runlim_stats.py +++ b/.scripts/runlim_diff_old.py @@ -1,3 +1,9 @@ +# +# NOTE: This file may soon be removed from the repository, +# superseded by runlim_diff.py that provides an easier +# API to compare two runlim runs +# + from array import array import re @@ -162,4 +168,8 @@ def main(): generate_scatter_plot(sorted_time_lines, file1, file2, "time") +# +# SEE THE NOTE AT THE BEGINNING OF THE FILE +# + main()