diff --git a/examples/exception.ec b/examples/exception.ec new file mode 100644 index 000000000..44276a186 --- /dev/null +++ b/examples/exception.ec @@ -0,0 +1,102 @@ +require import AllCore. + +exception assume. +exception assert. + +module M' ={ + proc truc (x:int) : int = { + if (x = 3) { + raise assume; + } + if (x=3) { + raise assert; + } + return x; + } +}. + +op p7: bool. +op p1: bool. +op p2: bool. +op p3: bool. + +lemma assume_assert : +hoare [M'.truc : p7 ==> p3 | assume:p1 |assert: p2 ]. + proof. +admitted. + +op p8: bool. +op q1: bool. +op q2: bool. +op q3: bool. + +lemma assert_assume : +hoare [M'.truc : p8 ==> q3 | assume:q1 |assert: q2 ]. + proof. +admitted. + +op p4: bool. +op p5: bool. +op p6: bool. +op p9: bool. +lemma assert_assume' : +hoare [M'.truc : p9 ==> p4 | assume:p6 |assert: p5 ]. + proof. + conseq (assume_assert) (assert_assume). + + admit. + + admit. + + admit. + + admit. + qed. + +exception e1. +exception e2. +exception e3. + +module M ={ + proc f1 (x:int) : int = { + if (x = 3) { + raise e1; + } else { + x <- 5; + } + return x; + } + + proc f2 (x:int) : int = { + if (x = 3) { + x <- x; + x <@ f1(x); + } else { + x <@ f1(x); + } + return x; + } +}. + + +(* Multiple time post for the same exception *) + +lemma l_f1 (_x: int): +hoare [M.f1 : _x = x ==> (4 < res) | e1:_x = 3 | e2:false ]. + proof. + proc. + conseq (: _ ==> x = 5 | e1: p1 | e2: p2 ). + + auto. + + auto. admit. + + admit. + + admit. + qed. + +lemma l_f2 (_x: int): +hoare [M.f2 : p9 ==> p4 | e1: p2 | e2:p1 | e3: p6 ]. + proof. + proc. + if. + + call (: p5 ==> p3 | e1 : p9 | e2: p7). + proc. + wp. + admit. + + admit. + + call (l_f1 _x);auto. admit. + qed. diff --git a/src/ecAst.ml b/src/ecAst.ml index 2b30c3568..a908c1883 100644 --- a/src/ecAst.ml +++ b/src/ecAst.ml @@ -115,7 +115,7 @@ and instr_node = | Sif of expr * stmt * stmt | Swhile of expr * stmt | Smatch of expr * ((EcIdent.t * ty) list * stmt) list - | Sassert of expr + | Sraise of EcPath.path * expr list | Sabstract of EcIdent.t and stmt = { @@ -230,18 +230,22 @@ and equivS = { es_sr : stmt; es_po : form; } +and post = (EcPath.path * form) list + and sHoareF = { hf_pr : form; hf_f : EcPath.xpath; hf_po : form; + hf_poe : post } and sHoareS = { - hs_m : memenv; - hs_pr : form; - hs_s : stmt; - hs_po : form; } - + hs_m : memenv; + hs_pr : form; + hs_s : stmt; + hs_po : form; + hs_poe : post +} and eHoareF = { ehf_pr : form; @@ -615,13 +619,25 @@ let b_hash (bs : bindings) = Why3.Hashcons.combine_list b1_hash 0 bs (*-------------------------------------------------------------------- *) + +let post_equal (e1, f1) (e2,f2) = + EcPath.p_equal e1 e2 && + f_equal f1 f2 + +let posts_equal posts1 posts2 = + List.equal post_equal posts1 posts2 + +(*-------------------------------------------------------------------- *) + let hf_equal hf1 hf2 = f_equal hf1.hf_pr hf2.hf_pr + && posts_equal hf1.hf_poe hf2.hf_poe && f_equal hf1.hf_po hf2.hf_po && EcPath.x_equal hf1.hf_f hf2.hf_f let hs_equal hs1 hs2 = f_equal hs1.hs_pr hs2.hs_pr + && posts_equal hs1.hs_poe hs2.hs_poe && f_equal hs1.hs_po hs2.hs_po && s_equal hs1.hs_s hs2.hs_s && me_equal hs1.hs_m hs2.hs_m @@ -680,14 +696,27 @@ let pr_equal pr1 pr2 = && f_equal pr1.pr_event pr2.pr_event && f_equal pr1.pr_args pr2.pr_args +(*-------------------------------------------------------------------- *) + +let post_hash (e, f) = +Why3.Hashcons.combine + (EcPath.p_hash e) + (f_hash f) + +let posts_hash posts = + Why3.Hashcons.combine_list post_hash 0 posts + (* -------------------------------------------------------------------- *) let hf_hash hf = Why3.Hashcons.combine2 - (f_hash hf.hf_pr) (f_hash hf.hf_po) (EcPath.x_hash hf.hf_f) + (f_hash hf.hf_pr) + (Why3.Hashcons.combine (f_hash hf.hf_po) (posts_hash hf.hf_poe)) + (EcPath.x_hash hf.hf_f) let hs_hash hs = Why3.Hashcons.combine3 - (f_hash hs.hs_pr) (f_hash hs.hs_po) + (f_hash hs.hs_pr) + (Why3.Hashcons.combine (f_hash hs.hs_po) (posts_hash hs.hs_poe)) (s_hash hs.hs_s) (me_hash hs.hs_m) @@ -1036,6 +1065,11 @@ module Hsform = Why3.Hashcons.Make (struct let fv_mlr = Sid.add mleft (Sid.singleton mright) + let posts_fv init posts = + List.fold + (fun acc (_,f) -> fv_union (f_fv f) acc) + init posts + let fv_node f = let union ex nodes = List.fold_left (fun s a -> fv_union s (ex a)) Mid.empty nodes @@ -1063,11 +1097,13 @@ module Hsform = Why3.Hashcons.Make (struct fv_union (f_fv f1) fv2 | FhoareF hf -> - let fv = fv_union (f_fv hf.hf_pr) (f_fv hf.hf_po) in + let fv = f_fv hf.hf_po in + let fv = fv_union (f_fv hf.hf_pr) (posts_fv fv hf.hf_poe) in EcPath.x_fv (Mid.remove mhr fv) hf.hf_f | FhoareS hs -> - let fv = fv_union (f_fv hs.hs_pr) (f_fv hs.hs_po) in + let fv = f_fv hs.hs_po in + let fv = fv_union (f_fv hs.hs_pr) (posts_fv fv hs.hs_poe) in fv_union (s_fv hs.hs_s) (Mid.remove (fst hs.hs_m) fv) | FeHoareF hf -> @@ -1162,8 +1198,9 @@ module Hinstr = Why3.Hashcons.Make (struct in List.all2 forbs bs1 bs2 && s_equal s1 s2 in e_equal e1 e2 && List.all2 forb b1 b2 - | Sassert e1, Sassert e2 -> - (e_equal e1 e2) + | Sraise (e1, es1), Sraise (e2, es2) -> + (EcPath.p_equal e1 e2) + && (List.all2 e_equal es1 es2) | Sabstract id1, Sabstract id2 -> EcIdent.id_equal id1 id2 @@ -1201,7 +1238,10 @@ module Hinstr = Why3.Hashcons.Make (struct in Why3.Hashcons.combine_list forbs (s_hash s) bds in Why3.Hashcons.combine_list forb (e_hash e) b - | Sassert e -> e_hash e + | Sraise (e, tys) -> + Why3.Hashcons.combine_list e_hash + (EcPath.p_hash e) + tys | Sabstract id -> EcIdent.id_hash id @@ -1235,8 +1275,10 @@ module Hinstr = Why3.Hashcons.Make (struct (fun s b -> EcIdent.fv_union s (forb b)) (e_fv e) b - | Sassert e -> - e_fv e + | Sraise (_, args) -> + List.fold_left + (fun s a -> EcIdent.fv_union s (e_fv a)) + Mid.empty args | Sabstract id -> EcIdent.fv_singleton id diff --git a/src/ecAst.mli b/src/ecAst.mli index 53c542bdc..a6d3074f8 100644 --- a/src/ecAst.mli +++ b/src/ecAst.mli @@ -111,7 +111,7 @@ and instr_node = | Sif of expr * stmt * stmt | Swhile of expr * stmt | Smatch of expr * ((EcIdent.t * ty) list * stmt) list - | Sassert of expr + | Sraise of EcPath.path * expr list | Sabstract of EcIdent.t and stmt = private { @@ -225,17 +225,22 @@ and equivS = { es_sr : stmt; es_po : form; } +and post = (EcPath.path * form) list + and sHoareF = { hf_pr : form; hf_f : EcPath.xpath; hf_po : form; + hf_poe : post; } and sHoareS = { - hs_m : memenv; - hs_pr : form; - hs_s : stmt; - hs_po : form; } + hs_m : memenv; + hs_pr : form; + hs_s : stmt; + hs_po : form; + hs_poe : post; +} and eHoareF = { diff --git a/src/ecCallbyValue.ml b/src/ecCallbyValue.ml index aee423acb..4deb99b3f 100644 --- a/src/ecCallbyValue.ml +++ b/src/ecCallbyValue.ml @@ -222,7 +222,7 @@ and try_reduce_record_projection try if not ( - st.st_ri.iota + st.st_ri.iota && EcEnv.Op.is_projection st.st_env p && not (Args.isempty args) ) then raise Bailout; @@ -483,7 +483,8 @@ and cbv (st : state) (s : subst) (f : form) (args : args) : form = let hf_pr = norm st s hf.hf_pr in let hf_po = norm st s hf.hf_po in let hf_f = norm_xfun st s hf.hf_f in - f_hoareF_r { hf_pr; hf_f; hf_po } + let hf_poe = List.map (fun (e,f) -> e, norm st s f) hf.hf_poe in + f_hoareF_r { hf_pr; hf_f; hf_po; hf_poe} | FhoareS hs -> assert (Args.isempty args); @@ -492,7 +493,8 @@ and cbv (st : state) (s : subst) (f : form) (args : args) : form = let hs_po = norm st s hs.hs_po in let hs_s = norm_stmt s hs.hs_s in let hs_m = norm_me s hs.hs_m in - f_hoareS_r { hs_pr; hs_po; hs_s; hs_m } + let hs_poe = List.map (fun (e,f) -> e, norm st s f) hs.hs_poe in + f_hoareS_r { hs_pr; hs_po; hs_s; hs_m; hs_poe} | FeHoareF hf -> assert (Args.isempty args); diff --git a/src/ecCommands.ml b/src/ecCommands.ml index f2da6001d..147c08377 100644 --- a/src/ecCommands.ml +++ b/src/ecCommands.ml @@ -262,17 +262,17 @@ module HiPrinting = struct let pr_axioms (fmt : Format.formatter) (env : EcEnv.env) = let ax = EcEnv.Ax.all ~check:(fun _ ax -> EcDecl.is_axiom ax.ax_kind) env in let ppe0 = EcPrinting.PPEnv.ofenv env in - EcPrinting.pp_by_theory ppe0 (EcPrinting.pp_axiom) fmt ax + EcPrinting.pp_by_theory ppe0 (EcPrinting.pp_axiom) fmt ax let pr_hint_solve (fmt : Format.formatter) (env : EcEnv.env) = let hint_solve = EcEnv.Auto.all env in let hint_solve = List.map (fun (p, mode) -> let ax = EcEnv.Ax.by_path p env in (p, (ax, mode)) - ) hint_solve in - + ) hint_solve in + let ppe = EcPrinting.PPEnv.ofenv env in - + let pp_hint_solve ppe fmt = (fun (p, (ax, mode)) -> let mode = match mode with @@ -281,8 +281,8 @@ module HiPrinting = struct Format.fprintf fmt "%a %s" (EcPrinting.pp_axiom ppe) (p, ax) mode ) in - - EcPrinting.pp_by_theory ppe pp_hint_solve fmt hint_solve + + EcPrinting.pp_by_theory ppe pp_hint_solve fmt hint_solve (* ------------------------------------------------------------------ *) let pr_hint_rewrite (fmt : Format.formatter) (env : EcEnv.env) = @@ -302,13 +302,13 @@ module HiPrinting = struct if List.is_empty elems then Format.fprintf fmt "%s (empty)@." (EcPath.basename p) else - Format.fprintf fmt "@[%s = @\n%a@]@\n" (EcPath.basename p) - (EcPrinting.pp_list "@\n" (fun fmt p -> + Format.fprintf fmt "@[%s = @\n%a@]@\n" (EcPath.basename p) + (EcPrinting.pp_list "@\n" (fun fmt p -> Format.fprintf fmt "%a" pp_path p)) (EcPath.Sp.ntr_elements sp) ) in - + EcPrinting.pp_by_theory ppe pp_hint_rewrite fmt hint_rewrite (* ------------------------------------------------------------------ *) @@ -317,17 +317,17 @@ module HiPrinting = struct let (hint_simplify: (EcEnv.Reduction.topsym * rule list) list) = EcEnv.Reduction.all env in - let hint_simplify = List.filter_map (fun (ts, rl) -> + let hint_simplify = List.filter_map (fun (ts, rl) -> match ts with - | `Path p -> Some (p, rl) + | `Path p -> Some (p, rl) | _ -> None - ) hint_simplify + ) hint_simplify in - + let ppe = EcPrinting.PPEnv.ofenv env in - + let pp_hint_simplify ppe fmt = (fun (p, (rls : rule list)) -> - Format.fprintf fmt "@[%s:@\n%a@]" (EcPath.basename p) + Format.fprintf fmt "@[%s:@\n%a@]" (EcPath.basename p) (EcPrinting.pp_list "@\n" (fun fmt rl -> begin match rl.rl_cond with | [] -> Format.fprintf fmt "Conditions: None@\n" @@ -340,7 +340,7 @@ module HiPrinting = struct rls ) in - + EcPrinting.pp_by_theory ppe pp_hint_simplify fmt hint_simplify end @@ -454,6 +454,15 @@ and process_operator (scope : EcScope.scope) (pop : poperator located) = List.iter (fun s -> EcScope.notify scope `Info "added axiom: `%s'" s) axs; scope +(* -------------------------------------------------------------------- *) +and process_exception (scope : EcScope.scope) (ed : pexception_decl located) = + EcScope.check_state `InTop "exception" scope; + let e, scope = EcScope.Except.add scope ed in + let ppe = EcPrinting.PPEnv.ofenv (EcScope.env scope) in + EcScope.notify scope `Info "added exception %s %a" + (unloc ed.pl_desc.pe_name) (EcPrinting.pp_added_except ppe) e; + scope + (* -------------------------------------------------------------------- *) and process_procop (scope : EcScope.scope) (pop : pprocop located) = EcScope.check_state `InTop "operator" scope; @@ -557,7 +566,7 @@ and process_th_require1 ld scope (nm, (sysname, thname), io) = try_finally (fun () -> let commands = EcIo.parseall (EcIo.from_file filename) in let commands = - List.fold_left + List.fold_left (fun scope g -> process_internal subld scope g.gl_action) iscope commands in commands) @@ -761,6 +770,7 @@ and process (ld : Loader.loader) (scope : EcScope.scope) g = | Gmodule m -> `Fct (fun scope -> process_module scope m) | Ginterface i -> `Fct (fun scope -> process_interface scope i) | Goperator o -> `Fct (fun scope -> process_operator scope (mk_loc loc o)) + | Gexception e -> `Fct (fun scope -> process_exception scope (mk_loc loc e)) | Gprocop o -> `Fct (fun scope -> process_procop scope (mk_loc loc o)) | Gpredicate p -> `Fct (fun scope -> process_predicate scope (mk_loc loc p)) | Gnotation n -> `Fct (fun scope -> process_notation scope (mk_loc loc n)) diff --git a/src/ecCoreFol.ml b/src/ecCoreFol.ml index 962125360..8c1b9bc6b 100644 --- a/src/ecCoreFol.ml +++ b/src/ecCoreFol.ml @@ -274,11 +274,11 @@ let f_eqs fs1 fs2 = let f_hoareS_r hs = mk_form (FhoareS hs) tbool let f_hoareF_r hf = mk_form (FhoareF hf) tbool -let f_hoareS hs_m hs_pr hs_s hs_po = - f_hoareS_r { hs_m; hs_pr; hs_s; hs_po; } +let f_hoareS hs_m hs_pr hs_s hs_po hs_poe = + f_hoareS_r { hs_m; hs_pr; hs_s; hs_po; hs_poe} -let f_hoareF hf_pr hf_f hf_po = - f_hoareF_r { hf_pr; hf_f; hf_po; } +let f_hoareF hf_pr hf_f hf_po hf_poe = + f_hoareF_r { hf_pr; hf_f; hf_po; hf_poe} (* -------------------------------------------------------------------- *) let f_eHoareS_r hs = mk_form (FeHoareS hs) tbool @@ -629,8 +629,8 @@ let decompose_exists ?(bound : int option) (f : form) = decompose_binder ?bound ~quantif:Lexists f let decompose_lambda ?(bound : int option) (f : form) = - decompose_binder ?bound ~quantif:Llambda f - + decompose_binder ?bound ~quantif:Llambda f + (* -------------------------------------------------------------------- *) let destr_binder ?(bound : int option) ~quantif:quantif (f : form) = let bds, f = decompose_binder ?bound ~quantif f in @@ -641,10 +641,10 @@ let destr_binder ?(bound : int option) ~quantif:quantif (f : form) = let destr_forall ?(bound : int option) (f : form) = destr_binder ?bound ~quantif:Lforall f - + let destr_exists ?(bound : int option) (f : form) = destr_binder ?bound ~quantif:Lexists f - + let destr_lambda ?(bound : int option) (f : form) = destr_binder ?bound ~quantif:Llambda f @@ -659,10 +659,10 @@ let destr_forall1 (f : form) = let destr_exists1 (f : form) = destr_binder1 ~quantif:Lexists f - + let destr_lambda1 (f : form) = destr_binder1 ~quantif:Llambda f - + (* -------------------------------------------------------------------- *) let destr_let f = match f.f_node with diff --git a/src/ecCoreFol.mli b/src/ecCoreFol.mli index 07f61851d..09d47c8ba 100644 --- a/src/ecCoreFol.mli +++ b/src/ecCoreFol.mli @@ -117,8 +117,8 @@ val f_forall_mems : (EcIdent.t * memtype) list -> form -> form val f_hoareF_r : sHoareF -> form val f_hoareS_r : sHoareS -> form -val f_hoareF : form -> xpath -> form -> form -val f_hoareS : memenv -> form -> stmt -> form -> form +val f_hoareF : form -> xpath -> form -> (EcPath.path * form) list -> form +val f_hoareS : memenv -> form -> stmt -> form -> (EcPath.path * form) list -> form (* soft-constructors - expected hoare *) val f_eHoareF_r : eHoareF -> form diff --git a/src/ecCoreModules.ml b/src/ecCoreModules.ml index 6f96118ab..71718a175 100644 --- a/src/ecCoreModules.ml +++ b/src/ecCoreModules.ml @@ -93,7 +93,7 @@ let i_call (lv, m, tys) = mk_instr (Scall (lv, m, tys)) let i_if (c, s1, s2) = mk_instr (Sif (c, s1, s2)) let i_while (c, s) = mk_instr (Swhile (c, s)) let i_match (e, b) = mk_instr (Smatch (e, b)) -let i_assert e = mk_instr (Sassert e) +let i_raise (e,es) = mk_instr (Sraise (e,es)) let i_abstract id = mk_instr (Sabstract id) let s_seq s1 s2 = stmt (s1.s_node @ s2.s_node) @@ -105,7 +105,7 @@ let s_call arg = stmt [i_call arg] let s_if arg = stmt [i_if arg] let s_while arg = stmt [i_while arg] let s_match arg = stmt [i_match arg] -let s_assert arg = stmt [i_assert arg] +let s_raise arg = stmt [i_raise arg] let s_abstract arg = stmt [i_abstract arg] (* -------------------------------------------------------------------- *) @@ -133,8 +133,8 @@ let get_match = function | { i_node = Smatch (e, b) } -> Some (e, b) | _ -> None -let get_assert = function - | { i_node = Sassert e } -> Some e +let get_raise = function + | { i_node = Sraise(e,es) } -> Some (e,es) | _ -> raise Not_found (* -------------------------------------------------------------------- *) @@ -147,7 +147,7 @@ let destr_call = _destr_of_get get_call let destr_if = _destr_of_get get_if let destr_while = _destr_of_get get_while let destr_match = _destr_of_get get_match -let destr_assert = _destr_of_get get_assert +let destr_raise = _destr_of_get get_raise (* -------------------------------------------------------------------- *) let _is_of_get (get : instr -> 'a option) (i : instr) = @@ -159,7 +159,7 @@ let is_call = _is_of_get get_call let is_if = _is_of_get get_if let is_while = _is_of_get get_while let is_match = _is_of_get get_match -let is_assert = _is_of_get get_assert +let is_raise = _is_of_get get_raise (* -------------------------------------------------------------------- *) module Uninit = struct (* FIXME: generalize this for use in ecPV *) @@ -223,8 +223,9 @@ and i_get_uninit_read (w : Ssym.t) (i : instr) = let ws, rs = List.split wrs in (Ssym.union w (Ssym.big_inter ws), Ssym.big_union (r :: rs)) - | Sassert e -> - (w, Ssym.diff (Uninit.e_pv e) w) + | Sraise (_, args) -> + let r1 = Ssym.diff (Ssym.big_union (List.map (Uninit.e_pv) args)) w in + (w, r1) | Sabstract (_ : EcIdent.t) -> (w, Ssym.empty) diff --git a/src/ecCoreModules.mli b/src/ecCoreModules.mli index 1b84c0df2..5227b4dcc 100644 --- a/src/ecCoreModules.mli +++ b/src/ecCoreModules.mli @@ -42,7 +42,7 @@ val i_call : lvalue option * xpath * expr list -> instr val i_if : expr * stmt * stmt -> instr val i_while : expr * stmt -> instr val i_match : expr * ((EcIdent.t * ty) list * stmt) list -> instr -val i_assert : expr -> instr +val i_raise : EcPath.path * expr list -> instr val i_abstract : EcIdent.t -> instr val s_asgn : lvalue * expr -> stmt @@ -51,7 +51,7 @@ val s_call : lvalue option * xpath * expr list -> stmt val s_if : expr * stmt * stmt -> stmt val s_while : expr * stmt -> stmt val s_match : expr * ((EcIdent.t * ty) list * stmt) list -> stmt -val s_assert : expr -> stmt +val s_raise : EcPath.path * expr list -> stmt val s_abstract : EcIdent.t -> stmt val s_seq : stmt -> stmt -> stmt val s_empty : stmt @@ -66,7 +66,7 @@ val destr_call : instr -> lvalue option * xpath * expr list val destr_if : instr -> expr * stmt * stmt val destr_while : instr -> expr * stmt val destr_match : instr -> expr * ((EcIdent.t * ty) list * stmt) list -val destr_assert : instr -> expr +val destr_raise : instr -> EcPath.path * expr list val is_asgn : instr -> bool val is_rnd : instr -> bool @@ -74,7 +74,7 @@ val is_call : instr -> bool val is_if : instr -> bool val is_while : instr -> bool val is_match : instr -> bool -val is_assert : instr -> bool +val is_raise : instr -> bool (* -------------------------------------------------------------------- *) val get_uninit_read : stmt -> Sx.t diff --git a/src/ecCorePrinting.ml b/src/ecCorePrinting.ml index d906a61e3..2ba5b69b5 100644 --- a/src/ecCorePrinting.ml +++ b/src/ecCorePrinting.ml @@ -61,7 +61,7 @@ module type PrinterAPI = sig val pp_tyvar : PPEnv.t -> ident pp val pp_tyunivar : PPEnv.t -> EcUid.uid pp val pp_path : path pp - + (* ------------------------------------------------------------------ *) val shorten_path : PPEnv.t -> (path -> qsymbol -> bool) -> path -> qsymbol * qsymbol option @@ -79,21 +79,22 @@ module type PrinterAPI = sig | `Glob of EcIdent.t * EcMemory.memory | `PVar of EcTypes.prog_var * EcMemory.memory ] - + val pp_vsubst : PPEnv.t -> vsubst pp (* ------------------------------------------------------------------ *) - val pp_typedecl : PPEnv.t -> (path * tydecl ) pp - val pp_opdecl : ?long:bool -> PPEnv.t -> (path * operator ) pp - val pp_added_op : PPEnv.t -> operator pp - val pp_axiom : ?long:bool -> PPEnv.t -> (path * axiom ) pp - val pp_theory : PPEnv.t -> (path * ctheory ) pp - val pp_modtype1 : PPEnv.t -> (module_type ) pp - val pp_modtype : PPEnv.t -> (module_type ) pp - val pp_modexp : PPEnv.t -> (mpath * module_expr ) pp - val pp_moditem : PPEnv.t -> (mpath * module_item ) pp - val pp_modsig : ?long:bool -> PPEnv.t -> (path * module_sig) pp - val pp_modsig_smpl : PPEnv.t -> (path * module_smpl_sig ) pp + val pp_typedecl : PPEnv.t -> (path * tydecl ) pp + val pp_opdecl : ?long:bool -> PPEnv.t -> (path * operator ) pp + val pp_added_op : PPEnv.t -> operator pp + val pp_added_except : PPEnv.t -> excep pp + val pp_axiom : ?long:bool -> PPEnv.t -> (path * axiom ) pp + val pp_theory : PPEnv.t -> (path * ctheory ) pp + val pp_modtype1 : PPEnv.t -> (module_type ) pp + val pp_modtype : PPEnv.t -> (module_type ) pp + val pp_modexp : PPEnv.t -> (mpath * module_expr ) pp + val pp_moditem : PPEnv.t -> (mpath * module_item ) pp + val pp_modsig : ?long:bool -> PPEnv.t -> (path * module_sig) pp + val pp_modsig_smpl : PPEnv.t -> (path * module_smpl_sig ) pp (* ------------------------------------------------------------------ *) val pp_hoareS : PPEnv.t -> ?prpo:prpo_display -> sHoareS pp @@ -113,10 +114,10 @@ module type PrinterAPI = sig val pp_goal : PPEnv.t -> prpo_display -> ppgoal pp (* ------------------------------------------------------------------ *) - val pp_by_theory : PPEnv.t -> (PPEnv.t -> (EcPath.path * 'a) pp) -> ((EcPath.path * 'a) list) pp + val pp_by_theory : PPEnv.t -> (PPEnv.t -> (EcPath.path * 'a) pp) -> ((EcPath.path * 'a) list) pp (* ------------------------------------------------------------------ *) - val pp_rule_pattern : PPEnv.t -> EcTheory.rule_pattern pp + val pp_rule_pattern : PPEnv.t -> EcTheory.rule_pattern pp (* ------------------------------------------------------------------ *) module ObjectInfo : sig diff --git a/src/ecCoreSubst.ml b/src/ecCoreSubst.ml index 7e0253c63..3091a9863 100644 --- a/src/ecCoreSubst.ml +++ b/src/ecCoreSubst.ml @@ -332,8 +332,9 @@ let rec s_subst_top (s : f_subst) : stmt -> stmt = i_match (e_subst e, List.Smart.map forb b) - | Sassert e -> - i_assert (e_subst e) + | Sraise (e,es) -> + let es' = List.Smart.map e_subst es in + i_raise (e, es') | Sabstract _ -> i @@ -455,14 +456,16 @@ module Fsubst = struct let s = f_rem_mem s mhr in let hf_pr = f_subst ~tx s hf.hf_pr in let hf_po = f_subst ~tx s hf.hf_po in - f_hoareF hf_pr hf_f hf_po + let hf_poe = List.map (fun (e,f) -> e, f_subst ~tx s f) hf.hf_poe in + f_hoareF hf_pr hf_f hf_po hf_poe | FhoareS hs -> let hs_s = s_subst s hs.hs_s in let s, hs_m = add_me_binding s hs.hs_m in let hs_pr = f_subst ~tx s hs.hs_pr in let hs_po = f_subst ~tx s hs.hs_po in - f_hoareS hs_m hs_pr hs_s hs_po + let hs_poe = List.map (fun (e,f) -> e, f_subst ~tx s f) hs.hs_poe in + f_hoareS hs_m hs_pr hs_s hs_po hs_poe | FeHoareF hf -> let hf_f = x_subst s hf.ehf_f in diff --git a/src/ecDecl.ml b/src/ecDecl.ml index c22745a78..9008ff64a 100644 --- a/src/ecDecl.ml +++ b/src/ecDecl.ml @@ -145,6 +145,13 @@ type axiom = { let is_axiom (x : axiom_kind) = match x with `Axiom _ -> true | _ -> false let is_lemma (x : axiom_kind) = match x with `Lemma -> true | _ -> false +(* -------------------------------------------------------------------- *) + +type excep = { + e_typargs : ty_params; + e_loca : locality; + } + (* -------------------------------------------------------------------- *) let op_ty op = op.op_ty @@ -210,6 +217,13 @@ let mk_op ?clinline ?unfold ~opaque tparams ty body lc = let kind = OB_oper body in gen_op ?clinline ?unfold ~opaque tparams ty kind lc +let gen_excep tparams lc = { + e_typargs = tparams; + e_loca = lc; +} + +let mk_except tparams lc = gen_excep tparams lc + let mk_abbrev ?(ponly = false) tparams xs (codom, body) lc = let kind = { ont_args = xs; diff --git a/src/ecDecl.mli b/src/ecDecl.mli index 6bde7e114..4475a8449 100644 --- a/src/ecDecl.mli +++ b/src/ecDecl.mli @@ -154,6 +154,14 @@ val axiomatized_op : -> locality -> axiom +(* -------------------------------------------------------------------- *) +type excep = { + e_typargs : ty_params; + e_loca : locality; +} + +val mk_except : ty_params -> locality -> excep + (* -------------------------------------------------------------------- *) type typeclass = { tc_prt : EcPath.path option; diff --git a/src/ecEnv.ml b/src/ecEnv.ml index c663f5ba9..2895cb26f 100644 --- a/src/ecEnv.ml +++ b/src/ecEnv.ml @@ -87,6 +87,7 @@ type mc = { mc_modsigs : (ipath * top_module_sig) MMsym.t; mc_tydecls : (ipath * EcDecl.tydecl) MMsym.t; mc_operators : (ipath * EcDecl.operator) MMsym.t; + mc_except : (ipath * EcDecl.excep) MMsym.t; mc_axioms : (ipath * EcDecl.axiom) MMsym.t; mc_theories : (ipath * ctheory) MMsym.t; mc_typeclasses: (ipath * typeclass) MMsym.t; @@ -272,6 +273,7 @@ let empty_mc params = { mc_modsigs = MMsym.empty; mc_tydecls = MMsym.empty; mc_operators = MMsym.empty; + mc_except = MMsym.empty; mc_axioms = MMsym.empty; mc_theories = MMsym.empty; mc_variables = MMsym.empty; @@ -510,6 +512,7 @@ module MC = struct let _downpath_for_tydecl = _downpath_for_th let _downpath_for_modsig = _downpath_for_th let _downpath_for_operator = _downpath_for_th + let _downpath_for_except = _downpath_for_th let _downpath_for_axiom = _downpath_for_th let _downpath_for_typeclass = _downpath_for_th let _downpath_for_rwbase = _downpath_for_th @@ -706,6 +709,20 @@ module MC = struct let import_axiom p ax env = import (_up_axiom true) (IPPath p) ax env + (* -------------------------------------------------------------------- *) + let lookup_except qnx env = + match lookup (fun mc -> mc.mc_except) qnx env with + | None -> lookup_error (`QSymbol qnx) + | Some (p, (args, obj)) -> (_downpath_for_except env p args, obj) + + let _up_except candup mc x obj = + if not candup && MMsym.last x mc.mc_except <> None then + raise (DuplicatedBinding x); + { mc with mc_except = MMsym.add x obj mc.mc_except } + + let import_except p e env = + import (_up_except true) (IPPath p) e env + (* -------------------------------------------------------------------- *) let lookup_operator qnx env = match lookup (fun mc -> mc.mc_operators) qnx env with @@ -1081,6 +1098,9 @@ module MC = struct | Th_operator (xop, op) -> (add2mc _up_operator xop op mc, None) + | Th_exception (xop, e) -> + (add2mc _up_except xop e mc, None) + | Th_axiom (xax, ax) -> (add2mc _up_axiom xax ax mc, None) @@ -1180,6 +1200,9 @@ module MC = struct and bind_axiom x ax env = bind _up_axiom x ax env + and bind_except x e env = + bind _up_except x e env + and bind_operator x op env = bind _up_operator x op env @@ -1444,7 +1467,7 @@ module BaseRw = struct env_item = mkitem ~import (Th_addrw (p, l, lc)) :: env.env_item; } let all env = - List.filter_map (fun (ip, sp) -> + List.filter_map (fun (ip, sp) -> match ip with | IPPath p -> Some (p, sp) | _ -> None) @@ @@ -1500,10 +1523,10 @@ module Reduction = struct |> odfl [] (* FIXME: handle other cases, right now only used for print hint *) - let all (env : env) = - List.map (fun (ts, mr) -> + let all (env : env) = + List.map (fun (ts, mr) -> (ts, Lazy.force mr.ri_list)) - (Mrd.bindings env.env_redbase) + (Mrd.bindings env.env_redbase) end (* -------------------------------------------------------------------- *) @@ -1552,7 +1575,7 @@ module Auto = struct let getall (bases : symbol list) (env : env) : atbase0 list = let dbs = List.map (fun base -> get_core ~base env) bases in - let dbs = + let dbs = List.fold_left (fun db mi -> Mint.union (fun _ sp1 sp2 -> Some (sp1 @ sp2)) db mi) Mint.empty dbs @@ -2749,6 +2772,36 @@ module Op = struct gen_all (fun mc -> mc.mc_operators) MC.lookup_operators ?check ?name env end +(* -------------------------------------------------------------------- *) +module Except = struct + type t = excep + + let by_path_opt (p : EcPath.path) (env : env) = + omap + check_not_suspended + (MC.by_path (fun mc -> mc.mc_except) (IPPath p) env) + + let by_path (p : EcPath.path) (env : env) = + match by_path_opt p env with + | None -> lookup_error (`Path p) + | Some obj -> obj + + let add (p : EcPath.path) (env : env) = + let obj = by_path p env in + MC.import_except p obj env + + let lookup qname (env : env) = + MC.lookup_except qname env + + let lookup_opt name env = + try_lf (fun () -> lookup name env) + + let bind ?(import = true) name e env = + let env = MC.bind_except name e env in + { env with env_item = mkitem ~import (Th_exception (name, e)) :: env.env_item } + +end + (* -------------------------------------------------------------------- *) module Ax = struct type t = axiom @@ -3055,6 +3108,9 @@ module Theory = struct | Th_operator (x, op) -> MC.import_operator (xpath x) op env + | Th_exception (x, e) -> + MC.import_except (xpath x) e env + | Th_axiom (x, ax) -> MC.import_axiom (xpath x) ax env diff --git a/src/ecEnv.mli b/src/ecEnv.mli index b5c51c9f9..0ee5c24cc 100644 --- a/src/ecEnv.mli +++ b/src/ecEnv.mli @@ -149,6 +149,16 @@ module Var : sig val bindall_pvglob : (EcSymbols.symbol * EcTypes.ty) list -> env -> env end +(* -------------------------------------------------------------------- *) +module Except : sig + type t = excep + val by_path : path -> env -> t + val by_path_opt : path -> env -> t option + val add : path -> env -> env + val lookup : qsymbol -> env -> path * t + val lookup_opt : qsymbol -> env -> (path * t) option + val bind : ?import:bool -> symbol -> t -> env -> env +end (* -------------------------------------------------------------------- *) module Ax : sig @@ -402,7 +412,7 @@ module BaseRw : sig val add : ?import:bool -> symbol -> is_local -> env -> env val addto : ?import:bool -> path -> path list -> is_local -> env -> env - val all : env -> (path * Sp.t) list + val all : env -> (path * Sp.t) list end (* -------------------------------------------------------------------- *) diff --git a/src/ecLexer.mll b/src/ecLexer.mll index 0e6f9a02e..a4b2982da 100644 --- a/src/ecLexer.mll +++ b/src/ecLexer.mll @@ -53,7 +53,7 @@ "match" , MATCH ; (* KW: prog *) "for" , FOR ; (* KW: prog *) "while" , WHILE ; (* KW: prog *) - "assert" , ASSERT ; (* KW: prog *) + "raise" , RAISE ; (* KW: prog *) "return" , RETURN ; (* KW: prog *) "res" , RES ; (* KW: prog *) "equiv" , EQUIV ; (* KW: prog *) @@ -188,6 +188,7 @@ "of" , OF ; (* KW: global *) "const" , CONST ; (* KW: global *) "op" , OP ; (* KW: global *) + "exception" , EXCEP ; (* KW: global *) "pred" , PRED ; (* KW: global *) "inductive" , INDUCTIVE ; (* KW: global *) "notation" , NOTATION ; (* KW: global *) diff --git a/src/ecLowPhlGoal.ml b/src/ecLowPhlGoal.ml index 79244af03..be6700703 100644 --- a/src/ecLowPhlGoal.ml +++ b/src/ecLowPhlGoal.ml @@ -95,7 +95,7 @@ let pf_first_call pe st = pf_first_gen "call" destr_call pe st let pf_first_if pe st = pf_first_gen "if" destr_if pe st let pf_first_match pe st = pf_first_gen "match" destr_match pe st let pf_first_while pe st = pf_first_gen "while" destr_while pe st -let pf_first_assert pe st = pf_first_gen "assert" destr_assert pe st +let pf_first_raise pe st = pf_first_gen "raise" destr_raise pe st (* -------------------------------------------------------------------- *) let pf_last_asgn pe st = pf_last_gen "asgn" destr_asgn pe st @@ -104,7 +104,7 @@ let pf_last_call pe st = pf_last_gen "call" destr_call pe st let pf_last_if pe st = pf_last_gen "if" destr_if pe st let pf_last_match pe st = pf_last_gen "match" destr_match pe st let pf_last_while pe st = pf_last_gen "while" destr_while pe st -let pf_last_assert pe st = pf_last_gen "assert" destr_assert pe st +let pf_last_raise pe st = pf_last_gen "raise" destr_raise pe st (* -------------------------------------------------------------------- *) let tc1_first_asgn tc st = pf_first_asgn !!tc st @@ -113,7 +113,7 @@ let tc1_first_call tc st = pf_first_call !!tc st let tc1_first_if tc st = pf_first_if !!tc st let tc1_first_match tc st = pf_first_match !!tc st let tc1_first_while tc st = pf_first_while !!tc st -let tc1_first_assert tc st = pf_first_assert !!tc st +let tc1_first_raise tc st = pf_first_raise !!tc st (* -------------------------------------------------------------------- *) let tc1_last_asgn tc st = pf_last_asgn !!tc st @@ -122,7 +122,7 @@ let tc1_last_call tc st = pf_last_call !!tc st let tc1_last_if tc st = pf_last_if !!tc st let tc1_last_match tc st = pf_last_match !!tc st let tc1_last_while tc st = pf_last_while !!tc st -let tc1_last_assert tc st = pf_last_assert !!tc st +let tc1_last_raise tc st = pf_last_raise !!tc st (* -------------------------------------------------------------------- *) (* TODO: use in change pos *) @@ -138,7 +138,7 @@ let pf_pos_last_call pe s = pf_pos_last_gen "call" is_call pe s let pf_pos_last_if pe s = pf_pos_last_gen "if" is_if pe s let pf_pos_last_match pe s = pf_pos_last_gen "match" is_match pe s let pf_pos_last_while pe s = pf_pos_last_gen "while" is_while pe s -let pf_pos_last_assert pe s = pf_pos_last_gen "assert" is_assert pe s +let pf_pos_last_raise pe s = pf_pos_last_gen "raise" is_raise pe s let tc1_pos_last_asgn tc s = pf_pos_last_asgn !!tc s @@ -147,7 +147,7 @@ let tc1_pos_last_call tc s = pf_pos_last_call !!tc s let tc1_pos_last_if tc s = pf_pos_last_if !!tc s let tc1_pos_last_match tc s = pf_pos_last_match !!tc s let tc1_pos_last_while tc s = pf_pos_last_while !!tc s -let tc1_pos_last_assert tc s = pf_pos_last_assert !!tc s +let tc1_pos_last_raise tc s = pf_pos_last_raise !!tc s (* -------------------------------------------------------------------- *) let pf_as_hoareF pe c = as_phl (`Hoare `Pred) (fun () -> destr_hoareF c) pe @@ -251,10 +251,22 @@ let tc1_get_post tc = | None -> tc_error_noXhl ~kinds:hlkinds_Xhl !!tc | Some f -> f +(* -------------------------------------------------------------------- *) +let get_poste f = + match f.f_node with + | FhoareF hf -> Some (hf.hf_poe ) + | FhoareS hs -> Some (hs.hs_poe ) + | _ -> None + +let tc1_get_poste tc = + match get_poste (FApi.tc1_goal tc) with + | None -> tc_error_noXhl ~kinds:hlkinds_Xhl !!tc + | Some f -> f + (* -------------------------------------------------------------------- *) let set_pre ~pre f = match f.f_node with - | FhoareF hf -> f_hoareF pre hf.hf_f hf.hf_po + | FhoareF hf -> f_hoareF pre hf.hf_f hf.hf_po hf.hf_poe | FhoareS hs -> f_hoareS_r { hs with hs_pr = pre } | FeHoareF hf -> f_eHoareF_r { hf with ehf_pr = pre } | FeHoareS hs -> f_eHoareS_r { hs with ehs_pr = pre } diff --git a/src/ecPV.ml b/src/ecPV.ml index 2126bb0f2..f58a4a2c7 100644 --- a/src/ecPV.ml +++ b/src/ecPV.ml @@ -129,7 +129,7 @@ module Mpv = struct | Sif (c, s1, s2) -> i_if (esubst c, ssubst s1, ssubst s2) | Swhile (e, stmt) -> i_while (esubst e, ssubst stmt) | Smatch (e, b) -> i_match (esubst e, List.Smart.map (snd_map ssubst) b) - | Sassert e -> i_assert (esubst e) + | Sraise (e,es) -> i_raise (e, List.map esubst es) | Sabstract _ -> i and issubst env (s : esubst) (is : instr list) = @@ -517,7 +517,7 @@ and i_write_r ?(except=Sx.empty) env w i = match i.i_node with | Sasgn (lp, _) -> lp_write_r env w lp | Srnd (lp, _) -> lp_write_r env w lp - | Sassert _ -> w + | Sraise _ -> w | Scall(lp,f,_) -> if Sx.mem f except then w else @@ -575,7 +575,8 @@ and i_read_r env r i = match i.i_node with | Sasgn (_lp, e) -> e_read_r env r e | Srnd (_lp, e) -> e_read_r env r e - | Sassert e -> e_read_r env r e + | Sraise (_e, es) -> + List.fold_left (e_read_r env) r es | Scall (_lp, f, es) -> let r = List.fold_left (e_read_r env) r es in @@ -1052,7 +1053,9 @@ and i_eqobs_in_refl env i eqo = let eqs = List.fold_left PV.union PV.empty eqs in add_eqs_refl env eqs e - | Sassert e -> add_eqs_refl env eqo e + | Sraise (_,es) -> + List.fold_left (add_eqs_refl env) eqo es + | Sabstract _ -> assert false and eqobs_inF_refl env f' eqo = diff --git a/src/ecParser.mly b/src/ecParser.mly index 81c38f90b..6ce0cc2f2 100644 --- a/src/ecParser.mly +++ b/src/ecParser.mly @@ -384,7 +384,7 @@ %token AMP %token APPLY %token AS -%token ASSERT +%token RAISE %token ASSUMPTION %token ASYNC %token AT @@ -442,6 +442,7 @@ %token EQUIV %token ETA %token EXACT +%token EXCEP %token EXFALSO %token EXIST %token EXIT @@ -1221,9 +1222,14 @@ hoare_bd_cmp : | EQ { EcAst.FHeq } | GE { EcAst.FHge } +epost(P): x=lident COLON post=form_r(P) + { let loc = EcLocation.make $startpos $endpos in + (EcLocation.mk_loc loc ([], x),post)} + hoare_body(P): - mp=loc(fident) COLON pre=form_r(P) LONGARROW post=form_r(P) - { PFhoareF (pre, mp, post) } + mp=loc(fident) COLON pre=form_r(P) LONGARROW post=form_r(P) + epost=splist(epost(P), PIPE) + { PFhoareF (pre, mp, post, epost) } ehoare_body(P): mp=loc(fident) COLON pre=form_r(P) LONGARROW @@ -1344,8 +1350,9 @@ base_instr: | f=loc(fident) LPAREN es=loc(plist0(expr, COMMA)) RPAREN { PScall (None, f, es) } -| ASSERT LPAREN c=expr RPAREN - { PSassert c } + | RAISE x=lident + {let loc = EcLocation.make $startpos $endpos in + PSraise (EcLocation.mk_loc loc ([], x), EcLocation.mk_loc loc []) } instr: | bi=base_instr SEMICOLON @@ -1752,6 +1759,7 @@ tyci_ax: | PROOF x=ident BY tg=tactic_core { (x, tg) } + (* -------------------------------------------------------------------- *) (* Operator definitions *) @@ -1858,6 +1866,14 @@ procop: | locality=locality PROC OP x=ident EQ f=loc(fident) { { ppo_name = x; ppo_target = f; ppo_locality = locality; } } +(* -------------------------------------------------------------------- *) +(* Exceptions *) +excep: +| locality=locality EXCEP x=oident + { { pe_name = x; + pe_typargs = None; + pe_locality = locality; } } + (* -------------------------------------------------------------------- *) (* Predicate definitions *) predicate: @@ -2507,17 +2523,36 @@ conseq: | UNDERSCORE LONGARROW f2=form { None, Some f2 } | f1=form LONGARROW f2=form { Some f1, Some f2 } +epost_xt: +| x=lident COLON f=form +{ let loc = EcLocation.make $startpos $endpos in + (EcLocation.mk_loc loc ([], x), f)} + +epostl: +| PIPE p=epost_xt { [ p ] } +| PIPE p=epost_xt a=epostl { p :: a } + +conseq_epost: +| empty { None } +| c=epostl { Some c } + conseq_xt: -| c=conseq { c, None } -| c=conseq COLON cmp=hoare_bd_cmp? bd=sform { c, Some (CQI_bd (cmp, bd)) } -| UNDERSCORE COLON cmp=hoare_bd_cmp? bd=sform { (None, None), Some (CQI_bd (cmp, bd)) } +| c=conseq d=conseq_epost + { (fst c, snd c, d), None } +| c=conseq COLON cmp=hoare_bd_cmp? bd=sform + { (fst c, snd c, None), Some (CQI_bd (cmp, bd)) } +| UNDERSCORE COLON cmp=hoare_bd_cmp? bd=sform + { (None, None, None), Some (CQI_bd (cmp, bd)) } +call_epost: +| empty { [] } +| c=epostl { c } call_info: -| f1=form LONGARROW f2=form { CI_spec (f1, f2) } -| f=form { CI_inv f } -| bad=form COMMA p=form { CI_upto (bad,p,None) } -| bad=form COMMA p=form COMMA q=form { CI_upto (bad,p,Some q) } +| f1=form LONGARROW f2=form poe=call_epost { CI_spec (f1, f2, poe) } +| f=form poe=call_epost { CI_inv (f, poe) } +| bad=form COMMA p=form { CI_upto (bad,p,None) } +| bad=form COMMA p=form COMMA q=form { CI_upto (bad,p,Some q) } tac_dir: | BACKS { Backs } @@ -3841,6 +3876,7 @@ global_action: | typeclass { Gtypeclass $1 } | tycinstance { Gtycinstance $1 } | operator { Goperator $1 } +| excep { Gexception $1 } | procop { Gprocop $1 } | predicate { Gpredicate $1 } | notation { Gnotation $1 } @@ -3929,6 +3965,11 @@ iplist1_r(X, S): %inline empty: | /**/ { () } +(* -------------------------------------------------------------------- *) +%inline splist(X, S): +| /* empty */ { [] } +| S xs=iplist1_r(X, S) { xs } + (* -------------------------------------------------------------------- *) __rlist1(X, S): (* left-recursive *) | x=X { [x] } diff --git a/src/ecParsetree.ml b/src/ecParsetree.ml index bc090f693..f607d2fa4 100644 --- a/src/ecParsetree.ml +++ b/src/ecParsetree.ml @@ -160,7 +160,7 @@ and pinstr_r = | PSif of pscond * pscond list * pstmt | PSwhile of pscond | PSmatch of pexpr * psmatch - | PSassert of pexpr + | PSraise of pgamepath * (pexpr list) located and psmatch = [ | `Full of (ppattern * pstmt) list @@ -202,7 +202,7 @@ and pformula_r = | PFlsless of pgamepath | PFscope of pqsymbol * pformula - | PFhoareF of pformula * pgamepath * pformula + | PFhoareF of pformula * pgamepath * pformula * (pgamepath * pformula) list | PFehoareF of pformula * pgamepath * pformula | PFequivF of pformula * (pgamepath * pgamepath) * pformula | PFeagerF of pformula * (pstmt * pgamepath * pgamepath * pstmt) * pformula @@ -429,6 +429,12 @@ and pprocop = { ppo_locality : locality; } +type pexception_decl = { + pe_name : psymbol; + pe_typargs : (psymbol * pqsymbol list) list option; + pe_locality : locality; + } + type ppred_def = | PPabstr of pty list | PPconcr of ptybindings * pformula @@ -531,9 +537,11 @@ type pipattern = and pspattern = unit +type pfel_spec_preds = (pgamepath * pformula) list + type call_info = - | CI_spec of (pformula * pformula) - | CI_inv of pformula + | CI_spec of (pformula * pformula * pfel_spec_preds) + | CI_inv of (pformula * pfel_spec_preds) | CI_upto of (pformula * pformula * pformula option) type p_app_xt_info = @@ -554,8 +562,6 @@ type psemrndpos = (bool * pcodepos1) doption type tac_dir = Backs | Fwds -type pfel_spec_preds = (pgamepath * pformula) list - (* -------------------------------------------------------------------- *) type pim_repeat_kind = | IM_R_Repeat of int option pair @@ -685,7 +691,9 @@ type deno_ppterm = (pformula option pair) gppterm type conseq_info = | CQI_bd of phoarecmp option * pformula -type conseq_ppterm = ((pformula option pair) * (conseq_info) option) gppterm +type conseq_contra = pformula option * pformula option * pfel_spec_preds option + +type conseq_ppterm = (conseq_contra * (conseq_info) option) gppterm (* -------------------------------------------------------------------- *) type sim_info = { @@ -1267,6 +1275,7 @@ type global_action = | Gmodule of pmodule_def_or_decl | Ginterface of pinterface | Goperator of poperator + | Gexception of pexception_decl | Gprocop of pprocop | Gpredicate of ppredicate | Gnotation of pnotation diff --git a/src/ecPath.ml b/src/ecPath.ml index 091840b6b..d703f4145 100644 --- a/src/ecPath.ml +++ b/src/ecPath.ml @@ -59,6 +59,14 @@ let rec p_ntr_compare (p1 : path) (p2 : path) = | n -> n end +let rec p_ntr_equal (p1 : path) (p2 : path) = + match p1.p_node, p2.p_node with + | Psymbol x1, Psymbol x2 -> + String.equal x1 x2 + | Pqname (p1, x1), Pqname (p2, x2) -> + p_ntr_equal p1 p2 && String.equal x1 x2 + | _, _ -> false + (* -------------------------------------------------------------------- *) module Mp = Path.M module Hp = Path.H @@ -268,7 +276,7 @@ let is_abstract (mp : mpath) = let is_concrete (mp : mpath) = match mp.m_top with `Concrete _ -> true | _ -> false - + let m_functor mp = let top = match mp.m_top with diff --git a/src/ecPath.mli b/src/ecPath.mli index ef2d2e8c0..1555e9d1a 100644 --- a/src/ecPath.mli +++ b/src/ecPath.mli @@ -22,6 +22,7 @@ val poappend : path -> path option -> path val p_equal : path -> path -> bool val p_compare : path -> path -> int +val p_ntr_equal : path -> path -> bool val p_ntr_compare : path -> path -> int val p_hash : path -> int diff --git a/src/ecPrinting.ml b/src/ecPrinting.ml index f7da6ecd6..88e462dfe 100644 --- a/src/ecPrinting.ml +++ b/src/ecPrinting.ml @@ -1536,9 +1536,8 @@ and pp_instr_for_form (ppe : PPEnv.t) fmt i = (pp_funname ppe) xp (pp_list ",@ " (pp_expr ppe)) args - | Sassert e -> - Format.fprintf fmt "assert %a;" - (pp_expr ppe) e + | Sraise _ -> + Format.fprintf fmt "raise TODO;" | Swhile (e, _) -> Format.fprintf fmt "while (%a) {...}" @@ -1928,10 +1927,19 @@ and pp_form_core_r let mepr, mepo = EcEnv.Fun.hoareF_memenv hf.hf_f ppe.PPEnv.ppe_env in let ppepr = PPEnv.create_and_push_mem ppe ~active:true mepr in let ppepo = PPEnv.create_and_push_mem ppe ~active:true mepo in - Format.fprintf fmt "hoare[@[@ %a :@ @[%a ==>@ %a@]@]]" + let ppepoe fmt l = + let aux fmt (e, f) = + Format.fprintf fmt " %a: %a" pp_path e (pp_form ppepo) f + in + match l with + | [] -> Format.fprintf fmt "" + | _ -> Format.fprintf fmt "| %a" (pp_list "|" aux) l + in + Format.fprintf fmt "hoare[@[@ %a :@ @[%a ==>@ %a@ %a]@]]" (pp_funname ppe) hf.hf_f (pp_form ppepr) hf.hf_pr (pp_form ppepo) hf.hf_po + ppepoe hf.hf_poe | FhoareS hs -> let ppe = PPEnv.push_mem ppe ~active:true hs.hs_m in @@ -2553,6 +2561,12 @@ let pp_added_op (ppe : PPEnv.t) fmt op = | ts -> Format.fprintf fmt "@[%a :@ %a.@]" (pp_tyvarannot ppe) ts (pp_type ppe) op.op_ty +(* -------------------------------------------------------------------- *) +let pp_added_except (ppe : PPEnv.t) fmt e = + let ppe = PPEnv.add_locals ppe (List.map fst e.e_typargs) in + match e.e_typargs with + | [] -> Format.fprintf fmt ": @[@]" + | ts -> Format.fprintf fmt "@[%a@]" (pp_tyvarannot ppe) ts (* -------------------------------------------------------------------- *) let pp_opname (ppe : PPEnv.t) fmt (p : EcPath.path) = @@ -2610,7 +2624,7 @@ let pp_axiom ?(long=false) (ppe : PPEnv.t) fmt (x, ax) = (* -------------------------------------------------------------------- *) type ppnode1 = [ | `Asgn of (EcModules.lvalue * EcTypes.expr) - | `Assert of (EcTypes.expr) + | `Raise of (EcPath.path * EcTypes.expr list) | `Call of (EcModules.lvalue option * P.xpath * EcTypes.expr list) | `Rnd of (EcModules.lvalue * EcTypes.expr) | `Abstract of EcIdent.t @@ -2633,7 +2647,7 @@ let at (ppe : PPEnv.t) n i = | Sasgn (lv, e) , 0 -> Some (`Asgn (lv, e) , `P, []) | Srnd (lv, e) , 0 -> Some (`Rnd (lv, e) , `P, []) | Scall (lv, f, es), 0 -> Some (`Call (lv, f, es), `P, []) - | Sassert e , 0 -> Some (`Assert e , `P, []) + | Sraise (e,es) , 0 -> Some (`Raise (e,es) , `P, []) | Sabstract id , 0 -> Some (`Abstract id , `P, []) | Swhile (e, s), 0 -> Some (`While e, `P, s.s_node) @@ -2721,8 +2735,8 @@ let pp_i_asgn (ppe : PPEnv.t) fmt (lv, e) = Format.fprintf fmt "%a <-@ %a" (pp_lvalue ppe) lv (pp_expr ppe) e -let pp_i_assert (ppe : PPEnv.t) fmt e = - Format.fprintf fmt "assert (%a)" (pp_expr ppe) e +let pp_i_raise (_ppe : PPEnv.t) fmt ((e,_es):EcPath.path * expr list) = + Format.fprintf fmt "raise %a" pp_path e let pp_i_call (ppe : PPEnv.t) fmt (lv, xp, args) = match lv with @@ -2769,7 +2783,7 @@ let pp_i_abstract (_ppe : PPEnv.t) fmt id = let c_ppnode1 ~width ppe (pp1 : ppnode1) = match pp1 with | `Asgn x -> c_split ~width (pp_i_asgn ppe) x - | `Assert x -> c_split ~width (pp_i_assert ppe) x + | `Raise x -> c_split ~width (pp_i_raise ppe) x | `Call x -> c_split ~width (pp_i_call ppe) x | `Rnd x -> c_split ~width (pp_i_rnd ppe) x | `Abstract x -> c_split ~width (pp_i_abstract ppe) x @@ -2939,6 +2953,14 @@ let pp_post (ppe : PPEnv.t) ?prpo fmt post = (omap (fun x -> x.prpo_po) prpo |> odfl false) fmt post None +(* -------------------------------------------------------------------- *) +let pp_poe (ppe : PPEnv.t) ?prpo (fmt: Format.formatter) poe = + List.iter (fun ((e:EcPath.path),f) -> + pp_prpo ppe (EcPath.basename e) + (omap (fun x -> x.prpo_po) prpo |> odfl false) + fmt f None + ) poe + (* -------------------------------------------------------------------- *) let pp_hoareF (ppe : PPEnv.t) ?prpo fmt hf = let mepr, mepo = EcEnv.Fun.hoareF_memenv hf.hf_f ppe.PPEnv.ppe_env in @@ -2947,22 +2969,25 @@ let pp_hoareF (ppe : PPEnv.t) ?prpo fmt hf = Format.fprintf fmt "%a@\n%!" (pp_pre ppepr ?prpo) hf.hf_pr; Format.fprintf fmt " %a@\n%!" (pp_funname ppe) hf.hf_f; - Format.fprintf fmt "@\n%a%!" (pp_post ppepo ?prpo) hf.hf_po + Format.fprintf fmt "@\n%a%!" (pp_post ppepo ?prpo) hf.hf_po; + Format.fprintf fmt "@\n%a%!" (pp_poe ppepo ?prpo) hf.hf_poe (* -------------------------------------------------------------------- *) let pp_hoareS (ppe : PPEnv.t) ?prpo fmt hs = let ppef = PPEnv.push_mem ppe ~active:true hs.hs_m in let ppnode = collect2_s ppef hs.hs_s.s_node [] in - let ppnode = c_ppnode ~width:ppe.PPEnv.ppe_width ppef ppnode - in - Format.fprintf fmt "Context : %a@\n%!" (pp_memtype ppe) (snd hs.hs_m); - Format.fprintf fmt "@\n%!"; - Format.fprintf fmt "%a%!" (pp_pre ppef ?prpo) hs.hs_pr; - Format.fprintf fmt "@\n%!"; - Format.fprintf fmt "%a" (pp_node `Left) ppnode; - Format.fprintf fmt "@\n%!"; - Format.fprintf fmt "%a%!" (pp_post ppef ?prpo) hs.hs_po + let ppnode = c_ppnode ~width:ppe.PPEnv.ppe_width ppef ppnode in + + Format.fprintf fmt "Context : %a@\n%!" (pp_memtype ppe) (snd hs.hs_m); + Format.fprintf fmt "@\n%!"; + Format.fprintf fmt "%a%!" (pp_pre ppef ?prpo) hs.hs_pr; + Format.fprintf fmt "@\n%!"; + Format.fprintf fmt "%a" (pp_node `Left) ppnode; + Format.fprintf fmt "@\n%!"; + Format.fprintf fmt "%a%!" (pp_post ppef ?prpo) hs.hs_po; + Format.fprintf fmt "@\n%!"; + Format.fprintf fmt "%a%!" (pp_poe ppef ?prpo) hs.hs_poe (* -------------------------------------------------------------------- *) let pp_eHoareF (ppe : PPEnv.t) ?prpo fmt hf = @@ -3348,9 +3373,8 @@ let rec pp_instr_r (ppe : PPEnv.t) fmt i = (pp_funname ppe) xp (pp_list ",@ " (pp_expr ppe)) args - | Sassert e -> - Format.fprintf fmt "@[assert %a@];" - (pp_expr ppe) e + | Sraise _ -> + Format.fprintf fmt "@[assert TODO@];" | Swhile (e, s) -> Format.fprintf fmt "@[while (@[%a@])%a@]" @@ -3394,7 +3418,7 @@ and pp_block ppe fmt s = | [] -> Format.fprintf fmt "{}" - | [ { i_node = (Sasgn _ | Srnd _ | Scall _ | Sassert _ | Sabstract _) } as i ] -> + | [ { i_node = (Sasgn _ | Srnd _ | Scall _ | Sraise _ | Sabstract _) } as i ] -> Format.fprintf fmt "@;<1 2>%a" (pp_instr ppe) i | _ -> @@ -3510,6 +3534,9 @@ let rec pp_theory ppe (fmt : Format.formatter) (path, cth) = | EcTheory.Th_operator (id, op) -> pp_opdecl ppe fmt (EcPath.pqname p id, op) + | EcTheory.Th_exception (_id, _e) -> + Format.fprintf fmt "exception ." + | EcTheory.Th_axiom (id, ax) -> pp_axiom ppe fmt (EcPath.pqname p id, ax) diff --git a/src/ecProcSem.ml b/src/ecProcSem.ml index 808ea8674..e643a4ba1 100644 --- a/src/ecProcSem.ml +++ b/src/ecProcSem.ml @@ -191,7 +191,7 @@ let rec translate_i (env : senv) (cont : senv -> mode * expr) (i : instr) = | Swhile _ | Smatch _ - | Sassert _ + | Sraise _ | Sabstract _ | Scall _ -> raise SemNotSupported; diff --git a/src/ecReduction.ml b/src/ecReduction.ml index 2458e51d7..2d367f4e1 100644 --- a/src/ecReduction.ml +++ b/src/ecReduction.ml @@ -242,8 +242,8 @@ end) = struct with E.NotConv -> false end - | Sassert a1, Sassert a2 -> - for_expr env alpha ~norm a1 a2 + | Sraise (_,es1), Sraise (_,es2) -> + List.all2 (for_expr env alpha ~norm) es1 es2 | Sabstract id1, Sabstract id2 -> EcIdent.id_equal id1 id2 diff --git a/src/ecScope.ml b/src/ecScope.ml index 95d7674e0..ec71d6512 100644 --- a/src/ecScope.ml +++ b/src/ecScope.ml @@ -1316,7 +1316,7 @@ module Op = struct tyop, List.rev !axs, scope - + (*-------------------------------------------------------------------------*) let add_opsem (scope : scope) (op : pprocop located) = let module Sem = EcProcSem in @@ -1427,7 +1427,9 @@ module Op = struct (f_app (f_op oppath [] opdecl.op_ty) (List.map (fun (x, ty) -> f_local x ty) locs) - sig_.fs_ret))) + sig_.fs_ret)) + [] + ) in let prax = EcDecl.{ @@ -1447,6 +1449,28 @@ module Op = struct scope end +(* -------------------------------------------------------------------- *) +module Except = struct + module TT = EcTyping + + let bind ?(import = true) (scope : scope) ((x, e) : _ * excep) = + assert (scope.sc_pr_uc = None); + let item = EcTheory.mkitem ~import (EcTheory.Th_exception (x, e)) in + { scope with sc_env = EcSection.add_item item scope.sc_env; } + + let add (scope : scope) (pe : pexception_decl located) = + assert (scope.sc_pr_uc = None); + let pe = pe.pl_desc and loc = pe.pl_loc in + let eenv = env scope in + let ue = TT.transtyvars eenv (loc, pe.pe_typargs) in + let lc = pe.pe_locality in + let tparams = EcUnify.UniEnv.tparams ue in + let e = EcDecl.mk_except tparams lc in + let scope = bind scope (unloc pe.pe_name, e) in + e, scope + +end + (* -------------------------------------------------------------------- *) module Pred = struct module TT = EcTyping @@ -1838,7 +1862,7 @@ module Cloning = struct | Some pt -> let t = { pt_core = pt; pt_intros = []; } in let t = { pl_loc = pt.pl_loc; pl_desc = Pby (Some [t]); } in - let t = { pt_core = t; pt_intros = []; } in + let t = { pt_core = t; pt_intros = []; } in let (x, ax) = axc.C.axc_axiom in let pucflags = { puc_smt = true; puc_local = false; } in @@ -1895,7 +1919,7 @@ module Cloning = struct | `Include -> scope) scope in - + if is_none thcl.pthc_local && oth.cth_loca = `Local then notify scope `Info "Theory `%s` has inherited `local` visibility. \ @@ -1984,14 +2008,14 @@ module Ty = struct let carrier = let ue = EcUnify.UniEnv.create None in transty tp_tydecl env ue subtype.pst_carrier in - + let pred = let x = EcIdent.create (fst subtype.pst_pred).pl_desc in let env = EcEnv.Var.bind_local x carrier env in let ue = EcUnify.UniEnv.create None in let pred = EcTyping.trans_prop env ue (snd subtype.pst_pred) in if not (EcUnify.UniEnv.closed ue) then - hierror ~loc:(snd subtype.pst_pred).pl_loc + hierror ~loc:(snd subtype.pst_pred).pl_loc "the predicate contains free type variables"; let uidmap = EcUnify.UniEnv.close ue in let fs = Tuni.subst uidmap in @@ -2013,12 +2037,12 @@ module Ty = struct ev_bynames = Msym.empty; ev_global = [ (None, Some [`Include, "prove"]) ] } } in - + let cname = Option.map unloc subtype.pst_cname in let npath = ofold ((^~) EcPath.pqname) (EcEnv.root env) cname in let cpath = EcPath.fromqsymbol ([EcCoreLib.i_top], "Subtype") in let theory = EcEnv.Theory.by_path ~mode:`Abstract cpath env in - + let renames = match subtype.pst_rename with | None -> [] @@ -2041,7 +2065,7 @@ module Ty = struct ) in let proofs = Cloning.replay_proofs scope `Check proofs in - + Ax.add_defer scope proofs (* ------------------------------------------------------------------ *) diff --git a/src/ecScope.mli b/src/ecScope.mli index bc3c2812d..8709187cf 100644 --- a/src/ecScope.mli +++ b/src/ecScope.mli @@ -98,6 +98,11 @@ module Op : sig val add_opsem : scope -> pprocop located -> scope end +(* -------------------------------------------------------------------- *) +module Except : sig + val add : scope -> pexception_decl located -> EcDecl.excep * scope +end + (* -------------------------------------------------------------------- *) module Pred : sig val add : scope -> ppredicate located -> EcDecl.operator * scope diff --git a/src/ecSection.ml b/src/ecSection.ml index 8dd953fe7..aeba1f7c3 100644 --- a/src/ecSection.ml +++ b/src/ecSection.ml @@ -18,6 +18,7 @@ type cbarg = [ | `Type of path | `Op of path | `Ax of path + | `Ex of path | `Module of mpath | `ModuleType of path | `Typeclass of path @@ -39,6 +40,7 @@ let pp_cbarg env fmt (who : cbarg) = | `Type p -> Format.fprintf fmt "type %a" (EcPrinting.pp_tyname ppe) p | `Op p -> Format.fprintf fmt "operator %a" (EcPrinting.pp_opname ppe) p | `Ax p -> Format.fprintf fmt "lemma/axiom %a" (EcPrinting.pp_axname ppe) p + | `Ex _p -> Format.fprintf fmt "TODO" | `Module mp -> let ppe = match mp.m_top with @@ -157,8 +159,8 @@ let rec on_instr (cb : cb) (i : instr)= on_lv cb lv; on_expr cb e - | Sassert e -> - on_expr cb e + | Sraise (_,es) -> + List.iter (on_expr cb) es | Scall (lv, f, args) -> lv |> oiter (on_lv cb); @@ -425,6 +427,10 @@ let on_axiom (cb : cb) (ax : axiom) = on_typarams cb ax.ax_tparams; on_form cb ax.ax_spec +(* -------------------------------------------------------------------- *) +let on_except (cb : cb) (e : excep) = + on_typarams cb e.e_typargs + (* -------------------------------------------------------------------- *) let on_modsig (cb:cb) (ms:module_sig) = List.iter (fun (_,mt) -> on_modty cb mt) ms.mis_params; @@ -507,6 +513,7 @@ let locality (env : EcEnv.env) (who : cbarg) = | `Type p -> (EcEnv. Ty.by_path p env).tyd_loca | `Op p -> (EcEnv. Op.by_path p env).op_loca | `Ax p -> (EcEnv. Ax.by_path p env).ax_loca + | `Ex _p -> assert false | `Typeclass p -> ((EcEnv.TypeClass.by_path p env).tc_loca :> locality) | `Module mp -> begin match EcEnv.Mod.by_mpath_opt mp env with @@ -1042,6 +1049,7 @@ let rec set_lc_item lc_override item = match item.ti_item with | Th_type (s,ty) -> Th_type (s, { ty with tyd_loca = set_lc lc_override ty.tyd_loca }) | Th_operator (s,op) -> Th_operator (s, { op with op_loca = set_lc lc_override op.op_loca }) + | Th_exception (s,e) -> Th_exception (s, { e with e_loca = set_lc lc_override e.e_loca }) | Th_axiom (s,ax) -> Th_axiom (s, { ax with ax_loca = set_lc lc_override ax.ax_loca }) | Th_modtype (s,ms) -> Th_modtype (s, { ms with tms_loca = set_lc lc_override ms.tms_loca }) | Th_module me -> Th_module { me with tme_loca = set_lc lc_override me.tme_loca } @@ -1101,6 +1109,7 @@ type can_depend = { d_ty : locality list; d_op : locality list; d_ax : locality list; + d_ex : locality list; d_sc : locality list; d_mod : locality list; d_modty : locality list; @@ -1111,6 +1120,7 @@ let cd_glob = { d_ty = [`Global]; d_op = [`Global]; d_ax = [`Global]; + d_ex = [`Global]; d_sc = [`Global]; d_mod = [`Global]; d_modty = [`Global]; @@ -1121,6 +1131,7 @@ let can_depend (cd : can_depend) = function | `Type _ -> cd.d_ty | `Op _ -> cd.d_op | `Ax _ -> cd.d_ax + | `Ex _ -> cd.d_ex | `Sc _ -> cd.d_sc | `Module _ -> cd.d_mod | `ModuleType _ -> cd.d_modty @@ -1153,6 +1164,7 @@ let check_tyd scenv prefix name tyd = d_ty = [`Declare; `Global]; d_op = [`Global]; d_ax = []; + d_ex = []; d_sc = []; d_mod = [`Global]; d_modty = []; @@ -1199,6 +1211,7 @@ let check_op scenv prefix name op = d_ty = [`Declare; `Global]; d_op = [`Declare; `Global]; d_ax = []; + d_ex = []; d_sc = []; d_mod = [`Declare; `Global]; d_modty = []; @@ -1211,6 +1224,7 @@ let check_op scenv prefix name op = d_ty = [`Declare; `Global]; d_op = [`Declare; `Global]; d_ax = []; + d_ex = []; d_sc = []; d_mod = [`Global]; d_modty = []; @@ -1230,6 +1244,7 @@ let check_ax (scenv : scenv) (prefix : path) (name : symbol) (ax : axiom) = d_ty = [`Declare; `Global]; d_op = [`Declare; `Global]; d_ax = []; + d_ex = []; d_sc = []; d_mod = [`Declare; `Global]; d_modty = [`Global]; @@ -1260,6 +1275,33 @@ let check_ax (scenv : scenv) (prefix : path) (name : symbol) (ax : axiom) = doit ax end +let check_except (scenv : scenv) (prefix : path) (name : symbol) (e : excep) = + let path = EcPath.pqname prefix name in + let from = e.e_loca, `Ex path in + let cd = { + d_ty = [`Declare; `Global]; + d_op = [`Declare; `Global]; + d_ax = []; + d_ex = []; + d_sc = []; + d_mod = [`Declare; `Global]; + d_modty = [`Global]; + d_tc = [`Global]; + } in + let doit = on_except (cb scenv from cd) in + + match e.e_loca with + | `Local -> + check_section scenv from + + | `Declare -> + check_section scenv from; + check_polymorph scenv from e.e_typargs; + doit e + + | `Global -> + if scenv.sc_insec then doit e + let check_modtype scenv prefix name ms = let path = pqname prefix name in let from = ((ms.tms_loca :> locality), `ModuleType path) in @@ -1282,6 +1324,7 @@ let check_module scenv prefix tme = { d_ty = [`Global]; d_op = [`Global]; d_ax = []; + d_ex = []; d_sc = []; d_mod = [`Global]; (* FIXME section: add local *) d_modty = [`Global]; @@ -1343,6 +1386,7 @@ let add_item_ ?(override_locality=None) (item : theory_item) (scenv:scenv) = | Th_type (s,tyd) -> EcEnv.Ty.bind ~import s tyd env | Th_operator (s,op) -> EcEnv.Op.bind ~import s op env | Th_axiom (s, ax) -> EcEnv.Ax.bind ~import s ax env + | Th_exception (s,e) -> EcEnv.Except.bind ~import s e env | Th_modtype (s, ms) -> EcEnv.ModTy.bind ~import s ms env | Th_module me -> EcEnv.Mod.bind ~import me.tme_expr.me_name me env | Th_typeclass(s,tc) -> EcEnv.TypeClass.bind ~import s tc env @@ -1371,6 +1415,7 @@ let rec generalize_th_item (to_gen : to_gen) (prefix : path) (th_item : theory_i match th_item.ti_item with | Th_type tydecl -> generalize_tydecl to_gen prefix tydecl | Th_operator opdecl -> generalize_opdecl to_gen prefix opdecl + | Th_exception _ -> assert false | Th_axiom ax -> generalize_axiom to_gen prefix ax | Th_modtype ms -> generalize_modtype to_gen ms | Th_module me -> generalize_module to_gen prefix me @@ -1442,7 +1487,7 @@ let genenv_of_scenv (scenv : scenv) : to_gen = ; tg_params = [] ; tg_binds = [] ; tg_subst = EcSubst.empty - ; tg_clear = empty_locals } + ; tg_clear = empty_locals } let generalize_lc_items scenv = let togen = @@ -1451,7 +1496,7 @@ let generalize_lc_items scenv = (EcEnv.root scenv.sc_env) (List.rev scenv.sc_items) in togen.tg_env - + (* -----------------------------------------------------------*) let import p scenv = { scenv with sc_env = EcEnv.Theory.import p scenv.sc_env } @@ -1475,6 +1520,7 @@ let check_item scenv item = match item.ti_item with | Th_type (s,tyd) -> check_tyd scenv prefix s tyd | Th_operator (s,op) -> check_op scenv prefix s op + | Th_exception (s,e) -> check_except scenv prefix s e | Th_axiom (s, ax) -> check_ax scenv prefix s ax | Th_modtype (s, ms) -> check_modtype scenv prefix s ms | Th_module me -> check_module scenv prefix me @@ -1520,6 +1566,7 @@ let add_decl_mod id mt scenv = d_ty = [`Global]; d_op = [`Global]; d_ax = []; + d_ex = []; d_sc = []; d_mod = [`Declare; `Global]; d_modty = [`Global]; diff --git a/src/ecSubst.ml b/src/ecSubst.ml index f667dcccf..67b39eeef 100644 --- a/src/ecSubst.ml +++ b/src/ecSubst.ml @@ -426,8 +426,9 @@ and subst_instr (s : subst) (i : instr) : instr = i_match (e, bs) - | Sassert e -> - i_assert (subst_expr s e) + | Sraise (e,es) -> + let es' = List.map (subst_expr s) es in + i_raise (e,es') | Sabstract _ -> i @@ -529,23 +530,21 @@ let rec subst_form (s : subst) (f : form) = let ty = subst_ty s f.f_ty in f_op p tys ty - | FhoareF { hf_pr; hf_f; hf_po } -> - let hf_pr, hf_po = - let s = add_memory s mhr mhr in - let hf_pr = subst_form s hf_pr in - let hf_po = subst_form s hf_po in - (hf_pr, hf_po) in - let hf_f = subst_xpath s hf_f in - f_hoareF hf_pr hf_f hf_po - - | FhoareS { hs_m; hs_pr; hs_s; hs_po } -> - let hs_m, (hs_pr, hs_po) = - let s, hs_m = subst_memtype s hs_m in - let hs_pr = subst_form s hs_pr in - let hs_po = subst_form s hs_po in - hs_m, (hs_pr, hs_po) in - let hs_s = subst_stmt s hs_s in - f_hoareS hs_m hs_pr hs_s hs_po + | FhoareF { hf_pr; hf_f; hf_po; hf_poe } -> + let s = add_memory s mhr mhr in + let hf_pr = subst_form s hf_pr in + let hf_po = subst_form s hf_po in + let hf_poe = List.map (fun (e,f) -> e, subst_form s f) hf_poe in + let hf_f = subst_xpath s hf_f in + f_hoareF hf_pr hf_f hf_po hf_poe + + | FhoareS { hs_m; hs_pr; hs_s; hs_po; hs_poe } -> + let s, hs_m = subst_memtype s hs_m in + let hs_pr = subst_form s hs_pr in + let hs_po = subst_form s hs_po in + let hs_poe = List.map (fun (e,f) -> e, subst_form s f) hs_poe in + let hs_s = subst_stmt s hs_s in + f_hoareS hs_m hs_pr hs_s hs_po hs_poe | FbdHoareF { bhf_pr; bhf_f; bhf_po; bhf_cmp; bhf_bd } -> let bhf_pr, bhf_po = @@ -970,6 +969,14 @@ let subst_op (s : subst) (op : operator) = op_clinline = op.op_clinline; op_unfold = op.op_unfold ; } +(* -------------------------------------------------------------------- *) +let subst_excep (s : subst) (e : excep) = + let _, tparams = fresh_tparams s e.e_typargs in + + { e_typargs = tparams ; + e_loca = e.e_loca ; + } + (* -------------------------------------------------------------------- *) let subst_ax (s : subst) (ax : axiom) = let s, tparams = fresh_tparams s ax.ax_tparams in @@ -1042,6 +1049,9 @@ let rec subst_theory_item_r (s : subst) (item : theory_item_r) = | Th_operator (x, op) -> Th_operator (x, subst_op s op) + | Th_exception (e, es) -> + Th_exception (e, subst_excep s es) + | Th_axiom (x, ax) -> Th_axiom (x, subst_ax s ax) diff --git a/src/ecThCloning.ml b/src/ecThCloning.ml index a2f24e593..b344535b1 100644 --- a/src/ecThCloning.ml +++ b/src/ecThCloning.ml @@ -304,7 +304,7 @@ end = struct (fun evc -> if Msym.mem x evc.evc_ops then clone_error oc.oc_env (CE_DupOverride (OVK_Operator, name)); - { evc with evc_ops = + { evc with evc_ops = Msym.add x (mk_loc lc opd :> xop_override located) evc.evc_ops }) nm evc @@ -437,6 +437,8 @@ end = struct let ovrd = (ovrd, mode) in nt_ovrd oc (proofs, evc) (loced (xdth @ prefix, x)) ovrd + | Th_exception _ -> (proofs, evc) + | Th_axiom (x, _) -> let axd = loced (thd @ prefix, x) in let name = (loced (xdth @ prefix, x)) in diff --git a/src/ecTheory.ml b/src/ecTheory.ml index ffe226d06..4f5c4f798 100644 --- a/src/ecTheory.ml +++ b/src/ecTheory.ml @@ -21,6 +21,7 @@ and theory_item = { and theory_item_r = | Th_type of (symbol * tydecl) | Th_operator of (symbol * operator) + | Th_exception of (symbol * excep) | Th_axiom of (symbol * axiom) | Th_modtype of (symbol * top_module_sig) | Th_module of top_module_expr diff --git a/src/ecTheory.mli b/src/ecTheory.mli index f246ee3f4..2b8551ed4 100644 --- a/src/ecTheory.mli +++ b/src/ecTheory.mli @@ -17,6 +17,7 @@ and theory_item = { and theory_item_r = | Th_type of (symbol * tydecl) | Th_operator of (symbol * operator) + | Th_exception of (symbol * excep) | Th_axiom of (symbol * axiom) | Th_modtype of (symbol * top_module_sig) | Th_module of top_module_expr diff --git a/src/ecTheoryReplay.ml b/src/ecTheoryReplay.ml index 3bdb4682e..4aba50cd8 100644 --- a/src/ecTheoryReplay.ml +++ b/src/ecTheoryReplay.ml @@ -523,7 +523,7 @@ and replay_opd (ove : _ ovrenv) (subst, ops, proofs, scope) (import, x, oopd) = ~opaque:optransparent ~clinline:(opmode <> `Alias) [] body.f_ty (Some (OP_Plain body)) refop.op_loca in (newop, body) - + in match opmode with | `Alias -> @@ -714,6 +714,14 @@ and replay_ntd (ove : _ ovrenv) (subst, ops, proofs, scope) (import, x, oont) = (subst, ops, proofs, scope) end + +(* -------------------------------------------------------------------- *) +and replay_excep + (ove : _ ovrenv) (subst, ops, proofs, scope) (import, name, excep) += + let scope = ove.ovre_hooks.hadd_item scope ~import (Th_exception (name, excep)) in + (subst, ops, proofs, scope) + (* -------------------------------------------------------------------- *) and replay_axd (ove : _ ovrenv) (subst, ops, proofs, scope) (import, x, ax) = let scenv = ove.ovre_hooks.henv scope in @@ -1029,6 +1037,9 @@ and replay1 (ove : _ ovrenv) (subst, ops, proofs, scope) (hidden, item) = | Th_operator (x, ({ op_kind = OB_nott _} as oont)) -> replay_ntd ove (subst, ops, proofs, scope) (import, x, oont) + | Th_exception (x, e) -> + replay_excep ove (subst, ops, proofs, scope) (import, x, e) + | Th_axiom (x, ax) -> replay_axd ove (subst, ops, proofs, scope) (import, x, ax) diff --git a/src/ecTyping.ml b/src/ecTyping.ml index 39490a68d..d30503406 100644 --- a/src/ecTyping.ml +++ b/src/ecTyping.ml @@ -157,6 +157,7 @@ type tyerror = | UnknownModName of qsymbol | UnknownTyModName of qsymbol | UnknownFunName of qsymbol +| UnknownExceptionName of qsymbol | UnknownModVar of qsymbol | UnknownMemName of symbol | InvalidFunAppl of funapp_error @@ -287,8 +288,8 @@ let (_i_inuse, s_inuse, se_inuse) = let map = List.fold_left (fun map -> s_inuse map |- snd) map bs in map - | Sassert e -> - se_inuse map e + | Sraise (_,args) -> + List.fold_left se_inuse map args | Sabstract _ -> assert false (* FIXME *) @@ -1470,6 +1471,13 @@ let trans_gamepath (env : EcEnv.env) gp = tyerror gp.pl_loc env (UnknownFunName (modsymb, funsymb)); EcPath.xpath mpath funsymb +let except_genpath env name = + let modsymb = List.map (unloc -| fst) (fst (unloc name)) + and symb = unloc (snd (unloc name)) in + match EcEnv.Except.lookup_opt (modsymb, symb) env with + | None -> tyerror name.pl_loc env (UnknownExceptionName (modsymb, symb)) + | Some (p,_) -> symb, p + (* -------------------------------------------------------------------- *) let trans_oracle (env : EcEnv.env) (m,f) = let msymbol = mk_loc (loc m) [m,None] in @@ -1685,6 +1693,10 @@ let f_or_mod_ident_loc : f_or_mod_ident -> EcLocation.t = function | FM_FunOrVar x -> loc x | FM_Mod x -> loc x +(* -------------------------------------------------------------------- *) +let check_unique_epost epost = + List.is_unique ~eq:(fun (e1,_) (e2,_) -> EcPath.p_ntr_equal e1 e2) epost + (* -------------------------------------------------------------------- *) let trans_restr_mem env (r_mem : pmod_restr_mem) = @@ -2751,10 +2763,43 @@ and transinstr [ i_match (e, branches) ] end - | PSassert pe -> - let e, ety = transexp env `InProc ue pe in - unify_or_fail env ue pe.pl_loc ~expct:tbool ety; - [ i_assert e ] + | PSraise (name, args) -> + let _,path = except_genpath env name in + let esig = (EcEnv.Except.by_path path env).e_typargs in + let aux (_,s) = + match Sp.elements s with + | [recp] -> recp + | _ -> assert false + in + let paths = List.map aux esig in + let indty = + List.map (fun path -> + oget (EcEnv.Ty.by_path_opt path env) + ) paths + in + let ind = + List.map (fun typdecl -> + let x = oget (EcDecl.tydecl_as_datatype typdecl) in + match x.tydt_ctors with + | [(_,t)] -> t + | _ -> assert false + ) indty + in + let typs = List.flatten ind in + + let args = unloc args in + let loc = name.pl_loc in + let args = + if List.length args <> List.length typs then + tyerror loc env (InvalidFunAppl FAE_WrongArgCount); + List.map2 + (fun a ty -> + let loc = a.pl_loc in + let a, aty = transexp env `InProc ue a in + unify_or_fail env ue loc ~expct:ty aty; a) args typs + in + + [i_raise (path, args)] (* -------------------------------------------------------------------- *) and trans_pv env { pl_desc = x; pl_loc = loc } = @@ -3447,7 +3492,7 @@ and trans_form_or_pattern env mode ?mv ?ps ue pf tt = unify_or_fail env ue event.pl_loc ~expct:tbool event'.f_ty; f_pr memid fpath (f_tuple args) event' - | PFhoareF (pre, gp, post) -> + | PFhoareF (pre, gp, post, eposts) -> if mode <> `Form then tyerror f.pl_loc env (NotAnExpression `Logic); @@ -3455,9 +3500,21 @@ and trans_form_or_pattern env mode ?mv ?ps ue pf tt = let penv, qenv = EcEnv.Fun.hoareF fpath env in let pre' = transf penv pre in let post' = transf qenv post in - unify_or_fail penv ue pre.pl_loc ~expct:tbool pre' .f_ty; - unify_or_fail qenv ue post.pl_loc ~expct:tbool post'.f_ty; - f_hoareF pre' fpath post' + let epost' = + List.map (fun (e,f) -> + let _,p = except_genpath env e in + p,transf penv f + ) eposts + in + + unify_or_fail penv ue pre.pl_loc ~expct:tbool pre'.f_ty; + unify_or_fail qenv ue post.pl_loc ~expct:tbool post'.f_ty; + List.iter + (fun (_,f) -> unify_or_fail qenv ue post.pl_loc ~expct:tbool f.f_ty) + epost'; + if check_unique_epost epost' then + f_hoareF pre' fpath post' epost' + else assert false | PFehoareF (pre, gp, post) -> if mode <> `Form then diff --git a/src/ecTyping.mli b/src/ecTyping.mli index 01b6883a6..57b1e3758 100644 --- a/src/ecTyping.mli +++ b/src/ecTyping.mli @@ -149,6 +149,7 @@ type tyerror = | UnknownModName of qsymbol | UnknownTyModName of qsymbol | UnknownFunName of qsymbol +| UnknownExceptionName of qsymbol | UnknownModVar of qsymbol | UnknownMemName of symbol | InvalidFunAppl of funapp_error @@ -256,6 +257,9 @@ val trans_pattern : env -> ptnmap -> EcUnify.unienv -> pformula -> EcFol.form val trans_memtype : env -> EcUnify.unienv -> pmemtype -> EcMemory.memtype +(* -------------------------------------------------------------------- *) +val check_unique_epost: EcAst.post -> bool + (* -------------------------------------------------------------------- *) val trans_restr_for_modty : env -> module_type -> pmod_restr option -> mty_mr @@ -267,6 +271,7 @@ val transmod : attop:bool -> env -> pmodule_def -> module_expr val trans_topmsymbol : env -> pmsymbol located -> mpath val trans_msymbol : env -> pmsymbol located -> mpath * module_smpl_sig val trans_gamepath : env -> pgamepath -> xpath +val except_genpath : env -> pgamepath -> symbol * path val trans_oracle : env -> psymbol * psymbol -> xpath val trans_restr_mem : env -> pmod_restr_mem -> mod_restr diff --git a/src/ecUnifyProc.ml b/src/ecUnifyProc.ml index e7ef218a7..f6284c88d 100644 --- a/src/ecUnifyProc.ml +++ b/src/ecUnifyProc.ml @@ -170,8 +170,24 @@ let rec unify_instrs env umap i1 i2 = unify_stmts env umap b1 b2 ) umap bs1 bs2 - | Sassert e1, Sassert e2 -> - unify_exprs env umap e1 e2 + | Sraise (_,args1), Sraise (_,args2) -> + let args1, args2 = + match args1, args2 with + | _, _ when List.length args1 = List.length args2 -> args1, args2 + | [al], _ -> begin + match al.e_node with + | Etuple args1 -> args1, args2 + | _ -> raise (UnificationError (UE_InstrNotInLockstep (i1, i2))) + end + | _, [ar] -> begin + match ar.e_node with + | Etuple args2 -> args1, args2 + | _ -> raise (UnificationError (UE_InstrNotInLockstep (i1, i2))) + end + | _, _ -> raise (UnificationError (UE_InstrNotInLockstep (i1, i2))) + in + + List.fold_left (fun umap (e1, e2) -> unify_exprs env umap e1 e2) umap (List.combine args1 args2) | Sabstract x1, Sabstract x2 when EcIdent.id_equal x1 x2 -> umap @@ -245,4 +261,3 @@ let unify_func env mode fname s = lv_of_list pvs in s_call (alias, fname, args) - diff --git a/src/ecUserMessages.ml b/src/ecUserMessages.ml index 9c947c1b6..985616507 100644 --- a/src/ecUserMessages.ml +++ b/src/ecUserMessages.ml @@ -423,7 +423,10 @@ end = struct msg "unknown type name: %a" pp_qsymbol name | UnknownFunName name -> - msg "unknown procedure: %a" pp_qsymbol name + msg "unknown procedure: %a" pp_qsymbol name + + | UnknownExceptionName name -> + msg "unknown exception: %a" pp_qsymbol name | UnknownModVar x -> msg "unknown module-level variable: %a" pp_qsymbol x diff --git a/src/phl/ecPhlApp.ml b/src/phl/ecPhlApp.ml index 554aa2402..1cbb747b6 100644 --- a/src/phl/ecPhlApp.ml +++ b/src/phl/ecPhlApp.ml @@ -41,7 +41,7 @@ let t_bdhoare_app_r_low i (phi, pR, f1, f2, g1, g2) tc = let s1, s2 = s_split env i bhs.bhs_s in let s1, s2 = stmt s1, stmt s2 in let nR = f_not pR in - let cond_phi = f_hoareS bhs.bhs_m bhs.bhs_pr s1 phi in + let cond_phi = f_hoareS bhs.bhs_m bhs.bhs_pr s1 phi [] in let condf1 = f_bdHoareS_r { bhs with bhs_s = s1; bhs_po = pR; bhs_bd = f1; } in let condg1 = f_bdHoareS_r { bhs with bhs_s = s1; bhs_po = nR; bhs_bd = g1; } in let condf2 = f_bdHoareS_r @@ -62,7 +62,7 @@ let t_bdhoare_app_r_low i (phi, pR, f1, f2, g1, g2) tc = let eqs = f_and (f_eq f2 r1) (f_eq g2 r2) in f_forall [(ir1, GTty treal); (ir2, GTty treal)] - (f_hoareS bhs.bhs_m (f_and bhs.bhs_pr eqs) s1 eqs) in + (f_hoareS bhs.bhs_m (f_and bhs.bhs_pr eqs) s1 eqs []) in let conds = [f_forall_mems [bhs.bhs_m] condbd; condnm] in let conds = if f_equal g1 f_r0 @@ -86,7 +86,7 @@ let t_bdhoare_app_r_low i (phi, pR, f1, f2, g1, g2) tc = let t_bdhoare_app_r i info tc = let tactic tc = let hs = tc1_as_hoareS tc in - let tt1 = EcPhlConseq.t_hoareS_conseq_nm hs.hs_pr f_true in + let tt1 = EcPhlConseq.t_hoareS_conseq_nm hs.hs_pr f_true hs.hs_poe in let tt2 = EcPhlAuto.t_pl_trivial in FApi.t_seqs [tt1; tt2; t_fail] tc in diff --git a/src/phl/ecPhlCall.ml b/src/phl/ecPhlCall.ml index f2f43747d..804d3dbfb 100644 --- a/src/phl/ecPhlCall.ml +++ b/src/phl/ecPhlCall.ml @@ -56,14 +56,24 @@ let wp2_call f_anda_simpl (PVM.subst env spre fpre) post (* -------------------------------------------------------------------- *) -let t_hoare_call fpre fpost tc = +let find_poe epost e = + snd (List.find (fun (e',_) -> EcPath.p_ntr_equal e e') epost) + +let call_conde env spre poe1 poe2 = + List.map + (fun (e1,p1) -> + let p2 = find_poe poe2 e1 in + f_imp_simpl (PVM.subst env spre p1) (PVM.subst env spre p2)) + poe1 + +let t_hoare_call fpre fpost fepost tc = let env = FApi.tc1_env tc in let hs = tc1_as_hoareS tc in let (lp,f,args),s = tc1_last_call tc hs.hs_s in let m = EcMemory.memory hs.hs_m in let fsig = (Fun.by_xpath f env).f_sig in (* The function satisfies the specification *) - let f_concl = f_hoareF fpre f fpost in + let f_concl = f_hoareF fpre f fpost fepost in (* The wp *) let pvres = pv_res in let vres = EcIdent.create "result" in @@ -75,11 +85,11 @@ let t_hoare_call fpre fpost tc = let post = f_forall_simpl [(vres, GTty fsig.fs_ret)] post in let spre = subst_args_call env m (e_tuple args) PVM.empty in let post = f_anda_simpl (PVM.subst env spre fpre) post in - let concl = f_hoareS_r { hs with hs_s = s; hs_po=post} in - + let poe = call_conde env spre fepost hs.hs_poe in + let post = List.fold f_anda_simpl post poe in + let concl = f_hoareS_r { hs with hs_s = s; hs_po=post } in FApi.xmutate1 tc `HlCall [f_concl; concl] - (* -------------------------------------------------------------------- *) let ehoare_call_pre_post fpre fpost tc = let (env, hyps, _) = FApi.tc1_eflat tc in @@ -207,7 +217,7 @@ let t_bdhoare_call fpre fpost opt_bd tc = (* most of the above code is duplicated from t_hoare_call *) let concl = match bhs.bhs_cmp, opt_bd with | FHle, None -> - f_hoareS bhs.bhs_m bhs.bhs_pr s post + f_hoareS bhs.bhs_m bhs.bhs_pr s post [] | FHeq, Some bd -> f_bdHoareS_r { bhs with bhs_s = s; bhs_po = post; bhs_bd = f_real_div bhs.bhs_bd bd; } @@ -308,7 +318,7 @@ let t_call side ax tc = let (_, f, _), _ = tc1_last_call tc hs.hs_s in if not (EcEnv.NormMp.x_equal env hf.hf_f f) then call_error env tc hf.hf_f f; - t_hoare_call hf.hf_pr hf.hf_po tc + t_hoare_call hf.hf_pr hf.hf_po hf.hf_poe tc | FeHoareF hf, FeHoareS hs -> let (_, f, _), _ = tc1_last_call tc hs.ehs_s in @@ -385,14 +395,9 @@ let mk_inv_spec (_pf : proofenv) env inv fl fr = f_equivF pre fl fr post let process_call side info tc = - let process_spec tc side = + let process_spec_2 tc side = let (hyps, concl) = FApi.tc1_flat tc in match concl.f_node, side with - | FhoareS hs, None -> - let (_,f,_) = fst (tc1_last_call tc hs.hs_s) in - let penv, qenv = LDecl.hoareF f hyps in - (penv, qenv, tbool, fun pre post -> f_hoareF pre f post) - | FbdHoareS bhs, None -> let (_,f,_) = fst (tc1_last_call tc bhs.bhs_s) in let penv, qenv = LDecl.hoareF f hyps in @@ -422,17 +427,23 @@ let process_call side info tc = | _ -> tc_error !!tc "the conclusion is not a hoare or an equiv" in - let process_inv tc side = + let process_spec_1 tc side = + let (hyps, concl) = FApi.tc1_flat tc in + match concl.f_node, side with + | FhoareS hs, None -> + let (_,f,_) = fst (tc1_last_call tc hs.hs_s) in + let penv, qenv = LDecl.hoareF f hyps in + (penv, qenv, tbool, fun pre post poe -> f_hoareF pre f post poe) + + | _ -> tc_error !!tc "the conclusion is not a hoare" in + + + let process_inv_1 tc side = if not (is_none side) then tc_error !!tc "cannot specify side for call with invariants"; let hyps, concl = FApi.tc1_flat tc in match concl.f_node with - | FhoareS hs -> - let (_,f,_) = fst (tc1_last_call tc hs.hs_s) in - let penv = LDecl.inv_memenv1 hyps in - (penv, tbool, fun inv -> f_hoareF inv f inv) - | FeHoareS hs -> let (_,f,_) = fst (tc1_last_call tc hs.ehs_s) in let penv = LDecl.inv_memenv1 hyps in @@ -452,6 +463,20 @@ let process_call side info tc = | _ -> tc_error !!tc "the conclusion is not a hoare or an equiv" in + let process_inv_2 tc side = + if not (is_none side) then + tc_error !!tc "cannot specify side for call with invariants"; + + let hyps, concl = FApi.tc1_flat tc in + match concl.f_node with + | FhoareS hs -> + let (_,f,_) = fst (tc1_last_call tc hs.hs_s) in + let penv = LDecl.inv_memenv1 hyps in + (penv, tbool, fun inv poe -> f_hoareF inv f inv poe) + + | _ -> tc_error !!tc "the conclusion is not a hoare" in + + let process_upto tc side info = if not (is_none side) then tc_error !!tc "cannot specify side for call with invariants"; @@ -481,18 +506,61 @@ let process_call side info tc = let process_cut tc info = match info with - | CI_spec (pre, post) -> - let penv,qenv,ty,fmake = process_spec tc side in - let pre = TTC.pf_process_form !!tc penv ty pre in - let post = TTC.pf_process_form !!tc qenv ty post in - fmake pre post - - | CI_inv inv -> - let hyps, ty, fmake = process_inv tc side in - let inv = TTC.pf_process_form !!tc hyps ty inv in - subtactic := (fun tc -> - FApi.t_firsts t_trivial 2 (EcPhlFun.t_fun inv tc)); - fmake inv + | CI_spec (pre, post, poe) -> + begin + let hyps, concl = FApi.tc1_flat tc in + match concl.f_node with + | FhoareS _ -> + let penv,qenv,ty,fmake = process_spec_1 tc side in + let pre = TTC.pf_process_form !!tc penv ty pre in + let post = TTC.pf_process_form !!tc qenv ty post in + let env = LDecl.toenv hyps in + let poe = + List.map + (fun (e,f) -> + let _,e = EcTyping.except_genpath env e in + let f = TTC.pf_process_form !!tc penv ty f in + e,f) + poe + in + if EcTyping.check_unique_epost poe then + fmake pre post poe + else assert false + + | _ -> + let penv,qenv,ty,fmake = process_spec_2 tc side in + let pre = TTC.pf_process_form !!tc penv ty pre in + let post = TTC.pf_process_form !!tc qenv ty post in + fmake pre post + end + + | CI_inv (inv, poe) -> + begin + let _, concl = FApi.tc1_flat tc in + match concl.f_node with + | FhoareS _ -> + let hyps, ty, fmake = process_inv_2 tc side in + let inv = TTC.pf_process_form !!tc hyps ty inv in + let env = LDecl.toenv hyps in + let poe = + List.map + (fun (e,f) -> + let _,e = EcTyping.except_genpath env e in + let f = TTC.pf_process_form !!tc hyps ty f in + e,f) + poe + in + if EcTyping.check_unique_epost poe then + fmake inv poe + else assert false + + | _ -> + let hyps, ty, fmake = process_inv_1 tc side in + let inv = TTC.pf_process_form !!tc hyps ty inv in + subtactic := (fun tc -> + FApi.t_firsts t_trivial 2 (EcPhlFun.t_fun inv tc)); + fmake inv + end | CI_upto info -> let bad, p, q, form = process_upto tc side info in @@ -561,13 +629,13 @@ let process_call_concave (fc, info) tc = let process_cut tc info = match info with - | CI_spec (pre, post) -> + | CI_spec (pre, post, []) -> let penv,qenv,ty,fmake = process_spec tc in let pre = TTC.pf_process_form !!tc penv ty pre in let post = TTC.pf_process_form !!tc qenv ty post in fmake pre post - | CI_inv inv -> + | CI_inv (inv, []) -> let env, ty, fmake = process_inv tc in let inv = TTC.pf_process_form !!tc env ty inv in subtactic := (fun tc -> diff --git a/src/phl/ecPhlCall.mli b/src/phl/ecPhlCall.mli index 1242522d1..3bca6f3e5 100644 --- a/src/phl/ecPhlCall.mli +++ b/src/phl/ecPhlCall.mli @@ -13,7 +13,7 @@ val wp2_call : -> EcMemory.memory -> EcMemory.memory -> form -> EcEnv.LDecl.hyps -> form -val t_hoare_call : form -> form -> backward +val t_hoare_call : form -> form -> (EcPath.path * form) list -> backward val t_bdhoare_call : form -> form -> form option -> backward val t_equiv_call : form -> form -> backward val t_equiv_call1 : side -> form -> form -> backward diff --git a/src/phl/ecPhlConseq.ml b/src/phl/ecPhlConseq.ml index 5627676a7..63595cfe8 100644 --- a/src/phl/ecPhlConseq.ml +++ b/src/phl/ecPhlConseq.ml @@ -48,24 +48,35 @@ let bd_goal tc fcmp fbd cmp bd = | Some fp -> fp (* -------------------------------------------------------------------- *) -let t_hoareF_conseq pre post tc = + +let find_poe epost e = + snd (List.find (fun (e',_) -> EcPath.p_ntr_equal e e') epost) + +let conseq_conde poe1 poe2 = + List.map (fun (e1,p1) -> f_imp (find_poe poe2 e1) p1) poe1 + +let t_hoareF_conseq pre post epost tc = let env = FApi.tc1_env tc in let hf = tc1_as_hoareF tc in let mpr,mpo = EcEnv.Fun.hoareF_memenv hf.hf_f env in let cond1, cond2 = conseq_cond hf.hf_pr hf.hf_po pre post in let concl1 = f_forall_mems [mpr] cond1 in let concl2 = f_forall_mems [mpo] cond2 in - let concl3 = f_hoareF pre hf.hf_f post in - FApi.xmutate1 tc `Conseq [concl1; concl2; concl3] + let cond2e = conseq_conde hf.hf_poe epost in + let concl2e = List.map (f_forall_mems [mpo]) cond2e in + let concl3 = f_hoareF pre hf.hf_f post epost in + FApi.xmutate1 tc `Conseq ([concl1; concl2; concl3] @ concl2e) (* -------------------------------------------------------------------- *) -let t_hoareS_conseq pre post tc = +let t_hoareS_conseq pre post epost tc = let hs = tc1_as_hoareS tc in let cond1, cond2 = conseq_cond hs.hs_pr hs.hs_po pre post in let concl1 = f_forall_mems [hs.hs_m] cond1 in let concl2 = f_forall_mems [hs.hs_m] cond2 in - let concl3 = f_hoareS_r { hs with hs_pr = pre; hs_po = post } in - FApi.xmutate1 tc `HlConseq [concl1; concl2; concl3] + let cond2e = conseq_conde hs.hs_poe epost in + let concl2e = List.map (f_forall_mems [hs.hs_m]) cond2e in + let concl3 = f_hoareS_r { hs with hs_pr = pre; hs_po = post; hs_poe = epost } in + FApi.xmutate1 tc `HlConseq ([concl1; concl2; concl3] @ concl2e) (* -------------------------------------------------------------------- *) let t_ehoareF_conseq pre post tc = @@ -174,8 +185,8 @@ let t_equivS_conseq pre post tc = (* -------------------------------------------------------------------- *) let t_conseq pre post tc = match (FApi.tc1_goal tc).f_node with - | FhoareF _ -> t_hoareF_conseq pre post tc - | FhoareS _ -> t_hoareS_conseq pre post tc + | FhoareF _ -> t_hoareF_conseq pre post [] tc + | FhoareS _ -> t_hoareS_conseq pre post [] tc | FbdHoareF _ -> t_bdHoareF_conseq pre post tc | FbdHoareS _ -> t_bdHoareS_conseq pre post tc | FeHoareF _ -> t_ehoareF_conseq pre post tc @@ -282,11 +293,17 @@ let cond_hoareF_notmod ?(mk_other=false) tc cond = else [] in cond, bmem, bother -let t_hoareF_notmod post tc = +let t_hoareF_notmod post epost tc = let hf = tc1_as_hoareF tc in let cond1, _, _ = cond_hoareF_notmod tc (f_imp post hf.hf_po) in - let cond2 = f_hoareF_r { hf with hf_po = post } in - FApi.xmutate1 tc `HlNotmod [cond1; cond2] + let econd1 = List.map (fun (e1,p1) -> + let f = f_imp (find_poe epost e1) p1 in + let cond1 , _, _ = cond_hoareF_notmod tc f in + cond1 + ) hf.hf_poe + in + let cond2 = f_hoareF_r { hf with hf_po = post; hf_poe = epost } in + FApi.xmutate1 tc `HlNotmod ([cond1; cond2] @ econd1) (* -------------------------------------------------------------------- *) let cond_hoareS_notmod ?(mk_other=false) tc cond = @@ -304,11 +321,17 @@ let cond_hoareS_notmod ?(mk_other=false) tc cond = else [] in cond, bmem, bother -let t_hoareS_notmod post tc = +let t_hoareS_notmod post epost tc = let hs = tc1_as_hoareS tc in let cond1, _, _ = cond_hoareS_notmod tc (f_imp post hs.hs_po) in - let cond2 = f_hoareS_r {hs with hs_po = post} in - FApi.xmutate1 tc `HlNotmod [cond1; cond2] + let econd1 = List.map (fun (e1,p1) -> + let f = f_imp (find_poe epost e1) p1 in + let cond1 , _, _ = cond_hoareS_notmod tc f in + cond1 + ) hs.hs_poe + in + let cond2 = f_hoareS_r {hs with hs_po = post; hs_poe = epost} in + FApi.xmutate1 tc `HlNotmod ([cond1; cond2] @ econd1) (* -------------------------------------------------------------------- *) let cond_bdHoareF_notmod ?(mk_other=false) tc cond = @@ -379,8 +402,19 @@ let gen_conseq_nm tnm tc pre post = FApi.t_swap_goals 0 1 gs ) -let t_hoareF_conseq_nm = gen_conseq_nm t_hoareF_notmod t_hoareF_conseq -let t_hoareS_conseq_nm = gen_conseq_nm t_hoareS_notmod t_hoareS_conseq +let gen_conseq_nm_e tnm tc pre post epost = + let t_trivials = List.map (fun _ -> fun a -> t_trivial a) epost in + let t_ids = List.map (fun _ -> t_id) epost in + FApi.t_internal ~info:"generic-conseq-nm" (fun g -> + let back = tc pre post epost @+ ([t_id; t_trivial; t_id] @ t_trivials) in + let gs = + (tnm post epost @+ ([t_id] @ [back] @ t_ids)) g + in + FApi.t_swap_goals 0 1 gs + ) + +let t_hoareF_conseq_nm = gen_conseq_nm_e t_hoareF_notmod t_hoareF_conseq +let t_hoareS_conseq_nm = gen_conseq_nm_e t_hoareS_notmod t_hoareS_conseq let t_equivF_conseq_nm = gen_conseq_nm t_equivF_notmod t_equivF_conseq let t_equivS_conseq_nm = gen_conseq_nm t_equivS_notmod t_equivS_conseq let t_bdHoareF_conseq_nm = gen_conseq_nm t_bdHoareF_notmod t_bdHoareF_conseq @@ -610,8 +644,8 @@ let t_hoareF_conseq_conj pre post pre' post' tc = tc_error !!tc "invalid pre-condition"; if not (f_equal hf.hf_po (f_and post' post)) then tc_error !!tc "invalid post-condition"; - let concl1 = f_hoareF pre hf.hf_f post in - let concl2 = f_hoareF pre' hf.hf_f post' in + let concl1 = f_hoareF pre hf.hf_f post hf.hf_poe in + let concl2 = f_hoareF pre' hf.hf_f post' hf.hf_poe in FApi.xmutate1 tc `HlConseqBd [concl1; concl2] (* -------------------------------------------------------------------- *) @@ -621,7 +655,7 @@ let t_bdHoareS_conseq_conj ~add post post' tc = let posth = if add then post' else f_and post' post in if not (f_equal hs.bhs_po postc) then tc_error !!tc "invalid post-condition"; - let concl1 = f_hoareS hs.bhs_m hs.bhs_pr hs.bhs_s post in + let concl1 = f_hoareS hs.bhs_m hs.bhs_pr hs.bhs_s post [] in let concl2 = f_bdHoareS_r { hs with bhs_po = posth } in FApi.xmutate1 tc `HlConseqBd [concl1; concl2] @@ -632,7 +666,7 @@ let t_bdHoareF_conseq_conj ~add post post' tc = let posth = if add then post' else f_and post' post in if not (f_equal hs.bhf_po postc) then tc_error !!tc "invalid post-condition"; - let concl1 = f_hoareF hs.bhf_pr hs.bhf_f post in + let concl1 = f_hoareF hs.bhf_pr hs.bhf_f post [] in let concl2 = f_bdHoareF_r { hs with bhf_po = posth } in FApi.xmutate1 tc `HlConseqBd [concl1; concl2] @@ -649,8 +683,8 @@ let t_equivS_conseq_conj pre1 post1 pre2 post2 pre' post' tc = tc_error !!tc "invalid pre-condition"; if not (f_equal es.es_po (f_ands [post';post1';post2'])) then tc_error !!tc "invalid post-condition"; - let concl1 = f_hoareS (mhr,snd es.es_ml) pre1 es.es_sl post1 in - let concl2 = f_hoareS (mhr,snd es.es_mr) pre2 es.es_sr post2 in + let concl1 = f_hoareS (mhr,snd es.es_ml) pre1 es.es_sl post1 [] in + let concl2 = f_hoareS (mhr,snd es.es_mr) pre2 es.es_sr post2 [] in let concl3 = f_equivS_r {es with es_pr = pre'; es_po = post'} in FApi.xmutate1 tc `HlConseqConj [concl1; concl2; concl3] @@ -667,8 +701,8 @@ let t_equivF_conseq_conj pre1 post1 pre2 post2 pre' post' tc = tc_error !!tc "invalid pre-condition"; if not (f_equal ef.ef_po (f_ands [post';post1';post2'])) then tc_error !!tc "invalid post-condition"; - let concl1 = f_hoareF pre1 ef.ef_fl post1 in - let concl2 = f_hoareF pre2 ef.ef_fr post2 in + let concl1 = f_hoareF pre1 ef.ef_fl post1 [] in + let concl2 = f_hoareF pre2 ef.ef_fr post2 [] in let concl3 = f_equivF pre' ef.ef_fl ef.ef_fr post' in FApi.xmutate1 tc `HlConseqConj [concl1; concl2; concl3] @@ -724,7 +758,7 @@ let t_hoareF_conseq_equiv f2 p q p2 q2 tc = let env, hyps, _ = FApi.tc1_eflat tc in let hf1 = tc1_as_hoareF tc in let ef = f_equivF p hf1.hf_f f2 q in - let hf2 = f_hoareF p2 f2 q2 in + let hf2 = f_hoareF p2 f2 q2 hf1.hf_poe in let (prml, _prmr), (poml, pomr) = Fun.equivF_memenv hf1.hf_f f2 env in let (cond1, cond2) = transitivity_side_cond hyps prml poml pomr p q p2 q2 hf1.hf_pr hf1.hf_po in @@ -799,12 +833,20 @@ let rec t_hi_conseq notmod f1 f2 f3 tc = let concl = FApi.tc1_goal tc in + let efind e poe = + List.find + (fun (e',_) -> + String.equal (EcPath.basename e) (EcPath.basename e')) + poe + in + + match concl.f_node, f1, f2, f3 with (* ------------------------------------------------------------------ *) (* hoareS / hoareS / ⊥ / ⊥ *) | FhoareS _, Some ((_, {f_node = FhoareS hs}) as nf1), None, None -> let tac = if notmod then t_hoareS_conseq_nm else t_hoareS_conseq in - t_on1 2 (t_apply_r nf1) (tac hs.hs_pr hs.hs_po tc) + t_on1 2 (t_apply_r nf1) (tac hs.hs_pr hs.hs_po hs.hs_poe tc) (* ------------------------------------------------------------------ *) (* hoareS / hoareS / hoareS / ⊥ *) @@ -816,8 +858,17 @@ let rec t_hi_conseq notmod f1 f2 f3 tc = let hs2 = pf_as_hoareS !!tc f2 in let tac = if notmod then t_hoareS_conseq_nm else t_hoareS_conseq in + let epost = + List.map + (fun (e, f1) -> + let _,f2 = efind e hs2.hs_poe in + (e, f_and f1 f2) + ) + hs.hs_poe + in + t_on1seq 2 - (tac (f_and hs.hs_pr hs2.hs_pr) (f_and hs.hs_po hs2.hs_po)) + (tac (f_and hs.hs_pr hs2.hs_pr) (f_and hs.hs_po hs2.hs_po) epost) (FApi.t_seqsub (t_hoareS_conseq_conj hs2.hs_pr hs2.hs_po hs.hs_pr hs.hs_po) [t_apply_r nf2; t_apply_r nf1]) @@ -834,14 +885,14 @@ let rec t_hi_conseq notmod f1 f2 f3 tc = t_hoareS_conseq_bdhoare (t_on1seq 1 (t_bdHoareS_conseq_bd hs.bhs_cmp hs.bhs_bd) - (t_on1seq 2 (tac hs.bhs_pr hs.bhs_po) (t_apply_r nf1))) + (t_on1seq 3 (tac hs.bhs_pr hs.bhs_po) (t_apply_r nf1))) tc (* ------------------------------------------------------------------ *) (* hoareF / hoareF / ⊥ / ⊥ *) | FhoareF _, Some ((_, {f_node = FhoareF hs}) as nf1), None, None -> let tac = if notmod then t_hoareF_conseq_nm else t_hoareF_conseq in - t_on1 2 (t_apply_r nf1) (tac hs.hf_pr hs.hf_po tc) + t_on1 2 (t_apply_r nf1) (tac hs.hf_pr hs.hf_po hs.hf_poe tc) (* ------------------------------------------------------------------ *) (* hoareF / hoareF / hoareF / ⊥ *) @@ -853,15 +904,24 @@ let rec t_hi_conseq notmod f1 f2 f3 tc = let hs2 = pf_as_hoareF !!tc f2 in let tac = if notmod then t_hoareF_conseq_nm else t_hoareF_conseq in + let epost = + List.map + (fun (e, f1) -> + let (_,f2) = efind e hs2.hf_poe in + (e, f_and f1 f2) + ) + hs.hf_poe + in + t_on1seq 2 - (tac (f_and hs.hf_pr hs2.hf_pr) (f_and hs.hf_po hs2.hf_po)) + (tac (f_and hs.hf_pr hs2.hf_pr) (f_and hs.hf_po hs2.hf_po) epost) (FApi.t_seqsub (t_hoareF_conseq_conj hs2.hf_pr hs2.hf_po hs.hf_pr hs.hf_po) [t_apply_r nf2; t_apply_r nf1]) tc (* ------------------------------------------------------------------ *) - (* hoareF / bdhoareF / ⊥ / ⊥ *) + (* hoareF / bdhoareF / ⊥ / ⊥ *) | FhoareF _, Some ((_, {f_node = FbdHoareF hs}) as nf1), None, None -> let tac = if notmod then t_bdHoareF_conseq_nm else t_bdHoareF_conseq in @@ -883,19 +943,19 @@ let rec t_hi_conseq notmod f1 f2 f3 tc = [t_id; t_id; t_apply_r nef; t_apply_r nf2] tc (* ------------------------------------------------------------------ *) - (* ehoareS / ehoareS / ⊥ / ⊥ *) + (* ehoareS / ehoareS / ⊥ / ⊥ *) | FeHoareS _, Some ((_, {f_node = FeHoareS hs}) as nf1), None, None -> let tac = if notmod then t_ehoareS_conseq_nm else t_ehoareS_conseq in FApi.t_last (t_apply_r nf1) (tac hs.ehs_pr hs.ehs_po tc) (* ------------------------------------------------------------------ *) - (* ehoareF / ehoareF / ⊥ / ⊥ *) + (* ehoareF / ehoareF / ⊥ / ⊥ *) | FeHoareF _, Some ((_, {f_node = FeHoareF hf}) as nf1), None, None -> let tac = if notmod then t_ehoareF_conseq_nm else t_ehoareF_conseq in FApi.t_last (t_apply_r nf1) (tac hf.ehf_pr hf.ehf_po tc) (* ------------------------------------------------------------------ *) - (* ehoareF / equivF / ehoareF *) + (* ehoareF / equivF / ehoareF *) | FeHoareF _, Some ((_, {f_node = FequivF ef}) as nef), Some((_, f2) as nf2), _ -> let hf2 = pf_as_ehoareF !!tc f2 in @@ -937,7 +997,7 @@ let rec t_hi_conseq notmod f1 f2 f3 tc = [ t_id; (* subgoal 1 : pre *) t_intro_i hi @! t_cut (f_hoareS_r {hs2 with hs_pr = pre}) @+ [ - t_hoareS_conseq hs2.hs_pr hs2.hs_po @+ + t_hoareS_conseq hs2.hs_pr hs2.hs_po [] @+ [ EcLowGoal.t_trivial; t_mytrivial; t_clear hi (* subgoal 2 : hs2 *)]; @@ -945,7 +1005,7 @@ let rec t_hi_conseq notmod f1 f2 f3 tc = (t_bdHoareS_conseq_bd hs.bhs_cmp hs.bhs_bd @+ [ t_id; (* subgoal 3 : bound *) t_bdHoareS_conseq_conj ~add:false hs2.hs_po post1 @+ [ - t_hoareS_conseq pre hs2.hs_po @+ [ + t_hoareS_conseq pre hs2.hs_po [] @+ [ t_intros_i [m;h0] @! t_cutdef (ptlocal ~args:[pamemory m; palocal h0] hi) mpre @! EcLowGoal.t_trivial; @@ -991,7 +1051,6 @@ let rec t_hi_conseq notmod f1 f2 f3 tc = Some ((_, f2) as nf2), None -> - let hs2 = pf_as_hoareF !!tc f2 in let tac = if notmod then t_bdHoareF_conseq_nm else t_bdHoareF_conseq in let m,hi,hh, h0 = @@ -1008,7 +1067,7 @@ let rec t_hi_conseq notmod f1 f2 f3 tc = [ t_id; (* subgoal 1 : pre *) t_intro_i hi @! t_cut (f_hoareF_r {hs2 with hf_pr = pre}) @+ [ - t_hoareF_conseq hs2.hf_pr hs2.hf_po @+ + t_hoareF_conseq hs2.hf_pr hs2.hf_po [] @+ [ EcLowGoal.t_trivial; t_mytrivial; t_clear hi (* subgoal 2 : hs2 *)]; @@ -1016,7 +1075,7 @@ let rec t_hi_conseq notmod f1 f2 f3 tc = (t_bdHoareF_conseq_bd hs.bhf_cmp hs.bhf_bd @+ [ t_id; (* subgoal 3 : bound *) t_bdHoareF_conseq_conj ~add:false hs2.hf_po post1 @+ [ - t_hoareF_conseq pre hs2.hf_po @+ [ + t_hoareF_conseq pre hs2.hf_po [] @+ [ t_intros_i [m;h0] @! t_cutdef (ptlocal ~args:[pamemory m; palocal h0] hi) mpre @! EcLowGoal.t_trivial; @@ -1109,13 +1168,13 @@ let rec t_hi_conseq notmod f1 f2 f3 tc = (* ------------------------------------------------------------------ *) (* equivS / ? / ? / ⊥ *) | FequivS es, Some _, Some _, None -> - let f3 = f_hoareS (mhr, snd es.es_mr) f_true es.es_sr f_true in + let f3 = f_hoareS (mhr, snd es.es_mr) f_true es.es_sr f_true [] in t_hi_conseq notmod f1 f2 (Some (None, f3)) tc (* ------------------------------------------------------------------ *) (* equivS / ? / ⊥ / ? *) | FequivS es, Some _, None, Some _ -> - let f2 = f_hoareS (mhr, snd es.es_ml) f_true es.es_sl f_true in + let f2 = f_hoareS (mhr, snd es.es_ml) f_true es.es_sl f_true [] in t_hi_conseq notmod f1 (Some (None, f2)) f3 tc (* ------------------------------------------------------------------ *) @@ -1184,13 +1243,13 @@ let rec t_hi_conseq notmod f1 f2 f3 tc = (* ------------------------------------------------------------------ *) (* equivF / ? / ? / ⊥ *) | FequivF ef, Some _, Some _, None -> - let f3 = f_hoareF f_true ef.ef_fr f_true in + let f3 = f_hoareF f_true ef.ef_fr f_true [] in t_hi_conseq notmod f1 f2 (Some (None, f3)) tc (* ------------------------------------------------------------------ *) (* equivF / ? / ⊥ / ? *) | FequivF ef, Some _, None, Some _ -> - let f2 = f_hoareF f_true ef.ef_fl f_true in + let f2 = f_hoareF f_true ef.ef_fl f_true [] in t_hi_conseq notmod f1 (Some (None, f2)) f3 tc | _ -> @@ -1222,38 +1281,15 @@ type processed_conseq_info = let process_info pe hyps = function | CQI_bd (cmp, bd) -> PCI_bd (cmp, TTC.pf_process_form pe hyps treal bd) -let process_conseq notmod ((info1, info2, info3) : conseq_ppterm option tuple3) tc = +let process_conseq_2 notmod ((info1, info2, info3) : conseq_ppterm option tuple3) tc = let hyps, concl = FApi.tc1_flat tc in let ensure_none o = if not (is_none o) then tc_error !!tc "cannot give a bound here" in - let process_cut1 ((pre, post), bd) = + let process_cut1 ((pre, post, poe), bd) = let penv, qenv, gpre, gpost, ty, fmake = match concl.f_node with - | FhoareS hs -> - let env = LDecl.push_active hs.hs_m hyps in - - let fmake pre post c_or_bd = - match c_or_bd with - | None -> - f_hoareS_r { hs with hs_pr = pre; hs_po = post; } - | Some (PCI_bd (cmp, bd)) -> - f_bdHoareS hs.hs_m pre hs.hs_s post (oget cmp) bd - in (env, env, hs.hs_pr, hs.hs_po, tbool, fmake) - - | FhoareF hf -> - let penv, qenv = LDecl.hoareF hf.hf_f hyps in - - let fmake pre post c_or_bd = - match c_or_bd with - | None -> - f_hoareF pre hf.hf_f post - | Some (PCI_bd (cmp, bd)) -> - f_bdHoareF pre hf.hf_f post (oget cmp) bd - - in (penv, qenv, hf.hf_pr, hf.hf_po, tbool, fmake) - | FeHoareS hs -> let env = LDecl.push_active hs.ehs_m hyps in let fmake pre post bd = @@ -1317,36 +1353,20 @@ let process_conseq notmod ((info1, info2, info3) : conseq_ppterm option tuple3) | _ -> tc_error !!tc "conseq: not a phl/prhl judgement" in - let pre = pre |> omap (TTC.pf_process_form !!tc penv ty) |> odfl gpre in - let post = post |> omap (TTC.pf_process_form !!tc qenv ty) |> odfl gpost in - let bd = bd |> omap (process_info !!tc penv) in + match poe with + | Some _ -> assert false + | None -> + let pre = pre |> omap (TTC.pf_process_form !!tc penv ty) |> odfl gpre in + let post = post |> omap (TTC.pf_process_form !!tc qenv ty) |> odfl gpost in + let bd = bd |> omap (process_info !!tc penv) in - fmake pre post bd + fmake pre post bd in - let process_cut2 side f1 ((pre, post), c_or_bd) = + let process_cut2 side f1 ((pre, post, poe), c_or_bd) = let penv, qenv, gpre, gpost, ty, fmake = match concl.f_node with - | FhoareS hs -> - let env = LDecl.push_active hs.hs_m hyps in - let fmake pre post c_or_bd = - ensure_none c_or_bd; - f_hoareS_r { hs with hs_pr = pre; hs_po = post; } - in (env, env, hs.hs_pr, hs.hs_po, tbool, fmake) - - | FhoareF hf -> - let f, pr, po = match f1 with - | None -> hf.hf_f, hf.hf_pr, hf.hf_po - | Some f1 -> match (snd f1).f_node with - | FequivF ef when side = `Left -> ef.ef_fr, f_true, f_true - | _ -> hf.hf_f, hf.hf_pr, hf.hf_po - in - let penv, qenv = LDecl.hoareF f hyps in - let fmake pre post c_or_bd = - ensure_none c_or_bd; f_hoareF pre f post in - (penv, qenv, pr, po, tbool, fmake) - | FeHoareF hf -> let f, pr, po = match f1 with | None -> hf.ehf_f, hf.ehf_pr, hf.ehf_po @@ -1365,7 +1385,7 @@ let process_conseq notmod ((info1, info2, info3) : conseq_ppterm option tuple3) let env = LDecl.push_active bhs.bhs_m hyps in let fmake pre post c_or_bd = ensure_none c_or_bd; - f_hoareS bhs.bhs_m pre bhs.bhs_s post + f_hoareS bhs.bhs_m pre bhs.bhs_s post [] in (env, env, bhs.bhs_pr, bhs.bhs_po, tbool, fmake) | FbdHoareF bhf -> @@ -1377,7 +1397,7 @@ let process_conseq notmod ((info1, info2, info3) : conseq_ppterm option tuple3) in let penv, qenv = LDecl.hoareF f hyps in let fmake pre post c_or_bd = - ensure_none c_or_bd; f_hoareF pre f post in + ensure_none c_or_bd; f_hoareF pre f post [] in (penv, qenv, pr, po, tbool, fmake) | FequivF ef -> @@ -1385,7 +1405,7 @@ let process_conseq notmod ((info1, info2, info3) : conseq_ppterm option tuple3) let penv, qenv = LDecl.hoareF f hyps in let fmake pre post c_or_bd = ensure_none c_or_bd; - f_hoareF pre f post in + f_hoareF pre f post [] in (penv, qenv, f_true, f_true, tbool, fmake) | FequivS es -> @@ -1404,7 +1424,7 @@ let process_conseq notmod ((info1, info2, info3) : conseq_ppterm option tuple3) f_bdHoareS m pre f post cmp bd | _, None -> - f_hoareS m pre f post + f_hoareS m pre f post [] | _, Some (PCI_bd (cmp,bd)) -> let cmp = odfl FHeq cmp in @@ -1415,11 +1435,123 @@ let process_conseq notmod ((info1, info2, info3) : conseq_ppterm option tuple3) | _ -> tc_error !!tc "conseq: not a phl/prhl judgement" in + match poe with + | Some _ -> assert false + | None -> + let pre = pre |> omap (TTC.pf_process_form !!tc penv ty) |> odfl gpre in + let post = post |> omap (TTC.pf_process_form !!tc qenv ty) |> odfl gpost in + let c_or_bd = c_or_bd |> omap (process_info !!tc penv) in + + fmake pre post c_or_bd + + in + + if List.for_all is_none [info1; info2; info3] + then t_id tc + else + let f1 = info1 |> omap (PT.tc1_process_full_closed_pterm_cut + ~prcut:(process_cut1) tc) in + let f2 = info2 |> omap (PT.tc1_process_full_closed_pterm_cut + ~prcut:(process_cut2 `Left f1) tc) in + let f3 = info3 |> omap (PT.tc1_process_full_closed_pterm_cut + ~prcut:(process_cut2 `Right f1) tc) in + + let ofalse = omap (fun (x, y) -> (Some x, y)) in + + t_hi_conseq notmod (f1 |> ofalse) (f2 |> ofalse) (f3 |> ofalse) tc + +let process_conseq_1 notmod ((info1, info2, info3) : conseq_ppterm option tuple3) tc = + let hyps, concl = FApi.tc1_flat tc in + + let env = LDecl.toenv hyps in + + let ensure_none o = + if not (is_none o) then tc_error !!tc "cannot give a bound here" in + + let process_conseq_e tc penv ty = + let aux (e,f) = + let _,e = EcTyping.except_genpath env e in + let f = TTC.pf_process_form !!tc penv ty f in + e,f + in + List.map aux + in + + let process_cut1 ((pre, post, poe), bd) = + let penv, qenv, gpre, gpost, gpoe, ty, fmake = + match concl.f_node with + | FhoareS hs -> + let env = LDecl.push_active hs.hs_m hyps in + + let fmake pre post poe c_or_bd = + match c_or_bd with + | None -> + f_hoareS_r { hs with hs_pr = pre; hs_po = post; hs_poe = poe } + | Some (PCI_bd (cmp, bd)) -> + f_bdHoareS hs.hs_m pre hs.hs_s post (oget cmp) bd + in (env, env, hs.hs_pr, hs.hs_po, hs.hs_poe, tbool, fmake) + + | FhoareF hf -> + let penv, qenv = LDecl.hoareF hf.hf_f hyps in + + let fmake pre post poe c_or_bd = + match c_or_bd with + | None -> + f_hoareF pre hf.hf_f post poe + | Some (PCI_bd (cmp, bd)) -> + f_bdHoareF pre hf.hf_f post (oget cmp) bd + + in (penv, qenv, hf.hf_pr, hf.hf_po, hf.hf_poe, tbool, fmake) + + | _ -> tc_error !!tc "conseq: not a phl/prhl judgement" + in + + let pre = pre |> omap (TTC.pf_process_form !!tc penv ty) |> odfl gpre in + let post = post |> omap (TTC.pf_process_form !!tc qenv ty) |> odfl gpost in + let poe = poe |> omap ( process_conseq_e tc penv ty) |> odfl gpoe in + + let bd = bd |> omap (process_info !!tc penv) in + + if EcTyping.check_unique_epost poe then + fmake pre post poe bd + else assert false + + in + + let process_cut2 side f1 ((pre, post, poe), c_or_bd) = + let penv, qenv, gpre, gpost, gpoe, ty, fmake = + match concl.f_node with + | FhoareS hs -> + let env = LDecl.push_active hs.hs_m hyps in + let fmake pre post poe c_or_bd = + ensure_none c_or_bd; + f_hoareS_r { hs with hs_pr = pre; hs_po = post; hs_poe = poe } + in (env, env, hs.hs_pr, hs.hs_po, hs.hs_poe, tbool, fmake) + + | FhoareF hf -> + let f, pr, po, poe = match f1 with + | None -> hf.hf_f, hf.hf_pr, hf.hf_po, hf.hf_poe + | Some f1 -> match (snd f1).f_node with + | FequivF ef when side = `Left -> ef.ef_fr, f_true, f_true, [] + | _ -> hf.hf_f, hf.hf_pr, hf.hf_po, hf.hf_poe + in + let penv, qenv = LDecl.hoareF f hyps in + let fmake pre post poe c_or_bd = + ensure_none c_or_bd; f_hoareF pre f post poe in + (penv, qenv, pr, po, poe, tbool, fmake) + + | _ -> tc_error !!tc "conseq: not a phl/prhl judgement" + in + let pre = pre |> omap (TTC.pf_process_form !!tc penv ty) |> odfl gpre in let post = post |> omap (TTC.pf_process_form !!tc qenv ty) |> odfl gpost in + let poe = poe |> omap (process_conseq_e tc penv ty) |> odfl gpoe in + let c_or_bd = c_or_bd |> omap (process_info !!tc penv) in - fmake pre post c_or_bd + if EcTyping.check_unique_epost poe then + fmake pre post poe c_or_bd + else assert false in @@ -1437,9 +1569,17 @@ let process_conseq notmod ((info1, info2, info3) : conseq_ppterm option tuple3) t_hi_conseq notmod (f1 |> ofalse) (f2 |> ofalse) (f3 |> ofalse) tc +let process_conseq notmod ((info1, info2, info3) : conseq_ppterm option tuple3) tc = + let _, concl = FApi.tc1_flat tc in + + match concl.f_node with + | FhoareS _ | FhoareF _ -> process_conseq_1 notmod (info1, info2, info3) tc + | _ -> process_conseq_2 notmod (info1, info2, info3) tc + + (* -------------------------------------------------------------------- *) let process_bd_equiv side (pr, po) tc = - let info = FPCut ((Some pr, Some po), None) in + let info = FPCut ((Some pr, Some po, None), None) in let info = Some { fp_mode = `Implicit; fp_head = info; fp_args = []; } in let info2, info3 = sideif side (info, None) (None, info) in process_conseq true (None, info2, info3) tc @@ -1469,10 +1609,13 @@ let process_conseq_opt cqopt infos tc = let t_conseqauto ?(delta = true) ?tsolve tc = let (hyps, concl), mk_other = FApi.tc1_flat tc, true in + let t_hoareF_notmod f = t_hoareF_notmod f [] in + let t_hoareS_notmod f = t_hoareS_notmod f [] in + let todo = match concl.f_node with - | FhoareF hf -> Some (t_hoareF_notmod, cond_hoareF_notmod ~mk_other tc hf.hf_po) - | FhoareS hs -> Some (t_hoareS_notmod, cond_hoareS_notmod ~mk_other tc hs.hs_po ) + | FhoareF hf -> Some (t_hoareF_notmod, cond_hoareF_notmod ~mk_other tc hf.hf_po) + | FhoareS hs -> Some (t_hoareS_notmod, cond_hoareS_notmod ~mk_other tc hs.hs_po ) | FbdHoareF hf -> Some (t_bdHoareF_notmod, cond_bdHoareF_notmod ~mk_other tc hf.bhf_po) | FbdHoareS hs -> Some (t_bdHoareS_notmod, cond_bdHoareS_notmod ~mk_other tc hs.bhs_po) | FequivF ef -> Some (t_equivF_notmod, cond_equivF_notmod ~mk_other tc ef.ef_po) diff --git a/src/phl/ecPhlConseq.mli b/src/phl/ecPhlConseq.mli index 4400622dc..c546a73d2 100644 --- a/src/phl/ecPhlConseq.mli +++ b/src/phl/ecPhlConseq.mli @@ -11,8 +11,10 @@ open EcCoreGoal val t_equivF_conseq : form -> form -> FApi.backward val t_equivS_conseq : form -> form -> FApi.backward val t_eagerF_conseq : form -> form -> FApi.backward -val t_hoareF_conseq : form -> form -> FApi.backward -val t_hoareS_conseq : form -> form -> FApi.backward +val t_hoareF_conseq : + form -> form -> (EcPath.path * form) list -> FApi.backward +val t_hoareS_conseq : + form -> form -> (EcPath.path * form) list -> FApi.backward val t_bdHoareF_conseq : form -> form -> FApi.backward val t_bdHoareS_conseq : form -> form -> FApi.backward @@ -24,8 +26,10 @@ val t_bdHoareF_conseq_bd : hoarecmp -> form -> FApi.backward (* -------------------------------------------------------------------- *) val t_equivF_conseq_nm : form -> form -> FApi.backward val t_equivS_conseq_nm : form -> form -> FApi.backward -val t_hoareF_conseq_nm : form -> form -> FApi.backward -val t_hoareS_conseq_nm : form -> form -> FApi.backward +val t_hoareF_conseq_nm : + form -> form -> (EcPath.path * form) list -> FApi.backward +val t_hoareS_conseq_nm : + form -> form -> (EcPath.path * form) list -> FApi.backward val t_bdHoareF_conseq_nm : form -> form -> FApi.backward val t_bdHoareS_conseq_nm : form -> form -> FApi.backward (* -------------------------------------------------------------------- *) diff --git a/src/phl/ecPhlCoreView.ml b/src/phl/ecPhlCoreView.ml index ef34c91b1..e5780e167 100644 --- a/src/phl/ecPhlCoreView.ml +++ b/src/phl/ecPhlCoreView.ml @@ -9,7 +9,7 @@ let t_hoare_of_bdhoareS_r tc = let bhs = tc1_as_bdhoareS tc in if not (bhs.bhs_cmp = FHeq && f_equal bhs.bhs_bd f_r0) then tc_error !!tc "%s" "bound must be equal to 0%r"; - let concl = f_hoareS bhs.bhs_m bhs.bhs_pr bhs.bhs_s (f_not bhs.bhs_po) in + let concl = f_hoareS bhs.bhs_m bhs.bhs_pr bhs.bhs_s (f_not bhs.bhs_po) [] in FApi.xmutate1 tc `ViewBdHoare [concl] (* -------------------------------------------------------------------- *) @@ -17,7 +17,7 @@ let t_hoare_of_bdhoareF_r tc = let bhf = tc1_as_bdhoareF tc in if not (bhf.bhf_cmp = FHeq && f_equal bhf.bhf_bd f_r0) then tc_error !!tc "%s" "bound must be equal to 0%r"; - let concl = f_hoareF bhf.bhf_pr bhf.bhf_f (f_not bhf.bhf_po) in + let concl = f_hoareF bhf.bhf_pr bhf.bhf_f (f_not bhf.bhf_po) [] in FApi.xmutate1 tc `ViewBdHoare [concl] (* -------------------------------------------------------------------- *) diff --git a/src/phl/ecPhlEager.ml b/src/phl/ecPhlEager.ml index 65d1541a6..36d55dccc 100644 --- a/src/phl/ecPhlEager.ml +++ b/src/phl/ecPhlEager.ml @@ -162,7 +162,7 @@ let t_eager_if_r tc = f_forall [(m2, GTmem (snd es.es_mr)); (b, GTty tbool)] - (f_hoareS (mhr, snd es.es_ml) (f_and p eqb) s eqb) in + (f_hoareS (mhr, snd es.es_ml) (f_and p eqb) s eqb []) in let cT = let pre = f_and es.es_pr (f_eq fel f_true) in @@ -505,13 +505,17 @@ let eager pf env s s' inv eqIs eqXs c c' eqO = (* (h) is assumed *) (fhyps, eqi) - | Sassert el, Sassert er -> - check_args [el]; + | Sraise (_,argsl), Sraise (_,argsr) -> + check_args argsl; let eqnm = Mpv2.split_nmod env modi modi' eqo in let eqm = Mpv2.split_mod env modi modi' eqo in if not (Mpv2.subset eqm eqXs) then raise EqObsInError; - let eqi = Mpv2.union eqIs eqnm in - (fhyps, Mpv2.add_eqs env el er eqi) + let eqi = + List.fold_left2 + (fun eqs e1 e2 -> Mpv2.add_eqs env e1 e2 eqs) + eqnm argsl argsr + in + (fhyps,eqi) | Sabstract _, Sabstract _ -> assert false (* FIXME *) @@ -643,7 +647,7 @@ let process_fun_abs info eqI tc = let process_call info tc = let process_cut info = match info with - | EcParsetree.CI_spec (fpre, fpost) -> + | EcParsetree.CI_spec (fpre, fpost, []) -> let env, hyps, _ = FApi.tc1_eflat tc in let es = tc1_as_equivS tc in diff --git a/src/phl/ecPhlEqobs.ml b/src/phl/ecPhlEqobs.ml index 311db4c14..abc5a0868 100644 --- a/src/phl/ecPhlEqobs.ml +++ b/src/phl/ecPhlEqobs.ml @@ -277,7 +277,25 @@ and i_eqobs_in il ir sim local (eqo:Mpv2.t) = let eqs = List.fold_left2 doit Mpv2.empty bsl bsr in !rsim, add_eqs !rsim local eqs el er - | Sassert el, Sassert er -> sim, add_eqs sim local eqo el er + | Sraise (_,argsl), Sraise (_,argsr) -> + let argsl, argsr = + match argsl, argsr with + | _, _ when List.length argsl = List.length argsr -> argsl, argsr + | [al], _ -> begin + match al.e_node with + | Etuple argsl -> argsl, argsr + | _ -> raise EqObsInError + end + | _, [ar] -> begin + match ar.e_node with + | Etuple argsr -> argsl, argsr + | _ -> raise EqObsInError + end + | _ -> raise EqObsInError + in + let eqi = List.fold_left2 (add_eqs sim local) eqo argsl argsr in + sim, eqi + | _, _ -> raise EqObsInError and s_eqobs_in_full sl sr sim local eqo = diff --git a/src/phl/ecPhlExists.ml b/src/phl/ecPhlExists.ml index c0b4bab2b..c4297cbf6 100644 --- a/src/phl/ecPhlExists.ml +++ b/src/phl/ecPhlExists.ml @@ -107,7 +107,7 @@ let process_exists_intro ~(elim : bool) fs tc = | FbdHoareS bhs -> LDecl.push_active bhs.bhs_m hyps | FequivF ef -> fst (LDecl.equivF ef.ef_fl ef.ef_fr hyps) | FequivS es -> LDecl.push_all [es.es_ml; es.es_mr] hyps - | _ -> tc_error_noXhl ~kinds:hlkinds_Xhl !!tc + | _ -> tc_error_noXhl ~kinds:hlkinds_Xhl !!tc in let fs = diff --git a/src/phl/ecPhlFel.ml b/src/phl/ecPhlFel.ml index 0bfdfbb7a..dd04459c7 100644 --- a/src/phl/ecPhlFel.ml +++ b/src/phl/ecPhlFel.ml @@ -103,7 +103,7 @@ and callable_oracles_i env modv os i = | Smatch (_, b) -> callable_oracles_sx env modv os (List.map snd b) | Sif (_, s1, s2) -> callable_oracles_sx env modv os [s1; s2] - | Sasgn _ | Srnd _ | Sassert _ -> os + | Sasgn _ | Srnd _ | Sraise _ -> os | Sabstract _ -> assert false (* FIXME *) @@ -185,7 +185,7 @@ let t_failure_event_r (at_pos, cntr, ash, q, f_event, pred_specs, inv) tc = let pre = f_ands (eqparams :: (eqxs@eqgs)) in let p = f_and (f_not f_event) (f_eq cntr f_i0) in let p = f_and_simpl p inv in - f_hoareS memenv pre (stmt s_hd) p + f_hoareS memenv pre (stmt s_hd) p [] in let oracle_goal o = @@ -215,7 +215,7 @@ let t_failure_event_r (at_pos, cntr, ash, q, f_event, pred_specs, inv) tc = let pre = f_and_simpl pre inv in let post = f_int_lt old_cntr cntr in let post = f_and_simpl post inv in - f_forall_simpl [old_cntr_id,GTty tint] (f_hoareF pre o post) + f_forall_simpl [old_cntr_id,GTty tint] (f_hoareF pre o post []) in let cntr_stable_goal = let pre = f_ands [f_not some_p;f_eq f_event old_b;f_eq cntr old_cntr] in @@ -224,7 +224,7 @@ let t_failure_event_r (at_pos, cntr, ash, q, f_event, pred_specs, inv) tc = let post = f_and_simpl post inv in f_forall_simpl [old_b_id,GTty tbool; old_cntr_id,GTty tint] - (f_hoareF pre o post) + (f_hoareF pre o post []) in [not_F_to_F_goal;cntr_decr_goal;cntr_stable_goal] in diff --git a/src/phl/ecPhlFun.ml b/src/phl/ecPhlFun.ml index 98120db7c..85fb067f5 100644 --- a/src/phl/ecPhlFun.ml +++ b/src/phl/ecPhlFun.ml @@ -79,8 +79,11 @@ let t_hoareF_fun_def_r tc = let m = EcMemory.memory memenv in let fres = odfl f_tt (omap (form_of_expr m) fdef.f_ret) in let post = PVM.subst1 env pv_res m fres hf.hf_po in + let epost = + List.map (fun (e,f) -> e,PVM.subst1 env pv_res m fres f) hf.hf_poe + in let pre = PVM.subst env (subst_pre env fsig m PVM.empty) hf.hf_pr in - let concl' = f_hoareS memenv pre fdef.f_body post in + let concl' = f_hoareS memenv pre fdef.f_body post epost in FApi.xmutate1 tc `FunDef [concl'] (* ------------------------------------------------------------------ *) @@ -160,7 +163,7 @@ module FunAbsLow = struct let (top, _, oi, _) = EcLowPhlGoal.abstract_info env f in let fv = PV.fv env mhr inv in PV.check_depend env fv top; - let ospec o = f_hoareF inv o inv in + let ospec o = f_hoareF inv o inv [] in let sg = List.map ospec (OI.allowed oi) in (inv, inv, sg) @@ -248,7 +251,7 @@ let t_hoareF_abs_r inv tc = let pre, post, sg = FunAbsLow.hoareF_abs_spec !!tc env hf.hf_f inv in let tactic tc = FApi.xmutate1 tc `FunAbs sg in - FApi.t_last tactic (EcPhlConseq.t_hoareF_conseq pre post tc) + FApi.t_last tactic (EcPhlConseq.t_hoareF_conseq pre post hf.hf_poe tc) let t_ehoareF_abs_r inv tc = let env = FApi.tc1_env tc in @@ -423,7 +426,7 @@ let t_fun_to_code_hoare_r tc = let spo = ToCodeLow.add_var env pv_res mhr r m PVM.empty in let pre = PVM.subst env spr hf.hf_pr in let post = PVM.subst env spo hf.hf_po in - let concl = f_hoareS m pre st post in + let concl = f_hoareS m pre st post hf.hf_poe in FApi.xmutate1 tc `FunToCode [concl] diff --git a/src/phl/ecPhlInline.ml b/src/phl/ecPhlInline.ml index 4e7f6d027..db1c5167a 100644 --- a/src/phl/ecPhlInline.ml +++ b/src/phl/ecPhlInline.ml @@ -50,7 +50,7 @@ module LowSubst = struct | Sif (c, s1, s2) -> i_if (esubst c, ssubst s1, ssubst s2) | Swhile (e, stmt) -> i_while (esubst e, ssubst stmt) | Smatch (e, bs) -> i_match (esubst e, List.Smart.map (snd_map ssubst) bs) - | Sassert e -> i_assert (esubst e) + | Sraise (e,es) -> i_raise (e, List.map esubst es) | Sabstract _ -> i and issubst m (is : instr list) = diff --git a/src/phl/ecPhlLoopTx.ml b/src/phl/ecPhlLoopTx.ml index 973fe8a14..4487e2c8c 100644 --- a/src/phl/ecPhlLoopTx.ml +++ b/src/phl/ecPhlLoopTx.ml @@ -67,7 +67,7 @@ let check_dslc pf = | Smatch (_, bs) -> List.iter (doit_s |- snd) bs - | Srnd _ | Scall _ | Swhile _ | Sassert _ | Sabstract _ -> + | Srnd _ | Scall _ | Swhile _ | Sraise _ | Sabstract _ -> error () and doit_s c = @@ -304,7 +304,7 @@ let process_unroll_for side cpos tc = t_doit (i+1) (pos + blen) zs]) tc in let t_conseq_nm tc = - (EcPhlConseq.t_hoareS_conseq_nm (tc1_get_pre tc) f_true @+ + (EcPhlConseq.t_hoareS_conseq_nm (tc1_get_pre tc) f_true (tc1_get_poste tc) @+ [ t_trivial; t_trivial; EcPhlTAuto.t_hoare_true]) tc in let doi i tc = @@ -336,4 +336,4 @@ let process_unroll (side, cpos, for_) tc = else begin let cpos = EcProofTyping.tc1_process_codepos tc (side, cpos) in t_unroll side cpos tc - end \ No newline at end of file + end diff --git a/src/phl/ecPhlRCond.ml b/src/phl/ecPhlRCond.ml index e28a07939..5ca1d2e11 100644 --- a/src/phl/ecPhlRCond.ml +++ b/src/phl/ecPhlRCond.ml @@ -50,7 +50,7 @@ module Low = struct | o, pre :: _ when f_equal o fop_interp_ehoare_form -> pre | _ -> tc_error !!tc "the pre should have the form \"_ `|` _\"" in - let concl1 = f_hoareS hs.ehs_m pre hd e in + let concl1 = f_hoareS hs.ehs_m pre hd e [] in let concl2 = f_eHoareS_r { hs with ehs_s = s } in FApi.xmutate1 tc `RCond [concl1; concl2] @@ -60,7 +60,7 @@ module Low = struct let bhs = tc1_as_bdhoareS tc in let m = EcMemory.memory bhs.bhs_m in let hd,_,e,s = gen_rcond (!!tc, env) b m at_pos bhs.bhs_s in - let concl1 = f_hoareS bhs.bhs_m bhs.bhs_pr hd e in + let concl1 = f_hoareS bhs.bhs_m bhs.bhs_pr hd e [] in let concl2 = f_bdHoareS_r { bhs with bhs_s = s } in FApi.xmutate1 tc `RCond [concl1; concl2] @@ -80,7 +80,7 @@ module Low = struct let pre1 = Fsubst.f_subst s1 es.es_pr in let concl1 = f_forall_mems [mo', EcMemory.memtype mo] - (f_hoareS (EcFol.mhr, EcMemory.memtype m) pre1 hd e) in + (f_hoareS (EcFol.mhr, EcMemory.memtype m) pre1 hd e []) in let sl,sr = match side with `Left -> s, es.es_sr | `Right -> es.es_sl, s in let concl2 = f_equivS_r { es with es_sl = sl; es_sr = sr } in FApi.xmutate1 tc `RCond [concl1; concl2] @@ -243,7 +243,7 @@ module LowMatch = struct let pr = ofold f_and bhs.bhs_pr epr in - let concl1 = f_hoareS bhs.bhs_m bhs.bhs_pr hd po1 in + let concl1 = f_hoareS bhs.bhs_m bhs.bhs_pr hd po1 [] in let concl2 = f_bdHoareS_r { bhs with bhs_pr = pr; bhs_m = me; bhs_s = full; } in FApi.xmutate1 tc `RCondMatch [concl1; concl2] @@ -273,7 +273,7 @@ module LowMatch = struct let concl1 = f_forall_mems [mo', EcMemory.memtype mo] - (f_hoareS (EcFol.mhr, EcMemory.memtype m) pre1 hd po1) in + (f_hoareS (EcFol.mhr, EcMemory.memtype m) pre1 hd po1 []) in let (ml, mr), (sl, sr) = match side with diff --git a/src/phl/ecPhlRnd.ml b/src/phl/ecPhlRnd.ml index 780bf5708..9baa3eca9 100644 --- a/src/phl/ecPhlRnd.ml +++ b/src/phl/ecPhlRnd.ml @@ -206,7 +206,7 @@ module Core = struct let bounded_distr = f_real_le (f_mu env distr event) bound in let pre = f_and bhs.bhs_pr pre_bound in let post = f_anda bounded_distr (mk_event_cond event) in - let concl = f_hoareS bhs.bhs_m pre s post in + let concl = f_hoareS bhs.bhs_m pre s post [] in let concl = f_forall_simpl binders concl in [concl] | PNoRndParams, _ -> @@ -230,7 +230,7 @@ module Core = struct let bounded_distr = f_real_le (f_mu env distr event) bound in let pre = f_and bhs.bhs_pr pre_bound in let post = f_anda bounded_distr (mk_event_cond event) in - let concl = f_hoareS bhs.bhs_m pre s post in + let concl = f_hoareS bhs.bhs_m pre s post [] in let concl = f_forall_simpl binders concl in [concl] | PSingleRndParam event, _ -> @@ -665,7 +665,7 @@ let process_rnd side pos tac_info tc = Double ((b1, p1), (b2, p2)) ) in - + t_equiv_rnd side ?pos bij_info tc | _ -> tc_error !!tc "invalid arguments" diff --git a/src/phl/ecPhlUpto.ml b/src/phl/ecPhlUpto.ml index 82a8dfa70..34b5b7aea 100644 --- a/src/phl/ecPhlUpto.ml +++ b/src/phl/ecPhlUpto.ml @@ -62,7 +62,7 @@ let rec check_bad_true env bad s = | Smatch(_,bs) -> let check_branch (_, s) = check_bad_true env bad s.s_node in List.iter (check_branch) bs - | Sassert _ -> () + | Sraise _ -> assert false | Sabstract _ -> () end; check_bad_true env bad c @@ -142,8 +142,7 @@ and i_upto env alpha bad i1 i2 = with E.NotConv -> false end - | Sassert a1, Sassert a2 -> - EqTest.for_expr env ~alpha a1 a2 + | Sraise (_e1,_es1), Sraise (_e2,_es2) -> assert false | Sabstract id1, Sabstract id2 -> EcIdent.id_equal id1 id2 diff --git a/src/phl/ecPhlWhile.ml b/src/phl/ecPhlWhile.ml index 153266917..85f499d69 100644 --- a/src/phl/ecPhlWhile.ml +++ b/src/phl/ecPhlWhile.ml @@ -41,8 +41,7 @@ let while_info env e s = let f = EcEnv.NormMp.norm_xfun env f in (w, r, Sx.add f c) - | Sassert e -> - (w, e_read_r env r e, c) + | Sraise _ -> assert false | Sabstract id -> let add_pv x (pv,ty) = PV.add env pv ty x in @@ -70,7 +69,7 @@ let t_hoare_while_r inv tc = (* the body preserves the invariant *) let b_pre = f_and_simpl inv e in let b_post = inv in - let b_concl = f_hoareS hs.hs_m b_pre c b_post in + let b_concl = f_hoareS hs.hs_m b_pre c b_post [] in (* the wp of the while *) let post = f_imps_simpl [f_not_simpl e; inv] hs.hs_po in let modi = s_write env c in @@ -185,7 +184,7 @@ let t_bdhoare_while_rev_r inv tc = (f_eq bound f_r1) in let term_post = generalize_mod env (EcMemory.memory mem) modi term_post in let term_post = f_and inv term_post in - f_hoareS mem b_pre rem_s term_post + f_hoareS mem b_pre rem_s term_post [] in FApi.xmutate1_hyps tc `While [(hyps', body_concl); (hyps, rem_concl)] @@ -654,7 +653,7 @@ let process_async_while (winfos : EP.async_while_info) tc = let pre = f_ands [inv; form_of_expr mhr el; f_not p0; p1] in f_forall_mems [evs.es_mr] - (f_hoareS (mhr, EcMemory.memtype evs.es_ml) pre cl inv) + (f_hoareS (mhr, EcMemory.memtype evs.es_ml) pre cl inv []) and hr2 = let subst = Fsubst.f_bind_mem Fsubst.f_subst_id mr mhr in @@ -664,7 +663,7 @@ let process_async_while (winfos : EP.async_while_info) tc = let pre = f_ands [inv; form_of_expr mhr er; f_not p0; f_not p1] in f_forall_mems [evs.es_ml] - (f_hoareS (mhr, EcMemory.memtype evs.es_mr) pre cr inv) + (f_hoareS (mhr, EcMemory.memtype evs.es_mr) pre cr inv []) in (hr1, hr2) in diff --git a/src/phl/ecPhlWp.ml b/src/phl/ecPhlWp.ml index 9c7553f56..4e5ccd40d 100644 --- a/src/phl/ecPhlWp.ml +++ b/src/phl/ecPhlWp.ml @@ -11,32 +11,35 @@ open EcLowPhlGoal module LowInternal = struct exception No_wp + let find_poe epost e = + snd (List.find (fun (e',_) -> EcPath.p_ntr_equal e e') epost) + let wp_asgn_aux c_pre memenv lv e (lets, f) = let m = EcMemory.memory memenv in let let1 = lv_subst ?c_pre m lv (form_of_expr m e) in (let1::lets, f) let rec wp_stmt - onesided c_pre env memenv (stmt: EcModules.instr list) letsf + onesided c_pre env memenv (stmt: EcModules.instr list) letsf epost = match stmt with | [] -> (stmt, letsf) | i :: stmt' -> try - let letsf = wp_instr onesided c_pre env memenv i letsf in - wp_stmt onesided c_pre env memenv stmt' letsf + let letsf = wp_instr onesided c_pre env memenv i letsf epost in + wp_stmt onesided c_pre env memenv stmt' letsf epost with No_wp -> (stmt, letsf) - and wp_instr onesided c_pre env memenv i letsf = + and wp_instr onesided c_pre env memenv i letsf epost = match i.i_node with | Sasgn (lv,e) -> wp_asgn_aux c_pre memenv lv e letsf | Sif (e,s1,s2) -> let (r1,letsf1) = - wp_stmt onesided c_pre env memenv (List.rev s1.s_node) letsf in + wp_stmt onesided c_pre env memenv (List.rev s1.s_node) letsf epost in let (r2,letsf2) = - wp_stmt onesided c_pre env memenv (List.rev s2.s_node) letsf in + wp_stmt onesided c_pre env memenv (List.rev s2.s_node) letsf epost in if List.is_empty r1 && List.is_empty r2 then begin let post1 = mk_let_of_lv_substs env letsf1 in let post2 = mk_let_of_lv_substs env letsf2 in @@ -49,7 +52,7 @@ module LowInternal = struct | Smatch (e, bs) -> begin let wps = let do1 (_, s) = - wp_stmt onesided c_pre env memenv (List.rev s.s_node) letsf in + wp_stmt onesided c_pre env memenv (List.rev s.s_node) letsf epost in List.map do1 bs in @@ -70,10 +73,7 @@ module LowInternal = struct ([],post) end - | Sassert e when onesided -> - let phi = form_of_expr (EcMemory.memory memenv) e in - let lets, f = letsf in - (lets, EcFol.f_and_simpl phi f) + | Sraise (e,_) when onesided -> ([], find_poe epost e) | _ -> raise No_wp @@ -120,10 +120,10 @@ module LowInternal = struct end -let wp ?(uselet=true) ?(onesided=false) ?c_pre env m s post = +let wp ?(uselet=true) ?(onesided=false) ?c_pre env m s post epost = let (r, letsf) = LowInternal.wp_stmt - onesided c_pre env m (List.rev s.s_node) ([], post) + onesided c_pre env m (List.rev s.s_node) ([], post) epost in let pre = mk_let_of_lv_substs ~uselet env letsf in List.rev r, pre @@ -145,7 +145,7 @@ module TacInternal = struct let (s_hd, s_wp) = o_split env i hs.hs_s in let s_wp = EcModules.stmt s_wp in let s_wp, post = - wp ~uselet ~onesided:true env hs.hs_m s_wp hs.hs_po in + wp ~uselet ~onesided:true env hs.hs_m s_wp hs.hs_po hs.hs_poe in check_wp_progress tc i hs.hs_s s_wp; let s = EcModules.stmt (s_hd @ s_wp) in let concl = f_hoareS_r { hs with hs_s = s; hs_po = post} in @@ -167,7 +167,7 @@ module TacInternal = struct let bhs = tc1_as_bdhoareS tc in let (s_hd, s_wp) = o_split env i bhs.bhs_s in let s_wp = EcModules.stmt s_wp in - let s_wp,post = wp ~uselet env bhs.bhs_m s_wp bhs.bhs_po in + let s_wp,post = wp ~uselet env bhs.bhs_m s_wp bhs.bhs_po [] in check_wp_progress tc i bhs.bhs_s s_wp; let s = EcModules.stmt (s_hd @ s_wp) in let concl = f_bdHoareS_r { bhs with bhs_s = s; bhs_po = post} in @@ -182,8 +182,8 @@ module TacInternal = struct let meml, s_wpl = es.es_ml, EcModules.stmt s_wpl in let memr, s_wpr = es.es_mr, EcModules.stmt s_wpr in let post = es.es_po in - let s_wpl, post = wp ~uselet env meml s_wpl post in - let s_wpr, post = wp ~uselet env memr s_wpr post in + let s_wpl, post = wp ~uselet env meml s_wpl post [] in + let s_wpr, post = wp ~uselet env memr s_wpr post [] in check_wp_progress tc i es.es_sl s_wpl; check_wp_progress tc j es.es_sr s_wpr; let sl = EcModules.stmt (s_hdl @ s_wpl) in @@ -217,7 +217,7 @@ let t_wp ?(uselet=true) = FApi.t_low1 "wp" (t_wp_r ~uselet) (* -------------------------------------------------------------------- *) let typing_wp env m s f = - match wp ~onesided:true env m s f with + match wp ~onesided:true env m s f [] with | [], f -> Some f | _, _ -> None let () = EcTyping.wp := Some typing_wp