diff --git a/thys/Automated_Stateful_Protocol_Verification/Stateful_Protocol_Model.thy b/thys/Automated_Stateful_Protocol_Verification/Stateful_Protocol_Model.thy --- a/thys/Automated_Stateful_Protocol_Verification/Stateful_Protocol_Model.thy +++ b/thys/Automated_Stateful_Protocol_Verification/Stateful_Protocol_Model.thy @@ -1,9074 +1,9074 @@ (* Title: Stateful_Protocol_Model.thy Author: Andreas Viktor Hess, DTU Author: Sebastian A. Mödersheim, DTU Author: Achim D. Brucker, University of Exeter Author: Anders Schlichtkrull, DTU SPDX-License-Identifier: BSD-3-Clause *) section\Stateful Protocol Model\ theory Stateful_Protocol_Model imports Stateful_Protocol_Composition_and_Typing.Stateful_Compositionality Transactions Term_Abstraction begin subsection \Locale Setup\ locale stateful_protocol_model = fixes arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" assumes Ana\<^sub>f_assm1: "\f. let (K, M) = Ana\<^sub>f f in (\k \ subterms\<^sub>s\<^sub>e\<^sub>t (set K). is_Fun k \ (is_Fu (the_Fun k)) \ length (args k) = arity\<^sub>f (the_Fu (the_Fun k)))" and Ana\<^sub>f_assm2: "\f. let (K, M) = Ana\<^sub>f f in \i \ fv\<^sub>s\<^sub>e\<^sub>t (set K) \ set M. i < arity\<^sub>f f" and public\<^sub>f_assm: "\f. arity\<^sub>f f > (0::nat) \ public\<^sub>f f" and \\<^sub>f_assm: "\f. arity\<^sub>f f = (0::nat) \ \\<^sub>f f \ None" and label_witness_assm: "label_witness1 \ label_witness2" begin lemma Ana\<^sub>f_assm1_alt: assumes "Ana\<^sub>f f = (K,M)" "k \ subterms\<^sub>s\<^sub>e\<^sub>t (set K)" shows "(\x. k = Var x) \ (\h T. k = Fun (Fu h) T \ length T = arity\<^sub>f h)" proof (cases k) case (Fun g T) let ?P = "\k. is_Fun k \ is_Fu (the_Fun k) \ length (args k) = arity\<^sub>f (the_Fu (the_Fun k))" let ?Q = "\K M. \k \ subterms\<^sub>s\<^sub>e\<^sub>t (set K). ?P k" have "?Q (fst (Ana\<^sub>f f)) (snd (Ana\<^sub>f f))" using Ana\<^sub>f_assm1 split_beta[of ?Q "Ana\<^sub>f f"] by meson hence "?Q K M" using assms(1) by simp hence "?P k" using assms(2) by blast thus ?thesis using Fun by (cases g) auto qed simp lemma Ana\<^sub>f_assm2_alt: assumes "Ana\<^sub>f f = (K,M)" "i \ fv\<^sub>s\<^sub>e\<^sub>t (set K) \ set M" shows "i < arity\<^sub>f f" using Ana\<^sub>f_assm2 assms by fastforce subsection \Definitions\ fun arity where "arity (Fu f) = arity\<^sub>f f" | "arity (Set s) = arity\<^sub>s s" | "arity (Val _) = 0" | "arity (Abs _) = 0" | "arity Pair = 2" | "arity (Attack _) = 0" | "arity OccursFact = 2" | "arity OccursSec = 0" | "arity (PubConst _ _) = 0" fun public where "public (Fu f) = public\<^sub>f f" | "public (Set s) = (arity\<^sub>s s > 0)" | "public (Val n) = False" | "public (Abs _) = False" | "public Pair = True" | "public (Attack _) = False" | "public OccursFact = True" | "public OccursSec = False" | "public (PubConst _ _) = True" fun Ana where "Ana (Fun (Fu f) T) = ( if arity\<^sub>f f = length T \ arity\<^sub>f f > 0 then let (K,M) = Ana\<^sub>f f in (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M) else ([], []))" | "Ana _ = ([], [])" definition \\<^sub>v where "\\<^sub>v v \ ( if (\t \ subterms (fst v). case t of (TComp f T) \ arity f > 0 \ arity f = length T | _ \ True) then fst v else TAtom Bottom)" fun \ where "\ (Var v) = \\<^sub>v v" | "\ (Fun f T) = ( if arity f = 0 then case f of (Fu g) \ TAtom (case \\<^sub>f g of Some a \ Atom a | None \ Bottom) | (Val _) \ TAtom Value | (Abs _) \ TAtom AbsValue | (Set _) \ TAtom SetType | (Attack _) \ TAtom AttackType | OccursSec \ TAtom OccursSecType | (PubConst a _) \ TAtom a | _ \ TAtom Bottom else TComp f (map \ T))" lemma \_consts_simps[simp]: "arity\<^sub>f g = 0 \ \ (Fun (Fu g) []::('fun,'atom,'sets,'lbl) prot_term) = TAtom (case \\<^sub>f g of Some a \ Atom a | None \ Bottom)" "\ (Fun (Val n) []::('fun,'atom,'sets,'lbl) prot_term) = TAtom Value" "\ (Fun (Abs b) []::('fun,'atom,'sets,'lbl) prot_term) = TAtom AbsValue" "arity\<^sub>s s = 0 \ \ (Fun (Set s) []::('fun,'atom,'sets,'lbl) prot_term) = TAtom SetType" "\ (Fun (Attack x) []::('fun,'atom,'sets,'lbl) prot_term) = TAtom AttackType" "\ (Fun OccursSec []::('fun,'atom,'sets,'lbl) prot_term) = TAtom OccursSecType" "\ (Fun (PubConst a t) []::('fun,'atom,'sets,'lbl) prot_term) = TAtom a" by simp+ lemma \_Fu_simps[simp]: "arity\<^sub>f f \ 0 \ \ (Fun (Fu f) T) = TComp (Fu f) (map \ T)" (is "?A1 \ ?A2") "arity\<^sub>f f = 0 \ \\<^sub>f f = Some a \ \ (Fun (Fu f) T) = TAtom (Atom a)" (is "?B1 \ ?B2 \ ?B3") "arity\<^sub>f f = 0 \ \\<^sub>f f = None \ \ (Fun (Fu f) T) = TAtom Bottom" (is "?C1 \ ?C2 \ ?C3") "\ (Fun (Fu f) T) \ TAtom Value" (is ?D) "\ (Fun (Fu f) T) \ TAtom AttackType" (is ?E) "\ (Fun (Fu f) T) \ TAtom OccursSecType" (is ?F) proof - show "?A1 \ ?A2" by simp show "?B1 \ ?B2 \ ?B3" by simp show "?C1 \ ?C2 \ ?C3" by simp show ?D by (cases "\\<^sub>f f") simp_all show ?E by (cases "\\<^sub>f f") simp_all show ?F by (cases "\\<^sub>f f") simp_all qed lemma \_Set_simps[simp]: "arity\<^sub>s s \ 0 \ \ (Fun (Set s) T) = TComp (Set s) (map \ T)" "\ (Fun (Set s) T) = TAtom SetType \ \ (Fun (Set s) T) = TComp (Set s) (map \ T)" "\ (Fun (Set s) T) \ TAtom Value" "\ (Fun (Set s) T) \ TAtom (Atom a)" "\ (Fun (Set s) T) \ TAtom AttackType" "\ (Fun (Set s) T) \ TAtom OccursSecType" "\ (Fun (Set s) T) \ TAtom Bottom" by auto subsection \Locale Interpretations\ lemma Ana_Fu_cases: assumes "Ana (Fun f T) = (K,M)" and "f = Fu g" and "Ana\<^sub>f g = (K',M')" shows "(K,M) = (if arity\<^sub>f g = length T \ arity\<^sub>f g > 0 then (K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M') else ([],[]))" (is ?A) and "(K,M) = (K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M') \ (K,M) = ([],[])" (is ?B) proof - show ?A using assms by (cases "arity\<^sub>f g = length T \ arity\<^sub>f g > 0") auto thus ?B by metis qed lemma Ana_Fu_intro: assumes "arity\<^sub>f f = length T" "arity\<^sub>f f > 0" and "Ana\<^sub>f f = (K',M')" shows "Ana (Fun (Fu f) T) = (K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M')" using assms by simp lemma Ana_Fu_elim: assumes "Ana (Fun f T) = (K,M)" and "f = Fu g" and "Ana\<^sub>f g = (K',M')" and "(K,M) \ ([],[])" shows "arity\<^sub>f g = length T" (is ?A) and "(K,M) = (K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M')" (is ?B) proof - show ?A using assms by force moreover have "arity\<^sub>f g > 0" using assms by force ultimately show ?B using assms by auto qed lemma Ana_nonempty_inv: assumes "Ana t \ ([],[])" shows "\f T. t = Fun (Fu f) T \ arity\<^sub>f f = length T \ arity\<^sub>f f > 0 \ (\K M. Ana\<^sub>f f = (K, M) \ Ana t = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M))" using assms proof (induction t rule: Ana.induct) case (1 f T) hence *: "arity\<^sub>f f = length T" "0 < arity\<^sub>f f" "Ana (Fun (Fu f) T) = (case Ana\<^sub>f f of (K, M) \ (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M))" using Ana.simps(1)[of f T] unfolding Let_def by metis+ obtain K M where **: "Ana\<^sub>f f = (K, M)" by (metis surj_pair) hence "Ana (Fun (Fu f) T) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M)" using *(3) by simp thus ?case using ** *(1,2) by blast qed simp_all lemma assm1: assumes "Ana t = (K,M)" shows "fv\<^sub>s\<^sub>e\<^sub>t (set K) \ fv t" using assms proof (induction t rule: term.induct) case (Fun f T) have aux: "fv\<^sub>s\<^sub>e\<^sub>t (set K \\<^sub>s\<^sub>e\<^sub>t (!) T) \ fv\<^sub>s\<^sub>e\<^sub>t (set T)" when K: "\i \ fv\<^sub>s\<^sub>e\<^sub>t (set K). i < length T" for K::"(('fun,'atom,'sets,'lbl) prot_fun, nat) term list" proof fix x assume "x \ fv\<^sub>s\<^sub>e\<^sub>t (set K \\<^sub>s\<^sub>e\<^sub>t (!) T)" then obtain k where k: "k \ set K" "x \ fv (k \ (!) T)" by moura have "\i \ fv k. i < length T" using K k(1) by simp thus "x \ fv\<^sub>s\<^sub>e\<^sub>t (set T)" by (metis (no_types, lifting) k(2) contra_subsetD fv_set_mono image_subsetI nth_mem subst_apply_fv_unfold) qed { fix g assume f: "f = Fu g" and K: "K \ []" obtain K' M' where *: "Ana\<^sub>f g = (K',M')" by moura have "(K, M) \ ([], [])" using K by simp hence "(K, M) = (K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M')" "arity\<^sub>f g = length T" using Ana_Fu_cases(1)[OF Fun.prems f *] by presburger+ hence ?case using aux[of K'] Ana\<^sub>f_assm2_alt[OF *] by auto } thus ?case using Fun by (cases f) fastforce+ qed simp lemma assm2: assumes "Ana t = (K,M)" and "\g S'. Fun g S' \ t \ length S' = arity g" and "k \ set K" and "Fun f T' \ k" shows "length T' = arity f" using assms proof (induction t rule: term.induct) case (Fun g T) obtain h where 2: "g = Fu h" using Fun.prems(1,3) by (cases g) auto obtain K' M' where 1: "Ana\<^sub>f h = (K',M')" by moura have "(K,M) \ ([],[])" using Fun.prems(3) by auto hence "(K,M) = (K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M')" "\i. i \ fv\<^sub>s\<^sub>e\<^sub>t (set K') \ set M' \ i < length T" using Ana_Fu_cases(1)[OF Fun.prems(1) 2 1] Ana\<^sub>f_assm2_alt[OF 1] by presburger+ hence "K = K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T" and 3: "\i\fv\<^sub>s\<^sub>e\<^sub>t (set K'). i < length T" by simp_all then obtain k' where k': "k' \ set K'" "k = k' \ (!) T" using Fun.prems(3) by moura hence 4: "Fun f T' \ subterms (k' \ (!) T)" "fv k' \ fv\<^sub>s\<^sub>e\<^sub>t (set K')" using Fun.prems(4) by auto show ?case proof (cases "\i \ fv k'. Fun f T' \ subterms (T ! i)") case True hence "Fun f T' \ subterms\<^sub>s\<^sub>e\<^sub>t (set T)" using k' Fun.prems(4) 3 by auto thus ?thesis using Fun.prems(2) by auto next case False then obtain S where "Fun f S \ subterms k'" "Fun f T' = Fun f S \ (!) T" using k'(2) Fun.prems(4) subterm_subst_not_img_subterm by force thus ?thesis using Ana\<^sub>f_assm1_alt[OF 1, of "Fun f S"] k'(1) by (cases f) auto qed qed simp lemma assm4: assumes "Ana (Fun f T) = (K, M)" shows "set M \ set T" using assms proof (cases f) case (Fu g) obtain K' M' where *: "Ana\<^sub>f g = (K',M')" by moura have "M = [] \ (arity\<^sub>f g = length T \ M = map ((!) T) M')" using Ana_Fu_cases(1)[OF assms Fu *] by (meson prod.inject) thus ?thesis using Ana\<^sub>f_assm2_alt[OF *] by auto qed auto lemma assm5: "Ana t = (K,M) \ K \ [] \ M \ [] \ Ana (t \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" proof (induction t rule: term.induct) case (Fun f T) thus ?case proof (cases f) case (Fu g) obtain K' M' where *: "Ana\<^sub>f g = (K',M')" by moura have **: "K = K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T" "M = map ((!) T) M'" "arity\<^sub>f g = length T" "\i \ fv\<^sub>s\<^sub>e\<^sub>t (set K') \ set M'. i < arity\<^sub>f g" "0 < arity\<^sub>f g" using Fun.prems(2) Ana_Fu_cases(1)[OF Fun.prems(1) Fu *] Ana\<^sub>f_assm2_alt[OF *] by (meson prod.inject)+ have ***: "\i \ fv\<^sub>s\<^sub>e\<^sub>t (set K'). i < length T" "\i \ set M'. i < length T" using **(3,4) by auto have "K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) (map (\t. t \ \) T)" "M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = map ((!) (map (\t. t \ \) T)) M'" using subst_idx_map[OF ***(2), of \] subst_idx_map'[OF ***(1), of \] **(1,2) by fast+ thus ?thesis using Fu * **(3,5) by auto qed auto qed simp sublocale intruder_model arity public Ana apply unfold_locales by (metis assm1, metis assm2, rule Ana.simps, metis assm4, metis assm5) adhoc_overloading INTRUDER_SYNTH intruder_synth adhoc_overloading INTRUDER_DEDUCT intruder_deduct lemma assm6: "arity c = 0 \ \a. \X. \ (Fun c X) = TAtom a" by (cases c) auto lemma assm7: "0 < arity f \ \ (Fun f T) = TComp f (map \ T)" by auto lemma assm8: "infinite {c. \ (Fun c []::('fun,'atom,'sets,'lbl) prot_term) = TAtom a \ public c}" (is "?P a") proof - let ?T = "\f. (range f)::('fun,'atom,'sets,'lbl) prot_fun set" let ?A = "\f. \x::nat \ UNIV. \y::nat \ UNIV. (f x = f y) = (x = y)" let ?B = "\f. \x::nat \ UNIV. f x \ ?T f" let ?C = "\f. \y::('fun,'atom,'sets,'lbl) prot_fun \ ?T f. \x \ UNIV. y = f x" let ?D = "\f b. ?T f \ {c. \ (Fun c []::('fun,'atom,'sets,'lbl) prot_term) = TAtom b \ public c}" have sub_lmm: "?P b" when "?A f" "?C f" "?C f" "?D f b" for b f proof - have "\g::nat \ ('fun,'atom,'sets,'lbl) prot_fun. bij_betw g UNIV (?T f)" using bij_betwI'[of UNIV f "?T f"] that(1,2,3) by blast hence "infinite (?T f)" by (metis nat_not_finite bij_betw_finite) thus ?thesis using infinite_super[OF that(4)] by blast qed show ?thesis proof (cases a) case (Atom b) thus ?thesis using sub_lmm[of "PubConst (Atom b)" a] by force next case Value thus ?thesis using sub_lmm[of "PubConst Value" a] by force next case SetType thus ?thesis using sub_lmm[of "PubConst SetType" a] by fastforce next case AttackType thus ?thesis using sub_lmm[of "PubConst AttackType" a] by fastforce next case Bottom thus ?thesis using sub_lmm[of "PubConst Bottom" a] by fastforce next case OccursSecType thus ?thesis using sub_lmm[of "PubConst OccursSecType" a] by fastforce next case AbsValue thus ?thesis using sub_lmm[of "PubConst AbsValue" a] by force qed qed lemma assm9: "TComp f T \ \ t \ arity f > 0" proof (induction t rule: term.induct) case (Var x) hence "\ (Var x) \ TAtom Bottom" by force hence "\t \ subterms (fst x). case t of TComp f T \ arity f > 0 \ arity f = length T | _ \ True" using Var \.simps(1)[of x] unfolding \\<^sub>v_def by meson thus ?case using Var by (fastforce simp add: \\<^sub>v_def) next case (Fun g S) have "arity g \ 0" using Fun.prems Var_subtermeq assm6 by force thus ?case using Fun by (cases "TComp f T = TComp g (map \ S)") auto qed lemma assm10: "wf\<^sub>t\<^sub>r\<^sub>m (\ (Var x))" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by (auto simp add: \\<^sub>v_def) lemma assm11: "arity f > 0 \ public f" using public\<^sub>f_assm by (cases f) auto lemma assm12: "\ (Var (\, n)) = \ (Var (\, m))" by (simp add: \\<^sub>v_def) lemma assm13: "arity c = 0 \ Ana (Fun c T) = ([],[])" by (cases c) simp_all lemma assm14: assumes "Ana (Fun f T) = (K,M)" shows "Ana (Fun f T \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" proof - show ?thesis proof (cases "(K, M) = ([],[])") case True { fix g assume f: "f = Fu g" obtain K' M' where "Ana\<^sub>f g = (K',M')" by moura hence ?thesis using assms f True by auto } thus ?thesis using True assms by (cases f) auto next case False then obtain g where **: "f = Fu g" using assms by (cases f) auto obtain K' M' where *: "Ana\<^sub>f g = (K',M')" by moura have ***: "K = K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T" "M = map ((!) T) M'" "arity\<^sub>f g = length T" "\i \ fv\<^sub>s\<^sub>e\<^sub>t (set K') \ set M'. i < arity\<^sub>f g" using Ana_Fu_cases(1)[OF assms ** *] False Ana\<^sub>f_assm2_alt[OF *] by (meson prod.inject)+ have ****: "\i\fv\<^sub>s\<^sub>e\<^sub>t (set K'). i < length T" "\i\set M'. i < length T" using ***(3,4) by auto have "K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) (map (\t. t \ \) T)" "M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = map ((!) (map (\t. t \ \) T)) M'" using subst_idx_map[OF ****(2), of \] subst_idx_map'[OF ****(1), of \] ***(1,2) by auto thus ?thesis using assms * ** ***(3) by auto qed qed sublocale labeled_stateful_typing' arity public Ana \ Pair label_witness1 label_witness2 apply unfold_locales subgoal by (metis assm6) subgoal by (metis assm7) subgoal by (metis assm9) subgoal by (rule assm10) subgoal by (metis assm12) subgoal by (metis assm13) subgoal by (metis assm14) subgoal by (rule label_witness_assm) subgoal by (rule arity.simps(5)) subgoal by (metis assm14) subgoal by (metis assm8) subgoal by (metis assm11) done subsection \The Protocol Transition System, Defined in Terms of the Reachable Constraints\ definition transaction_decl_subst where "transaction_decl_subst (\::('fun,'atom,'sets,'lbl) prot_subst) T \ subst_domain \ = fst ` set (transaction_decl T ()) \ (\(x,cs) \ set (transaction_decl T ()). \c \ cs. \ x = Fun (Fu c) [] \ arity (Fu c::('fun,'atom,'sets,'lbl) prot_fun) = 0) \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" definition transaction_fresh_subst where "transaction_fresh_subst \ T M \ subst_domain \ = set (transaction_fresh T) \ (\t \ subst_range \. \c. t = Fun c [] \ \public c \ arity c = 0) \ (\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t M) \ (\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction T)) \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ inj_on \ (subst_domain \)" (* NB: We need the protocol P as a parameter for this definition---even though we will only apply \ to a single transaction T of P---because we have to ensure that \(fv(T)) is disjoint from the bound variables of P and \. *) definition transaction_renaming_subst where "transaction_renaming_subst \ P X \ \n \ max_var_set (\(vars_transaction ` set P) \ X). \ = var_rename n" definition (in intruder_model) constraint_model where "constraint_model \ \ \ constr_sem_stateful \ (unlabel \) \ interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" definition (in typed_model) welltyped_constraint_model where "welltyped_constraint_model \ \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ constraint_model \ \" text \ The set of symbolic constraints reachable in any symbolic run of the protocol \P\. \\\ instantiates the "declared variables" of transaction \T\ with ground terms. \\\ instantiates the fresh variables of transaction \T\ with fresh terms. \\\ is a variable-renaming whose range consists of fresh variables. \ inductive_set reachable_constraints:: "('fun,'atom,'sets,'lbl) prot \ ('fun,'atom,'sets,'lbl) prot_constr set" for P::"('fun,'atom,'sets,'lbl) prot" where init[simp]: "[] \ reachable_constraints P" | step: "\\ \ reachable_constraints P; T \ set P; transaction_decl_subst \ T; transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \); transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ \ \@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \) \ reachable_constraints P" subsection \Minor Lemmata\ lemma \\<^sub>v_TAtom[simp]: "\\<^sub>v (TAtom a, n) = TAtom a" unfolding \\<^sub>v_def by simp lemma \\<^sub>v_TAtom': assumes "a \ Bottom" shows "\\<^sub>v (\, n) = TAtom a \ \ = TAtom a" proof assume "\\<^sub>v (\, n) = TAtom a" thus "\ = TAtom a" by (metis (no_types, lifting) assms \\<^sub>v_def fst_conv term.inject(1)) qed simp lemma \\<^sub>v_TAtom_inv: "\\<^sub>v x = TAtom (Atom a) \ \m. x = (TAtom (Atom a), m)" "\\<^sub>v x = TAtom Value \ \m. x = (TAtom Value, m)" "\\<^sub>v x = TAtom SetType \ \m. x = (TAtom SetType, m)" "\\<^sub>v x = TAtom AttackType \ \m. x = (TAtom AttackType, m)" "\\<^sub>v x = TAtom OccursSecType \ \m. x = (TAtom OccursSecType, m)" by (metis \\<^sub>v_TAtom' surj_pair prot_atom.distinct(7), metis \\<^sub>v_TAtom' surj_pair prot_atom.distinct(18), metis \\<^sub>v_TAtom' surj_pair prot_atom.distinct(26), metis \\<^sub>v_TAtom' surj_pair prot_atom.distinct(32), metis \\<^sub>v_TAtom' surj_pair prot_atom.distinct(38)) lemma \\<^sub>v_TAtom'': "(fst x = TAtom (Atom a)) = (\\<^sub>v x = TAtom (Atom a))" (is "?A = ?A'") "(fst x = TAtom Value) = (\\<^sub>v x = TAtom Value)" (is "?B = ?B'") "(fst x = TAtom SetType) = (\\<^sub>v x = TAtom SetType)" (is "?C = ?C'") "(fst x = TAtom AttackType) = (\\<^sub>v x = TAtom AttackType)" (is "?D = ?D'") "(fst x = TAtom OccursSecType) = (\\<^sub>v x = TAtom OccursSecType)" (is "?E = ?E'") proof - have 1: "?A \ ?A'" "?B \ ?B'" "?C \ ?C'" "?D \ ?D'" "?E \ ?E'" by (metis \\<^sub>v_TAtom prod.collapse)+ have 2: "?A' \ ?A" "?B' \ ?B" "?C' \ ?C" "?D' \ ?D" "?E' \ ?E" using \\<^sub>v_TAtom \\<^sub>v_TAtom_inv(1) apply fastforce using \\<^sub>v_TAtom \\<^sub>v_TAtom_inv(2) apply fastforce using \\<^sub>v_TAtom \\<^sub>v_TAtom_inv(3) apply fastforce using \\<^sub>v_TAtom \\<^sub>v_TAtom_inv(4) apply fastforce using \\<^sub>v_TAtom \\<^sub>v_TAtom_inv(5) by fastforce show "?A = ?A'" "?B = ?B'" "?C = ?C'" "?D = ?D'" "?E = ?E'" using 1 2 by metis+ qed lemma \\<^sub>v_Var_image: "\\<^sub>v ` X = \ ` Var ` X" by force lemma \_Fu_const: assumes "arity\<^sub>f g = 0" shows "\a. \ (Fun (Fu g) T) = TAtom (Atom a)" proof - have "\\<^sub>f g \ None" using assms \\<^sub>f_assm by blast thus ?thesis using assms by force qed lemma Fun_Value_type_inv: fixes T::"('fun,'atom,'sets,'lbl) prot_term list" assumes "\ (Fun f T) = TAtom Value" shows "(\n. f = Val n) \ (\bs. f = Abs bs) \ (\n. f = PubConst Value n)" proof - have *: "arity f = 0" by (metis const_type_inv assms) show ?thesis using assms proof (cases f) case (Fu g) hence "arity\<^sub>f g = 0" using * by simp hence False using Fu \_Fu_const[of g T] assms by auto thus ?thesis by metis next case (Set s) hence "arity\<^sub>s s = 0" using * by simp hence False using Set assms by auto thus ?thesis by metis qed simp_all qed lemma Ana\<^sub>f_keys_not_val_terms: assumes "Ana\<^sub>f f = (K, T)" and "k \ set K" and "g \ funs_term k" shows "\is_Val g" and "\is_PubConstValue g" and "\is_Abs g" proof - { assume "is_Val g" then obtain n S where *: "Fun (Val n) S \ subterms\<^sub>s\<^sub>e\<^sub>t (set K)" using assms(2) funs_term_Fun_subterm[OF assms(3)] by (cases g) auto hence False using Ana\<^sub>f_assm1_alt[OF assms(1) *] by simp } moreover { assume "is_PubConstValue g" then obtain n S where *: "Fun (PubConst Value n) S \ subterms\<^sub>s\<^sub>e\<^sub>t (set K)" using assms(2) funs_term_Fun_subterm[OF assms(3)] unfolding is_PubConstValue_def by (cases g) auto hence False using Ana\<^sub>f_assm1_alt[OF assms(1) *] by simp } moreover { assume "is_Abs g" then obtain a S where *: "Fun (Abs a) S \ subterms\<^sub>s\<^sub>e\<^sub>t (set K)" using assms(2) funs_term_Fun_subterm[OF assms(3)] by (cases g) auto hence False using Ana\<^sub>f_assm1_alt[OF assms(1) *] by simp } ultimately show "\is_Val g" "\is_PubConstValue g" "\is_Abs g" by metis+ qed lemma Ana\<^sub>f_keys_not_pairs: assumes "Ana\<^sub>f f = (K, T)" and "k \ set K" and "g \ funs_term k" shows "g \ Pair" proof assume "g = Pair" then obtain S where *: "Fun Pair S \ subterms\<^sub>s\<^sub>e\<^sub>t (set K)" using assms(2) funs_term_Fun_subterm[OF assms(3)] by (cases g) auto show False using Ana\<^sub>f_assm1_alt[OF assms(1) *] by simp qed lemma Ana_Fu_keys_funs_term_subset: fixes K::"('fun,'atom,'sets,'lbl) prot_term list" assumes "Ana (Fun (Fu f) S) = (K, T)" and "Ana\<^sub>f f = (K', T')" shows "\(funs_term ` set K) \ \(funs_term ` set K') \ funs_term (Fun (Fu f) S)" proof - { fix k assume k: "k \ set K" then obtain k' where k': "k' \ set K'" "k = k' \ (!) S" "arity\<^sub>f f = length S" "subterms k' \ subterms\<^sub>s\<^sub>e\<^sub>t (set K')" using assms Ana_Fu_elim[OF assms(1) _ assms(2)] by fastforce have 1: "funs_term k' \ \(funs_term ` set K')" using k'(1) by auto have "i < length S" when "i \ fv k'" for i using that Ana\<^sub>f_assm2_alt[OF assms(2), of i] k'(1,3) by auto hence 2: "funs_term (S ! i) \ funs_term (Fun (Fu f) S)" when "i \ fv k'" for i using that by force have "funs_term k \ \(funs_term ` set K') \ funs_term (Fun (Fu f) S)" using funs_term_subst[of k' "(!) S"] k'(2) 1 2 by fast } thus ?thesis by blast qed lemma Ana_Fu_keys_not_pubval_terms: fixes k::"('fun,'atom,'sets,'lbl) prot_term" assumes "Ana (Fun (Fu f) S) = (K, T)" and "Ana\<^sub>f f = (K', T')" and "k \ set K" and "\g \ funs_term (Fun (Fu f) S). \is_PubConstValue g" shows "\g \ funs_term k. \is_PubConstValue g" using assms(3,4) Ana\<^sub>f_keys_not_val_terms(1,2)[OF assms(2)] Ana_Fu_keys_funs_term_subset[OF assms(1,2)] by blast lemma Ana_Fu_keys_not_abs_terms: fixes k::"('fun,'atom,'sets,'lbl) prot_term" assumes "Ana (Fun (Fu f) S) = (K, T)" and "Ana\<^sub>f f = (K', T')" and "k \ set K" and "\g \ funs_term (Fun (Fu f) S). \is_Abs g" shows "\g \ funs_term k. \is_Abs g" using assms(3,4) Ana\<^sub>f_keys_not_val_terms(3)[OF assms(2)] Ana_Fu_keys_funs_term_subset[OF assms(1,2)] by blast lemma Ana_Fu_keys_not_pairs: fixes k::"('fun,'atom,'sets,'lbl) prot_term" assumes "Ana (Fun (Fu f) S) = (K, T)" and "Ana\<^sub>f f = (K', T')" and "k \ set K" and "\g \ funs_term (Fun (Fu f) S). g \ Pair" shows "\g \ funs_term k. g \ Pair" using assms(3,4) Ana\<^sub>f_keys_not_pairs[OF assms(2)] Ana_Fu_keys_funs_term_subset[OF assms(1,2)] by blast lemma Ana_Fu_keys_length_eq: assumes "length T = length S" shows "length (fst (Ana (Fun (Fu f) T))) = length (fst (Ana (Fun (Fu f) S)))" proof (cases "arity\<^sub>f f = length T \ arity\<^sub>f f > 0") case True thus ?thesis using assms by (cases "Ana\<^sub>f f") auto next case False thus ?thesis using assms by force qed lemma Ana_key_PubConstValue_subterm_in_term: fixes k::"('fun,'atom,'sets,'lbl) prot_term" assumes KR: "Ana t = (K, R)" and k: "k \ set K" and n: "Fun (PubConst Value n) [] \ k" shows "Fun (PubConst Value n) [] \ t" proof (cases t) case (Var x) thus ?thesis using KR k n by force next case (Fun f ts) note t = this then obtain g where f: "f = Fu g" using KR k by (cases f) auto obtain K' R' where KR': "Ana\<^sub>f g = (K', R')" by fastforce have K: "K = K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) ts" using k Ana_Fu_elim(2)[OF KR[unfolded t] f KR'] by force obtain k' where k': "k' \ set K'" "k = k' \ (!) ts" using k K by auto have 0: "\(Fun (PubConst Value n) [] \ k')" proof assume *: "Fun (PubConst Value n) [] \ k'" have **: "PubConst Value n \ funs_term k'" using funs_term_Fun_subterm'[OF *] by (cases k') auto show False using Ana\<^sub>f_keys_not_val_terms(2)[OF KR' k'(1) **] unfolding is_PubConstValue_def by force qed hence "\i \ fv k'. Fun (PubConst Value n) [] \ ts ! i" by (metis n const_subterm_subst_var_obtain k'(2)) then obtain i where i: "i \ fv k'" "Fun (PubConst Value n) [] \ ts ! i" by blast have "i < length ts" using i(1) KR' k'(1) Ana\<^sub>f_assm2_alt[OF KR', of i] Ana_Fu_elim(1)[OF KR[unfolded t] f KR'] k by fastforce thus ?thesis using i(2) unfolding t by force qed lemma deduct_occurs_in_ik: fixes t::"('fun,'atom,'sets,'lbl) prot_term" assumes t: "M \ occurs t" and M: "\s \ subterms\<^sub>s\<^sub>e\<^sub>t M. OccursFact \ \(funs_term ` set (snd (Ana s)))" "\s \ subterms\<^sub>s\<^sub>e\<^sub>t M. OccursSec \ \(funs_term ` set (snd (Ana s)))" "Fun OccursSec [] \ M" shows "occurs t \ M" using private_fun_deduct_in_ik''[of M OccursFact "[Fun OccursSec [], t]" OccursSec] t M by fastforce lemma deduct_val_const_swap: fixes \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "M \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" and "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. (\n. \ x = Fun (Val n) []) \ (\n. \ x = Fun (PubConst Value n) [])" and "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. (\n. \ x = Fun (Val n) [])" and "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. (\n. \ x = Fun (PubConst Value n) []) \ \ x \ M \ N" and "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. (\n. \ x = Fun (Val n) []) \ \ x = \ x" and "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. \y \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. \ x = \ y \ \ x = \ y" and "\n. \(Fun (PubConst Value n) [] \\<^sub>s\<^sub>e\<^sub>t insert t M)" shows "(M \\<^sub>s\<^sub>e\<^sub>t \) \ N \ t \ \" proof - obtain n where n: "intruder_deduct_num (M \\<^sub>s\<^sub>e\<^sub>t \) n (t \ \)" using assms(1) deduct_num_if_deduct by blast hence "\m \ n. intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) m (t \ \)" using assms(2-) proof (induction n arbitrary: t rule: nat_less_induct) case (1 n) note prems = "1.prems" note IH = "1.IH" show ?case proof (cases "t \ \ \ M \\<^sub>s\<^sub>e\<^sub>t \") case True note 2 = this have 3: "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. \c. \ x = Fun c []" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. \c. \ x = Fun c []" using prems(2,3) by (blast, blast) have "t \ \ \ M \\<^sub>s\<^sub>e\<^sub>t \" using subst_const_swap_eq_mem[OF 2 _ 3 prems(6)] prems(2,5,7) by metis thus ?thesis using intruder_deduct_num.AxiomN by auto next case False then obtain n' where n: "n = Suc n'" using prems(1) deduct_zero_in_ik by (cases n) fast+ have M_subterms_eq: "subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) = subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" "subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) = subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" subgoal using prems(2) subterms_subst''[of M \] by blast subgoal using prems(3) subterms_subst''[of M \] by blast done from deduct_inv[OF prems(1)] show ?thesis proof (elim disjE) assume "t \ \ \ M \\<^sub>s\<^sub>e\<^sub>t \" thus ?thesis using False by argo next assume "\f ts. t \ \ = Fun f ts \ public f \ length ts = arity f \ (\t \ set ts. \l < n. intruder_deduct_num (M \\<^sub>s\<^sub>e\<^sub>t \) l t)" then obtain f ts where t: "t \ \ = Fun f ts" "public f" "length ts = arity f" "\t \ set ts. \l < n. intruder_deduct_num (M \\<^sub>s\<^sub>e\<^sub>t \) l t" by blast show ?thesis proof (cases t) case (Var x) hence ts: "ts = []" and f: "\c. f = PubConst Value c" using t(1,2) prems(2) by (force, auto) have "\ x \ M \ N" using prems(4) Var f ts t(1) by auto moreover have "fv (\ x) = {}" using prems(3) Var by auto hence "\ x \ M \\<^sub>s\<^sub>e\<^sub>t \" when "\ x \ M" using that subst_ground_ident[of "\ x" \] by force ultimately have "\ x \ (M \\<^sub>s\<^sub>e\<^sub>t \) \ N" by fast thus ?thesis using intruder_deduct_num.AxiomN Var by force next case (Fun g ss) hence f: "f = g" and ts: "ts = ss \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using t(1) by auto have ss: "\l < n. intruder_deduct_num (M \\<^sub>s\<^sub>e\<^sub>t \) l (s \ \)" when s: "s \ set ss" for s using t(4) ts s by auto have IH': "\l < n. intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) l (s \ \)" when s: "s \ set ss" for s proof - obtain l where l: "l < n" "intruder_deduct_num (M \\<^sub>s\<^sub>e\<^sub>t \) l (s \ \)" using ss s by blast have *: "fv s \ fv t" "subterms\<^sub>s\<^sub>e\<^sub>t (insert s M) \ subterms\<^sub>s\<^sub>e\<^sub>t (insert t M)" using s unfolding Fun f ts by auto have "\l' \ l. intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) l' (s \ \)" proof - have "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv s. (\n. \ x = Fun (Val n) []) \ (\n. \ x = Fun (PubConst Value n) [])" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv s. \n. \ x = Fun (Val n) []" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv s. (\n. \ x = Fun (PubConst Value n) []) \ \ x \ M \ N" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv s. (\n. \ x = Fun (Val n) []) \ \ x = \ x" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv s. \y \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv s. \ x = \ y \ \ x = \ y" "\n. Fun (PubConst Value n) [] \ subterms\<^sub>s\<^sub>e\<^sub>t (insert s M)" subgoal using prems(2) *(1) by blast subgoal using prems(3) *(1) by blast subgoal using prems(4) *(1) by blast subgoal using prems(5) *(1) by blast subgoal using prems(6) *(1) by blast subgoal using prems(7) *(2) by blast done thus ?thesis using IH l by presburger qed then obtain l' where l': "l' \ l" "intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) l' (s \ \)" by blast have "l' < n" using l'(1) l(1) by linarith thus ?thesis using l'(2) by blast qed have g: "length (ss \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) = arity g" "public g" using t(2,3) unfolding f ts by auto let ?P = "\s l. l < n \ intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) l s" define steps where "steps \ \s. SOME l. ?P s l" have 2: "steps (s \ \) < n" "intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) (steps (s \ \)) (s \ \)" when s: "s \ set ss" for s using someI_ex[OF IH'[OF s]] unfolding steps_def by (blast, blast) have 3: "Suc (Max (insert 0 (steps ` set (ss \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)))) \ n" proof (cases "ss = []") case True show ?thesis unfolding True n by simp next case False thus ?thesis using 2 Max_nat_finite_lt[of "set (ss \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" steps n] by (simp add: Suc_leI) qed show ?thesis using intruder_deduct_num.ComposeN[OF g, of "(M \\<^sub>s\<^sub>e\<^sub>t \) \ N" steps] 2(2) 3 unfolding Fun by auto qed next assume "\s \ subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \). (\l < n. intruder_deduct_num (M \\<^sub>s\<^sub>e\<^sub>t \) l s) \ (\k \ set (fst (Ana s)). \l < n. intruder_deduct_num (M \\<^sub>s\<^sub>e\<^sub>t \) l k) \ t \ \ \ set (snd (Ana s))" then obtain s l where s: "s \ subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" "\k \ set (fst (Ana s)). \l < n. intruder_deduct_num (M \\<^sub>s\<^sub>e\<^sub>t \) l k" "t \ \ \ set (snd (Ana s))" and l: "l < n" "intruder_deduct_num (M \\<^sub>s\<^sub>e\<^sub>t \) l s" by (metis (no_types, lifting) M_subterms_eq(1)) obtain u where u: "u \\<^sub>s\<^sub>e\<^sub>t M" "s = u \ \" using s(1) by blast have u_fv: "fv u \ fv\<^sub>s\<^sub>e\<^sub>t M" by (metis fv_subset_subterms u(1)) have "\x. u = Var x" proof assume "\x. u = Var x" then obtain x where x: "u = Var x" by blast then obtain c where c: "s = Fun c []" using u prems(2) u_fv by auto thus False using s(3) Ana_subterm by (cases "Ana s") force qed then obtain f ts where u': "u = Fun f ts" by (cases u) auto obtain K R where KR: "Ana u = (K,R)" by (metis surj_pair) have KR': "Ana s = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, R \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using KR Ana_subst'[OF KR[unfolded u'], of \] unfolding u(2) u' by blast hence s': "\k \ set K. \l < n. intruder_deduct_num (M \\<^sub>s\<^sub>e\<^sub>t \) l (k \ \)" "t \ \ \ set (R \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using s(2,3) by auto have IH1: "\l < n. intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) l (u \ \)" proof - have "subterms u \ subterms\<^sub>s\<^sub>e\<^sub>t M" using u(1) subterms_subset by auto hence "subterms\<^sub>s\<^sub>e\<^sub>t (insert u M) = subterms\<^sub>s\<^sub>e\<^sub>t M" by blast hence *: "subterms\<^sub>s\<^sub>e\<^sub>t (insert u M) \ subterms\<^sub>s\<^sub>e\<^sub>t (insert t M)" by auto have "\l' \ l. intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) l' (u \ \)" proof - have "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv u. (\n. \ x = Fun (Val n) []) \ (\n. \ x = Fun (PubConst Value n) [])" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv u. \n. \ x = Fun (Val n) []" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv u. (\n. \ x = Fun (PubConst Value n) []) \ \ x \ M \ N" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv u. (\n. \ x = Fun (Val n) []) \ \ x = \ x" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv u. \y \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv u. \ x = \ y \ \ x = \ y" "\n. Fun (PubConst Value n) [] \ subterms\<^sub>s\<^sub>e\<^sub>t (insert u M)" subgoal using prems(2) u_fv by blast subgoal using prems(3) u_fv by blast subgoal using prems(4) u_fv by blast subgoal using prems(5) u_fv by blast subgoal using prems(6) u_fv by blast subgoal using prems(7) * by blast done thus ?thesis using IH l unfolding u(2) by presburger qed then obtain l' where l': "l' \ l" "intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) l' (u \ \)" by blast have "l' < n" using l'(1) l(1) by linarith thus ?thesis using l'(2) by blast qed have IH2: "\l < n. intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) l (k \ \)" when k: "k \ set K" for k using k IH prems(2-) Ana\<^sub>f_keys_not_val_terms s'(1) KR u(1) proof - have *: "fv k \ fv\<^sub>s\<^sub>e\<^sub>t M" using k KR Ana_keys_fv u(1) fv_subset_subterms by blast have **: "Fun (PubConst Value n) [] \\<^sub>s\<^sub>e\<^sub>t M" when "Fun (PubConst Value n) [] \ k" for n using in_subterms_subset_Union[OF u(1)] Ana_key_PubConstValue_subterm_in_term[OF KR k that] by fast obtain lk where lk: "lk < n" "intruder_deduct_num (M \\<^sub>s\<^sub>e\<^sub>t \) lk (k \ \)" using s'(1) k by fast have "\l' \ lk. intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) l' (k \ \)" proof - have "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv k. (\n. \ x = Fun (Val n) []) \ (\n. \ x = Fun (PubConst Value n) [])" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv k. \n. \ x = Fun (Val n) []" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv k. (\n. \ x = Fun (PubConst Value n) []) \ \ x \ M \ N" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv k. (\n. \ x = Fun (Val n) []) \ \ x = \ x" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv k. \y \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv k. \ x = \ y \ \ x = \ y" "\n. Fun (PubConst Value n) [] \ subterms\<^sub>s\<^sub>e\<^sub>t (insert k M)" subgoal using prems(2) * by blast subgoal using prems(3) * by blast subgoal using prems(4) * by blast subgoal using prems(5) * by blast subgoal using prems(6) * by blast subgoal using prems(7) ** by blast done thus ?thesis using IH lk by presburger qed then obtain lk' where lk': "lk' \ lk" "intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) lk' (k \ \)" by blast have "lk' < n" using lk'(1) lk(1) by linarith thus ?thesis using lk'(2) by blast qed have KR'': "Ana (u \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, R \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Ana_subst' KR unfolding u' by blast obtain r where r: "r \ set R" "t \ \ = r \ \" using s'(2) by fastforce have r': "t \ \ \ set (R \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" proof - have r_subterm_u: "r \ u" using r(1) KR Ana_subterm by blast have r_fv: "fv r \ fv\<^sub>s\<^sub>e\<^sub>t M" by (meson r_subterm_u u(1) fv_subset_subterms in_mono in_subterms_subset_Union) have t_subterms_M: "subterms t \ subterms\<^sub>s\<^sub>e\<^sub>t (insert t M)" by blast have r_subterm_M: "subterms r \ subterms\<^sub>s\<^sub>e\<^sub>t (insert t M)" using subterms_subset[OF r_subterm_u] in_subterms_subset_Union[OF u(1)] by (auto intro: subterms\<^sub>s\<^sub>e\<^sub>t_mono) have *: "\x \ fv t \ fv r. \ x = \ x \ \(\ x \ t) \ \(\ x \ r)" proof fix x assume "x \ fv t \ fv r" hence "x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t" using r_fv by blast thus "\ x = \ x \ \(\ x \ t) \ \(\ x \ r)" using prems(2,5,7) r_subterm_M t_subterms_M by (metis (no_types, opaque_lifting) in_mono) qed have **: "\x \ fv t \ fv r. \c. \ x = Fun c []" "\x \ fv t \ fv r. \c. \ x = Fun c []" "\x \ fv t \ fv r. \y \ fv t \ fv r. \ x = \ y \ \ x = \ y" subgoal using prems(2) r_fv by blast subgoal using prems(3) r_fv by blast subgoal using prems(6) r_fv by blast done have "t \ \ = r \ \" by (rule subst_const_swap_eq'[OF r(2) * **]) thus ?thesis using r(1) by simp qed obtain l1 where l1: "l1 < n" "intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) l1 (u \ \)" using IH1 by blast let ?P = "\s l. l < n \ intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) l s" define steps where "steps \ \s. SOME l. ?P s l" have 2: "steps (k \ \) < n" "intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) (steps (k \ \)) (k \ \)" when k: "k \ set K" for k using someI_ex[OF IH2[OF k]] unfolding steps_def by (blast, blast) have 3: "Suc (Max (insert l1 (steps ` set (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)))) \ n" proof (cases "K = []") case True show ?thesis using l1(1) unfolding True n by simp next case False thus ?thesis using l1(1) 2 Max_nat_finite_lt[of "set (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" steps n] by (simp add: Suc_leI) qed have IH2': "intruder_deduct_num ((M \\<^sub>s\<^sub>e\<^sub>t \) \ N) (steps k) k" when k: "k \ set (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" for k using IH2 k 2 by auto show ?thesis using l1(1) intruder_deduct_num.DecomposeN[OF l1(2) KR'' IH2' r'] 3 by fast qed qed qed thus ?thesis using deduct_if_deduct_num by blast qed lemma constraint_model_Nil: assumes I: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" shows "constraint_model I []" using I unfolding constraint_model_def by simp lemma welltyped_constraint_model_Nil: assumes I: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" shows "welltyped_constraint_model I []" using I(1) constraint_model_Nil[OF I(2,3)] unfolding welltyped_constraint_model_def by simp lemma constraint_model_prefix: assumes "constraint_model I (A@B)" shows "constraint_model I A" by (metis assms strand_sem_append_stateful unlabel_append constraint_model_def) lemma welltyped_constraint_model_prefix: assumes "welltyped_constraint_model I (A@B)" shows "welltyped_constraint_model I A" by (metis assms constraint_model_prefix welltyped_constraint_model_def) lemma welltyped_constraint_model_deduct_append: assumes "welltyped_constraint_model I A" and "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ s \ I" shows "welltyped_constraint_model I (A@[(l,send\[s]\)])" using assms strand_sem_append_stateful[of "{}" "{}" "unlabel A" _ I] unfolding welltyped_constraint_model_def constraint_model_def by simp lemma welltyped_constraint_model_deduct_split: assumes "welltyped_constraint_model I (A@[(l,send\[s]\)])" shows "welltyped_constraint_model I A" and "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ s \ I" using assms strand_sem_append_stateful[of "{}" "{}" "unlabel A" _ I] unfolding welltyped_constraint_model_def constraint_model_def by simp_all lemma welltyped_constraint_model_deduct_iff: "welltyped_constraint_model I (A@[(l,send\[s]\)]) \ welltyped_constraint_model I A \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ s \ I" by (metis welltyped_constraint_model_deduct_append welltyped_constraint_model_deduct_split) lemma welltyped_constraint_model_attack_if_receive_attack: assumes I: "welltyped_constraint_model \ \" and rcv_attack: "receive\ts\ \ set (unlabel \)" "attack\n\ \ set ts" shows "welltyped_constraint_model \ (\@[(l, send\[attack\n\]\)])" proof - have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ attack\n\" using rcv_attack in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of "attack\n\" "unlabel \"] ideduct_subst[OF intruder_deduct.Axiom[of "attack\n\" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"], of \] by auto thus ?thesis using I strand_sem_append_stateful[of "{}" "{}" "unlabel \" "[send\[attack\n\]\]" \] unfolding welltyped_constraint_model_def constraint_model_def by auto qed lemma constraint_model_Val_is_Value_term: assumes "welltyped_constraint_model I A" and "t \ I = Fun (Val n) []" shows "t = Fun (Val n) [] \ (\m. t = Var (TAtom Value, m))" proof - have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" using assms(1) unfolding welltyped_constraint_model_def by simp moreover have "\ (Fun (Val n) []) = TAtom Value" by auto ultimately have *: "\ t = TAtom Value" by (metis (no_types) assms(2) wt_subst_trm'') show ?thesis proof (cases t) case (Var x) obtain \ m where x: "x = (\, m)" by (metis surj_pair) have "\\<^sub>v x = TAtom Value" using * Var by auto hence "\ = TAtom Value" using x \\<^sub>v_TAtom'[of Value \ m] by simp thus ?thesis using x Var by metis next case (Fun f T) thus ?thesis using assms(2) by auto qed qed lemma wellformed_transaction_sem_receives: fixes T::"('fun,'atom,'sets,'lbl) prot_transaction" assumes T_valid: "wellformed_transaction T" and \: "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \" and s: "receive\ts\ \ set (unlabel (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" shows "\t \ set ts. IK \ t \ \" proof - let ?R = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?S = "\A. unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?S' = "?S (transaction_receive T)" obtain l B s where B: "(l,send\ts\) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "prefix ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]) (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using s dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(2)[of ts "transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_set_prefix_obtain_subst[of "send\ts\" "transaction_receive T" \] by blast have 1: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \])) = unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))@[send\ts\]" using B(1) unlabel_append dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst singleton_lst_proj(4) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_snoc subst_lsst_append subst_lsst_singleton by (metis (no_types, lifting) subst_apply_labeled_stateful_strand_step.simps) have "strand_sem_stateful IK DB ?S' \" using \ strand_sem_append_stateful[of IK DB _ _ \] transaction_dual_subst_unfold[of T \] by fastforce hence "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))@[send\ts\]) \" using B 1 unfolding prefix_def unlabel_def by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def map_append strand_sem_append_stateful) hence t_deduct: "\t \ set ts. IK \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \\<^sub>s\<^sub>e\<^sub>t \) \ t \ \" using strand_sem_append_stateful[of IK DB "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "[send\ts\]" \] by simp have "\s \ set (unlabel (transaction_receive T)). \t. s = receive\t\" using T_valid wellformed_transaction_unlabel_cases(1)[OF T_valid] by auto moreover { fix A::"('fun,'atom,'sets,'lbl) prot_strand" and \ assume "\s \ set (unlabel A). \t. s = receive\t\" hence "\s \ set (unlabel (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)). \t. s = receive\t\" proof (induction A) case (Cons a A) thus ?case using subst_lsst_cons[of a A \] by (cases a) auto qed simp hence "\s \ set (unlabel (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)). \t. s = receive\t\" by (simp add: list.pred_set is_Receive_def) hence "\s \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))). \t. s = send\t\" by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_memberD dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_inv(2) unlabel_in unlabel_mem_has_label) } ultimately have "\s \ set ?R. \ts. s = send\ts\" by simp hence "ik\<^sub>s\<^sub>s\<^sub>t ?R = {}" unfolding unlabel_def ik\<^sub>s\<^sub>s\<^sub>t_def by fast hence "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = {}" using B(2) 1 ik\<^sub>s\<^sub>s\<^sub>t_append dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append by (metis (no_types, lifting) Un_empty map_append prefix_def unlabel_def) thus ?thesis using t_deduct by simp qed lemma wellformed_transaction_sem_pos_checks: assumes T_valid: "wellformed_transaction T" and \: "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \" shows "\ac: t \ u\ \ set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ (t \ \, u \ \) \ DB" and "\ac: t \ u\ \ set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ t \ \ = u \ \" proof - let ?s = "\ac: t \ u\" let ?s' = "\ac: t \ u\" let ?C = "set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?R = "transaction_receive T@transaction_checks T" let ?R' = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?S = "\A. unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?S' = "?S (transaction_receive T)@?S (transaction_checks T)" let ?P = "\a. is_Receive a \ is_Check_or_Assignment a" let ?Q = "\a. is_Send a \ is_Check_or_Assignment a" let ?dbupd = "\B. dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \ DB" have s_in: "?s \ ?C \ ?s \ set (unlabel (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "?s' \ ?C \ ?s' \ set (unlabel (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using subst_lsst_append[of "transaction_receive T"] unlabel_append[of "transaction_receive T"] by auto have 1: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \])) = unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))@[s']" when B: "(l,s') = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "s' = \ac: t \ u\ \ s' = \ac: t \ u\" for l s s' and B::"('fun,'atom,'sets,'lbl) prot_strand" using B unlabel_append dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst singleton_lst_proj(4) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_snoc subst_lsst_append subst_lsst_singleton by (metis (no_types, lifting) subst_apply_labeled_stateful_strand_step.simps) have 2: "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))@[s']) \" when B: "(l,s') = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "prefix ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]) (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "s' = \ac: t \ u\ \ s' = \ac: t \ u\" for l s s' and B::"('fun,'atom,'sets,'lbl) prot_strand" proof - have "strand_sem_stateful IK DB ?S' \" using \ strand_sem_append_stateful[of IK DB _ _ \] transaction_dual_subst_unfold[of T \] by fastforce thus ?thesis using B(2) 1[OF B(1,3)] strand_sem_append_stateful subst_lsst_append unfolding prefix_def unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by (metis (no_types) map_append) qed have s_sem: "?s \ ?C \ (l,?s) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ (t \ \, u \ \) \ ?dbupd B" "?s' \ ?C \ (l,?s') = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ t \ \ = u \ \" when B: "prefix ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]) (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" for l s and B::"('fun,'atom,'sets,'lbl) prot_strand" using 2[OF _ B] strand_sem_append_stateful[of IK DB "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" _ \] by (fastforce, fastforce) have 3: "\a \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))). \is_Insert a \ \is_Delete a" when B: "prefix ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]) (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" for l s and B::"('fun,'atom,'sets,'lbl) prot_strand" proof - have "\a \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))). ?Q a" proof fix a assume a: "a \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" have "?P a" when a: "a \ set (unlabel ?R)" for a using a wellformed_transaction_unlabel_cases(1,2)[OF T_valid] unfolding unlabel_def by fastforce hence "?P a" when a: "a \ set (unlabel (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" for a using a stateful_strand_step_cases_subst(2,11)[of _ \] subst_lsst_unlabel[of ?R \] unfolding subst_apply_stateful_strand_def by auto hence B_P: "\a \ set (unlabel (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)). ?P a" using unlabel_mono[OF set_mono_prefix[OF append_prefixD[OF B]]] by blast obtain l where "(l,a) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using a by (meson unlabel_mem_has_label) then obtain b where b: "(l,b) \ set (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (l,b) = (l,a)" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_memberD by blast hence "?P b" using B_P unfolding unlabel_def by fastforce thus "?Q a" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_inv[OF b(2)] by (cases b) auto qed thus ?thesis by fastforce qed show "(t \ \, u \ \) \ DB" when s: "?s \ ?C" proof - obtain l B s where B: "(l,?s) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "prefix ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]) (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using s_in(1)[OF s] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(6)[of _ t u] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_set_prefix_obtain_subst[of ?s ?R \] by blast show ?thesis using 3[OF B(2)] s_sem(1)[OF B(2) s B(1)] dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd[of "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" \ DB] by simp qed show "t \ \ = u \ \" when s: "?s' \ ?C" proof - obtain l B s where B: "(l,?s') = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "prefix ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]) (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using s_in(2)[OF s] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(3)[of _ t u] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_set_prefix_obtain_subst[of ?s' ?R \] by blast show ?thesis using 3[OF B(2)] s_sem(2)[OF B(2) s B(1)] dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd[of "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" \ DB] by simp qed qed lemma wellformed_transaction_sem_neg_checks: assumes T_valid: "wellformed_transaction T" and \: "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \" and "NegChecks X F G \ set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" shows "negchecks_model \ DB X F G" proof - let ?s = "NegChecks X F G" let ?R = "transaction_receive T@transaction_checks T" let ?R' = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?S = "\A. unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?S' = "?S (transaction_receive T)@?S (transaction_checks T)" let ?P = "\a. is_Receive a \ is_Check_or_Assignment a" let ?Q = "\a. is_Send a \ is_Check_or_Assignment a" let ?U = "\\. subst_domain \ = set X \ ground (subst_range \)" have s: "?s \ set (unlabel (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using assms(3) subst_lsst_append[of "transaction_receive T"] unlabel_append[of "transaction_receive T"] by auto obtain l B s where B: "(l,?s) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "prefix ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]) (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using s dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(7)[of X F G] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_set_prefix_obtain_subst[of ?s ?R \] by blast have 1: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \])) = unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))@[?s]" using B(1) unlabel_append dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst singleton_lst_proj(4) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_snoc subst_lsst_append subst_lsst_singleton by (metis (no_types, lifting) subst_apply_labeled_stateful_strand_step.simps) have "strand_sem_stateful IK DB ?S' \" using \ strand_sem_append_stateful[of IK DB _ _ \] transaction_dual_subst_unfold[of T \] by fastforce hence "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))@[?s]) \" using B 1 strand_sem_append_stateful subst_lsst_append unfolding prefix_def unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by (metis (no_types) map_append) hence s_sem: "negchecks_model \ (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \ DB) X F G" using strand_sem_append_stateful[of IK DB "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "[?s]" \] by fastforce have "\a \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))). ?Q a" proof fix a assume a: "a \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" have "?P a" when a: "a \ set (unlabel ?R)" for a using a wellformed_transaction_unlabel_cases(1,2,3)[OF T_valid] unfolding unlabel_def by fastforce hence "?P a" when a: "a \ set (unlabel (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" for a using a stateful_strand_step_cases_subst(2,11)[of _ \] subst_lsst_unlabel[of ?R \] unfolding subst_apply_stateful_strand_def by auto hence B_P: "\a \ set (unlabel (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)). ?P a" using unlabel_mono[OF set_mono_prefix[OF append_prefixD[OF B(2)]]] by blast obtain l where "(l,a) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using a by (meson unlabel_mem_has_label) then obtain b where b: "(l,b) \ set (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (l,b) = (l,a)" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_memberD by blast hence "?P b" using B_P unfolding unlabel_def by fastforce thus "?Q a" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_inv[OF b(2)] by (cases b) auto qed hence "\a \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))). \is_Insert a \ \is_Delete a" by fastforce thus ?thesis using dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd[of "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" \ DB] s_sem by simp qed lemma wellformed_transaction_sem_neg_checks': assumes T_valid: "wellformed_transaction T" and \: "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \" and c: "NegChecks X [] [(t,u)] \ set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" shows "\\. subst_domain \ = set X \ ground (subst_range \) \ (t \ \ \ \, u \ \ \ \) \ DB" (is ?A) and "X = [] \ (t \ \, u \ \) \ DB" (is "?B \ ?B'") proof - show ?A using wellformed_transaction_sem_neg_checks[OF T_valid \ c] unfolding negchecks_model_def by auto moreover have "\ = Var" "t \ \ = t" when "subst_domain \ = set []" for t and \::"('fun, 'atom, 'sets, 'lbl) prot_subst" using that by auto moreover have "subst_domain Var = set []" "range_vars Var = {}" by simp_all ultimately show "?B \ ?B'" unfolding range_vars_alt_def by metis qed lemma wellformed_transaction_sem_iff: fixes T \ defines "A \ unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" and "rm \ \X. rm_vars (set X)" assumes T: "wellformed_transaction T" and I: " interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" shows "strand_sem_stateful M D A I \ ( (\l ts. (l,receive\ts\) \ set (transaction_receive T) \ (\t \ set ts. M \ t \ \ \ I)) \ (\l ac t s. (l,\ac: t \ s\) \ set (transaction_checks T) \ t \ \ \ I = s \ \ \ I) \ (\l ac t s. (l,\ac: t \ s\) \ set (transaction_checks T) \ (t \ \ \ I, s \ \ \ I) \ D) \ (\l X F G. (l,\X\\\: F \\: G\) \ set (transaction_checks T) \ (\\. subst_domain \ = set X \ ground (subst_range \) \ (\(t,s) \ set F. t \ rm X \ \ \ \ I \ s \ rm X \ \ \ \ I) \ (\(t,s) \ set G. (t \ rm X \ \ \ \ I, s \ rm X \ \ \ \ I) \ D))))" (is "?A \ ?B") proof note 0 = A_def transaction_dual_subst_unlabel_unfold note 1 = wellformed_transaction_sem_receives[OF T, of M D \ I, unfolded A_def[symmetric]] wellformed_transaction_sem_pos_checks[OF T, of M D \ I, unfolded A_def[symmetric]] wellformed_transaction_sem_neg_checks[OF T, of M D \ I, unfolded A_def[symmetric]] note 2 = stateful_strand_step_subst_inI[OF unlabel_in] note 3 = unlabel_subst note 4 = strand_sem_append_stateful[of M D _ _ I] let ?C = "\T. unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?P = "\X \. subst_domain \ = set X \ ground (subst_range \)" let ?sem = "\M D T. strand_sem_stateful M D (?C T) I" let ?negchecks = "\X F G. \\. ?P X \ \ (\(t,s) \ set F. t \ rm X \ \ \ \ I \ s \ rm X \ \ \ \ I) \ (\(t,s) \ set G. (t \ rm X \ \ \ \ I, s \ rm X \ \ \ \ I) \ D)" have "list_all is_Receive (unlabel (transaction_receive T))" "list_all is_Check_or_Assignment (unlabel (transaction_checks T))" "list_all is_Update (unlabel (transaction_updates T))" "list_all is_Send (unlabel (transaction_send T))" using T unfolding wellformed_transaction_def by (blast, blast, blast, blast) hence 5: "list_all is_Send (?C (transaction_receive T))" "list_all is_Check_or_Assignment (?C (transaction_checks T))" "list_all is_Update (?C (transaction_updates T))" "list_all is_Receive (?C (transaction_send T))" by (metis (no_types) subst_sst_list_all(2) unlabel_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(1), metis (no_types) subst_sst_list_all(11) unlabel_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(11), metis (no_types) subst_sst_list_all(10) unlabel_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(10), metis (no_types) subst_sst_list_all(1) unlabel_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(2)) have "\a \ set (?C (transaction_receive T)). \is_Receive a \ \is_Insert a \ \is_Delete a" "\a \ set (?C (transaction_checks T)). \is_Receive a \ \is_Insert a \ \is_Delete a" using 5(1,2) unfolding list_all_iff by (blast,blast) hence 6: "M \ (ik\<^sub>s\<^sub>s\<^sub>t (?C (transaction_receive T)) \\<^sub>s\<^sub>e\<^sub>t I) = M" "dbupd\<^sub>s\<^sub>s\<^sub>t (?C (transaction_receive T)) I D = D" "M \ (ik\<^sub>s\<^sub>s\<^sub>t (?C (transaction_checks T)) \\<^sub>s\<^sub>e\<^sub>t I) = M" "dbupd\<^sub>s\<^sub>s\<^sub>t (?C (transaction_checks T)) I D = D" by (metis ik\<^sub>s\<^sub>s\<^sub>t_snoc_no_receive_empty sup_bot.right_neutral, metis dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd, metis ik\<^sub>s\<^sub>s\<^sub>t_snoc_no_receive_empty sup_bot.right_neutral, metis dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd) have ?B when A: ?A proof - have "M \ t \ \ \ I" when "(l, receive\ts\) \ set (transaction_receive T)" "t \ set ts" for l ts t using that(2) 1(1)[OF A, of "ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \"] 2(2)[OF that(1)] unfolding 3 by auto moreover have "t \ \ \ I = s \ \ \ I" when "(l, \ac: t \ s\) \ set (transaction_checks T)" for l ac t s using 1(3)[OF A] 2(3)[OF that] unfolding 3 by blast moreover have "(t \ \ \ I, s \ \ \ I) \ D" when "(l, \ac: t \ s\) \ set (transaction_checks T)" for l ac t s using 1(2)[OF A] 2(6)[OF that] unfolding 3 by blast moreover have "?negchecks X F G" when "(l, \X\\\: F \\: G\) \ set (transaction_checks T)" for l X F G using 1(4)[OF A 2(7)[OF that, of \, unfolded 3]] unfolding negchecks_model_def rm_def subst_apply_pairs_def by fastforce ultimately show ?B by blast qed thus "?A \ ?B" by fast have ?A when B: ?B proof - have 7: "\t \ set ts. M \ t \ I" when ts: "send\ts\ \ set (?C (transaction_receive T))" for ts proof - obtain l ss where "(l,receive\ss\) \ set (transaction_receive T)" "ts = ss \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" by (metis ts dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(2) subst_lsst_memD(1) unlabel_mem_has_label) thus ?thesis using B by auto qed have 8: "t \ I = s \ I" when ts: "\ac: t \ s\ \ set (?C (transaction_checks T))" for ac t s proof - obtain l t' s' where "(l,\ac: t' \ s'\) \ set (transaction_checks T)" "t = t' \ \" "s = s' \ \" by (metis ts dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(3) subst_lsst_memD(3) unlabel_mem_has_label) thus ?thesis using B by auto qed have 9: "(t \ I, s \ I) \ D" when ts: "\ac: t \ s\ \ set (?C (transaction_checks T))" for ac t s proof - obtain l t' s' where "(l,\ac: t' \ s'\) \ set (transaction_checks T)" "t = t' \ \" "s = s' \ \" by (metis ts dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(6) subst_lsst_memD(6) unlabel_mem_has_label) thus ?thesis using B by auto qed have 10: "negchecks_model I D X F G" when ts: "\X\\\: F \\: G\ \ set (?C (transaction_checks T))" for X F G proof - obtain l F' G' where *: "(l,\X\\\: F' \\: G'\) \ set (transaction_checks T)" "F = F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \" "G = G' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \" using unlabel_mem_has_label[OF iffD2[OF dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(7) ts]] subst_lsst_memD(7)[of _ X F G "transaction_checks T" \] by fast have "?negchecks X F' G'" using *(1) B by blast moreover have "\(t,s) \ set F. t \ \ \\<^sub>s I \ s \ \ \\<^sub>s I" when "(t,s) \ set F'" "t \ rm X \ \ \ \ I \ s \ rm X \ \ \ \ I" for \ t s using that unfolding rm_def *(2) subst_apply_pairs_def by force moreover have "\(t,s) \ set G. (t,s) \\<^sub>p \ \\<^sub>s I \ D" when "(t,s) \ set G'" "(t \ rm X \ \ \ \ I, s \ rm X \ \ \ \ I) \ D" for \ t s using that unfolding rm_def *(3) subst_apply_pairs_def by force ultimately show ?thesis unfolding negchecks_model_def by auto qed have "?sem M D (transaction_receive T)" using 7 strand_sem_stateful_if_sends_deduct[OF 5(1)] by blast moreover have "?sem M D (transaction_checks T)" using 8 9 10 strand_sem_stateful_if_checks[OF 5(2)] by blast moreover have "?sem M D (transaction_updates T)" for M D using 5(3) strand_sem_stateful_if_no_send_or_check unfolding list_all_iff by blast moreover have "?sem M D (transaction_send T)" for M D using 5(4) strand_sem_stateful_if_no_send_or_check unfolding list_all_iff by blast ultimately show ?thesis using 4[of "?C (transaction_receive T)" "?C (transaction_checks T)@?C (transaction_updates T)@?C (transaction_send T)"] 4[of "?C (transaction_checks T)" "?C (transaction_updates T)@?C (transaction_send T)"] 4[of "?C (transaction_updates T)" "?C (transaction_send T)"] unfolding 0 6 by blast qed thus "?B \ ?A" by fast qed lemma wellformed_transaction_unlabel_sem_iff: fixes T \ defines "A \ unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" and "rm \ \X. rm_vars (set X)" assumes T: "wellformed_transaction T" and I: " interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" shows "strand_sem_stateful M D A I \ ( (\ts. receive\ts\ \ set (unlabel (transaction_receive T)) \ (\t \ set ts. M \ t \ \ \ I)) \ (\ac t s. \ac: t \ s\ \ set (unlabel (transaction_checks T)) \ t \ \ \ I = s \ \ \ I) \ (\ac t s. \ac: t \ s\ \ set (unlabel (transaction_checks T)) \ (t \ \ \ I, s \ \ \ I) \ D) \ (\X F G. \X\\\: F \\: G\ \ set (unlabel (transaction_checks T)) \ (\\. subst_domain \ = set X \ ground (subst_range \) \ (\(t,s) \ set F. t \ rm X \ \ \ \ I \ s \ rm X \ \ \ \ I) \ (\(t,s) \ set G. (t \ rm X \ \ \ \ I, s \ rm X \ \ \ \ I) \ D))))" using wellformed_transaction_sem_iff[OF T I, of M D \] unlabel_in[of _ _ "transaction_receive T"] unlabel_mem_has_label[of _ "transaction_receive T"] unlabel_in[of _ _ "transaction_checks T"] unlabel_mem_has_label[of _ "transaction_checks T"] unfolding A_def[symmetric] rm_def by meson lemma dual_transaction_ik_is_transaction_send'': fixes \ \::"('a,'b,'c,'d) prot_subst" assumes "wellformed_transaction T" shows "(ik\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a = (trms\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a" (is "?A = ?B") using dual_transaction_ik_is_transaction_send[OF assms] subst_lsst_unlabel[of "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)" \] ik\<^sub>s\<^sub>s\<^sub>t_subst[of "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))" \] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst[of "transaction_strand T" \] by (auto simp add: abs_apply_terms_def) lemma while_prot_terms_fun_mono: "mono (\M'. M \ \(subterms ` M') \ \((set \ fst \ Ana) ` M'))" unfolding mono_def by fast lemma while_prot_terms_SMP_overapprox: fixes M::"('fun,'atom,'sets,'lbl) prot_terms" assumes N_supset: "M \ \(subterms ` N) \ \((set \ fst \ Ana) ` N) \ N" and Value_vars_only: "\x \ fv\<^sub>s\<^sub>e\<^sub>t N. \\<^sub>v x = TAtom Value" shows "SMP M \ {a \ \ | a \. a \ N \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)}" proof - define f where "f \ \M'. M \ \(subterms ` M') \ \((set \ fst \ Ana) ` M')" define S where "S \ {a \ \ | a \. a \ N \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)}" note 0 = Value_vars_only have "t \ S" when "t \ SMP M" for t using that proof (induction t rule: SMP.induct) case (MP t) hence "t \ N" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range Var)" using N_supset by auto hence "t \ Var \ S" unfolding S_def by blast thus ?case by simp next case (Subterm t t') then obtain \ a where a: "a \ \ = t" "a \ N" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (auto simp add: S_def) hence "\x \ fv a. \\. \ (Var x) = TAtom \" using 0 by auto hence *: "\x \ fv a. (\f. \ x = Fun f []) \ (\y. \ x = Var y)" using a(3) TAtom_term_cases[OF wf_trm_subst_rangeD[OF a(4)]] by (metis wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) obtain b where b: "b \ \ = t'" "b \ subterms a" using subterms_subst_subterm[OF *, of t'] Subterm.hyps(2) a(1) by fast hence "b \ N" using N_supset a(2) by blast thus ?case using a b(1) unfolding S_def by blast next case (Substitution t \) then obtain \ a where a: "a \ \ = t" "a \ N" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (auto simp add: S_def) have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" by (fact wt_subst_compose[OF a(3) Substitution.hyps(2)], fact wf_trms_subst_compose[OF a(4) Substitution.hyps(3)]) moreover have "t \ \ = a \ \ \\<^sub>s \" using a(1) subst_subst_compose[of a \ \] by simp ultimately show ?case using a(2) unfolding S_def by blast next case (Ana t K T k) then obtain \ a where a: "a \ \ = t" "a \ N" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (auto simp add: S_def) obtain Ka Ta where a': "Ana a = (Ka,Ta)" by moura have *: "K = Ka \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" proof (cases a) case (Var x) then obtain g U where gU: "t = Fun g U" using a(1) Ana.hyps(2,3) Ana_var by (cases t) simp_all have "\ (Var x) = TAtom Value" using Var a(2) 0 by auto hence "\ (Fun g U) = TAtom Value" using a(1,3) Var gU wt_subst_trm''[OF a(3), of a] by argo thus ?thesis using gU Fun_Value_type_inv Ana.hyps(2,3) by fastforce next case (Fun g U) thus ?thesis using a(1) a' Ana.hyps(2) Ana_subst'[of g U] by simp qed then obtain ka where ka: "k = ka \ \" "ka \ set Ka" using Ana.hyps(3) by auto have "ka \ set ((fst \ Ana) a)" using ka(2) a' by simp hence "ka \ N" using a(2) N_supset by auto thus ?case using ka a(3,4) unfolding S_def by blast qed thus ?thesis unfolding S_def by blast qed subsection \Admissible Transactions\ definition admissible_transaction_checks where "admissible_transaction_checks T \ \x \ set (unlabel (transaction_checks T)). (is_InSet x \ is_Var (the_elem_term x) \ is_Fun_Set (the_set_term x) \ fst (the_Var (the_elem_term x)) = TAtom Value) \ (is_NegChecks x \ bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p x = [] \ ((the_eqs x = [] \ length (the_ins x) = 1) \ (the_ins x = [] \ length (the_eqs x) = 1))) \ (is_NegChecks x \ the_eqs x = [] \ (let h = hd (the_ins x) in is_Var (fst h) \ is_Fun_Set (snd h) \ fst (the_Var (fst h)) = TAtom Value))" definition admissible_transaction_updates where "admissible_transaction_updates T \ \x \ set (unlabel (transaction_updates T)). is_Update x \ is_Var (the_elem_term x) \ is_Fun_Set (the_set_term x) \ fst (the_Var (the_elem_term x)) = TAtom Value" definition admissible_transaction_terms where "admissible_transaction_terms T \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)) \ (\f \ \(funs_term ` trms_transaction T). \is_Val f \ \is_Abs f \ \is_PubConst f \ f \ Pair) \ (\r \ set (unlabel (transaction_strand T)). (\f \ \(funs_term ` (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r)). is_Attack f) \ transaction_fresh T = [] \ is_Send r \ length (the_msgs r) = 1 \ is_Fun_Attack (hd (the_msgs r)))" definition admissible_transaction_send_occurs_form where "admissible_transaction_send_occurs_form T \ ( let snds = transaction_send T; frsh = transaction_fresh T in \t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t snds. OccursFact \ funs_term t \ OccursSec \ funs_term t \ (\x \ set frsh. t = occurs (Var x)) )" definition admissible_transaction_occurs_checks where "admissible_transaction_occurs_checks T \ ( let occ_in = \x S. occurs (Var x) \ set (the_msgs (hd (unlabel S))); rcvs = transaction_receive T; snds = transaction_send T; frsh = transaction_fresh T; fvs = fv_transaction T in admissible_transaction_send_occurs_form T \ ((\x \ fvs - set frsh. fst x = TAtom Value) \ ( rcvs \ [] \ is_Receive (hd (unlabel rcvs)) \ (\x \ fvs - set frsh. fst x = TAtom Value \ occ_in x rcvs))) \ (frsh \ [] \ ( snds \ [] \ is_Send (hd (unlabel snds)) \ (\x \ set frsh. occ_in x snds))) )" definition admissible_transaction_no_occurs_msgs where "admissible_transaction_no_occurs_msgs T \ ( let no_occ = \t. is_Fun t \ the_Fun t \ OccursFact; rcvs = transaction_receive T; snds = transaction_send T in list_all (\a. is_Receive (snd a) \ list_all no_occ (the_msgs (snd a))) rcvs \ list_all (\a. is_Send (snd a) \ list_all no_occ (the_msgs (snd a))) snds )" definition admissible_transaction' where "admissible_transaction' T \ ( wellformed_transaction T \ transaction_decl T () = [] \ list_all (\x. fst x = TAtom Value) (transaction_fresh T) \ (\x \ vars_transaction T. is_Var (fst x) \ (the_Var (fst x) = Value)) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) = {} \ set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (filter (is_Insert \ snd) (transaction_updates T)) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \ (\x \ fv_transaction T - set (transaction_fresh T). \y \ fv_transaction T - set (transaction_fresh T). x \ y \ \Var x != Var y\ \ set (unlabel (transaction_checks T)) \ \Var y != Var x\ \ set (unlabel (transaction_checks T))) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) - set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ (\r \ set (unlabel (transaction_checks T)). is_Equality r \ fv (the_rhs r) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (filter (\s. is_InSet (snd s) \ the_check (snd s) = Assign) (transaction_checks T)) \ list_all (\a. is_Receive (snd a) \ the_msgs (snd a) \ []) (transaction_receive T) \ list_all (\a. is_Send (snd a) \ the_msgs (snd a) \ []) (transaction_send T) \ admissible_transaction_checks T \ admissible_transaction_updates T \ admissible_transaction_terms T \ admissible_transaction_send_occurs_form T )" definition admissible_transaction where "admissible_transaction T \ admissible_transaction' T \ admissible_transaction_no_occurs_msgs T" definition has_initial_value_producing_transaction where "has_initial_value_producing_transaction P \ let f = \s. list_all (\T. list_all (\a. ((is_Delete a \ is_InSet a) \ the_set_term a \ \s\\<^sub>s) \ (is_NegChecks a \ list_all (\(_,t). t \ \s\\<^sub>s) (the_ins a))) (unlabel (transaction_checks T@transaction_updates T))) P in list_ex (\T. length (transaction_fresh T) = 1 \ transaction_receive T = [] \ transaction_checks T = [] \ length (transaction_send T) = 1 \ (let x = hd (transaction_fresh T); a = hd (transaction_send T); u = transaction_updates T in is_Send (snd a) \ Var x \ set (the_msgs (snd a)) \ fv\<^sub>s\<^sub>e\<^sub>t (set (the_msgs (snd a))) = {x} \ (u \ [] \ ( let b = hd u; c = snd b in tl u = [] \ is_Insert c \ the_elem_term c = Var x \ is_Fun_Set (the_set_term c) \ f (the_Set (the_Fun (the_set_term c)))))) ) P" lemma admissible_transaction_is_wellformed_transaction: assumes "admissible_transaction' T" shows "wellformed_transaction T" and "admissible_transaction_checks T" and "admissible_transaction_updates T" and "admissible_transaction_terms T" and "admissible_transaction_send_occurs_form T" using assms unfolding admissible_transaction'_def by blast+ lemma admissible_transaction_no_occurs_msgsE: assumes T: "admissible_transaction' T" "admissible_transaction_no_occurs_msgs T" shows "\ts. send\ts\ \ set (unlabel (transaction_strand T)) \ receive\ts\ \ set (unlabel (transaction_strand T)) \ (\t s. t \ set ts \ t \ occurs s)" proof - note 1 = admissible_transaction_is_wellformed_transaction(1)[OF T(1)] have 2: "send\ts\ \ set (unlabel (transaction_send T))" when "send\ts\ \ set (unlabel (transaction_strand T))" for ts using wellformed_transaction_strand_unlabel_memberD(8)[OF 1 that] by fast have 3: "receive\ts\ \ set (unlabel (transaction_receive T))" when "receive\ts\ \ set (unlabel (transaction_strand T))" for ts using wellformed_transaction_strand_unlabel_memberD(1)[OF 1 that] by fast show ?thesis using T(2) 2 3 wellformed_transaction_unlabel_cases(1,4)[OF 1] unfolding admissible_transaction_no_occurs_msgs_def Let_def list_all_iff by (metis sndI stateful_strand_step.discI(1,2) stateful_strand_step.sel(1,2) term.discI(2) term.sel(2) unlabel_mem_has_label) qed lemma admissible_transactionE: assumes T: "admissible_transaction' T" shows "transaction_decl T () = []" (is ?A) and "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" (is ?B) and "\x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T). \\<^sub>v x = TAtom Value" (is ?C) and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) = {}" (is ?D1) and "fv_transaction T \ bvars_transaction T = {}" (is ?D2) and "set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (filter (is_Insert \ snd) (transaction_updates T)) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" (is ?E) and "set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" (is ?F) and "\x \ fv_transaction T - set (transaction_fresh T). \y \ fv_transaction T - set (transaction_fresh T). x \ y \ \Var x != Var y\ \ set (unlabel (transaction_checks T)) \ \Var y != Var x\ \ set (unlabel (transaction_checks T))" (is ?G) and "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T). x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ (\t s. select\t,s\ \ set (unlabel (transaction_checks T)) \ x \ fv t \ fv s)" (is ?H) and "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) - set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" (is ?I) and "\x \ set (unlabel (transaction_checks T)). is_Equality x \ fv (the_rhs x) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" (is ?J) (* TODO: why do we need this requirement? *) and "set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) = {}" (is ?K1) and "set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) = {}" (is ?K2) and "list_all (\x. fst x = Var Value) (transaction_fresh T)" (is ?K3) and "\x \ vars_transaction T. \TAtom AttackType \ \\<^sub>v x" (is ?K4) and "\l ts. (l,receive\ts\) \ set (transaction_receive T) \ ts \ []" (is ?L1) and "\l ts. (l,send\ts\) \ set (transaction_send T) \ ts \ []" (is ?L2) proof - show ?A ?D1 ?D2 ?G ?I ?J ?K3 using T unfolding admissible_transaction'_def by (blast, blast, blast, blast, blast, blast, blast) have "list_all (\a. is_Receive (snd a) \ the_msgs (snd a) \ []) (transaction_receive T)" "list_all (\a. is_Send (snd a) \ the_msgs (snd a) \ []) (transaction_send T)" using T unfolding admissible_transaction'_def by auto thus ?L1 ?L2 unfolding list_all_iff by (force,force) have "list_all (\x. fst x = Var Value) (transaction_fresh T)" "\x\vars_transaction T. is_Var (fst x) \ the_Var (fst x) = Value" using T unfolding admissible_transaction'_def by (blast, blast) thus ?B ?C ?K4 using \\<^sub>v_TAtom''(2) unfolding list_all_iff by (blast, force, force) show ?E using T unfolding admissible_transaction'_def by argo thus ?F unfolding unlabel_def by auto show ?K1 ?K2 using T unfolding admissible_transaction'_def wellformed_transaction_def by (argo, argo) let ?selects = "filter (\s. is_InSet (snd s) \ the_check (snd s) = Assign) (transaction_checks T)" show ?H proof fix x assume "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" hence "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?selects" using T unfolding admissible_transaction'_def by blast thus "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ (\t s. select\t,s\ \ set (unlabel (transaction_checks T)) \ x \ fv t \ fv s)" proof assume "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?selects" then obtain r where r: "x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p r" "r \ set (unlabel (transaction_checks T))" "is_InSet r" "the_check r = Assign" unfolding unlabel_def by force thus ?thesis by (cases r) auto qed simp qed qed lemma admissible_transactionE': assumes T: "admissible_transaction T" shows "admissible_transaction' T" (is ?A) and "admissible_transaction_no_occurs_msgs T" (is ?B) and "\ts. send\ts\ \ set (unlabel (transaction_strand T)) \ receive\ts\ \ set (unlabel (transaction_strand T)) \ (\t s. t \ set ts \ t \ occurs s)" (is ?C) proof - show 0: ?A ?B using T unfolding admissible_transaction_def by (blast, blast) show ?C using admissible_transaction_no_occurs_msgsE[OF 0] by blast qed lemma transaction_inserts_are_Value_vars: assumes T_valid: "wellformed_transaction T" and "admissible_transaction_updates T" and "insert\t,s\ \ set (unlabel (transaction_strand T))" shows "\n. t = Var (TAtom Value, n)" and "\u. s = Fun (Set u) []" proof - let ?x = "insert\t,s\" have "?x \ set (unlabel (transaction_updates T))" using assms(3) wellformed_transaction_unlabel_cases[OF T_valid, of ?x] by (auto simp add: transaction_strand_def unlabel_def) hence *: "is_Var (the_elem_term ?x)" "fst (the_Var (the_elem_term ?x)) = TAtom Value" "is_Fun (the_set_term ?x)" "args (the_set_term ?x) = []" "is_Set (the_Fun (the_set_term ?x))" using assms(2) unfolding admissible_transaction_updates_def is_Fun_Set_def by fastforce+ show "\n. t = Var (TAtom Value, n)" using *(1,2) by (cases t) auto show "\u. s = Fun (Set u) []" using *(3,4,5) unfolding is_Set_def by (cases s) auto qed lemma transaction_deletes_are_Value_vars: assumes T_valid: "wellformed_transaction T" and "admissible_transaction_updates T" and "delete\t,s\ \ set (unlabel (transaction_strand T))" shows "\n. t = Var (TAtom Value, n)" and "\u. s = Fun (Set u) []" proof - let ?x = "delete\t,s\" have "?x \ set (unlabel (transaction_updates T))" using assms(3) wellformed_transaction_unlabel_cases[OF T_valid, of ?x] by (auto simp add: transaction_strand_def unlabel_def) hence *: "is_Var (the_elem_term ?x)" "fst (the_Var (the_elem_term ?x)) = TAtom Value" "is_Fun (the_set_term ?x)" "args (the_set_term ?x) = []" "is_Set (the_Fun (the_set_term ?x))" using assms(2) unfolding admissible_transaction_updates_def is_Fun_Set_def by fastforce+ show "\n. t = Var (TAtom Value, n)" using *(1,2) by (cases t) auto show "\u. s = Fun (Set u) []" using *(3,4,5) unfolding is_Set_def by (cases s) auto qed lemma transaction_selects_are_Value_vars: assumes T_valid: "wellformed_transaction T" and "admissible_transaction_checks T" and "select\t,s\ \ set (unlabel (transaction_strand T))" shows "\n. t = Var (TAtom Value, n) \ (TAtom Value, n) \ set (transaction_fresh T)" (is ?A) and "\u. s = Fun (Set u) []" (is ?B) proof - let ?x = "select\t,s\" have *: "?x \ set (unlabel (transaction_checks T))" using assms(3) wellformed_transaction_unlabel_cases[OF T_valid, of ?x] by (auto simp add: transaction_strand_def unlabel_def) have **: "is_Var (the_elem_term ?x)" "fst (the_Var (the_elem_term ?x)) = TAtom Value" "is_Fun (the_set_term ?x)" "args (the_set_term ?x) = []" "is_Set (the_Fun (the_set_term ?x))" using * assms(2) unfolding admissible_transaction_checks_def is_Fun_Set_def by fastforce+ have "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using * by force hence ***: "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x \ set (transaction_fresh T) = {}" using T_valid unfolding wellformed_transaction_def by fast show ?A using **(1,2) *** by (cases t) auto show ?B using **(3,4,5) unfolding is_Set_def by (cases s) auto qed lemma transaction_inset_checks_are_Value_vars: assumes T_valid: "admissible_transaction' T" and t: "\t in s\ \ set (unlabel (transaction_strand T))" shows "\n. t = Var (TAtom Value, n) \ (TAtom Value, n) \ set (transaction_fresh T)" (is ?A) and "\u. s = Fun (Set u) []" (is ?B) proof - let ?x = "\t in s\" note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_valid] note T_adm_checks = admissible_transaction_is_wellformed_transaction(2)[OF T_valid] have *: "?x \ set (unlabel (transaction_checks T))" using t wellformed_transaction_unlabel_cases[OF T_wf, of ?x] unfolding transaction_strand_def unlabel_def by fastforce have **: "is_Var (the_elem_term ?x)" "fst (the_Var (the_elem_term ?x)) = TAtom Value" "is_Fun (the_set_term ?x)" "args (the_set_term ?x) = []" "is_Set (the_Fun (the_set_term ?x))" using * T_adm_checks unfolding admissible_transaction_checks_def is_Fun_Set_def by fastforce+ have "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using * by force hence ***: "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x \ set (transaction_fresh T) = {}" using T_wf unfolding wellformed_transaction_def by fast show ?A using **(1,2) *** by (cases t) auto show ?B using **(3,4,5) unfolding is_Set_def by (cases s) auto qed lemma transaction_notinset_checks_are_Value_vars: assumes T_adm: "admissible_transaction' T" and FG: "\X\\\: F \\: G\ \ set (unlabel (transaction_strand T))" and t: "(t,s) \ set G" shows "\n. t = Var (TAtom Value, n) \ (TAtom Value, n) \ set (transaction_fresh T)" (is ?A) and "\u. s = Fun (Set u) []" (is ?B) and "F = []" (is ?C) and "G = [(t,s)]" (is ?D) proof - let ?x = "\X\\\: F \\: G\" note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] note T_adm_checks = admissible_transaction_is_wellformed_transaction(2)[OF T_adm] have 0: "?x \ set (unlabel (transaction_checks T))" using FG wellformed_transaction_unlabel_cases[OF T_wf, of ?x] by (auto simp add: transaction_strand_def unlabel_def) hence 1: "F = [] \ length G = 1" using T_adm_checks t unfolding admissible_transaction_checks_def by fastforce hence "hd G = (t,s)" using t by (cases "the_ins ?x") auto hence **: "is_Var t" "fst (the_Var t) = TAtom Value" "is_Fun s" "args s = []" "is_Set (the_Fun s)" using 1 Set.bspec[OF T_adm_checks[unfolded admissible_transaction_checks_def] 0] unfolding is_Fun_Set_def by auto show ?C using 1 by blast show ?D using 1 t by force have "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using 0 by force+ moreover have "set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) = {}" "set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) = {}" using T_wf unfolding wellformed_transaction_def by fast+ ultimately have "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x \ set (transaction_fresh T) = {}" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x) \ set (transaction_fresh T) = {}" using admissible_transactionE(7)[OF T_adm] wellformed_transaction_wf\<^sub>s\<^sub>s\<^sub>t(2)[OF T_wf] fv_transaction_unfold[of T] bvars_transaction_unfold[of T] by (blast, blast) hence ***: "fv t \ set (transaction_fresh T) = {}" using t by auto show ?A using **(1,2) *** by (cases t) auto show ?B using **(3,4,5) unfolding is_Set_def by (cases s) auto qed lemma transaction_noteqs_checks_case: assumes T_adm: "admissible_transaction' T" and FG: "\X\\\: F \\: G\ \ set (unlabel (transaction_strand T))" and G: "G = []" shows "\t s. F = [(t,s)]" (is ?A) proof - let ?x = "\X\\\: F \\: G\" note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] note T_adm_checks = admissible_transaction_is_wellformed_transaction(2)[OF T_adm] have "?x \ set (unlabel (transaction_checks T))" using FG wellformed_transaction_unlabel_cases[OF T_wf, of ?x] by (auto simp add: transaction_strand_def unlabel_def) hence "length F = 1" using T_adm_checks unfolding admissible_transaction_checks_def G by fastforce thus ?thesis by fast qed lemma admissible_transaction_fresh_vars_notin: assumes T: "admissible_transaction' T" and x: "x \ set (transaction_fresh T)" shows "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" (is ?A) and "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" (is ?B) and "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" (is ?C) and "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" (is ?D) and "x \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" (is ?E) and "x \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" (is ?F) proof - note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T] have 0: "set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) = {}" "set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) = {}" "fv_transaction T \ bvars_transaction T = {}" using admissible_transactionE[OF T] by argo+ have 1: "set (transaction_fresh T) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) = {}" using 0(1,4) fv_transaction_unfold[of T] bvars_transaction_unfold[of T] by blast have 2: "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) = {}" using bvars_wellformed_transaction_unfold[OF T_wf] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_receive T)"] by blast+ show ?A ?B ?C ?E ?F using 0 1 2 x by (fast, fast, fast, fast, fast) show ?D using 0(3) 1 x vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_checks T)"] by fast qed lemma admissible_transaction_fv_in_receives_or_selects: assumes T: "admissible_transaction' T" and x: "x \ fv_transaction T" "x \ set (transaction_fresh T)" shows "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ (x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ (\t s. select\t,s\ \ set (unlabel (transaction_checks T)) \ x \ fv t \ fv s))" proof - have "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" using x(1) fv\<^sub>s\<^sub>s\<^sub>t_append unlabel_append by (metis transaction_strand_def append_assoc) thus ?thesis using x(2) admissible_transactionE(9,10)[OF T] by blast qed lemma admissible_transaction_fv_in_receives_or_selects': assumes T: "admissible_transaction' T" and x: "x \ fv_transaction T" "x \ set (transaction_fresh T)" shows "(\ts. receive\ts\ \ set (unlabel (transaction_receive T)) \ x \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) \ (\s. select\Var x, s\ \ set (unlabel (transaction_checks T)))" proof (cases "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)") case True thus ?thesis using wellformed_transaction_unlabel_cases(1)[ OF admissible_transaction_is_wellformed_transaction(1)[OF T]] by force next case False then obtain t s where t: "select\t,s\ \ set (unlabel (transaction_checks T))" "x \ fv t \ fv s" using admissible_transaction_fv_in_receives_or_selects[OF T x] by blast have t': "select\t,s\ \ set (unlabel (transaction_strand T))" using t(1) unfolding transaction_strand_def by simp show ?thesis using t transaction_selects_are_Value_vars[ OF admissible_transaction_is_wellformed_transaction(1,2)[OF T] t'] by force qed lemma admissible_transaction_fv_in_receives_or_selects_subst: assumes T: "admissible_transaction' T" and x: "x \ fv_transaction T" "x \ set (transaction_fresh T)" shows "(\ts. receive\ts\ \ set (unlabel (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ \ x \\<^sub>s\<^sub>e\<^sub>t set ts) \ (\s. select\\ x, s\ \ set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" proof - note 0 = admissible_transaction_fv_in_receives_or_selects'[OF T x] have 1: "\ x \\<^sub>s\<^sub>e\<^sub>t set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" when ts: "x \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" for ts using that subst_mono_fv[of x _ \] by auto have 2: "receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\ \ set (A \\<^sub>s\<^sub>s\<^sub>t \)" when "receive\ts\ \ set A" for ts A using that by fast have 3: "select\t \ \,s \ \\ \ set (A \\<^sub>s\<^sub>s\<^sub>t \)" when "select\t,s\ \ set A" for t s A using that by fast show ?thesis using 0 1 2[of _ "unlabel (transaction_receive T)"] 3[of _ _ "unlabel (transaction_checks T)"] - unfolding unlabel_subst by (metis subst_apply_term.simps(1)) + unfolding unlabel_subst by (metis eval_term.simps(1)) qed lemma admissible_transaction_fv_in_receives_or_selects_dual_subst: defines "f \ \S. unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S)" assumes T: "admissible_transaction' T" and x: "x \ fv_transaction T" "x \ set (transaction_fresh T)" shows "(\ts. send\ts\ \ set (f (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ \ x \\<^sub>s\<^sub>e\<^sub>t set ts) \ (\s. select\\ x, s\ \ set (f (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" using admissible_transaction_fv_in_receives_or_selects_subst[OF T x, of \] by (metis (no_types, lifting) f_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(2) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(6)) lemma admissible_transaction_decl_subst_empty': assumes T: "transaction_decl T () = []" and \: "transaction_decl_subst \ T" shows "\ = Var" proof - have "subst_domain \ = {}" using \ T unfolding transaction_decl_subst_def by auto thus ?thesis by auto qed lemma admissible_transaction_decl_subst_empty: assumes T: "admissible_transaction' T" and \: "transaction_decl_subst \ T" shows "\ = Var" by (rule admissible_transaction_decl_subst_empty'[OF admissible_transactionE(1)[OF T] \]) lemma admissible_transaction_no_bvars: assumes "admissible_transaction' T" shows "fv_transaction T = vars_transaction T" and "bvars_transaction T = {}" using admissible_transactionE(4)[OF assms] bvars_wellformed_transaction_unfold vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t by (fast, fast) lemma admissible_transactions_fv_bvars_disj: assumes "\T \ set P. admissible_transaction' T" shows "(\T \ set P. fv_transaction T) \ (\T \ set P. bvars_transaction T) = {}" using assms admissible_transaction_no_bvars(2) by fast lemma admissible_transaction_occurs_fv_types: assumes "admissible_transaction' T" and "x \ vars_transaction T" shows "\a. \ (Var x) = TAtom a \ \ (Var x) \ TAtom OccursSecType" proof - have "is_Var (fst x)" "the_Var (fst x) = Value" using assms unfolding admissible_transaction'_def by blast+ thus ?thesis using \\<^sub>v_TAtom''(2)[of x] by force qed lemma admissible_transaction_Value_vars_are_fv: assumes "admissible_transaction' T" and "x \ vars_transaction T" and "\\<^sub>v x = TAtom Value" shows "x \ fv_transaction T" using assms(2,3) \\<^sub>v_TAtom''(2)[of x] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] admissible_transactionE(4)[OF assms(1)] by fast lemma transaction_receive_deduct: assumes T_wf: "wellformed_transaction T" and \: "constraint_model \ (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T M" and \: "transaction_renaming_subst \ P X" and t: "receive\ts\ \ set (unlabel (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" shows "\t \ set ts. ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" proof - define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" have t': "send\ts\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" using t dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(2) unfolding \_def by blast then obtain T1 T2 where T: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = T1@send\ts\#T2" using t' by (meson split_list) have "constr_sem_stateful \ (unlabel A@unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" using \ unlabel_append[of A] unfolding constraint_model_def \_def by simp hence "constr_sem_stateful \ (unlabel A@T1@[send\ts\])" using strand_sem_append_stateful[of "{}" "{}" "unlabel A@T1@[send\ts\]" _ \] transaction_dual_subst_unlabel_unfold[of T \] T by (metis append.assoc append_Cons append_Nil) hence "\t \ set ts. ik\<^sub>s\<^sub>s\<^sub>t (unlabel A@T1) \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" using strand_sem_append_stateful[of "{}" "{}" "unlabel A@T1" "[send\ts\]" \] T by force moreover have "\is_Receive x" when x: "x \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" for x proof - have *: "is_Receive a" when "a \ set (unlabel (transaction_receive T))" for a using T_wf Ball_set[of "unlabel (transaction_receive T)" is_Receive] that unfolding wellformed_transaction_def by blast obtain l where l: "(l,x) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using x unfolding unlabel_def by fastforce then obtain ly where ly: "ly \ set (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "(l,x) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ly" unfolding dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto obtain j y where j: "ly = (j,y)" by (metis surj_pair) hence "j = l" using ly(2) by (cases y) auto hence y: "(l,y) \ set (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "(l,x) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (l,y)" by (metis j ly(1), metis j ly(2)) obtain z where z: "z \ set (unlabel (transaction_receive T))" "(l,z) \ set (transaction_receive T)" "(l,y) = (l,z) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" using y(1) unfolding subst_apply_labeled_stateful_strand_def unlabel_def by force have "is_Receive y" using *[OF z(1)] z(3) by (cases z) auto thus "\is_Receive x" using l y by (cases y) auto qed hence "\is_Receive x" when "x \ set T1" for x using T that by simp hence "ik\<^sub>s\<^sub>s\<^sub>t T1 = {}" unfolding ik\<^sub>s\<^sub>s\<^sub>t_def is_Receive_def by fast hence "ik\<^sub>s\<^sub>s\<^sub>t (unlabel A@T1) = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using ik\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel A" T1] by simp ultimately show ?thesis by (simp add: \_def) qed lemma transaction_checks_db: assumes T: "admissible_transaction' T" and \: "constraint_model \ (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T M" and \: "transaction_renaming_subst \ P X" shows "\Var (TAtom Value, n) in Fun (Set s) []\ \ set (unlabel (transaction_checks T)) \ (\ (TAtom Value, n) \ \, Fun (Set s) []) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \)" (is "?A \ ?B") and "\Var (TAtom Value, n) not in Fun (Set s) []\ \ set (unlabel (transaction_checks T)) \ (\ (TAtom Value, n) \ \, Fun (Set s) []) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \)" (is "?C \ ?D") proof - let ?x = "\n. (TAtom Value, n)" let ?s = "Fun (Set s) []" let ?T = "transaction_receive T@transaction_checks T" let ?T' = "?T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" let ?S = "\S. transaction_receive T@S" let ?S' = "\S. ?S S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" note \_empty = admissible_transaction_decl_subst_empty[OF T \] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T] have "constr_sem_stateful \ (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)))" using \ unfolding constraint_model_def by simp moreover have "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S (T1@[c]) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T2@transaction_updates T@transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" when "transaction_checks T = T1@c#T2" for T1 T2 c \ using that dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append subst_lsst_append unfolding transaction_strand_def by (metis append.assoc append_Cons append_Nil) ultimately have T'_model: "constr_sem_stateful \ (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' (T1@[(l,c)]))))" when "transaction_checks T = T1@(l,c)#T2" for T1 T2 l c using strand_sem_append_stateful[of _ _ _ _ \] by (simp add: that transaction_strand_def) show "?A \ ?B" proof - assume a: ?A hence *: "\Var (?x n) in ?s\ \ set (unlabel ?T)" unfolding transaction_strand_def unlabel_def by simp then obtain l T1 T2 where T1: "transaction_checks T = T1@(l,\Var (?x n) in ?s\)#T2" by (metis a split_list unlabel_mem_has_label) have "?x n \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using a by force hence "?x n \ set (transaction_fresh T)" using a admissible_transaction_fresh_vars_notin[OF T] by fast hence "unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' (T1@[(l,\Var (?x n) in ?s\)]))) = unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' T1))@[\\ (?x n) in ?s\]" using T a \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append subst_lsst_append unlabel_append \_empty by (fastforce simp add: transaction_fresh_subst_def unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def - subst_apply_labeled_stateful_strand_def) + subst_apply_labeled_stateful_strand_def subst_compose) moreover have "db\<^sub>s\<^sub>s\<^sub>t (unlabel A) = db\<^sub>s\<^sub>s\<^sub>t (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' T1)))" by (simp add: T1 db\<^sub>s\<^sub>s\<^sub>t_transaction_prefix_eq[OF T_wf] del: unlabel_append) ultimately have "\M. strand_sem_stateful M (set (db\<^sub>s\<^sub>s\<^sub>t (unlabel A) \)) [\\ (?x n) in ?s\] \" using T'_model[OF T1] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of _ \] strand_sem_append_stateful[of _ _ _ _ \] by (simp add: db\<^sub>s\<^sub>s\<^sub>t_def del: unlabel_append) thus ?B by simp qed show "?C \ ?D" proof - assume a: ?C hence *: "\Var (?x n) not in ?s\ \ set (unlabel ?T)" unfolding transaction_strand_def unlabel_def by simp then obtain l T1 T2 where T1: "transaction_checks T = T1@(l,\Var (?x n) not in ?s\)#T2" by (metis a split_list unlabel_mem_has_label) have "?x n \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p \Var (?x n) not in ?s\" using vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_cases(9)[of "[]" "Var (?x n)" ?s] by auto hence "?x n \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using a unfolding vars\<^sub>s\<^sub>s\<^sub>t_def by force hence "?x n \ set (transaction_fresh T)" using a admissible_transaction_fresh_vars_notin[OF T] by fast hence "unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' (T1@[(l,\Var (?x n) not in ?s\)]))) = unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' T1))@[\\ (?x n) not in ?s\]" using T a \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append subst_lsst_append unlabel_append \_empty by (fastforce simp add: transaction_fresh_subst_def unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def - subst_apply_labeled_stateful_strand_def) + subst_apply_labeled_stateful_strand_def subst_compose) moreover have "db\<^sub>s\<^sub>s\<^sub>t (unlabel A) = db\<^sub>s\<^sub>s\<^sub>t (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' T1)))" by (simp add: T1 db\<^sub>s\<^sub>s\<^sub>t_transaction_prefix_eq[OF T_wf] del: unlabel_append) ultimately have "\M. strand_sem_stateful M (set (db\<^sub>s\<^sub>s\<^sub>t (unlabel A) \)) [\\ (?x n) not in ?s\] \" using T'_model[OF T1] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of _ \] strand_sem_append_stateful[of _ _ _ _ \] by (simp add: db\<^sub>s\<^sub>s\<^sub>t_def del: unlabel_append) thus ?D using stateful_strand_sem_NegChecks_no_bvars(1)[of _ _ _ ?s \] by simp qed qed lemma transaction_selects_db: assumes T: "admissible_transaction' T" and \: "constraint_model \ (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T M" and \: "transaction_renaming_subst \ P X" shows "select\Var (TAtom Value, n), Fun (Set s) []\ \ set (unlabel (transaction_checks T)) \ (\ (TAtom Value, n) \ \, Fun (Set s) []) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \)" (is "?A \ ?B") proof - let ?x = "\n. (TAtom Value, n)" let ?s = "Fun (Set s) []" let ?T = "transaction_receive T@transaction_checks T" let ?T' = "?T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" let ?S = "\S. transaction_receive T@S" let ?S' = "\S. ?S S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" note \_empty = admissible_transaction_decl_subst_empty[OF T \] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T] have "constr_sem_stateful \ (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)))" using \ unfolding constraint_model_def by simp moreover have "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S (T1@[c]) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T2@transaction_updates T@transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" when "transaction_checks T = T1@c#T2" for T1 T2 c \ using that dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append subst_lsst_append unfolding transaction_strand_def by (metis append.assoc append_Cons append_Nil) ultimately have T'_model: "constr_sem_stateful \ (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' (T1@[(l,c)]))))" when "transaction_checks T = T1@(l,c)#T2" for T1 T2 l c using strand_sem_append_stateful[of _ _ _ _ \] by (simp add: that transaction_strand_def) show "?A \ ?B" proof - assume a: ?A hence *: "select\Var (?x n), ?s\ \ set (unlabel ?T)" unfolding transaction_strand_def unlabel_def by simp then obtain l T1 T2 where T1: "transaction_checks T = T1@(l,select\Var (?x n), ?s\)#T2" by (metis a split_list unlabel_mem_has_label) have "?x n \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using a by force hence "?x n \ set (transaction_fresh T)" using a admissible_transaction_fresh_vars_notin[OF T] by fast hence "unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' (T1@[(l,select\Var (?x n), ?s\)]))) = unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' T1))@[select\\ (?x n), ?s\]" using T a \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append subst_lsst_append unlabel_append \_empty by (fastforce simp add: transaction_fresh_subst_def unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def - subst_apply_labeled_stateful_strand_def) + subst_apply_labeled_stateful_strand_def subst_compose) moreover have "db\<^sub>s\<^sub>s\<^sub>t (unlabel A) = db\<^sub>s\<^sub>s\<^sub>t (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' T1)))" by (simp add: T1 db\<^sub>s\<^sub>s\<^sub>t_transaction_prefix_eq[OF T_wf] del: unlabel_append) ultimately have "\M. strand_sem_stateful M (set (db\<^sub>s\<^sub>s\<^sub>t (unlabel A) \)) [\\ (?x n) in ?s\] \" using T'_model[OF T1] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of _ \] strand_sem_append_stateful[of _ _ _ _ \] by (simp add: db\<^sub>s\<^sub>s\<^sub>t_def del: unlabel_append) thus ?B by simp qed qed lemma admissible_transaction_terms_no_Value_consts: assumes "admissible_transaction_terms T" and "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))" shows "\a T. t = Fun (Val a) T" (is ?A) and "\a T. t = Fun (Abs a) T" (is ?B) and "\a T. t = Fun (PubConst Value a) T" (is ?C) proof - have "\is_Val f" "\is_Abs f" "\is_PubConstValue f" when "f \ \(funs_term ` (trms_transaction T))" for f using that assms(1)[unfolded admissible_transaction_terms_def] unfolding is_PubConstValue_def by (blast,blast,blast) moreover have "f \ \(funs_term ` (trms_transaction T))" when "f \ funs_term t" for f using that assms(2) funs_term_subterms_eq(2)[of "trms_transaction T"] by blast+ ultimately have *: "\is_Val f" "\is_Abs f" "\is_PubConstValue f" when "f \ funs_term t" for f using that by presburger+ show ?A using *(1) by force show ?B using *(2) by force show ?C using *(3) unfolding is_PubConstValue_def by force qed lemma admissible_transactions_no_Value_consts: assumes "admissible_transaction' T" and "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))" shows "\a T. t = Fun (Val a) T" (is ?A) and "\a T. t = Fun (Abs a) T" (is ?B) and "\a T. t = Fun (PubConst Value a) T" (is ?C) using admissible_transaction_terms_no_Value_consts[OF admissible_transaction_is_wellformed_transaction(4)[OF assms(1)] assms(2)] by auto lemma admissible_transactions_no_Value_consts': assumes "admissible_transaction' T" and "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)" shows "\a T. Fun (Val a) T \ subterms t" and "\a T. Fun (Abs a) T \ subterms t" using admissible_transactions_no_Value_consts[OF assms(1)] assms(2) by fast+ lemma admissible_transactions_no_Value_consts'': assumes "admissible_transaction' T" shows "\n. PubConst Value n \ \(funs_term ` trms_transaction T)" and "\n. Abs n \ \(funs_term ` trms_transaction T)" using assms unfolding admissible_transaction'_def admissible_transaction_terms_def by (meson prot_fun.discI(6), meson prot_fun.discI(4)) lemma admissible_transactions_no_PubConsts: assumes "admissible_transaction' T" and "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))" shows "\a n T. t = Fun (PubConst a n) T" proof - have "\is_PubConst f" when "f \ \(funs_term ` (trms_transaction T))" for f using that conjunct1[OF conjunct2[OF admissible_transaction_is_wellformed_transaction(4)[ OF assms(1), unfolded admissible_transaction_terms_def]]] by blast moreover have "f \ \(funs_term ` (trms_transaction T))" when "f \ funs_term t" for f using that assms(2) funs_term_subterms_eq(2)[of "trms_transaction T"] by blast+ ultimately have *: "\is_PubConst f" when "f \ funs_term t" for f using that by presburger+ show ?thesis using * by force qed lemma admissible_transactions_no_PubConsts': assumes "admissible_transaction' T" and "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)" shows "\a n T. Fun (PubConst a n) T \ subterms t" using admissible_transactions_no_PubConsts[OF assms(1)] assms(2) by fast+ lemma admissible_transaction_strand_step_cases: assumes T_adm: "admissible_transaction' T" shows "r \ set (unlabel (transaction_receive T)) \ \t. r = receive\t\" (is "?A \ ?A'") and "r \ set (unlabel (transaction_checks T)) \ (\x s. (r = \Var x in Fun (Set s) []\ \ r = select\Var x, Fun (Set s) []\ \ r = \Var x not in Fun (Set s) []\) \ fst x = TAtom Value \ x \ fv_transaction T - set (transaction_fresh T)) \ (\s t. r = \s == t\ \ r = \s := t\ \ r = \s != t\)" (is "?B \ ?B'") and "r \ set (unlabel (transaction_updates T)) \ \x s. (r = insert\Var x, Fun (Set s) []\ \ r = delete\Var x, Fun (Set s) []\) \ fst x = TAtom Value" (is "?C \ ?C'") and "r \ set (unlabel (transaction_send T)) \ \t. r = send\t\" (is "?D \ ?D'") proof - note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] show "?A \ ?A'" using T_wf Ball_set[of "unlabel (transaction_receive T)" is_Receive] unfolding wellformed_transaction_def is_Receive_def by blast show "?D \ ?D'" using T_wf Ball_set[of "unlabel (transaction_send T)" is_Send] unfolding wellformed_transaction_def is_Send_def by blast show "?B \ ?B'" proof - assume r: ?B note adm_checks = admissible_transaction_is_wellformed_transaction(1,2)[OF T_adm] have fv_r1: "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p r \ fv_transaction T" using r fv_transaction_unfold[of T] by auto have fv_r2: "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p r \ set (transaction_fresh T) = {}" using r T_wf unfolding wellformed_transaction_def by fastforce have "list_all is_Check_or_Assignment (unlabel (transaction_checks T))" using adm_checks(1) unfolding wellformed_transaction_def by blast hence "is_InSet r \ is_Equality r \ is_NegChecks r" using r unfolding list_all_iff by blast thus ?B' proof (elim disjE conjE) assume *: "is_InSet r" hence **: "is_Var (the_elem_term r)" "is_Fun (the_set_term r)" "is_Set (the_Fun (the_set_term r))" "args (the_set_term r) = []" "fst (the_Var (the_elem_term r)) = TAtom Value" using r adm_checks unfolding admissible_transaction_checks_def is_Fun_Set_def by fast+ obtain ac rt rs where r': "r = \ac: rt \ rs\" using * by (cases r) auto obtain x where x: "rt = Var x" "fst x = TAtom Value" using **(1,5) r' by auto obtain f S where fS: "rs = Fun f S" using **(2) r' by auto obtain s where s: "f = Set s" using **(3) fS r' by (cases f) auto hence S: "S = []" using **(4) fS r' by auto show ?B' using r' x fS s S fv_r1 fv_r2 by (cases ac) simp_all next assume *: "is_NegChecks r" hence **: "bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p r = []" "(the_eqs r = [] \ length (the_ins r) = 1) \ (the_ins r = [] \ length (the_eqs r) = 1)" using r adm_checks unfolding admissible_transaction_checks_def by fast+ show ?B' using **(2) proof (elim disjE conjE) assume ***: "the_eqs r = []" "length (the_ins r) = 1" then obtain t s where ts: "the_ins r = [(t,s)]" by (cases "the_ins r") auto hence "hd (the_ins r) = (t,s)" by simp hence ****: "is_Var (fst (t,s))" "is_Fun (snd (t,s))" "is_Set (the_Fun (snd (t,s)))" "args (snd (t,s)) = []" "fst (the_Var (fst (t,s))) = TAtom Value" using * ***(1) Set.bspec[OF adm_checks(2)[unfolded admissible_transaction_checks_def] r] unfolding is_Fun_Set_def by simp_all obtain x where x: "t = Var x" "fst x = TAtom Value" using ts ****(1,5) by (cases t) simp_all obtain f S where fS: "s = Fun f S" using ts ****(2) by (cases s) simp_all obtain ss where ss: "f = Set ss" using fS ****(3) by (cases f) simp_all have S: "S = []" using ts fS ss ****(4) by simp show ?B' using ts x fS ss S *** **(1) * fv_r1 fv_r2 by (cases r) auto next assume ***: "the_ins r = []" "length (the_eqs r) = 1" then obtain t s where "the_eqs r = [(t,s)]" by (cases "the_eqs r") auto thus ?B' using *** **(1) * by (cases r) auto qed qed (auto simp add: is_Equality_def the_check_def intro: poscheckvariant.exhaust) qed show "?C \ ?C'" proof - assume r: ?C note adm_upds = admissible_transaction_is_wellformed_transaction(3)[OF T_adm] have *: "is_Update r" "is_Var (the_elem_term r)" "is_Fun (the_set_term r)" "is_Set (the_Fun (the_set_term r))" "args (the_set_term r) = []" "fst (the_Var (the_elem_term r)) = TAtom Value" using r adm_upds unfolding admissible_transaction_updates_def is_Fun_Set_def by fast+ obtain t s where ts: "r = insert\t,s\ \ r = delete\t,s\" using *(1) by (cases r) auto obtain x where x: "t = Var x" "fst x = TAtom Value" using ts *(2,6) by (cases t) auto obtain f T where fT: "s = Fun f T" using ts *(3) by (cases s) auto obtain ss where ss: "f = Set ss" using ts fT *(4) by (cases f) fastforce+ have T: "T = []" using ts fT *(5) ss by (cases T) auto show ?C' using ts x fT ss T by blast qed qed lemma protocol_transaction_vars_TAtom_typed: assumes T_adm: "admissible_transaction' T" shows "\x \ vars_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" and "\x \ fv_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" and "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" proof - note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] show "\x \ vars_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" using admissible_transactionE(3)[OF T_adm] by fast thus "\x \ fv_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" using vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t by fast show "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" using admissible_transactionE(2)[OF T_adm] by argo qed lemma protocol_transactions_no_pubconsts: assumes "admissible_transaction' T" shows "Fun (Val n) S \ subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction T)" and "Fun (PubConst Value n) S \ subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction T)" using assms admissible_transactions_no_Value_consts(1,3) by (blast, blast) lemma protocol_transactions_no_abss: assumes "admissible_transaction' T" shows "Fun (Abs n) S \ subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction T)" using assms admissible_transactions_no_Value_consts(2) by fast lemma admissible_transaction_strand_sem_fv_ineq: assumes T_adm: "admissible_transaction' T" and \: "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \" and x: "x \ fv_transaction T - set (transaction_fresh T)" and y: "y \ fv_transaction T - set (transaction_fresh T)" and x_not_y: "x \ y" shows "\ x \ \ \ \ y \ \" proof - have "\Var x != Var y\ \ set (unlabel (transaction_checks T)) \ \Var y != Var x\ \ set (unlabel (transaction_checks T))" using x y x_not_y admissible_transactionE(8)[OF T_adm] by auto hence "\Var x != Var y\ \ set (unlabel (transaction_strand T)) \ \Var y != Var x\ \ set (unlabel (transaction_strand T))" unfolding transaction_strand_def unlabel_def by auto hence "\\ x != \ y\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \ \\ y != \ x\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" using stateful_strand_step_subst_inI(8)[of _ _ "unlabel (transaction_strand T)" \] subst_lsst_unlabel[of "transaction_strand T" \] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(7)[of "[]" _ "[]"] by force then obtain B where B: "prefix (B@[\\ x != \ y\]) (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \ prefix (B@[\\ y != \ x\]) (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" unfolding prefix_def by (metis (no_types, opaque_lifting) append.assoc append_Cons append_Nil split_list) thus ?thesis using \ strand_sem_append_stateful[of IK DB _ _ \] stateful_strand_sem_NegChecks_no_bvars(2) unfolding prefix_def by metis qed lemma admissible_transaction_sem_iff: fixes \ and T::"('fun,'atom,'sets,'lbl) prot_transaction" defines "A \ unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" assumes T: "admissible_transaction' T" and I: " interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" shows "strand_sem_stateful M D A I \ ( (\l ts. (l,receive\ts\) \ set (transaction_receive T) \ (\t \ set ts. M \ t \ \ \ I)) \ (\l ac t s. (l,\ac: t \ s\) \ set (transaction_checks T) \ t \ \ \ I = s \ \ \ I) \ (\l ac t s. (l,\ac: t \ s\) \ set (transaction_checks T) \ (t \ \ \ I, s \ \ \ I) \ D) \ (\l t s. (l,\t != s\) \ set (transaction_checks T) \ t \ \ \ I \ s \ \ \ I) \ (\l t s. (l,\t not in s\) \ set (transaction_checks T) \ (t \ \ \ I, s \ \ \ I) \ D))" (is "?A \ ?B") proof - define P where "P \ \X. \\::('fun,'atom,'sets,'lbl) prot_subst. subst_domain \ = set X \ ground (subst_range \)" define rm where "rm \ \X. \\::('fun,'atom,'sets,'lbl) prot_subst. rm_vars (set X) \" define chks where "chks \ transaction_checks T" define q1 where "q1 \ \t X \. t \ rm X \ \ \ \ I" define q2 where "q2 \ \t. t \ \ \ I" note 0 = admissible_transaction_is_wellformed_transaction[OF T] note 1 = wellformed_transaction_sem_iff[OF 0(1) I, of M D \, unfolded A_def[symmetric]] note 2 = admissible_transactionE[OF T] have 3: "rm X \ = \" when "X = []" for X using that unfolding rm_def by auto have 4: "P X \ \ \ = Var" when "X = []" for X and \ using that unfolding P_def by auto have 5: "\t s. \X\\\: F \\: G\ = \t != s\ \ \X\\\: F \\: G\ = \t not in s\" when X: "(l, \X\\\: F \\: G\) \ set chks" for l X F G proof - have *: "\X\\\: F \\: G\ \ set (unlabel (transaction_strand T))" using X transaction_strand_subsets(2)[of T] unlabel_in unfolding chks_def by fast hence **: "X = []" using 2(4) by auto note *** = transaction_notinset_checks_are_Value_vars(3,4)[OF T *] transaction_noteqs_checks_case[OF T *] show ?thesis proof (cases "G = []") case True thus ?thesis using ** ***(3) by blast next case False then obtain t s where g: "(t,s) \ set G" by (cases G) auto show ?thesis using ** ***(1,2)[OF g] by blast qed qed have 6: "q1 t X \ = q2 t" when "P X \" "X = []" for X \ t using that 3 4 unfolding q1_def q2_def by simp let ?negcheck_sem = "\X F G. \\. P X \ \ (\(t,s) \ set F. q1 t X \ \ q1 s X \) \ (\(t,s) \ set G. (q1 t X \, q1 s X \) \ D)" have "(\l X F G. (l,\X\\\: F \\: G\) \ set chks \ ?negcheck_sem X F G) \ ((\l t s. (l,\t != s\) \ set chks \ q2 t \ q2 s) \ (\l t s. (l,\t not in s\) \ set chks \ (q2 t, q2 s) \ D))" (is "?A \ ?B") proof have "q2 t \ q2 s" when t: "(l,\t != s\) \ set chks" ?A for l t s proof - have "?negcheck_sem [] [(t,s)] []" using t by blast thus ?thesis using 4[of "[]"] 6[of "[]"] by force qed moreover have "(q2 t, q2 s) \ D" when t: "(l,\t not in s\) \ set chks" ?A for l t s proof - have "?negcheck_sem [] [] [(t,s)]" using t by blast thus ?thesis using 4[of "[]"] 6[of "[]"] by force qed ultimately show "?A \ ?B" by blast have "?negcheck_sem X F G" when t: "(l,\X\\\: F \\: G\) \ set chks" ?B for l X F G proof - obtain t s where ts: "(X = [] \ F = [(t,s)] \ G = []) \ (X = [] \ F = [] \ G = [(t,s)])" using 5[OF t(1)] by blast hence "(X = [] \ F = [(t,s)] \ G = [] \ q2 t \ q2 s) \ (X = [] \ F = [] \ G = [(t,s)] \ (q2 t, q2 s) \ D)" using t by blast thus ?thesis using 4[of "[]"] 6[of "[]"] by fastforce qed thus "?B \ ?A" by simp qed thus ?thesis using 1 unfolding rm_def chks_def P_def q1_def q2_def by simp qed lemma admissible_transaction_terms_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: assumes "admissible_transaction_terms T" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms_transaction T)" by (rule conjunct1[OF assms[unfolded admissible_transaction_terms_def wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[symmetric]]]) lemma admissible_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: assumes "admissible_transaction' T" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms_transaction T)" proof - have "admissible_transaction_terms T" using assms[unfolded admissible_transaction'_def] by fast thus ?thesis by (metis admissible_transaction_terms_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s) qed lemma admissible_transaction_no_Ana_Attack: assumes "admissible_transaction_terms T" and "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction T)" shows "attack\n\ \ set (snd (Ana t))" proof - obtain r where r: "r \ set (unlabel (transaction_strand T))" "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r)" using assms(2) by force obtain K M where t: "Ana t = (K, M)" by (metis surj_pair) show ?thesis proof assume n: "attack\n\ \ set (snd (Ana t))" hence "attack\n\ \ set M" using t by simp hence n': "attack\n\ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r)" using Ana_subterm[OF t] r(2) subterms_subset by fast hence "\f \ \(funs_term ` trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r). is_Attack f" using funs_term_Fun_subterm' unfolding is_Attack_def by fast hence "is_Send r" "length (the_msgs r) = 1" "is_Fun (hd (the_msgs r))" "is_Attack (the_Fun (hd (the_msgs r)))" "args (hd (the_msgs r)) = []" using assms(1) r(1) unfolding admissible_transaction_terms_def is_Fun_Attack_def by metis+ hence "t = attack\n\" using n' r(2) unfolding is_Send_def is_Attack_def by (cases "the_msgs r") auto thus False using n by fastforce qed qed lemma admissible_transaction_Value_vars: assumes T: "admissible_transaction' T" and x: "x \ fv_transaction T" shows "\\<^sub>v x = TAtom Value" proof - have "x \ vars_transaction T" using x vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] by blast thus "\\<^sub>v x = TAtom Value" using admissible_transactionE(3)[OF T] by simp qed lemma admissible_transaction_occurs_checksE1: assumes T: "admissible_transaction_occurs_checks T" and x: "x \ fv_transaction T - set (transaction_fresh T)" "\\<^sub>v x = TAtom Value" obtains l ts S where "transaction_receive T = (l,receive\ts\)#S" "occurs (Var x) \ set ts" proof - let ?rcvs = "transaction_receive T" let ?frsh = "transaction_fresh T" let ?fvs = "fv_transaction T" have *: "?rcvs \ []" "is_Receive (hd (unlabel ?rcvs))" "\x \ ?fvs - set ?frsh. \\<^sub>v x = TAtom Value \ occurs (Var x) \ set (the_msgs (hd (unlabel ?rcvs)))" using T x unfolding admissible_transaction_occurs_checks_def \\<^sub>v_TAtom''(2) by meson+ obtain r S where S: "?rcvs = r#S" using *(1) by (cases ?rcvs) auto obtain l ts where r: "r = (l,receive\ts\)" by (metis *(1,2) S list.map_sel(1) list.sel(1) prod.collapse is_Receive_def unlabel_def) have 0: "occurs (Var x) \ set ts" using *(3) x S r by fastforce show ?thesis using that[unfolded S r, of l ts S] 0 by blast qed lemma admissible_transaction_occurs_checksE2: assumes T: "admissible_transaction_occurs_checks T" and x: "x \ set (transaction_fresh T)" obtains l ts S where "transaction_send T = (l,send\ts\)#S" "occurs (Var x) \ set ts" proof - let ?snds = "transaction_send T" let ?frsh = "transaction_fresh T" let ?fvs = "fv_transaction T" define ts where "ts \ the_msgs (hd (unlabel ?snds))" let ?P = "\x \ set ?frsh. occurs (Var x) \ set ts" have "?frsh \ []" using x by auto hence *: "?snds \ []" "is_Send (hd (unlabel ?snds))" "?P" using T x unfolding admissible_transaction_occurs_checks_def ts_def by meson+ obtain r S where S: "?snds = r#S" using *(1) by (cases ?snds) auto obtain l where r: "r = (l,send\ts\)" by (metis *(1,2) S list.map_sel(1) list.sel(1) prod.collapse unlabel_def ts_def stateful_strand_step.collapse(1)) have ts: "occurs (Var x) \ set ts" using x *(3) unfolding S by auto show ?thesis using that[unfolded S r, of l ts S] ts by blast qed lemma admissible_transaction_occurs_checksE3: assumes T: "admissible_transaction_occurs_checks T" and t: "OccursFact \ funs_term t \ OccursSec \ funs_term t" "t \ set ts" and ts: "send\ts\ \ set (unlabel (transaction_send T))" obtains x where "t = occurs (Var x)" "x \ set (transaction_fresh T)" proof - let ?P = "\t. \x \ set (transaction_fresh T). t = occurs (Var x)" have "?P t" when "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "OccursFact \ funs_term t \ OccursSec \ funs_term t" for t using assms that unfolding admissible_transaction_occurs_checks_def admissible_transaction_send_occurs_form_def by metis moreover have "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" using t(2) ts unfolding trms\<^sub>s\<^sub>s\<^sub>t_def by fastforce ultimately have "?P t" using t(1) by blast thus thesis using that by blast qed lemma admissible_transaction_occurs_checksE4: assumes T: "admissible_transaction_occurs_checks T" and ts: "send\ts\ \ set (unlabel (transaction_send T))" and t: "occurs t \ set ts" obtains x where "t = Var x" "x \ set (transaction_fresh T)" using admissible_transaction_occurs_checksE3[OF T _ t ts] by auto lemma admissible_transaction_occurs_checksE5: assumes T: "admissible_transaction_occurs_checks T" shows "Fun OccursSec [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" proof - have "\x \ set (transaction_fresh T). t = occurs (Var x)" when "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "OccursFact \ funs_term t \ OccursSec \ funs_term t" for t using assms that unfolding admissible_transaction_occurs_checks_def admissible_transaction_send_occurs_form_def by metis thus ?thesis by fastforce qed lemma admissible_transaction_occurs_checksE6: assumes T: "admissible_transaction_occurs_checks T" and t: "t \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" shows "Fun OccursSec [] \ set (snd (Ana t))" (is ?A) and "occurs k \ set (snd (Ana t))" (is ?B) proof - obtain t' where t': "t' \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "t \ t'" using t by blast have "?A \ ?B" proof (rule ccontr) assume *: "\(?A \ ?B)" hence "OccursSec \ funs_term t' \ OccursFact \ funs_term t'" by (meson t'(2) Ana_subterm' funs_term_Fun_subterm' term.order.trans) then obtain x where x: "x \ set (transaction_fresh T)" "t' = occurs (Var x)" using t'(1) T unfolding admissible_transaction_occurs_checks_def admissible_transaction_send_occurs_form_def by metis have "t = occurs (Var x) \ t = Var x \ t = Fun OccursSec []" using x(2) t'(2) by auto thus False using * by fastforce qed thus ?A ?B by simp_all qed lemma has_initial_value_producing_transactionE: fixes P::"('fun,'atom,'sets,'lbl) prot" assumes P: "has_initial_value_producing_transaction P" and P_adm: "\T \ set P. admissible_transaction' T" obtains T x s ts upds l l' where "\\<^sub>v x = TAtom Value" "Var x \ set ts" "fv\<^sub>s\<^sub>e\<^sub>t (set ts) = {x}" "\n. \(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t set ts)" "T \ set P" "T = Transaction (\(). []) [x] [] [] upds [(l, send\ts\)]" "upds = [] \ (upds = [(l', insert\Var x, \s\\<^sub>s\)] \ (\T \ set P. \(l,a) \ set (transaction_strand T). \t. a \ select\t, \s\\<^sub>s\ \ a \ \t in \s\\<^sub>s\ \ a \ \t not in \s\\<^sub>s\ \ a \ delete\t, \s\\<^sub>s\)) \ T = Transaction (\(). []) [x] [] [] [] [(l, send\ts\)]" proof - define f where "f \ \s. list_all (\T. list_all (\a. ((is_Delete a \ is_InSet a) \ the_set_term a \ \s\\<^sub>s) \ (is_NegChecks a \ list_all (\(_,t). t \ \s\\<^sub>s) (the_ins a))) (unlabel (transaction_checks T@transaction_updates T))) P" obtain T where T0: "T \ set P" "length (transaction_fresh T) = 1" "transaction_receive T = []" "transaction_checks T = []" "length (transaction_send T) = 1" "let x = hd (transaction_fresh T); a = hd (transaction_send T); u = transaction_updates T in is_Send (snd a) \ Var x \ set (the_msgs (snd a)) \ fv\<^sub>s\<^sub>e\<^sub>t (set (the_msgs (snd a))) = {x} \ (u \ [] \ ( let b = hd u; c = snd b in tl u = [] \ is_Insert c \ the_elem_term c = Var x \ is_Fun_Set (the_set_term c) \ f (the_Set (the_Fun (the_set_term c)))))" using P unfolding has_initial_value_producing_transaction_def Let_def list_ex_iff f_def by blast obtain x upds ts h s l l' where T1: "T = Transaction h [x] [] [] upds [(l, send\ts\)]" "Var x \ set ts" "fv\<^sub>s\<^sub>e\<^sub>t (set ts) = {x}" "upds = [] \ (upds = [(l', insert\Var x, \s\\<^sub>s\)] \ f s)" proof (cases T) case T: (Transaction A B C D E F) obtain x where B: "B = [x]" using T0(2) unfolding T by (cases B) auto have C: "C = []" using T0(3) unfolding T by simp have D: "D = []" using T0(4) unfolding T by simp obtain l a where F: "F = [(l,a)]" using T0(5) unfolding T by fastforce obtain ts where ts: "a = send\ts\" using T0(6) unfolding T F by (cases a) auto obtain k u where E: "E = [] \ E = [(k,u)]" using T0(6) unfolding T by (cases E) fastforce+ have x: "Var x \ set ts" "fv\<^sub>s\<^sub>e\<^sub>t (set ts) = {x}" using T0(6) unfolding T B F ts by auto from E show ?thesis proof assume E': "E = [(k,u)]" obtain t t' where u: "u = insert\t,t'\" using T0(6) unfolding T E' by (cases u) auto have t: "t = Var x" using T0(6) unfolding T B E' u Let_def by simp obtain s where t': "t' = \s\\<^sub>s" and s: "f s" using T0(6) unfolding T B E' u Let_def by auto show ?thesis using that[OF T[unfolded B C D F ts E' u t t'] x] s by blast qed (use that[OF T[unfolded B C D F ts] x] in blast) qed note T_adm = bspec[OF P_adm T0(1)] have "x \ set (transaction_fresh T)" using T1(1) by fastforce hence x: "\\<^sub>v x = TAtom Value" using admissible_transactionE(2)[OF T_adm] by fast have "set ts \ trms_transaction T" unfolding T1(1) trms_transaction_unfold by simp hence ts: "\n. \(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t set ts)" using admissible_transactions_no_Value_consts[OF T_adm] by blast have "a \ select\t, \s\\<^sub>s\ \ a \ \t in \s\\<^sub>s\ \ a \ \t not in \s\\<^sub>s\ \ a \ delete\t, \s\\<^sub>s\" when upds: "upds = [(k, insert\Var x,\s\\<^sub>s\)]" and T': "T' \ set P" and la: "(l,a) \ set (transaction_strand T')" for T' l k a t proof - note T'_wf = admissible_transaction_is_wellformed_transaction(1)[OF bspec[OF P_adm T']] have "a \ set (unlabel (transaction_checks T'@transaction_updates T'))" when a': "is_Check_or_Assignment a \ is_Update a" using that wellformed_transaction_strand_unlabel_memberD[OF T'_wf unlabel_in[OF la]] by (cases a) auto note 0 = this T1(4) T' note 1 = upds f_def list_all_iff show ?thesis proof (cases a) case (Delete t' s') thus ?thesis using 0 unfolding 1 by fastforce next case (InSet ac t' s') thus ?thesis using 0 unfolding 1 by fastforce next case (NegChecks X F G) thus ?thesis using 0 unfolding 1 by fastforce qed auto qed hence s: "\T \ set P. \(l,a) \ set (transaction_strand T). \t. a \ select\t, \s\\<^sub>s\ \ a \ \t in \s\\<^sub>s\ \ a \ \t not in \s\\<^sub>s\ \ a \ delete\t, \s\\<^sub>s\" when upds: "upds = [(k, insert\Var x,\s\\<^sub>s\)]" for k using upds by force have h: "h = (\(). [])" proof - have "transaction_decl T = h" using T1(1) by fastforce hence "h a = []" for a using admissible_transactionE(1)[OF T_adm] by simp thus ?thesis using ext[of h "\(). []"] by (metis case_unit_Unity) qed show ?thesis using that[OF x T1(2,3) ts T0(1)] T1(1,4) s unfolding h by auto qed lemma has_initial_value_producing_transaction_update_send_ex_filter: fixes P::"('a,'b,'c,'d) prot" defines "f \ \T. transaction_fresh T = [] \ list_ex (\a. is_Update (snd a) \ is_Send (snd a)) (transaction_strand T)" assumes P: "has_initial_value_producing_transaction P" shows "has_initial_value_producing_transaction (filter f P)" proof - define g where "g \ \P::('a,'b,'c,'d) prot. \s. list_all (\T. list_all (\a. ((is_Delete a \ is_InSet a) \ the_set_term a \ \s\\<^sub>s) \ (is_NegChecks a \ list_all (\(_,t). t \ \s\\<^sub>s) (the_ins a))) (unlabel (transaction_checks T@transaction_updates T))) P" let ?Q = "\P T. let x = hd (transaction_fresh T); a = hd (transaction_send T); u = transaction_updates T in is_Send (snd a) \ Var x \ set (the_msgs (snd a)) \ fv\<^sub>s\<^sub>e\<^sub>t (set (the_msgs (snd a))) = {x} \ (u \ [] \ ( let b = hd u; c = snd b in tl u = [] \ is_Insert c \ the_elem_term c = Var x \ is_Fun_Set (the_set_term c) \ g P (the_Set (the_Fun (the_set_term c)))))" have "set (filter f P) \ set P" by simp hence "list_all h P \ list_all h (filter f P)" for h unfolding list_all_iff by blast hence g_f_subset: "g P s \ g (filter f P) s" for s unfolding g_def by blast obtain T where T: "T \ set P" "length (transaction_fresh T) = 1" "transaction_receive T = []" "transaction_checks T = []" "length (transaction_send T) = 1" "?Q P T" using P unfolding has_initial_value_producing_transaction_def Let_def list_ex_iff g_def by blast obtain x where x: "transaction_fresh T = [x]" using T(2) by blast obtain a where a: "transaction_send T = [a]" using T(5) by blast obtain l b where b: "a = (l,b)" by (cases a) auto obtain ts where ts: "b = send\ts\" using T(6) unfolding Let_def a b by (cases b) auto have "T \ set (filter f P)" using T(1) x a unfolding b ts f_def by auto moreover have "?Q (filter f P) T" using T(6) g_f_subset by meson ultimately show ?thesis using T(2-5) unfolding has_initial_value_producing_transaction_def Let_def list_ex_iff g_def by blast qed subsection \Lemmata: Renaming, Declaration, and Fresh Substitutions\ lemma transaction_decl_subst_empty_inv: assumes "transaction_decl_subst Var T" shows "transaction_decl T () = []" using assms unfolding transaction_decl_subst_def subst_domain_Var by blast lemma transaction_decl_subst_domain: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_decl_subst \ T" shows "subst_domain \ = fst ` set (transaction_decl T ())" using assms unfolding transaction_decl_subst_def by argo lemma transaction_decl_subst_grounds_domain: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_decl_subst \ T" and "x \ fst ` set (transaction_decl T ())" shows "fv (\ x) = {}" proof - obtain c where "\ x = Fun c []" using assms unfolding transaction_decl_subst_def by moura thus ?thesis by simp qed lemma transaction_decl_subst_range_vars_empty: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_decl_subst \ T" shows "range_vars \ = {}" using assms unfolding transaction_decl_subst_def range_vars_def by auto lemma transaction_decl_subst_wt: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_decl_subst \ T" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using assms unfolding transaction_decl_subst_def by blast lemma transaction_decl_subst_is_wf_trm: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_decl_subst \ P" shows "wf\<^sub>t\<^sub>r\<^sub>m (\ v)" proof (cases "v \ subst_domain \") case True thus ?thesis using assms unfolding transaction_decl_subst_def by fastforce qed auto lemma transaction_decl_subst_range_wf_trms: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_decl_subst \ P" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (metis transaction_decl_subst_is_wf_trm[OF assms] wf_trm_subst_range_iff) lemma transaction_renaming_subst_is_renaming: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes \: "transaction_renaming_subst \ P A" shows "\m. \\ n. \ (\,n) = Var (\,n+Suc m)" (is ?A) and "\y. \ x = Var y" (is ?B) and "\ x \ Var x" (is ?C) and "subst_domain \ = UNIV" (is ?D) and "subst_range \ \ range Var" (is ?E) and "fv (t \ \) \ range_vars \" (is ?F) proof - show 0: ?A using \ unfolding transaction_renaming_subst_def var_rename_def by force show ?B using \ unfolding transaction_renaming_subst_def var_rename_def by blast show ?C using 0 by (cases x) auto show 1: ?D using 0 by fastforce show ?E using 0 by auto show ?F by (induct t) (auto simp add: 1 subst_dom_vars_in_subst subst_fv_imgI) qed lemma transaction_renaming_subst_is_injective: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_renaming_subst \ P A" shows "inj \" proof (intro injI) fix x y::"('fun,'atom,'sets,'lbl) prot_var" obtain \x nx where x: "x = (\x,nx)" by (metis surj_pair) obtain \y ny where y: "y = (\y,ny)" by (metis surj_pair) obtain m where m: "\\. \n. \ (\, n) = Var (\, n + Suc m)" using transaction_renaming_subst_is_renaming(1)[OF assms] by blast assume "\ x = \ y" hence "\x = \y" "nx = ny" using x y m by simp_all thus "x = y" using x y by argo qed lemma transaction_renaming_subst_vars_disj: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" shows "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (\(vars_transaction ` set P))) \ (\(vars_transaction ` set P)) = {}" (is ?A) and "fv\<^sub>s\<^sub>e\<^sub>t (\ ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = {}" (is ?B) and "T \ set P \ vars_transaction T \ range_vars \ = {}" (is "T \ set P \ ?C1") and "T \ set P \ bvars_transaction T \ range_vars \ = {}" (is "T \ set P \ ?C2") and "T \ set P \ fv_transaction T \ range_vars \ = {}" (is "T \ set P \ ?C3") and "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ range_vars \ = {}" (is ?D1) and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ range_vars \ = {}" (is ?D2) and "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ range_vars \ = {}" (is ?D3) proof - define X where "X \ \(vars_transaction ` set P) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" have 1: "finite X" by (simp add: X_def) obtain n where n: "n \ max_var_set X" "\ = var_rename n" using assms unfolding transaction_renaming_subst_def X_def by moura hence 2: "\x \ X. snd x < Suc n" using less_Suc_max_var_set[OF _ 1] unfolding var_rename_def by fastforce have 3: "x \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` X)" "fv (\ x) \ X = {}" "x \ range_vars \" when x: "x \ X" for x using 2 x n unfolding var_rename_def by force+ show ?A ?B using 3(1,2) unfolding X_def by auto show ?C1 when T: "T \ set P" using T 3(3) unfolding X_def by blast thus ?C2 ?C3 when T: "T \ set P" using T by (simp_all add: disjoint_iff_not_equal vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t) show ?D1 using 3(3) unfolding X_def by auto thus ?D2 ?D3 by (simp_all add: disjoint_iff_not_equal vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t) qed lemma transaction_renaming_subst_wt: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_renaming_subst \ P X" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" proof - { fix x::"('fun,'atom,'sets,'lbl) prot_var" obtain \ n where x: "x = (\,n)" by moura then obtain m where m: "\ x = Var (\,m)" using assms transaction_renaming_subst_is_renaming(1) by moura hence "\ (\ x) = \\<^sub>v x" using x by (simp add: \\<^sub>v_def) } thus ?thesis by (simp add: wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) qed lemma transaction_renaming_subst_is_wf_trm: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_renaming_subst \ P X" shows "wf\<^sub>t\<^sub>r\<^sub>m (\ v)" proof - obtain \ n where "v = (\, n)" by moura then obtain m where "\ v = Var (\, n + Suc m)" using transaction_renaming_subst_is_renaming(1)[OF assms] by moura thus ?thesis by (metis wf_trm_Var) qed lemma transaction_renaming_subst_range_wf_trms: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_renaming_subst \ P X" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (metis transaction_renaming_subst_is_wf_trm[OF assms] wf_trm_subst_range_iff) lemma transaction_renaming_subst_range_notin_vars: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" shows "\y. \ x = Var y \ y \ \(vars_transaction ` set P) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" proof - obtain \ n where x: "x = (\,n)" by (metis surj_pair) define y where "y \ \m. (\,n+Suc m)" have "\m \ max_var_set (\(vars_transaction ` set P) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A). \ x = Var (y m)" using assms x by (auto simp add: y_def transaction_renaming_subst_def var_rename_def) moreover have "finite (\(vars_transaction ` set P) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" by auto ultimately show ?thesis using x unfolding y_def by force qed lemma transaction_renaming_subst_var_obtain: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes \: "transaction_renaming_subst \ P X" shows "x \ fv\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \) \ \y. \ y = Var x" (is "?A1 \ ?B1") and "x \ fv (t \ \) \ \y \ fv t. \ y = Var x" (is "?A2 \ ?B2") proof - assume x: ?A1 obtain y where y: "y \ fv\<^sub>s\<^sub>s\<^sub>t S" "x \ fv (\ y)" using fv\<^sub>s\<^sub>s\<^sub>t_subst_obtain_var[OF x] by moura thus ?B1 using transaction_renaming_subst_is_renaming(2)[OF \, of y] by fastforce next assume x: ?A2 obtain y where y: "y \ fv t" "x \ fv (\ y)" using fv_subst_obtain_var[OF x] by moura thus ?B2 using transaction_renaming_subst_is_renaming(2)[OF \, of y] by fastforce qed lemma transaction_renaming_subst_set_eq: assumes "set P1 = set P2" shows "transaction_renaming_subst \ P1 X = transaction_renaming_subst \ P2 X" (is "?A = ?B") using assms unfolding transaction_renaming_subst_def by presburger lemma transaction_renaming_subst_vars_transaction_neq: assumes T: "T \ set P" and \: "transaction_renaming_subst \ P vars" and vars:"finite vars" and x: "x \ vars_transaction T" shows "\ y \ Var x" proof - have "\n. \ = var_rename n \ n \ max_var_set (\(vars_transaction ` set P))" using T \ vars x unfolding transaction_renaming_subst_def by auto then obtain n where n_p: "\ = var_rename n" "n \ max_var_set (\(vars_transaction ` set P))" by blast moreover have "\(vars_transaction ` set P) \ vars_transaction T" using T by blast ultimately have n_gt: "n \ max_var_set (vars_transaction T)" by auto obtain a b where ab: "x = (a,b)" by (cases x) auto obtain c d where cd: "y = (c,d)" by (cases y) auto have nb: "n \ b" using n_gt x ab by auto have "\ y = \ (c, d)" using cd by auto moreover have "... = Var (c, Suc (d + n))" unfolding n_p(1) unfolding var_rename_def by simp moreover have "... \ Var x" using nb ab by auto ultimately show ?thesis by auto qed lemma transaction_renaming_subst_fv_disj: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" shows "fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = {}" proof - have "fv\<^sub>s\<^sub>e\<^sub>t (\ ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = {}" using assms transaction_renaming_subst_vars_disj(2) by blast moreover have "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" by (simp add: vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t) ultimately show ?thesis by auto qed lemma transaction_fresh_subst_is_wf_trm: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_fresh_subst \ T X" shows "wf\<^sub>t\<^sub>r\<^sub>m (\ v)" proof (cases "v \ subst_domain \") case True then obtain c where "\ v = Fun c []" "arity c = 0" using assms unfolding transaction_fresh_subst_def by moura thus ?thesis by auto qed auto lemma transaction_fresh_subst_wt: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_fresh_subst \ T X" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using assms unfolding transaction_fresh_subst_def by blast lemma transaction_fresh_subst_domain: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_fresh_subst \ T X" shows "subst_domain \ = set (transaction_fresh T)" using assms unfolding transaction_fresh_subst_def by fast lemma transaction_fresh_subst_range_wf_trms: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_fresh_subst \ T X" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (metis transaction_fresh_subst_is_wf_trm[OF assms] wf_trm_subst_range_iff) lemma transaction_fresh_subst_range_fresh: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_fresh_subst \ T M" shows "\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t M" and "\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))" using assms unfolding transaction_fresh_subst_def by meson+ lemma transaction_fresh_subst_sends_to_val: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes \: "transaction_fresh_subst \ T X" and y: "y \ set (transaction_fresh T)" "\\<^sub>v y = TAtom Value" obtains n where "\ y = Fun (Val n) []" "Fun (Val n) [] \ subst_range \" proof - have "\ y \ subst_range \" using assms unfolding transaction_fresh_subst_def by simp obtain c where c: "\ y = Fun c []" "\public c" "arity c = 0" using \ y(1) unfolding transaction_fresh_subst_def by fastforce have "\ (\ y) = TAtom Value" using \ y(2) \\<^sub>v_TAtom''(2)[of y] wt_subst_trm''[of \ "Var y"] unfolding transaction_fresh_subst_def by simp then obtain n where "c = Val n" using c by (cases c) (auto split: option.splits) thus ?thesis using c that unfolding transaction_fresh_subst_def by fastforce qed lemma transaction_fresh_subst_sends_to_val': fixes \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_fresh_subst \ T X" and "y \ set (transaction_fresh T)" "\\<^sub>v y = TAtom Value" obtains n where "(\ \\<^sub>s \) y \ \ = Fun (Val n) []" "Fun (Val n) [] \ subst_range \" proof - obtain n where "\ y = Fun (Val n) []" "Fun (Val n) [] \ subst_range \" using transaction_fresh_subst_sends_to_val[OF assms] by moura thus ?thesis using that by (fastforce simp add: subst_compose_def) qed lemma transaction_fresh_subst_grounds_domain: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_fresh_subst \ T X" and "y \ set (transaction_fresh T)" shows "fv (\ y) = {}" proof - obtain c where "\ y = Fun c []" using assms unfolding transaction_fresh_subst_def by moura thus ?thesis by simp qed lemma transaction_fresh_subst_range_vars_empty: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_fresh_subst \ T X" shows "range_vars \ = {}" proof - have "fv t = {}" when "t \ subst_range \" for t using assms that unfolding transaction_fresh_subst_def by fastforce thus ?thesis unfolding range_vars_def by blast qed lemma transaction_decl_fresh_renaming_substs_range: fixes \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T M" and \: "transaction_renaming_subst \ P X" shows "x \ fst ` set (transaction_decl T ()) \ \c. (\ \\<^sub>s \ \\<^sub>s \) x = Fun c [] \ arity c = 0" and "x \ fst ` set (transaction_decl T ()) \ x \ set (transaction_fresh T) \ \c. (\ \\<^sub>s \ \\<^sub>s \) x = Fun c [] \ \public c \ arity c = 0" and "x \ fst ` set (transaction_decl T ()) \ x \ set (transaction_fresh T) \ fst x = TAtom Value \ \n. (\ \\<^sub>s \ \\<^sub>s \) x = Fun (Val n) []" and "x \ fst ` set (transaction_decl T ()) \ x \ set (transaction_fresh T) \ \y. (\ \\<^sub>s \ \\<^sub>s \) x = Var y" proof - assume "x \ fst ` set (transaction_decl T ())" then obtain c where c: "\ x = Fun c []" "arity c = 0" using \ unfolding transaction_decl_subst_def by fastforce thus "\c. (\ \\<^sub>s \ \\<^sub>s \) x = Fun c [] \ arity c = 0" using subst_compose[of "\ \\<^sub>s \" \ x] subst_compose[of \ \ x] by simp next assume x: "x \ fst ` set (transaction_decl T ())" "x \ set (transaction_fresh T)" have *: "(\ \\<^sub>s \) x = \ x" using x(1) \ unfolding transaction_decl_subst_def by (metis (no_types, opaque_lifting) subst_comp_notin_dom_eq) then obtain c where c: "(\ \\<^sub>s \) x = Fun c []" "\public c" "arity c = 0" using \ x(2) unfolding transaction_fresh_subst_def by fastforce thus "\c. (\ \\<^sub>s \ \\<^sub>s \) x = Fun c [] \ \public c \ arity c = 0" using subst_compose[of "\ \\<^sub>s \" \ x] subst_compose[of \ \ x] by simp assume "fst x = TAtom Value" hence "\ ((\ \\<^sub>s \) x) = TAtom Value" using * \ \\<^sub>v_TAtom''(2)[of x] wt_subst_trm''[of \ "Var x"] unfolding transaction_fresh_subst_def by simp then obtain n where "c = Val n" using c by (cases c) (auto split: option.splits) thus "\n. (\ \\<^sub>s\ \\<^sub>s \) x = Fun (Val n) []" using c subst_compose[of "\ \\<^sub>s \" \ x] subst_compose[of \ \ x] by simp next assume "x \ fst ` set (transaction_decl T ())" "x \ set (transaction_fresh T)" hence "(\ \\<^sub>s \) x = Var x" using \ \ unfolding transaction_decl_subst_def transaction_fresh_subst_def by (metis (no_types, opaque_lifting) subst_comp_notin_dom_eq subst_domI) thus "\y. (\ \\<^sub>s \ \\<^sub>s \) x = Var y" using transaction_renaming_subst_is_renaming(1)[OF \] subst_compose[of "\ \\<^sub>s \" \ x] subst_compose[of \ \ x] by (cases x) force qed lemma transaction_decl_fresh_renaming_substs_range': fixes \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T M" and \: "transaction_renaming_subst \ P X" and t: "t \ subst_range (\ \\<^sub>s \ \\<^sub>s \)" shows "(\c. t = Fun c [] \ arity c = 0) \ (\x. t = Var x)" and "\ = Var \ (\c. t = Fun c [] \ \public c \ arity c = 0) \ (\x. t = Var x)" and "\ = Var \ \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\n. t = Fun (Val n) []) \ (\x. t = Var x)" and "\ = Var \ is_Fun t \ t \ subst_range \" proof - obtain x where x: "x \ subst_domain (\ \\<^sub>s \ \\<^sub>s \)" "(\ \\<^sub>s \ \\<^sub>s \) x = t" using t by auto note 0 = x transaction_decl_fresh_renaming_substs_range[OF \ \ \, of x] show "(\c. t = Fun c [] \ arity c = 0) \ (\x. t = Var x)" using 0 unfolding \\<^sub>v_TAtom'' by auto assume 1: "\ = Var" note 2 = transaction_decl_subst_empty_inv[OF \[unfolded 1]] show 3: "(\c. t = Fun c [] \ \public c \ arity c = 0) \ (\x. t = Var x)" using 0 2 unfolding \\<^sub>v_TAtom'' by auto show "(\n. t = Fun (Val n) []) \ (\x. t = Var x)" when "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" using 0 2 that unfolding \\<^sub>v_TAtom'' by auto show "t \ subst_range \" when t': "is_Fun t" proof - obtain x where x: "(\ \\<^sub>s \) x = t" using t 1 by auto show ?thesis proof (cases "x \ subst_domain \") case True thus ?thesis by (metis subst_dom_vars_in_subst subst_ground_ident_compose(1) subst_imgI x transaction_fresh_subst_grounds_domain[OF \] transaction_fresh_subst_domain[OF \]) next case False thus ?thesis by (metis (no_types, lifting) subst_compose_def subst_domI term.disc(1) that transaction_renaming_subst_is_renaming(5)[OF \] var_renaming_is_Fun_iff x) qed qed qed lemma transaction_decl_fresh_renaming_substs_range'': fixes \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and y: "y \ fv ((\ \\<^sub>s \ \\<^sub>s \) x)" shows "\ x = Var x" and "\ x = Var x" and "\ x = Var y" and "(\ \\<^sub>s \ \\<^sub>s \) x = Var y" proof - have "\z. z \ fv (\ x)" by (metis y subst_compose_fv') hence "x \ subst_domain \" using y transaction_decl_subst_domain[OF \] transaction_decl_subst_grounds_domain[OF \, of x] by blast thus 0: "\ x = Var x" by blast hence "y \ fv ((\ \\<^sub>s \) x)" using y by (simp add: subst_compose) hence "\z. z \ fv (\ x)" by (metis subst_compose_fv') hence "x \ subst_domain \" using y transaction_fresh_subst_domain[OF \] transaction_fresh_subst_grounds_domain[OF \, of x] by blast thus 1: "\ x = Var x" by blast show "\ x = Var y" "(\ \\<^sub>s \ \\<^sub>s \) x = Var y" using 0 1 y transaction_renaming_subst_is_renaming(2)[OF \, of x] unfolding subst_compose_def by (fastforce,fastforce) qed lemma transaction_decl_fresh_renaming_substs_vars_subset: fixes \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "\(fv_transaction ` set P) \ subst_domain (\ \\<^sub>s \ \\<^sub>s \)" (is ?A) and "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ subst_domain (\ \\<^sub>s \ \\<^sub>s \)" (is ?B) and "T' \ set P \ fv_transaction T' \ subst_domain (\ \\<^sub>s \ \\<^sub>s \)" (is "T' \ set P \ ?C") and "T' \ set P \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ \\<^sub>s \ \\<^sub>s \)) \ range_vars (\ \\<^sub>s \ \\<^sub>s \)" (is "T' \ set P \ ?D") proof - have *: "x \ subst_domain (\ \\<^sub>s \ \\<^sub>s \)" for x proof (cases "x \ subst_domain \") case True thus ?thesis using transaction_decl_subst_domain[OF \] transaction_decl_subst_grounds_domain[OF \] by (simp add: subst_domI subst_dom_vars_in_subst subst_ground_ident_compose(1)) next case False hence \_x_eq: "(\ \\<^sub>s \ \\<^sub>s \) x = (\ \\<^sub>s \) x" by (auto simp add: subst_compose) show ?thesis proof (cases "x \ subst_domain \") case True hence "x \ {x. \y. \ x = Var y \ \ y = Var x}" using transaction_fresh_subst_domain[OF \] transaction_fresh_subst_grounds_domain[OF \, of x] by auto hence "x \ subst_domain (\ \\<^sub>s \)" using subst_domain_subst_compose[of \ \] by blast thus ?thesis using \_x_eq subst_dom_vars_in_subst by fastforce next case False hence "(\ \\<^sub>s \) x = \ x" unfolding subst_compose_def by fastforce moreover have "\ x \ Var x" using transaction_renaming_subst_is_renaming(1)[OF \] by (cases x) auto ultimately show ?thesis using \_x_eq by fastforce qed qed show ?A ?B using * by blast+ show ?C when T: "T' \ set P" using T * by blast hence "fv\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T') \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \) \ range_vars (\ \\<^sub>s \ \\<^sub>s \)" when T: "T' \ set P" using T fv\<^sub>s\<^sub>s\<^sub>t_subst_subset_range_vars_if_subset_domain by blast thus ?D when T: "T' \ set P" by (metis T unlabel_subst) qed lemma transaction_decl_fresh_renaming_substs_vars_disj: fixes \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \ \\<^sub>s \) ` (\(vars_transaction ` set P))) \ (\(vars_transaction ` set P)) = {}" (is ?A) and "x \ \(vars_transaction ` set P) \ fv ((\ \\<^sub>s \ \\<^sub>s \) x) \ (\(vars_transaction ` set P)) = {}" (is "?B' \ ?B") and "T' \ set P \ vars_transaction T' \ range_vars (\ \\<^sub>s \ \\<^sub>s \) = {}" (is "T' \ set P \ ?C1") and "T' \ set P \ bvars_transaction T' \ range_vars (\ \\<^sub>s \ \\<^sub>s \) = {}" (is "T' \ set P \ ?C2") and "T' \ set P \ fv_transaction T' \ range_vars (\ \\<^sub>s \ \\<^sub>s \) = {}" (is "T' \ set P \ ?C3") and "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ range_vars (\ \\<^sub>s \ \\<^sub>s \) = {}" (is ?D1) and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ range_vars (\ \\<^sub>s \ \\<^sub>s \) = {}" (is ?D2) and "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ range_vars (\ \\<^sub>s \ \\<^sub>s \) = {}" (is ?D3) and "range_vars \ = {}" (is ?E1) and "range_vars \ = {}" (is ?E2) and "range_vars (\ \\<^sub>s \ \\<^sub>s \) \ range_vars \" (is ?E3) proof - note 0 = transaction_renaming_subst_vars_disj[OF \] show ?A proof (cases "fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \ \\<^sub>s \) ` (\(vars_transaction ` set P))) = {}") case False hence "\x \ (\(vars_transaction ` set P)). (\ \\<^sub>s \ \\<^sub>s \) x = \ x \ fv ((\ \\<^sub>s \ \\<^sub>s \) x) = {}" using transaction_decl_fresh_renaming_substs_range''[OF \ \ \] by auto thus ?thesis using 0(1) by force qed blast thus "?B' \ ?B" by auto show ?E1 ?E2 using transaction_fresh_subst_grounds_domain[OF \] transaction_decl_subst_grounds_domain[OF \] unfolding transaction_fresh_subst_domain[OF \, symmetric] transaction_decl_subst_domain[OF \, symmetric] by (fastforce, fastforce) thus 1: ?E3 using range_vars_subst_compose_subset[of \ \] range_vars_subst_compose_subset[of "\ \\<^sub>s \" \] by blast show ?C1 ?C2 ?C3 when T: "T' \ set P" using T 1 0(3,4,5)[of T'] by blast+ show ?D1 ?D2 ?D3 using 1 0(6,7,8) by blast+ qed lemma transaction_decl_fresh_renaming_substs_trms: fixes \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ subst_domain \ = {}" and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ subst_domain \ = {}" and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ subst_domain \ = {}" shows "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ \\<^sub>s \ \\<^sub>s \))) = subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S) \\<^sub>s\<^sub>e\<^sub>t (\ \\<^sub>s \ \\<^sub>s \)" proof - have 1: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S). (\f. (\ \\<^sub>s \ \\<^sub>s \) x = Fun f []) \ (\y. (\ \\<^sub>s \ \\<^sub>s \) x = Var y)" using transaction_decl_fresh_renaming_substs_range'[OF \ \ \] by blast have 2: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ subst_domain (\ \\<^sub>s \ \\<^sub>s \) = {}" using assms(4-6) subst_domain_compose[of \ \] subst_domain_compose[of "\ \\<^sub>s \" \] by blast show ?thesis using subterms_subst_lsst[OF 1 2] by simp qed lemma transaction_decl_fresh_renaming_substs_wt: fixes \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_decl_subst \ T" "transaction_fresh_subst \ T M" "transaction_renaming_subst \ P X" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \ \\<^sub>s \)" using transaction_renaming_subst_wt[OF assms(3)] transaction_fresh_subst_wt[OF assms(2)] transaction_decl_subst_wt[OF assms(1)] by (metis wt_subst_compose) lemma transaction_decl_fresh_renaming_substs_range_wf_trms: fixes \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_decl_subst \ T" "transaction_fresh_subst \ T M" "transaction_renaming_subst \ P X" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \ \\<^sub>s \))" using transaction_renaming_subst_range_wf_trms[OF assms(3)] transaction_fresh_subst_range_wf_trms[OF assms(2)] transaction_decl_subst_range_wf_trms[OF assms(1)] wf_trms_subst_compose[of \ \] wf_trms_subst_compose[of "\ \\<^sub>s \" \] by metis lemma transaction_decl_fresh_renaming_substs_fv: fixes \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T M" and \: "transaction_renaming_subst \ P X" and x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" shows "\y \ fv_transaction T - set (transaction_fresh T). (\ \\<^sub>s \ \\<^sub>s \) y = Var x" proof - have "x \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T) \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using x fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] by argo then obtain y where "y \ fv_transaction T" "x \ fv ((\ \\<^sub>s \ \\<^sub>s \) y)" by (metis fv\<^sub>s\<^sub>s\<^sub>t_subst_obtain_var) thus ?thesis using transaction_decl_fresh_renaming_substs_range[OF \ \ \, of y] by (cases "y \ set (transaction_fresh T)") force+ qed lemma transaction_decl_fresh_renaming_substs_range_no_attack_const: fixes \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T M" and \: "transaction_renaming_subst \ P X" and T: "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" and t: "t \ subst_range (\ \\<^sub>s \ \\<^sub>s \)" shows "\n. t = attack\n\" proof - note \\\_wt = transaction_decl_fresh_renaming_substs_wt[OF \ \ \] obtain x where x: "(\ \\<^sub>s \ \\<^sub>s \) x = t" using t by auto have x_type: "\ (Var x) = \ (Var x \ \)" "\ (Var x) = \ (Var x \ \ \\<^sub>s \ \\<^sub>s \)" using \ wt_subst_trm''[of \ "Var x"] wt_subst_trm''[OF \\\_wt, of "Var x"] unfolding transaction_decl_subst_def by (blast, blast) show ?thesis proof (cases t) case (Fun f S) hence "x \ set (transaction_fresh T) \ x \ fst ` set (transaction_decl T ())" using transaction_decl_fresh_renaming_substs_range[OF \ \ \, of x] x by force thus ?thesis proof assume "x \ set (transaction_fresh T)" hence "\ t = TAtom Value \ (\a. \ t = TAtom (Atom a))" - using T x_type(2) x by (metis \.simps(1) subst_apply_term.simps(1)) + using T x_type(2) x by (metis \.simps(1) eval_term.simps(1)) thus ?thesis by auto next assume "x \ fst ` set (transaction_decl T ())" then obtain c where c: "\ x = Fun (Fu c) []" "arity\<^sub>f c = 0" using \ unfolding transaction_decl_subst_def by auto have "\ t = TAtom Bottom \ (\a. \ t = TAtom (Atom a))" using c(1) \_consts_simps(1)[OF c(2)] x x_type - subst_apply_term.simps(1)[of x \] subst_apply_term.simps(1)[of x "\ \\<^sub>s \ \\<^sub>s \"] + eval_term.simps(1)[of _ x \] eval_term.simps(1)[of _ x "\ \\<^sub>s \ \\<^sub>s \"] by (cases "\\<^sub>f c") simp_all thus ?thesis by auto qed qed simp qed lemma transaction_decl_fresh_renaming_substs_occurs_fact_send_receive: fixes t::"('fun,'atom,'sets,'lbl) prot_term" assumes \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T M" and \: "transaction_renaming_subst \ P X" and T: "admissible_transaction' T" and t: "occurs t \ set ts" shows "send\ts\ \ set (unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \ \ts' s. send\ts'\ \ set (unlabel (transaction_send T)) \ occurs s \ set ts' \ t = s \ \ \\<^sub>s \ \\<^sub>s \" (is "?A \ ?A'") and "receive\ts\ \ set (unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \ \ts' s. receive\ts'\ \ set (unlabel (transaction_receive T)) \ occurs s \ set ts' \ t = s \ \ \\<^sub>s \ \\<^sub>s \" (is "?B \ ?B'") proof - assume ?A then obtain s ts' where s: "s \ set ts'" "send\ts'\ \ set (unlabel (transaction_strand T))" "occurs t = s \ \ \\<^sub>s \ \\<^sub>s \" using t stateful_strand_step_mem_substD(1)[ of ts "unlabel (transaction_strand T)" "\ \\<^sub>s \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] by auto note \_empty = admissible_transaction_decl_subst_empty[OF T \] have T_decl_notin: "x \ fst ` set (transaction_decl T ())" for x using transaction_decl_subst_empty_inv[OF \[unfolded \_empty]] by simp note 0 = s(3) transaction_decl_fresh_renaming_substs_range[OF \ \ \] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T] note T_fresh = admissible_transactionE(14)[OF T] have "\u. s = occurs u" proof (cases s) case (Var x) hence "(\c. s \ \ \\<^sub>s \ \\<^sub>s \ = Fun c []) \ (\y. s \ \ \\<^sub>s \ \\<^sub>s \ = Var y)" using 0(2-5)[of x] \_empty by (auto simp del: subst_subst_compose) thus ?thesis using 0(1) by simp next case (Fun f T) hence 1: "f = OccursFact" "length T = 2" "T ! 0 \ \ \\<^sub>s \ \\<^sub>s \ = Fun OccursSec []" "T ! 1 \ \ \\<^sub>s \ \\<^sub>s \ = t" using 0(1) by auto have "T ! 0 = Fun OccursSec []" proof (cases "T ! 0") case (Var x) thus ?thesis using 0(2-5)[of x] 1(3) T_fresh T_decl_notin unfolding list_all_iff by (auto simp del: subst_subst_compose) qed (use 1(3) in simp) thus ?thesis using Fun 1 0(1) by (auto simp del: subst_subst_compose) qed then obtain u where u: "s = occurs u" by moura hence "t = u \ \ \\<^sub>s \ \\<^sub>s \" using s(3) by fastforce thus ?A' using s u wellformed_transaction_strand_unlabel_memberD(8)[OF T_wf] by metis next assume ?B then obtain s ts' where s: "s \ set ts'" "receive\ts'\ \ set (unlabel (transaction_strand T))" "occurs t = s \ \ \\<^sub>s \ \\<^sub>s \" using t stateful_strand_step_mem_substD(2)[ of ts "unlabel (transaction_strand T)" "\ \\<^sub>s \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] by auto note \_empty = admissible_transaction_decl_subst_empty[OF T \] have T_decl_notin: "x \ fst ` set (transaction_decl T ())" for x using transaction_decl_subst_empty_inv[OF \[unfolded \_empty]] by simp note 0 = s(3) transaction_decl_fresh_renaming_substs_range[OF \ \ \] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T] note T_fresh = admissible_transactionE(14)[OF T] have "\u. s = occurs u" proof (cases s) case (Var x) hence "(\c. s \ \ \\<^sub>s \ \\<^sub>s \ = Fun c []) \ (\y. s \ \ \\<^sub>s \ \\<^sub>s \ = Var y)" using 0(2-5)[of x] \_empty by (auto simp del: subst_subst_compose) thus ?thesis using 0(1) by simp next case (Fun f T) hence 1: "f = OccursFact" "length T = 2" "T ! 0 \ \ \\<^sub>s \ \\<^sub>s \ = Fun OccursSec []" "T ! 1 \ \ \\<^sub>s \ \\<^sub>s \ = t" using 0(1) by auto have "T ! 0 = Fun OccursSec []" proof (cases "T ! 0") case (Var x) thus ?thesis using 0(2-5)[of x] 1(3) T_fresh T_decl_notin unfolding list_all_iff by (auto simp del: subst_subst_compose) qed (use 1(3) in simp) thus ?thesis using Fun 1 0(1) by (auto simp del: subst_subst_compose) qed then obtain u where u: "s = occurs u" by moura hence "t = u \ \ \\<^sub>s \ \\<^sub>s \" using s(3) by fastforce thus ?B' using s u wellformed_transaction_strand_unlabel_memberD(1)[OF T_wf] by metis qed lemma transaction_decl_subst_proj: assumes "transaction_decl_subst \ T" shows "transaction_decl_subst \ (transaction_proj n T)" using assms transaction_proj_decl_eq[of n T] unfolding transaction_decl_subst_def by presburger lemma transaction_fresh_subst_proj: assumes "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" shows "transaction_fresh_subst \ (transaction_proj n T) (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A))" using assms transaction_proj_fresh_eq[of n T] contra_subsetD[OF subterms\<^sub>s\<^sub>e\<^sub>t_mono[OF transaction_proj_trms_subset[of n T]]] contra_subsetD[OF subterms\<^sub>s\<^sub>e\<^sub>t_mono[OF trms\<^sub>s\<^sub>s\<^sub>t_proj_subset(1)[of n A]]] unfolding transaction_fresh_subst_def by metis lemma transaction_renaming_subst_proj: assumes "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" shows "transaction_renaming_subst \ (map (transaction_proj n) P) (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A))" proof - let ?X = "\P A. \(vars_transaction ` set P) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" define Y where "Y \ ?X (map (transaction_proj n) P) (proj n A)" define Z where "Z \ ?X P A" have "Y \ Z" using sst_vars_proj_subset(3)[of n A] transaction_proj_vars_subset[of n] unfolding Y_def Z_def by fastforce hence "insert 0 (snd ` Y) \ insert 0 (snd ` Z)" by blast moreover have "finite (insert 0 (snd ` Z))" "finite (insert 0 (snd ` Y))" unfolding Y_def Z_def by auto ultimately have 0: "max_var_set Y \ max_var_set Z" using Max_mono by blast have "\n\max_var_set Z. \ = var_rename n" using assms unfolding transaction_renaming_subst_def Z_def by blast hence "\n\max_var_set Y. \ = var_rename n" using 0 le_trans by fast thus ?thesis unfolding transaction_renaming_subst_def Y_def by blast qed lemma transaction_decl_fresh_renaming_substs_wf_sst: fixes \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" assumes T: "wf'\<^sub>s\<^sub>s\<^sub>t (fst ` set (transaction_decl T ()) \ set (transaction_fresh T)) (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "wf'\<^sub>s\<^sub>s\<^sub>t {} (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)))" proof - have 0: "range_vars \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)) = {}" "range_vars \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = {}" "ground (\ ` (fst ` set (transaction_decl T ())))" "ground (\ ` set (transaction_fresh T))" "ground (\ ` {})" using transaction_decl_subst_domain[OF \] transaction_decl_subst_grounds_domain[OF \] transaction_decl_subst_range_vars_empty[OF \] transaction_fresh_subst_range_vars_empty[OF \] transaction_fresh_subst_domain[OF \] transaction_fresh_subst_grounds_domain[OF \] by (simp, simp, simp, simp, simp) have 1: "fv\<^sub>s\<^sub>e\<^sub>t (\ ` set (transaction_fresh T)) \ set (transaction_fresh T)" (is "?A \ ?B") proof fix x assume x: "x \ ?A" then obtain y where y: "y \ set (transaction_fresh T)" "x \ fv (\ y)" by auto hence "y \ subst_domain \" using transaction_decl_subst_domain[OF \] transaction_decl_subst_grounds_domain[OF \] by fast thus "x \ ?B" using x y by auto qed let ?X = "fst ` set (transaction_decl T ()) \ set (transaction_fresh T)" have "fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>s\<^sub>e\<^sub>t (\ ` ?X))) = {}" using 0(3-5) 1 by auto hence "wf'\<^sub>s\<^sub>s\<^sub>t {} (((unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)) \\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>s\<^sub>t \)" by (metis wf\<^sub>s\<^sub>s\<^sub>t_subst_apply[OF wf\<^sub>s\<^sub>s\<^sub>t_subst_apply[OF wf\<^sub>s\<^sub>s\<^sub>t_subst_apply[OF T]]]) thus ?thesis using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst labeled_stateful_strand_subst_comp[OF 0(1), of "\ \\<^sub>s \"] labeled_stateful_strand_subst_comp[OF 0(2), of \] subst_compose_assoc[of \ \ \] by metis qed lemma admissible_transaction_decl_fresh_renaming_subst_not_occurs: fixes \ \ \ defines "\ \ \ \\<^sub>s \ \\<^sub>s \" assumes T_adm: "admissible_transaction' T" and \\\: "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "\t. \ x = occurs t" and "\ x \ Fun OccursSec []" proof - note \_empty = admissible_transaction_decl_subst_empty[OF T_adm \\\(1)] note T_fresh_val = admissible_transactionE(2)[OF T_adm] show "\t. \ x = occurs t" for x using transaction_decl_fresh_renaming_substs_range'(1)[OF \\\] unfolding \_def[symmetric] by (cases "x \ subst_domain \") (force,force) show "\ x \ Fun OccursSec []" for x using transaction_decl_fresh_renaming_substs_range'(3)[ OF \\\ _ \_empty T_fresh_val, of "\ x"] unfolding \_def[symmetric] by (cases "x \ subst_domain \") auto qed subsection \Lemmata: Reachable Constraints\ lemma reachable_constraints_as_transaction_lists: fixes f defines "f \ \(T,\,\,\). dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" and "g \ concat \ map f" assumes A: "A \ reachable_constraints P" obtains Ts where "A = g Ts" and "\B. prefix B Ts \ g B \ reachable_constraints P" and "\B T \ \ \. prefix (B@[(T,\,\,\)]) Ts \ T \ set P \ transaction_decl_subst \ T \ transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B)) \ transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B))" proof - let ?P1 = "\A Ts. A = g Ts" let ?P2 = "\Ts. \B. prefix B Ts \ g B \ reachable_constraints P" let ?P3 = "\Ts. \B T \ \ \. prefix (B@[(T,\,\,\)]) Ts \ T \ set P \ transaction_decl_subst \ T \ transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B)) \ transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B))" have "\Ts. ?P1 A Ts \ ?P2 Ts \ ?P3 Ts" using A proof (induction A rule: reachable_constraints.induct) case init have "?P1 [] []" "?P2 []" "?P3 []" unfolding g_def f_def by simp_all thus ?case by blast next case (step A T \ \ \) let ?A' = "A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" obtain Ts where Ts: "?P1 A Ts" "?P2 Ts" "?P3 Ts" using step.IH by blast have 1: "?P1 ?A' (Ts@[(T,\,\,\)])" using Ts(1) unfolding g_def f_def by simp have 2: "?P2 (Ts@[(T,\,\,\)])" proof (intro allI impI) fix B assume "prefix B (Ts@[(T,\,\,\)])" hence "prefix B Ts \ B = Ts@[(T,\,\,\)]" by fastforce thus "g B \ reachable_constraints P " using Ts(1,2) reachable_constraints.step[OF step.hyps] unfolding g_def f_def by auto qed have 3: "?P3 (Ts@[(T,\,\,\)])" using Ts(1,3) step.hyps(2-5) unfolding g_def f_def by auto show ?case using 1 2 3 by blast qed thus thesis using that by blast qed lemma reachable_constraints_transaction_action_obtain: assumes A: "A \ reachable_constraints P" and a: "a \ set A" obtains T b B \ \ \ where "prefix (B@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) A" and "T \ set P" "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" and "b \ set (transaction_strand T)" "a = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \ \\<^sub>s \" "fst a = fst b" proof - define f where "f \ \(T,\,\::('fun,'atom,'sets,'lbl) prot_subst,\). dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define g where "g \ concat \ map f" obtain Ts where Ts: "A = g Ts" "\B. prefix B Ts \ g B \ reachable_constraints P" "\B T \ \ \. prefix (B@[(T,\,\,\)]) Ts \ T \ set P \ transaction_decl_subst \ T \ transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B)) \ transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B))" using reachable_constraints_as_transaction_lists[OF A] unfolding g_def f_def by blast obtain T \ \ \ where T: "(T,\,\,\) \ set Ts" "a \ set (f (T,\,\,\))" using Ts(1) a unfolding g_def by auto obtain B where B: "prefix (B@[(T,\,\,\)]) Ts" using T(1) by (meson prefix_snoc_in_iff) obtain b where b: "b \ set (transaction_strand T)" "a = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \ \\<^sub>s \" "fst a = fst b" using T(2) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_memberD'[of a "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \" thesis] unfolding f_def by simp have 0: "prefix (g B@f (T, \, \, \)) A" using concat_map_mono_prefix[OF B, of f] unfolding g_def Ts(1) by simp have 1: "T \ set P" "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B))" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B))" using B Ts(3) by (blast,blast,blast,blast) show thesis using 0[unfolded f_def] that[OF _ 1 b] by fast qed lemma reachable_constraints_unlabel_eq: defines "transaction_unlabel_eq \ \T1 T2. transaction_decl T1 = transaction_decl T2 \ transaction_fresh T1 = transaction_fresh T2 \ unlabel (transaction_receive T1) = unlabel (transaction_receive T2) \ unlabel (transaction_checks T1) = unlabel (transaction_checks T2) \ unlabel (transaction_updates T1) = unlabel (transaction_updates T2) \ unlabel (transaction_send T1) = unlabel (transaction_send T2)" assumes Peq: "list_all2 transaction_unlabel_eq P1 P2" shows "unlabel ` reachable_constraints P1 = unlabel ` reachable_constraints P2" (is "?A = ?B") proof (intro antisym subsetI) have "transaction_unlabel_eq T2 T1 = transaction_unlabel_eq T1 T2" for T1 T2 unfolding transaction_unlabel_eq_def by argo hence Peq': "list_all2 transaction_unlabel_eq P2 P1" using Peq list_all2_sym by metis have 0: "unlabel (transaction_strand T1) = unlabel (transaction_strand T2)" when "transaction_unlabel_eq T1 T2" for T1 T2 using that unfolding transaction_unlabel_eq_def transaction_strand_def by force have "vars_transaction T1 = vars_transaction T2" when "transaction_unlabel_eq T1 T2" for T1 T2 using 0[OF that] by simp hence "vars_transaction (P1 ! i) = vars_transaction (P2 ! i)" when "i < length P1" for i using that Peq list_all2_conv_all_nth by blast moreover have "length P1 = length P2" using Peq unfolding list_all2_iff by argo ultimately have 1: "\(vars_transaction ` set P1) = \(vars_transaction ` set P2)" using in_set_conv_nth[of _ P1] in_set_conv_nth[of _ P2] by fastforce have 2: "transaction_decl_subst \ T1 \ transaction_decl_subst \ T2" (is "?A1 \ ?A2") "transaction_fresh_subst \ T1 (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ transaction_fresh_subst \ T2 (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" (is "?B1 \ ?B2") "transaction_renaming_subst \ P1 (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ transaction_renaming_subst \ P2 (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" (is "?C1 \ ?C2") "transaction_renaming_subst \ P2 (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ transaction_renaming_subst \ P1 (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" (is "?D1 \ ?D2") when "transaction_unlabel_eq T1 T2" "unlabel \ = unlabel \" for T1 T2::"('fun,'atom,'sets,'lbl) prot_transaction" and \ \::"('fun,'atom,'sets,'lbl) prot_strand" and \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" proof - have *: "transaction_decl T1 = transaction_decl T2" "transaction_fresh T1 = transaction_fresh T2" "trms_transaction T1 = trms_transaction T2" using that unfolding transaction_unlabel_eq_def transaction_strand_def by force+ show "?A1 \ ?A2" using *(1) unfolding transaction_decl_subst_def by argo show "?B1 \ ?B2" using that(2) *(2,3) unfolding transaction_fresh_subst_def by force show "?C1 \ ?C2" using that(2) 1 unfolding transaction_renaming_subst_def by metis show "?D1 \ ?D2" using that(2) 1 unfolding transaction_renaming_subst_def by metis qed have 3: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T1 \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T2 \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" when "transaction_unlabel_eq T1 T2" for T1 T2 \ using 0[OF that] unlabel_subst[of _ \] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_cong by metis have "\\ \ reachable_constraints P2. unlabel \ = unlabel \" when "\ \ reachable_constraints P1" for \ using that proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) obtain \ where IH: "\ \ reachable_constraints P2" "unlabel \ = unlabel \" by (meson step.IH) obtain T' where T': "T' \ set P2" "transaction_unlabel_eq T T'" using list_all2_in_set_ex[OF Peq step.hyps(2)] by auto show ?case using 3[OF T'(2), of "\ \\<^sub>s \ \\<^sub>s \"] IH(2) reachable_constraints.step[OF IH(1) T'(1)] 2[OF T'(2) IH(2)] step.hyps(3-5) by (metis unlabel_append[of \] unlabel_append[of \]) qed (simp add: unlabel_def) thus "\ \ ?A \ \ \ ?B" for \ by fast have "\\ \ reachable_constraints P1. unlabel \ = unlabel \" when "\ \ reachable_constraints P2" for \ using that proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) obtain \ where IH: "\ \ reachable_constraints P1" "unlabel \ = unlabel \" by (meson step.IH) obtain T' where T': "T' \ set P1" "transaction_unlabel_eq T T'" using list_all2_in_set_ex[OF Peq' step.hyps(2)] by auto show ?case using 3[OF T'(2), of "\ \\<^sub>s \ \\<^sub>s \"] IH(2) reachable_constraints.step[OF IH(1) T'(1)] 2[OF T'(2) IH(2)] step.hyps(3-5) by (metis unlabel_append[of \] unlabel_append[of \]) qed (simp add: unlabel_def) thus "\ \ ?B \ \ \ ?A" for \ by fast qed lemma reachable_constraints_set_eq: assumes "set P1 = set P2" shows "reachable_constraints P1 = reachable_constraints P2" (is "?A = ?B") proof (intro antisym subsetI) note 0 = assms transaction_renaming_subst_set_eq[OF assms] note 1 = reachable_constraints.intros show "\ \ ?A \ \ \ ?B" for \ by (induct \ rule: reachable_constraints.induct) (auto simp add: 0 intro: 1) show "\ \ ?B \ \ \ ?A" for \ by (induct \ rule: reachable_constraints.induct) (auto simp add: 0 intro: 1) qed lemma reachable_constraints_set_subst: assumes "set P1 = set P2" and "Q (reachable_constraints P1)" shows "Q (reachable_constraints P2)" by (rule subst[of _ _ Q, OF reachable_constraints_set_eq[OF assms(1)] assms(2)]) lemma reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: assumes "\T \ set P. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms_transaction T)" and "\ \ reachable_constraints P" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using assms(2) proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms_transaction T)" using assms(1) step.hyps(2) by blast hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms_transaction T \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using transaction_decl_fresh_renaming_substs_range_wf_trms[OF step.hyps(3-5)] by (metis wf_trms_subst) hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" using wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_trms\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] by metis hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)))" using trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq by blast thus ?case using step.IH unlabel_append[of \] trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \"] by auto qed simp lemma reachable_constraints_var_types_in_transactions: fixes \::"('fun,'atom,'sets,'lbl) prot_constr" assumes \: "\ \ reachable_constraints P" and P: "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" shows "\\<^sub>v ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ (\T \ set P. \\<^sub>v ` fv_transaction T)" (is "?A \") and "\\<^sub>v ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ (\T \ set P. \\<^sub>v ` bvars_transaction T)" (is "?B \") and "\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ (\T \ set P. \\<^sub>v ` vars_transaction T)" (is "?C \") using \ proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" note 2 = transaction_decl_fresh_renaming_substs_wt[OF step.hyps(3-5)] have 3: "\t \ subst_range (\ \\<^sub>s \ \\<^sub>s \). fv t = {} \ (\x. t = Var x)" using transaction_decl_fresh_renaming_substs_range'(1)[OF step.hyps(3-5)] by fastforce have "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" unfolding T'_def by (metis fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq, metis bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq, metis vars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq) hence "\ ` Var ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ \ ` Var ` fv_transaction T" "\ ` Var ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = \ ` Var ` bvars_transaction T" "\ ` Var ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ \ ` Var ` vars_transaction T" using wt_subst_lsst_vars_type_subset[OF 2 3, of "transaction_strand T"] by argo+ hence "\\<^sub>v ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ \\<^sub>v ` fv_transaction T" "\\<^sub>v ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = \\<^sub>v ` bvars_transaction T" "\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ \\<^sub>v ` vars_transaction T" by (metis \\<^sub>v_Var_image)+ hence 4: "\\<^sub>v ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ (\T \ set P. \\<^sub>v ` fv_transaction T)" "\\<^sub>v ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ (\T \ set P. \\<^sub>v ` bvars_transaction T)" "\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ (\T \ set P. \\<^sub>v ` vars_transaction T)" using step.hyps(2) by fast+ have 5: "\\<^sub>v ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ @ T') = (\\<^sub>v ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ (\\<^sub>v ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" "\\<^sub>v ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ @ T') = (\\<^sub>v ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ (\\<^sub>v ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" "\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ @ T') = (\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ (\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" using unlabel_append[of \ T'] fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel T'"] bvars\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel T'"] vars\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel T'"] by auto { case 1 thus ?case using step.IH(1) 4(1) 5(1) unfolding T'_def by (simp del: subst_subst_compose fv\<^sub>s\<^sub>s\<^sub>t_def) } { case 2 thus ?case using step.IH(2) 4(2) 5(2) unfolding T'_def by (simp del: subst_subst_compose bvars\<^sub>s\<^sub>s\<^sub>t_def) } { case 3 thus ?case using step.IH(3) 4(3) 5(3) unfolding T'_def by (simp del: subst_subst_compose) } qed simp_all lemma reachable_constraints_no_bvars: assumes \: "\ \ reachable_constraints P" and P: "\T \ set P. bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) = {}" shows "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" using assms proof (induction) case init then show ?case unfolding unlabel_def by auto next case (step \ T \ \ \) then have "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" by metis moreover have "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) = {}" using step by (metis bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq) ultimately show ?case using bvars\<^sub>s\<^sub>s\<^sub>t_append unlabel_append by (metis sup_bot.left_neutral) qed lemma reachable_constraints_fv_bvars_disj: fixes \::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "\ \ reachable_constraints P" and P: "\S \ set P. admissible_transaction' S" shows "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" proof - let ?X = "\T \ set P. bvars_transaction T" note 0 = admissible_transactions_fv_bvars_disj[OF P] have 1: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ ?X" using \_reach proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) have "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) = bvars_transaction T" using bvars\<^sub>s\<^sub>s\<^sub>t_subst[of "unlabel (transaction_strand T)" "\ \\<^sub>s \ \\<^sub>s \"] bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \"] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] by argo hence "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \ ?X" using step.hyps(2) by blast thus ?case using step.IH bvars\<^sub>s\<^sub>s\<^sub>t_append by auto qed (simp add: unlabel_def) have 2: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ ?X = {}" using \_reach proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) have "x \ y" when x: "x \ ?X" and y: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" for x y proof - obtain y' where y': "y' \ fv_transaction T" "y \ fv ((\ \\<^sub>s \ \\<^sub>s \) y')" using y unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] by (metis fv\<^sub>s\<^sub>s\<^sub>t_subst_obtain_var) have "y \ \(vars_transaction ` set P)" using transaction_decl_fresh_renaming_substs_range''[OF step.hyps(3-5) y'(2)] transaction_renaming_subst_range_notin_vars[OF step.hyps(5), of y'] by auto thus ?thesis using x vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t by fast qed hence "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \) \ ?X = {}" by blast thus ?case using step.IH fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \"] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)"] unlabel_append[of \ "transaction_strand T"] by force qed (simp add: unlabel_def) show ?thesis using 0 1 2 by blast qed lemma reachable_constraints_vars_TAtom_typed: fixes \::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "\ \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and x: "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" shows "\\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" proof - have \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (metis reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s admissible_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P \_reach) have T_adm: "admissible_transaction' T" when "T \ set P" for T by (meson that Ball_set P) have "\T\set P. \x\set (transaction_fresh T). \\<^sub>v x = TAtom Value" using protocol_transaction_vars_TAtom_typed(3) P by blast hence *: "\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ (\T\set P. \\<^sub>v ` vars_transaction T)" using reachable_constraints_var_types_in_transactions[of \ P, OF \_reach] by auto have "\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ TAtom ` insert Value (range Atom)" proof - have "\\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" when "T \ set P" "x \ vars_transaction T" for T x using that protocol_transaction_vars_TAtom_typed(1)[of T] P admissible_transactionE(5) by blast hence "(\T\set P. \\<^sub>v ` vars_transaction T) \ TAtom ` insert Value (range Atom)" using P by blast thus "\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ TAtom ` insert Value (range Atom)" using * by auto qed thus ?thesis using x by auto qed lemma reachable_constraints_vars_not_attack_typed: fixes \::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "\ \ reachable_constraints P" and P: "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" "\T \ set P. \x \ vars_transaction T. \TAtom AttackType \ \\<^sub>v x" and x: "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" shows "\TAtom AttackType \ \\<^sub>v x" using reachable_constraints_var_types_in_transactions(3)[OF \_reach P(1)] P(2) x by fastforce lemma reachable_constraints_Value_vars_are_fv: assumes \_reach: "\ \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and x: "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" and "\\<^sub>v x = TAtom Value" shows "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" proof - have "\T\set P. bvars_transaction T = {}" using P admissible_transactionE(4) by metis hence \_no_bvars: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" using reachable_constraints_no_bvars[OF \_reach] by metis thus ?thesis using x vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] by blast qed lemma reachable_constraints_subterms_subst: assumes \_reach: "\ \ reachable_constraints P" and \: "welltyped_constraint_model \ \" and P: "\T \ set P. admissible_transaction' T" shows "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = (subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \\<^sub>s\<^sub>e\<^sub>t \" proof - have \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (metis reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s admissible_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P \_reach) from \ have \': "welltyped_constraint_model \ \" using welltyped_constraint_model_prefix by auto have 1: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \). (\f. \ x = Fun f []) \ (\y. \ x = Var y)" proof fix x assume xa: "x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" have "\f T. \ x = Fun f T" using \ interpretation_grounds[of \ "Var x"] unfolding welltyped_constraint_model_def constraint_model_def by (cases "\ x") auto then obtain f T where fT_p: "\ x = Fun f T" by auto hence "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" using \ unfolding welltyped_constraint_model_def constraint_model_def using wf_trm_subst_rangeD by metis moreover have "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using xa var_subterm_trms\<^sub>s\<^sub>s\<^sub>t_is_vars\<^sub>s\<^sub>s\<^sub>t[of x "unlabel \"] vars_iff_subtermeq[of x] by auto hence "\a. \\<^sub>v x = TAtom a" using reachable_constraints_vars_TAtom_typed[OF \_reach P] by blast hence "\a. \ (Var x) = TAtom a" by simp hence "\a. \ (Fun f T) = TAtom a" by (metis (no_types, opaque_lifting) \' welltyped_constraint_model_def fT_p wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) ultimately show "(\f. \ x = Fun f []) \ (\y. \ x = Var y)" using TAtom_term_cases fT_p by metis qed have "\T\set P. bvars_transaction T = {}" using assms admissible_transactionE(4) by metis then have "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" using reachable_constraints_no_bvars assms by metis then have 2: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ subst_domain \ = {}" by auto show ?thesis using subterms_subst_lsst[OF _ 2] 1 by simp qed lemma reachable_constraints_val_funs_private': fixes \::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "\ \ reachable_constraints P" and P: "\T \ set P. admissible_transaction_terms T" "\T \ set P. transaction_decl T () = []" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" and f: "f \ \(funs_term ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "\is_PubConstValue f" and "\is_Abs f" proof - have "\is_PubConstValue f \ \is_Abs f" using \_reach f proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) let ?T' = "unlabel (transaction_strand T) \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" let ?T'' = "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" note \_empty = admissible_transaction_decl_subst_empty'[OF bspec[OF P(2) step.hyps(2)] step.hyps(3)] have T: "admissible_transaction_terms T" using P(1) step.hyps(2) by metis have T_fresh: "\x \ set (transaction_fresh T). fst x = TAtom Value" when "T \ set P" for T using P that admissible_transactionE(14) unfolding list_all_iff \\<^sub>v_TAtom'' by fast show ?thesis using step proof (cases "f \ \(funs_term ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)") case False then obtain t where t: "t \ trms\<^sub>s\<^sub>s\<^sub>t ?T'" "f \ funs_term t" using step.prems trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of ?T''] trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?T'')"] unlabel_append[of \ "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?T''"] unlabel_subst[of "transaction_strand T"] by fastforce show ?thesis using trms\<^sub>s\<^sub>s\<^sub>t_funs_term_cases[OF t] proof assume "\u \ trms_transaction T. f \ funs_term u" thus ?thesis using conjunct1[OF conjunct2[OF T[unfolded admissible_transaction_terms_def]]] unfolding is_PubConstValue_def by blast next assume "\x \ fv_transaction T. f \ funs_term ((\ \\<^sub>s \ \\<^sub>s \) x)" then obtain x where "x \ fv_transaction T" "f \ funs_term ((\ \\<^sub>s \ \\<^sub>s \) x)" by moura thus ?thesis using transaction_decl_fresh_renaming_substs_range'(3)[ OF step.hyps(3-5) _ \_empty T_fresh[OF step.hyps(2), unfolded \\<^sub>v_TAtom''(2)]] unfolding is_PubConstValue_def by (metis (no_types, lifting) funs_term_Fun_subterm prot_fun.disc(30,48) subst_imgI subtermeq_Var_const(2) term.distinct(1) term.inject(2) term.set_cases(1)) qed qed simp qed simp thus "\is_PubConstValue f" "\is_Abs f" by simp_all qed lemma reachable_constraints_val_funs_private: fixes \::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "\ \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and f: "f \ \(funs_term ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "\is_PubConstValue f" and "\is_Abs f" using P reachable_constraints_val_funs_private'[OF \_reach _ _ _ f] admissible_transaction_is_wellformed_transaction(4) admissible_transactionE(1,14) unfolding list_all_iff \\<^sub>v_TAtom'' by (blast,fast) lemma reachable_constraints_occurs_fact_ik_case: fixes A::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and occ: "occurs t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" shows "\n. t = Fun (Val n) []" using \_reach occ proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" have T_adm: "admissible_transaction' T" using P step.hyps(2) by blast hence T: "wellformed_transaction T" "admissible_transaction_occurs_checks T" using admissible_transaction_is_wellformed_transaction(1) P_occ step.hyps(2) by (blast,blast) have T_fresh: "\x \ set (transaction_fresh T). fst x = TAtom Value" using admissible_transactionE(14)[OF T_adm] unfolding list_all_iff by fast note \_empty = admissible_transaction_decl_subst_empty[OF T_adm step.hyps(3)] have \_dom_empty: "z \ fst ` set (transaction_decl T ())" for z using transaction_decl_subst_empty_inv[OF step.hyps(3)[unfolded \_empty]] by simp show ?case proof (cases "occurs t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A") case False hence "occurs t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using step.prems unfolding \_def by simp hence "\ts. occurs t \ set ts \ receive\ts\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" unfolding ik\<^sub>s\<^sub>s\<^sub>t_def by force hence "\ts. occurs t \ set ts \ send\ts\ \ set (unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(1) by blast then obtain ts s where s: "s \ set ts" "send\ts\ \ set (unlabel (transaction_strand T))" "s \ \ = occurs t" using stateful_strand_step_mem_substD(1)[of _ "unlabel (transaction_strand T)" \] unlabel_subst[of "transaction_strand T" \] by force note 0 = transaction_decl_fresh_renaming_substs_range[OF step.hyps(3-5)] have 1: "send\ts\ \ set (unlabel (transaction_send T))" using s(2) wellformed_transaction_strand_unlabel_memberD(8)[OF T(1)] by blast have 2: "is_Send (send\ts\)" unfolding is_Send_def by simp have 3: "\u. s = occurs u" proof - { fix z have "(\n. \ z = Fun (Val n) []) \ (\y. \ z = Var y)" using 0(3,4) T_fresh \_dom_empty unfolding \_def by blast hence "\u. \ z = occurs u" "\ z \ Fun OccursSec []" by auto } note * = this obtain u u' where T: "s = Fun OccursFact [u,u']" using *(1) s(3) by (cases s) auto thus ?thesis using *(2) s(3) by (cases u) auto qed obtain x where x: "x \ set (transaction_fresh T)" "s = occurs (Var x)" using 3 s(1) admissible_transaction_occurs_checksE4[OF T(2) 1] by metis have "t = \ x" using s(3) x(2) by auto thus ?thesis using 0(3)[OF \_dom_empty x(1)] x(1) T_fresh unfolding \_def by fast qed (simp add: step.IH) qed simp lemma reachable_constraints_occurs_fact_send_ex: fixes A::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and x: "\\<^sub>v x = TAtom Value" "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" shows "\ts. occurs (Var x) \ set ts \ send\ts\ \ set (unlabel A)" using \_reach x(2) proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) note \_empty = admissible_transaction_decl_subst_empty[OF bspec[OF P step.hyps(2)] step.hyps(3)] note T = bspec[OF P_occ step.hyps(2)] show ?case proof (cases "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A") case True show ?thesis using step.IH[OF True] unlabel_append[of A] by auto next case False then obtain y where y: "y \ fv_transaction T - set (transaction_fresh T)" "(\ \\<^sub>s \ \\<^sub>s \) y = Var x" using transaction_decl_fresh_renaming_substs_fv[OF step.hyps(3-5), of x] step.prems(1) fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel A"] unlabel_append[of A] by auto have "\ y = Var y" using y(1) step.hyps(4) unfolding transaction_fresh_subst_def by auto hence "\ y = Var x" using y(2) unfolding subst_compose_def \_empty by simp hence y_val: "fst y = TAtom Value" "\\<^sub>v y = TAtom Value" using x(1) \\<^sub>v_TAtom''[of x] \\<^sub>v_TAtom''[of y] wt_subst_trm''[OF transaction_renaming_subst_wt[OF step.hyps(5)], of "Var y"] by force+ obtain ts where ts: "occurs (Var y) \ set ts" "receive\ts\ \ set (unlabel (transaction_receive T))" using admissible_transaction_occurs_checksE1[OF T y(1) y_val(2)] by (metis list.set_intros(1) unlabel_Cons(1)) hence "receive\ts\ \ set (unlabel (transaction_strand T))" using transaction_strand_subsets(5) by blast hence *: "receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \\ \ set (unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" using unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] stateful_strand_step_subst_inI(2)[of _ _ "\ \\<^sub>s \ \\<^sub>s \"] by force have "occurs (Var y) \ \ \\<^sub>s \ \\<^sub>s \ = occurs (Var x)" using y(2) by (auto simp del: subst_subst_compose) hence **: "occurs (Var x) \ set ts \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" using ts(1) by force have "send\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)))" using * dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(2) by blast thus ?thesis using ** unlabel_append[of A] by force qed qed simp lemma reachable_constraints_db\<^sub>l\<^sub>s\<^sub>s\<^sub>t_set_args_empty: assumes \: "\ \ reachable_constraints P" and PP: "list_all wellformed_transaction P" and admissible_transaction_updates: "let f = (\T. \x \ set (unlabel (transaction_updates T)). is_Update x \ is_Var (the_elem_term x) \ is_Fun_Set (the_set_term x) \ fst (the_Var (the_elem_term x)) = TAtom Value) in list_all f P" and d: "(t, s) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" shows "\ss. s = Fun (Set ss) []" using \ d proof (induction) case (step \ TT \ \ \) let ?TT = "transaction_strand TT \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" let ?TTu = "unlabel ?TT" let ?TTd = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?TT" let ?TTdu = "unlabel ?TTd" from step(6) have "(t, s) \ set (db'\<^sub>s\<^sub>s\<^sub>t ?TTdu \ (db'\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ []))" by (metis db\<^sub>s\<^sub>s\<^sub>t_append db\<^sub>s\<^sub>s\<^sub>t_def step.prems unlabel_append) hence "(t, s) \ set (db'\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ []) \ (\t' s'. insert\t',s'\ \ set ?TTdu \ t = t' \ \ \ s = s' \ \)" using db\<^sub>s\<^sub>s\<^sub>t_in_cases[of t "s" ?TTdu \] by metis thus ?case proof assume "\t' s'. insert\t',s'\ \ set ?TTdu \ t = t' \ \ \ s = s' \ \" then obtain t' s' where t's'_p: "insert\t',s'\ \ set ?TTdu" "t = t' \ \" "s = s' \ \" by metis then obtain lll where "(lll, insert\t',s'\) \ set ?TTd" by (meson unlabel_mem_has_label) hence "(lll, insert\t',s'\) \ set (transaction_strand TT \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(4) by blast hence "insert\t',s'\ \ set ?TTu" by (meson unlabel_in) hence "insert\t',s'\ \ set ((unlabel (transaction_strand TT)) \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" by (simp add: subst_lsst_unlabel) hence "insert\t',s'\ \ (\x. x \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \ \\<^sub>s \) ` set (unlabel (transaction_strand TT))" unfolding subst_apply_stateful_strand_def by auto then obtain u where "u \ set (unlabel (transaction_strand TT)) \ u \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \ \\<^sub>s \ = insert\t',s'\" by auto hence "\t'' s''. insert\t'',s''\ \ set (unlabel (transaction_strand TT)) \ t' = t'' \ \ \\<^sub>s \ \\<^sub>s \ \ s' = s'' \ \ \\<^sub>s \ \\<^sub>s \" by (cases u) auto then obtain t'' s'' where t''s''_p: "insert\t'',s''\ \ set (unlabel (transaction_strand TT)) \ t' = t'' \ \ \\<^sub>s \ \\<^sub>s \ \ s' = s'' \ \ \\<^sub>s \ \\<^sub>s \" by auto hence "insert\t'',s''\ \ set (unlabel (transaction_updates TT))" using is_Update_in_transaction_updates[of "insert\t'',s''\" TT] using PP step(2) unfolding list_all_iff by auto moreover have "\x\set (unlabel (transaction_updates TT)). is_Fun_Set (the_set_term x)" using step(2) admissible_transaction_updates unfolding is_Fun_Set_def list_all_iff by auto ultimately have "is_Fun_Set (the_set_term (insert\t'',s''\))" by auto moreover have "s' = s'' \ \ \\<^sub>s \ \\<^sub>s \" using t''s''_p by blast ultimately have "is_Fun_Set (the_set_term (insert\t',s'\))" by (auto simp add: is_Fun_Set_subst) hence "is_Fun_Set s" by (simp add: t's'_p(3) is_Fun_Set_subst) thus ?case using is_Fun_Set_exi by auto qed (auto simp add: step db\<^sub>s\<^sub>s\<^sub>t_def) qed auto lemma reachable_constraints_occurs_fact_ik_ground: fixes A::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and t: "occurs t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" shows "fv (occurs t) = {}" proof - have 0: "admissible_transaction' T" when "T \ set P" for T using P that unfolding list_all_iff by simp note 1 = admissible_transaction_is_wellformed_transaction(1)[OF 0] bspec[OF P_occ] have 2: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \)" when "T \ set P" for T \ and A::"('fun,'atom,'sets,'lbl) prot_constr" using dual_transaction_ik_is_transaction_send'[OF 1(1)[OF that]] by fastforce show ?thesis using \_reach t proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) note \_empty = admissible_transaction_decl_subst_empty[OF 0[OF step.hyps(2)] step.hyps(3)] from step show ?case proof (cases "occurs t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A") case False hence "occurs t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" using 2[OF step.hyps(2)] step.prems \_empty by blast then obtain ts where ts: "occurs t \ set ts" "send\ts\ \ set (unlabel (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" using wellformed_transaction_send_receive_subst_trm_cases(2)[OF 1(1)[OF step.hyps(2)]] by blast then obtain ts' s where s: "occurs s \ set ts'" "send\ts'\ \ set (unlabel (transaction_send T))" "t = s \ \ \\<^sub>s \ \\<^sub>s \" using transaction_decl_fresh_renaming_substs_occurs_fact_send_receive(1)[ OF step.hyps(3-5) 0[OF step.hyps(2)] ts(1)] transaction_strand_subst_subsets(8)[of T "\ \\<^sub>s \ \\<^sub>s \"] by blast obtain x where x: "x \ set (transaction_fresh T)" "s = Var x" using admissible_transaction_occurs_checksE4[OF 1(2)[OF step.hyps(2)] s(2,1)] by metis have "fv t = {}" using transaction_decl_fresh_renaming_substs_range(2)[OF step.hyps(3-5) _ x(1)] s(3) x(2) transaction_decl_subst_empty_inv[OF step.hyps(3)[unfolded \_empty]] by (auto simp del: subst_subst_compose) thus ?thesis by simp qed simp qed simp qed lemma reachable_constraints_occurs_fact_ik_funs_terms: fixes A::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "A \ reachable_constraints P" and \: "welltyped_constraint_model I A" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" shows "\s \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I). OccursFact \ \(funs_term ` set (snd (Ana s)))" (is "?A A") and "\s \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I). OccursSec \ \(funs_term ` set (snd (Ana s)))" (is "?B A") and "Fun OccursSec [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" (is "?C A") and "\x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. I x \ Fun OccursSec []" (is "?D A") proof - have T_adm: "admissible_transaction' T" when "T \ set P" for T using P that unfolding list_all_iff by simp note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] note T_occ = bspec[OF P_occ] note \_empty = admissible_transaction_decl_subst_empty[OF T_adm] have \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" by (metis \ welltyped_constraint_model_def) have \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" by (metis \ welltyped_constraint_model_def constraint_model_def) have \_grounds: "fv (I x) = {}" "\f T. I x = Fun f T" for x using \ interpretation_grounds[of I, of "Var x"] empty_fv_exists_fun[of "I x"] unfolding welltyped_constraint_model_def constraint_model_def by auto have 00: "fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \ vars_transaction T" "fv\<^sub>s\<^sub>e\<^sub>t (subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))) = fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" for T::"('fun,'atom,'sets,'lbl) prot_transaction" using fv_trms\<^sub>s\<^sub>s\<^sub>t_subset(1)[of "unlabel (transaction_send T)"] vars_transaction_unfold fv_subterms_set[of "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)"] by blast+ have 0: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). \a. \ (Var x) = TAtom a" "\x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). \ (Var x) \ TAtom OccursSecType" "\x \ fv\<^sub>s\<^sub>e\<^sub>t (subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))). \a. \ (Var x) = TAtom a" "\x \ fv\<^sub>s\<^sub>e\<^sub>t (subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))). \ (Var x) \ TAtom OccursSecType" "\x \ vars_transaction T. \a. \ (Var x) = TAtom a" "\x \ vars_transaction T. \ (Var x) \ TAtom OccursSecType" when "T \ set P" for T using admissible_transaction_occurs_fv_types[OF T_adm[OF that]] 00 by blast+ note T_fresh_type = admissible_transactionE(2)[OF T_adm] have 1: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \\<^sub>s\<^sub>e\<^sub>t I = (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I) \ (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t I)" when "T \ set P" for T \ and A::"('fun,'atom,'sets,'lbl) prot_constr" using dual_transaction_ik_is_transaction_send'[OF T_wf[OF that]] by fastforce have 2: "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t I) = subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t I" when "T \ set P" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" for T \ using wt_subst_TAtom_subterms_set_subst[OF wt_subst_compose[OF \(1) \_wt] 0(1)[OF that(1)]] wf_trm_subst_rangeD[OF wf_trms_subst_compose[OF \(2) \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s]] by auto have 3: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \ \\<^sub>s \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \ \\<^sub>s \))" when "T \ set P" "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" for \ \ \ and T::"('fun,'atom,'sets,'lbl) prot_transaction" and A::"('fun,'atom,'sets,'lbl) prot_constr" using protocol_transaction_vars_TAtom_typed(3)[of T] P that(1) transaction_decl_fresh_renaming_substs_wt[OF that(2-4)] transaction_decl_fresh_renaming_substs_range_wf_trms[OF that(2-4)] wf_trms_subst_compose by simp_all have 4: "\s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). OccursFact \ \(funs_term ` set (snd (Ana s))) \ OccursSec \ \(funs_term ` set (snd (Ana s)))" when T: "T \ set P" for T proof fix t assume t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" then obtain ts s where s: "send\ts\ \ set (unlabel (transaction_send T))" "s \ set ts" "t \ subterms s" using wellformed_transaction_unlabel_cases(4)[OF T_wf[OF T]] by fastforce have s_occ: "\x. s = occurs (Var x)" when "OccursFact \ funs_term t \ OccursSec \ funs_term t" using that s(1) subtermeq_imp_funs_term_subset[OF s(3)] admissible_transaction_occurs_checksE3[OF T_occ[OF T] _ s(2)] by blast obtain K T' where K: "Ana t = (K,T')" by moura show "OccursFact \ \(funs_term ` set (snd (Ana t))) \ OccursSec \ \(funs_term ` set (snd (Ana t)))" proof (rule ccontr) assume "\(OccursFact \ \(funs_term ` set (snd (Ana t))) \ OccursSec \ \(funs_term ` set (snd (Ana t))))" hence a: "OccursFact \ \(funs_term ` set (snd (Ana t))) \ OccursSec \ \(funs_term ` set (snd (Ana t)))" by simp hence "OccursFact \ \(funs_term ` set T') \ OccursSec \ \(funs_term ` set T')" using K by simp hence "OccursFact \ funs_term t \ OccursSec \ funs_term t" using Ana_subterm[OF K] funs_term_subterms_eq(1)[of t] by blast then obtain x where x: "t \ subterms (occurs (Var x))" using s(3) s_occ by blast thus False using a by fastforce qed qed have 5: "OccursFact \ \(funs_term ` subst_range (\ \\<^sub>s \ \\<^sub>s \))" "OccursSec \ \(funs_term ` subst_range (\ \\<^sub>s \ \\<^sub>s \))" when \\\: "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" and T: "T \ set P" for \ \ \ and T::"('fun,'atom,'sets,'lbl) prot_transaction" and A::"('fun,'atom,'sets,'lbl) prot_constr" proof - have "OccursFact \ funs_term t" "OccursSec \ funs_term t" when "t \ subst_range (\ \\<^sub>s \ \\<^sub>s \)" for t using transaction_decl_fresh_renaming_substs_range'(3)[ OF \\\ that \_empty[OF T \\\(1)] T_fresh_type[OF T]] by auto thus "OccursFact \ \(funs_term ` subst_range (\ \\<^sub>s \ \\<^sub>s \))" "OccursSec \ \(funs_term ` subst_range (\ \\<^sub>s \ \\<^sub>s \))" by blast+ qed have 6: "I x \ Fun OccursSec []" "\t. I x = occurs t" "\a. \ (I x) = TAtom a \ a \ OccursSecType" when T: "T \ set P" and \\\: "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" and x: "Var x \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" for x \ \ \ and T::"('fun,'atom,'sets,'lbl) prot_transaction" and A::"('fun,'atom,'sets,'lbl) prot_constr" proof - obtain t where t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "t \ (\ \\<^sub>s \ \\<^sub>s \) = Var x" using x by moura then obtain y where y: "t = Var y" by (cases t) auto have "\a. \ t = TAtom a \ a \ OccursSecType" using 0(1,2)[OF T] t(1) y by force thus "\a. \ (I x) = TAtom a \ a \ OccursSecType" using wt_subst_trm''[OF 3(1)[OF T \\\]] wt_subst_trm''[OF \_wt] t(2) - by (metis subst_apply_term.simps(1)) + by (metis eval_term.simps(1)) thus "I x \ Fun OccursSec []" "\t. I x = occurs t" by auto qed have 7: "I x \ Fun OccursSec []" "\t. I x = occurs t" "\a. \ (I x) = TAtom a \ a \ OccursSecType" when T: "T \ set P" and \\\: "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" and x: "x \ fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \ \\<^sub>s \) ` vars_transaction T)" for x \ \ \ and T::"('fun,'atom,'sets,'lbl) prot_transaction" and A::"('fun,'atom,'sets,'lbl) prot_constr" proof - obtain y where y: "y \ vars_transaction T" "x \ fv ((\ \\<^sub>s \ \\<^sub>s \) y)" using x by auto hence y': "(\ \\<^sub>s \ \\<^sub>s \) y = Var x" using transaction_decl_fresh_renaming_substs_range'(3)[ OF \\\ _ \_empty[OF T \\\(1)] T_fresh_type[OF T]] by (cases "(\ \\<^sub>s \ \\<^sub>s \) y \ subst_range (\ \\<^sub>s \ \\<^sub>s \)") force+ have "\a. \ (Var y) = TAtom a \ a \ OccursSecType" using 0(5,6)[OF T] y by force thus "\a. \ (I x) = TAtom a \ a \ OccursSecType" using wt_subst_trm''[OF 3(1)[OF T \\\]] wt_subst_trm''[OF \_wt] y' - by (metis subst_apply_term.simps(1)) + by (metis eval_term.simps(1)) thus "I x \ Fun OccursSec []" "\t. I x = occurs t" by auto qed have 8: "I x \ Fun OccursSec []" "\t. I x = occurs t" "\a. \ (I x) = TAtom a \ a \ OccursSecType" when T: "T \ set P" and \\\: "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" and x: "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" for x \ \ \ and T::"('fun,'atom,'sets,'lbl) prot_transaction" and A::"('fun,'atom,'sets,'lbl) prot_constr" proof - obtain t where t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" "t \ (\ \\<^sub>s \ \\<^sub>s \) = Var x" using x by moura then obtain y where y: "t = Var y" by (cases t) auto have "\a. \ t = TAtom a \ a \ OccursSecType" using 0(3,4)[OF T] t(1) y by force thus "\a. \ (I x) = TAtom a \ a \ OccursSecType" using wt_subst_trm''[OF 3(1)[OF T \\\]] wt_subst_trm''[OF \_wt] t(2) - by (metis subst_apply_term.simps(1)) + by (metis eval_term.simps(1)) thus "I x \ Fun OccursSec []" "\t. I x = occurs t" by auto qed have s_fv: "fv s \ fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \ \\<^sub>s \) ` vars_transaction T)" when s: "s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" and T: "T \ set P" for s and \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" and T::"('fun,'atom,'sets,'lbl) prot_transaction" proof fix x assume "x \ fv s" hence "x \ fv\<^sub>s\<^sub>e\<^sub>t (subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using s by auto hence *: "x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using fv_subterms_set_subst' by fast have **: "list_all is_Send (unlabel (transaction_send T))" using T_wf[OF T] unfolding wellformed_transaction_def by blast have "x \ fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \ \\<^sub>s \) ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" proof - obtain t where t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "x \ fv (t \ \ \\<^sub>s \ \\<^sub>s \)" using * by fastforce hence "fv t \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" using fv_trms\<^sub>s\<^sub>s\<^sub>t_subset(1)[of "unlabel (transaction_send T)"] by auto thus ?thesis using t(2) subst_apply_fv_subset by fast qed thus "x \ fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \ \\<^sub>s \) ` vars_transaction T)" using vars_transaction_unfold[of T] by fastforce qed show "?A A" using \_reach proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) have *: "\s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). OccursFact \ \(funs_term ` set (snd (Ana s)))" using 4[OF step.hyps(2)] by blast have "\s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \ \\<^sub>s\<^sub>e\<^sub>t I. OccursFact \ \(funs_term ` set (snd (Ana s)))" proof fix t assume t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \ \\<^sub>s\<^sub>e\<^sub>t I" then obtain s u where su: "s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" "s \ I = t" "u \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" "u \ \ \\<^sub>s \ \\<^sub>s \ = s" by force obtain Ku Tu where KTu: "Ana u = (Ku,Tu)" by moura have *: "OccursFact \ \(funs_term ` set Tu)" "OccursFact \ \(funs_term ` subst_range (\ \\<^sub>s \ \\<^sub>s \))" "OccursFact \ \(funs_term ` \(((set \ snd \ Ana) ` subst_range (\ \\<^sub>s \ \\<^sub>s \))))" using transaction_decl_fresh_renaming_substs_range'(3)[ OF step.hyps(3-5) _ \_empty[OF step.hyps(2,3)] T_fresh_type[OF step.hyps(2)]] 4[OF step.hyps(2)] su(3) KTu by (fastforce,fastforce,fastforce) have "OccursFact \ \(funs_term ` set (Tu \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" proof - { fix f assume f: "f \ \(funs_term ` set (Tu \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" then obtain tf where tf: "tf \ set Tu" "f \ funs_term (tf \ \ \\<^sub>s \ \\<^sub>s \)" by moura hence "f \ funs_term tf \ f \ \(funs_term ` subst_range (\ \\<^sub>s \ \\<^sub>s \))" using funs_term_subst[of tf "\ \\<^sub>s \ \\<^sub>s \"] by force hence "f \ OccursFact" using *(1,2) tf(1) by blast } thus ?thesis by metis qed hence **: "OccursFact \ \(funs_term ` set (snd (Ana s)))" proof (cases u) case (Var xu) - hence "s = (\ \\<^sub>s \ \\<^sub>s \) xu" using su(4) by (metis subst_apply_term.simps(1)) + hence "s = (\ \\<^sub>s \ \\<^sub>s \) xu" using su(4) by (metis eval_term.simps(1)) thus ?thesis using *(3) by fastforce qed (use su(4) KTu Ana_subst'[of _ _ Ku Tu "\ \\<^sub>s \ \\<^sub>s \"] in simp) show "OccursFact \ \(funs_term ` set (snd (Ana t)))" proof (cases s) case (Var sx) then obtain a where a: "\ (I sx) = Var a" using su(1) 8(3)[OF step.hyps(2-5), of sx] by fast hence "Ana (I sx) = ([],[])" by (metis \_grounds(2) const_type_inv[THEN Ana_const]) thus ?thesis using Var su(2) by simp next case (Fun f S) hence snd_Ana_t: "snd (Ana t) = snd (Ana s) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t I" using su(2) Ana_subst'[of f S _ "snd (Ana s)" I] by (cases "Ana s") simp_all { fix g assume "g \ \(funs_term ` set (snd (Ana t)))" hence "g \ \(funs_term ` set (snd (Ana s))) \ (\x \ fv\<^sub>s\<^sub>e\<^sub>t (set (snd (Ana s))). g \ funs_term (I x))" using snd_Ana_t funs_term_subst[of _ I] by auto hence "g \ OccursFact" proof assume "\x \ fv\<^sub>s\<^sub>e\<^sub>t (set (snd (Ana s))). g \ funs_term (I x)" then obtain x where x: "x \ fv\<^sub>s\<^sub>e\<^sub>t (set (snd (Ana s)))" "g \ funs_term (I x)" by moura have "x \ fv s" using x(1) Ana_vars(2)[of s] by (cases "Ana s") auto hence "x \ fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \ \\<^sub>s \) ` vars_transaction T)" using s_fv[OF su(1) step.hyps(2)] by blast then obtain a h U where h: "I x = Fun h U" "\ (I x) = Var a" "a \ OccursSecType" "arity h = 0" using \_grounds(2) 7(3)[OF step.hyps(2-5)] const_type_inv by metis hence "h \ OccursFact" by auto moreover have "U = []" using h(1,2,4) const_type_inv_wf[of h U a] \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s by fastforce ultimately show ?thesis using h(1) x(2) by auto qed (use ** in blast) } thus ?thesis by blast qed qed thus ?case using step.IH step.prems 1[OF step.hyps(2), of A "\ \\<^sub>s \ \\<^sub>s \"] 2[OF step.hyps(2) 3[OF step.hyps(2-5)]] by auto qed simp show "?B A" using \_reach proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) have "\s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \ \\<^sub>s\<^sub>e\<^sub>t I. OccursSec \ \(funs_term ` set (snd (Ana s)))" proof fix t assume t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \ \\<^sub>s\<^sub>e\<^sub>t I" then obtain s u where su: "s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" "s \ I = t" "u \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" "u \ \ \\<^sub>s \ \\<^sub>s \ = s" by force obtain Ku Tu where KTu: "Ana u = (Ku,Tu)" by moura have *: "OccursSec \ \(funs_term ` set Tu)" "OccursSec \ \(funs_term ` subst_range (\ \\<^sub>s \ \\<^sub>s \))" "OccursSec \ \(funs_term ` \(((set \ snd \ Ana) ` subst_range (\ \\<^sub>s \ \\<^sub>s \))))" using transaction_decl_fresh_renaming_substs_range'(3)[ OF step.hyps(3-5) _ \_empty[OF step.hyps(2,3)] T_fresh_type[OF step.hyps(2)]] 4[OF step.hyps(2)] su(3) KTu by (fastforce,fastforce,fastforce) have "OccursSec \ \(funs_term ` set (Tu \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" proof - { fix f assume f: "f \ \(funs_term ` set (Tu \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" then obtain tf where tf: "tf \ set Tu" "f \ funs_term (tf \ \ \\<^sub>s \ \\<^sub>s \)" by moura hence "f \ funs_term tf \ f \ \(funs_term ` subst_range (\ \\<^sub>s \ \\<^sub>s \))" using funs_term_subst[of tf "\ \\<^sub>s \ \\<^sub>s \"] by force hence "f \ OccursSec" using *(1,2) tf(1) by blast } thus ?thesis by metis qed hence **: "OccursSec \ \(funs_term ` set (snd (Ana s)))" proof (cases u) case (Var xu) - hence "s = (\ \\<^sub>s \ \\<^sub>s \) xu" using su(4) by (metis subst_apply_term.simps(1)) + hence "s = (\ \\<^sub>s \ \\<^sub>s \) xu" using su(4) by (metis eval_term.simps(1)) thus ?thesis using *(3) by fastforce qed (use su(4) KTu Ana_subst'[of _ _ Ku Tu "\ \\<^sub>s \ \\<^sub>s \"] in simp) show "OccursSec \ \(funs_term ` set (snd (Ana t)))" proof (cases s) case (Var sx) then obtain a where a: "\ (I sx) = Var a" using su(1) 8(3)[OF step.hyps(2-5), of sx] by fast hence "Ana (I sx) = ([],[])" by (metis \_grounds(2) const_type_inv[THEN Ana_const]) thus ?thesis using Var su(2) by simp next case (Fun f S) hence snd_Ana_t: "snd (Ana t) = snd (Ana s) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t I" using su(2) Ana_subst'[of f S _ "snd (Ana s)" I] by (cases "Ana s") simp_all { fix g assume "g \ \(funs_term ` set (snd (Ana t)))" hence "g \ \(funs_term ` set (snd (Ana s))) \ (\x \ fv\<^sub>s\<^sub>e\<^sub>t (set (snd (Ana s))). g \ funs_term (I x))" using snd_Ana_t funs_term_subst[of _ I] by auto hence "g \ OccursSec" proof assume "\x \ fv\<^sub>s\<^sub>e\<^sub>t (set (snd (Ana s))). g \ funs_term (I x)" then obtain x where x: "x \ fv\<^sub>s\<^sub>e\<^sub>t (set (snd (Ana s)))" "g \ funs_term (I x)" by moura have "x \ fv s" using x(1) Ana_vars(2)[of s] by (cases "Ana s") auto hence "x \ fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \ \\<^sub>s \) ` vars_transaction T)" using s_fv[OF su(1) step.hyps(2)] by blast then obtain a h U where h: "I x = Fun h U" "\ (I x) = Var a" "a \ OccursSecType" "arity h = 0" using \_grounds(2) 7(3)[OF step.hyps(2-5)] const_type_inv by metis hence "h \ OccursSec" by auto moreover have "U = []" using h(1,2,4) const_type_inv_wf[of h U a] \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s by fastforce ultimately show ?thesis using h(1) x(2) by auto qed (use ** in blast) } thus ?thesis by blast qed qed thus ?case using step.IH step.prems 1[OF step.hyps(2), of A "\ \\<^sub>s \ \\<^sub>s \"] 2[OF step.hyps(2) 3[OF step.hyps(2-5)]] by auto qed simp show "?C A" using \_reach proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) have *: "Fun OccursSec [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" using admissible_transaction_occurs_checksE5[OF T_occ[OF step.hyps(2)]] by blast have **: "Fun OccursSec [] \ subst_range (\ \\<^sub>s \ \\<^sub>s \)" using transaction_decl_fresh_renaming_substs_range'(3)[ OF step.hyps(3-5) _ \_empty[OF step.hyps(2,3)] T_fresh_type[OF step.hyps(2)]] by auto have "Fun OccursSec [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \ \\<^sub>s\<^sub>e\<^sub>t I" proof assume "Fun OccursSec [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \ \\<^sub>s\<^sub>e\<^sub>t I" then obtain s where "s \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" "s \ I = Fun OccursSec []" by moura moreover have "Fun OccursSec [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" proof assume "Fun OccursSec [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" then obtain u where "u \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "u \ \ \\<^sub>s \ \\<^sub>s \ = Fun OccursSec []" by moura thus False using * ** by (cases u) (force simp del: subst_subst_compose)+ qed ultimately show False using 6[OF step.hyps(2-5)] by (cases s) auto qed thus ?case using step.IH step.prems 1[OF step.hyps(2), of A "\ \\<^sub>s \ \\<^sub>s \"] by fast qed simp show "?D A" using \_reach proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) { fix x assume x: "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" hence x': "x \ vars\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T) \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" by (metis vars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unlabel_subst) hence "x \ vars_transaction T \ x \ fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \ \\<^sub>s \) ` vars_transaction T)" using vars\<^sub>s\<^sub>s\<^sub>t_subst_cases[OF x'] by metis moreover have "I x \ Fun OccursSec []" when "x \ vars_transaction T" using that 0(5,6)[OF step.hyps(2)] wt_subst_trm''[OF \_wt, of "Var x"] by fastforce ultimately have "I x \ Fun OccursSec []" using 7(1)[OF step.hyps(2-5), of x] by blast } thus ?case using step.IH by auto qed simp qed lemma reachable_constraints_occurs_fact_ik_subst_aux: assumes \_reach: "A \ reachable_constraints P" and \: "welltyped_constraint_model I A" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and t: "t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "t \ I = occurs s" shows "\u. t = occurs u" proof - have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" using \ unfolding welltyped_constraint_model_def constraint_model_def by metis hence 0: "\ t = \ (occurs s)" using t(2) wt_subst_trm'' by metis have 1: "\\<^sub>v ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ (\T \ set P. \\<^sub>v ` fv_transaction T)" "\T \ set P. \x \ fv_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" using reachable_constraints_var_types_in_transactions(1)[OF \_reach] protocol_transaction_vars_TAtom_typed(2,3) P by fast+ show ?thesis proof (cases t) case (Var x) thus ?thesis using 0 1 t(1) var_subterm_ik\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of x "unlabel A"] by fastforce next case (Fun f T) hence 2: "f = OccursFact" "length T = Suc (Suc 0)" "T ! 0 \ I = Fun OccursSec []" using t(2) by auto have "T ! 0 = Fun OccursSec []" proof (cases "T ! 0") case (Var y) hence "I y = Fun OccursSec []" using Fun 2(3) by simp moreover have "Var y \ set T" using Var 2(2) length_Suc_conv[of T 1] by auto hence "y \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" using Fun t(1) by force hence "y \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using fv_ik_subset_fv_sst'[of "unlabel A"] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel A"] by blast ultimately have False using reachable_constraints_occurs_fact_ik_funs_terms(4)[OF \_reach \ P P_occ] by blast thus ?thesis by simp qed (use 2(3) in simp) moreover have "\u u'. T = [u,u']" using iffD1[OF length_Suc_conv 2(2)] iffD1[OF length_Suc_conv[of _ 0]] length_0_conv by fast ultimately show ?thesis using Fun 2(1,2) by force qed qed lemma reachable_constraints_occurs_fact_ik_subst: assumes \_reach: "A \ reachable_constraints P" and \: "welltyped_constraint_model I A" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and t: "occurs t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" shows "occurs t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" proof - have \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" using \ unfolding welltyped_constraint_model_def constraint_model_def by metis obtain s where s: "s \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "s \ I = occurs t" using t by auto hence u: "\u. s = occurs u" using \_wt reachable_constraints_occurs_fact_ik_subst_aux[OF \_reach \ P P_occ] by blast hence "fv s = {}" using reachable_constraints_occurs_fact_ik_ground[OF \_reach P P_occ] s by fast thus ?thesis using s u subst_ground_ident[of s I] by argo qed lemma reachable_constraints_occurs_fact_send_in_ik: assumes \_reach: "A \ reachable_constraints P" and \: "welltyped_constraint_model I A" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and x: "occurs (Var x) \ set ts" "send\ts\ \ set (unlabel A)" shows "occurs (I x) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using \_reach \ x proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" have T_adm: "admissible_transaction' T" using P step.hyps(2) unfolding list_all_iff by blast note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] note T_adm_occ = bspec[OF P_occ] have \_is_T_model: "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I) (set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I)) (unlabel T') I" using step.prems unlabel_append[of A T'] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of "unlabel A" I "[]"] strand_sem_append_stateful[of "{}" "{}" "unlabel A" "unlabel T'" I] by (simp add: T'_def \_def welltyped_constraint_model_def constraint_model_def db\<^sub>s\<^sub>s\<^sub>t_def) show ?case proof (cases "send\ts\ \ set (unlabel A)") case False hence "send\ts\ \ set (unlabel T')" using step.prems(3) unfolding T'_def \_def by simp hence "receive\ts\ \ set (unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(2) unfolding T'_def by blast then obtain y ts' where y: "receive\ts'\ \ set (unlabel (transaction_receive T))" "\ y = Var x" "occurs (Var y) \ set ts'" using transaction_decl_fresh_renaming_substs_occurs_fact_send_receive(2)[ OF step.hyps(3-5) T_adm] subst_to_var_is_var[of _ \ x] step.prems(2) - unfolding \_def by (metis subst_apply_term.simps(1)) + unfolding \_def by (metis eval_term.simps(1)) hence "occurs (Var y) \ \ \ set ts' \\<^sub>s\<^sub>e\<^sub>t \" "receive\ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\ \ set (unlabel (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using subst_lsst_unlabel_member[of "receive\ts'\" "transaction_receive T" \] by fastforce+ hence "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ occurs (Var y) \ \ \ I" using wellformed_transaction_sem_receives[ OF T_wf, of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" "set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I)" \ I "ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \"] \_is_T_model unfolding T'_def list_all_iff by fastforce hence *: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ occurs (\ y \ I)" by auto have "occurs (\ y \ I) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using deduct_occurs_in_ik[OF *] reachable_constraints_occurs_fact_ik_subst[ OF step.hyps(1) welltyped_constraint_model_prefix[OF step.prems(1)] P P_occ, of "\ y \ I"] reachable_constraints_occurs_fact_ik_funs_terms[ OF step.hyps(1) welltyped_constraint_model_prefix[OF step.prems(1)] P P_occ] by blast thus ?thesis using y(2) by simp qed (simp add: step.IH[OF welltyped_constraint_model_prefix[OF step.prems(1)]] step.prems(2)) qed simp lemma reachable_constraints_occurs_fact_deduct_in_ik: assumes \_reach: "A \ reachable_constraints P" and \: "welltyped_constraint_model I A" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and k: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ occurs k" shows "occurs k \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" and "occurs k \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using reachable_constraints_occurs_fact_ik_funs_terms(1-3)[OF \_reach \ P P_occ] reachable_constraints_occurs_fact_ik_subst[OF \_reach \ P P_occ] deduct_occurs_in_ik[OF k] by (presburger, presburger) lemma reachable_constraints_fv_bvars_subset: assumes A: "A \ reachable_constraints P" shows "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ (\T \ set P. bvars_transaction T)" using assms proof (induction A rule: reachable_constraints.induct) case (step \ T \ \ \) let ?T' = "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" show ?case using step.IH step.hyps(2) bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of ?T'] bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] bvars\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?T')"] unlabel_append[of \ "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?T'"] by (metis (no_types, lifting) SUP_upper Un_subset_iff) qed simp lemma reachable_constraints_fv_disj: fixes A::"('fun,'atom,'sets,'lbl) prot_constr" assumes A: "A \ reachable_constraints P" shows "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ (\T \ set P. bvars_transaction T) = {}" using A proof (induction A rule: reachable_constraints.induct) case (step \ T \ \ \) define T' where "T' \ transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" define X where "X \ \T \ set P. bvars_transaction T" have "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ X = {}" using transaction_decl_fresh_renaming_substs_vars_disj(4)[OF step.hyps(3-5)] transaction_decl_fresh_renaming_substs_vars_subset(4)[OF step.hyps(3-5,2)] unfolding T'_def X_def by blast hence "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T') \ X = {}" using step.IH[unfolded X_def[symmetric]] fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of T'] by auto thus ?case unfolding T'_def X_def by blast qed simp (* TODO: this lemma subsumes reachable_constraints_fv_bvars_disj *) lemma reachable_constraints_fv_bvars_disj': fixes A::"('fun,'atom,'sets,'lbl) prot_constr" assumes P: "\T \ set P. wellformed_transaction T" and A: "A \ reachable_constraints P" shows "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = {}" using A proof (induction A rule: reachable_constraints.induct) case (step \ T \ \ \) define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" note 0 = transaction_decl_fresh_renaming_substs_vars_disj[OF step.hyps(3-5)] note 1 = transaction_decl_fresh_renaming_substs_vars_subset[OF step.hyps(3-5)] have 2: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = {}" using 0(7) 1(4)[OF step.hyps(2)] fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unfolding T'_def by (metis (no_types) disjoint_iff_not_equal subset_iff) have "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ \(bvars_transaction ` set P)" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ \(bvars_transaction ` set P) = {}" using reachable_constraints_fv_bvars_subset[OF reachable_constraints.step[OF step.hyps]] reachable_constraints_fv_disj[OF reachable_constraints.step[OF step.hyps]] unfolding T'_def by auto hence 3: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = {}" by blast have "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \) \ bvars_transaction T = {}" using 0(4)[OF step.hyps(2)] 1(4)[OF step.hyps(2)] by blast hence 4: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = {}" by (metis (no_types) T'_def fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unlabel_subst bvars\<^sub>s\<^sub>s\<^sub>t_subst) have "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') = {}" using 2 3 4 step.IH unfolding unlabel_append[of \ T'] fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel T'"] bvars\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel T'"] by fast thus ?case unfolding T'_def by blast qed simp lemma reachable_constraints_wf: assumes P: "\T \ set P. wellformed_transaction T" "\T \ set P. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms_transaction T)" and A: "A \ reachable_constraints P" shows "wf\<^sub>s\<^sub>s\<^sub>t (unlabel A)" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" proof - let ?X = "\T. fst ` set (transaction_decl T ()) \ set (transaction_fresh T)" have "wellformed_transaction T" when "T \ set P" for T using P(1) that by fast+ hence 0: "wf'\<^sub>s\<^sub>s\<^sub>t (?X T) (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)))" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)) = {}" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms_transaction T)" when T: "T \ set P" for T unfolding admissible_transaction_terms_def by (metis T wellformed_transaction_wf\<^sub>s\<^sub>s\<^sub>t(1), metis T wellformed_transaction_wf\<^sub>s\<^sub>s\<^sub>t(2) fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq, metis T wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code P(2)) from A have "wf\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) let ?T' = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" have IH: "wf'\<^sub>s\<^sub>s\<^sub>t {} (unlabel A)" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = {}" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" using step.IH by metis+ have 1: "wf'\<^sub>s\<^sub>s\<^sub>t {} (unlabel (A@?T'))" using transaction_decl_fresh_renaming_substs_wf_sst[OF 0(1)[OF step.hyps(2)] step.hyps(3-5)] wf\<^sub>s\<^sub>s\<^sub>t_vars_mono[of "{}"] wf\<^sub>s\<^sub>s\<^sub>t_append[OF IH(1)] by simp have 2: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@?T') \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@?T') = {}" using reachable_constraints_fv_bvars_disj'[OF P(1)] reachable_constraints.step[OF step.hyps] by blast have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?T')" using trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unlabel_subst wf_trms_subst[ OF transaction_decl_fresh_renaming_substs_range_wf_trms[OF step.hyps(3-5)], THEN wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_trms\<^sub>s\<^sub>s\<^sub>t_subst, OF 0(3)[OF step.hyps(2)]] by metis hence 3: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@?T'))" using IH(3) by auto show ?case using 1 2 3 by force qed simp thus "wf\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" by metis+ qed lemma reachable_constraints_no_Ana_attack: assumes \: "\ \ reachable_constraints P" and P: "\T \ set P. wellformed_transaction T" "\T \ set P. admissible_transaction_terms T" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" and t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "attack\n\ \ set (snd (Ana t))" proof - have T_adm_term: "admissible_transaction_terms T" when "T \ set P" for T using P that by blast have T_wf: "wellformed_transaction T" when "T \ set P" for T using P that by blast have T_fresh: "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" when "T \ set P" for T using P(3) that by fast show ?thesis using \ t proof (induction \ rule: reachable_constraints.induct) case (step A T \ \ \) thus ?case proof (cases "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)") case False hence "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)))" using step.prems by simp hence "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using dual_transaction_ik_is_transaction_send'[OF T_wf[OF step.hyps(2)]] by metis hence "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" using transaction_decl_fresh_renaming_substs_trms[ OF step.hyps(3-5), of "transaction_send T"] wellformed_transaction_unlabel_cases(4)[OF T_wf[OF step.hyps(2)]] by fastforce then obtain s where s: "s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" "t = s \ \ \\<^sub>s \ \\<^sub>s \" by moura hence s': "attack\n\ \ set (snd (Ana s))" using admissible_transaction_no_Ana_Attack[OF T_adm_term[OF step.hyps(2)]] trms_transaction_unfold[of T] by blast note * = transaction_decl_fresh_renaming_substs_range'(1-3)[OF step.hyps(3-5)] transaction_decl_fresh_renaming_substs_range_no_attack_const[ OF step.hyps(3-5) T_fresh[OF step.hyps(2)]] show ?thesis proof assume n: "attack\n\ \ set (snd (Ana t))" thus False proof (cases s) case (Var x) hence "(\c. t = Fun c []) \ (\y. t = Var y)" using *(1)[of t] n s(2) by (force simp del: subst_subst_compose) thus ?thesis using n Ana_subterm' by fastforce next case (Fun f S) hence "attack\n\ \ set (snd (Ana s)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" using Ana_subst'[of f S _ "snd (Ana s)" "\ \\<^sub>s \ \\<^sub>s \"] s(2) s' n by (cases "Ana s") auto hence "attack\n\ \ set (snd (Ana s)) \ attack\n\ \ subst_range (\ \\<^sub>s \ \\<^sub>s \)" using const_mem_subst_cases' by fast thus ?thesis using *(4) s' by fast qed qed qed simp qed simp qed lemma reachable_constraints_receive_attack_if_attack: assumes \: "\ \ reachable_constraints P" and P: "\T \ set P. wellformed_transaction T" "\T \ set P. admissible_transaction_terms T" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" "\T \ set P. \x \ vars_transaction T. \TAtom AttackType \ \\<^sub>v x" and \: "welltyped_constraint_model \ \" and l: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ attack\l\" shows "attack\l\ \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" and "receive\[attack\l\]\ \ set (unlabel \)" and "\T \ set P. \s \ set (transaction_strand T). is_Send (snd s) \ length (the_msgs (snd s)) = 1 \ is_Fun_Attack (hd (the_msgs (snd s))) \ the_Attack_label (the_Fun (hd (the_msgs (snd s)))) = fst s \ (l,receive\[attack\l\]\) \ set \" (is "?Q \ (l,receive\[attack\l\]\) \ set \") proof - have \': "constr_sem_stateful \ (unlabel \)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using \ unfolding welltyped_constraint_model_def constraint_model_def by metis+ have 0: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \)" when \: "\ \ reachable_constraints P" for \ using reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s[OF _ \] admissible_transaction_terms_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P(2) ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset[of "unlabel \"] wf_trms_subst[OF \'(3)] by fast have 1: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \). \TAtom AttackType \ \\<^sub>v x" when \: "\ \ reachable_constraints P" for \ using reachable_constraints_vars_not_attack_typed[OF \ P(3,4)] fv_ik_subset_vars_sst'[of "unlabel \"] by fast have 2: "attack\l\ \ set (snd (Ana t)) \\<^sub>s\<^sub>e\<^sub>t \" when t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" for t proof assume "attack\l\ \ set (snd (Ana t)) \\<^sub>s\<^sub>e\<^sub>t \" then obtain s where s: "s \ set (snd (Ana t))" "s \ \ = attack\l\" by moura obtain x where x: "s = Var x" by (cases s) (use s reachable_constraints_no_Ana_attack[OF \ P(1-3) t] in auto) have "x \ fv t" using x Ana_subterm'[OF s(1)] vars_iff_subtermeq by force hence "x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using t fv_subterms by fastforce hence "\\<^sub>v x \ TAtom AttackType" using 1[OF \] by fastforce thus False using s(2) x wt_subst_trm''[OF \'(4), of "Var x"] by fastforce qed have 3: "attack\l\ \ set (snd (Ana t))" when t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \)" for t proof assume "attack\l\ \ set (snd (Ana t))" then obtain s where s: "s \ subterms\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "attack\l\ \ set (snd (Ana s))" using Ana_subst_subterms_cases[OF t] 2 by fast then obtain x where x: "x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "s \ \ x" by moura hence "\ x \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \)" using var_is_subterm[of x] subterms_subst_subset'[of \ "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] by force hence *: "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" "wf\<^sub>t\<^sub>r\<^sub>m s" using wf_trms_subterms[OF 0[OF \]] wf_trm_subtermeq[OF _ x(2)] by auto show False using term.order_trans[ OF subtermeq_imp_subtermtypeeq[OF *(2) Ana_subterm'[OF s(2)]] subtermeq_imp_subtermtypeeq[OF *(1) x(2)]] 1[OF \] x(1) wt_subst_trm''[OF \'(4), of "Var x"] by force qed have 4: "t = attack\n\" when t: "t \ \ \\<^sub>s \ \\<^sub>s \ = attack\n\" and hyps: "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and T: "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" for t n and T::"('fun, 'atom, 'sets, 'lbl) prot_transaction" and \ \ \::"('fun, 'atom, 'sets, 'lbl) prot_subst" and \::"('fun, 'atom, 'sets, 'lbl) prot_strand" proof (cases t) case (Var x) hence "attack\n\ \ subst_range (\ \\<^sub>s \ \\<^sub>s \)" - by (metis (no_types, lifting) t subst_apply_term.simps(1) subst_imgI term.distinct(1)) + by (metis (no_types, lifting) t eval_term.simps(1) subst_imgI term.distinct(1)) thus ?thesis using transaction_decl_fresh_renaming_substs_range_no_attack_const[OF hyps T] by blast qed (use t in simp) have 5: "\ts'. ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \ (l,send\ts'\) \ set (transaction_strand T)" when ts: "(l,receive\ts\) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" for l ts \ and T::"('fun, 'atom, 'sets, 'lbl) prot_transaction" using subst_lsst_memD(2)[OF ts[unfolded dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(1)[symmetric]]] by auto have 6: "l' = l" when "(l',receive\[attack\l\]\) \ set \" and Q: "?Q" for l' using \ that(1) proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) show ?case proof (cases "(l',receive\[attack\l\]\) \ set \") case False hence *: "(l',receive\[attack\l\]\) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" using step.prems by simp have "(l',send\[attack\l\]\) \ set (transaction_strand T)" using 4[OF _ step.hyps(3-5)] P(3) step.hyps(2) 5[OF *] by force thus ?thesis using Q step.hyps(2) unfolding is_Fun_Attack_def by fastforce qed (use step.IH in simp) qed simp have 7: "\t. ts = [t] \ t = attack\l\" when ts: "receive\ts\ \ set (unlabel \)" "attack\l\ \ set ts \\<^sub>s\<^sub>e\<^sub>t \" for ts using \ ts(1) proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) obtain t where t: "t \ set ts" "attack\l\ = t \ \" using ts(2) by blast hence t_in_ik: "t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ @ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" using step.prems(1) in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of t] by blast have t_attack_eq: "t = attack\l\" proof (cases t) case (Var x) hence "TAtom AttackType \ subterms (\ t)" using t_in_ik 1[OF reachable_constraints.step[OF step.hyps]] by fastforce thus ?thesis using t(2) wt_subst_trm''[OF \'(4), of t] by force qed (use t(2) in simp) show ?case proof (cases "receive\ts\ \ set (unlabel \)") case False then obtain l' where l': "(l', receive\ts\) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" using step.prems(1) unfolding unlabel_def by force then obtain ts' where ts': "ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" "(l', send\ts'\) \ set (transaction_strand T)" using 5 by meson then obtain t' where t': "t' \ set ts'" "t' \ \ \\<^sub>s \ \\<^sub>s \ = attack\l\" using t(1) t_attack_eq by force note * = t'(1) 4[OF t'(2) step.hyps(3-5)] have "send\ts'\ \ set (unlabel (transaction_strand T))" using ts'(2) step.hyps(2) P(2) unfolding unlabel_def by force hence "length ts' = 1" using step.hyps(2) P(2,3) * unfolding admissible_transaction_terms_def by fastforce hence "ts' = [attack\l\]" using * P(3) step.hyps(2) by (cases ts') auto thus ?thesis by (simp add: ts'(1)) qed (use step.IH in simp) qed simp show "attack\l\ \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" using private_const_deduct[OF _ l] 3 by simp then obtain ts where ts: "receive\ts\ \ set (unlabel \)" "attack\l\ \ set ts \\<^sub>s\<^sub>e\<^sub>t \" using in_ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t_iff[of _ \] unfolding unlabel_def by force then obtain t where "ts = [t]" "t = attack\l\" using 7 by blast thus "receive\[attack\l\]\ \ set (unlabel \)" using ts(1) by blast hence "\l'. (l', receive\[attack\l\]\) \ set \" unfolding unlabel_def by fastforce thus "(l,receive\[attack\l\]\) \ set \" when ?Q using that 6 by fast qed lemma reachable_constraints_receive_attack_if_attack': assumes \: "\ \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and \: "welltyped_constraint_model \ \" and n: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ attack\n\" shows "attack\n\ \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" and "receive\[attack\n\]\ \ set (unlabel \)" proof - have P': "\T \ set P. wellformed_transaction T" "\T \ set P. admissible_transaction_terms T" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" "\T \ set P. \x \ vars_transaction T. \TAtom AttackType \ \\<^sub>v x" using admissible_transaction_is_wellformed_transaction(1,4)[OF bspec[OF P]] admissible_transactionE(2,15)[OF bspec[OF P]] by (blast, blast, blast, blast) show "attack\n\ \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" "receive\[attack\n\]\ \ set (unlabel \)" using reachable_constraints_receive_attack_if_attack(1,2)[OF \ P'(1,2) _ P'(4) \ n] P'(3) by (metis, metis) qed lemma constraint_model_Value_term_is_Val: assumes \_reach: "A \ reachable_constraints P" and \: "welltyped_constraint_model I A" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and x: "\\<^sub>v x = TAtom Value" "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" shows "\n. I x = Fun (Val n) []" using reachable_constraints_occurs_fact_send_ex[OF \_reach P P_occ x] reachable_constraints_occurs_fact_send_in_ik[OF \_reach \ P P_occ] reachable_constraints_occurs_fact_ik_case[OF \_reach P P_occ] by fast lemma constraint_model_Value_term_is_Val': assumes \_reach: "A \ reachable_constraints P" and \: "welltyped_constraint_model I A" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and x: "(TAtom Value, m) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" shows "\n. I (TAtom Value, m) = Fun (Val n) []" using constraint_model_Value_term_is_Val[OF \_reach \ P P_occ _ x] by simp (* We use this lemma to show that fresh constants first occur in \(\) at the point where they were generated *) lemma constraint_model_Value_var_in_constr_prefix: assumes \_reach: "\ \ reachable_constraints P" and \: "welltyped_constraint_model \ \" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" shows "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \\<^sub>v x = TAtom Value \ (\B. prefix B \ \ x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \ \ x \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" (is "\x \ ?X \. ?R x \ ?Q x \") using \_reach \ proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) let ?P = "\\. \x \ ?X \. ?R x \ ?Q x \" define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" have IH: "?P \" using step welltyped_constraint_model_prefix by fast note \_empty = admissible_transaction_decl_subst_empty[OF bspec[OF P step.hyps(2)] step.hyps(3)] have T_adm: "admissible_transaction' T" by (metis P step.hyps(2)) note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] have \_is_T_model: "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) (set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)) (unlabel T') \" using step.prems unlabel_append[of \ T'] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of "unlabel \" \ "[]"] strand_sem_append_stateful[of "{}" "{}" "unlabel \" "unlabel T'" \] by (simp add: T'_def welltyped_constraint_model_def constraint_model_def db\<^sub>s\<^sub>s\<^sub>t_def) have \_interp: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (metis \ welltyped_constraint_model_def constraint_model_def, metis \ welltyped_constraint_model_def, metis \ welltyped_constraint_model_def constraint_model_def) have 1: "?Q x \" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "\\<^sub>v x = TAtom Value" for x proof - obtain n where n: "\ x = Fun n []" "is_Val n" "\public n" using constraint_model_Value_term_is_Val[ OF reachable_constraints.step[OF step.hyps] step.prems P P_occ x(2)] x(1) fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel T'"] unlabel_append[of \ T'] unfolding T'_def by moura have "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using x(1) fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unfolding T'_def by fastforce then obtain y where y: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)" "x \ fv ((\ \\<^sub>s \ \\<^sub>s \) y)" using fv\<^sub>s\<^sub>s\<^sub>t_subst_obtain_var[of x "unlabel (transaction_strand T)" "\ \\<^sub>s \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] by auto have y_x: "(\ \\<^sub>s \ \\<^sub>s \) y = Var x" and y_not_fresh: "y \ set (transaction_fresh T)" using y(2) transaction_decl_fresh_renaming_substs_range[OF step.hyps(3-5), of y] by (force, fastforce) have "\ ((\ \\<^sub>s \ \\<^sub>s \) y) = TAtom Value" using x(2) y_x by simp moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \ \\<^sub>s \)" by (rule transaction_decl_fresh_renaming_substs_wt[OF step.hyps(3-5)]) ultimately have y_val: "\\<^sub>v y = TAtom Value" by (metis wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def \.simps(1)) have "Fun n [] = (\ \\<^sub>s \ \\<^sub>s \) y \ \" using n y_x by simp hence y_n: "Fun n [] = (\ \\<^sub>s \ \\<^sub>s \ \\<^sub>s \) y" - by (metis subst_subst_compose[of "Var y" "\ \\<^sub>s \ \\<^sub>s \" \] subst_apply_term.simps(1)) + by (metis subst_subst_compose[of "Var y" "\ \\<^sub>s \ \\<^sub>s \" \] eval_term.simps(1)) have \_ik_\_vals: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \). \f. \ x = Fun f []" proof - have "\a. \ (\ x) = Var a" when "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" for x using that reachable_constraints_vars_TAtom_typed[OF step.hyps(1) P, of x] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] wt_subst_trm''[OF \_wt, of "Var x"] by force hence "\f. \ x = Fun f []" when "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" for x using that wf_trm_subst[OF \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s, of "Var x"] wf_trm_Var[of x] const_type_inv_wf empty_fv_exists_fun[OF interpretation_grounds[OF \_interp], of "Var x"] - by (metis subst_apply_term.simps(1)[of x \]) + by (metis eval_term.simps(1)[of _ x \]) thus ?thesis using fv_ik_subset_fv_sst'[of "unlabel \"] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] by blast qed hence \_subterms_subst_cong: "subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) = subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \" by (metis ik\<^sub>s\<^sub>s\<^sub>t_subst[of "unlabel \" \] unlabel_subst[of \ \] subterms_subst_lsst_ik[of \ \]) have x_nin_\: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" proof - have "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using x(1) fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unfolding T'_def by fast hence "x \ fv\<^sub>s\<^sub>s\<^sub>t ((unlabel (transaction_strand T) \\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>s\<^sub>t \)" using transaction_fresh_subst_grounds_domain[OF step.hyps(4)] step.hyps(4) labeled_stateful_strand_subst_comp[of \ "transaction_strand T" \] unlabel_subst[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" \] unlabel_subst[of "transaction_strand T" \] by (simp add: \_empty transaction_fresh_subst_def range_vars_alt_def) then obtain y where "\ y = Var x" using transaction_renaming_subst_var_obtain(1)[OF step.hyps(5)] by blast thus ?thesis using transaction_renaming_subst_range_notin_vars[OF step.hyps(5), of y] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] by auto qed from admissible_transaction_fv_in_receives_or_selects[OF T_adm y(1) y_not_fresh] have n_cases: "Fun n [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ (\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \\<^sub>v z = TAtom Value \ \ z = Fun n [])" proof assume y_in: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" then obtain ts where ts: "receive\ts\ \ set (unlabel (transaction_receive T))" "y \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" using admissible_transaction_strand_step_cases(1)[OF T_adm] by force hence ts_deduct: "list_all (\t. ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \ \\<^sub>s \ \\<^sub>s \ \ \) ts" using wellformed_transaction_sem_receives[ OF T_wf, of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" "set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" "\ \\<^sub>s \ \\<^sub>s \" \ "ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \"] \_is_T_model subst_lsst_unlabel_member[of "receive\ts\" "transaction_receive T" "\ \\<^sub>s \ \\<^sub>s \"] unfolding T'_def list_all_iff by force obtain ty where ty: "ty \ set ts" "y \ fv ty" using ts(2) by fastforce have "Fun n [] \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ (\z \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \). \\<^sub>v z = TAtom Value \ \ z = Fun n [])" proof - have "Fun n [] \ ty \ \ \\<^sub>s \ \\<^sub>s \ \ \" using imageI[of "Var y" "subterms ty" "\x. x \ \ \\<^sub>s \ \\<^sub>s \ \\<^sub>s \"] var_is_subterm[OF ty(2)] subterms_subst_subset[of "\ \\<^sub>s \ \\<^sub>s \ \\<^sub>s \" ty] subst_subst_compose[of ty "\ \\<^sub>s \ \\<^sub>s \" \] y_n by (auto simp del: subst_subst_compose) hence "Fun n [] \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" using ty(1) private_fun_deduct_in_ik[of _ _ n "[]"] n(2,3) ts_deduct unfolding is_Val_def is_Abs_def list_all_iff by fast hence "Fun n [] \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ (\z \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \). \ z = Fun n [])" using const_subterm_subst_cases[of n _ \] \_ik_\_vals by fastforce thus ?thesis using \_wt n(2) unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def is_Val_def is_Abs_def by force qed thus ?thesis using fv_ik_subset_fv_sst' ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset[of "unlabel \"] \_subterms_subst_cong by fast next assume y_in: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ (\t s. select\t,s\ \ set (unlabel (transaction_checks T)) \ y \ fv t \ fv s)" then obtain s where s: "select\Var y,Fun (Set s) []\ \ set (unlabel (transaction_checks T))" using admissible_transaction_strand_step_cases(2)[OF T_adm] by force hence "select\(\ \\<^sub>s \ \\<^sub>s \) y, Fun (Set s) []\ \ set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" using subst_lsst_unlabel_member by fastforce hence n_in_db: "(Fun n [], Fun (Set s) []) \ set (db'\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ [])" using wellformed_transaction_sem_pos_checks(1)[ OF T_wf, of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" "set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" "\ \\<^sub>s \ \\<^sub>s \" \ assign "(\ \\<^sub>s \ \\<^sub>s \) y" "Fun (Set s) []"] \_is_T_model n y_x unfolding T'_def db\<^sub>s\<^sub>s\<^sub>t_def by fastforce obtain tn sn where tsn: "insert\tn,sn\ \ set (unlabel \)" "Fun n [] = tn \ \" using db\<^sub>s\<^sub>s\<^sub>t_in_cases[OF n_in_db] by force have "Fun n [] = tn \ (\z. \\<^sub>v z = TAtom Value \ tn = Var z)" using \_wt tsn(2) n(2) unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def is_Val_def is_Abs_def by (cases tn) auto moreover have "tn \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "fv tn \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using tsn(1) in_subterms_Union by force+ ultimately show ?thesis using tsn(2) by auto qed from n_cases show ?thesis proof assume "\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \\<^sub>v z = TAtom Value \ \ z = Fun n []" then obtain B where B: "prefix B \" "Fun n [] \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" by (metis IH n(1)) thus ?thesis using n x_nin_\ trms\<^sub>s\<^sub>s\<^sub>t_unlabel_prefix_subset(1)[of B] by (metis (no_types, opaque_lifting) self_append_conv subset_iff subterms\<^sub>s\<^sub>e\<^sub>t_mono prefix_def) qed (use n x_nin_\ in fastforce) qed have "?P (\@T')" proof (intro ballI impI) fix x assume x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T')" "\\<^sub>v x = TAtom Value" show "?Q x (\@T')" proof (cases "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \") case False hence "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" using x(1) unlabel_append[of \] fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \"] by simp then obtain B where B: "prefix B \" "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t B" "\ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" using x(2) 1 by moura thus ?thesis using prefix_prefix by fast qed (use x(2) IH prefix_prefix in fast) qed thus ?case unfolding T'_def by blast qed simp lemma constraint_model_Val_const_in_constr_prefix: assumes \_reach: "\ \ reachable_constraints P" and \: "welltyped_constraint_model \ \" and P: "\T \ set P. wellformed_transaction T" "\T \ set P. admissible_transaction_terms T" and n: "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" shows "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" proof - have *: "wf\<^sub>s\<^sub>s\<^sub>t (unlabel \)" "constr_sem_stateful \ (unlabel \)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using reachable_constraints_wf(1)[OF P(1) _ \_reach] admissible_transaction_terms_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s \ P(2) n unfolding welltyped_constraint_model_def constraint_model_def wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code by fast+ show ?thesis using constraint_model_priv_const_in_constr_prefix[OF * _ _ n] by simp qed lemma constraint_model_Val_const_in_constr_prefix': assumes \_reach: "\ \ reachable_constraints P" and \: "welltyped_constraint_model \ \" and P: "\T \ set P. admissible_transaction' T" and n: "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" shows "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using constraint_model_Val_const_in_constr_prefix[OF \_reach \ _ _ n] P admissible_transaction_is_wellformed_transaction(1,4) by fast lemma constraint_model_Value_in_constr_prefix_fresh_action': fixes P::"('fun, 'atom, 'sets, 'lbl) prot_transaction list" assumes A: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction_terms T" "\T \ set P. transaction_decl T () = []" "\T \ set P. bvars_transaction T = {}" and n: "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" obtains B T \ \ \ where "prefix (B@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) A" and "B \ reachable_constraints P" "T \ set P" "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" and "Fun (Val n) [] \ subst_range \" proof - define f where "f \ \(T::('fun, 'atom, 'sets, 'lbl) prot_transaction, \::('fun, 'atom, 'sets, 'lbl) prot_subst, \::('fun, 'atom, 'sets, 'lbl) prot_subst, \::('fun, 'atom, 'sets, 'lbl) prot_subst). dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define g where "g \ concat \ map f" obtain Ts where Ts: "A = g Ts" "\B. prefix B Ts \ g B \ reachable_constraints P" "\B T \ \ \. prefix (B@[(T,\,\,\)]) Ts \ T \ set P \ transaction_decl_subst \ T \ transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B)) \ transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B))" using reachable_constraints_as_transaction_lists[OF A] unfolding g_def f_def by blast obtain T \ \ \ where T: "(T, \, \, \) \ set Ts" "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using n trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unlabel_subst unfolding Ts(1) g_def f_def unlabel_def trms\<^sub>s\<^sub>s\<^sub>t_def by fastforce obtain B where B: "prefix (B@[(T, \, \, \)]) Ts" "g B \ reachable_constraints P" "T \ set P" "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B))" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g B))" proof - obtain B where "\C. B@(T, \, \, \)#C = Ts" by (metis T(1) split_list) thus ?thesis using Ts(2-) that[of B] by auto qed note T_adm_terms = bspec[OF P(1) B(3)] note T_decl_empty = bspec[OF P(2) B(3)] note T_no_bvars = bspec[OF P(3) B(3)] note \_empty = admissible_transaction_decl_subst_empty'[OF T_decl_empty B(4)] have "trms\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T) \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \) = trms_transaction T \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" using trms\<^sub>s\<^sub>s\<^sub>t_subst[of _ "\ \\<^sub>s \ \\<^sub>s \"] T_no_bvars by blast hence "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms_transaction T \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" by (metis T(2) unlabel_subst) hence "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t subst_range (\ \\<^sub>s \ \\<^sub>s \)" using admissible_transaction_terms_no_Value_consts(1)[OF T_adm_terms] const_subterms_subst_cases'[of "Val n" "\ \\<^sub>s \ \\<^sub>s \" "trms_transaction T"] by blast then obtain tn where tn: "tn \ subst_range (\ \\<^sub>s \ \\<^sub>s \)" "Fun (Val n) [] \ tn" "is_Fun tn" by fastforce have "Fun (Val n) [] \ subst_range \" using tn(1-) transaction_decl_fresh_renaming_substs_range'(2,4)[OF B(4-6) tn(1) \_empty] by fastforce moreover have "prefix (g B@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) A" using Ts(1) B(1) unfolding g_def f_def prefix_def by fastforce ultimately show thesis using that B(2-) by blast qed lemma constraint_model_Value_in_constr_prefix_fresh_action: fixes P::"('fun, 'atom, 'sets, 'lbl) prot_transaction list" assumes A: "A \ reachable_constraints P" and P_adm: "\T \ set P. admissible_transaction' T" and n: "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" obtains B T \ \ \ where "prefix (B@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) A" and "B \ reachable_constraints P" "T \ set P" "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" and "Fun (Val n) [] \ subst_range \" proof - have P: "\T \ set P. admissible_transaction_terms T" "\T \ set P. transaction_decl T () = []" "\T \ set P. bvars_transaction T = {}" using P_adm admissible_transactionE(1) admissible_transaction_no_bvars(2) admissible_transaction_is_wellformed_transaction(4) by (blast,blast,blast) show ?thesis using that constraint_model_Value_in_constr_prefix_fresh_action'[OF A P n] by blast qed lemma reachable_constraints_occurs_fact_ik_case': fixes A::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and val: "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" shows "occurs (Fun (Val n) []) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" proof - obtain B T \ \ \ where B: "prefix (B@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) A" "B \ reachable_constraints P" "T \ set P" "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "Fun (Val n) [] \ subst_range \" using constraint_model_Value_in_constr_prefix_fresh_action[OF \_reach P val] by blast define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" have T_adm: "admissible_transaction' T" using P B(3) by blast hence T_wf: "wellformed_transaction T" "admissible_transaction_occurs_checks T" using admissible_transaction_is_wellformed_transaction(1) bspec[OF P_occ B(3)] by (blast,blast) obtain x where x: "x \ set (transaction_fresh T)" "\ x = Fun (Val n) []" using transaction_fresh_subst_domain[OF B(5)] B(7) admissible_transaction_decl_subst_empty[OF T_adm B(4)] by (force simp add: subst_compose \_def) obtain ts where ts: "send\ts\ \ set (unlabel (transaction_send T))" "occurs (Var x) \ set ts" using admissible_transaction_occurs_checksE2[OF T_wf(2) x(1)] by (metis (mono_tags, lifting) list.set_intros(1) unlabel_Cons(1)) have "occurs (Var x) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" using ts by force hence "occurs (Var x) \ \ \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using dual_transaction_ik_is_transaction_send'[OF T_wf(1), of \] by fast hence "occurs (Fun (Val n) []) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using x(2) by simp thus ?thesis using B(1)[unfolded \_def[symmetric]] unlabel_append[of B "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)"] ik\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel B" "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))"] unfolding prefix_def by force qed lemma reachable_constraints_occurs_fact_ik_case'': fixes A::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "A \ reachable_constraints P" and \: "welltyped_constraint_model \ A" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and val: "Fun (Val n) [] \ t" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ \ t" shows "occurs (Fun (Val n) []) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" proof - obtain f ts where t: "t = Fun f ts" using val(1) by (cases t) simp_all show ?thesis using private_fun_deduct_in_ik[OF val(2,1)[unfolded t]] constraint_model_Val_const_in_constr_prefix'[OF \_reach \ P, of n] reachable_constraints_occurs_fact_ik_case'[OF \_reach P P_occ, of n] by fastforce qed lemma admissible_transaction_occurs_checks_prop: assumes \_reach: "\ \ reachable_constraints P" and \: "welltyped_constraint_model \ \" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and f: "f \ \(funs_term ` (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" shows "\is_PubConstValue f" and "\is_Abs f" proof - obtain x where x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "f \ funs_term (\ x)" using f by moura obtain T where T: "Fun f T \ \ x" using funs_term_Fun_subterm[OF x(2)] by moura have \_interp: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (metis \ welltyped_constraint_model_def constraint_model_def, metis \ welltyped_constraint_model_def, metis \ welltyped_constraint_model_def constraint_model_def) note 0 = x(1) reachable_constraints_vars_TAtom_typed[OF \_reach P, of x] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] have 1: "\ (Var x) = \ (\ x)" using wt_subst_trm''[OF \_wt, of "Var x"] by simp hence "\a. \ (\ x) = Var a" using 0 by force hence "\f. \ x = Fun f []" using x(1) wf_trm_subst[OF \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s, of "Var x"] wf_trm_Var[of x] const_type_inv_wf empty_fv_exists_fun[OF interpretation_grounds[OF \_interp], of "Var x"] - by (metis subst_apply_term.simps(1)[of x \]) + by (metis eval_term.simps(1)[of _ x \]) hence 2: "\ x = Fun f []" using x(2) by force have 3: "\\<^sub>v x \ TAtom AbsValue" using 0 by force have "\is_PubConstValue f \ \is_Abs f" proof (cases "\\<^sub>v x = TAtom Value") case True then obtain B where B: "prefix B \" "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t B" "\ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" using constraint_model_Value_var_in_constr_prefix[OF \_reach \ P P_occ] x(1) by fast have "\ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using B(1,3) trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel B"] unlabel_append[of B] unfolding prefix_def by auto hence "f \ \(funs_term ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using x(2) funs_term_subterms_eq(2)[of "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] by blast thus ?thesis using reachable_constraints_val_funs_private[OF \_reach P] by blast+ next case False thus ?thesis using x 1 2 3 unfolding is_PubConstValue_def by (cases f) auto qed thus "\is_PubConstValue f" "\is_Abs f" by metis+ qed lemma admissible_transaction_occurs_checks_prop': assumes \_reach: "\ \ reachable_constraints P" and \: "welltyped_constraint_model \ \" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and f: "f \ \(funs_term ` (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" shows "\n. f = PubConst Value n" and "\n. f = Abs n" using admissible_transaction_occurs_checks_prop[OF \_reach \ P P_occ f] unfolding is_PubConstValue_def by auto lemma transaction_var_becomes_Val: assumes \_reach: "\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \) \ reachable_constraints P" and \: "welltyped_constraint_model \ (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and T: "T \ set P" and x: "x \ fv_transaction T" "fst x = TAtom Value" shows "\n. Fun (Val n) [] = (\ \\<^sub>s \ \\<^sub>s \) x \ \" proof - obtain m where m: "x = (TAtom Value, m)" by (metis x(2) eq_fst_iff) note \_empty = admissible_transaction_decl_subst_empty[OF bspec[OF P T] \] have x_not_bvar: "x \ bvars_transaction T" "fv ((\ \\<^sub>s \ \\<^sub>s \) x) \ bvars_transaction T = {}" using x(1) admissible_transactions_fv_bvars_disj[OF P] T transaction_decl_fresh_renaming_substs_vars_disj(2)[OF \ \ \, of x] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] by (blast, blast) have \x_type: "\ (\ x) = TAtom Value" using \ x \\<^sub>v_TAtom''(2)[of x] wt_subst_trm''[of \ "Var x"] unfolding transaction_fresh_subst_def by simp show ?thesis proof (cases "x \ subst_domain \") case True then obtain c where c: "\ x = Fun c []" "\public c" "arity c = 0" using \ unfolding transaction_fresh_subst_def by fastforce then obtain n where n: "c = Val n" using \x_type by (cases c) (auto split: option.splits) show ?thesis using c n subst_compose[of \ \ x] \_empty by simp next case False hence "\ x = Var x" by auto then obtain n where n: "(\ \\<^sub>s \) x = Var (TAtom Value, n)" using m transaction_renaming_subst_is_renaming(1)[OF \] subst_compose[of \ \ x] by force hence "(TAtom Value, n) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using x_not_bvar fv\<^sub>s\<^sub>s\<^sub>t_subst_fv_subset[OF x(1), of "\ \\<^sub>s \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] \_empty by force hence "\n'. \ (TAtom Value, n) = Fun (Val n') []" using constraint_model_Value_term_is_Val'[OF \_reach \ P P_occ, of n] x fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \"] fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \"] unlabel_append[of \] by fastforce thus ?thesis using n \_empty by simp qed qed lemma reachable_constraints_SMP_subset: assumes \: "\ \ reachable_constraints P" shows "SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ SMP (\T \ set P. trms_transaction T)" (is "?A \") and "SMP (pair`setops\<^sub>s\<^sub>s\<^sub>t (unlabel \)) \ SMP (\T\set P. pair`setops_transaction T)" (is "?B \") proof - have "?A \ \ ?B \" using \ proof (induction \ rule: reachable_constraints.induct) case (step A T \ \ \) define T' where "T' \ transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" define M where "M \ \T \ set P. trms_transaction T" define N where "N \ \T \ set P. pair ` setops_transaction T" let ?P = "\t. \s \. s \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = s \ \" let ?Q = "\t. \s \. s \ N \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = s \ \" have IH: "SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ SMP M" "SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)) \ SMP N" using step.IH by (metis M_def, metis N_def) note \\\_wt = transaction_decl_fresh_renaming_substs_wt[OF step.hyps(3-5)] note \\\_wf = transaction_decl_fresh_renaming_substs_range_wf_trms[OF step.hyps(3-5)] have 0: "SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')) = SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" "SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'))) = SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)) \ SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel T'))" using trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of T'] setops\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of T'] trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel A" "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')"] setops\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel A" "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')"] unlabel_append[of A "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'"] image_Un[of pair "setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "setops\<^sub>s\<^sub>s\<^sub>t (unlabel T')"] SMP_union[of "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'"] SMP_union[of "pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel T')"] by argo+ have 1: "SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T') \ SMP M" proof (intro SMP_subset_I ballI) fix t show "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ ?P t" using trms\<^sub>s\<^sub>s\<^sub>t_wt_subst_ex[OF \\\_wt \\\_wf, of t "unlabel (transaction_strand T)"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] step.hyps(2) unfolding T'_def M_def by auto qed have 2: "SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel T')) \ SMP N" proof (intro SMP_subset_I ballI) fix t show "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel T') \ ?Q t" using setops\<^sub>s\<^sub>s\<^sub>t_wt_subst_ex[OF \\\_wt \\\_wf, of t "unlabel (transaction_strand T)"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] step.hyps(2) unfolding T'_def N_def by auto qed have "SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')) \ SMP M" "SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'))) \ SMP N" using 0 1 2 IH by blast+ thus ?case unfolding T'_def M_def N_def by blast qed (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) thus "?A \" "?B \" by metis+ qed lemma reachable_constraints_no_Pair_fun': assumes A: "A \ reachable_constraints P" and P: "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" "\T \ set P. transaction_decl T () = []" "\T \ set P. admissible_transaction_terms T" "\T \ set P. \x \ vars_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = \a\\<^sub>\\<^sub>a)" shows "Pair \ \(funs_term ` SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" using A proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" note T_fresh_type = bspec[OF P(1) step.hyps(2)] note \_empty = admissible_transaction_decl_subst_empty'[OF bspec[OF P(2) step.hyps(2)] step.hyps(3)] note T_adm_terms = bspec[OF P(3) step.hyps(2)] note \\\_wt = transaction_decl_fresh_renaming_substs_wt[OF step.hyps(3-5)] note \\\_wf = transaction_decl_fresh_renaming_substs_range_wf_trms[OF step.hyps(3-5)] have 0: "SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T')) = SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" using SMP_union[of "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'"] unlabel_append[of A T'] trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel A" "unlabel T'"] by simp have 1: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" using reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s[OF _ reachable_constraints.step[OF step.hyps]] admissible_transaction_terms_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P(3) trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel A"] unlabel_append[of A] unfolding T'_def by force have 2: "Pair \ \(funs_term ` (subst_range (\ \\<^sub>s \ \\<^sub>s \)))" using transaction_decl_fresh_renaming_substs_range'(3)[ OF step.hyps(3-5) _ \_empty T_fresh_type] by force have "Pair \ \(funs_term ` (trms_transaction T))" using T_adm_terms unfolding admissible_transaction_terms_def by blast hence "Pair \ funs_term t" when t: "t \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T) \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" for t using 2 trms\<^sub>s\<^sub>s\<^sub>t_funs_term_cases[OF t] by force hence 3: "Pair \ funs_term t" when t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" for t using t unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \"] unfolding T'_def by metis have "\a. \\<^sub>v x = TAtom a" when "x \ vars_transaction T" for x using that protocol_transaction_vars_TAtom_typed(1) bspec[OF P(4) step.hyps(2)] by fast hence "\a. \\<^sub>v x = TAtom a" when "x \ vars\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T) \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" for x using wt_subst_fv\<^sub>s\<^sub>e\<^sub>t_termtype_subterm[OF _ \\\_wt \\\_wf, of x "vars_transaction T"] vars\<^sub>s\<^sub>s\<^sub>t_subst_cases[OF that] by fastforce hence "\a. \\<^sub>v x = TAtom a" when "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" for x using that unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] vars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \"] unfolding T'_def by simp hence "\a. \\<^sub>v x = TAtom a" when "x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" for x using that fv_trms\<^sub>s\<^sub>s\<^sub>t_subset(1) by fast hence "Pair \ funs_term (\ (Var x))" when "x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" for x using that by fastforce moreover have "Pair \ funs_term s" when s: "Ana s = (K, M)" "Pair \ \(funs_term ` set K)" for s::"('fun,'atom,'sets,'lbl) prot_term" and K M proof (cases s) case (Fun f S) thus ?thesis using s Ana_Fu_keys_not_pairs[of _ S K M] by (cases f) force+ qed (use s in simp) ultimately have "Pair \ funs_term t" when t: "t \ SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" for t using t 3 SMP_funs_term[OF t _ _ 1, of Pair] funs_term_type_iff by fastforce thus ?case using 0 step.IH(1) unfolding T'_def by blast qed simp lemma reachable_constraints_no_Pair_fun: assumes A: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" shows "Pair \ \(funs_term ` SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" using reachable_constraints_no_Pair_fun'[OF A] P admissible_transactionE(1,2,3) admissible_transaction_is_wellformed_transaction(4) by blast lemma reachable_constraints_setops_form: assumes A: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and t: "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" shows "\c s. t = pair (c, Fun (Set s) []) \ \ c = TAtom Value" using A t proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) have T_adm: "admissible_transaction' T" when "T \ set P" for T using P that unfolding list_all_iff by simp note T_adm' = admissible_transaction_is_wellformed_transaction(2,3)[OF T_adm] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] note \\\_wt = transaction_decl_fresh_renaming_substs_wt[OF step.hyps(3-5)] note \\\_wf = transaction_decl_fresh_renaming_substs_range_wf_trms[OF step.hyps(3-5)] show ?case using step.IH proof (cases "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)") case False hence "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T) \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using step.prems setops\<^sub>s\<^sub>s\<^sub>t_append unlabel_append setops\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] by fastforce then obtain t' \ where t': "t' \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T))" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t = t' \ \" using setops\<^sub>s\<^sub>s\<^sub>t_wt_subst_ex[OF \\\_wt \\\_wf] by blast then obtain s s' where s: "t' = pair (s,s')" using setops\<^sub>s\<^sub>s\<^sub>t_are_pairs by fastforce moreover have "InSet ac s s' = InSet Assign s s' \ InSet ac s s' = InSet Check s s'" for ac by (cases ac) simp_all ultimately have "\n. s = Var (Var Value, n)" "\u. s' = Fun (Set u) []" using t'(1) setops\<^sub>s\<^sub>s\<^sub>t_member_iff[of s s' "unlabel (transaction_strand T)"] pair_in_pair_image_iff[of s s'] transaction_inserts_are_Value_vars[ OF T_wf[OF step.hyps(2)] T_adm'(2)[OF step.hyps(2)], of s s'] transaction_deletes_are_Value_vars[ OF T_wf[OF step.hyps(2)] T_adm'(2)[OF step.hyps(2)], of s s'] transaction_selects_are_Value_vars[ OF T_wf[OF step.hyps(2)] T_adm'(1)[OF step.hyps(2)], of s s'] transaction_inset_checks_are_Value_vars[ OF T_adm[OF step.hyps(2)], of s s'] transaction_notinset_checks_are_Value_vars[ OF T_adm[OF step.hyps(2)], of _ _ _ s s'] by metis+ then obtain ss n where ss: "t = pair (\ (Var Value, n), Fun (Set ss) [])" using t'(4) s unfolding pair_def by force have "\ (\ (Var Value, n)) = TAtom Value" "wf\<^sub>t\<^sub>r\<^sub>m (\ (Var Value, n))" using t'(2) wt_subst_trm''[OF t'(2), of "Var (Var Value, n)"] apply simp using t'(3) by (cases "(Var Value, n) \ subst_domain \") auto thus ?thesis using ss by blast qed simp qed (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) lemma reachable_constraints_insert_delete_form: assumes A: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and t: "insert\t,s\ \ set (unlabel A) \ delete\t,s\ \ set (unlabel A)" (is "?Q t s A") shows "\k. s = Fun (Set k) []" (is ?A) and "\ t = TAtom Value" (is ?B) and "(\x. t = Var x) \ (\n. t = Fun (Val n) [])" (is ?C) proof - have 0: "pair (t,s) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" using t unfolding setops\<^sub>s\<^sub>s\<^sub>t_def by force show 1: ?A ?B using reachable_constraints_setops_form[OF A P 0] by (fast,fast) show ?C using A t proof (induction A rule: reachable_constraints.induct) case (step \ T \ \ \) let ?T' = "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" note T_adm = bspec[OF P step.hyps(2)] note T_wf = admissible_transaction_is_wellformed_transaction(1,3)[OF T_adm] have "?Q t s \ \ ?Q t s ?T'" using step.prems dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(4,5)[of t s ?T'] unfolding unlabel_append by auto thus ?case proof assume "?Q t s ?T'" then obtain u v where u: "?Q u v (transaction_strand T)" "t = u \ \ \\<^sub>s \ \\<^sub>s \" by (metis (no_types, lifting) stateful_strand_step_mem_substD(4,5) unlabel_subst) obtain x where x: "u = Var x" using u(1) transaction_inserts_are_Value_vars(1)[OF T_wf, of u v] transaction_deletes_are_Value_vars(1)[OF T_wf, of u v] by fastforce show ?case using u(2) x transaction_decl_fresh_renaming_substs_range'(3)[ OF step.hyps(3,4,5) _ admissible_transaction_decl_subst_empty[OF T_adm step.hyps(3)] admissible_transactionE(2)[OF T_adm], of t] by (cases "t \ subst_range (\ \\<^sub>s \ \\<^sub>s \)") - (blast, metis subst_apply_term.simps(1) subst_imgI) + (blast, metis eval_term.simps(1) subst_imgI) qed (use step.IH in fastforce) qed simp qed lemma reachable_constraints_setops_type: fixes t::"('fun,'atom,'sets,'lbl) prot_term" assumes A: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and t: "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" shows "\ t = TComp Pair [TAtom Value, TAtom SetType]" proof - obtain s c where s: "t = pair (c, Fun (Set s) [])" "\ c = TAtom Value" using reachable_constraints_setops_form[OF A P t] by moura hence "(Fun (Set s) []::('fun,'atom,'sets,'lbl) prot_term) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using t setops\<^sub>s\<^sub>s\<^sub>t_member_iff[of c "Fun (Set s) []" "unlabel A"] by force hence "wf\<^sub>t\<^sub>r\<^sub>m (Fun (Set s) []::('fun,'atom,'sets,'lbl) prot_term)" using reachable_constraints_wf(2) P A admissible_transaction_is_wellformed_transaction(1,4) unfolding admissible_transaction_terms_def by blast hence "arity (Set s) = 0" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by simp thus ?thesis using s unfolding pair_def by fastforce qed lemma reachable_constraints_setops_same_type_if_unifiable: assumes A: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" shows "\s \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A). \t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A). (\\. Unifier \ s t) \ \ s = \ t" (is "?P A") using reachable_constraints_setops_type[OF A P] by simp lemma reachable_constraints_setops_unifiable_if_wt_instance_unifiable: assumes A: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" shows "\s \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A). \t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A). (\\ \ \. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ Unifier \ (s \ \) (t \ \)) \ (\\. Unifier \ s t)" proof (intro ballI impI) fix s t assume st: "s \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" and "\\ \ \. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ Unifier \ (s \ \) (t \ \)" then obtain \ \ \ where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "Unifier \ (s \ \) (t \ \)" by moura obtain fs ft cs ct where c: "s = pair (cs, Fun (Set fs) [])" "t = pair (ct, Fun (Set ft) [])" "\ cs = TAtom Value" "\ ct = TAtom Value" using reachable_constraints_setops_form[OF A P st(1)] reachable_constraints_setops_form[OF A P st(2)] by moura have "cs \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "ct \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" using c(1,2) setops_subterm_trms[OF st(1), of cs] setops_subterm_trms[OF st(2), of ct] Fun_param_is_subterm[of cs "args s"] Fun_param_is_subterm[of ct "args t"] unfolding pair_def by simp_all moreover have "\T \ set P. wellformed_transaction T" "\T \ set P. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms_transaction T)" using P admissible_transaction_is_wellformed_transaction(1,4) unfolding admissible_transaction_terms_def by fast+ ultimately have *: "wf\<^sub>t\<^sub>r\<^sub>m cs" "wf\<^sub>t\<^sub>r\<^sub>m ct" using reachable_constraints_wf(2)[OF _ _ A] wf_trms_subterms by blast+ have "(\x. cs = Var x) \ (\c d. cs = Fun c [])" using const_type_inv_wf c(3) *(1) by (cases cs) auto moreover have "(\x. ct = Var x) \ (\c d. ct = Fun c [])" using const_type_inv_wf c(4) *(2) by (cases ct) auto ultimately show "\\. Unifier \ s t" using reachable_constraints_setops_form[OF A P] reachable_constraints_setops_type[OF A P] st \ c unfolding pair_def by auto qed lemma reachable_constraints_tfr: assumes M: "M \ \T \ set P. trms_transaction T" "has_all_wt_instances_of \ M N" "finite N" "tfr\<^sub>s\<^sub>e\<^sub>t N" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and P: "\T \ set P. admissible_transaction T" "\T \ set P. list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" and \: "\ \ reachable_constraints P" shows "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel \)" using \ proof (induction \ rule: reachable_constraints.induct) case (step A T \ \ \) define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" have P': "\T \ set P. admissible_transaction' T" using P(1) admissible_transactionE'(1) by blast have AT'_reach: "A@T' \ reachable_constraints P" using reachable_constraints.step[OF step.hyps] unfolding T'_def by metis note T_adm = bspec[OF P(1) step.hyps(2)] note T_adm' = bspec[OF P'(1) step.hyps(2)] note \_empty = admissible_transaction_decl_subst_empty[OF T_adm' step.hyps(3)] note \\\_wt = transaction_decl_fresh_renaming_substs_wt[OF step.hyps(3-5)] note \\\_wf = transaction_decl_fresh_renaming_substs_range_wf_trms[OF step.hyps(3-5)] have \\\_bvars_disj: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) \ range_vars (\ \\<^sub>s \ \\<^sub>s \) = {}" using transaction_decl_fresh_renaming_substs_vars_disj(4)[OF step.hyps(3,4,5,2)] \_empty by simp have wf_trms_M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" using admissible_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P'(1) unfolding M(1) by blast have "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T'))" using reachable_constraints_SMP_subset(1)[OF AT'_reach] tfr_subset(3)[OF M(4), of "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T')"] SMP_SMP_subset[of M N] SMP_I'[OF wf_trms_M M(5,2)] unfolding M(1) by blast moreover have "\p. Ana (pair p) = ([],[])" unfolding pair_def by auto ultimately have 1: "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T') \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@T')))" using tfr_setops_if_tfr_trms[of "unlabel (A@T')"] reachable_constraints_no_Pair_fun[OF AT'_reach P'] reachable_constraints_setops_same_type_if_unifiable[OF AT'_reach P'] reachable_constraints_setops_unifiable_if_wt_instance_unifiable[OF AT'_reach P'] by blast have "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" using step.hyps(2) P(2) tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p unfolding comp_tfr\<^sub>s\<^sub>s\<^sub>t_def tfr\<^sub>s\<^sub>s\<^sub>t_def by fastforce hence "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel T')" using tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_all_wt_subst_apply[OF _ \\\_wt \\\_wf \\\_bvars_disj] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] unfolding T'_def by argo hence 2: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (A@T'))" using step.IH unlabel_append unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by auto have "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel (A@T'))" using 1 2 by (metis tfr\<^sub>s\<^sub>s\<^sub>t_def) thus ?case by (metis T'_def) qed simp lemma reachable_constraints_tfr': assumes M: "M \ \T \ set P. trms_transaction T \ pair' Pair ` setops_transaction T" "has_all_wt_instances_of \ M N" "finite N" "tfr\<^sub>s\<^sub>e\<^sub>t N" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and P: "\T \ set P. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms_transaction T)" "\T \ set P. list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" and \: "\ \ reachable_constraints P" shows "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel \)" using \ proof (induction \ rule: reachable_constraints.induct) case (step A T \ \ \) define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" have AT'_reach: "A@T' \ reachable_constraints P" using reachable_constraints.step[OF step.hyps] unfolding T'_def by metis note \\\_wt = transaction_decl_fresh_renaming_substs_wt[OF step.hyps(3-5)] note \\\_wf = transaction_decl_fresh_renaming_substs_range_wf_trms[OF step.hyps(3-5)] have \\\_bvars_disj: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) \ range_vars (\ \\<^sub>s \ \\<^sub>s \) = {}" by (rule transaction_decl_fresh_renaming_substs_vars_disj(4)[OF step.hyps(3,4,5,2)]) have wf_trms_M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" using P(1) setops\<^sub>s\<^sub>s\<^sub>t_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s(2) unfolding M(1) pair_code wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[symmetric] by fast have "SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T')) \ SMP M" "SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@T'))) \ SMP M" using reachable_constraints_SMP_subset[OF AT'_reach] SMP_mono[of "\T \ set P. trms_transaction T" M] SMP_mono[of "\T \ set P. pair ` setops_transaction T" M] unfolding M(1) pair_code[symmetric] by blast+ hence 1: "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T') \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@T')))" using tfr_subset(3)[OF M(4), of "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T') \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@T'))"] SMP_union[of "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T')" "pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@T'))"] SMP_SMP_subset[of M N] SMP_I'[OF wf_trms_M M(5,2)] by blast have "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" using step.hyps(2) P(2) tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p unfolding comp_tfr\<^sub>s\<^sub>s\<^sub>t_def tfr\<^sub>s\<^sub>s\<^sub>t_def by fastforce hence "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel T')" using tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_all_wt_subst_apply[OF _ \\\_wt \\\_wf \\\_bvars_disj] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] unfolding T'_def by argo hence 2: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (A@T'))" using step.IH unlabel_append unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by auto have "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel (A@T'))" using 1 2 by (metis tfr\<^sub>s\<^sub>s\<^sub>t_def) thus ?case by (metis T'_def) qed simp lemma reachable_constraints_typing_cond\<^sub>s\<^sub>s\<^sub>t: assumes M: "M \ \T \ set P. trms_transaction T \ pair' Pair ` setops_transaction T" "has_all_wt_instances_of \ M N" "finite N" "tfr\<^sub>s\<^sub>e\<^sub>t N" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and P: "\T \ set P. wellformed_transaction T" "\T \ set P. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms_transaction T)" "\T \ set P. list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" and \: "\ \ reachable_constraints P" shows "typing_cond\<^sub>s\<^sub>s\<^sub>t (unlabel \)" using reachable_constraints_wf[OF P(1,2) \] reachable_constraints_tfr'[OF M P(2,3) \] unfolding typing_cond\<^sub>s\<^sub>s\<^sub>t_def by blast context begin private lemma reachable_constraints_typing_result_aux: assumes 0: "wf\<^sub>s\<^sub>s\<^sub>t (unlabel \)" "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "wf\<^sub>s\<^sub>s\<^sub>t (unlabel (\@[(l,send\[attack\n\]\)]))" "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel (\@[(l,send\[attack\n\]\)]))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@[(l,send\[attack\n\]\)]))" proof - let ?n = "[(l,send\[attack\n\]\)]" let ?A = "\@?n" show "wf\<^sub>s\<^sub>s\<^sub>t (unlabel ?A)" using 0(1) wf\<^sub>s\<^sub>s\<^sub>t_append_suffix'[of "{}" "unlabel \" "unlabel ?n"] unlabel_append[of \ ?n] by simp show "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?A)" using 0(3) trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel ?n"] unlabel_append[of \ ?n] by fastforce have "\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?n \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel ?n). \c. t = Fun c []" "\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?n \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel ?n). Ana t = ([],[])" by (simp_all add: setops\<^sub>s\<^sub>s\<^sub>t_def) hence "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?n \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel ?n)))" using 0(2) tfr_consts_mono unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by blast hence "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@?n) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (\@?n)))" using unlabel_append[of \ ?n] trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel ?n"] setops\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel ?n"] by (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) thus "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel ?A)" using 0(2) unlabel_append[of ?A ?n] unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by auto qed lemma reachable_constraints_typing_result: fixes P assumes M: "has_all_wt_instances_of \ (\T \ set P. trms_transaction T) N" "finite N" "tfr\<^sub>s\<^sub>e\<^sub>t N" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and P: "\T \ set P. admissible_transaction T" "\T \ set P. list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" and A: "\ \ reachable_constraints P" and \: "constraint_model \ (\@[(l, send\[attack\n\]\)])" shows "\\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[(l, send\[attack\n\]\)])" proof - have I: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "constr_sem_stateful \ (unlabel (\@[(l, send\[attack\n\]\)]))" using \ unfolding constraint_model_def by metis+ note 0 = admissible_transaction_is_wellformed_transaction(1,4)[OF admissible_transactionE'(1)] have 1: "\T \ set P. wellformed_transaction T" using P(1) 0(1) by blast have 2: "\T \ set P. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms_transaction T)" using P(1) 0(2) unfolding admissible_transaction_terms_def by blast have 3: "wf\<^sub>s\<^sub>s\<^sub>t (unlabel \)" "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using reachable_constraints_tfr[OF _ M P A] reachable_constraints_wf[OF 1(1) 2 A] by metis+ show ?thesis using stateful_typing_result[OF reachable_constraints_typing_result_aux[OF 3] I(1,3)] by (metis welltyped_constraint_model_def constraint_model_def) qed lemma reachable_constraints_typing_result': fixes P assumes M: "M \ \T \ set P. trms_transaction T \ pair' Pair ` setops_transaction T" "has_all_wt_instances_of \ M N" "finite N" "tfr\<^sub>s\<^sub>e\<^sub>t N" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and P: "\T \ set P. wellformed_transaction T" "\T \ set P. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms_transaction T)" "\T \ set P. list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" and A: "\ \ reachable_constraints P" and \: "constraint_model \ (\@[(l, send\[attack\n\]\)])" shows "\\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[(l, send\[attack\n\]\)])" proof - have I: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "constr_sem_stateful \ (unlabel (\@[(l, send\[attack\n\]\)]))" using \ unfolding constraint_model_def by metis+ have 0: "wf\<^sub>s\<^sub>s\<^sub>t (unlabel \)" "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using reachable_constraints_tfr'[OF M P(2-3) A] reachable_constraints_wf[OF P(1,2) A] by metis+ show ?thesis using stateful_typing_result[OF reachable_constraints_typing_result_aux[OF 0] I(1,3)] by (metis welltyped_constraint_model_def constraint_model_def) qed end lemma reachable_constraints_transaction_proj: assumes "\ \ reachable_constraints P" shows "proj n \ \ reachable_constraints (map (transaction_proj n) P)" using assms proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) show ?case using step.hyps(2) reachable_constraints.step[OF step.IH _ transaction_decl_subst_proj[OF step.hyps(3)] transaction_fresh_subst_proj[OF step.hyps(4)] transaction_renaming_subst_proj[OF step.hyps(5)]] by (simp add: proj_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t proj_subst transaction_strand_proj) qed (simp add: reachable_constraints.init) context begin private lemma reachable_constraints_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_aux: fixes P defines "Ts \ concat (map transaction_strand P)" assumes A: "A \ reachable_constraints P" shows "\b \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A). \a \ set Ts. \\. b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ (\t \ subst_range \. (\x. t = Var x) \ (\c. t = Fun c []))" (is "\b \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A). \a \ set Ts. ?P b a") using A proof (induction A rule: reachable_constraints.induct) case (step \ T \ \ \) define Q where "Q \ ?P" define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" let ?R = "\A Ts. \b \ set A. \a \ set Ts. Q b a" have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "\t \ subst_range \. (\x. t = Var x) \ (\c. t = Fun c [])" using transaction_decl_fresh_renaming_substs_wt[OF step.hyps(3-5)] transaction_decl_fresh_renaming_substs_range_wf_trms[OF step.hyps(3-5)] transaction_decl_fresh_renaming_substs_range'(1)[OF step.hyps(3-5)] unfolding \_def by (metis,metis,fastforce) hence "?R (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) (transaction_strand T)" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_self_inverse[of "transaction_strand T"] by (auto simp add: Q_def subst_apply_labeled_stateful_strand_def) hence "?R (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) (transaction_strand T)" by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst) hence "?R (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) Ts" using step.hyps(2) unfolding Ts_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by fastforce thus ?case using step.IH unfolding Q_def \_def by auto qed simp lemma reachable_constraints_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t: fixes P defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and "Ts \ concat (map transaction_strand P)" assumes P_pc: "comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ Pair Ts M S" and A: "A \ reachable_constraints P" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A ((f S) - {m. intruder_synth {} m})" using par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_if_comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t'[OF P_pc, of "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A", THEN par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t] reachable_constraints_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_aux[OF A, unfolded Ts_def[symmetric]] unfolding f_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_self_inverse by fast end lemma reachable_constraints_par_comp_constr: fixes P f S defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and "Ts \ concat (map transaction_strand P)" and "Sec \ f S - {m. intruder_synth {} m}" and "M \ \T \ set P. trms_transaction T \ pair' Pair ` setops_transaction T" assumes M: "has_all_wt_instances_of \ M N" "finite N" "tfr\<^sub>s\<^sub>e\<^sub>t N" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and P: "\T \ set P. wellformed_transaction T" "\T \ set P. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms_transaction T)" "\T \ set P. list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" "comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ Pair Ts M_fun S" and \: "\ \ reachable_constraints P" and \: "constraint_model \ \" shows "\\\<^sub>\. welltyped_constraint_model \\<^sub>\ \ \ ((\n. welltyped_constraint_model \\<^sub>\ (proj n \)) \ (\\' l t. prefix \' \ \ suffix [(l, receive\t\)] \' \ strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' Sec \\<^sub>\))" proof - have \': "constr_sem_stateful \ (unlabel \)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using \ unfolding constraint_model_def by blast+ show ?thesis using reachable_constraints_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t[OF P(4)[unfolded Ts_def] \] reachable_constraints_typing_cond\<^sub>s\<^sub>s\<^sub>t[OF M_def M P(1-3) \] par_comp_constr_stateful[OF _ _ \', of Sec] unfolding f_def Sec_def welltyped_constraint_model_def constraint_model_def by blast qed lemma reachable_constraints_component_leaks_if_composed_leaks: fixes Sec Q defines "leaks \ \\. \\\<^sub>\ \'. Q \\<^sub>\ \ prefix \' \ \ (\l' t. suffix [(l', receive\t\)] \') \ strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' Sec \\<^sub>\" assumes Sec: "\s \ Sec. \{} \\<^sub>c s" "ground Sec" and composed_leaks: "\\ \ reachable_constraints Ps. leaks \" shows "\l. \\ \ reachable_constraints (map (transaction_proj l) Ps). leaks \" proof - from composed_leaks obtain \ \\<^sub>\ \' s n where \: "\ \ reachable_constraints Ps" and \': "prefix \' \" "constr_sem_stateful \\<^sub>\ (proj_unl n \'@[send\[s]\])" and \\<^sub>\: "Q \\<^sub>\" and s: "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\" unfolding leaks_def strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by fast have "\{} \\<^sub>c s" "s \ \\<^sub>\ = s" using s Sec by auto then obtain B k' u where "constr_sem_stateful \\<^sub>\ (proj_unl n B@[send\[s]\])" "prefix (proj n B) (proj n \)" "suffix [(k', receive\u\)] (proj n B)" "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n B) \\<^sub>\" using constr_sem_stateful_proj_priv_term_prefix_obtain[OF \' s] unfolding welltyped_constraint_model_def constraint_model_def by metis thus ?thesis using \\<^sub>\ reachable_constraints_transaction_proj[OF \, of n] proj_idem[of n B] unfolding leaks_def strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by metis qed lemma reachable_constraints_preserves_labels: assumes \: "\ \ reachable_constraints P" shows "\a \ set \. \T \ set P. \b \ set (transaction_strand T). fst b = fst a" (is "\a \ set \. \T \ set P. ?P T a") using \ proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) have "\a \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)). ?P T a" proof fix a assume a: "a \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" then obtain b where b: "b \ set (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" "a = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p b" unfolding dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto then obtain c where c: "c \ set (transaction_strand T)" "b = c \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \ \\<^sub>s \" unfolding subst_apply_labeled_stateful_strand_def by auto have "?P T c" using c(1) by blast hence "?P T b" using c(2) by (simp add: subst_lsstp_fst_eq) thus "?P T a" using b(2) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_fst_eq[of b] by presburger qed thus ?case using step.IH step.hyps(2) by (metis Un_iff set_append) qed simp lemma reachable_constraints_preserves_labels': assumes P: "\T \ set P. \a \ set (transaction_strand T). has_LabelN l a \ has_LabelS a" and \: "\ \ reachable_constraints P" shows "\a \ set \. has_LabelN l a \ has_LabelS a" using reachable_constraints_preserves_labels[OF \] P by fastforce lemma reachable_constraints_transaction_proj_proj_eq: assumes \: "\ \ reachable_constraints (map (transaction_proj l) P)" shows "proj l \ = \" and "prefix \' \ \ proj l \' = \'" using \ proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) let ?T = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" note A = reachable_constraints.step[OF step.hyps] have P: "\T \ set (map (transaction_proj l) P). \a \ set (transaction_strand T). has_LabelN l a \ has_LabelS a" using transaction_proj_labels[of l] unfolding list_all_iff by auto note * = reachable_constraints_preserves_labels'[OF P A] have **: "\a \ set \'. has_LabelN l a \ has_LabelS a" when "\a \ set B. has_LabelN l a \ has_LabelS a" "prefix \' B" for B using that assms unfolding prefix_def by auto note *** = proj_ident[unfolded list_all_iff] { case 1 thus ?case using *[THEN ***] by blast } { case 2 thus ?case using *[THEN **, THEN ***] by blast } qed (simp_all add: reachable_constraints.init) lemma reachable_constraints_transaction_proj_star_proj: assumes \: "\ \ reachable_constraints (map (transaction_proj l) P)" and k_neq_l: "k \ l" shows "proj k \ \ reachable_constraints (map transaction_star_proj P)" using \ proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) have "map (transaction_proj k) (map (transaction_proj l) P) = map transaction_star_proj P" using transaction_star_proj_negates_transaction_proj(2)[OF k_neq_l] by fastforce thus ?case using reachable_constraints_transaction_proj[OF reachable_constraints.step[OF step.hyps], of k] by argo qed (simp add: reachable_constraints.init) lemma reachable_constraints_aligned_prefix_ex: fixes P defines "f \ \T. list_all is_Receive (unlabel (transaction_receive T)) \ list_all is_Check_or_Assignment (unlabel (transaction_checks T)) \ list_all is_Update (unlabel (transaction_updates T)) \ list_all is_Send (unlabel (transaction_send T))" assumes P: "list_all f P" "list_all ((list_all (Not \ has_LabelS)) \ tl \ transaction_send) P" and s: "\{} \\<^sub>c s" "fv s = {}" and A: "A \ reachable_constraints P" "prefix B A" and B: "\l ts. suffix [(l, receive\ts\)] B" "constr_sem_stateful \ (unlabel B@[send\[s]\])" shows "\C \ reachable_constraints P. prefix C A \ (\l ts. suffix [(l, receive\ts\)] C) \ declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \ = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t C \ \ constr_sem_stateful \ (unlabel C@[send\[s]\])" using A proof (induction A rule: reachable_constraints.induct) case (step A T \ \ \) define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" let ?T = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" note AT_reach = reachable_constraints.step[OF step.hyps] obtain lb tsb B' where B': "B = B'@[(lb, receive\tsb\)]" using B(1) unfolding suffix_def by blast define decl_ik where "decl_ik \ \S::('fun,'atom,'sets,'lbl) prot_strand. \{set ts |ts. \\, receive\ts\\ \ set S} \\<^sub>s\<^sub>e\<^sub>t \" have decl_ik_append: "decl_ik (M@N) = decl_ik M \ decl_ik N" for M N unfolding decl_ik_def by fastforce have "\\, receive\ts\\ \ set N" when "\ \ fst ` set N" for ts and N::"('fun, 'atom, 'sets, 'lbl) prot_strand" using that by force hence decl_ik_star: "decl_ik M = decl_ik (M@N)" when "\ \ fst ` set N" for M N using that unfolding decl_ik_def by simp have decl_decl_ik: "declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ = {t. decl_ik S \ t}" for S unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_alt_def decl_ik_def by blast have "f T" using P(1) step.hyps(2) by (simp add: list_all_iff) hence "list_all is_Send (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" "list_all is_Check_or_Assignment (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" "list_all is_Update (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" "list_all is_Receive (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" using subst_sst_list_all(2)[of "unlabel (transaction_receive T)" \] subst_sst_list_all(11)[of "unlabel (transaction_checks T)" \] subst_sst_list_all(10)[of "unlabel (transaction_updates T)" \] subst_sst_list_all(1)[of "unlabel (transaction_send T)" \] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(1)[of "transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(11)[of "transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(10)[of "transaction_updates T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(2)[of "transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] unfolding f_def by (metis unlabel_subst[of _ \])+ hence "\list_ex is_Receive (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" "\list_ex is_Receive (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" "\list_ex is_Receive (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" "list_all is_Receive (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" unfolding list_ex_iff list_all_iff by blast+ then obtain TA TB where T: "?T = TA@TB" "\list_ex is_Receive (unlabel TA)" "list_all is_Receive (unlabel TB)" "TB = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using transaction_dual_subst_unfold[of T \] unfolding \_def by fastforce have 0: "prefix B (A@TA@TB)" using step.prems B' T by argo have 1: "prefix B A" when "prefix B (A@TA)" using that T(2) B' prefix_prefix_inv unfolding list_ex_iff unlabel_def by fastforce have 2: "\ \ fst ` set TB2" when "TB = TB1@(l,x)#TB2" for TB1 l x TB2 proof - have "k \ \" when "k \ set (map fst (tl (transaction_send T)))" for k using that P(2) step.hyps(2) unfolding list_all_iff by auto hence "k \ \" when "k \ set (map fst (tl TB))" for k using that subst_lsst_map_fst_eq[of "tl (transaction_send T)" \] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_map_fst_eq[of "tl (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)"] unfolding T(4) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_tl subst_lsst_tl by simp moreover have "set TB2 \ set (tl TB)" using that by (metis (no_types, lifting) append_eq_append_conv2 order.eq_iff list.sel(3) self_append_conv set_mono_suffix suffix_appendI suffix_tl tl_append2) ultimately show ?thesis by auto qed have 3: "declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t TB \ = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t (TB1@[(l,x)]) \" when "TB = TB1@(l,x)#TB2" for TB1 l x TB2 using decl_ik_star[OF 2[OF that], of "TB1@[(l,x)]"] unfolding that decl_decl_ik by simp show ?case proof (cases "prefix B A") case False hence 4: "\prefix B (A@TA)" using 1 by blast have 5: "\l ts. suffix [(l, receive\ts\)] (A@?T)" proof - have "(lb, receive\tsb\) \ set TB" using 0 4 prefix_prefix_inv[OF _ suffixI[OF B'], of "A@TA" TB] by (metis append_assoc) hence "receive\tsb\ \ set (unlabel TB)" unfolding unlabel_def by force hence "\ts. suffix [receive\ts\] (unlabel TB)" using T(3) unfolding list_all_iff is_Receive_def suffix_def by (metis in_set_conv_decomp list.distinct(1) list.set_cases rev_exhaust) then obtain TB' ts where "unlabel TB = TB'@[receive\ts\]" unfolding suffix_def by blast then obtain TB'' x where "TB = TB''@[x]" "snd x = receive\ts\" by (metis (no_types, opaque_lifting) append1_eq_conv list.distinct(1) rev_exhaust rotate1.simps(2) rotate1_is_Nil_conv unlabel_Cons(2) unlabel_append unlabel_nil) then obtain l where "suffix [(l, receive\ts\)] TB" by (metis surj_pair prod.sel(2) suffix_def) thus ?thesis using T(4) transaction_dual_subst_unfold[of T \] suffix_append[of "[(l, receive\ts\)]"] unfolding \_def by metis qed obtain TB1 where TB: "B = A@TA@TB1@[(lb, receive\tsb\)]" "prefix (TB1@[(lb, receive\tsb\)]) TB" using 0 4 B' prefix_snoc_obtain[of B' "(lb, receive\tsb\)" "A@TA" TB thesis] by (metis append_assoc) then obtain TB2 where TB2: "TB = TB1@(lb, receive\tsb\)#TB2" unfolding prefix_def by fastforce hence TB2': "list_all is_Receive (unlabel TB2)" using T(3) unfolding list_all_iff is_Receive_def proj_def unlabel_def by auto have 6: "constr_sem_stateful \ (unlabel B)" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \\<^sub>s\<^sub>e\<^sub>t \ \ s \ \" using B(2) strand_sem_append_stateful[of "{}" "{}" "unlabel B" "[send\[s]\]" \] by auto have "constr_sem_stateful \ (unlabel (A@TA@TB1@[(lb, receive\tsb\)]))" using 6(1) TB(1) by blast hence "constr_sem_stateful \ (unlabel (A@?T))" using T(1) TB2 strand_sem_receive_prepend_stateful[ of "{}" "{}" "unlabel (A@TA@TB1@[(lb, receive\tsb\)])" \, OF _ TB2'] by auto moreover have "set (unlabel B) \ set (unlabel (A@?T))" using step.prems(1) unfolding prefix_def by force hence "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@?T) \\<^sub>s\<^sub>e\<^sub>t \ \ s \ \" using ideduct_mono[OF 6(2)] subst_all_mono[of _ _ \] ik\<^sub>s\<^sub>s\<^sub>t_set_subset[of "unlabel B" "unlabel (A@?T)"] by meson ultimately have 7: "constr_sem_stateful \ (unlabel (A@?T)@[send\[s]\])" using strand_sem_append_stateful[of "{}" "{}" "unlabel (A@?T)" "[send\[s]\]" \] by auto have "declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \ = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@?T) \" proof - have "declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t TB \ = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t (TB1@[(lb, receive\tsb\)]) \" using 3[of _ lb "receive\tsb\"] TB(2) unfolding prefix_def by auto hence "(decl_ik TB \ t) \ decl_ik (TB1@[(lb,receive\tsb\)]) \ t" for t unfolding TB(1) T(1) decl_decl_ik by blast hence "(decl_ik (A@TA@TB) \ t) \ decl_ik (A@TA@TB1@[(lb,receive\tsb\)]) \ t" for t using ideduct_mono_eq[of "decl_ik TB" "decl_ik (TB1@[(lb,receive\tsb\)])" "decl_ik (A@TA)"] by (metis decl_ik_append[of "A@TA"] Un_commute[of _ "decl_ik (A@TA)"] append_assoc) thus ?thesis unfolding TB(1) T(1) decl_decl_ik by blast qed thus ?thesis using step.prems AT_reach B(1) 5 7 by force qed (use step.IH prefix_append in blast) qed (use B(1) suffix_def in simp) lemma reachable_constraints_secure_if_filter_secure_case: fixes f l n and P::"('fun,'atom,'sets,'lbl) prot_transaction list" defines "has_attack \ \P. \\ \ reachable_constraints P. \\. constraint_model \ (\@[(l, send\[attack\n\]\)])" and "f \ \T. list_ex (\a. is_Update (snd a) \ is_Send (snd a)) (transaction_strand T)" and "g \ \T. transaction_fresh T = [] \ f T" assumes att: "has_attack P" shows "has_attack (filter g P)" proof - let ?attack = "\A I. constraint_model I (A@[(l, send\[attack\n\]\)])" define constr' where "constr' \ \(T::('fun,'atom,'sets,'lbl) prot_transaction,\::('fun,'atom,'sets,'lbl) prot_subst, \::('fun,'atom,'sets,'lbl) prot_subst,\::('fun,'atom,'sets,'lbl) prot_subst). dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define constr where "constr \ \Ts. concat (map constr' Ts)" define h where "h \ \(T::('fun,'atom,'sets,'lbl) prot_transaction,_::('fun,'atom,'sets,'lbl) prot_subst, _::('fun,'atom,'sets,'lbl) prot_subst,_::('fun,'atom,'sets,'lbl) prot_subst). g T" obtain A I where A: "A \ reachable_constraints P" "?attack A I" using att unfolding has_attack_def by blast obtain Ts where Ts: "A = constr Ts" "\B. prefix B Ts \ constr B \ reachable_constraints P" "\B T \ \ \. prefix (B@[(T,\,\,\)]) Ts \ T \ set P \ transaction_decl_subst \ T \ transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr B)) \ transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr B))" using reachable_constraints_as_transaction_lists[OF A(1)] constr_def constr'_def by auto define B where "B \ constr (filter h Ts)" have Ts': "T \ set P" when "(T,\,\,\) \ set Ts" for T \ \ \ using that Ts(3) by (meson prefix_snoc_in_iff) have constr_Cons: "constr (p#Ts) = constr' p@constr Ts" for p Ts unfolding constr_def by simp have constr_snoc: "constr (Ts@[p]) = constr Ts@constr' p" for p Ts unfolding constr_def by simp have 0: "?attack B I" when A_att: "?attack A I" proof - have not_f_T_case: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr' p) = {}" "\D. dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (constr' p)) I D = D" when not_f_T: "\(f T)" and p: "p = (T,\,\,\)" for p T \ \ \ proof - have constr_p: "constr' p = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" unfolding p constr'_def by fast have "\is_Receive a" when a: "(l,a) \ set (constr' p)" for l a proof assume "is_Receive a" then obtain ts where ts: "a = receive\ts\" by (cases a) auto then obtain ts' where ts': "(l,send\ts'\) \ set (transaction_strand T)" "ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" using a dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(1)[of l ts] subst_lsst_memD(2)[of l _ "transaction_strand T"] unfolding constr_p by blast thus False using not_f_T unfolding f_def list_ex_iff by fastforce qed thus "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr' p) = {}" using in_ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t_iff by fastforce have "\is_Update a" when a: "(l,a) \ set (constr' p)" for l a proof assume "is_Update a" then obtain t s where ts: "a = insert\t,s\ \ a = delete\t,s\" by (cases a) auto then obtain t' s' where ts': "(l,insert\t',s'\) \ set (transaction_strand T) \ (l,delete\t',s'\) \ set (transaction_strand T)" "t = t' \ \ \\<^sub>s \ \\<^sub>s \" "s = s' \ \ \\<^sub>s \ \\<^sub>s \" using a dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(4,5)[of l] subst_lsst_memD(4,5)[of l _ _ "transaction_strand T"] unfolding constr_p by blast thus False using not_f_T unfolding f_def list_ex_iff by fastforce qed thus "\D. dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (constr' p)) I D = D" using dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd[of "unlabel (constr' p)" I] by (meson unlabel_mem_has_label) qed have *: "strand_sem_stateful M D (unlabel B) I" when "strand_sem_stateful M D (unlabel A) I" for M D using that Ts' unfolding Ts(1) B_def proof (induction Ts arbitrary: M D) case (Cons p Ts) obtain T \ \ \ where p: "p = (T,\,\,\)" by (cases p) simp have T_in: "T \ set P" using Cons.prems(2) unfolding p by fastforce let ?M' = "M \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr' p) \\<^sub>s\<^sub>e\<^sub>t I)" let ?D' = "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (constr' p)) I D" have p_sem: "strand_sem_stateful M D (unlabel (constr' p)) I" and IH: "strand_sem_stateful ?M' ?D' (unlabel (constr (filter h Ts))) I" using Cons.IH[of ?M' ?D'] Cons.prems strand_sem_append_stateful[of M D "unlabel (constr' p)" "unlabel (constr Ts)" I] unfolding constr_Cons unlabel_append by fastforce+ show ?case proof (cases "T \ set (filter g P)") case True hence h_p: "filter h (p#Ts) = p#filter h Ts" unfolding h_def p by simp show ?thesis using p_sem IH strand_sem_append_stateful[of M D "unlabel (constr' p)" _ I] unfolding h_p constr_Cons unlabel_append by blast next case False hence not_f: "\(f T)" and not_h: "\(h p)" using T_in unfolding g_def h_def p by auto show ?thesis using not_h not_f_T_case[OF not_f p] IH unfolding constr_Cons unlabel_append by auto qed qed simp have **: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t B = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" proof show "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" unfolding Ts(1) B_def constr_def by (induct Ts) (auto simp add: ik\<^sub>s\<^sub>s\<^sub>t_def) show "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t B" using Ts' unfolding Ts(1) B_def proof (induction Ts) case (Cons p Ts) obtain T \ \ \ where p: "p = (T,\,\,\)" by (cases p) simp have T_in: "T \ set P" using Cons.prems unfolding p by fastforce have IH: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr Ts) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr (filter h Ts))" using Cons.IH Cons.prems by auto show ?case proof (cases "T \ set (filter g P)") case True hence h_p: "filter h (p#Ts) = p#filter h Ts" unfolding h_def p by simp show ?thesis using IH unfolding h_p constr_Cons unlabel_append ik\<^sub>s\<^sub>s\<^sub>t_append by blast next case False hence not_f: "\(f T)" and not_h: "\(h p)" using T_in unfolding g_def h_def p by auto show ?thesis using not_h not_f_T_case[OF not_f p] IH unfolding constr_Cons unlabel_append by auto qed qed simp qed show ?thesis using A_att *[of "{}" "{}"] ** strand_sem_stateful_if_sends_deduct strand_sem_append_stateful[of "{}" "{}" _ "unlabel [(l, send\[attack\n\]\)]" I] unfolding constraint_model_def unlabel_append by force qed have 1: "B \ reachable_constraints (filter g P)" using A(1) Ts(2,3) unfolding Ts(1) B_def proof (induction Ts rule: List.rev_induct) case (snoc p Ts) obtain T \ \ \ where p: "p = (T,\,\,\)" by (cases p) simp have constr_p: "constr' p = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" unfolding constr'_def p by fastforce have T_in: "T \ set P" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr Ts))" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr Ts))" using snoc.prems(3) unfolding p by fast+ have "transaction_fresh_subst s t bb" when "transaction_fresh_subst s t aa" "bb \ aa" for s t bb aa using that unfolding transaction_fresh_subst_def by fast have "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr (filter h Ts)) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr Ts)" unfolding constr_def unlabel_def by fastforce hence \': "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr (filter h Ts)))" using \ unfolding transaction_fresh_subst_def by fast have "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr (filter h Ts)) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr Ts)" unfolding constr_def unlabel_def vars\<^sub>s\<^sub>s\<^sub>t_def by auto hence \': "transaction_renaming_subst \ (filter g P) (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (constr (filter h Ts)))" using \ unfolding transaction_renaming_subst_def by auto have IH: "constr (filter h Ts) \ reachable_constraints (filter g P)" using snoc.prems snoc.IH by simp show ?case proof (cases "h p") case True hence h_p: "filter h (Ts@[p]) = filter h Ts@[p]" by fastforce have T_in': "T \ set (filter g P)" using T_in True unfolding h_def p by fastforce show ?thesis using IH reachable_constraints.step[OF IH T_in' \ \' \'] unfolding h_p constr_snoc constr_p by fast next case False thus ?thesis using IH by fastforce qed qed (simp add: constr_def) show ?thesis using 0 1 A(2) unfolding has_attack_def by blast qed lemma reachable_constraints_fv_Value_typed: assumes P: "\T \ set P. admissible_transaction' T" and A: "\ \ reachable_constraints P" and x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" shows "\\<^sub>v x = TAtom Value" proof - obtain T where T: "T \ set P" "\\<^sub>v x \ \\<^sub>v ` fv_transaction T" using x P(1) reachable_constraints_var_types_in_transactions(1)[OF A(1)] admissible_transactionE(2) by blast show ?thesis using T(2) admissible_transactionE(3)[OF bspec[OF P(1) T(1)]] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] by force qed lemma reachable_constraints_fv_Value_const_cases: assumes P: "\T \ set P. admissible_transaction' T" and A: "\ \ reachable_constraints P" and I: "welltyped_constraint_model \ \" and x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" shows "(\n. \ x = Fun (Val n) []) \ (\n. \ x = Fun (PubConst Value n) [])" proof - have x': "\ (\ x) = TAtom Value" "fv (\ x) = {}" "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" using reachable_constraints_fv_Value_typed[OF P A x] I wt_subst_trm''[of \ "Var x"] unfolding welltyped_constraint_model_def constraint_model_def by auto obtain f where f: "arity f = 0" "\ x = Fun f []" using TAtom_term_cases[OF x'(3,1)] x' const_type_inv_wf[of _ _ Value] by (cases "\ x") force+ show ?thesis proof (cases f) case (Fu g) thus ?thesis by (metis f(2) x'(1) \_Fu_simps(4)[of g "[]"]) qed (use f x'(1) in auto) qed lemma reachable_constraints_receive_attack_if_attack'': assumes P: "\T \ set P. admissible_transaction' T" and A: "\ \ reachable_constraints P" and wt_attack: "welltyped_constraint_model \ (\@[(l, send\[attack\n\]\)])" shows "receive\[attack\n\]\ \ set (unlabel \)" proof - have I: "welltyped_constraint_model \ \" using welltyped_constraint_model_prefix wt_attack by blast show ?thesis using wt_attack strand_sem_append_stateful[of "{}" "{}" "unlabel \" "[send\[attack\n\]\]" \] reachable_constraints_receive_attack_if_attack'(2)[OF A(1) P I] unfolding welltyped_constraint_model_def constraint_model_def by simp qed context begin private lemma reachable_constraints_initial_value_transaction_aux: fixes P::"('fun,'atom,'sets,'lbl) prot" and N::"nat set" assumes P: "\T \ set P. admissible_transaction' T" and A: "A \ reachable_constraints P" and P': "\T \ set P. \(l,a) \ set (transaction_strand T). \t. a \ select\t,\k\\<^sub>s\ \ a \ \t in \k\\<^sub>s\ \ a \ \t not in \k\\<^sub>s\ \ a \ delete\t,\k\\<^sub>s\" shows "(l,\ac: t \ s\) \ set A \ (\u. s = \u\\<^sub>s \ u \ k)" (is "?A A \ ?Q A") and "(l,\t not in s\) \ set A \ (\u. s = \u\\<^sub>s \ u \ k)" (is "?B A \ ?Q A") and "(l,delete\t,s\) \ set A \ (\u. s = \u\\<^sub>s \ u \ k)" (is "?C A \ ?Q A") proof - have "(?A A \ ?Q A) \ (?B A \ ?Q A) \ (?C A \ ?Q A)" (is "?D A") using A proof (induction A rule: reachable_constraints.induct) case (step \ T \ \ \) define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" let ?T' = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" note T_adm = bspec[OF P step.hyps(2)] note T_wf = admissible_transaction_is_wellformed_transaction[OF T_adm] note T_P' = bspec[OF P' step.hyps(2)] have 0: "?Q ?T'" when A: "?A ?T'" proof - obtain t' s' where t: "(l, \ac: t' \ s'\) \ set (transaction_strand T)" "t = t' \ \" "s = s' \ \" using A dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(6) subst_lsst_memD(6) by blast obtain u where u: "s' = \u\\<^sub>s" using transaction_selects_are_Value_vars[OF T_wf(1,2), of t' s'] transaction_inset_checks_are_Value_vars[OF T_adm, of t' s'] unlabel_in[OF t(1)] by (cases ac) auto show ?thesis using T_P' t(1,3) unfolding u by (cases ac) auto qed have 1: "?Q ?T'" when B: "?B ?T'" proof - obtain t' s' where t: "(l, \t' not in s'\) \ set (transaction_strand T)" "t = t' \ \" "s = s' \ \" using B dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(7) subst_lsst_memD(9) by blast obtain u where u: "s' = \u\\<^sub>s" using transaction_notinset_checks_are_Value_vars(2)[OF T_adm unlabel_in[OF t(1)]] by fastforce show ?thesis using T_P' t(1,3) unfolding u by auto qed have 2: "?Q ?T'" when C: "?C ?T'" proof - obtain t' s' where t: "(l, delete\t', s'\) \ set (transaction_strand T)" "t = t' \ \" "s = s' \ \" using C dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(5) subst_lsst_memD(5) by blast obtain u where u: "s' = \u\\<^sub>s" using transaction_deletes_are_Value_vars(2)[OF T_wf(1,3) unlabel_in[OF t(1)]] by blast show ?thesis using T_P' t(1,3) unfolding u by auto qed show ?case using 0 1 2 step.IH unfolding \_def by auto qed simp thus "?A A \ ?Q A" "?B A \ ?Q A" "?C A \ ?Q A" by fast+ qed lemma reachable_constraints_initial_value_transaction: fixes P::"('fun,'atom,'sets,'lbl) prot" and N::"nat set" and k A T_upds defines "checks_not_k \ \B. T_upds \ [] \ ( (\l t s. (l,\t in s\) \ set (A@B) \ (\u. s = \u\\<^sub>s \ u \ k)) \ (\l t s. (l,\t not in s\) \ set (A@B) \ (\u. s = \u\\<^sub>s \ u \ k)) \ (\l t s. (l,delete\t,s\) \ set (A@B) \ (\u. s = \u\\<^sub>s \ u \ k)))" assumes P: "\T \ set P. admissible_transaction' T" and A: "A \ reachable_constraints P" and N: "finite N" "\n \ N. \(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" and T: "T \ set P" "Var x \ set T_ts" "\\<^sub>v x = TAtom Value" "fv\<^sub>s\<^sub>e\<^sub>t (set T_ts) = {x}" "\n. \(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t set T_ts)" "T = Transaction (\(). []) [x] [] [] T_upds [(l1,send\T_ts\)]" "T_upds = [] \ (T_upds = [(l2,insert\Var x, \k\\<^sub>s\)] \ (\T \ set P. \(l,a) \ set (transaction_strand T). \t. a \ select\t,\k\\<^sub>s\ \ a \ \t in \k\\<^sub>s\ \ a \ \t not in \k\\<^sub>s\ \ a \ delete\t,\k\\<^sub>s\))" shows "\B. A@B \ reachable_constraints P \ B \ reachable_constraints P \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t B = {} \ (T_upds = [] \ list_all is_Receive (unlabel B)) \ (T_upds \ [] \ list_all (\a. is_Insert a \ is_Receive a) (unlabel B)) \ (\n. Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ Fun (Val n) [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t B) \ (\n. Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \ Fun (Val n) [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t B) \ N = {n. Fun (Val n) [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t B} \ checks_not_k B \ (\l a. (l,a) \ set B \ is_Insert a \ (l = l2 \ (\n. a = insert\Fun (Val n) [],\k\\<^sub>s\)))" (is "\B. A@B \ ?reach P \ B \ ?reach P \ ?Q1 B \ ?Q2 B \ ?Q3 B \ ?Q4 B \ ?Q5 B \ ?Q6 N B \ checks_not_k B \ ?Q7 B") using N proof (induction N rule: finite_induct) case empty define B where "B \ []::('fun,'atom,'sets,'lbl) prot_constr" have 0: "A@B \ ?reach P" "B \ ?reach P" using A unfolding B_def by auto have 1: "?Q1 B" "?Q2 B" "?Q3 B" "?Q4 B" "?Q6 {} B" unfolding B_def by auto have 2: "checks_not_k B" using reachable_constraints_initial_value_transaction_aux[OF P 0(1)] T(7) unfolding checks_not_k_def by presburger have 3: "?Q5 B" "?Q7 B" unfolding B_def by simp_all show ?case using 0 1 2 3 by blast next case (insert n N) obtain B where B: "A@B \ reachable_constraints P" "B \ reachable_constraints P" "?Q1 B" "?Q2 B" "?Q3 B" "?Q4 B" "?Q5 B" "?Q6 N B" "checks_not_k B" "?Q7 B" using insert.IH insert.prems by blast define \ where "\ \ Var::('fun,'atom,'sets,'lbl) prot_subst" define \ where "\ \ Var(x := Fun (Val n) [])::('fun,'atom,'sets,'lbl) prot_subst" have \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B))" proof (unfold transaction_fresh_subst_def; intro conjI) have "subst_range \ = {Fun (Val n) []}" unfolding \_def by simp moreover have "Fun (Val n) [] \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B))" using insert.prems insert.hyps(2) B(7,8) ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset[of "unlabel B"] unfolding unlabel_append trms\<^sub>s\<^sub>s\<^sub>t_append by blast ultimately show "\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B))" by fastforce next show "subst_domain \ = set (transaction_fresh T)" using T(6) unfolding \_def by auto next show "\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction T)" using T(5,7) unfolding \_def T(6) by fastforce qed (force simp add: T(3) \_def wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def)+ hence \': "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" unfolding transaction_fresh_subst_def by fastforce have \: "transaction_decl_subst \ T" using T(6) unfolding \_def transaction_decl_subst_def by fastforce obtain \::"('fun,'atom,'sets,'lbl) prot_subst" where \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B))" unfolding transaction_renaming_subst_def by blast hence \': "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" unfolding transaction_renaming_subst_def by auto define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" define C where "C \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" have \x: "\ x = Fun (Val n) []" unfolding \_def \_def \_def subst_compose by force have "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = []" "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = []" using T(6) unfolding \_def \_def \_def by auto moreover have "(T_upds = [] \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = []) \ (T_upds \ [] \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = [(l2,insert\\ x, \k\\<^sub>s\)])" using subst_lsst_cons[of "(l2,insert\Var x,\k\\<^sub>s\)" "[]" \] T(6,7) unfolding dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst by auto hence "(T_upds = [] \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = []) \ (T_upds \ [] \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = [(l2,insert\Fun (Val n) [],\k\\<^sub>s\)])" - using subst_apply_term.simps(1)[of x \] unfolding \_def \_def \_def by auto + unfolding \_def \_def \_def by (auto simp: subst_compose) moreover have "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = [(l1, receive\T_ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\)]" using subst_lsst_cons[of "(l1, receive\T_ts\)" "[]" \] T(6) unfolding dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst by auto hence "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = [(l1, receive\T_ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\)]" - using subst_apply_term.simps(1)[of x \] by auto + by auto ultimately have C: "(T_upds = [] \ C = [(l1, receive\T_ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\)]) \ (T_upds \ [] \ C = [(l2, insert\Fun (Val n) [],\k\\<^sub>s\), (l1, receive\T_ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\)])" unfolding C_def transaction_dual_subst_unfold by force have C': "Fun (Val n) [] \ set (T_ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" "Fun (Val n) [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t C" using T(2) in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of _ "unlabel C"] C unfolding \_def \_def \_def by (force, force) have "fv (t \ \) = {}" when t: "t \\<^sub>s\<^sub>e\<^sub>t set T_ts" for t proof - have "fv t \ {x}" using t T(4) fv_subset_subterms by blast hence "fv (t \ \ \\<^sub>s \) = {}" unfolding \_def \_def by (induct t) auto thus ?thesis unfolding \_def by (metis subst_ground_ident_compose(2)) qed hence 1: "ground (set (T_ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \))" by auto have 2: "m = n" when m: "Fun (Val m) [] \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t C" for m proof - have "Fun (Val m) [] \\<^sub>s\<^sub>e\<^sub>t set (T_ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using m C in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of _ "unlabel C"] by fastforce hence *: "Fun (Val m) [] \\<^sub>s\<^sub>e\<^sub>t set T_ts \\<^sub>s\<^sub>e\<^sub>t \" by simp show ?thesis using const_subterms_subst_cases[OF *] T(4,5) \x by fastforce qed have C_trms: "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t C \ {Fun (Val n) [],\k\\<^sub>s} \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t C" using C in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of _ "unlabel C"] by fastforce have 3: "m = n" when m: "Fun (Val m) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t C" for m using m 2[of m] C_trms by fastforce have Q1: "?Q1 (B@C)" using B(3) C 1 by auto have Q2: "?Q2 (B@C)" using B(4) C by force have Q3: "?Q3 (B@C)" using B(5) C by force have Q4: "?Q4 (B@C)" using B(6) insert.prems C 2 unfolding unlabel_append ik\<^sub>s\<^sub>s\<^sub>t_append by blast have Q5: "?Q5 (B@C)" using B(7) C' 3 unfolding unlabel_append ik\<^sub>s\<^sub>s\<^sub>t_append trms\<^sub>s\<^sub>s\<^sub>t_append by blast have Q6: "?Q6 (insert n N) (B@C)" using B(8) C' 2 unfolding unlabel_append ik\<^sub>s\<^sub>s\<^sub>t_append by blast have Q7: "?Q7 (B@C)" using B(10) C by fastforce have Q8: "checks_not_k (B@C)" using B(9) C unfolding checks_not_k_def by force have "B@C \ reachable_constraints P" "A@B@C \ reachable_constraints P" using reachable_constraints.step[OF B(1) T(1) \ \ \] reachable_constraints.step[OF B(2) T(1) \ \' \'] unfolding \_def[symmetric] C_def[symmetric] by simp_all thus ?case using Q1 Q2 Q3 Q4 Q5 Q6 Q7 Q8 by blast qed end subsection \Equivalence Between the Symbolic Protocol Model and a Ground Protocol Model\ context begin subsubsection \Intermediate Step: Equivalence to a Ground Protocol Model with Renaming\ private definition "consts_of X = {t. t \\<^sub>s\<^sub>e\<^sub>t X \ (\c. t = Fun c [])}" private fun mk_symb where "mk_symb (\, \, I, T, \) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t((transaction_strand T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" private fun T_symb :: " _ \ ('fun,'atom,'sets,'lbl) prot_constr" where "T_symb w = concat (map mk_symb w)" private definition "narrow \ S = (\x. if x \ S then \ x else Var x)" private fun mk_inv\I where "mk_inv\I n (\, \, I, T) = narrow ((var_rename_inv n) \\<^sub>s I) (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s var_rename n))" private fun inv\I where "inv\I ns w = foldl (\\<^sub>s) Var (map2 mk_inv\I ns w)" private fun mk_I where "mk_I (\, \, I, T, \) = narrow I (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" private fun comb_I where "comb_I w = fold (\\<^sub>s) (map mk_I w) (\x. Fun OccursSec [])" private abbreviation "ground_term t \ ground {t}" private lemma ground_term_def2: "ground_term t \ (fv t = {})" by auto private definition "ground_strand s \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t s = {}" private fun ground_step :: "(_, _) stateful_strand_step \ bool" where "ground_step s \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p s = {}" private fun ground_lstep :: "_ strand_label \ (_, _) stateful_strand_step \ bool" where "ground_lstep (l,s) \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p s = {}" private inductive_set ground_protocol_states_aux:: "('fun,'atom,'sets,'lbl) prot \ (('fun,'atom,'sets,'lbl) prot_terms \ (('fun,'atom,'sets,'lbl) prot_term \ ('fun,'atom,'sets,'lbl) prot_term) set \ _ set \ _ set \ _ list) set" for P::"('fun,'atom,'sets,'lbl) prot" where init: "({},{},{},{},[]) \ ground_protocol_states_aux P" | step: "\(IK,DB,trms,vars,w) \ ground_protocol_states_aux P; T \ set P; transaction_decl_subst \ T; transaction_fresh_subst \ T trms; transaction_renaming_subst \ P vars; A = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \); strand_sem_stateful IK DB (unlabel A) I; interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I; wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I) \ \ (IK \ ((ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t I), dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I DB, trms \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A, vars \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A, w@[(\, \, I, T, \)]) \ ground_protocol_states_aux P" private lemma T_symb_append': " T_symb (w@w') = T_symb w @ T_symb w'" proof (induction w arbitrary: w') case Nil then show ?case by auto next case (Cons a w) then show ?case by auto qed private lemma T_symb_append: "T_symb (w@[(\, \, I, T, \)]) = T_symb w @ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t((transaction_strand T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using T_symb_append'[of w "[(\, \, I, T,\)]"] by auto private lemma ground_step_subst: assumes "ground_step a" shows "a = a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" using assms proof (induction a) case (NegChecks Y F F') then have FY: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set Y = {}" unfolding ground_step.simps unfolding fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p.simps by auto { have "\t s. (t,s) \ set F \ t \ (rm_vars (set Y) \) = t" proof (rule, rule, rule) fix t s assume "(t, s) \ set F" then show "t \ rm_vars (set Y) \ = t" using FY by fastforce qed moreover have "\t s. (t,s) \ set F \ s \ (rm_vars (set Y) \) = s" proof (rule, rule, rule) fix t s assume "(t, s) \ set F" then show "s \ rm_vars (set Y) \ = s" using FY by fastforce qed ultimately have "\f \ set F. f \\<^sub>p (rm_vars (set Y) \) = f" by auto then have "F = F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set Y) \" by (metis (no_types, lifting) map_idI split_cong subst_apply_pairs_def) } moreover from NegChecks have F'Y: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' - set Y = {}" unfolding ground_step.simps unfolding fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p.simps by auto { have "\t s. (t,s) \ set F' \ t \ (rm_vars (set Y) \) = t" proof (rule, rule, rule) fix t s assume "(t, s) \ set F'" then show "t \ rm_vars (set Y) \ = t" using F'Y by fastforce qed moreover have "\t s. (t,s) \ set F' \ s \ (rm_vars (set Y) \) = s" proof (rule, rule, rule) fix t s assume "(t, s) \ set F'" then show "s \ rm_vars (set Y) \ = s" using F'Y by fastforce qed ultimately have "\f \ set F'. f \\<^sub>p (rm_vars (set Y) \) = f" by auto then have "F' = F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set Y) \" by (simp add: map_idI subst_apply_pairs_def) } ultimately show ?case by simp qed (auto simp add: map_idI subst_ground_ident) private lemma ground_lstep_subst: assumes "ground_lstep a" shows "a = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" using assms by (cases a) (auto simp add: ground_step_subst) private lemma subst_apply_term_rm_vars_swap: assumes "\x\fv t - set X. I x = I' x" shows "t \ rm_vars (set X) I = t \ rm_vars (set X) I'" using assms by (induction t) auto private lemma subst_apply_pairs_rm_vars_swap: assumes "\x\\ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r ` set ps) - set X. I x = I' x" shows "ps \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) I = ps \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) I'" proof - have "\p \ set ps. p \\<^sub>p rm_vars (set X) I = p \\<^sub>p rm_vars (set X) I'" proof fix p assume "p \ set ps" obtain t s where "p = (t,s)" by (cases p) auto have "\x\fv t - set X. I x = I' x" by (metis DiffD1 DiffD2 DiffI \p = (t, s)\ \p \ set ps\ assms fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.elims fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_inI(4)) then have "t \ rm_vars (set X) I = t \ rm_vars (set X) I'" using subst_apply_term_rm_vars_swap by blast have "\x\fv s - set X. I x = I' x" by (metis DiffD1 DiffD2 DiffI \p = (t, s)\ \p \ set ps\ assms fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.elims fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_inI(5)) then have "s \ rm_vars (set X) I = s \ rm_vars (set X) I'" using subst_apply_term_rm_vars_swap by blast show "p \\<^sub>p rm_vars (set X) I = p \\<^sub>p rm_vars (set X) I'" using \p = (t, s)\ \s \ rm_vars (set X) I = s \ rm_vars (set X) I'\ \t \ rm_vars (set X) I = t \ rm_vars (set X) I'\ by fastforce qed then show ?thesis unfolding subst_apply_pairs_def by auto qed private lemma subst_apply_stateful_strand_step_swap: assumes "\x\fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p T. I x = I' x" shows "T \\<^sub>s\<^sub>s\<^sub>t\<^sub>p I = T \\<^sub>s\<^sub>s\<^sub>t\<^sub>p I'" using assms proof (induction T) case (Send ts) then show ?case using term_subst_eq by fastforce next case (NegChecks X F G) then have "\x \ \(fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r ` set F) \ \(fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r ` set G) - set X. I x = I' x" by auto then show ?case using subst_apply_pairs_rm_vars_swap[of F] subst_apply_pairs_rm_vars_swap[of G] by auto qed (simp_all add: term_subst_eq_conv) private lemma subst_apply_labeled_stateful_strand_step_swap: assumes "\x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd T). I x = I' x" shows "T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p I = T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p I'" using assms subst_apply_stateful_strand_step_swap by (metis prod.exhaust_sel subst_apply_labeled_stateful_strand_step.simps) private lemma subst_apply_labeled_stateful_strand_swap: assumes "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T. I x = I' x" shows "T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I = T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I'" using assms proof (induction T) case Nil then show ?case by auto next case (Cons a T) then show ?case using subst_apply_labeled_stateful_strand_step_swap by (metis UnCI fv\<^sub>s\<^sub>s\<^sub>t_Cons subst_lsst_cons unlabel_Cons(2)) qed private lemma transaction_renaming_subst_not_in_fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t: fixes \ \ \::"('fun,'atom,'sets,'lbl) prot_subst" and A::"('fun,'atom,'sets,'lbl) prot_constr" assumes "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" assumes "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" shows "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" proof - have 0: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" when x: "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" for x and A::"('fun,'atom,'sets,'lbl) prot_constr" and \ \::"('fun,'atom,'sets,'lbl) prot_subst" proof - have "x \ range_vars \" using x \ transaction_renaming_subst_vars_disj(6) by blast moreover have "subst_domain \ = UNIV" using \ transaction_renaming_subst_is_renaming(4) by blast ultimately show ?thesis using subst_fv_dom_img_subset[of _ \] fv\<^sub>s\<^sub>s\<^sub>t_subst_obtain_var subst_compose unlabel_subst by (metis (no_types, opaque_lifting) subset_iff top_greatest) qed have 1: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" for x and A::"('fun,'atom,'sets,'lbl) prot_constr" and \ \::"('fun,'atom,'sets,'lbl) prot_subst" using \ x 0 by (metis Un_iff vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t) show ?thesis using 1 assms by metis qed private lemma wf_comb_I_Nil: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (comb_I []))" by auto private lemma comb_I_append: "comb_I (w @ [(\, \, I, T, \)]) = (mk_I (\, \, I, T, \) \\<^sub>s (comb_I w))" by auto private lemma reachable_constraints_if_ground_protocol_states_aux: assumes "(IK, DB, trms, vars, w) \ ground_protocol_states_aux P" shows "T_symb w \ reachable_constraints P \ constr_sem_stateful (comb_I w) (unlabel (T_symb w)) \ IK = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((T_symb w) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t (comb_I w)) \ DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel ((T_symb w))) (comb_I w) {} \ trms = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \ vars = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \ interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (comb_I w) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (comb_I w))" using assms proof (induction rule: ground_protocol_states_aux.induct) case init show ?case using wf_comb_I_Nil by auto next case (step IK DB trms vars w T \ \ \ A I) then have step': "T_symb w \ reachable_constraints P" "strand_sem_stateful {} {} (unlabel (T_symb w)) (comb_I w)" "IK = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w)" "DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w) {}" "trms = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w)" "vars = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (comb_I w)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (comb_I w))" by auto define w' where "w' = w @ [(\, \, I, T, \)]" have interps_w: "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w). (comb_I w) x = (comb_I w') x" proof fix x assume "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w)" then have "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" using step(5) transaction_renaming_subst_not_in_fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t unfolding step'(6) by blast then have "mk_I (\, \, I, T, \) x = Var x" unfolding mk_I.simps narrow_def by metis then have "comb_I w x = (mk_I (\, \, I, T, \) \\<^sub>s (comb_I (w))) x" by (simp add: subst_compose) then show "comb_I w x = comb_I w' x" unfolding w'_def by auto qed have interps_T: "\x \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel (mk_symb (\, \, I, T, \))). I x = (comb_I w') x" proof fix x assume "x \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel (mk_symb (\, \, I, T, \)))" then have a: "x \ (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" by (metis fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq mk_symb.simps) have "(comb_I w') x = (mk_I (\, \, I, T, \) \\<^sub>s (comb_I (w))) x" unfolding w'_def by auto also have "... = ((mk_I (\, \, I, T, \)) x) \ comb_I w" unfolding subst_compose by auto also have "... = (narrow I (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) x) \ comb_I w" using a by auto also have "... = (I x) \ comb_I w" by (metis a narrow_def) also have "... = I x" by (metis UNIV_I ground_subst_range_empty_fv step.hyps(8) subst_compose subst_ground_ident_compose(1)) finally show "I x = (comb_I w') x" by auto qed have "T_symb w' \ reachable_constraints P" proof - have "T_symb w \ reachable_constraints P" using step'(1) . moreover have "T \ set P" using step(2) by auto moreover have "transaction_decl_subst \ T" using step(3) by auto moreover have "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w))" using step(4) step'(5) by auto moreover have "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w))" using step(5) step'(6) by auto ultimately have "(T_symb w) @ mk_symb (\, \, I, T, \) \ reachable_constraints P" using reachable_constraints.step[of "T_symb w" P T \ \ \] by auto then show "T_symb w' \ reachable_constraints P" unfolding w'_def by auto qed moreover have "strand_sem_stateful {} {} (unlabel (T_symb w')) (comb_I w')" proof - have "strand_sem_stateful {} {} (unlabel (T_symb w)) (comb_I w')" proof - have "strand_sem_stateful {} {} (unlabel (T_symb w)) (comb_I w)" using step'(2) by auto then show "strand_sem_stateful {} {} (unlabel (T_symb w)) (comb_I w')" using interps_w strand_sem_model_swap by blast qed moreover have "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \\<^sub>s\<^sub>e\<^sub>t comb_I w') (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w') {}) (unlabel (mk_symb (\, \, I, T, \))) (comb_I w')" proof - have "A = (mk_symb (\, \, I, T, \))" unfolding step(6) by auto moreover have "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w)) (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w) {}) (unlabel A) I" using step'(3) step'(4) step.hyps(7) by force moreover { have "\x\fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w)). comb_I w x = comb_I w' x" using interps_w by (metis fv_ik\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t) then have "\x\fv t. comb_I w x = comb_I w' x" when t: "t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w)" for t using t by auto then have "t \ comb_I w' = t \ comb_I w" when t: "t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w)" for t using t term_subst_eq[of t "comb_I w'" "comb_I w"] by metis then have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \\<^sub>s\<^sub>e\<^sub>t comb_I w' = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \\<^sub>s\<^sub>e\<^sub>t comb_I w" by auto also have "... = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w)" by (metis ik\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst) finally have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \\<^sub>s\<^sub>e\<^sub>t comb_I w' = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w)" by auto } moreover { have "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w) {} = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w') {}" by (metis db\<^sub>s\<^sub>s\<^sub>t_subst_swap[OF interps_w] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t empty_set) } ultimately have "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \\<^sub>s\<^sub>e\<^sub>t comb_I w') (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w') {}) (unlabel (mk_symb (\, \, I, T, \))) I" by force then show "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \\<^sub>s\<^sub>e\<^sub>t comb_I w') (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w') {}) (unlabel (mk_symb (\, \, I, T, \))) (comb_I w')" using interps_T strand_sem_model_swap[of "unlabel (mk_symb (\, \, I, T, \))" I "comb_I w'"] by force qed ultimately show "strand_sem_stateful {} {} (unlabel (T_symb w')) (comb_I w')" using strand_sem_append_stateful[ of "{}" "{}" "unlabel (T_symb w)" "unlabel (mk_symb (\, \, I, T, \))" "(comb_I w')"] unfolding w'_def by auto qed moreover have "IK \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I) = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w')" proof - have AI: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I) = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" by (metis ik\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst) have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w') = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w') \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb [(\, \, I, T, \)] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w')" unfolding w'_def by (simp add: subst_lsst_append) also have "... = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w') \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb [(\, \, I, T, \)] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)" by (metis T_symb_append T_symb_append' interps_T mk_symb.simps self_append_conv2 subst_apply_labeled_stateful_strand_swap) also have "... = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb [(\, \, I, T, \)] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)" by (metis interps_w subst_apply_labeled_stateful_strand_swap) also have "... = IK \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)" using step'(3) step.hyps(6) by auto also have "... = IK \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I)" unfolding AI by auto finally show "IK \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I) = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w')" using step'(3) step(6) T_symb.simps mk_symb.simps by auto qed moreover have "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w')) (comb_I w') {}" proof - have "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w) {})" using step'(4) by auto moreover have "... = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))) I (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w) {})" using step(6) by auto moreover have "... = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (mk_symb (\, \, I, T, \))) I (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w) {})" by auto moreover have "... = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (mk_symb (\, \, I, T, \))) (comb_I w') (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w) {})" by (metis (no_types, lifting) db\<^sub>s\<^sub>s\<^sub>t_subst_swap[OF interps_T] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t empty_set) moreover have "... = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (mk_symb (\, \, I, T, \))) (comb_I w') (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w') {})" by (metis (no_types, lifting) db\<^sub>s\<^sub>s\<^sub>t_subst_swap[OF interps_w] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t empty_set) moreover have "... = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w) @ unlabel (mk_symb (\, \, I, T, \))) (comb_I w') {}" using dbupd\<^sub>s\<^sub>s\<^sub>t_append by metis moreover have "... = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel ((T_symb w) @ mk_symb (\, \, I, T, \))) (comb_I w') {}" by auto moreover have "... = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w')) (comb_I w') {}" unfolding w'_def by auto ultimately show "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w')) (comb_I w') {}" by auto qed moreover have "trms \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w')" proof - have "trms \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using step'(5) by auto moreover have "... = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" using step(6) by auto moreover have "... = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb [(\, \, I, T, \)])" by auto moreover have "... = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w @ T_symb [(\, \, I, T, \)])" by auto moreover have "... = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w')" unfolding w'_def by auto ultimately show "trms \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w')" by auto qed moreover have "vars \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w')" proof - have "vars \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using step'(6) by fastforce moreover have "... = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" using step(6) by auto moreover have "... = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb [(\, \, I, T, \)])" by auto moreover have "... = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w @ T_symb [(\, \, I, T, \)])" by auto moreover have "... = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w')" unfolding w'_def by auto ultimately show "vars \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w')" using step(6) by auto qed moreover have interp_comb_I_w': "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (comb_I w')" using interpretation_comp(1) step'(7) unfolding w'_def by auto moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (comb_I w'))" proof fix t assume "t \ subst_range (comb_I w')" then have "\x. x \ subst_domain (comb_I w') \ t = comb_I w' x" by auto then obtain x where "x \ subst_domain (comb_I w')" "t = comb_I w' x" by auto then show "wf\<^sub>t\<^sub>r\<^sub>m t" by (metis (no_types, lifting) w'_def interp_comb_I_w' comb_I_append ground_subst_dom_iff_img mk_I.simps narrow_def step'(8) step.hyps(8) step.hyps(9) subst_compose_def wf_trm_Var wf_trm_subst) qed ultimately show ?case unfolding w'_def by auto qed private lemma ground_protocol_states_aux_if_reachable_constraints: assumes "A \ reachable_constraints P" assumes "constr_sem_stateful I (unlabel A)" assumes "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" shows "\w. (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I, dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I {}, trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A, vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A, w) \ ground_protocol_states_aux P" using assms proof (induction rule: reachable_constraints.induct) case init then show ?case using ground_protocol_states_aux.init by auto next case (step \ T \ \ \) have "\w. (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t I, dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) I {}, trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \, vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \, w) \ ground_protocol_states_aux P" by (metis local.step(6,7,8,9) step.IH strand_sem_append_stateful unlabel_append) then obtain w where w_p: "(ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t I, dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) I {}, trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \, vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \, w) \ ground_protocol_states_aux P" by auto define w' where "w' = w@[(\, \, I, T, \)]" define \' where "\' = \@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" let ?T = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" have "T \ set P" using step.hyps(2) . moreover have "transaction_decl_subst \ T" using step.hyps(3) . moreover have "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using step.hyps(4) . moreover have "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using step.hyps(5) . moreover have "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t I) (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) I {}) ?T I" using step(7) strand_sem_append_stateful[of "{}" "{}" "unlabel \" ?T I] by auto moreover have "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" using assms(3) . moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" using step.prems(3) by fastforce ultimately have "((ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t I) \ (ik\<^sub>s\<^sub>s\<^sub>t ?T \\<^sub>s\<^sub>e\<^sub>t I), dbupd\<^sub>s\<^sub>s\<^sub>t ?T I (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) I {}), trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ trms\<^sub>s\<^sub>s\<^sub>t ?T, vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ vars\<^sub>s\<^sub>s\<^sub>t ?T, w@[(\, \, I, T, \)]) \ ground_protocol_states_aux P" using ground_protocol_states_aux.step[ OF w_p, of T \ \ \ "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" I] by metis moreover have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>s\<^sub>e\<^sub>t I = (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t I) \ (ik\<^sub>s\<^sub>s\<^sub>t ?T \\<^sub>s\<^sub>e\<^sub>t I)" unfolding \'_def by auto moreover have "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \') I {} = dbupd\<^sub>s\<^sub>s\<^sub>t ?T I (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) I {})" unfolding \'_def by (simp add: dbupd\<^sub>s\<^sub>s\<^sub>t_append) moreover have "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ trms\<^sub>s\<^sub>s\<^sub>t ?T" unfolding \'_def by auto moreover have "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ vars\<^sub>s\<^sub>s\<^sub>t ?T" unfolding \'_def by auto ultimately have "(ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>s\<^sub>e\<^sub>t I, dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \') I {}, trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \', vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \', w') \ ground_protocol_states_aux P" using w'_def by auto then show ?case unfolding \'_def w'_def by auto qed private lemma protocol_model_equivalence_aux1: "{(IK, DB) | IK DB. \w trms vars. (IK, DB, trms, vars, w) \ ground_protocol_states_aux P} = {(ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I), dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I {}) | A I. A \ reachable_constraints P \ strand_sem_stateful {} {} (unlabel A) I \ interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)}" proof (rule; rule; rule) fix IK DB assume "(IK, DB) \ {(IK, DB) | IK DB. \w trms vars. (IK, DB, trms, vars, w) \ ground_protocol_states_aux P}" then have "\w trms vars. (IK, DB, trms, vars, w) \ ground_protocol_states_aux P" by auto then obtain w trms vars where "(IK, DB, trms, vars, w) \ ground_protocol_states_aux P" by auto then have reachable: "T_symb w \ reachable_constraints P" "strand_sem_stateful {} {} (unlabel (T_symb w)) (comb_I w)" "IK = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w \\<^sub>l\<^sub>s\<^sub>s\<^sub>t comb_I w)" "DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w) {}" "trms = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w)" "vars = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (comb_I w)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (comb_I w))" using reachable_constraints_if_ground_protocol_states_aux[of IK DB trms vars w P] by auto then have "IK = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w \\<^sub>l\<^sub>s\<^sub>s\<^sub>t (comb_I w))" "DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (T_symb w)) (comb_I w) {}" "T_symb w \ reachable_constraints P" "strand_sem_stateful {} {} (unlabel (T_symb w)) (comb_I w)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (comb_I w) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (comb_I w))" by auto then show "\A I. (IK, DB) = (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I), dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I {}) \ A \ reachable_constraints P \ strand_sem_stateful {} {} (unlabel A) I \ interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" by blast next fix IK DB assume "(IK, DB) \ {(ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I), dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I {}) | A I. A \ reachable_constraints P \ strand_sem_stateful {} {} (unlabel A) I \ interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)}" then obtain A I where A_I_p: "IK = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)" "DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I {}" "A \ reachable_constraints P" "strand_sem_stateful {} {} (unlabel A) I" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" by auto then have "\w. (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I, dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I {}, trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A, vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A, w) \ ground_protocol_states_aux P" using ground_protocol_states_aux_if_reachable_constraints[of A P I] by auto then have "\w. (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I, DB, trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A, vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A, w) \ ground_protocol_states_aux P" using A_I_p by blast then have "\w. (ik\<^sub>s\<^sub>s\<^sub>t (unlabel A \\<^sub>s\<^sub>s\<^sub>t I), DB, trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A, vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A, w) \ ground_protocol_states_aux P" by (simp add: ik\<^sub>s\<^sub>s\<^sub>t_subst) then have "(\w trms vars. (IK, DB, trms, vars, w) \ ground_protocol_states_aux P)" by (metis (no_types) A_I_p(1) unlabel_subst) then show "\IK' DB'. (IK, DB) = (IK', DB') \ (\w trms vars. (IK', DB', trms, vars, w) \ ground_protocol_states_aux P)" by auto qed subsubsection \The Protocol Model Equivalence Proof\ private lemma subst_ground_term_ident: assumes "ground_term t" shows "t \ I = t" using assms by (simp add: subst_ground_ident) private lemma subst_comp_rm_vars_eq: fixes \ :: "('fun,'atom,'sets,'lbl) prot_subst" fixes \ :: "('fun,'atom,'sets,'lbl) prot_subst" fixes I :: "('fun,'atom,'sets,'lbl) prot_subst" assumes "subst_domain \ = set X \ ground (subst_range \)" shows "(\ \\<^sub>s \) = (\ \\<^sub>s (rm_vars (set X) \))" proof (rule ext) fix x show "(\ \\<^sub>s \) x = (\ \\<^sub>s rm_vars (set X) \) x" proof (cases "x \ set X") case True have gt: "ground_term (\ x)" using True assms by auto have "(\ \\<^sub>s \) x = (\ x) \ \" using subst_compose by metis also have "... = \ x" using gt subst_ground_term_ident by blast also have "... = (\ x) \ (rm_vars (set X) \)" using gt subst_ground_term_ident by fastforce also have "... = (\ \\<^sub>s (rm_vars (set X) \)) x" using subst_compose by metis ultimately show ?thesis by auto next case False have delta_x: "\ x = Var x" using False assms by blast have "(rm_vars (set X) \) x = \ x" using False by auto have "(\ \\<^sub>s \) x = (\ x) \ \" using subst_compose by metis also have "... = (Var x) \ \" using delta_x by presburger also have "... = (Var x) \ (rm_vars (set X) \)" using False by force also have "... = (\ x) \ (rm_vars (set X) \)" using delta_x by presburger also have "... = (\ \\<^sub>s (rm_vars (set X) \)) x" using subst_compose by metis finally show ?thesis by auto qed qed private lemma subst_comp_rm_vars_commute: assumes "\x\set X. \y. \ y \ Var x" assumes "subst_range \ \ range Var" assumes "subst_domain \ = set X" assumes "ground (subst_range \)" shows "(\ \\<^sub>s (rm_vars (set X) \)) = (rm_vars (set X) \ \\<^sub>s \)" proof (rule ext) fix x show "(\ \\<^sub>s rm_vars (set X) \) x = (rm_vars (set X) \ \\<^sub>s \) x" proof (cases "x \ set X") case True then have gt: "ground_term (\ x)" using True assms(3,4) by auto have "(\ \\<^sub>s (rm_vars (set X) \)) x = \ x \ rm_vars (set X) \" by (simp add: subst_compose) also have "... = \ x" using gt by auto also have "... = ((rm_vars (set X) \) x) \ \" by (simp add: True) also have "... = (rm_vars (set X) \ \\<^sub>s \) x" by (simp add: subst_compose) finally show ?thesis . next case False have \_x: "\ x = Var x" using False assms(3) by blast obtain y where y_p: "\ x = Var y" by (meson assms(2) image_iff subsetD subst_imgI) then have "y \ set X" using assms(1) by blast then show ?thesis using assms(3,4) subst_domI False \_x y_p by (metis (mono_tags, lifting) subst_comp_notin_dom_eq subst_compose) qed qed private lemma negchecks_model_substitution_lemma_1: fixes \ :: "('fun,'atom,'sets,'lbl) prot_subst" fixes I :: "('fun,'atom,'sets,'lbl) prot_subst" assumes "negchecks_model (\ \\<^sub>s I) DB X F F'" assumes "subst_range \ \ range Var" assumes "\x\set X. \y. \ y \ Var x" shows "negchecks_model I DB X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" unfolding negchecks_model_def proof (rule, rule) fix \ :: "('fun,'atom,'sets,'lbl) prot_subst" assume a: "subst_domain \ = set X \ ground (subst_range \)" have "(\(t, s)\set F. t \ \ \\<^sub>s (\ \\<^sub>s I) \ s \ \ \\<^sub>s (\ \\<^sub>s I)) \ (\(t, s)\set F'. (t, s) \\<^sub>p \ \\<^sub>s (\ \\<^sub>s I) \ DB)" using a assms(1) unfolding negchecks_model_def by auto then show "(\(t, s)\set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \). t \ \ \\<^sub>s I \ s \ \ \\<^sub>s I) \ (\(t, s)\set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \). (t, s) \\<^sub>p \ \\<^sub>s I \ DB)" proof assume "\(t, s)\set F. t \ \ \\<^sub>s (\ \\<^sub>s I) \ s \ \ \\<^sub>s (\ \\<^sub>s I)" then obtain t s where t_s_p: "(t, s)\set F" "t \ \ \\<^sub>s (\ \\<^sub>s I) \ s \ \ \\<^sub>s (\ \\<^sub>s I)" by auto from this(2) have "t \ \ \\<^sub>s ((rm_vars (set X) \) \\<^sub>s I) \ s \ \ \\<^sub>s ((rm_vars (set X) \) \\<^sub>s I)" using assms(3) a using subst_comp_rm_vars_eq[of \ X \] subst_compose_assoc by (metis (no_types, lifting)) then have "t \ (rm_vars (set X) \) \\<^sub>s ( \ \\<^sub>s I) \ s \ (rm_vars (set X) \) \\<^sub>s (\ \\<^sub>s I)" using subst_comp_rm_vars_commute[of X \ \, OF assms(3) assms(2)] a by (metis (no_types, lifting) subst_compose_assoc[symmetric]) then have "t \ (rm_vars (set X) \) \ ( \ \\<^sub>s I) \ s \ (rm_vars (set X) \) \ (\ \\<^sub>s I)" by auto moreover have "(t \ rm_vars (set X) \, s \ rm_vars (set X) \) \ set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" using subst_apply_pairs_pset_subst t_s_p(1) by fastforce ultimately have "\(t, s)\set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \). t \ \ \\<^sub>s I \ s \ \ \\<^sub>s I" by auto then show ?thesis by auto next assume "\(t, s)\set F'. (t, s) \\<^sub>p \ \\<^sub>s (\ \\<^sub>s I) \ DB" then obtain t s where t_s_p: "(t, s) \ set F'" "(t, s) \\<^sub>p \ \\<^sub>s (\ \\<^sub>s I) \ DB" by auto from this(2) have "(t, s) \\<^sub>p \ \\<^sub>s (rm_vars (set X) \ \\<^sub>s I) \ DB" using assms(3) a subst_comp_rm_vars_eq[OF a] by (metis (no_types, lifting) case_prod_conv subst_subst_compose) then have "(t, s) \\<^sub>p rm_vars (set X) \ \\<^sub>s (\ \\<^sub>s I) \ DB" using a subst_comp_rm_vars_commute[of X \ \, OF assms(3) assms(2)] by (metis (no_types, lifting) case_prod_conv subst_compose_assoc) then have "(t \ rm_vars (set X) \, s \ rm_vars (set X) \) \\<^sub>p \ \\<^sub>s I \ DB" by auto moreover have "(t \ rm_vars (set X) \, s \ rm_vars (set X) \) \ set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" using t_s_p(1) subst_apply_pairs_pset_subst by fastforce ultimately have "(\(t, s) \ set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \). (t, s) \\<^sub>p \ \\<^sub>s I \ DB)" by auto then show ?thesis by auto qed qed private lemma negchecks_model_substitution_lemma_2: fixes \ :: "('fun,'atom,'sets,'lbl) prot_subst" fixes I :: "('fun,'atom,'sets,'lbl) prot_subst" assumes "negchecks_model I DB X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" assumes "subst_range \ \ range Var" assumes "\x\set X. \y. \ y \ Var x" shows "negchecks_model (\ \\<^sub>s I) DB X F F'" unfolding negchecks_model_def proof (rule, rule) fix \ :: "('fun,'atom,'sets,'lbl) prot_subst" assume a: "subst_domain \ = set X \ ground (subst_range \)" have "(\(t, s)\set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \). t \ \ \\<^sub>s I \ s \ \ \\<^sub>s (I)) \ (\(t, s)\set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \). (t, s) \\<^sub>p \ \\<^sub>s I \ DB)" using a assms(1)unfolding negchecks_model_def by auto then show "(\(t, s)\set F. t \ \ \\<^sub>s (\ \\<^sub>s I) \ s \ \ \\<^sub>s (\ \\<^sub>s I)) \ (\(t, s)\set F'. (t, s) \\<^sub>p \ \\<^sub>s (\ \\<^sub>s I) \ DB)" proof assume "\(t, s)\set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \). t \ \ \\<^sub>s I \ s \ \ \\<^sub>s (I)" then obtain t s where t_s_p: "(t, s) \ set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" "t \ \ \\<^sub>s I \ s \ \ \\<^sub>s I" by auto then have "\t' s'. t = t' \ rm_vars (set X) \ \ s = s' \ rm_vars (set X) \ \ (t',s') \ set F" unfolding subst_apply_pairs_def by auto then obtain t' s' where t'_s'_p: "t = t' \ rm_vars (set X) \" "s = s' \ rm_vars (set X) \" "(t',s') \ set F" by auto then have "t' \ rm_vars (set X) \ \ \ \\<^sub>s I \ s' \ rm_vars (set X) \ \ \ \\<^sub>s I" using t_s_p by auto then have "t' \ \ \ rm_vars (set X) \ \\<^sub>s I \ s' \ \ \ rm_vars (set X) \ \\<^sub>s I" using a subst_comp_rm_vars_commute[OF assms(3,2)] by (metis (no_types, lifting) subst_subst) then have "t' \ \ \ \ \\<^sub>s I \ s' \ \ \ \ \\<^sub>s I" using subst_comp_rm_vars_eq[OF a] by (metis (no_types, lifting) subst_subst) moreover from t_s_p(1) have "(t', s') \ set F" using subst_apply_pairs_pset_subst t'_s'_p by fastforce ultimately have "\(t, s)\set F. t \ \ \\<^sub>s (\ \\<^sub>s I) \ s \ \ \\<^sub>s (\ \\<^sub>s I)" by auto then show ?thesis by auto next assume "\(t, s)\set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \). (t, s) \\<^sub>p \ \\<^sub>s I \ DB" then obtain t s where t_s_p: "(t, s) \ set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" "(t \ \ \\<^sub>s I, s \ \ \\<^sub>s I) \ DB" by auto then have "\t' s'. t = t' \ rm_vars (set X) \ \ s = s' \ rm_vars (set X) \ \ (t',s') \ set F'" unfolding subst_apply_pairs_def by auto then obtain t' s' where t'_s'_p: "t = t' \ rm_vars (set X) \" "s = s' \ rm_vars (set X) \" "(t',s') \ set F'" by auto then have "(t' \ rm_vars (set X) \ \ \ \\<^sub>s I, s' \ rm_vars (set X) \ \ \ \\<^sub>s I) \ DB" using t_s_p by auto then have "(t' \ \ \ rm_vars (set X) \ \\<^sub>s I, s' \ \ \ rm_vars (set X) \ \\<^sub>s I) \ DB" using a subst_comp_rm_vars_commute[OF assms(3,2)] by (metis (no_types, lifting) subst_subst) then have "(t' \ \ \ \ \\<^sub>s I , s' \ \ \ \ \\<^sub>s I) \ DB" using subst_comp_rm_vars_eq[OF a] by (metis (no_types, lifting) subst_subst) moreover from t_s_p(1) have "(t', s') \ set F'" using subst_apply_pairs_pset_subst t'_s'_p by fastforce ultimately have "\(t, s)\set F'. (t, s) \\<^sub>p \ \\<^sub>s (\ \\<^sub>s I) \ DB" by auto then show ?thesis by auto qed qed private lemma negchecks_model_substitution_lemma: fixes \ :: "('fun,'atom,'sets,'lbl) prot_subst" fixes I :: "('fun,'atom,'sets,'lbl) prot_subst" assumes "subst_range \ \ range Var" assumes "\x\set X. \y. \ y \ Var x" shows "negchecks_model (\ \\<^sub>s I) DB X F F' \ negchecks_model I DB X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" using assms negchecks_model_substitution_lemma_1[of \ I DB X F F'] negchecks_model_substitution_lemma_2[of I DB X F \ F'] assms by auto private lemma strand_sem_stateful_substitution_lemma: fixes \ :: "('fun,'atom,'sets,'lbl) prot_subst" fixes I :: "('fun,'atom,'sets,'lbl) prot_subst" assumes "subst_range \ \ range Var" assumes "\x \ bvars\<^sub>s\<^sub>s\<^sub>t T. \y. \ y \ Var x" shows "strand_sem_stateful IK DB (T \\<^sub>s\<^sub>s\<^sub>t \) I = strand_sem_stateful IK DB T (\ \\<^sub>s I)" using assms proof (induction T arbitrary: IK DB) case Nil then show ?case by auto next case (Cons a T) then show ?case proof (induction a) case (Receive ts) have "((\x. x \ \ \ I) ` (set ts)) \ IK = ((\t. t \ \) ` set ts \\<^sub>s\<^sub>e\<^sub>t I) \ IK" by blast then show ?case using Receive by (force simp add: subst_sst_cons) next case (NegChecks X F F') have bounds: "\x\bvars\<^sub>s\<^sub>s\<^sub>t T. \y. \ y \ Var x" using NegChecks by auto have "\x\bvars\<^sub>s\<^sub>s\<^sub>t ([\X\\\: F \\: F'\]). \y. \ y \ Var x" using NegChecks by auto then have bounds2: "\x\set X. \y. \ y \ Var x" by simp have "negchecks_model I DB X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ negchecks_model (\ \\<^sub>s I) DB X F F'" using NegChecks.prems(2) bounds2 negchecks_model_substitution_lemma by blast moreover have "strand_sem_stateful IK DB (T \\<^sub>s\<^sub>s\<^sub>t \) I \ strand_sem_stateful IK DB T (\ \\<^sub>s I)" using Cons NegChecks(2) bounds by blast ultimately show ?case by (simp add: subst_sst_cons) qed (force simp add: subst_sst_cons)+ qed private lemma ground_subst_rm_vars_subst_compose_dist: assumes "ground (subst_range \\)" shows "(rm_vars (set X) (\\ \\<^sub>s \)) = (rm_vars (set X) \\ \\<^sub>s rm_vars (set X) \)" proof (rule ext) fix x show "rm_vars (set X) (\\ \\<^sub>s \) x = (rm_vars (set X) \\ \\<^sub>s rm_vars (set X) \) x" proof (cases "x \ set X") case True then show ?thesis by (simp add: subst_compose) next case False note False_outer = False show ?thesis proof (cases "x \ subst_domain \\") case True then show ?thesis by (metis (mono_tags, lifting) False assms ground_subst_range_empty_fv subst_ground_ident_compose(1)) next case False have "\\ x = Var x" using False by blast then show ?thesis using False_outer by (simp add: subst_compose) qed qed qed private lemma stateful_strand_ground_subst_comp: assumes "ground (subst_range \\)" shows "T \\<^sub>s\<^sub>s\<^sub>t \\ \\<^sub>s \ = (T \\<^sub>s\<^sub>s\<^sub>t \\) \\<^sub>s\<^sub>s\<^sub>t \" using assms by (meson disjoint_iff ground_subst_no_var stateful_strand_subst_comp) private lemma labelled_stateful_strand_ground_subst_comp: assumes "ground (subst_range \\)" shows "T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \\ \\<^sub>s \ = (T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \\) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using assms by (metis Int_empty_left ground_range_vars labeled_stateful_strand_subst_comp) private lemma transaction_fresh_subst_ground_subst_range: assumes "transaction_fresh_subst \ T trms" shows "ground (subst_range \)" using assms by (metis range_vars_alt_def transaction_fresh_subst_range_vars_empty) private lemma transaction_decl_subst_ground_subst_range: assumes "transaction_decl_subst \ T" shows "ground (subst_range \)" proof - have \_ground: "\x \ subst_domain \. ground_term (\ x)" using assms transaction_decl_subst_domain transaction_decl_subst_grounds_domain by force show ?thesis proof (rule ccontr) assume "fv\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ {}" then have "\x \ subst_domain \. fv (\ x) \ {}" by auto then obtain x where x_p: "x \ subst_domain \ \ fv (\ x) \ {}" by meson moreover have "ground_term (\ x)" using \_ground x_p by auto ultimately show "False" by auto qed qed private lemma fresh_transaction_decl_subst_ground_subst_range: assumes "transaction_fresh_subst \ T trms" assumes "transaction_decl_subst \ T" shows "ground (subst_range (\ \\<^sub>s \))" proof - have "ground (subst_range \)" using assms transaction_decl_subst_ground_subst_range by blast moreover have "ground (subst_range \)" using assms using transaction_fresh_subst_ground_subst_range by blast ultimately show "ground (subst_range (\ \\<^sub>s \))" by (metis (no_types, opaque_lifting) Diff_iff all_not_in_conv empty_iff empty_subsetI range_vars_alt_def range_vars_subst_compose_subset subset_antisym sup_bot.right_neutral) qed private lemma strand_sem_stateful_substitution_lemma': assumes "transaction_renaming_subst \ P vars" assumes "transaction_fresh_subst \ T trms" assumes "transaction_decl_subst \ T" assumes "finite vars" assumes "T \ set P" shows "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))) I \ strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \))) (\ \\<^sub>s I)" proof - have \_Var: "subst_range \ \ range Var" using assms(1) transaction_renaming_subst_is_renaming(5) by blast have "(\x\vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T). \y. \ y \ Var x)" using assms(4,2) transaction_renaming_subst_vars_transaction_neq assms(1) assms(5) by blast then have "(\x\bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T). \y. \ y \ Var x)" by (metis UnCI vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t) then have T_Vars: "(\x\bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)). \y. \ y \ Var x)" by (metis bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq) have ground_\_\: "ground (subst_range (\ \\<^sub>s \))" using fresh_transaction_decl_subst_ground_subst_range using assms(2) assms(3) by blast from assms(1) ground_\_\ have "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \) = unlabel ((dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using stateful_strand_ground_subst_comp[of _ "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))"] by (simp add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst) then show ?thesis using strand_sem_stateful_substitution_lemma \_Var T_Vars by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst subst_lsst_unlabel) qed inductive_set ground_protocol_states:: "('fun,'atom,'sets,'lbl) prot \ (('fun,'atom,'sets,'lbl) prot_terms \ (('fun,'atom,'sets,'lbl) prot_term \ ('fun,'atom,'sets,'lbl) prot_term) set \ _ set ) set" (* TODO: write up the type nicer *) for P::"('fun,'atom,'sets,'lbl) prot" where init: "({},{},{}) \ ground_protocol_states P" | step: "\(IK,DB,consts) \ ground_protocol_states P; T \ set P; transaction_decl_subst \ T; transaction_fresh_subst \ T consts; A = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \); strand_sem_stateful IK DB (unlabel A) I; interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I; wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I) \ \ (IK \ ((ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t I), dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I DB, consts \ {t. t \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ (\c. t = Fun c [])}) \ ground_protocol_states P" private lemma transaction_fresh_subst_consts_of_iff: "transaction_fresh_subst \ T (consts_of trms) \ transaction_fresh_subst \ T trms" proof (cases "transaction_fresh_subst \ T (consts_of trms) \ transaction_fresh_subst \ T trms") case True then have "\t \ subst_range \. \c. t = Fun c []" unfolding transaction_fresh_subst_def by auto have "(\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t (consts_of trms) \ t \ subterms\<^sub>s\<^sub>e\<^sub>t trms)" proof fix t assume "t \ subst_range \" then obtain c where c_p: "t = Fun c []" using \\t\subst_range \. \c. t = Fun c []\ by blast have "Fun c [] \ subterms\<^sub>s\<^sub>e\<^sub>t (consts_of trms) \ Fun c [] \ subterms\<^sub>s\<^sub>e\<^sub>t trms" unfolding consts_of_def by auto then show "t \ subterms\<^sub>s\<^sub>e\<^sub>t (consts_of trms) \ t \ subterms\<^sub>s\<^sub>e\<^sub>t trms" using c_p by auto qed then show ?thesis using transaction_fresh_subst_def by force next case False then show ?thesis by auto qed private lemma transaction_renaming_subst_inv: assumes "transaction_renaming_subst \ P X " shows "\\inv. \ \\<^sub>s \inv = Var \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \inv)" using var_rename_inv_comp transaction_renaming_subst_def assms subst_apply_term_empty subst_term_eqI by (metis var_rename_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_range(2)) private lemma consts_of_union_distr: "consts_of (trms1 \ trms2) = consts_of trms1 \ consts_of trms2" unfolding consts_of_def by auto private lemma ground_protocol_states_aux_finite_vars: assumes "(IK,DB,trms,vars,w) \ ground_protocol_states_aux P" shows "finite vars" using assms by (induction rule: ground_protocol_states_aux.induct) auto private lemma dbupd\<^sub>s\<^sub>s\<^sub>t_substitution_lemma: "dbupd\<^sub>s\<^sub>s\<^sub>t T (\ \\<^sub>s I) DB = dbupd\<^sub>s\<^sub>s\<^sub>t (T \\<^sub>s\<^sub>s\<^sub>t \) I DB" proof (induction T arbitrary: DB) case Nil then show ?case by auto next case (Cons a T) then show ?case by (induction a) (simp_all add: subst_apply_stateful_strand_def) qed private lemma subst_Var_const_subterm_subst: assumes "subst_range \ \ range Var" shows "Fun c [] \ t \ Fun c [] \ t \ \" using assms proof (induction t) case (Var x) then show ?case by (metis is_Var_def subtermeq_Var_const(1) term.discI(2) var_renaming_is_Fun_iff) next case (Fun f ts) then show ?case by auto qed private lemma subst_Var_consts_of: assumes "subst_range \ \ range Var" shows "consts_of T = consts_of (T \\<^sub>s\<^sub>e\<^sub>t \)" proof (rule antisym; rule subsetI) fix x assume "x \ consts_of T" then obtain t c where t_c_p: "t \ T \ x \ t \ x = Fun c []" unfolding consts_of_def by auto moreover have "x \ t \ \" using t_c_p by (meson assms subst_Var_const_subterm_subst) ultimately show "x \ consts_of (T \\<^sub>s\<^sub>e\<^sub>t \)" unfolding consts_of_def by auto next fix x assume "x \ consts_of (T \\<^sub>s\<^sub>e\<^sub>t \)" then obtain t c where t_c_p: "t \ T \ x \ t \ \ \ x = Fun c []" unfolding consts_of_def by auto moreover have "x \ t" using t_c_p by (meson assms subst_Var_const_subterm_subst) ultimately show "x \ consts_of T" unfolding consts_of_def by auto qed private lemma fst_set_subst_apply_set: "fst ` set F \\<^sub>s\<^sub>e\<^sub>t \ = fst ` (set F \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" by (induction F) auto private lemma snd_set_subst_apply_set: "snd ` set F \\<^sub>s\<^sub>e\<^sub>t \ = snd ` (set F \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" by (induction F) auto private lemma trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_fst_snd: "trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F = fst ` set F \ snd ` set F" by (auto simp add: rev_image_eqI) private lemma consts_of_trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_range_Var: assumes "subst_range \ \ range Var" shows "consts_of (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) = consts_of (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \))" using assms proof (induction a) case (NegChecks X F F') have \_subs_rng_Var: "subst_range (rm_vars (set X) \) \ range Var" using assms by auto have "consts_of (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) = consts_of (fst ` set F \ snd ` set F)" using trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_fst_snd by metis also have "... = consts_of (fst ` set F) \ consts_of (snd ` set F)" using consts_of_union_distr by blast also have "... = consts_of ((fst ` set F) \\<^sub>s\<^sub>e\<^sub>t rm_vars (set X) \) \ consts_of ((snd ` set F) \\<^sub>s\<^sub>e\<^sub>t rm_vars (set X) \)" using \_subs_rng_Var subst_Var_consts_of[of "rm_vars (set X) \" "fst ` set F"] subst_Var_consts_of[of "rm_vars (set X) \" "snd ` set F"] by auto also have "... = consts_of (((fst ` set F) \\<^sub>s\<^sub>e\<^sub>t rm_vars (set X) \) \ ((snd ` set F) \\<^sub>s\<^sub>e\<^sub>t rm_vars (set X) \))" using consts_of_union_distr by auto also have "... = consts_of (fst ` set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ snd ` set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \))" unfolding subst_apply_pairs_def fst_set_subst_apply_set snd_set_subst_apply_set by simp also have "... = consts_of (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \))" using trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_fst_snd[of "F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \"] by metis finally have 1: "consts_of (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) = consts_of (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \))" by auto have "consts_of (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F') = consts_of (fst ` set F' \ snd ` set F')" using trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_fst_snd by metis also have "... = consts_of (fst ` set F') \ consts_of (snd ` set F')" using consts_of_union_distr by blast also have "... = consts_of ((fst ` set F') \\<^sub>s\<^sub>e\<^sub>t rm_vars (set X) \) \ consts_of ((snd ` set F') \\<^sub>s\<^sub>e\<^sub>t rm_vars (set X) \)" using subst_Var_consts_of[of "rm_vars (set X) \" "fst ` set F'"] \_subs_rng_Var subst_Var_consts_of[of "rm_vars (set X) \" "snd ` set F'"] by auto also have "... = consts_of ((fst ` set F' \\<^sub>s\<^sub>e\<^sub>t rm_vars (set X) \) \ (snd ` set F' \\<^sub>s\<^sub>e\<^sub>t rm_vars (set X) \))" using consts_of_union_distr by auto also have "... = consts_of (fst ` set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ snd ` set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \))" unfolding subst_apply_pairs_def fst_set_subst_apply_set snd_set_subst_apply_set by simp also have "... = consts_of (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \))" using trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_fst_snd[of "F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \"] by metis finally have 2: "consts_of (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F') = consts_of (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \))" by auto show ?case using 1 2 by (simp add: consts_of_union_distr) qed (use subst_Var_consts_of[of _ "{_, _}", OF assms] subst_Var_consts_of[OF assms] in auto) private lemma consts_of_trms\<^sub>s\<^sub>s\<^sub>t_range_Var: assumes "subst_range \ \ range Var" shows "consts_of (trms\<^sub>s\<^sub>s\<^sub>t T) = consts_of (trms\<^sub>s\<^sub>s\<^sub>t (T \\<^sub>s\<^sub>s\<^sub>t \))" proof (induction T) case Nil then show ?case by auto next case (Cons a T) have "consts_of (trms\<^sub>s\<^sub>s\<^sub>t (a # T)) = consts_of (trms\<^sub>s\<^sub>s\<^sub>t [a] \ trms\<^sub>s\<^sub>s\<^sub>t T)" by simp also have "... = consts_of (trms\<^sub>s\<^sub>s\<^sub>t [a]) \ consts_of (trms\<^sub>s\<^sub>s\<^sub>t T)" using consts_of_union_distr by simp also have "... = consts_of (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) \ consts_of (trms\<^sub>s\<^sub>s\<^sub>t T)" by simp also have "... = consts_of (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)) \ consts_of (trms\<^sub>s\<^sub>s\<^sub>t T)" using consts_of_trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_range_Var[OF assms] by simp also have "... = consts_of (trms\<^sub>s\<^sub>s\<^sub>t ([a] \\<^sub>s\<^sub>s\<^sub>t \)) \ consts_of (trms\<^sub>s\<^sub>s\<^sub>t T)" by (simp add: subst_apply_stateful_strand_def) also have "... = consts_of (trms\<^sub>s\<^sub>s\<^sub>t ([a] \\<^sub>s\<^sub>s\<^sub>t \)) \ consts_of (trms\<^sub>s\<^sub>s\<^sub>t (T \\<^sub>s\<^sub>s\<^sub>t \))" using local.Cons by simp also have "... = consts_of (trms\<^sub>s\<^sub>s\<^sub>t (a # T \\<^sub>s\<^sub>s\<^sub>t \))" by (simp add: consts_of_union_distr subst_sst_cons) finally show ?case by simp qed private lemma consts_of_trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t_range_Var: assumes "subst_range \ \ range Var" shows "consts_of (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T) = consts_of (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using consts_of_trms\<^sub>s\<^sub>s\<^sub>t_range_Var[of \ "unlabel T"] by (metis assms unlabel_subst) private lemma transaction_renaming_subst_range: assumes "transaction_renaming_subst \ P vars" shows "subst_range \ \ range Var" using assms unfolding transaction_renaming_subst_def var_rename_def by auto private lemma protocol_models_equiv3': assumes "(IK,DB,trms,vars,w) \ ground_protocol_states_aux P" shows "(IK,DB, consts_of trms) \ ground_protocol_states P" using assms proof (induction rule: ground_protocol_states_aux.induct) case init then show ?case using ground_protocol_states.init unfolding consts_of_def by force next case (step IK DB trms vars w T \ \ \ A I) have fin_vars: "finite vars" using ground_protocol_states_aux_finite_vars step by auto have ground_\\: "ground (subst_range (\ \\<^sub>s \))" using fresh_transaction_decl_subst_ground_subst_range using step.hyps(3) step.hyps(4) by blast have \_Var: "subst_range \ \ range Var" using step(5) transaction_renaming_subst_range by metis define I' where "I' = \ \\<^sub>s I" define A' where "A' = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" have "(IK, DB, consts_of trms) \ ground_protocol_states P" using step by force moreover have T_in_P: "T \ set P" using step by force moreover have "transaction_decl_subst \ T" using step by force moreover have "transaction_fresh_subst \ T (consts_of trms)" using step transaction_fresh_subst_consts_of_iff by force moreover have "A' = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" using A'_def . moreover have "strand_sem_stateful IK DB (unlabel A') I'" using step(7) step(4) step(5) step(3) T_in_P fin_vars unfolding A'_def I'_def step(6) using strand_sem_stateful_substitution_lemma' by auto moreover have "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I'" using step(8) unfolding I'_def by (meson interpretation_comp(1)) moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I')" using step(9) unfolding I'_def using step.hyps(5) transaction_renaming_subst_range_wf_trms wf_trms_subst_compose by blast ultimately have "(IK \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A' \\<^sub>s\<^sub>e\<^sub>t I'), dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A') I' DB, consts_of trms \ consts_of (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A')) \ ground_protocol_states P" using ground_protocol_states.step[of IK DB "consts_of trms" P T \ \ A' I'] unfolding consts_of_def by blast moreover have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A' \\<^sub>s\<^sub>e\<^sub>t I' = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" proof - have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A' \\<^sub>s\<^sub>e\<^sub>t I' = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A' \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s I" unfolding A'_def I'_def step(6) by auto also have "... = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((A' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)" unfolding unlabel_subst[symmetric] ik\<^sub>s\<^sub>s\<^sub>t_subst by auto also have "... = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)" unfolding A'_def step(6) using labelled_stateful_strand_ground_subst_comp[of _ "transaction_strand T", OF ground_\\] by (simp add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst) also have "... = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" by (metis ik\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst) finally show "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A' \\<^sub>s\<^sub>e\<^sub>t I' = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" by auto qed moreover have "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A') I' DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I DB" proof - have "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A') I' DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A') (\ \\<^sub>s I) DB" unfolding A'_def I'_def step(6) using step by auto also have "... = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A' \\<^sub>s\<^sub>s\<^sub>t \) I DB" using dbupd\<^sub>s\<^sub>s\<^sub>t_substitution_lemma by metis also have "... = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I DB" unfolding A'_def step(6) using stateful_strand_ground_subst_comp[of _ "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))"] ground_\\ by (simp add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_unlabel) finally show "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A') I' DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I DB" by auto qed moreover have "consts_of (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A') = consts_of (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" by (metis (no_types, lifting) A'_def \_Var consts_of_trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t_range_Var ground_\\ labelled_stateful_strand_ground_subst_comp step.hyps(6) trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq) ultimately show ?case using consts_of_union_distr by metis qed private lemma protocol_models_equiv4': assumes "(IK, DB, csts) \ ground_protocol_states P" shows "\trms w vars. (IK,DB,trms,vars,w) \ ground_protocol_states_aux P \ csts = consts_of trms \ vars = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w)" using assms proof (induction rule: ground_protocol_states.induct) case init have "({}, {}, {}, {}, []) \ ground_protocol_states_aux P" using ground_protocol_states_aux.init by blast moreover have "{} = consts_of {}" unfolding consts_of_def by auto moreover have "{} = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb [])" by auto ultimately show ?case by metis next case (step IK DB "consts" T \ \ A I) then obtain trms w vars where trms_w_vars_p: "(IK, DB, trms, vars, w) \ ground_protocol_states_aux P" "consts = consts_of trms" "vars = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w)" by auto have "\\. transaction_renaming_subst \ P vars" unfolding transaction_renaming_subst_def by blast then obtain \ :: "('fun,'atom,'sets,'lbl) prot_subst" where \_p: "transaction_renaming_subst \ P vars" by blast then obtain \inv where \inv_p: "\ \\<^sub>s \inv = Var \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \inv)" using transaction_renaming_subst_inv[of \ P vars] by auto define A' where "A' = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define I' where "I' = \inv \\<^sub>s I" define trms' where "trms' = trms \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A'" define vars' where "vars' = vars \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A'" define w' where "w' = w @ [(\, \, I', T, \)]" define IK' where "IK' = IK \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I)" define DB' where "DB' = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I DB" have P_state: "(IK', DB' , trms', vars', w') \ ground_protocol_states_aux P" proof - have 1: "(IK, DB, trms, vars, w) \ ground_protocol_states_aux P" using \(IK, DB, trms, vars, w) \ ground_protocol_states_aux P\ by blast moreover have "T \ set P" using step(2) . moreover have "transaction_decl_subst \ T" using step(3) . moreover have fresh_\: "transaction_fresh_subst \ T trms" using step(4) trms_w_vars_p(2) using transaction_fresh_subst_consts_of_iff by auto moreover have "transaction_renaming_subst \ P vars" using \_p . moreover have "A' = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" unfolding A'_def by auto moreover have "strand_sem_stateful IK DB (unlabel A') I'" proof - have fin_vars: "finite vars" using 1 ground_protocol_states_aux_finite_vars by blast show "strand_sem_stateful IK DB (unlabel A') I'" using step(6) strand_sem_stateful_substitution_lemma'[OF \_p fresh_\ step(3) fin_vars] step(2) unfolding A'_def step(5) by (metis (no_types, lifting) I'_def \inv_p subst_compose_assoc var_comp(2)) qed moreover have "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I'" using step(7) by (simp add: I'_def interpretation_substI subst_compose) moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I')" using I'_def \inv_p step.hyps(8) wf_trms_subst_compose by blast moreover have "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A') I' DB" proof - have "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I DB = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) (\ \\<^sub>s \inv \\<^sub>s I) DB" by (simp add: \inv_p) also have "... = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A') (\inv \\<^sub>s I) DB" unfolding A'_def step(5) by (metis (no_types, lifting) dbupd\<^sub>s\<^sub>s\<^sub>t_substitution_lemma dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_unlabel subst_compose_assoc) also have "... = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A') I' DB" unfolding A'_def I'_def by auto finally show ?thesis by auto qed moreover have "IK \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I) = IK \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A' \\<^sub>s\<^sub>e\<^sub>t I')" proof - have "IK \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I) = IK \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t (\ \\<^sub>s \inv) \\<^sub>s I)" using \inv_p by auto also have "... = IK \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A' \\<^sub>s\<^sub>e\<^sub>t I')" unfolding A'_def step(5) unlabel_subst[symmetric] ik\<^sub>s\<^sub>s\<^sub>t_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst I'_def by auto finally show ?thesis by auto qed ultimately show "(IK', DB' , trms', vars', w') \ ground_protocol_states_aux P" using ground_protocol_states_aux.step[of IK DB trms vars w P T \ \ \ A' I'] unfolding trms'_def vars'_def w'_def IK'_def DB'_def by auto qed moreover have "consts \ consts_of (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) = consts_of trms'" proof - have \_Var: "subst_range \ \ range Var" using \_p transaction_renaming_subst_range by blast have ground_\\: "ground (subst_range (\ \\<^sub>s \))" using fresh_transaction_decl_subst_ground_subst_range using step.hyps(3) step.hyps(4) by blast have "consts \ consts_of (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) = (consts_of trms) \ consts_of (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" using trms_w_vars_p(2) by blast also have "... = (consts_of trms) \ consts_of (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using consts_of_trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t_range_Var[of \, OF \_Var, of A] by auto also have "... = (consts_of trms) \ consts_of (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A')" using step(5) A'_def ground_\\ using labelled_stateful_strand_ground_subst_comp[of _ "(dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))"] by (simp add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst) also have "... = consts_of (trms \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A')" using consts_of_union_distr by blast also have "... = consts_of trms'" unfolding trms'_def by auto finally show "?thesis" by auto qed moreover have "vars' = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T_symb w')" using P_state reachable_constraints_if_ground_protocol_states_aux by auto ultimately show ?case unfolding DB'_def IK'_def consts_of_def[symmetric] by metis qed private lemma protocol_model_equivalence_aux2: "{(IK, DB) | IK DB. \csts. (IK, DB, csts) \ ground_protocol_states P} = {(IK, DB) | IK DB. \w trms vars. (IK, DB, trms, vars, w) \ ground_protocol_states_aux P}" using protocol_models_equiv4' protocol_models_equiv3' by meson theorem protocol_model_equivalence: "{(IK, DB) | IK DB. \csts. (IK, DB, csts) \ ground_protocol_states P} = {(ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I), dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I {}) | A I. A \ reachable_constraints P \ strand_sem_stateful {} {} (unlabel A) I \ interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)}" using protocol_model_equivalence_aux2 protocol_model_equivalence_aux1 by auto end end end diff --git a/thys/Automated_Stateful_Protocol_Verification/Stateful_Protocol_Verification.thy b/thys/Automated_Stateful_Protocol_Verification/Stateful_Protocol_Verification.thy --- a/thys/Automated_Stateful_Protocol_Verification/Stateful_Protocol_Verification.thy +++ b/thys/Automated_Stateful_Protocol_Verification/Stateful_Protocol_Verification.thy @@ -1,11393 +1,11391 @@ (* Title: Stateful_Protocol_Verification.thy Author: Andreas Viktor Hess, DTU Author: Sebastian A. Mödersheim, DTU Author: Achim D. Brucker, University of Exeter Author: Anders Schlichtkrull, DTU SPDX-License-Identifier: BSD-3-Clause *) section\Stateful Protocol Verification\ theory Stateful_Protocol_Verification imports Stateful_Protocol_Model Term_Implication begin subsection \Fixed-Point Intruder Deduction Lemma\ context stateful_protocol_model begin abbreviation pubval_terms::"('fun,'atom,'sets,'lbl) prot_terms" where "pubval_terms \ {t. \f \ funs_term t. is_PubConstValue f}" abbreviation abs_terms::"('fun,'atom,'sets,'lbl) prot_terms" where "abs_terms \ {t. \f \ funs_term t. is_Abs f}" definition intruder_deduct_GSMP:: "[('fun,'atom,'sets,'lbl) prot_terms, ('fun,'atom,'sets,'lbl) prot_terms, ('fun,'atom,'sets,'lbl) prot_term] \ bool" ("\_;_\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P _" 50) where "\M; T\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t \ intruder_deduct_restricted M (\t. t \ GSMP T - (pubval_terms \ abs_terms)) t" lemma intruder_deduct_GSMP_induct[consumes 1, case_names AxiomH ComposeH DecomposeH]: assumes "\M; T\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t" "\t. t \ M \ P M t" "\S f. \length S = arity f; public f; \s. s \ set S \ \M; T\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P s; \s. s \ set S \ P M s; Fun f S \ GSMP T - (pubval_terms \ abs_terms) \ \ P M (Fun f S)" "\t K T' t\<^sub>i. \\M; T\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t; P M t; Ana t = (K, T'); \k. k \ set K \ \M; T\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P k; \k. k \ set K \ P M k; t\<^sub>i \ set T'\ \ P M t\<^sub>i" shows "P M t" proof - let ?Q = "\t. t \ GSMP T - (pubval_terms \ abs_terms)" show ?thesis using intruder_deduct_restricted_induct[of M ?Q t "\M Q t. P M t"] assms unfolding intruder_deduct_GSMP_def by blast qed lemma pubval_terms_subst: assumes "t \ \ \ pubval_terms" "\ ` fv t \ pubval_terms = {}" shows "t \ pubval_terms" using assms(1,2) proof (induction t) case (Fun f T) let ?P = "\f. is_PubConstValue f" from Fun show ?case proof (cases "?P f") case False then obtain t where t: "t \ set T" "t \ \ \ pubval_terms" using Fun.prems by auto hence "\ ` fv t \ pubval_terms = {}" using Fun.prems(2) by auto thus ?thesis using Fun.IH[OF t] t(1) by auto qed force qed simp lemma abs_terms_subst: assumes "t \ \ \ abs_terms" "\ ` fv t \ abs_terms = {}" shows "t \ abs_terms" using assms(1,2) proof (induction t) case (Fun f T) let ?P = "\f. is_Abs f" from Fun show ?case proof (cases "?P f") case False then obtain t where t: "t \ set T" "t \ \ \ abs_terms" using Fun.prems by auto hence "\ ` fv t \ abs_terms = {}" using Fun.prems(2) by auto thus ?thesis using Fun.IH[OF t] t(1) by auto qed force qed simp lemma pubval_terms_subst': assumes "t \ \ \ pubval_terms" "\n. PubConst Value n \ \(funs_term ` (\ ` fv t))" shows "t \ pubval_terms" proof - have False when fs: "f \ funs_term s" "s \ subterms\<^sub>s\<^sub>e\<^sub>t (\ ` fv t)" "is_PubConstValue f" for f s proof - obtain T where T: "Fun f T \ subterms s" using funs_term_Fun_subterm[OF fs(1)] by moura hence "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (\ ` fv t)" using fs(2) in_subterms_subset_Union by blast thus ?thesis using assms(2) funs_term_Fun_subterm'[of f T] fs(3) unfolding is_PubConstValue_def by (cases f) force+ qed thus ?thesis using pubval_terms_subst[OF assms(1)] by auto qed lemma abs_terms_subst': assumes "t \ \ \ abs_terms" "\n. Abs n \ \(funs_term ` (\ ` fv t))" shows "t \ abs_terms" proof - have "\is_Abs f" when fs: "f \ funs_term s" "s \ subterms\<^sub>s\<^sub>e\<^sub>t (\ ` fv t)" for f s proof - obtain T where T: "Fun f T \ subterms s" using funs_term_Fun_subterm[OF fs(1)] by moura hence "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (\ ` fv t)" using fs(2) in_subterms_subset_Union by blast thus ?thesis using assms(2) funs_term_Fun_subterm'[of f T] by (cases f) auto qed thus ?thesis using abs_terms_subst[OF assms(1)] by force qed lemma pubval_terms_subst_range_disj: "subst_range \ \ pubval_terms = {} \ \ ` fv t \ pubval_terms = {}" proof (induction t) case (Var x) thus ?case by (cases "x \ subst_domain \") auto qed auto lemma abs_terms_subst_range_disj: "subst_range \ \ abs_terms = {} \ \ ` fv t \ abs_terms = {}" proof (induction t) case (Var x) thus ?case by (cases "x \ subst_domain \") auto qed auto lemma pubval_terms_subst_range_comp: assumes "subst_range \ \ pubval_terms = {}" "subst_range \ \ pubval_terms = {}" shows "subst_range (\ \\<^sub>s \) \ pubval_terms = {}" proof - { fix t f assume t: "t \ subst_range (\ \\<^sub>s \)" "f \ funs_term t" "is_PubConstValue f" then obtain x where x: "(\ \\<^sub>s \) x = t" by auto have "\ x \ pubval_terms" using assms(1) by (cases "\ x \ subst_range \") force+ hence "(\ \\<^sub>s \) x \ pubval_terms" using assms(2) pubval_terms_subst[of "\ x" \] pubval_terms_subst_range_disj by (metis (mono_tags, lifting) subst_compose_def) hence False using t(2,3) x by blast } thus ?thesis by fast qed lemma pubval_terms_subst_range_comp': assumes "(\ ` X) \ pubval_terms = {}" "(\ ` fv\<^sub>s\<^sub>e\<^sub>t (\ ` X)) \ pubval_terms = {}" shows "((\ \\<^sub>s \) ` X) \ pubval_terms = {}" proof - { fix t f assume t: "t \ (\ \\<^sub>s \) ` X" "f \ funs_term t" "is_PubConstValue f" then obtain x where x: "(\ \\<^sub>s \) x = t" "x \ X" by auto have "\ x \ pubval_terms" using assms(1) x(2) by force moreover have "fv (\ x) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` X)" using x(2) by (auto simp add: fv_subset) hence "\ ` fv (\ x) \ pubval_terms = {}" using assms(2) by auto ultimately have "(\ \\<^sub>s \) x \ pubval_terms" using pubval_terms_subst[of "\ x" \] by (metis (mono_tags, lifting) subst_compose_def) hence False using t(2,3) x by blast } thus ?thesis by fast qed lemma abs_terms_subst_range_comp: assumes "subst_range \ \ abs_terms = {}" "subst_range \ \ abs_terms = {}" shows "subst_range (\ \\<^sub>s \) \ abs_terms = {}" proof - { fix t f assume t: "t \ subst_range (\ \\<^sub>s \)" "f \ funs_term t" "is_Abs f" then obtain x where x: "(\ \\<^sub>s \) x = t" by auto have "\ x \ abs_terms" using assms(1) by (cases "\ x \ subst_range \") force+ hence "(\ \\<^sub>s \) x \ abs_terms" using assms(2) abs_terms_subst[of "\ x" \] abs_terms_subst_range_disj by (metis (mono_tags, lifting) subst_compose_def) hence False using t(2,3) x by blast } thus ?thesis by fast qed lemma abs_terms_subst_range_comp': assumes "(\ ` X) \ abs_terms = {}" "(\ ` fv\<^sub>s\<^sub>e\<^sub>t (\ ` X)) \ abs_terms = {}" shows "((\ \\<^sub>s \) ` X) \ abs_terms = {}" proof - { fix t f assume t: "t \ (\ \\<^sub>s \) ` X" "f \ funs_term t" "is_Abs f" then obtain x where x: "(\ \\<^sub>s \) x = t" "x \ X" by auto have "\ x \ abs_terms" using assms(1) x(2) by force moreover have "fv (\ x) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` X)" using x(2) by (auto simp add: fv_subset) hence "\ ` fv (\ x) \ abs_terms = {}" using assms(2) by auto ultimately have "(\ \\<^sub>s \) x \ abs_terms" using abs_terms_subst[of "\ x" \] by (metis (mono_tags, lifting) subst_compose_def) hence False using t(2,3) x by blast } thus ?thesis by fast qed context begin private lemma Ana_abs_aux1: fixes \::"(('fun,'atom,'sets,'lbl) prot_fun, nat, ('fun,'atom,'sets,'lbl) prot_var) gsubst" and \::"nat \ 'sets set" assumes "Ana\<^sub>f f = (K,T)" shows "(K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (\n. \ n \\<^sub>\ \)" proof - { fix k assume "k \ set K" hence "k \ subterms\<^sub>s\<^sub>e\<^sub>t (set K)" by force hence "k \ \ \\<^sub>\ \ = k \ (\n. \ n \\<^sub>\ \)" proof (induction k) case (Fun g S) have "\s. s \ set S \ s \ \ \\<^sub>\ \ = s \ (\n. \ n \\<^sub>\ \)" using Fun.IH in_subterms_subset_Union[OF Fun.prems] Fun_param_in_subterms[of _ S g] by (meson contra_subsetD) thus ?case using Ana\<^sub>f_assm1_alt[OF assms Fun.prems] by (cases g) auto qed simp } thus ?thesis unfolding abs_apply_list_def by force qed private lemma Ana_abs_aux2: fixes \::"nat \ 'sets set" and K::"(('fun,'atom,'sets,'lbl) prot_fun, nat) term list" and M::"nat list" and T::"('fun,'atom,'sets,'lbl) prot_term list" assumes "\i \ fv\<^sub>s\<^sub>e\<^sub>t (set K) \ set M. i < length T" and "(K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T) \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (\n. T ! n \\<^sub>\ \)" shows "(K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T) \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) (map (\s. s \\<^sub>\ \) T)" (is "?A1 = ?A2") and "(map ((!) T) M) \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = map ((!) (map (\s. s \\<^sub>\ \) T)) M" (is "?B1 = ?B2") proof - have "T ! i \\<^sub>\ \ = (map (\s. s \\<^sub>\ \) T) ! i" when "i \ fv\<^sub>s\<^sub>e\<^sub>t (set K)" for i using that assms(1) by auto hence "k \ (\i. T ! i \\<^sub>\ \) = k \ (\i. (map (\s. s \\<^sub>\ \) T) ! i)" when "k \ set K" for k using that term_subst_eq_conv[of k "\i. T ! i \\<^sub>\ \" "\i. (map (\s. s \\<^sub>\ \) T) ! i"] by auto thus "?A1 = ?A2" using assms(2) by (force simp add: abs_apply_terms_def) have "T ! i \\<^sub>\ \ = map (\s. s \\<^sub>\ \) T ! i" when "i \ set M" for i using that assms(1) by auto thus "?B1 = ?B2" by (force simp add: abs_apply_list_def) qed private lemma Ana_abs_aux1_set: fixes \::"(('fun,'atom,'sets,'lbl) prot_fun, nat, ('fun,'atom,'sets,'lbl) prot_var) gsubst" and \::"nat \ 'sets set" assumes "Ana\<^sub>f f = (K,T)" shows "(set K \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t \ = set K \\<^sub>s\<^sub>e\<^sub>t (\n. \ n \\<^sub>\ \)" proof - { fix k assume "k \ set K" hence "k \ subterms\<^sub>s\<^sub>e\<^sub>t (set K)" by force hence "k \ \ \\<^sub>\ \ = k \ (\n. \ n \\<^sub>\ \)" proof (induction k) case (Fun g S) have "\s. s \ set S \ s \ \ \\<^sub>\ \ = s \ (\n. \ n \\<^sub>\ \)" using Fun.IH in_subterms_subset_Union[OF Fun.prems] Fun_param_in_subterms[of _ S g] by (meson contra_subsetD) thus ?case using Ana\<^sub>f_assm1_alt[OF assms Fun.prems] by (cases g) auto qed simp } thus ?thesis unfolding abs_apply_terms_def by force qed private lemma Ana_abs_aux2_set: fixes \::"nat \ 'sets set" and K::"(('fun,'atom,'sets,'lbl) prot_fun, nat) terms" and M::"nat set" and T::"('fun,'atom,'sets,'lbl) prot_term list" assumes "\i \ fv\<^sub>s\<^sub>e\<^sub>t K \ M. i < length T" and "(K \\<^sub>s\<^sub>e\<^sub>t (!) T) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t \ = K \\<^sub>s\<^sub>e\<^sub>t (\n. T ! n \\<^sub>\ \)" shows "(K \\<^sub>s\<^sub>e\<^sub>t (!) T) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t \ = K \\<^sub>s\<^sub>e\<^sub>t (!) (map (\s. s \\<^sub>\ \) T)" (is "?A1 = ?A2") and "((!) T ` M) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t \ = (!) (map (\s. s \\<^sub>\ \) T) ` M" (is "?B1 = ?B2") proof - have "T ! i \\<^sub>\ \ = (map (\s. s \\<^sub>\ \) T) ! i" when "i \ fv\<^sub>s\<^sub>e\<^sub>t K" for i using that assms(1) by auto hence "k \ (\i. T ! i \\<^sub>\ \) = k \ (\i. (map (\s. s \\<^sub>\ \) T) ! i)" when "k \ K" for k using that term_subst_eq_conv[of k "\i. T ! i \\<^sub>\ \" "\i. (map (\s. s \\<^sub>\ \) T) ! i"] by auto thus "?A1 = ?A2" using assms(2) by (force simp add: abs_apply_terms_def) have "T ! i \\<^sub>\ \ = map (\s. s \\<^sub>\ \) T ! i" when "i \ M" for i using that assms(1) by auto thus "?B1 = ?B2" by (force simp add: abs_apply_terms_def) qed lemma Ana_abs: fixes t::"('fun,'atom,'sets,'lbl) prot_term" assumes "Ana t = (K, T)" shows "Ana (t \\<^sub>\ \) = (K \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, T \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using assms proof (induction t rule: Ana.induct) case (1 f S) obtain K' T' where *: "Ana\<^sub>f f = (K',T')" by moura show ?case using 1 proof (cases "arity\<^sub>f f = length S \ arity\<^sub>f f > 0") case True hence "K = K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) S" "T = map ((!) S) T'" and **: "arity\<^sub>f f = length (map (\s. s \\<^sub>\ \) S)" "arity\<^sub>f f > 0" using 1 * by auto hence "K \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) (map (\s. s \\<^sub>\ \) S)" "T \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = map ((!) (map (\s. s \\<^sub>\ \) S)) T'" using Ana\<^sub>f_assm2_alt[OF *] Ana_abs_aux2[OF _ Ana_abs_aux1[OF *], of T' S \] unfolding abs_apply_list_def by auto moreover have "Fun (Fu f) S \\<^sub>\ \ = Fun (Fu f) (map (\s. s \\<^sub>\ \) S)" by simp ultimately show ?thesis using Ana_Fu_intro[OF ** *] by metis qed (auto simp add: abs_apply_list_def) qed (simp_all add: abs_apply_list_def) end lemma deduct_FP_if_deduct: fixes M IK FP::"('fun,'atom,'sets,'lbl) prot_terms" assumes IK: "IK \ GSMP M - (pubval_terms \ abs_terms)" "\t \ IK \\<^sub>\\<^sub>s\<^sub>e\<^sub>t \. FP \\<^sub>c t" and t: "IK \ t" "t \ GSMP M - (pubval_terms \ abs_terms)" shows "FP \ t \\<^sub>\ \" proof - let ?P = "\f. \is_PubConstValue f" let ?GSMP = "GSMP M - (pubval_terms \ abs_terms)" have 1: "\m \ IK. m \ ?GSMP" using IK(1) by blast have 2: "\t t'. t \ ?GSMP \ t' \ t \ t' \ ?GSMP" proof (intro allI impI) fix t t' assume t: "t \ ?GSMP" "t' \ t" hence "t' \ GSMP M" using ground_subterm unfolding GSMP_def by auto moreover have "\is_PubConstValue f" when "f \ funs_term t" for f using t(1) that by auto hence "\is_PubConstValue f" when "f \ funs_term t'" for f using that subtermeq_imp_funs_term_subset[OF t(2)] by auto moreover have "\is_Abs f" when "f \ funs_term t" for f using t(1) that by auto hence "\is_Abs f" when "f \ funs_term t'" for f using that subtermeq_imp_funs_term_subset[OF t(2)] by auto ultimately show "t' \ ?GSMP" by simp qed have 3: "\t K T k. t \ ?GSMP \ Ana t = (K, T) \ k \ set K \ k \ ?GSMP" proof (intro allI impI) fix t K T k assume t: "t \ ?GSMP" "Ana t = (K, T)" "k \ set K" hence "k \ GSMP M" using GSMP_Ana_key by blast moreover have "\f \ funs_term t. ?P f" using t(1) by auto with t(2,3) have "\f \ funs_term k. ?P f" proof (induction t arbitrary: k rule: Ana.induct) case 1 thus ?case by (metis Ana_Fu_keys_not_pubval_terms surj_pair) qed auto moreover have "\f \ funs_term t. \is_Abs f" using t(1) by auto with t(2,3) have "\f \ funs_term k. \is_Abs f" proof (induction t arbitrary: k rule: Ana.induct) case 1 thus ?case by (metis Ana_Fu_keys_not_abs_terms surj_pair) qed auto ultimately show "k \ ?GSMP" by simp qed have "\IK; M\ \\<^sub>G\<^sub>S\<^sub>M\<^sub>P t" unfolding intruder_deduct_GSMP_def by (rule restricted_deduct_if_deduct'[OF 1 2 3 t]) thus ?thesis proof (induction t rule: intruder_deduct_GSMP_induct) case (AxiomH t) show ?case using IK(2) abs_in[OF AxiomH.hyps] by force next case (ComposeH T f) have *: "Fun f T \\<^sub>\ \ = Fun f (map (\t. t \\<^sub>\ \) T)" using ComposeH.hyps(2,4) by (cases f) auto have **: "length (map (\t. t \\<^sub>\ \) T) = arity f" using ComposeH.hyps(1) by auto show ?case using intruder_deduct.Compose[OF ** ComposeH.hyps(2)] ComposeH.IH(1) * by auto next case (DecomposeH t K T' t\<^sub>i) have *: "Ana (t \\<^sub>\ \) = (K \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, T' \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Ana_abs[OF DecomposeH.hyps(2)] by metis have **: "t\<^sub>i \\<^sub>\ \ \ set (T' \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using DecomposeH.hyps(4) abs_in abs_list_set_is_set_abs_set[of T'] by auto have ***: "FP \ k" when k: "k \ set (K \\<^sub>\\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" for k proof - obtain k' where k': "k' \ set K" "k = k' \\<^sub>\ \" by (metis (no_types) k abs_apply_terms_def imageE abs_list_set_is_set_abs_set) show "FP \ k" using DecomposeH.IH k' by blast qed show ?case using intruder_deduct.Decompose[OF _ * _ **] DecomposeH.IH(1) ***(1) by blast qed qed end subsection \Computing and Checking Term Implications and Messages\ context stateful_protocol_model begin abbreviation (input) "absc s \ (Fun (Abs s) []::('fun,'atom,'sets,'lbl) prot_term)" fun absdbupd where "absdbupd [] _ a = a" | "absdbupd (insert\Var y, Fun (Set s) T\#D) x a = ( if x = y then absdbupd D x (insert s a) else absdbupd D x a)" | "absdbupd (delete\Var y, Fun (Set s) T\#D) x a = ( if x = y then absdbupd D x (a - {s}) else absdbupd D x a)" | "absdbupd (_#D) x a = absdbupd D x a" lemma absdbupd_cons_cases: "absdbupd (insert\Var x, Fun (Set s) T\#D) x d = absdbupd D x (insert s d)" "absdbupd (delete\Var x, Fun (Set s) T\#D) x d = absdbupd D x (d - {s})" "t \ Var x \ (\s T. u = Fun (Set s) T) \ absdbupd (insert\t,u\#D) x d = absdbupd D x d" "t \ Var x \ (\s T. u = Fun (Set s) T) \ absdbupd (delete\t,u\#D) x d = absdbupd D x d" proof - assume *: "t \ Var x \ (\s T. u = Fun (Set s) T)" let ?P = "absdbupd (insert\t,u\#D) x d = absdbupd D x d" let ?Q = "absdbupd (delete\t,u\#D) x d = absdbupd D x d" { fix y f T assume "t = Fun f T \ u = Var y" hence ?P ?Q by auto } moreover { fix y f T assume "t = Var y" "u = Fun f T" hence ?P using * by (cases f) auto } moreover { fix y f T assume "t = Var y" "u = Fun f T" hence ?Q using * by (cases f) auto } ultimately show ?P ?Q by (metis term.exhaust)+ qed simp_all lemma absdbupd_filter: "absdbupd S x d = absdbupd (filter is_Update S) x d" by (induction S x d rule: absdbupd.induct) simp_all lemma absdbupd_append: "absdbupd (A@B) x d = absdbupd B x (absdbupd A x d)" proof (induction A arbitrary: d) case (Cons a A) thus ?case proof (cases a) case (Insert t u) thus ?thesis proof (cases "t \ Var x \ (\s T. u = Fun (Set s) T)") case False then obtain s T where "t = Var x" "u = Fun (Set s) T" by moura thus ?thesis by (simp add: Insert Cons.IH absdbupd_cons_cases(1)) qed (simp_all add: Cons.IH absdbupd_cons_cases(3)) next case (Delete t u) thus ?thesis proof (cases "t \ Var x \ (\s T. u = Fun (Set s) T)") case False then obtain s T where "t = Var x" "u = Fun (Set s) T" by moura thus ?thesis by (simp add: Delete Cons.IH absdbupd_cons_cases(2)) qed (simp_all add: Cons.IH absdbupd_cons_cases(4)) qed simp_all qed simp lemma absdbupd_wellformed_transaction: assumes T: "wellformed_transaction T" shows "absdbupd (unlabel (transaction_strand T)) = absdbupd (unlabel (transaction_updates T))" proof - define S0 where "S0 \ unlabel (transaction_strand T)" define S1 where "S1 \ unlabel (transaction_receive T)" define S2 where "S2 \ unlabel (transaction_checks T)" define S3 where "S3 \ unlabel (transaction_updates T)" define S4 where "S4 \ unlabel (transaction_send T)" note S_defs = S0_def S1_def S2_def S3_def S4_def have 0: "list_all is_Receive S1" "list_all is_Check_or_Assignment S2" "list_all is_Update S3" "list_all is_Send S4" using T unfolding wellformed_transaction_def S_defs by metis+ have "filter is_Update S1 = []" "filter is_Update S2 = []" "filter is_Update S3 = S3" "filter is_Update S4 = []" using list_all_filter_nil[OF 0(1), of is_Update] list_all_filter_nil[OF 0(2), of is_Update] list_all_filter_eq[OF 0(3)] list_all_filter_nil[OF 0(4), of is_Update] by blast+ moreover have "S0 = S1@S2@S3@S4" unfolding S_defs transaction_strand_def unlabel_def by auto ultimately have "filter is_Update S0 = S3" using filter_append[of is_Update] list_all_append[of is_Update] by simp thus ?thesis using absdbupd_filter[of S0] unfolding S_defs by presburger qed fun abs_substs_set:: "[('fun,'atom,'sets,'lbl) prot_var list, 'sets set list, ('fun,'atom,'sets,'lbl) prot_var \ 'sets set, ('fun,'atom,'sets,'lbl) prot_var \ 'sets set, ('fun,'atom,'sets,'lbl) prot_var \ 'sets set \ bool] \ ((('fun,'atom,'sets,'lbl) prot_var \ 'sets set) list) list" where "abs_substs_set [] _ _ _ _ = [[]]" | "abs_substs_set (x#xs) as posconstrs negconstrs msgconstrs = ( let bs = filter (\a. posconstrs x \ a \ a \ negconstrs x = {} \ msgconstrs x a) as; \ = abs_substs_set xs as posconstrs negconstrs msgconstrs in concat (map (\b. map (\\. (x, b)#\) \) bs))" definition abs_substs_fun:: "[(('fun,'atom,'sets,'lbl) prot_var \ 'sets set) list, ('fun,'atom,'sets,'lbl) prot_var] \ 'sets set" where "abs_substs_fun \ x = (case find (\b. fst b = x) \ of Some (_,a) \ a | None \ {})" lemmas abs_substs_set_induct = abs_substs_set.induct[case_names Nil Cons] fun transaction_poschecks_comp:: "(('fun,'atom,'sets,'lbl) prot_fun, ('fun,'atom,'sets,'lbl) prot_var) stateful_strand \ (('fun,'atom,'sets,'lbl) prot_var \ 'sets set)" where "transaction_poschecks_comp [] = (\_. {})" | "transaction_poschecks_comp (\_: Var x \ Fun (Set s) []\#T) = ( let f = transaction_poschecks_comp T in f(x := insert s (f x)))" | "transaction_poschecks_comp (_#T) = transaction_poschecks_comp T" fun transaction_negchecks_comp:: "(('fun,'atom,'sets,'lbl) prot_fun, ('fun,'atom,'sets,'lbl) prot_var) stateful_strand \ (('fun,'atom,'sets,'lbl) prot_var \ 'sets set)" where "transaction_negchecks_comp [] = (\_. {})" | "transaction_negchecks_comp (\Var x not in Fun (Set s) []\#T) = ( let f = transaction_negchecks_comp T in f(x := insert s (f x)))" | "transaction_negchecks_comp (_#T) = transaction_negchecks_comp T" definition transaction_check_pre where "transaction_check_pre FPT T \ \ let (FP, _, TI) = FPT; C = set (unlabel (transaction_checks T)); xs = fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)); \ = \\ x. if fst x = TAtom Value then (absc \ \) x else Var x in (\x \ set (transaction_fresh T). \ x = {}) \ (\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T). intruder_synth_mod_timpls FP TI (t \ \ \)) \ (\u \ C. (is_InSet u \ ( let x = the_elem_term u; s = the_set_term u in (is_Var x \ is_Fun_Set s) \ the_Set (the_Fun s) \ \ (the_Var x))) \ ((is_NegChecks u \ bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p u = [] \ the_eqs u = [] \ length (the_ins u) = 1) \ ( let x = fst (hd (the_ins u)); s = snd (hd (the_ins u)) in (is_Var x \ is_Fun_Set s) \ the_Set (the_Fun s) \ \ (the_Var x))))" definition transaction_check_post where "transaction_check_post FPT T \ \ let (FP, _, TI) = FPT; xs = fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)); \ = \\ x. if fst x = TAtom Value then (absc \ \) x else Var x; u = \\ x. absdbupd (unlabel (transaction_updates T)) x (\ x) in (\x \ set xs - set (transaction_fresh T). \ x \ u \ x \ List.member TI (\ x, u \ x)) \ (\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T). intruder_synth_mod_timpls FP TI (t \ \ (u \)))" definition fun_point_inter where "fun_point_inter f g \ \x. f x \ g x" definition fun_point_union where "fun_point_union f g \ \x. f x \ g x" definition fun_point_Inter where "fun_point_Inter fs \ \x. \f \ fs. f x" definition fun_point_Union where "fun_point_Union fs \ \x. \f \ fs. f x" definition fun_point_Inter_list where "fun_point_Inter_list fs \ \x. \(set (map (\f. f x) fs))" definition fun_point_Union_list where "fun_point_Union_list fs \ \x. \(set (map (\f. f x) fs))" definition ticl_abs where "ticl_abs TI a \ set (a#map snd (filter (\p. fst p = a) TI))" definition ticl_abss where "ticl_abss TI as \ \a \ as. ticl_abs TI a" lemma fun_point_Inter_set_eq: "fun_point_Inter (set fs) = fun_point_Inter_list fs" unfolding fun_point_Inter_def fun_point_Inter_list_def by simp lemma fun_point_Union_set_eq: "fun_point_Union (set fs) = fun_point_Union_list fs" unfolding fun_point_Union_def fun_point_Union_list_def by simp lemma ticl_abs_refl_in: "x \ ticl_abs TI x" unfolding ticl_abs_def by simp lemma ticl_abs_iff: assumes TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" shows "ticl_abs TI a = {b. (a,b) \ (set TI)\<^sup>*}" proof (intro order_antisym subsetI) fix x assume x: "x \ {b. (a, b) \ (set TI)\<^sup>*}" hence "x = a \ (x \ a \ (a,x) \ (set TI)\<^sup>+)" by (metis mem_Collect_eq rtranclD) moreover have "ticl_abs TI a = {a} \ {b. (a,b) \ set TI}" unfolding ticl_abs_def by force ultimately show "x \ ticl_abs TI a" using TI by blast qed (fastforce simp add: ticl_abs_def) lemma ticl_abs_Inter: assumes xs: "\(ticl_abs TI ` xs) \ {}" and TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" shows "\(ticl_abs TI ` \(ticl_abs TI ` xs)) \ \(ticl_abs TI ` xs)" proof fix x assume x: "x \ \(ticl_abs TI ` \(ticl_abs TI ` xs))" have *: "\(ticl_abs TI ` xs) = {b. \a \ xs. (a,b) \ (set TI)\<^sup>*}" unfolding ticl_abs_iff[OF TI] by blast have "(b,x) \ (set TI)\<^sup>*" when b: "\a \ xs. (a,b) \ (set TI)\<^sup>*" for b using x b unfolding ticl_abs_iff[OF TI] by blast hence "(a,x) \ (set TI)\<^sup>*" when "a \ xs" for a using that xs rtrancl.rtrancl_into_rtrancl[of a _ "(set TI)\<^sup>*" x] unfolding * rtrancl_idemp[of "set TI"] by blast thus "x \ \(ticl_abs TI ` xs)" unfolding * by blast qed function (sequential) match_abss' ::"(('a,'b,'c,'d) prot_fun, 'e) term \ (('a,'b,'c,'d) prot_fun, 'e) term \ ('e \ 'c set set) option" where "match_abss' (Var x) (Fun (Abs a) _) = Some ((\_. {})(x := {a}))" | "match_abss' (Fun f ts) (Fun g ss) = ( if f = g \ length ts = length ss then map_option fun_point_Union_list (those (map2 match_abss' ts ss)) else None)" | "match_abss' _ _ = None" by pat_completeness auto termination proof - let ?m = "measures [size \ fst]" have 0: "wf ?m" by simp show ?thesis apply (standard, use 0 in fast) by (metis (no_types) comp_def fst_conv measures_less Fun_zip_size_lt(1)) qed definition match_abss where "match_abss OCC TI t s \ ( let xs = fv t; OCC' = set OCC; f = \\ x. if x \ xs then \ x else OCC'; g = \\ x. \(ticl_abs TI ` \ x) in case match_abss' t s of Some \ \ let \' = g \ in if \x \ xs. \' x \ {} then Some (f \') else None | None \ None)" lemma match_abss'_Var_inv: assumes \: "match_abss' (Var x) t = Some \" shows "\a ts. t = Fun (Abs a) ts \ \ = (\_. {})(x := {a})" proof - obtain f ts where t: "t = Fun f ts" using \ by (cases t) auto then obtain a where a: "f = Abs a" using \ by (cases f) auto show ?thesis using \ unfolding t a by simp qed lemma match_abss'_Fun_inv: assumes "match_abss' (Fun f ts) (Fun g ss) = Some \" shows "f = g" (is ?A) and "length ts = length ss" (is ?B) and "\\. Some \ = those (map2 match_abss' ts ss) \ \ = fun_point_Union_list \" (is ?C) and "\(t,s) \ set (zip ts ss). \\. match_abss' t s = Some \" (is ?D) proof - note 0 = assms match_abss'.simps(2)[of f ts g ss] option.distinct(1) show ?A by (metis 0) show ?B by (metis 0) show ?C by (metis (no_types, opaque_lifting) 0 map_option_eq_Some) thus ?D using map2_those_Some_case[of match_abss' ts ss] by fastforce qed lemma match_abss'_FunI: assumes \: "\i. i < length T \ match_abss' (U ! i) (T ! i) = Some (\ i)" and T: "length T = length U" shows "match_abss' (Fun f U) (Fun f T) = Some (fun_point_Union_list (map \ [0.. [0.. T those_map2_SomeI by metis ultimately show ?thesis by simp qed lemma match_abss'_Fun_param_subset: assumes "match_abss' (Fun f ts) (Fun g ss) = Some \" and "(t,s) \ set (zip ts ss)" and "match_abss' t s = Some \" shows "\ x \ \ x" proof - obtain \ where \: "those (map2 match_abss' ts ss) = Some \" "\ = fun_point_Union_list \" using match_abss'_Fun_inv[OF assms(1)] by metis have "\ \ set \" using \(1) assms(2-) those_Some_iff[of "map2 match_abss' ts ss" \] by force thus ?thesis using \(2) unfolding fun_point_Union_list_def by auto qed lemma match_abss'_fv_is_nonempty: assumes "match_abss' t s = Some \" and "x \ fv t" shows "\ x \ {}" (is "?P \") using assms proof (induction t s arbitrary: \ rule: match_abss'.induct) case (2 f ts g ss) note prems = "2.prems" note IH = "2.IH" have 0: "\(t,s) \ set (zip ts ss). \\. match_abss' t s = Some \" "f = g" "length ts = length ss" using match_abss'_Fun_inv[OF prems(1)] by simp_all obtain t where t: "t \ set ts" "x \ fv t" using prems(2) by auto then obtain s where s: "s \ set ss" "(t,s) \ set (zip ts ss)" by (meson 0(3) in_set_impl_in_set_zip1 in_set_zipE) then obtain \ where \: "match_abss' t s = Some \" using 0(1) by fast show ?case using IH[OF conjI[OF 0(2,3)] s(2) _ \] t(2) match_abss'_Fun_param_subset[OF prems(1) s(2) \] by auto qed auto lemma match_abss'_nonempty_is_fv: fixes s t::"(('a,'b,'c,'d) prot_fun, 'v) term" assumes "match_abss' s t = Some \" and "\ x \ {}" shows "x \ fv s" using assms proof (induction s t arbitrary: \ rule: match_abss'.induct) case (2 f ts g ss) note prems = "2.prems" note IH = "2.IH" obtain \ where \: "Some \ = those (map2 match_abss' ts ss)" "\ = fun_point_Union_list \" and fg: "f = g" "length ts = length ss" using match_abss'_Fun_inv[OF prems(1)] by fast have "\\ \ set \. \ x \ {}" using fg(2) prems \ unfolding fun_point_Union_list_def by auto then obtain t' s' \ where ts': "(t',s') \ set (zip ts ss)" "match_abss' t' s' = Some \" "\ x \ {}" using those_map2_SomeD[OF \(1)[symmetric]] by blast show ?case using ts'(3) IH[OF conjI[OF fg] ts'(1) _ ts'(2)] set_zip_leftD[OF ts'(1)] by force qed auto lemma match_abss'_Abs_in_funs_term: fixes s t::"(('a,'b,'c,'d) prot_fun, 'v) term" assumes "match_abss' s t = Some \" and "a \ \ x" shows "Abs a \ funs_term t" using assms proof (induction s t arbitrary: a \ rule: match_abss'.induct) case (1 y b ts) show ?case using match_abss'_Var_inv[OF "1.prems"(1)] "1.prems"(2) by (cases "x = y") simp_all next case (2 f ts g ss) note prems = "2.prems" note IH = "2.IH" obtain \ where \: "Some \ = those (map2 match_abss' ts ss)" "\ = fun_point_Union_list \" and fg: "f = g" "length ts = length ss" using match_abss'_Fun_inv[OF prems(1)] by fast obtain t' s' \ where ts': "(t',s') \ set (zip ts ss)" "match_abss' t' s' = Some \" "a \ \ x" using fg(2) prems \ those_map2_SomeD[OF \(1)[symmetric]] unfolding fun_point_Union_list_def by fastforce show ?case using ts'(1) IH[OF conjI[OF fg] ts'(1) _ ts'(2,3)] by (meson set_zip_rightD term.set_intros(2)) qed auto lemma match_abss'_subst_fv_ex_abs: assumes "match_abss' s (s \ \) = Some \" and TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" shows "\x \ fv s. \a ts. \ x = Fun (Abs a) ts \ \ x = {a}" (is "?P s \") using assms(1) proof (induction s "s \ \" arbitrary: \ rule: match_abss'.induct) case (2 f ts g ss) note prems = "2.prems" note hyps = "2.hyps" obtain \ where \: "Some \ = those (map2 match_abss' ts ss)" "\ = fun_point_Union_list \" and fg: "f = g" "length ts = length ss" "ss = ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" and ts: "\(t,s) \ set (zip ts ss). \\. match_abss' t s = Some \" using match_abss'_Fun_inv[OF prems(1)[unfolded hyps(2)[symmetric]]] hyps(2) by fastforce have 0: "those (map (\t. match_abss' t (t \ \)) ts) = Some \" using \(1) map2_map_subst unfolding fg(3) by metis have 1: "\t \ set ts. \\. match_abss' t (t \ \) = Some \" using ts zip_map_subst[of ts \] unfolding fg(3) by simp have 2: "\' \ set \" when t: "t \ set ts" "match_abss' t (t \ \) = Some \'" for t \' using t 0 those_Some_iff[of "map (\t. match_abss' t (t \ \)) ts" \] by force have 3: "?P t \'" "\' x \ {}" when t: "t \ set ts" "x \ fv t" "match_abss' t (t \ \) = Some \'" for t \' x using t hyps(1)[OF conjI[OF fg(1,2)], of "(t, t \ \)" t \'] zip_map_subst[of ts \] match_abss'_fv_is_nonempty[of t "t \ \" \' x] unfolding fg(3) by auto have 4: "\' x = {}" when t: "x \ fv t" "match_abss' t (t \ \) = Some \'" for t \' x by (meson t match_abss'_nonempty_is_fv) show ?case proof fix x assume "x \ fv (Fun f ts)" then obtain t \' where t: "t \ set ts" "x \ fv t" and \': "match_abss' t (t \ \) = Some \'" using 1 by auto then obtain a tsa where a: "\ x = Fun (Abs a) tsa" using 3[OF t \'] by fast have "\'' x = {a} \ \'' x = {}" when "\'' \ set \" for \'' using that a 0 3[of _ x] 4[of x] unfolding those_Some_iff by fastforce thus "\a ts. \ x = Fun (Abs a) ts \ \ x = {a}" using a 2[OF t(1) \'] 3[OF t \'] unfolding \(2) fun_point_Union_list_def by auto qed qed auto lemma match_abss'_subst_disj_nonempty: assumes TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and "match_abss' s (s \ \) = Some \" and "x \ fv s" shows "\(ticl_abs TI ` \ x) \ {} \ (\a tsa. \ x = Fun (Abs a) tsa \ \ x = {a})" (is "?P \") using assms(2,3) proof (induction s "s \ \" arbitrary: \ rule: match_abss'.induct) case (1 x a ts) thus ?case unfolding ticl_abs_def by force next case (2 f ts g ss) note prems = "2.prems" note hyps = "2.hyps" obtain \ where \: "Some \ = those (map2 match_abss' ts ss)" "\ = fun_point_Union_list \" and fg: "f = g" "length ts = length ss" "ss = ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" and ts: "\(t,s) \ set (zip ts ss). \\. match_abss' t s = Some \" using match_abss'_Fun_inv[OF prems(1)[unfolded hyps(2)[symmetric]]] hyps(2) by fastforce define ts' where "ts' \ filter (\t. x \ fv t) ts" define \' where "\' \ map (\t. (t, the (match_abss' t (t \ \)))) ts" define \'' where "\'' \ map (\t. the (match_abss' t (t \ \))) ts'" have 0: "those (map (\t. match_abss' t (t \ \)) ts) = Some \" using \(1) map2_map_subst unfolding fg(3) by metis have 1: "\t \ set ts. \\. match_abss' t (t \ \) = Some \" using ts zip_map_subst[of ts \] unfolding fg(3) by simp have ts_not_nil: "ts \ []" using prems(2) by fastforce hence "\t \ set ts. x \ fv t" using prems(2) by simp then obtain a tsa where a: "\ x = Fun (Abs a) tsa" using 1 match_abss'_subst_fv_ex_abs[OF _ TI, of _ \] by metis hence a': "\' x = {a}" when "t \ set ts" "x \ fv t" "match_abss' t (t \ \) = Some \'" for t \' using that match_abss'_subst_fv_ex_abs[OF _ TI, of _ \] by fastforce have "ts' \ []" using prems(2) unfolding ts'_def by (simp add: filter_empty_conv) hence \''_not_nil: "\'' \ []" unfolding \''_def by simp have 2: "\' \ set \" when t: "t \ set ts" "match_abss' t (t \ \) = Some \'" for t \' using t 0 those_Some_iff[of "map (\t. match_abss' t (t \ \)) ts" \] by force have 3: "?P \'" "\' x \ {}" when t: "t \ set ts'" "match_abss' t (t \ \) = Some \'" for t \' using t hyps(1)[OF conjI[OF fg(1,2)], of "(t, t \ \)" t \'] zip_map_subst[of ts \] match_abss'_fv_is_nonempty[of t "t \ \" \' x] unfolding fg(3) ts'_def by (force, force) have 4: "\' x = {}" when t: "x \ fv t" "match_abss' t (t \ \) = Some \'" for t \' by (meson t match_abss'_nonempty_is_fv) have 5: "\ = map snd \'" using 0 1 unfolding \'_def by (induct ts arbitrary: \) auto have "fun_point_Union_list (map snd \') x = fun_point_Union_list (map snd (filter (\(t,_). x \ fv t) \')) x" using 1 4 unfolding \'_def fun_point_Union_list_def by fastforce hence 6: "fun_point_Union_list \ x = fun_point_Union_list \'' x" using 0 1 4 unfolding 5 \'_def \''_def fun_point_Union_list_def ts'_def by auto have 7: "?P \'" "\' x \ {}" when \': "\' \ set \''" for \' using that 1 3 unfolding \''_def ts'_def by auto have "\' x = {a}" when \': "\' \ set \''" for \' using \' a' 1 unfolding \''_def ts'_def by fastforce hence "fun_point_Union_list \'' x = {b | b \'. \' \ set \'' \ b \ {a}}" using \''_not_nil unfolding fun_point_Union_list_def by auto hence 8: "fun_point_Union_list \'' x = {a}" using \''_not_nil by auto show ?case using 8 a unfolding \(2) 6 ticl_abs_iff[OF TI] by auto qed simp_all lemma match_abssD: fixes OCC TI s defines "f \ (\\ x. if x \ fv s then \ x else set OCC)" and "g \ (\\ x. \(ticl_abs TI ` \ x))" assumes \': "match_abss OCC TI s t = Some \'" shows "\\. match_abss' s t = Some \ \ \' = f (g \) \ (\x \ fv s. \ x \ {} \ f (g \) x \ {}) \ (set OCC \ {} \ (\x. f (g \) x \ {}))" proof - obtain \ where \: "match_abss' s t = Some \" using \' unfolding match_abss_def by force hence "Some \' = (if \x \ fv s. g \ x \ {} then Some (f (g \)) else None)" using \' unfolding match_abss_def f_def g_def Let_def by simp hence "\' = f (g \)" "\x \ fv s. \ x \ {} \ f (g \) x \ {}" by (metis (no_types, lifting) option.inject option.distinct(1), metis (no_types, lifting) f_def option.distinct(1) match_abss'_fv_is_nonempty[OF \]) thus ?thesis using \ unfolding f_def by force qed lemma match_abss_ticl_abs_Inter_subset: assumes TI: "set TI = {(a,b). (a,b) \ (set TI)\<^sup>+ \ a \ b}" and \: "match_abss OCC TI s t = Some \" and x: "x \ fv s" shows "\(ticl_abs TI ` \ x) \ \ x" proof - let ?h1 = "\\ x. if x \ fv s then \ x else set OCC" let ?h2 = "\\ x. \(ticl_abs TI ` \ x)" obtain \' where \': "match_abss' s t = Some \'" "\ = ?h1 (?h2 \')" "\x \ fv s. \' x \ {} \ \ x \ {}" using match_abssD[OF \] by blast have "\ x = \(ticl_abs TI ` \' x)" "\' x \ {}" "\ x \ {}" using x \'(2,3) by auto thus ?thesis using ticl_abs_Inter TI by simp qed lemma match_abss_fv_has_abs: assumes "match_abss OCC TI s t = Some \" and "x \ fv s" shows "\ x \ {}" using assms match_abssD by fast lemma match_abss_OCC_if_not_fv: fixes s t::"(('a,'b,'c,'d) prot_fun, 'v) term" assumes \': "match_abss OCC TI s t = Some \'" and x: "x \ fv s" shows "\' x = set OCC" proof - define f where "f \ \s::(('a,'b,'c,'d) prot_fun, 'v) term. \\ x. if x \ fv s then \ x else set OCC" define g where "g \ \\. \x::'v. \(ticl_abs TI ` \ x)" obtain \ where \: "match_abss' s t = Some \" "\' = f s (g \)" using match_abssD[OF \'] unfolding f_def g_def by blast show ?thesis using x \(2) unfolding f_def by presburger qed inductive synth_abs_substs_constrs_rel for FP OCC TI where SolveNil: "synth_abs_substs_constrs_rel FP OCC TI [] (\_. set OCC)" | SolveCons: "ts \ [] \ \t \ set ts. synth_abs_substs_constrs_rel FP OCC TI [t] (\ t) \ synth_abs_substs_constrs_rel FP OCC TI ts (fun_point_Inter (\ ` set ts))" | SolvePubConst: "arity c = 0 \ public c \ synth_abs_substs_constrs_rel FP OCC TI [Fun c []] (\_. set OCC)" | SolvePrivConstIn: "arity c = 0 \ \public c \ Fun c [] \ set FP \ synth_abs_substs_constrs_rel FP OCC TI [Fun c []] (\_. set OCC)" | SolvePrivConstNotin: "arity c = 0 \ \public c \ Fun c [] \ set FP \ synth_abs_substs_constrs_rel FP OCC TI [Fun c []] (\_. {})" | SolveValueVar: "\ = ((\_. set OCC)(x := ticl_abss TI {a \ set OCC. \a\\<^sub>a\<^sub>b\<^sub>s \ set FP})) \ synth_abs_substs_constrs_rel FP OCC TI [Var x] \" | SolveComposed: "arity f > 0 \ length ts = arity f \ \\. \ \ \ \ (\s \ set FP. match_abss OCC TI (Fun f ts) s = Some \) \ \1 = fun_point_Union \ \ synth_abs_substs_constrs_rel FP OCC TI ts \2 \ synth_abs_substs_constrs_rel FP OCC TI [Fun f ts] (fun_point_union \1 \2)" fun synth_abs_substs_constrs_aux where "synth_abs_substs_constrs_aux FP OCC TI (Var x) = ( (\_. set OCC)(x := ticl_abss TI (set (filter (\a. \a\\<^sub>a\<^sub>b\<^sub>s \ set FP) OCC))))" | "synth_abs_substs_constrs_aux FP OCC TI (Fun f ts) = ( if ts = [] then if \public f \ Fun f ts \ set FP then (\_. {}) else (\_. set OCC) else let \ = map the (filter (\\. \ \ None) (map (match_abss OCC TI (Fun f ts)) FP)); \1 = fun_point_Union_list \; \2 = fun_point_Inter_list (map (synth_abs_substs_constrs_aux FP OCC TI) ts) in fun_point_union \1 \2)" definition synth_abs_substs_constrs where "synth_abs_substs_constrs FPT T \ let (FP,OCC,TI) = FPT; ts = trms_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_receive T)); f = fun_point_Inter_list \ map (synth_abs_substs_constrs_aux FP OCC TI) in if ts = [] then (\_. set OCC) else f ts" (* definition synth_abs_substs_constrs where "synth_abs_substs_constrs FPT T \ let (FP,OCC,TI) = FPT; negsy = Not \ intruder_synth_mod_timpls FP TI; \ = \\ x. let as = \ x in if as \ {} then as else set OCC; C = unlabel (transaction_checks T); poss = transaction_poschecks_comp C; negs = transaction_negchecks_comp C; ts = trms_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_receive T)); f = \t. let \ = map the (filter (\\. \ \ None) (map (match_abss OCC TI t) FP)) in fun_point_Union_list (map \ \); g = \t. if is_Fun t \ args t \ [] then let s = hd (args t) in case fv_list s of [] \ if negsy s then Some (f t) else None | [x] \ let bs = filter (\a. poss x \ a \ a \ negs x = {}) OCC in if list_all (\b. negsy (s \ Var(x := \b\\<^sub>a\<^sub>b\<^sub>s))) bs then Some (f t) else None | _ \ None else None; h = \t. case g t of Some d \ d | None \ synth_abs_substs_constrs_aux FP OCC TI t in if ts = [] then (\_. set OCC) else fun_point_Inter_list (map h ts)" *) (* poss = transaction_poschecks_comp (C A); negs = transaction_negchecks_comp (C A); bs = filter (\a. poss PK \ a \ a \ negs PK = {}) OCC in if list_all (Not \ sy \ s A) bs then Some (map the (filter (\\. \ \ None) (map (match_abss OCC TI (t' A)) FP))) else None *) definition transaction_check_comp:: "[('fun,'atom,'sets,'lbl) prot_var \ 'sets set \ bool, ('fun,'atom,'sets,'lbl) prot_term list \ 'sets set list \ ('sets set \ 'sets set) list, ('fun,'atom,'sets,'lbl) prot_transaction] \ ((('fun,'atom,'sets,'lbl) prot_var \ 'sets set) list) list" where "transaction_check_comp msgcs FPT T \ let (_, OCC, _) = FPT; S = unlabel (transaction_strand T); C = unlabel (transaction_checks T); xs = filter (\x. x \ set (transaction_fresh T) \ fst x = TAtom Value) (fv_list\<^sub>s\<^sub>s\<^sub>t S); posconstrs = transaction_poschecks_comp C; negconstrs = transaction_negchecks_comp C; pre_check = transaction_check_pre FPT T; \ = abs_substs_set xs OCC posconstrs negconstrs msgcs in filter (\\. pre_check (abs_substs_fun \)) \" definition transaction_check':: "[('fun,'atom,'sets,'lbl) prot_var \ 'sets set \ bool, ('fun,'atom,'sets,'lbl) prot_term list \ 'sets set list \ ('sets set \ 'sets set) list, ('fun,'atom,'sets,'lbl) prot_transaction] \ bool" where "transaction_check' msgcs FPT T \ list_all (\\. transaction_check_post FPT T (abs_substs_fun \)) (transaction_check_comp msgcs FPT T)" definition transaction_check:: "[('fun,'atom,'sets,'lbl) prot_term list \ 'sets set list \ ('sets set \ 'sets set) list, ('fun,'atom,'sets,'lbl) prot_transaction] \ bool" where "transaction_check \ transaction_check' (\_ _. True)" definition transaction_check_coverage_rcv:: "[('fun,'atom,'sets,'lbl) prot_term list \ 'sets set list \ ('sets set \ 'sets set) list, ('fun,'atom,'sets,'lbl) prot_transaction] \ bool" where "transaction_check_coverage_rcv FPT T \ let msgcs = synth_abs_substs_constrs FPT T in transaction_check' (\x a. a \ msgcs x) FPT T" lemma abs_subst_fun_cons: "abs_substs_fun ((x,b)#\) = (abs_substs_fun \)(x := b)" unfolding abs_substs_fun_def by fastforce lemma abs_substs_cons: assumes "\ \ set (abs_substs_set xs as poss negs msgcs)" "b \ set as" "poss x \ b" "b \ negs x = {}" "msgcs x b" shows "(x,b)#\ \ set (abs_substs_set (x#xs) as poss negs msgcs)" using assms by auto lemma abs_substs_cons': assumes \: "\ \ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)" and b: "b \ set as" "poss x \ b" "b \ negs x = {}" "msgcs x b" shows "\(x := b) \ abs_substs_fun ` set (abs_substs_set (x#xs) as poss negs msgcs)" proof - obtain \ where \: "\ = abs_substs_fun \" "\ \ set (abs_substs_set xs as poss negs msgcs)" using \ by moura have "abs_substs_fun ((x, b)#\) \ abs_substs_fun ` set (abs_substs_set (x#xs) as poss negs msgcs)" using abs_substs_cons[OF \(2) b] by blast thus ?thesis using \(1) abs_subst_fun_cons[of x b \] by argo qed lemma abs_substs_has_abs: assumes "\x. x \ set xs \ \ x \ set as" and "\x. x \ set xs \ poss x \ \ x" and "\x. x \ set xs \ \ x \ negs x = {}" and "\x. x \ set xs \ msgcs x (\ x)" and "\x. x \ set xs \ \ x = {}" shows "\ \ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)" using assms proof (induction xs arbitrary: \) case (Cons x xs) define \ where "\ \ \y. if y \ set xs then \ y else {}" have "\ \ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)" using Cons.prems Cons.IH by (simp add: \_def) moreover have "\ x \ set as" "poss x \ \ x" "\ x \ negs x = {}" "msgcs x (\ x)" by (simp_all add: Cons.prems(1,2,3,4)) ultimately have 0: "\(x := \ x) \ abs_substs_fun ` set (abs_substs_set (x#xs) as poss negs msgcs)" by (metis abs_substs_cons') have "\ = \(x := \ x)" proof fix y show "\ y = (\(x := \ x)) y" proof (cases "y \ set (x#xs)") case False thus ?thesis using Cons.prems(5) by (fastforce simp add: \_def) qed (auto simp add: \_def) qed thus ?case by (metis 0) qed (auto simp add: abs_substs_fun_def) lemma abs_substs_abss_bounded: assumes "\ \ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)" and "x \ set xs" shows "\ x \ set as" and "poss x \ \ x" and "\ x \ negs x = {}" and "msgcs x (\ x)" using assms proof (induct xs as poss negs msgcs arbitrary: \ rule: abs_substs_set_induct) case (Cons y xs as poss negs msgcs) { case 1 thus ?case using Cons.hyps(1) unfolding abs_substs_fun_def by fastforce } { case 2 thus ?case proof (cases "x = y") case False then obtain \' where \': "\' \ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)" "\' x = \ x" using 2 unfolding abs_substs_fun_def by force moreover have "x \ set xs" using 2(2) False by simp moreover have "\b. b \ set as \ poss y \ b \ b \ negs y = {}" using 2 False by auto ultimately show ?thesis using Cons.hyps(2) by fastforce qed (auto simp add: abs_substs_fun_def) } { case 3 thus ?case proof (cases "x = y") case False then obtain \' where \': "\' \ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)" "\' x = \ x" using 3 unfolding abs_substs_fun_def by force moreover have "x \ set xs" using 3(2) False by simp moreover have "\b. b \ set as \ poss y \ b \ b \ negs y = {}" using 3 False by auto ultimately show ?thesis using Cons.hyps(3) by fastforce qed (auto simp add: abs_substs_fun_def) } { case 4 thus ?case proof (cases "x = y") case False then obtain \' where \': "\' \ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)" "\' x = \ x" using 4 unfolding abs_substs_fun_def by force moreover have "x \ set xs" using 4(2) False by simp moreover have "\b. b \ set as \ poss y \ b \ b \ negs y = {}" using 4 False by auto ultimately show ?thesis using Cons.hyps(4) by fastforce qed (auto simp add: abs_substs_fun_def) } qed (simp_all add: abs_substs_fun_def) lemma abs_substs_abss_bounded': assumes "\ \ abs_substs_fun ` set (abs_substs_set xs as poss negs msgcs)" and "x \ set xs" shows "\ x = {}" using assms unfolding abs_substs_fun_def by (induct xs as poss negs msgcs arbitrary: \ rule: abs_substs_set_induct) (force, fastforce) lemma transaction_poschecks_comp_unfold: "transaction_poschecks_comp C x = {s. \a. \a: Var x \ Fun (Set s) []\ \ set C}" proof (induction C) case (Cons c C) thus ?case proof (cases "\a y s. c = \a: Var y \ Fun (Set s) []\") case True then obtain a y s where c: "c = \a: Var y \ Fun (Set s) []\" by moura define f where "f \ transaction_poschecks_comp C" have "transaction_poschecks_comp (c#C) = f(y := insert s (f y))" using c by (simp add: f_def Let_def) moreover have "f x = {s. \a. \a: Var x \ Fun (Set s) []\ \ set C}" using Cons.IH unfolding f_def by blast ultimately show ?thesis using c by auto next case False hence "transaction_poschecks_comp (c#C) = transaction_poschecks_comp C" (is ?P) using transaction_poschecks_comp.cases[of "c#C" ?P] by force thus ?thesis using False Cons.IH by auto qed qed simp lemma transaction_poschecks_comp_notin_fv_empty: assumes "x \ fv\<^sub>s\<^sub>s\<^sub>t C" shows "transaction_poschecks_comp C x = {}" using assms transaction_poschecks_comp_unfold[of C x] by fastforce lemma transaction_negchecks_comp_unfold: "transaction_negchecks_comp C x = {s. \Var x not in Fun (Set s) []\ \ set C}" proof (induction C) case (Cons c C) thus ?case proof (cases "\y s. c = \Var y not in Fun (Set s) []\") case True then obtain y s where c: "c = \Var y not in Fun (Set s) []\" by moura define f where "f \ transaction_negchecks_comp C" have "transaction_negchecks_comp (c#C) = f(y := insert s (f y))" using c by (simp add: f_def Let_def) moreover have "f x = {s. \Var x not in Fun (Set s) []\ \ set C}" using Cons.IH unfolding f_def by blast ultimately show ?thesis using c by auto next case False hence "transaction_negchecks_comp (c#C) = transaction_negchecks_comp C" (is ?P) using transaction_negchecks_comp.cases[of "c#C" ?P] by force thus ?thesis using False Cons.IH by fastforce qed qed simp lemma transaction_negchecks_comp_notin_fv_empty: assumes "x \ fv\<^sub>s\<^sub>s\<^sub>t C" shows "transaction_negchecks_comp C x = {}" using assms transaction_negchecks_comp_unfold[of C x] by fastforce lemma transaction_check_preI[intro]: fixes T defines "\ \ \\ x. if fst x = TAtom Value then (absc \ \) x else Var x" and "C \ set (unlabel (transaction_checks T))" assumes a0: "\x \ set (transaction_fresh T). \ x = {}" and a1: "\x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value \ \ x \ set OCC" and a2: "\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T). intruder_synth_mod_timpls FP TI (t \ \ \)" and a3: "\a x s. \a: Var x \ Fun (Set s) []\ \ C \ s \ \ x" and a4: "\x s. \Var x not in Fun (Set s) []\ \ C \ s \ \ x" shows "transaction_check_pre (FP, OCC, TI) T \" proof - let ?P = "\u. is_InSet u \ ( let x = the_elem_term u; s = the_set_term u in (is_Var x \ is_Fun_Set s) \ the_Set (the_Fun s) \ \ (the_Var x))" let ?Q = "\u. (is_NegChecks u \ bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p u = [] \ the_eqs u = [] \ length (the_ins u) = 1) \ ( let x = fst (hd (the_ins u)); s = snd (hd (the_ins u)) in (is_Var x \ is_Fun_Set s) \ the_Set (the_Fun s) \ \ (the_Var x))" have 1: "?P u" when u: "u \ C" for u apply (unfold Let_def, intro impI, elim conjE) using u a3 Fun_Set_InSet_iff[of u] by metis have 2: "?Q u" when u: "u \ C" for u apply (unfold Let_def, intro impI, elim conjE) using u a4 Fun_Set_NotInSet_iff[of u] by metis show ?thesis using a0 a1 a2 1 2 fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] unfolding transaction_check_pre_def \_def C_def Let_def by blast qed lemma transaction_check_pre_InSetE: assumes T: "transaction_check_pre FPT T \" and u: "u = \a: Var x \ Fun (Set s) []\" "u \ set (unlabel (transaction_checks T))" shows "s \ \ x" proof - have "is_InSet u \ is_Var (the_elem_term u) \ is_Fun_Set (the_set_term u) \ the_Set (the_Fun (the_set_term u)) \ \ (the_Var (the_elem_term u))" using T u unfolding transaction_check_pre_def Let_def by blast thus ?thesis using Fun_Set_InSet_iff[of u a x s] u by argo qed lemma transaction_check_pre_NotInSetE: assumes T: "transaction_check_pre FPT T \" and u: "u = \Var x not in Fun (Set s) []\" "u \ set (unlabel (transaction_checks T))" shows "s \ \ x" proof - have "is_NegChecks u \ bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p u = [] \ the_eqs u = [] \ length (the_ins u) = 1 \ is_Var (fst (hd (the_ins u))) \ is_Fun_Set (snd (hd (the_ins u))) \ the_Set (the_Fun (snd (hd (the_ins u)))) \ \ (the_Var (fst (hd (the_ins u))))" using T u unfolding transaction_check_pre_def Let_def by blast thus ?thesis using Fun_Set_NotInSet_iff[of u x s] u by argo qed lemma transaction_check_pre_ReceiveE: defines "\ \ \\ x. if fst x = TAtom Value then (absc \ \) x else Var x" assumes T: "transaction_check_pre (FP, OCC, TI) T \" and t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" shows "intruder_synth_mod_timpls FP TI (t \ \ \)" using T t unfolding transaction_check_pre_def Let_def \_def by blast lemma transaction_check_compI[intro]: assumes T: "transaction_check_pre (FP, OCC, TI) T \" and T_adm: "admissible_transaction' T" and x1: "\x. (x \ fv_transaction T - set (transaction_fresh T) \ fst x = TAtom Value) \ \ x \ set OCC \ msgcs x (\ x)" and x2: "\x. (x \ fv_transaction T - set (transaction_fresh T) \ fst x \ TAtom Value) \ \ x = {}" shows "\ \ abs_substs_fun ` set (transaction_check_comp msgcs (FP, OCC, TI) T)" proof - define S where "S \ unlabel (transaction_strand T)" define C where "C \ unlabel (transaction_checks T)" let ?xs = "fv_list\<^sub>s\<^sub>s\<^sub>t S" define poss where "poss \ transaction_poschecks_comp C" define negs where "negs \ transaction_negchecks_comp C" define ys where "ys \ filter (\x. x \ set (transaction_fresh T) \ fst x = TAtom Value) ?xs" have ys: "{x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value} = set ys" using fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of S] unfolding ys_def S_def by force have "\ x \ set OCC" "msgcs x (\ x)" when x: "x \ set ys" for x using x1 x ys by (blast, blast) moreover have "\ x = {}" when x: "x \ set ys" for x using x2 x ys by blast moreover have "poss x \ \ x" when x: "x \ set ys" for x proof - have "s \ \ x" when u: "u = \a: Var x \ Fun (Set s) []\" "u \ set C" for u a s using T u transaction_check_pre_InSetE[of "(FP, OCC, TI)" T \] unfolding C_def by blast thus ?thesis using transaction_poschecks_comp_unfold[of C x] unfolding poss_def by blast qed moreover have "\ x \ negs x = {}" when x: "x \ set ys" for x proof (cases "x \ fv\<^sub>s\<^sub>s\<^sub>t C") case True hence "s \ \ x" when u: "u = \Var x not in Fun (Set s) []\" "u \ set C" for u s using T u transaction_check_pre_NotInSetE[of "(FP, OCC, TI)" T \] unfolding C_def by blast thus ?thesis using transaction_negchecks_comp_unfold[of C x] unfolding negs_def by blast next case False hence "negs x = {}" using x transaction_negchecks_comp_notin_fv_empty unfolding negs_def by blast thus ?thesis by blast qed ultimately have "\ \ abs_substs_fun ` set (abs_substs_set ys OCC poss negs msgcs)" using abs_substs_has_abs[of ys \ OCC poss negs msgcs] by fast thus ?thesis using T unfolding transaction_check_comp_def Let_def S_def C_def ys_def poss_def negs_def by fastforce qed context begin private lemma transaction_check_comp_in_aux: fixes T defines "C \ set (unlabel (transaction_checks T))" assumes T_adm: "admissible_transaction' T" and a1: "\x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value \ (\s. select\Var x, Fun (Set s) []\ \ C \ s \ \ x)" and a2: "\x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value \ (\s. \Var x in Fun (Set s) []\ \ C \ s \ \ x)" and a3: "\x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value \ (\s. \Var x not in Fun (Set s) []\ \ C \ s \ \ x)" shows "\a x s. \a: Var x \ Fun (Set s) []\ \ C \ s \ \ x" (is ?A) and "\x s. \Var x not in Fun (Set s) []\ \ C \ s \ \ x" (is ?B) proof - note * = admissible_transaction_strand_step_cases(2,3)[OF T_adm] have 1: "fst x = TAtom Value" "x \ fv_transaction T - set (transaction_fresh T)" when x: "\a: Var x \ Fun (Set s) []\ \ C" for a x s using * x unfolding C_def by fast+ have 2: "fst x = TAtom Value" "x \ fv_transaction T - set (transaction_fresh T)" when x: "\Var x not in Fun (Set s) []\ \ C" for x s using * x unfolding C_def by fast+ show ?A proof (intro allI impI) fix a x s assume u: "\a: Var x \ Fun (Set s) []\ \ C" thus "s \ \ x" using 1 a1 a2 by (cases a) metis+ qed show ?B proof (intro allI impI) fix x s assume u: "\Var x not in Fun (Set s) []\ \ C" thus "s \ \ x" using 2 a3 by meson qed qed lemma transaction_check_comp_in: fixes T defines "\ \ \\ x. if fst x = TAtom Value then (absc \ \) x else Var x" and "C \ set (unlabel (transaction_checks T))" assumes T_adm: "admissible_transaction' T" and a1: "\x \ set (transaction_fresh T). \ x = {}" and a2: "\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T). intruder_synth_mod_timpls FP TI (t \ \ \)" and a3: "\x \ fv_transaction T - set (transaction_fresh T). \s. select\Var x, Fun (Set s) []\ \ C \ s \ \ x" and a4: "\x \ fv_transaction T - set (transaction_fresh T). \s. \Var x in Fun (Set s) []\ \ C \ s \ \ x" and a5: "\x \ fv_transaction T - set (transaction_fresh T). \s. \Var x not in Fun (Set s) []\ \ C \ s \ \ x" and a6: "\x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value \ \ x \ set OCC" and a7: "\x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value \ msgcs x (\ x)" shows "\\ \ abs_substs_fun ` set (transaction_check_comp msgcs (FP, OCC, TI) T). \x \ fv_transaction T. fst x = TAtom Value \ \ x = \ x" proof - let ?xs = "fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T))" let ?ys = "filter (\x. x \ set (transaction_fresh T)) ?xs" define \' where "\' \ \x. if x \ fv_transaction T - set (transaction_fresh T) \ fst x = TAtom Value then \ x else {}" note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] have \\_Fun: "is_Fun (t \ \ \) \ is_Fun (t \ \ \')" for t unfolding \'_def \_def by (induct t) auto have "\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T). intruder_synth_mod_timpls FP TI (t \ \ \')" proof (intro ballI impI) fix t assume t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" have 1: "intruder_synth_mod_timpls FP TI (t \ \ \)" using t a2 by auto obtain r where r: "r \ set (unlabel (transaction_receive T))" "t \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r" using t by auto hence "\ts. r = receive\ts\ \ t \ set ts" using wellformed_transaction_unlabel_cases(1)[OF T_wf] by fastforce hence 2: "fv t \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" using r by force have "fv t \ fv_transaction T" by (metis (no_types, lifting) 2 transaction_strand_def sst_vars_append_subset(1) unlabel_append subset_Un_eq sup.bounded_iff) moreover have "fv t \ set (transaction_fresh T) = {}" using 2 T_wf vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_receive T)"] unfolding wellformed_transaction_def by fast ultimately have "\ \ x = \ \' x" when "x \ fv t" for x using that unfolding \'_def \_def by fastforce hence 3: "t \ \ \ = t \ \ \'" using term_subst_eq by blast show "intruder_synth_mod_timpls FP TI (t \ \ \')" using 1 3 by simp qed moreover have "\x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value \ (\s. select\Var x, Fun (Set s) []\ \ C \ s \ \' x)" "\x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value \ (\s. \Var x in Fun (Set s) []\ \ C \ s \ \' x)" "\x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value \ (\s. \Var x not in Fun (Set s) []\ \ C \ s \ \' x)" using a3 a4 a5 unfolding \'_def \_def C_def by meson+ hence "\a x s. \a: Var x \ Fun (Set s) []\ \ C \ s \ \' x" "\x s. \Var x not in Fun (Set s) []\ \ C \ s \ \' x" using transaction_check_comp_in_aux[OF T_adm, of \'] unfolding C_def by (fast, fast) ultimately have 4: "transaction_check_pre (FP, OCC, TI) T \'" using a6 transaction_check_preI[of T \' OCC FP TI] unfolding \'_def \_def C_def by simp have 5: "\x \ fv_transaction T. fst x = TAtom Value \ \ x = \' x" using a1 by (auto simp add: \'_def) have 6: "\' \ abs_substs_fun ` set (transaction_check_comp msgcs (FP, OCC, TI) T)" using transaction_check_compI[OF 4 T_adm, of msgcs] a6 a7 unfolding \'_def by auto show ?thesis using 5 6 by blast qed end lemma transaction_check_trivial_case: assumes "transaction_updates T = []" and "transaction_send T = []" shows "transaction_check FPT T" using assms by (simp add: list_all_iff transaction_check_def transaction_check'_def transaction_check_post_def) end subsection \Soundness of the Occurs-Message Transformation\ context stateful_protocol_model begin context begin text \The occurs-message transformation, \add_occurs_msgs\, extends a transaction \T\ with additional message-transmission steps such that the following holds: 1. for each fresh variable \x\ of \T\ the message \occurs (Var x)\ now occurs in a send-step, 2. for each of the remaining free variables \x\ of \T\ the message \occurs (Var x)\ now occurs in a receive-step.\ definition add_occurs_msgs where "add_occurs_msgs T \ let frsh = transaction_fresh T; xs = filter (\x. x \ set frsh) (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T))); f = map (\x. occurs (Var x)); g = \C. if xs = [] then C else \\, receive\f xs\\#C; h = \F. if frsh = [] then F else if F \ [] \ fst (hd F) = \ \ is_Send (snd (hd F)) then \\,send\f frsh@the_msgs (snd (hd F))\\#tl F else \\,send\f frsh\\#F in case T of Transaction A B C D E F \ Transaction A B (g C) D E (h F)" private fun rm_occurs_msgs_constr where "rm_occurs_msgs_constr [] = []" | "rm_occurs_msgs_constr ((l,receive\ts\)#A) = ( if \t. occurs t \ set ts then if \t \ set ts. \s. t \ occurs s then (l,receive\filter (\t. \s. t \ occurs s) ts\)#rm_occurs_msgs_constr A else rm_occurs_msgs_constr A else (l,receive\ts\)#rm_occurs_msgs_constr A)" | "rm_occurs_msgs_constr ((l,send\ts\)#A) = ( if \t. occurs t \ set ts then if \t \ set ts. \s. t \ occurs s then (l,send\filter (\t. \s. t \ occurs s) ts\)#rm_occurs_msgs_constr A else rm_occurs_msgs_constr A else (l,send\ts\)#rm_occurs_msgs_constr A)" | "rm_occurs_msgs_constr (a#A) = a#rm_occurs_msgs_constr A" private lemma add_occurs_msgs_cases: fixes T C frsh xs f defines "T' \ add_occurs_msgs T" and "frsh \ transaction_fresh T" and "xs \ filter (\x. x \ set frsh) (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)))" and "xs' \ fv_transaction T - set frsh" and "f \ map (\x. occurs (Var x))" and "C' \ if xs = [] then C else \\, receive\f xs\\#C" and "ts' \ f frsh" assumes T: "T = Transaction A B C D E F" shows "F = \\,send\ts\\#F' \ T' = Transaction A B C' D E (\\,send\ts'@ts\\#F')" (is "?A ts F' \ ?A' ts F'") and "\ts' F'. F = \\,send\ts'\\#F' \ frsh \ [] \ T' = Transaction A B C' D E (\\,send\ts'\\#F)" (is "?B \ ?B' \ ?B''") and "frsh = [] \ T' = Transaction A B C' D E F" (is "?C \ ?C'") and "transaction_decl T' = transaction_decl T" and "transaction_fresh T' = transaction_fresh T" and "xs = [] \ transaction_receive T' = transaction_receive T" and "xs \ [] \ transaction_receive T' = \\,receive\f xs\\#transaction_receive T" and "transaction_checks T' = transaction_checks T" and "transaction_updates T' = transaction_updates T" and "transaction_send T = \\,send\ts\\#F' \ transaction_send T' = \\,send\ts'@ts\\#F'" (is "?D ts F' \ ?D' ts F'") and "\ts' F'. transaction_send T = \\,send\ts'\\#F' \ frsh \ [] \ transaction_send T' = \\,send\ts'\\#transaction_send T" (is "?E \ ?E' \ ?E''") and "frsh = [] \ transaction_send T' = transaction_send T" (is "?F \ ?F'") and "(xs' \ {} \ transaction_receive T' = \\, receive\f xs\\#transaction_receive T) \ (xs' = {} \ transaction_receive T' = transaction_receive T)" (is ?G) and "(frsh \ [] \ (\ts F'. transaction_send T = \\,send\ts\\#F' \ transaction_send T' = \\,send\ts'@ts\\#F')) \ (frsh \ [] \ transaction_send T' = \\,send\ts'\\#transaction_send T) \ (frsh = [] \ transaction_send T' = transaction_send T)" (is ?H) proof - note defs = T'_def T frsh_def xs_def xs'_def f_def C'_def ts'_def add_occurs_msgs_def Let_def show 0: "?A ts F' \ ?A' ts F'" for ts F' unfolding defs by simp have "F = [] \ fst (hd F) \ \ \ \is_Send (snd (hd F))" when ?B using that unfolding is_Send_def by (cases F) auto thus 1: "?B \ ?B' \ ?B''" unfolding defs by force show "?C \ ?C'" unfolding defs by auto show "transaction_decl T' = transaction_decl T" "transaction_fresh T' = transaction_fresh T" "transaction_checks T' = transaction_checks T" "transaction_updates T' = transaction_updates T" unfolding defs by simp_all show "xs = [] \ transaction_receive T' = transaction_receive T" "xs \ [] \ transaction_receive T' = \\, receive\f xs\\#transaction_receive T" unfolding defs by simp_all moreover have "xs = [] \ xs' = {}" using filter_empty_conv[of "\x. x \ set frsh"] fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] unfolding xs_def xs'_def by blast ultimately show ?G by blast show 2: "?D ts F' \ ?D' ts F'" for ts F' using 0 unfolding T by simp show 3: "?E \ ?E' \ ?E''" using 1 unfolding T by force show 4: "?F \ ?F'" unfolding defs by simp show ?H proof (cases "frsh = []") case False thus ?thesis using 2 3[OF _ False] by (cases "\ts F'. transaction_send T = \\,send\ts\\#F'") (blast,blast) qed (simp add: 4) qed private lemma add_occurs_msgs_transaction_strand_set: fixes T C frsh xs f defines "frsh \ transaction_fresh T" and "xs \ filter (\x. x \ set frsh) (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)))" and "f \ map (\x. occurs (Var x))" assumes T: "T = Transaction A B C D E F" shows "F = \\,send\ts\\#F' \ set (transaction_strand (add_occurs_msgs T)) \ set (transaction_strand T) \ {\\,receive\f xs\\,\\,send\f frsh@ts\\}" (is "?A \ ?A'") and "F = \\,send\ts\\#F' \ set (unlabel (transaction_strand (add_occurs_msgs T))) \ set (unlabel (transaction_strand T)) \ {receive\f xs\,send\f frsh@ts\}" (is "?B \ ?B'") and "\ts' F'. F = \\,send\ts'\\#F' \ set (transaction_strand (add_occurs_msgs T)) \ set (transaction_strand T) \ {\\,receive\f xs\\,\\,send\f frsh\\}" (is "?C \ ?C'") and "\ts' F'. F = \\,send\ts'\\#F' \ set (unlabel (transaction_strand (add_occurs_msgs T))) \ set (unlabel (transaction_strand T)) \ {receive\f xs\,send\f frsh\}" (is "?D \ ?D'") proof - note 0 = add_occurs_msgs_cases[ OF T, unfolded frsh_def[symmetric] xs_def[symmetric] f_def[symmetric]] show "?A \ ?A'" using 0(1,3) unfolding T transaction_strand_def by (cases "frsh = []") auto thus "?B \ ?B'" unfolding unlabel_def by force show "?C \ ?C'" using 0(2,3) unfolding T transaction_strand_def by (cases "frsh = []") auto thus "?D \ ?D'" unfolding unlabel_def by auto qed private lemma add_occurs_msgs_transaction_strand_cases: fixes T T'::"('a,'b,'c,'d) prot_transaction" and C frsh xs f \ defines "T' \ add_occurs_msgs T" and "S \ transaction_strand T" and "S' \ transaction_strand T'" and "frsh \ transaction_fresh T" and "xs \ filter (\x. x \ set frsh) (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)))" and "f \ map (\x. occurs (Var x))" and "C \ transaction_receive T" and "D \ transaction_checks T" and "E \ transaction_updates T" and "F \ transaction_send T" and "C' \ if xs = [] then C else \\,receive\f xs\\#C" and "C'' \ if xs = [] then dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t C else \\,send\f xs\\#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t C" and "C''' \ if xs = [] then dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (C \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) else \\,send\f xs \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (C \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "frsh = [] \ S' = C'@D@E@F" (is "?A \ ?A'") and "frsh \ [] \ \ts F'. F = \\,send\ts\\#F' \ S' = C'@D@E@(\\,send\f frsh\\#F)" (is "?B \ ?B' \ ?B''") and "frsh \ [] \ \ts F'. F = \\,send\ts\\#F' \ \ts F'. F = \\,send\ts\\#F' \ S' = C'@D@E@(\\,send\f frsh@ts\\#F')" (is "?C \ ?C' \ ?C''") and "frsh = [] \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S' = C''@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t D@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t E@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t F" (is "?D \ ?D'") and "frsh \ [] \ \ts F'. F = \\,send\ts\\#F' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S' = C''@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t D@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t E@(\\,receive\f frsh\\#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t F)" (is "?E \ ?E' \ ?E''") and "frsh \ [] \ \ts F'. F = \\,send\ts\\#F' \ \ts F'. F = \\,send\ts\\#F' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S' = C''@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t D@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t E@(\\,receive\f frsh@ts\\#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t F')" (is "?F \ ?F' \ ?F''") and "frsh = [] \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = C'''@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (D \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (E \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (F \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" (is "?G \ ?G'") and "frsh \ [] \ \ts F'. F = \\,send\ts\\#F' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = C'''@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (D \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (E \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@ (\\,receive\f frsh \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (F \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" (is "?H \ ?H' \ ?H''") and "frsh \ [] \ \ts F'. F = \\,send\ts\\#F' \ \ts F'. F = \\,send\ts\\#F' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = C'''@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (D \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (E \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@ (\\,receive\f frsh@ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (F' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" (is "?I \ ?I' \ ?I''") proof - obtain A' B' CC' D' E' F' where T: "T = Transaction A' B' CC' D' E' F'" by (cases T) simp note 0 = add_occurs_msgs_cases[ OF T, unfolded frsh_def[symmetric] xs_def[symmetric] f_def[symmetric] T'_def[symmetric]] note defs = S'_def C_def D_def E_def F_def C'_def C''_def T transaction_strand_def show A: "?A \ ?A'" using 0(3) unfolding defs by simp show B: "?B \ ?B' \ ?B''" using 0(2) unfolding defs by simp show C: "?C \ ?C' \ ?C''" using 0(1) unfolding defs by force have 1: "C''' = C'' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using subst_lsst_cons[of "\\, send\f xs\\" "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t C" \] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst[of C \] unfolding C'''_def C''_def by (cases "xs = []") auto have 2: "(\\, receive\ts\\#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t G) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = \\, receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (G \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" for ts and G::"('a,'b,'c,'d) prot_strand" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst[of G \] subst_lsst_cons[of "\\, receive\ts\\" "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t G" \] by simp note 3 = subst_lsst_append[of _ _ \] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst[of _ \] show "?D \ ?D'" using A unfolding defs by fastforce thus "?G \ ?G'" unfolding 1 by (metis 3) show "?E \ ?E' \ ?E''" using B unfolding defs by fastforce thus "?H \ ?H' \ ?H''" unfolding 1 by (metis 2 3) show "?F \ ?F' \ ?F''" using C unfolding defs by fastforce thus "?I \ ?I' \ ?I''" unfolding 1 by (metis 2 3) qed private lemma add_occurs_msgs_trms_transaction: fixes T::"('a,'b,'c,'d) prot_transaction" shows "trms_transaction (add_occurs_msgs T) = trms_transaction T \ (\x. occurs (Var x))`(fv_transaction T \ set (transaction_fresh T))" (is "?A = ?B") proof let ?occs = "(\x. occurs (Var x)) ` (fv_transaction T \ set (transaction_fresh T))" define frsh where "frsh \ transaction_fresh T" define xs where "xs \ filter (\x. x \ set frsh) (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)))" define f where "f \ map (\x. occurs (Var x)::('a,'b,'c,'d) prot_term)" obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp note 0 = add_occurs_msgs_transaction_strand_set(2,4)[ OF T, unfolded f_def[symmetric] frsh_def[symmetric] xs_def[symmetric]] note 1 = add_occurs_msgs_transaction_strand_cases(1,2,3)[ of T, unfolded f_def[symmetric] frsh_def[symmetric] xs_def[symmetric]] have 2: "set (f xs) \ set (f frsh) = ?occs" proof - define ys where "ys \ fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T))" let ?ys' = "fv_transaction T - set frsh" define g where "g \ filter (\x. x \ set frsh)" have "set (g ys) = ?ys'" using fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] unfolding ys_def g_def by auto hence "set (f (g ys)) = (\x. occurs (Var x)) ` ?ys'" unfolding f_def by force moreover have "set (f frsh) = (\x. occurs (Var x)) ` set frsh" unfolding f_def by force ultimately show ?thesis unfolding xs_def frsh_def[symmetric] ys_def[symmetric] g_def[symmetric] by blast qed have 3: "set (f []) = {}" unfolding f_def by blast have "trms_transaction (add_occurs_msgs T) \ trms_transaction T \ set (f xs) \ set (f frsh)" proof (cases "\ts F'. F = \\, send\ts\\#F'") case True then obtain ts F' where F: "F = \\, send\ts\\#F'" by blast have "set ts \ trms_transaction T" unfolding T F trms_transaction_unfold by auto thus ?thesis using 0(1)[OF F] by force next case False show ?thesis using 0(2)[OF False] by force qed thus "?A \ ?B" using 2 by blast have "trms_transaction T \ set (f xs) \ set (f frsh) \ trms_transaction (add_occurs_msgs T)" proof (cases "frsh = []") case True show ?thesis using 1(1)[OF True] 3 unfolding True by (cases xs) (fastforce,force) next case False note * = 1(2-)[OF False] show ?thesis proof (cases "\ts F'. transaction_send T = \\, send\ts\\#F'") case True show ?thesis using *(2)[OF True] 3 by force next case False show ?thesis using *(1)[OF False] 3 by force qed qed thus "?B \ ?A" using 2 by blast qed private lemma add_occurs_msgs_vars_eq: fixes T::"('fun,'var,'sets,'lbl) prot_transaction" assumes T_adm: "admissible_transaction' T" shows "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive (add_occurs_msgs T)) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" (is ?A) and "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send (add_occurs_msgs T)) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \ set (transaction_fresh T)" (is ?B) and "fv_transaction (add_occurs_msgs T) = fv_transaction T" (is ?C) and "bvars_transaction (add_occurs_msgs T) = bvars_transaction T" (is ?D) and "vars_transaction (add_occurs_msgs T) = vars_transaction T" (is ?E) and "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand (add_occurs_msgs T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" (is ?F) and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand (add_occurs_msgs T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" (is ?G) and "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand (add_occurs_msgs T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" (is ?H) and "set (transaction_fresh (add_occurs_msgs T)) = set (transaction_fresh T)" (is ?I) proof - obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp have T_fresh: "set (transaction_fresh T) \ fv_transaction T" using admissible_transactionE(7)[OF T_adm] unfolding fv_transaction_unfold by blast note 0 = add_occurs_msgs_cases[OF T] define xs where "xs \ filter (\x. x \ set (transaction_fresh T)) (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)))" show D: ?D proof - have "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive (add_occurs_msgs T)) = bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" using 0(6,7) by (cases "xs = []") (auto simp add: xs_def) moreover have "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send (add_occurs_msgs T)) = bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" proof (cases "\ts' F'. F = \\, send\ts'\\#F'") case True thus ?thesis using 0(1) unfolding T by force next case False show ?thesis using 0(2)[OF False] 0(3) unfolding T by (cases "B = []") auto qed ultimately show ?thesis using 0(8,9) unfolding bvars_transaction_unfold by argo qed have T_no_bvars: "bvars_transaction T = {}" "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) = {}" "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) = {}" "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) = {}" "bvars_transaction (add_occurs_msgs T) = {}" using admissible_transactionE(4)[OF T_adm] D unfolding bvars_transaction_unfold by (blast,blast,blast,blast,blast) have T_fv_subst: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv_transaction T)" (is ?Q1) "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T))" (is ?Q2) "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T))" (is ?Q3) "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" (is ?Q4) "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand (add_occurs_msgs T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand (add_occurs_msgs T)))" (is ?Q5) "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive (add_occurs_msgs T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive (add_occurs_msgs T)))" (is ?Q6) for \ proof - note * = fv\<^sub>s\<^sub>s\<^sub>t_subst_if_no_bvars have **: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive (add_occurs_msgs T)) = {}" using T_no_bvars(5) unfolding bvars_transaction_unfold by fast show ?Q1 using *[OF T_no_bvars(1)] unfolding unlabel_subst by blast show ?Q2 using *[OF T_no_bvars(2)] unfolding unlabel_subst by blast show ?Q3 using *[OF T_no_bvars(3)] unfolding unlabel_subst by blast show ?Q4 using *[OF T_no_bvars(4)] unfolding unlabel_subst by blast show ?Q5 using *[OF T_no_bvars(5)] unfolding unlabel_subst by blast show ?Q6 using *[OF **] unfolding unlabel_subst by blast qed have A: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive (add_occurs_msgs T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" for \ proof - define rcv_trms where "rcv_trms \ map (\x. occurs (Var x)::('fun,'var,'sets,'lbl) prot_term) xs" have "fv\<^sub>s\<^sub>e\<^sub>t (set rcv_trms) = fv_transaction T - set (transaction_fresh T)" "rcv_trms = [] \ xs = []" using fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] unfolding rcv_trms_def xs_def by auto hence 1: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive (add_occurs_msgs T)) = (fv_transaction T - set (transaction_fresh T)) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" using 0(6,7)[unfolded rcv_trms_def[symmetric] xs_def[symmetric]] by (cases "xs = []") auto have 2: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ fv_transaction T - set (transaction_fresh T)" using admissible_transactionE(12)[OF T_adm] unfolding fv_transaction_unfold by fast have 3: "fv_transaction T - set (transaction_fresh T) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using admissible_transactionE(7,10,12,13)[OF T_adm] unfolding fv_transaction_unfold by blast show ?thesis using 1 2 3 T_fv_subst(2,3,6)[of \] by force qed show ?A using A[of Var] unfolding subst_lsst_id_subst by blast show B: ?B using 0(14) by fastforce have B': "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send (add_occurs_msgs T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` set (transaction_fresh T))" for \ proof - note * = fv\<^sub>s\<^sub>s\<^sub>t_subst_if_no_bvars[of _ \] have **: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send (add_occurs_msgs T)) = {}" using T_no_bvars(5) unfolding bvars_transaction_unfold by fast show ?thesis using B *[OF T_no_bvars(4)] *[OF **] unfolding unlabel_subst by simp qed show C: ?C using A[of Var] B T_fresh unfolding fv_transaction_unfold 0(8,9) subst_lsst_id_subst by blast show ?E using C D vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t by metis have "fv\<^sub>s\<^sub>e\<^sub>t (\ ` set (transaction_fresh T)) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using T_fresh unfolding fv\<^sub>s\<^sub>s\<^sub>t_subst_if_no_bvars[OF T_no_bvars(1), of \, unfolded unlabel_subst] by auto thus F: ?F using A[of \] B'[of \] fv\<^sub>s\<^sub>s\<^sub>t_append fv\<^sub>s\<^sub>s\<^sub>t_subst_if_no_bvars[OF T_no_bvars(1), of \, unfolded unlabel_subst] fv\<^sub>s\<^sub>s\<^sub>t_subst_if_no_bvars[OF T_no_bvars(5), of \, unfolded unlabel_subst C] unfolding transaction_strand_def by argo show G: ?G using D bvars\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst by metis show ?H using F G vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t by metis show ?I using 0(5) by argo qed private lemma add_occurs_msgs_trms: "trms_transaction (add_occurs_msgs T) = trms_transaction T \ (\x. occurs (Var x)) ` (set (transaction_fresh T) \ fv_transaction T)" proof - let ?f = "\x. occurs (Var x)" let ?xs = "filter (\x. x \ set (transaction_fresh T)) (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)))" obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp note 0 = add_occurs_msgs_cases[OF T] have "set ?xs = fv_transaction T - set (transaction_fresh T)" using fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] by auto hence 1: "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive (add_occurs_msgs T)) = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ ?f ` (fv_transaction T - set (transaction_fresh T))" using 0(6,7) by (cases "?xs = []") auto have 2: "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send (add_occurs_msgs T)) = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \ ?f ` set (transaction_fresh T)" using 0(10,11,12) by (cases "transaction_fresh T = []") (simp,fastforce) have 3: "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive (add_occurs_msgs T)) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send (add_occurs_msgs T)) = trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \ ?f ` (set (transaction_fresh T) \ fv_transaction T)" using 1 2 by blast show ?thesis using 3 unfolding trms_transaction_unfold 0(8,9) by blast qed lemma add_occurs_msgs_admissible_occurs_checks: fixes T::"('fun,'atom,'sets,'lbl) prot_transaction" assumes T_adm: "admissible_transaction' T" shows "admissible_transaction' (add_occurs_msgs T)" (is ?A) and "admissible_transaction_occurs_checks (add_occurs_msgs T)" (is ?B) proof - let ?T' = "add_occurs_msgs T" obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp note defs = T add_occurs_msgs_def Let_def admissible_transaction'_def admissible_transaction_occurs_checks_def note defs' = admissible_transaction_terms_def wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[symmetric] note 1 = add_occurs_msgs_cases[OF T] note 2 = add_occurs_msgs_vars_eq[OF T_adm] note 3 = add_occurs_msgs_trms[of T] note 4 = add_occurs_msgs_transaction_strand_set[OF T] have occurs_wf: "wf\<^sub>t\<^sub>r\<^sub>m (occurs (Var x))" for x::"('fun,'atom,'sets,'lbl) prot_var" by fastforce have occurs_funs: "funs_term (occurs (Var x)) = {OccursFact, OccursSec}" for x::"('fun,'atom,'sets,'lbl) prot_var" by force have occurs_funs_not_attack: "\(\f \ \(funs_term ` trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r). is_Attack f)" when "r = receive\map (\x. occurs (Var x)) xs\ \ r = send\map (\x. occurs (Var x)) ys\" for r:: "(('fun,'atom,'sets,'lbl) prot_fun, ('fun,'atom,'sets,'lbl) prot_var) stateful_strand_step" and xs ys::"('fun,'atom,'sets,'lbl) prot_var list" using that by fastforce have occurs_funs_not_attack': "\(\f \ \(funs_term ` trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r). is_Attack f)" when "r = send\map (\x. occurs (Var x)) xs@ts\" and "\(\f \ \(funs_term ` trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (send\ts\)). is_Attack f)" for r:: "(('fun,'atom,'sets,'lbl) prot_fun, ('fun,'atom,'sets,'lbl) prot_var) stateful_strand_step" and xs::"('fun,'atom,'sets,'lbl) prot_var list" and ts using that by fastforce let ?P1 = "\T. wellformed_transaction T" let ?P2 = "\T. transaction_decl T () = []" let ?P3 = "\T. list_all (\x. fst x = TAtom Value) (transaction_fresh T)" let ?P4 = "\T. \x \ vars_transaction T. is_Var (fst x) \ (the_Var (fst x) = Value)" let ?P5 = "\T. bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) = {}" let ?P6 = "\T. set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (filter (is_Insert \ snd) (transaction_updates T)) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" let ?P7 = "\T. \x \ fv_transaction T - set (transaction_fresh T). \y \ fv_transaction T - set (transaction_fresh T). x \ y \ \Var x != Var y\ \ set (unlabel (transaction_checks T)) \ \Var y != Var x\ \ set (unlabel (transaction_checks T))" let ?P8 = "\T. fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) - set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" let ?P9 = "\T. \r \ set (unlabel (transaction_checks T)). is_Equality r \ fv (the_rhs r) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" let ?P10 = "\T. fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (filter (\s. is_InSet (snd s) \ the_check (snd s) = Assign) (transaction_checks T))" let ?P11 = "\T. admissible_transaction_checks T" let ?P12 = "\T. admissible_transaction_updates T" let ?P13 = "\T. admissible_transaction_terms T" let ?P14 = "\T. admissible_transaction_send_occurs_form T" let ?P15 = "\T. list_all (\a. is_Receive (snd a) \ the_msgs (snd a) \ []) (transaction_receive T)" let ?P16 = "\T. list_all (\a. is_Send (snd a) \ the_msgs (snd a) \ []) (transaction_send T)" have T_props: "?P1 T" "?P2 T" "?P3 T" "?P4 T" "?P5 T" "?P6 T" "?P7 T" "?P8 T" "?P9 T" "?P10 T" "?P11 T" "?P12 T" "?P13 T" "?P14 T" "?P15 T" "?P16 T" using T_adm unfolding defs by meson+ have 5: "wf'\<^sub>s\<^sub>s\<^sub>t (X \ Y) (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand (add_occurs_msgs T))))" when X: "X = fst ` set (transaction_decl T ())" and Y: "Y = set (transaction_fresh T)" and T_wf: "wf'\<^sub>s\<^sub>s\<^sub>t (X \ Y) (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)))" for X Y proof - define frsh where "frsh \ transaction_fresh T" define xs where "xs \ fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T))" define ys where "ys \ filter (\x. x \ set frsh) xs" let ?snds = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T))" let ?snds' = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive (add_occurs_msgs T)))" let ?chks = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T))" let ?chks' = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks (add_occurs_msgs T)))" let ?upds = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T))" let ?upds' = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates (add_occurs_msgs T)))" let ?rcvs = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" let ?rcvs' = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send (add_occurs_msgs T)))" have p0: "set ?snds \ set ?snds'" using 1(13) by auto have p1: "?chks = ?chks'" "?upds = ?upds'" using 1(8,9) by (argo,argo) have p2: "wfvarsoccs\<^sub>s\<^sub>s\<^sub>t ?snds \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t ?snds'" "wfvarsoccs\<^sub>s\<^sub>s\<^sub>t (?snds@?chks@?upds) \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t (?snds'@?chks'@?upds')" "X \ Y \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t (?snds@?chks@?upds) \ X \ Y \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t (?snds'@?chks'@?upds')" using p0 p1 unfolding wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def by auto have "wf'\<^sub>s\<^sub>s\<^sub>t (X \ Y \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t (?snds@?chks@?upds)) ?rcvs" using T_wf wf\<^sub>s\<^sub>s\<^sub>t_append_exec[of "X \ Y" "?snds@?chks@?upds" ?rcvs] unfolding transaction_strand_unlabel_dual_unfold by simp hence r0: "wf'\<^sub>s\<^sub>s\<^sub>t (X \ Y \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t (?snds'@?chks'@?upds')) ?rcvs" using wf\<^sub>s\<^sub>s\<^sub>t_vars_mono'[OF _ p2(3)] by blast have "list_all is_Send (unlabel (transaction_send T))" using admissible_transaction_is_wellformed_transaction(1)[OF T_adm] unfolding wellformed_transaction_def by blast hence "list_all is_Receive ?rcvs" by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(2)) hence r1: "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t ?rcvs \ X \ Y \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t (?snds'@?chks'@?upds')" using wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_receives_only_eq wf\<^sub>s\<^sub>s\<^sub>t_receives_only_fv_subset[OF r0] by blast have "fv\<^sub>s\<^sub>e\<^sub>t ((\x. occurs (Var x)) ` set (transaction_fresh T)) \ Y" unfolding Y by auto hence r2: "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t ?rcvs' \ X \ Y \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t (?snds'@?chks'@?upds')" using 1(14) r1 unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by fastforce (* TODO: find faster proof *) have r3: "wf'\<^sub>s\<^sub>s\<^sub>t (X \ Y) (?snds'@?chks'@?upds')" proof - have *: "wf'\<^sub>s\<^sub>s\<^sub>t (X \ Y) (?snds@?chks'@?upds')" using T_wf wf\<^sub>s\<^sub>s\<^sub>t_prefix[of "X \ Y" "?snds@?chks@?upds" ?rcvs] p1 unfolding transaction_strand_unlabel_dual_unfold by simp have "?snds' = ?snds \ (\ts. ?snds' = send\ts\#?snds)" using 1(13) by auto thus ?thesis proof assume "?snds' = ?snds" thus ?thesis using * by simp next assume "\ts. ?snds' = send\ts\#?snds" then obtain ts where "?snds' = send\ts\#?snds" by blast thus ?thesis using wf\<^sub>s\<^sub>s\<^sub>t_sends_only_prepend[OF *, of "[send\ts\]"] by simp qed qed have "wf'\<^sub>s\<^sub>s\<^sub>t (X \ Y) (?snds'@?chks'@?upds'@?rcvs')" using wf\<^sub>s\<^sub>s\<^sub>t_append_suffix''[OF r3] r2 by auto thus ?thesis using unlabel_append dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append unfolding transaction_strand_def by auto qed have T'_props_1: "?P1 ?T'" unfolding wellformed_transaction_def apply (intro conjI) subgoal using 1(13) T_props(1) unfolding wellformed_transaction_def by force subgoal using 1(8) T_props(1) unfolding wellformed_transaction_def by simp subgoal using 1(9) T_props(1) unfolding wellformed_transaction_def by simp subgoal using 1(14) T_props(1) unfolding wellformed_transaction_def by force subgoal using 1(4) T_props(1) unfolding wellformed_transaction_def by simp subgoal using 1(5) T_props(1) unfolding wellformed_transaction_def by simp subgoal using 1(4,5) T_props(1) unfolding wellformed_transaction_def by simp subgoal using T_props(1) unfolding 2(1) 1(5) wellformed_transaction_def by blast subgoal using 1(5,8) T_props(1) unfolding wellformed_transaction_def by simp subgoal using 1(5) 2(4) T_props(1) unfolding wellformed_transaction_def by simp subgoal using 2(3,4) T_props(1) unfolding wellformed_transaction_def by simp subgoal using 1(4,5) 5 T_props(1) unfolding wellformed_transaction_def by simp done have T'_props_2_12: "?P2 ?T'" "?P3 ?T'" "?P4 ?T'" "?P5 ?T'" "?P6 ?T'" "?P7 ?T'" "?P8 ?T'" "?P9 ?T'" "?P10 ?T'" "?P11 ?T'" "?P12 ?T'" subgoal using T_props(2) unfolding defs by force subgoal using T_props(3) unfolding defs by force subgoal using T_props(4) 2(5) by argo subgoal using T_props(5) 2(4) by argo subgoal using T_props(6) 1(5,8) 2(2) by auto subgoal using T_props(7) 1(5,8) 2(3) by presburger subgoal using T_props(8) 1(5,9) 2(1,2) by auto subgoal using T_props(9) 1(8) 2(1) by auto subgoal using T_props(10) 1(8) 2(1) by auto subgoal using T_props(11) 1(8) unfolding admissible_transaction_checks_def by argo subgoal using T_props(12) 1(9) unfolding admissible_transaction_updates_def by argo done (* TODO: clean up? *) have T'_props_13_aux: "transaction_fresh ?T' = []" (is ?Q1) "is_Send r" (is ?Q2) "length (the_msgs r) = 1" (is ?Q3) "is_Fun_Attack (hd (the_msgs r))" (is ?Q4) when r: "r \ set (unlabel (transaction_strand (add_occurs_msgs T)))" "\f \ \(funs_term ` trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r). is_Attack f" (is "?Q' (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r)") for r proof - note q0 = conjunct2[OF conjunct2[OF T_props(13)[unfolded defs']]] let ?Q'' = "\ts' F'. F = \\, send\ts'\\#F'" let ?f = "map (\x. occurs (Var x))" let ?frsh = "transaction_fresh T" let ?xs = "fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T))" have q1: "r \ send\?f ?frsh\" "r \ receive\?f (filter (\x. x \ set ?frsh) ?xs)\" "\f \ \(funs_term ` set (?f ?frsh)). \is_Attack f" using r(2) by (fastforce,fastforce,simp) have q2: "send\ts'\ \ set (unlabel (transaction_strand T))" "r = send\?f ?frsh@ts'\ \ r \ set (unlabel (transaction_strand T))" when "?Q'' ts' F'" for ts' F' subgoal using that unfolding T transaction_strand_def by force subgoal using that r(1) 4(2)[OF that] q1 unfolding T transaction_strand_def by fast done have q3: "?Q' (set ts')" when r': "?Q'' ts' F'" "r \ set (unlabel (transaction_strand T))" for ts' F' proof - have "r = send\?f ?frsh@ts'\" using q2(2)[OF r'(1)] r'(2) by argo thus ?thesis using r(2) by fastforce qed have q4: "r \ set (unlabel (transaction_strand T))" when "\ts' F'. ?Q'' ts' F'" using 4(4)[OF that] r(1) q1(1,2) by blast have "\r' \ set (unlabel (transaction_strand T)). ?Q' (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r')" when "?Q'' ts' F'" for ts' F' apply (cases "r \ set (unlabel (transaction_strand T))") subgoal using q2(2)[OF that] r(2) by metis subgoal using q2(1)[OF that] q3[OF that] trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p.simps(1)[of ts'] by metis done hence "?frsh = []" when "?Q'' ts' F'" for ts' F' using q0 that by blast hence "r = send\ts'\ \ r \ set (unlabel (transaction_strand T))" when "?Q'' ts' F'" for ts' F' using q2(2)[OF that] that by blast hence "r \ set (unlabel (transaction_strand T))" using q2(1) q4 by fast thus ?Q1 ?Q2 ?Q3 ?Q4 using r(2) q0 unfolding 1(5) by auto qed have T'_props_13: "?P13 ?T'" unfolding defs' 3 apply (intro conjI) subgoal using conjunct1[OF T_props(13)[unfolded defs']] occurs_wf by fast subgoal using conjunct1[OF conjunct2[OF T_props(13)[unfolded defs']]] occurs_funs by auto subgoal using T'_props_13_aux by meson done have T'_props_14: "?P14 ?T'" proof (cases "\ts' F'. transaction_send T = \\,send\ts'\\#F'") case True then obtain ts' F' where F': "transaction_send T = \\,send\ts'\\#F'" by meson show ?thesis using T_props(14) 1(10)[OF F'] F' 1(5,12) unfolding admissible_transaction_send_occurs_form_def Let_def by (cases "transaction_fresh T = []") auto next case False show ?thesis using T_props(14) 1(11)[OF False] 1(5,12) unfolding admissible_transaction_send_occurs_form_def Let_def by (cases "transaction_fresh T = []") auto qed let ?xs = "fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T))" have T'_props_15: "?P15 ?T'" using T_props(15) 1(6,7) unfolding Let_def by (cases "filter (\x. x \ set (transaction_fresh T)) ?xs = []") (simp,fastforce) have T'_props_16: "?P16 ?T'" proof (cases "\ts' F'. transaction_send T = \\,send\ts'\\#F'") case True then obtain ts' F' where F': "transaction_send T = \\,send\ts'\\#F'" by meson show ?thesis using T_props(16) 1(10)[OF F'] F' 1(5,12) unfolding Let_def by (cases "transaction_fresh T = []") auto next case False show ?thesis using T_props(16) 1(11)[OF False] 1(5,12) unfolding Let_def by (cases "transaction_fresh T = []") auto qed note T'_props = T'_props_1 T'_props_2_12 T'_props_13 T'_props_14 T'_props_15 T'_props_16 show ?A using T'_props unfolding admissible_transaction'_def by meson have 5: "set (filter (\x. x \ set (transaction_fresh T)) (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)))) = fv_transaction T - set (transaction_fresh T)" using fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t by fastforce have "transaction_receive ?T' \ []" and "is_Receive (hd (unlabel (transaction_receive ?T')))" and "\x \ fv_transaction ?T' - set (transaction_fresh ?T'). fst x = TAtom Value \ occurs (Var x) \ set (the_msgs (hd (unlabel (transaction_receive ?T'))))" when x: "x \ fv_transaction ?T' - set (transaction_fresh ?T')" "fst x = TAtom Value" for x using 1(13) 5 x unfolding 1(5) 2(3) by (force,force,force) moreover have "transaction_send ?T' \ []" (is ?C) and "is_Send (hd (unlabel (transaction_send ?T')))" (is ?D) and "\x \ set (transaction_fresh ?T'). occurs (Var x) \ set (the_msgs (hd (unlabel (transaction_send ?T'))))" (is ?E) when T'_frsh: "transaction_fresh ?T' \ []" using 1(14) T'_frsh unfolding 1(5) by auto ultimately show ?B using T'_props_14 unfolding admissible_transaction_occurs_checks_def Let_def by blast qed private lemma add_occurs_msgs_in_trms_subst_cases: fixes T::"('fun,'atom,'sets,'lbl) prot_transaction" assumes T_adm: "admissible_transaction' T" and t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand (add_occurs_msgs T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ (\x \ fv_transaction T. t = occurs (\ x))" proof - define frsh where "frsh \ transaction_fresh T" define xs where "xs \ filter (\x. x \ set frsh) (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)))" define f where "f \ map (\x. occurs (Var x)::('fun,'atom,'sets,'lbl) prot_term)" obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp note T'_adm = add_occurs_msgs_admissible_occurs_checks(1)[OF T_adm] have 0: "set (transaction_fresh T) \ fv_transaction T" using admissible_transactionE(7)[OF T_adm] unfolding fv_transaction_unfold by blast hence 00: "set (f xs) \ set (f frsh) = (\x. occurs (Var x)) ` fv_transaction T" using fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] unfolding f_def xs_def frsh_def by auto note 1 = add_occurs_msgs_transaction_strand_set[OF T] have 2: "set (transaction_strand (add_occurs_msgs T)) \ set (transaction_strand T) \ {\\,receive\f xs\\,\\,send\f frsh\\}" when "\ts F'. F = \\,send\ts\\#F'" using 1(3,4)[OF that] unfolding f_def[symmetric] frsh_def[symmetric] xs_def[symmetric] by blast have 3: "trms_transaction (add_occurs_msgs T) = trms_transaction T \ (\x. occurs (Var x)) ` fv_transaction T" using 0 add_occurs_msgs_trms_transaction[of T] by blast have 4: "bvars_transaction T \ subst_domain \ = {}" "bvars_transaction (add_occurs_msgs T) \ subst_domain \ = {}" using admissible_transactionE(4)[OF T_adm] admissible_transactionE(4)[OF T'_adm] by (blast,blast) note 5 = trms\<^sub>s\<^sub>s\<^sub>t_subst[OF 4(1), unfolded unlabel_subst] trms\<^sub>s\<^sub>s\<^sub>t_subst[OF 4(2), unfolded unlabel_subst] note 6 = fv\<^sub>s\<^sub>s\<^sub>t_is_subterm_trms\<^sub>s\<^sub>s\<^sub>t_subst[ OF _ 4(1), unfolded add_occurs_msgs_admissible_occurs_checks(1)[OF T_adm] unlabel_subst] show ?thesis using t 6 unfolding 3 5 by fastforce qed private lemma add_occurs_msgs_updates_send_filter_iff: fixes f defines "f \ \T. list_ex (\a. is_Send (snd a) \ is_Update (snd a)) (transaction_strand T)" and "g \ \T. transaction_fresh T = [] \ f T" shows "map add_occurs_msgs (filter g P) = filter g (map add_occurs_msgs P)" proof - have "g T \ g (add_occurs_msgs T)" for T proof - obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp_all note 0 = add_occurs_msgs_cases[OF T] show ?thesis using 0(6,7,12) unfolding g_def f_def transaction_strand_def 0(5,8,9) by fastforce qed thus ?thesis by (induct P) simp_all qed lemma add_occurs_msgs_updates_send_filter_iff': fixes f defines "f \ \T. list_ex (\a. is_Send (snd a) \ is_Update (snd a)) (transaction_strand T)" and "g \ \T. transaction_fresh T = [] \ transaction_updates T \ [] \ transaction_send T \ []" shows "map add_occurs_msgs (filter g P) = filter g (map add_occurs_msgs P)" proof - have "g T \ g (add_occurs_msgs T)" for T proof - obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp_all note 0 = add_occurs_msgs_cases[OF T] show ?thesis using 0(6,7,12) unfolding g_def f_def transaction_strand_def 0(5,8,9) by argo qed thus ?thesis by (induct P) simp_all qed private lemma rm_occurs_msgs_constr_Cons: defines "f \ rm_occurs_msgs_constr" shows "\is_Receive a \ \is_Send a \ f ((l,a)#A) = (l,a)#f A" "is_Receive a \ \t. occurs t \ set (the_msgs a) \ f ((l,a)#A) = (l,a)#f A" "is_Receive a \ \t. occurs t \ set (the_msgs a) \ \t \ set (the_msgs a). \s. t \ occurs s \ f ((l,a)#A) = (l,receive\filter (\t. \s. t \ occurs s) (the_msgs a)\)#f A" "is_Receive a \ \t. occurs t \ set (the_msgs a) \ \t \ set (the_msgs a). \s. t = occurs s \ f ((l,a)#A) = f A" "is_Send a \ \t. occurs t \ set (the_msgs a) \ f ((l,a)#A) = (l,a)#f A" "is_Send a \ \t. occurs t \ set (the_msgs a) \ \t \ set (the_msgs a). \s. t \ occurs s \ f ((l,a)#A) = (l,send\filter (\t. \s. t \ occurs s) (the_msgs a)\)#f A" "is_Send a \ \t. occurs t \ set (the_msgs a) \ \t \ set (the_msgs a). \s. t = occurs s \ f ((l,a)#A) = f A" unfolding f_def by (cases a; auto)+ private lemma rm_occurs_msgs_constr_Cons': defines "f \ rm_occurs_msgs_constr" and "g \ filter (\t. \s. t \ occurs s)" assumes a: "is_Receive a \ is_Send a" shows "\t. occurs t \ set (the_msgs a) \ f ((l,a)#A) = (l,a)#f A" "\t. occurs t \ set (the_msgs a) \ \t \ set (the_msgs a). \s. t \ occurs s \ is_Send a \ f ((l,a)#A) = (l,send\g (the_msgs a)\)#f A" "\t. occurs t \ set (the_msgs a) \ \t \ set (the_msgs a). \s. t \ occurs s \ is_Receive a \ f ((l,a)#A) = (l,receive\g (the_msgs a)\)#f A" "\t. occurs t \ set (the_msgs a) \ \t \ set (the_msgs a). \s. t = occurs s \ f ((l,a)#A) = f A" using a unfolding f_def g_def by (cases a; auto)+ private lemma rm_occurs_msgs_constr_Cons'': defines "f \ rm_occurs_msgs_constr" and "g \ filter (\t. \s. t \ occurs s)" assumes a: "a = receive\ts\ \ a = send\ts\" shows "f ((l,a)#A) = (l,a)#f A \ f ((l,a)#A) = (l,receive\g ts\)#f A \ f ((l,a)#A) = (l,send\g ts\)#f A \ f ((l,a)#A) = f A" using rm_occurs_msgs_constr_Cons(2-)[of a l A] a unfolding f_def g_def by (cases a) auto private lemma rm_occurs_msgs_constr_ik_subset: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (rm_occurs_msgs_constr A) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" proof (induction A) case (Cons a A) let ?f = "filter (\t. \s. t \ occurs s)" note IH = Cons.IH obtain l b where a: "a = (l,b)" by (metis surj_pair) have 0: "set (unlabel A) \ set (unlabel (a#A))" by auto note 1 = rm_occurs_msgs_constr_Cons[of b l A] note 2 = in_ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t_iff note 3 = ik\<^sub>s\<^sub>s\<^sub>t_set_subset[OF 0] note 4 = ik\<^sub>s\<^sub>s\<^sub>t_append note 5 = 4[of "unlabel [a]" "unlabel A"] 4[of "unlabel [a]" "unlabel (rm_occurs_msgs_constr A)"] show ?case proof (cases "is_Send b \ is_Receive b") case True note b_cases = this define ts where "ts \ the_msgs b" have ts_cases: "is_Send b \ b = send\ts\" "is_Receive b \ b = receive\ts\" unfolding ts_def by simp_all have 6: "is_Send b \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t [(l,b)] = {}" "is_Send b \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t [(l,send\the_msgs b\)] = {}" "is_Send b \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t [(l,send\?f (the_msgs b)\)] = {}" "is_Receive b \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t [(l,b)] = set ts" "is_Receive b \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t [(l,receive\the_msgs b\)] = set ts" "is_Receive b \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t [(l,receive\?f (the_msgs b)\)] = set (?f ts)" using 2[of _ "[(l, send\the_msgs b\)]"] 2[of _ "[(l, send\?f (the_msgs b)\)]"] 2[of _ "[(l, receive\the_msgs b\)]"] 2[of _ "[(l, receive\?f (the_msgs b)\)]"] b_cases ts_cases by auto have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (rm_occurs_msgs_constr (a#A)) = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (rm_occurs_msgs_constr A)" when b: "is_Send b" proof (cases "\t. occurs t \ set (the_msgs b)") case True note 7 = 1(6,7)[OF b True] show ?thesis proof (cases "\t \ set (the_msgs b). \s. t \ occurs s") case True show ?thesis using 4[of "unlabel [(l,send\?f (the_msgs b)\)]" "unlabel (rm_occurs_msgs_constr A)"] unfolding a 7(1)[OF True] 6(3)[OF b] by simp next case False hence F: "\t \ set (the_msgs b). \s. t = occurs s" by simp show ?thesis using 4[of "unlabel [(l,send\the_msgs b\)]" "unlabel (rm_occurs_msgs_constr A)"] unfolding a 7(2)[OF F] 6(2)[OF b] by simp qed next case False show ?thesis using 4[of "unlabel [(l,b)]" "unlabel (rm_occurs_msgs_constr A)"] unfolding a 1(5)[OF b False] 6(1)[OF b] by auto qed moreover have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (rm_occurs_msgs_constr (a#A)) \ set ts \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (rm_occurs_msgs_constr A)" when b: "is_Receive b" proof (cases "\t. occurs t \ set (the_msgs b)") case True note 8 = 1(3,4)[OF b True] show ?thesis proof (cases "\t \ set (the_msgs b). \s. t \ occurs s") case True show ?thesis using 4[of "unlabel [(l,receive\?f (the_msgs b)\)]" "unlabel (rm_occurs_msgs_constr A)"] unfolding a 8(1)[OF True] 6(6)[OF b] by auto next case False hence F: "\t \ set (the_msgs b). \s. t = occurs s" by simp show ?thesis using 4[of "unlabel [(l,receive\the_msgs b\)]" "unlabel (rm_occurs_msgs_constr A)"] unfolding a 8(2)[OF F] 6(5)[OF b] by simp qed next case False show ?thesis using 4[of "unlabel [(l,b)]" "unlabel (rm_occurs_msgs_constr A)"] unfolding a 1(2)[OF b False] 6(4)[OF b] by auto qed moreover have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) = set ts \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" when b: "is_Receive b" using ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t_Cons(2)[of l ts A] unfolding a ts_cases(2)[OF b] by blast ultimately show ?thesis using IH 3 b_cases by blast qed (use 1(1) IH 5 a in auto) qed simp private lemma rm_occurs_msgs_constr_append: "rm_occurs_msgs_constr (A@B) = rm_occurs_msgs_constr A@rm_occurs_msgs_constr B" by (induction A rule: rm_occurs_msgs_constr.induct) auto private lemma rm_occurs_msgs_constr_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t: "rm_occurs_msgs_constr (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (rm_occurs_msgs_constr A)" proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) show ?case using Cons.IH unfolding a by (cases b) auto qed simp private lemma rm_occurs_msgs_constr_dbupd\<^sub>s\<^sub>s\<^sub>t_eq: "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (rm_occurs_msgs_constr A)) I D = dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I D" proof (induction A arbitrary: I D) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) show ?case proof (cases "is_Receive b \ is_Send b") case True then obtain ts where b: "b = receive\ts\ \ b = send\ts\" by (cases b) simp_all show ?thesis using rm_occurs_msgs_constr_Cons''[OF b, of l A] Cons.IH b unfolding a by fastforce next case False thus ?thesis using Cons.IH unfolding a by (cases b) auto qed qed simp private lemma rm_occurs_msgs_constr_subst: fixes A::"('a,'b,'c,'d) prot_strand" and \::"('a,'b,'c,'d) prot_subst" assumes "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \t. \ x = occurs t" "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \ x \ Fun OccursSec []" shows "rm_occurs_msgs_constr (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = (rm_occurs_msgs_constr A) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" (is "?f (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = (?f A) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \") using assms proof (induction A) case (Cons a A) note 0 = rm_occurs_msgs_constr_Cons note 1 = rm_occurs_msgs_constr_Cons' define f where "f \ ?f" define not_occ where "not_occ \ \t::('a,'b,'c,'d) prot_term. \s. t \ occurs s" define flt where "flt \ filter not_occ" obtain l b where a: "a = (l,b)" by (metis surj_pair) have 2: "\t. \ x = occurs t" "\ x \ Fun OccursSec []" when b: "is_Receive b \ is_Send b" and t: "t \ set (the_msgs b)" and x: "x \ fv t" for x t using Cons.prems x t b unfolding a by (cases b; auto)+ have IH: "f (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = (f A) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using Cons.prems Cons.IH unfolding f_def by simp show ?case proof (cases "is_Receive b \ is_Send b") case True note T = this then obtain ts where ts: "b = receive\ts\ \ b = send\ts\" by (cases b) simp_all hence ts': "b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\ \ b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = send\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\" by auto have the_msgs_b: "the_msgs b = ts" "the_msgs (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using ts ts' by auto have 4: "is_Receive (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ is_Send (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" using T by (cases b) simp_all note 6 = 1[OF T, of l A, unfolded f_def[symmetric]] note 7 = 1[OF 4, of l "A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \", unfolded f_def[symmetric]] note 8 = ts IH subst_lsst_cons[of _ _ \] have 9: "t \ \ \ set (the_msgs (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \))" "not_occ (t \ \)" when t: "t \ set (the_msgs b)" "not_occ t" for t proof - show "t \ \ \ set (the_msgs (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \))" using t ts ts' by auto moreover have "not_occ (t \ \)" when "t = Var x" for x using 2[OF T t(1)] t(2) unfolding that not_occ_def by simp moreover have "not_occ (t \ \)" when "t = Fun g ss" "g \ OccursFact" for g ss using 2[OF T t(1)] t(2) that(2) unfolding that(1) not_occ_def by simp moreover have "not_occ (t \ \)" when "t = Fun OccursFact ss" "\s1 s2. ss = [s1,s2]" for ss using 2[OF T t(1)] t(2) that(2) unfolding that(1) not_occ_def by auto moreover have "not_occ (t \ \)" when "t = Fun OccursFact [s1,s2]" for s1 s2 using 2[OF T t(1)] t(2) unfolding that not_occ_def by (cases s1) auto ultimately show "not_occ (t \ \)" by (cases t) (metis, metis) qed have 10: "not_occ t" when t: "t \ set (the_msgs b)" "not_occ (t \ \)" for t proof - have "t \ \ \ set (the_msgs (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \))" using t ts ts' by auto moreover have "not_occ t" when "t = Var x" for x using 2[OF T t(1)] t(2) unfolding that not_occ_def by simp moreover have "not_occ t" when "t = Fun g ss" "g \ OccursFact" for g ss using 2[OF T t(1)] t(2) that(2) unfolding that(1) not_occ_def by simp moreover have "not_occ t" when "t = Fun OccursFact ss" "\s1 s2. ss = [s1,s2]" for ss using 2[OF T t(1)] t(2) that(2) unfolding that(1) not_occ_def by auto moreover have "not_occ t" when "t = Fun OccursFact [s1,s2]" for s1 s2 using 2[OF T t(1)] t(2) unfolding that not_occ_def by (cases s1) auto ultimately show "not_occ t" unfolding not_occ_def by force qed have 11: "not_occ (t \ \) \ not_occ t" when "t \ set ts" for t using that 9 10 unfolding the_msgs_b by blast have 5: "(\t. occurs t \ set ts) \ (\t. occurs t \ set ts \\<^sub>s\<^sub>e\<^sub>t \)" using 11 image_iff unfolding not_occ_def by fastforce have 12: "flt (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) = (flt ts) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using 11 proof (induction ts) case (Cons t ts) hence "not_occ (t \ \) = not_occ t" "flt (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) = (flt ts) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" by auto thus ?case unfolding flt_def by auto qed (metis flt_def filter.simps(1) map_is_Nil_conv) show ?thesis proof (cases "\t. occurs t \ set (the_msgs b)") case True note T1 = this hence T2: "\t. occurs t \ set (the_msgs (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \))" using 5 unfolding the_msgs_b by simp show ?thesis proof (cases "\t \ set (the_msgs b). \s. t \ occurs s") case True note T1' = this have T2': "\t \ set (the_msgs (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)). \s. t \ occurs s" using T1' 11 unfolding the_msgs_b not_occ_def by auto show ?thesis using T proof assume b: "is_Receive b" hence b\: "is_Receive (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" using ts by fastforce show ?thesis using 6(3)[OF T1 T1' b] 7(3)[OF T2 T2' b\] IH 12 unfolding f_def[symmetric] a flt_def[symmetric] not_occ_def[symmetric] the_msgs_b by (simp add: subst_lsst_cons) next assume b: "is_Send b" hence b\: "is_Send (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" using ts by fastforce show ?thesis using 6(2)[OF T1 T1' b] 7(2)[OF T2 T2' b\] IH 12 unfolding f_def[symmetric] a flt_def[symmetric] not_occ_def[symmetric] the_msgs_b by (simp add: subst_lsst_cons) qed next case False hence F: "\t \ set (the_msgs b). \s. t = occurs s" by blast hence F': "\t \ set (the_msgs (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)). \s. t = occurs s" unfolding the_msgs_b by auto have *: "\t. occurs t \ set (the_msgs b)" when "the_msgs b \ []" using that F by (cases "the_msgs b") auto hence **: "\t. occurs t \ set (the_msgs (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \))" when "the_msgs b \ []" using that 5 unfolding the_msgs_b by simp show ?thesis proof (cases "ts = []") case True hence ***: "\t. occurs t \ set (the_msgs b)" "\t. occurs t \ set (the_msgs (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \))" unfolding the_msgs_b by simp_all show ?thesis using IH 6(1)[OF ***(1)] 7(1)[OF ***(2)] unfolding a f_def[symmetric] True by (simp add: subst_lsst_cons) next case False thus ?thesis using IH 6(4)[OF * F] 7(4)[OF ** F'] unfolding f_def[symmetric] not_occ_def[symmetric] a the_msgs_b by (simp add: subst_lsst_cons) qed qed next case False note F = this have F': "\t. occurs t \ set (the_msgs (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \))" using F 11 unfolding not_occ_def the_msgs_b by fastforce show ?thesis using IH 6(1)[OF F] 7(1)[OF F'] unfolding a f_def[symmetric] True by (simp add: subst_lsst_cons) qed next case False hence *: "\is_Receive b" "\is_Send b" "\is_Receive (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "\is_Send (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" by (cases b; auto)+ show ?thesis using IH 0(1)[OF *(1,2), of l A] 0(1)[OF *(3,4), of l "A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] subst_lsst_cons[of a _ \] unfolding a f_def by simp qed qed simp private lemma rm_occurs_msgs_constr_transaction_strand: assumes T_adm: "admissible_transaction' T" shows "rm_occurs_msgs_constr (transaction_checks T) = transaction_checks T" (is ?A) and "rm_occurs_msgs_constr (transaction_updates T) = transaction_updates T" (is ?B) and "admissible_transaction_no_occurs_msgs T \ rm_occurs_msgs_constr (transaction_receive T) = transaction_receive T" (is "?C \ ?C'") and "admissible_transaction_no_occurs_msgs T \ rm_occurs_msgs_constr (transaction_send T) = transaction_send T" (is "?D \ ?D'") proof - note 0 = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] note 1 = wellformed_transaction_cases[OF 0] have 2: "\ts. b = receive\ts\ \ (\t. occurs t \ set ts)" when "admissible_transaction_no_occurs_msgs T" "(l,b) \ set (transaction_receive T)" for l b using that 1(1)[OF that(2)] unfolding admissible_transaction_no_occurs_msgs_def Let_def list_all_iff by fastforce have 3: "\ts. b = send\ts\ \ (\t. occurs t \ set ts)" when "admissible_transaction_no_occurs_msgs T" "(l,b) \ set (transaction_send T)" for l b using that 1(4)[OF that(2)] unfolding admissible_transaction_no_occurs_msgs_def Let_def list_all_iff by fastforce define A where "A \ transaction_receive T" define B where "B \ transaction_checks T" define C where "C \ transaction_updates T" define D where "D \ transaction_send T" show ?A using 1(2) unfolding B_def[symmetric] proof (induction B) case (Cons a A) hence IH: "rm_occurs_msgs_constr A = A" by (meson list.set_intros(2)) obtain l b where a: "a = (l,b)" by (metis surj_pair) show ?case using Cons.prems IH unfolding a by (cases b) auto qed simp show ?B using 1(3) unfolding C_def[symmetric] proof (induction C) case (Cons a A) hence IH: "rm_occurs_msgs_constr A = A" by (meson list.set_intros(2)) obtain l b where a: "a = (l,b)" by (metis surj_pair) show ?case using Cons.prems IH unfolding a by (cases b) auto qed simp show ?C' when ?C using 2[OF that] unfolding A_def[symmetric] proof (induction A) case (Cons a A) hence IH: "rm_occurs_msgs_constr A = A" by (meson list.set_intros(2)) obtain l b where a: "a = (l,b)" by (metis surj_pair) obtain ts where b: "b = receive\ts\" using Cons.prems unfolding a by auto show ?case using Cons.prems IH unfolding a b by fastforce qed simp show ?D' when ?D using 3[OF that] unfolding D_def[symmetric] proof (induction D) case (Cons a A) hence IH: "rm_occurs_msgs_constr A = A" by (meson list.set_intros(2)) obtain l b where a: "a = (l,b)" by (metis surj_pair) obtain ts where b: "b = send\ts\" using Cons.prems unfolding a by auto show ?case using Cons.prems IH unfolding a b by fastforce qed simp qed private lemma rm_occurs_msgs_constr_transaction_strand': fixes T::"('fun,'atom,'sets,'lbl) prot_transaction" assumes T_adm: "admissible_transaction' T" and T_no_occ: "admissible_transaction_no_occurs_msgs T" shows "rm_occurs_msgs_constr (transaction_strand (add_occurs_msgs T)) = transaction_strand T" (is "?f (?g (?h T)) = ?g T") proof - obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp have B: "B = transaction_fresh T" unfolding T by simp have F: "F = transaction_send T" unfolding T by simp define xs where "xs \ filter (\x. x \ set B) (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)))" note 0 = rm_occurs_msgs_constr_transaction_strand note 1 = add_occurs_msgs_admissible_occurs_checks[OF T_adm] note 2 = 0(3,4)[OF T_adm T_no_occ] note 3 = add_occurs_msgs_cases[OF T] note 4 = 0(1,2)[OF 1(1)] have 5: "?f (transaction_checks (?h T)) = transaction_checks T" "?f (transaction_updates (?h T)) = transaction_updates T" using 4 3(8,9) by (argo, argo) have 6: "?f (transaction_receive (?h T)) = transaction_receive T" proof (cases "xs = []") case True show ?thesis using 3(6)[OF True[unfolded xs_def B]] 2(1) by simp next case False show ?thesis using False 3(7)[OF False[unfolded xs_def B]] 2(1) rm_occurs_msgs_constr_Cons(4)[ of "receive\map (\x. occurs (Var x)) xs\" \ "transaction_receive T"] unfolding B[symmetric] xs_def[symmetric] by (cases xs) (blast, auto) qed have 7: "?f (transaction_send (?h T)) = transaction_send T" proof (cases "\ts' F'. F = \\, send\ts'\\#F'") case True then obtain ts' F' where F': "F = \\, send\ts'\\#F'" by blast have *: "transaction_send (?h T) = \\, send\map (\x. occurs (Var x)) B@ts'\\#F'" using 3(1)[OF F'] unfolding T by fastforce have **: "ts' \ []" using admissible_transactionE(17)[OF T_adm] unfolding T F' by auto have ***: "\s. t \ occurs s" when t: "t \ set ts'" for t using that T_no_occ unfolding T F' admissible_transaction_no_occurs_msgs_def Let_def list_all_iff by auto let ?ts = "map (\x. occurs (Var x)) B@ts'" have "\t \ set ?ts. \s. t \ occurs s" using ** *** by (cases ts') auto moreover have "filter (\t. \s. t \ occurs s) ?ts = ts'" using *** by simp moreover have "\t. occurs t \ set ?ts" when "B \ []" using that by (cases B) auto moreover have "?f [\\, send\ts'\\] = [\\, send\ts'\\]" using 2(2) ** *** unfolding F[symmetric] F' by force hence "?f F' = F'" using 2(2) rm_occurs_msgs_constr_append[of "[\\, send\ts'\\]" F'] unfolding F[symmetric] F' by fastforce ultimately have "?f (\\, send\?ts\\#F') = \\, send\ts'\\#F'" using 2(2) 3(10)[OF F'[unfolded F]] 3(12) rm_occurs_msgs_constr.simps(3)[of \ ts' F'] rm_occurs_msgs_constr_append[of "[\\, send\ts'\\]" F'] unfolding F[symmetric] F' B[symmetric] by auto thus ?thesis using F * unfolding F' by argo next case False show ?thesis using 3(2)[OF False] 3(3) 2(2) unfolding B[symmetric] xs_def[symmetric] F[symmetric] by (cases B) auto qed show ?thesis using 5 6 7 rm_occurs_msgs_constr_append unfolding transaction_strand_def by metis qed private lemma rm_occurs_msgs_constr_transaction_strand'': fixes T::"('fun,'atom,'sets,'lbl) prot_transaction" assumes T_adm: "admissible_transaction' T" and T_no_occ: "admissible_transaction_no_occurs_msgs T" and \: "\x \ fv_transaction (add_occurs_msgs T). \t. \ x = occurs t" "\x \ fv_transaction (add_occurs_msgs T). \ x \ Fun OccursSec []" shows "rm_occurs_msgs_constr (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand (add_occurs_msgs T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using rm_occurs_msgs_constr_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t[of "transaction_strand (add_occurs_msgs T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] rm_occurs_msgs_constr_subst[OF \] rm_occurs_msgs_constr_transaction_strand'[OF T_adm T_no_occ] by argo private lemma rm_occurs_msgs_constr_bvars_subst_eq: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (rm_occurs_msgs_constr A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" proof - have "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (rm_occurs_msgs_constr A) = bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) show ?case using Cons.IH unfolding a by (cases b) auto qed simp thus ?thesis by (metis bvars\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst) qed private lemma rm_occurs_msgs_constr_reachable_constraints_fv_eq: assumes P: "\T \ set P. admissible_transaction' T" "\T \ set P. admissible_transaction_no_occurs_msgs T" and A: "A \ reachable_constraints (map add_occurs_msgs P)" shows "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (rm_occurs_msgs_constr A) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using A proof (induction A rule: reachable_constraints.induct) case (step \ T \ \ \) let ?f = rm_occurs_msgs_constr let ?B = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" obtain T' where T': "T' \ set P" "T = add_occurs_msgs T'" using step.hyps(2) by auto have T_adm: "admissible_transaction' T" using add_occurs_msgs_admissible_occurs_checks(1) step.hyps(2) P by auto have T'_adm: "admissible_transaction' T'" and T'_no_occ: "admissible_transaction_no_occurs_msgs T'" using T'(1) P by (blast,blast) have "?f (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using rm_occurs_msgs_constr_transaction_strand''[OF T'_adm T'_no_occ, of \] admissible_transaction_decl_fresh_renaming_subst_not_occurs[OF T_adm step.hyps(3,4,5)] unfolding T'(2) \_def[symmetric] by blast moreover have "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using add_occurs_msgs_vars_eq(6)[OF T'_adm, of \] unfolding T'(2) by blast ultimately have "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?f ?B) = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?B" using fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unfolding T'(2) \_def[symmetric] by metis thus ?case using step.IH fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel ?B"] rm_occurs_msgs_constr_append[of \ ?B] by force qed simp private lemma rm_occurs_msgs_constr_reachable_constraints_vars_eq: assumes P: "\T \ set P. admissible_transaction' T" "\T \ set P. admissible_transaction_no_occurs_msgs T" and A: "A \ reachable_constraints (map add_occurs_msgs P)" shows "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (rm_occurs_msgs_constr A) = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using rm_occurs_msgs_constr_bvars_subst_eq[of _ Var] rm_occurs_msgs_constr_reachable_constraints_fv_eq[OF P A] by (metis vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t subst_lsst_id_subst) private lemma rm_occurs_msgs_constr_reachable_constraints_trms_cases_aux: assumes A: "x \ fv\<^sub>s\<^sub>s\<^sub>t A" "bvars\<^sub>s\<^sub>s\<^sub>t A = {}" and t: "t = occurs (\ x)" and \: "(\y. \ x = Var y) \ (\c. \ x = Fun c [])" shows "(\x \ fv\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>s\<^sub>t \). t = occurs (Var x)) \ (\c. Fun c [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>s\<^sub>t \) \ t = occurs (Fun c []))" using A proof (induction A) case (Cons a A) have 0: "bvars\<^sub>s\<^sub>s\<^sub>t A = {}" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) = {}" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) \ subst_domain \ = {}" using Cons.prems(2) by auto note 1 = fv\<^sub>s\<^sub>s\<^sub>t_Cons[of a A] trms\<^sub>s\<^sub>s\<^sub>t_cons[of a A] subst_sst_cons[of a A \] show ?case proof (cases "x \ fv\<^sub>s\<^sub>s\<^sub>t A") case False hence x: "x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" using Cons.prems(1) by simp note 2 = x t \ have 3: "\ x \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" using subst_subterms[OF fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_subterm_trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p[OF x]] trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst[OF 0(3)] by auto have "Fun c [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" when "\ x = Fun c []" for c using that 3 t by argo moreover have "y \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" when "\ x = Var y" for y using that 3 var_subterm_trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p[of y "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \"] 0(2) unfolding vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst by simp ultimately have "(\x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \). t = occurs (Var x)) \ (\c. Fun c [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ t = occurs (Fun c []))" using t \ by fast thus ?thesis using 1 by auto qed (use Cons.IH[OF _ 0(1)] 1 in force) qed simp private lemma rm_occurs_msgs_constr_reachable_constraints_trms_cases: assumes P: "\T \ set P. admissible_transaction' T" "\T \ set P. admissible_transaction_no_occurs_msgs T" and A: "A = rm_occurs_msgs_constr B" and B: "B \ reachable_constraints (map add_occurs_msgs P)" and t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B" shows "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ (\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. t = occurs (Var x)) \ (\c. Fun c [] \\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ t = occurs (Fun c []))" (is "?A A \ ?B A \ ?C A") proof - define rm_occs where "rm_occs \ \A::('fun,'atom,'sets,'lbl) prot_strand. rm_occurs_msgs_constr A" define Q where "Q \ \A. ?A A \ ?B A \ ?C A" have 0: "Q B" when "Q A" "set A \ set B" for A B using that unfolding Q_def fv\<^sub>s\<^sub>s\<^sub>t_def trms\<^sub>s\<^sub>s\<^sub>t_def unlabel_def by auto have "Q A" using B t unfolding A proof (induction rule: reachable_constraints.induct) case (step \ T \ \ \) define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" define \ where "\ \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" obtain T' where T': "T' \ set P" "T = add_occurs_msgs T'" using step.hyps(2) by auto note T'_adm = bspec[OF P(1) T'(1)] bspec[OF P(2) T'(1)] note T_adm = add_occurs_msgs_admissible_occurs_checks[OF T'_adm(1), unfolded T'(2)[symmetric]] note 1 = \_def[symmetric] \_def[symmetric] rm_occs_def[symmetric] note 2 = rm_occurs_msgs_constr_append[of \ \, unfolded rm_occs_def[symmetric]] note 3 = admissible_transaction_decl_fresh_renaming_subst_not_occurs[ OF T_adm(1) step.hyps(3,4,5)] have 4: "rm_occs (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using 3 rm_occurs_msgs_constr_transaction_strand''[OF T'_adm] unfolding T'(2) 1 by blast have 5: "(\y. \ x = Var y) \ (\c. \ x = Fun c [])" for x using transaction_decl_fresh_renaming_substs_range'(1)[OF step.hyps(3,4,5)] unfolding \_def[symmetric] by blast show ?case proof (cases "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \") case True show ?thesis using 0[OF step.IH[OF True]] unfolding 1 2 by simp next case False hence "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using step.prems unfolding \_def \_def by simp hence "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ (\x \ fv_transaction T'. t = occurs (\ x))" using add_occurs_msgs_in_trms_subst_cases[OF T'_adm(1), of t \] unfolding \_def trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq T'(2) by blast moreover have "(\y. \ x = Var y) \ (\c. \ x = Fun c [])" for x using transaction_decl_fresh_renaming_substs_range'(1)[OF step.hyps(3,4,5)] unfolding \_def[symmetric] by blast ultimately have "Q (rm_occs \)" using rm_occurs_msgs_constr_reachable_constraints_trms_cases_aux[ of _ "unlabel (transaction_strand T')" t \] admissible_transactionE(4)[OF T'_adm(1)] unfolding Q_def \_def 4 trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unlabel_subst by fast thus ?thesis using 0[of "rm_occs \"] unfolding 1 2 by auto qed qed simp thus ?thesis unfolding Q_def by blast qed private lemma rm_occurs_msgs_constr_receive_attack_iff: fixes A::"('a,'b,'c,'d) prot_strand" shows "(\ts. attack\n\ \ set ts \ receive\ts\ \ set (unlabel A)) \ (\ts. attack\n\ \ set ts \ receive\ts\ \ set (unlabel (rm_occurs_msgs_constr A)))" (is "(\ts. attack\n\ \ set ts \ ?A A ts) \ (\ts. attack\n\ \ set ts \ ?B A ts)") proof let ?att = "\ts. attack\n\ \ set ts" define f where "f \ \ts::('a,'b,'c,'d) prot_term list. filter (\t. \s. t \ occurs s) ts" have 0: "?att ts \ ?att (f ts)" "?att ts \ \t. occurs t \ set ts \ \t \ set ts. \s. t \ occurs s" "\t. occurs t \ set ts \ f ts = ts" for ts::"('a,'b,'c,'d) prot_term list" unfolding f_def subgoal by simp subgoal by auto subgoal by (induct ts) auto done have "?B A (f ts)" when A: "?A A ts" and ts: "?att ts" for ts using A proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) show ?case proof (cases "?A A ts") case True thus ?thesis using Cons.IH unfolding a by (cases b) simp_all next case False hence b: "b = receive\ts\" using Cons.prems unfolding a by simp show ?thesis using 0(2)[OF ts] 0(3) unfolding a b f_def by simp qed qed simp thus "(\ts. ?att ts \ ?A A ts) \ (\ts. ?att ts \ ?B A ts)" using 0(1) by fast have "\ts'. ts = f ts' \ ?A A ts'" when B: "?B A ts" and ts: "?att ts" for ts using B proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) note 1 = rm_occurs_msgs_constr_Cons have 2: "receive\ts\ \ set (unlabel (rm_occurs_msgs_constr A))" when rcv_ts: "receive\ts\ \ set (unlabel (rm_occurs_msgs_constr ((l,send\ts'\)#A)))" for l ts ts' and A::"('a,'b,'c,'d) prot_strand" proof - have *: "is_Send (send\ts'\)" by simp have "set (unlabel (rm_occurs_msgs_constr [(l, send\ts'\)])) \ {send\ts'\, send\f ts'\}" using 1(5-7)[OF *, of l "[]"] unfolding f_def by auto thus ?thesis using rcv_ts rm_occurs_msgs_constr_append[of "[(l,send\ts'\)]" A] by auto qed show ?case proof (cases "?B A ts") case True thus ?thesis using Cons.IH by auto next case False hence 3: "receive\ts\ \ set (unlabel (rm_occurs_msgs_constr [a]))" using rm_occurs_msgs_constr_append[of "[a]" A] Cons.prems by simp obtain ts' where b: "b = receive\ts'\" and b': "is_Receive b" using 2[of ts l _ A] Cons.prems False unfolding a by (cases b) auto have ts': "the_msgs (receive\ts'\) = ts'" by simp have "\t \ set (the_msgs b). \s. t \ occurs s" when "\t. occurs t \ set (the_msgs b)" using that 3 1(4)[OF b' that, of l "[]"] unfolding a by force hence "ts = f ts'" using 3 0(3)[of ts'] 1(3)[OF b', of l "[]", unfolded rm_occurs_msgs_constr.simps(1)] unfolding a b ts' f_def[symmetric] by fastforce thus ?thesis unfolding a b by auto qed qed simp thus "(\ts. ?att ts \ ?B A ts) \ (\ts. ?att ts \ ?A A ts)" using 0 by fast qed private lemma add_occurs_msgs_soundness_aux1: fixes P::"('fun,'atom,'sets,'lbl) prot" defines "wt_attack \ \\ \ l n. welltyped_constraint_model \ (\@[(l, send\[attack\n\]\)])" assumes P: "\T \ set P. admissible_transaction' T" and P_val: "has_initial_value_producing_transaction P" and A: "\ \ reachable_constraints P" "wt_attack \ \ l n" shows "\\ \ reachable_constraints P. \\. wt_attack \ \ l n \ (\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \n. \ x = Fun (Val n) [])" proof - let ?f = "\(T,\,\,\). dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" let ?g = "concat \ map ?f" let ?rcv_att = "\A n. receive\[attack\n\]\ \ set (unlabel A)" let ?wt_model = welltyped_constraint_model define valconst_cases where "valconst_cases \ \I::('fun,'atom,'sets,'lbl) prot_subst. \x. (\n. I x = Fun (Val n) []) \ (\n. I x = Fun (PubConst Value n) [])" define valconsts_only where "valconsts_only \ \X. \I::('fun,'atom,'sets,'lbl) prot_subst. \x \ X. \n. I x = Fun (Val n) []" define db_eq where "db_eq \ \A B::('fun,'atom,'sets,'lbl) prot_constr. \s. \upds::('fun,'atom,'sets,'lbl) prot_constr. let f = filter is_Update \ unlabel; g = filter (\a. \l t ts. a = (l,insert\t,Fun (Set s) ts\)) in (upds = [] \ f A = f B) \ (upds \ [] \ f (g A) = f (g B))" define db_upds_consts_fresh where "db_upds_consts_fresh \ \A::('fun,'atom,'sets,'lbl) prot_constr. \X. \J::('fun,'atom,'sets,'lbl) prot_subst. \x \ X. (\n. \ x = Fun (PubConst Value n) []) \ (\n s. insert\Fun (Val n) [],s\ \ set (unlabel A) \ delete\Fun (Val n) [],s\ \ set (unlabel A) \ J x \ Fun (Val n) [])" define subst_eq_on_privvals where "subst_eq_on_privvals \ \X \. \x \ X. (\n. \ x = Fun (Val n) []) \ \ x = \ x" define subst_in_ik_if_subst_pubval where "subst_in_ik_if_subst_pubval \ \X \. \\::('fun,'atom,'sets,'lbl) prot_constr. \x \ X. (\n. \ x = Fun (PubConst Value n) []) \ \ x \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" define subst_eq_iff where "subst_eq_iff \ \X. \\::('fun,'atom,'sets,'lbl) prot_subst. \x \ X. \y \ X. \ x = \ y \ \ x = \ y" obtain x_val T_val T_upds s_val ts_val l1_val l2_val where x_val: "T_val \ set P" "Var x_val \ set ts_val" "\\<^sub>v x_val = TAtom Value" "fv\<^sub>s\<^sub>e\<^sub>t (set ts_val) = {x_val}" "\n. \(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t set ts_val)" "T_val = Transaction (\(). []) [x_val] [] [] T_upds [(l1_val,send\ts_val\)]" "T_upds = [] \ (T_upds = [(l2_val,insert\Var x_val,\s_val\\<^sub>s\)] \ (\T \ set P. \(l,a) \ set (transaction_strand T). \t. a \ select\t,\s_val\\<^sub>s\ \ a \ \t in \s_val\\<^sub>s\ \ a \ \t not in \s_val\\<^sub>s\ \ a \ delete\t,\s_val\\<^sub>s\))" using has_initial_value_producing_transactionE[OF P_val P, of thesis] by (auto simp add: disj_commute) have 0: "\n. Fun (PubConst Value n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" when "\ \ reachable_constraints P" for \ using reachable_constraints_val_funs_private(1)[OF that P(1)] funs_term_Fun_subterm' unfolding is_PubConstValue_def by fastforce have I: "?wt_model \ \" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using welltyped_constraint_model_prefix[OF A(2)[unfolded wt_attack_def]] A(2)[unfolded wt_attack_def welltyped_constraint_model_def constraint_model_def] by blast+ have 1: "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. valconst_cases \ x" using reachable_constraints_fv_Value_const_cases[OF P(1) A(1) I(1)] unfolding valconst_cases_def by blast have 2: "?rcv_att \ n" using A(2) strand_sem_append_stateful[of "{}" "{}" "unlabel \" "[send\[attack\n\]\]" \] reachable_constraints_receive_attack_if_attack'(2)[OF A(1) P(1) I(1)] unfolding wt_attack_def welltyped_constraint_model_def constraint_model_def by simp note \_empty = admissible_transaction_decl_subst_empty[OF bspec[OF P(1)]] have lmm: "\\ \ reachable_constraints P. \\. ?wt_model \ \ \ valconsts_only (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X) \ \ (?rcv_att \ n \ ?rcv_att \ n) \ subst_eq_on_privvals (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X) \ \ subst_in_ik_if_subst_pubval (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X) \ \ \ subst_eq_iff (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X) \ \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ (\n \ N. \(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ db_eq \ \ s_val T_upds \ db_upds_consts_fresh \ (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X) \" when "finite N" "\n \ N. \(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "X \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" "finite X" "\x \ X. valconst_cases \ x" "\x \ X. \\<^sub>v x = TAtom Value" for N X using A(1) I(1) 1 that proof (induction arbitrary: N X rule: reachable_constraints.induct) case init define pubvals where "pubvals \ {n | n x. x \ X \ \ x = Fun (PubConst Value n) []}" define X_vals where "X_vals \ {n | n x. x \ X \ \ x = Fun (Val n) []}" have X_vals_finite: "finite X_vals" using finite_surj[OF init.prems(6), of X_vals "\x. THE n. \ x = Fun (Val n) []"] unfolding X_vals_def by force have pubvals_finite: "finite pubvals" using finite_surj[OF init.prems(6), of pubvals "\x. THE n. \ x = Fun (PubConst Value n) []"] unfolding pubvals_def by force obtain T_val_fresh_vals and \::"nat \ nat" where T_val_fresh_vals: "T_val_fresh_vals \ (N \ X_vals) = {}" and \: "inj \" "\ ` pubvals = T_val_fresh_vals" using ex_finite_disj_nat_inj[OF pubvals_finite finite_UnI[OF init.prems(3) X_vals_finite]] by blast have T_val_fresh_vals_finite: "finite T_val_fresh_vals" using pubvals_finite \(2) by blast obtain \::"('fun,'atom,'sets,'lbl) prot_constr" where B: "\ \ reachable_constraints P" "T_upds = [] \ list_all is_Receive (unlabel \)" "T_upds \ [] \ list_all (\a. is_Insert a \ is_Receive a) (unlabel \)" "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" "\n. Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ Fun (Val n) [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "T_val_fresh_vals = {n. Fun (Val n) [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \}" "\l a. (l,a) \ set \ \ is_Insert a \ (l = l2_val \ (\n. a = insert\Fun (Val n) [],\s_val\\<^sub>s\))" using reachable_constraints_initial_value_transaction[ OF P reachable_constraints.init T_val_fresh_vals_finite _ x_val] by auto define \ where "\ \ \x. if x \ X \ (\n. \ x = Fun (PubConst Value n) []) then Fun (Val (\ (THE n. \ x = Fun (PubConst Value n) []))) [] else \ x" have 0: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "?rcv_att [] n \ ?rcv_att \ n" "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using B(4) vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] by auto have 1: "db_eq [] \ s_val T_upds" using B(2,3,7) proof (induction \) case (Cons a B) then obtain l b where a: "a = (l,b)" by (metis surj_pair) have IH: "db_eq [] B s_val T_upds" using Cons.prems Cons.IH by auto show ?case proof (cases "T_upds = []") case True hence "is_Receive b" using a Cons.prems(1) by simp thus ?thesis using IH unfolding a db_eq_def Let_def by auto next case False hence "is_Insert b \ is_Receive b" using a Cons.prems(2) by simp hence "\t. a = (l2_val,insert\t,\s_val\\<^sub>s\)" when b: "\is_Receive b" using b Cons.prems(3) unfolding a by (metis list.set_intros(1)) thus ?thesis using IH False unfolding a db_eq_def Let_def by auto qed qed (simp add: db_eq_def) have 2: "?wt_model \ \" unfolding welltyped_constraint_model_def constraint_model_def proof (intro conjI) show "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using I(4) init.prems(8) unfolding \_def wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by fastforce show "strand_sem_stateful {} {} (unlabel \) \" using B(2,3) strand_sem_stateful_if_no_send_or_check[of "unlabel \" "{}" "{}" \] unfolding list_all_iff by blast show "subst_domain \ = UNIV" "ground (subst_range \)" using I(2) unfolding \_def subst_domain_def by auto show "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using I(3) unfolding \_def by fastforce qed have 3: "valconsts_only (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] \ X) \" using init.prems(7) unfolding \_def valconsts_only_def valconst_cases_def by fastforce have 4: "subst_eq_on_privvals (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] \ X) \" unfolding subst_eq_on_privvals_def \_def by force have 5: "subst_in_ik_if_subst_pubval (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] \ X) \ \" proof (unfold subst_in_ik_if_subst_pubval_def; intro ballI impI) fix x assume x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] \ X" and "\n. \ x = Fun (PubConst Value n) []" then obtain n where n: "\ x = Fun (PubConst Value n) []" by blast have "n \ pubvals" using x n unfolding pubvals_def by fastforce hence "\ n \ T_val_fresh_vals" using \(2) by fast hence "Fun (Val (\ n)) [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using B(6) by fast thus "\ x \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using x n unfolding \_def by simp qed have 6: "subst_eq_iff (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] \ X) \" proof (unfold subst_eq_iff_def; intro ballI) fix x y assume "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] \ X" "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] \ X" hence x: "x \ X" and y: "y \ X" by auto show "\ x = \ y \ \ x = \ y" proof show "\ x = \ y \ \ x = \ y" using x y unfolding \_def by presburger next assume J_eq: "\ x = \ y" show "\ x = \ y" proof (cases "\n. \ x = Fun (PubConst Value n) []") case True then obtain n where n: "\ x = Fun (PubConst Value n) []" by blast hence J_x: "\ x = Fun (Val (\ n)) []" using x unfolding \_def by simp show ?thesis proof (cases "\m. \ y = Fun (PubConst Value m) []") case True then obtain m where m: "\ y = Fun (PubConst Value m) []" by blast have J_y: "\ y = Fun (Val (\ m)) []" using y m unfolding \_def by simp show ?thesis using J_eq J_x J_y injD[OF \(1), of n m] n m by auto next case False then obtain m where m: "\ y = Fun (Val m) []" using init.prems(7) y unfolding valconst_cases_def by blast moreover have "\ n \ T_val_fresh_vals" using \(2) x n unfolding pubvals_def by blast moreover have "m \ X_vals" using y m unfolding X_vals_def by blast ultimately have "\ x \ \ y" using m J_x T_val_fresh_vals by auto moreover have "\ y = \ y" using m unfolding \_def by simp ultimately show ?thesis using J_eq by argo qed next case False then obtain n where n: "\ x = Fun (Val n) []" using init.prems(7) x unfolding valconst_cases_def by blast hence J_x: "\ x = \ x" unfolding \_def by auto show ?thesis proof (cases "\m. \ y = Fun (PubConst Value m) []") case False then obtain m where m: "\ y = Fun (Val m) []" using init.prems(7) y unfolding valconst_cases_def by blast have J_y: "\ y = \ y" using y m unfolding \_def by simp show ?thesis using J_x J_y J_eq by presburger next case True then obtain m where m: "\ y = Fun (PubConst Value m) []" by blast hence "\ y = Fun (Val (\ m)) []" using y unfolding \_def by fastforce moreover have "\ m \ T_val_fresh_vals" using \(2) y m unfolding pubvals_def by blast moreover have "n \ X_vals" using x n unfolding X_vals_def by blast ultimately have "\ y \ \ x" using n J_x T_val_fresh_vals by auto thus ?thesis using J_x J_eq by argo qed qed qed qed have 7: "\n \ N. Fun (Val n) [] \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using B(5,6) T_val_fresh_vals by blast have 8: "db_upds_consts_fresh [] (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] \ X) \" unfolding db_upds_consts_fresh_def by simp show ?case using B(1) 0 1 2 3 4 5 6 7 8 by blast next case (step \ T \ \ \ N X') define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" define T'_pubvals where "T'_pubvals \ {n. \x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'. \ x = Fun (PubConst Value n) []}" define \_vals where "\_vals \ {n. Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \}" define \_vals where "\_vals \ {n. \x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'. \ x = Fun (Val n) []}" define \_vals where "\_vals \ {n. Fun (Val n) [] \ subst_range \}" have 3: "welltyped_constraint_model \ \" "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. valconst_cases \ x" "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'. valconst_cases \ x" "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}) (unlabel T') \" "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. valconst_cases \ x" using step.prems(2) welltyped_constraint_model_prefix[OF step.prems(1)] step.prems(1)[unfolded welltyped_constraint_model_def constraint_model_def] strand_sem_append_stateful[of "{}" "{}" "unlabel \" "unlabel T'" \] step.prems(7) unfolding \_def[symmetric] T'_def[symmetric] unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by (blast,blast,blast,simp,blast) note T_adm = bspec[OF P step.hyps(2)] note T_wf = admissible_transaction_is_wellformed_transaction[OF T_adm] note \_empty = admissible_transaction_decl_subst_empty[OF T_adm step.hyps(3)] note 4 = admissible_transaction_sem_iff note 5 = iffD1[OF 4[OF T_adm I(2,3), of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" \, unfolded T'_def[symmetric]] 3(4)] note \_dom = transaction_fresh_subst_domain[OF step.hyps(4)] have \_ran: "\n. t = Fun (Val n) []" when t: "t \ subst_range \" for t proof - obtain x where x: "x \ set (transaction_fresh T)" "t = \ x" using \_dom t by auto show ?thesis using x(1) admissible_transactionE(2)[OF T_adm] transaction_fresh_subst_sends_to_val[OF step.hyps(4) x(1)] unfolding x(2) by meson qed have T'_vals_in_\: "Fun (Val k) [] \ subst_range \" when k: "Fun (Val k) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" for k proof - have "Fun (Val k) [] \ (subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction T)) \\<^sub>s\<^sub>e\<^sub>t \" using k admissible_transactionE(4)[OF T_adm] transaction_decl_fresh_renaming_substs_trms[ OF step.hyps(3,4,5), of "transaction_strand T"] unfolding T'_def \_def[symmetric] trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq by fast then obtain t where t: "t \\<^sub>s\<^sub>e\<^sub>t trms_transaction T" "t \ \ = Fun (Val k) []" by force hence "Fun (Val k) [] \ subst_range \" using admissible_transactions_no_Value_consts(1)[OF T_adm] by (cases t) force+ thus ?thesis using transaction_decl_fresh_renaming_substs_range'(4)[OF step.hyps(3,4,5)] \_empty unfolding \_def[symmetric] by blast qed have \_vals_is_T'_vals: "k \ \_vals \ Fun (Val k) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" for k proof show "k \ \_vals" when "Fun (Val k) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" using that T'_vals_in_\ unfolding \_vals_def by blast show "Fun (Val k) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" when k: "k \ \_vals" proof - have "Fun (Val k) [] \ subst_range \" using k unfolding \_vals_def by fast then obtain x where x: "x \ fv_transaction T" "\ x = Fun (Val k) []" using admissible_transactionE(7)[OF T_adm] transaction_fresh_subst_domain[OF step.hyps(4)] unfolding fv_transaction_unfold by fastforce have "\ x = Fun (Val k) []" using x(2) unfolding \_def \_empty subst_compose_def by auto thus ?thesis using fv\<^sub>s\<^sub>s\<^sub>t_is_subterm_trms\<^sub>s\<^sub>s\<^sub>t_subst[OF x(1), of \] admissible_transactionE(4)[OF T_adm] unfolding T'_def trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unlabel_subst by simp qed qed have \_vals_N_disj: "N \ \_vals = {}" using step.prems(4) \_vals_is_T'_vals unfolding \_def[symmetric] T'_def[symmetric] unlabel_append trms\<^sub>s\<^sub>s\<^sub>t_append by blast have T'_pubvals_finite: "finite T'_pubvals" using finite_surj[OF fv\<^sub>s\<^sub>s\<^sub>t_finite[of "unlabel T'"], of T'_pubvals "\x. THE n. \ x = Fun (PubConst Value n) []"] unfolding T'_pubvals_def by force have \_vals_finite: "finite \_vals" proof - have *: "finite (subst_range \)" using transaction_fresh_subst_domain[OF step.hyps(4)] by simp show ?thesis using finite_surj[OF *, of \_vals "\t. THE n. t = Fun (Val n) []"] unfolding \_vals_def by force qed have \_vals_finite: "finite \_vals" proof - have *: "\_vals \ (\t. THE n. t = Fun (Val n) []) ` subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" unfolding \_vals_def by force show ?thesis by (rule finite_surj[OF subterms_union_finite[OF trms\<^sub>s\<^sub>s\<^sub>t_finite] *]) qed have \_vals_finite: "finite \_vals" proof - define X where "X \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" have *: "finite X" using fv\<^sub>s\<^sub>s\<^sub>t_finite step.prems(6) unfolding X_def by blast show ?thesis using finite_surj[OF *, of \_vals "\x. THE n. \ x = Fun (Val n) []"] unfolding \_vals_def X_def[symmetric] by force qed obtain T_val_fresh_vals and \::"nat \ nat" where T_val_fresh_vals: "T_val_fresh_vals \ (N \ \_vals \ \_vals \ \_vals) = {}" and \: "inj \" "\ ` T'_pubvals = T_val_fresh_vals" using step.prems(3) T'_pubvals_finite \_vals_finite \_vals_finite \_vals_finite by (metis finite_UnI ex_finite_disj_nat_inj) define N' where "N' \ N \ \_vals \ T_val_fresh_vals" have T_val_fresh_vals_finite: "finite T_val_fresh_vals" using T'_pubvals_finite \(2) by blast have N'_finite: "finite N'" using step.prems(3) T'_pubvals_finite T_val_fresh_vals_finite \_vals_finite unfolding N'_def by auto have \_vals_trms_in: "n \ \_vals" when "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" for n using that unfolding \_vals_def by blast have N'_notin_\: "\(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" when n: "n \ N'" for n proof - have ?thesis when n': "n \ N" using n' step.prems(4) unfolding N'_def unlabel_append trms\<^sub>s\<^sub>s\<^sub>t_append by blast moreover have ?thesis when n': "n \ \_vals" using n' step.hyps(4) unfolding \_vals_def transaction_fresh_subst_def by blast moreover have ?thesis when n': "n \ T_val_fresh_vals" using n' T_val_fresh_vals \_vals_trms_in by blast ultimately show ?thesis using n unfolding N'_def by blast qed have T'_fv_\_disj: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = {}" using transaction_decl_fresh_renaming_substs_vars_disj(8)[OF step.hyps(3,4,5)] transaction_decl_fresh_renaming_substs_vars_subset(4)[OF step.hyps(3,4,5,2)] unfolding \_def[symmetric] T'_def fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq by blast have X'_disj: "X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" "X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = {}" using step.prems(5) unfolding \_def[symmetric] T'_def[symmetric] unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by (blast, blast) have X'_disj': "(X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T') \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" using X'_disj(1) T'_fv_\_disj by blast have X'_finite: "finite (X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" using step.prems(6) fv\<^sub>s\<^sub>s\<^sub>t_finite by blast have \_X'_valconstcases: "\x \ X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'. valconst_cases \ x" using 3(3,5) by blast have T'_value_vars: "\\<^sub>v x = TAtom Value" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" for x using x reachable_constraints_fv_Value_typed[ OF P reachable_constraints.step[OF step.hyps]] unfolding \_def[symmetric] T'_def[symmetric] unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast have X'_T'_value_vars: "\x \ X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'. \\<^sub>v x = TAtom Value" using step.prems(8) T'_value_vars by blast have N'_not_subterms_\: "\n \ N'. \(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using N'_notin_\ by blast obtain \ \ where B: "\ \ reachable_constraints P" "?wt_model \ \" "valconsts_only (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T') \" "?rcv_att \ n \ ?rcv_att \ n" "subst_eq_on_privvals (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T') \" "subst_in_ik_if_subst_pubval (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T') \ \" "subst_eq_iff (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T') \" "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "\n \ N'. \(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "db_eq \ \ s_val T_upds" "db_upds_consts_fresh \ (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T') \" using step.IH[OF 3(1,2) N'_finite N'_not_subterms_\ X'_disj' X'_finite \_X'_valconstcases X'_T'_value_vars] unfolding Un_assoc by fast have J: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "constr_sem_stateful \ (unlabel \)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using B(2) unfolding welltyped_constraint_model_def constraint_model_def by blast+ have T_val_fresh_vals_notin_\: "\(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" when "n \ T_val_fresh_vals" for n using that B(12) unfolding N'_def by blast hence "\n \ T_val_fresh_vals. \(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by blast then obtain T_val_constr::"('fun,'atom,'sets,'lbl) prot_constr" where T_val_constr: "\@T_val_constr \ reachable_constraints P" "T_val_constr \ reachable_constraints P" "T_upds = [] \ list_all is_Receive (unlabel T_val_constr)" "T_upds \ [] \ list_all (\a. is_Insert a \ is_Receive a) (unlabel T_val_constr)" "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr = {}" "\n. Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ Fun (Val n) [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr" "\n. Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr \ Fun (Val n) [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr" "T_val_fresh_vals = {n. Fun (Val n) [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr}" "\l a. (l,a) \ set T_val_constr \ is_Insert a \ (l = l2_val \ (\n. a = insert\Fun (Val n) [],\s_val\\<^sub>s\))" using reachable_constraints_initial_value_transaction[ OF P B(1) T_val_fresh_vals_finite _ x_val] by blast have T_val_constr_no_upds_if_no_T_upds: "filter is_Update (unlabel T_val_constr) = []" when "T_upds = []" using T_val_constr(3)[OF that] by (induct T_val_constr) auto have T_val_fresh_vals_is_T_val_constr_vals: "k \ T_val_fresh_vals \ Fun (Val k) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr" for k using that T_val_constr(7,8) ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset by fast have T_val_constr_no_fv: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr = {}" using T_val_constr(5) vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t by fast have T_val_\: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T_val_constr))" proof - have "\(t \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T_val_constr))" when t: "t \ subst_range \" for t proof - obtain k where k: "t = Fun (Val k) []" using t \_ran by fast have "k \ \_vals" using t unfolding k \_vals_def by blast thus ?thesis using B(12) T_val_fresh_vals T_val_constr(7,8) unfolding N'_def k unlabel_append trms\<^sub>s\<^sub>s\<^sub>t_append by blast qed thus ?thesis using step.hyps(4) unfolding transaction_fresh_subst_def by fast qed have T_val_\: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T_val_constr))" using step.hyps B(8) T_val_constr(5) by auto define \' where "\' \ \@T_val_constr@T'" define \ where "\ \ \x. if x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' then if \n. \ x = Fun (PubConst Value n) [] then if \y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ y = \ x then \ (SOME y. y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X' \ \ y = \ x) else Fun (Val (\ (THE n. \ x = Fun (PubConst Value n) []))) [] else \ x else \ x" have \_ground_ran: "ground (subst_range \)" "range_vars \ = {}" and \_ran_bvars_disj: "range_vars \ \ bvars_transaction T = {}" using transaction_fresh_subst_domain[OF step.hyps(4)] transaction_fresh_subst_range_vars_empty[OF step.hyps(4)] transaction_decl_subst_range_vars_empty[OF step.hyps(3)] by (metis range_vars_alt_def, argo, blast) have \_T'_fv_disj: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = {}" using T'_fv_\_disj unfolding B(9) by argo hence \_\_fv_\_eq: "\ x = \ x" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" for x using x X'_disj unfolding \_def by auto have B'1: "\' \ reachable_constraints P" using reachable_constraints.step[OF T_val_constr(1) step.hyps(2,3) T_val_\ T_val_\] unfolding \'_def T'_def \_def by simp have "\n. \ x = Fun (Val n) []" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ X'" for x proof (cases "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'") case True note T = this show ?thesis proof (cases "\n. \ x = Fun (PubConst Value n) []") case True thus ?thesis using T B(3,9) someI_ex[of "\y. y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X' \ \ y = \ x"] unfolding \_def valconsts_only_def by (cases "\y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ y = \ x") (meson, auto) next case False thus ?thesis using T 3(3) unfolding \_def valconst_cases_def by fastforce qed next case False thus ?thesis using x B(3) unfolding \_def valconsts_only_def by auto qed hence B'3: "valconsts_only (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ X') \" unfolding valconsts_only_def by blast have B'4: "?rcv_att \' n" when "?rcv_att (\@T') n" using that B(4) unfolding \'_def by auto have "\ x = \ x" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ X'" "\ x = Fun (Val n) []" for x n proof - have "\ x = \ x" when "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" using that unfolding \_def by meson moreover have "\ x = \ x" when "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" using that x unfolding \_def by simp ultimately show ?thesis using B(5) x unfolding subst_eq_on_privvals_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by (cases "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'") auto qed hence B'5: "subst_eq_on_privvals (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ X') \" unfolding subst_eq_on_privvals_def by blast have \_fv_\_eq_\: "\ x = \ x" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" for x proof - have "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" using x T'_fv_\_disj X'_disj by blast thus ?thesis unfolding \_def by argo qed have T'_fv_\_val_\_eq_\: "\ x = \ x" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "\n. \ x = Fun (PubConst Value n) []" for x using x B'5 3(3) unfolding unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append valconst_cases_def \_def by meson have T'_fv_\_pubval_\_eq_\_fresh_val: "\ x = Fun (Val (\ n)) []" "\ n \ T_val_fresh_vals" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "\ x = Fun (PubConst Value n) []" "\y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ y \ \ x" for x n proof - show "\ x = Fun (Val (\ n)) []" using x unfolding \_def by auto show "\ n \ T_val_fresh_vals" using \(2) x unfolding T'_pubvals_def by blast qed have T'_fv_\_pubval_\_eq_\_val: "\y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \m. \ y = \ x \ \ x = \ y \ \ x = Fun (Val m) []" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "\ x = Fun (PubConst Value n) []" "\y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ y = \ x" for x n proof - have "\ x = \ (SOME y. y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X' \ \ y = \ x)" using x unfolding \_def by meson then obtain y where y: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" "\ y = \ x" "\ x = \ y" using x(3) someI_ex[of "\y. y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X' \ \ y = \ x"] by blast thus ?thesis using B(3,9) unfolding valconsts_only_def by auto qed have T'_fv_\_pubval_\_eq_val: "\n. \ x = Fun (Val n) []" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "\ x = Fun (PubConst Value n) []" for x n using T'_fv_\_pubval_\_eq_\_fresh_val[OF x] T'_fv_\_pubval_\_eq_\_val[OF x] by auto have B'6': "\ x \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" "\ x = Fun (PubConst Value n) []" for x n using x B(6) \_fv_\_eq_\ x(2) unfolding B(8) subst_in_ik_if_subst_pubval_def by simp have B'6'': "\ x \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "\ x = Fun (PubConst Value n) []" for x n proof (cases "\y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ y = \ x") case True thus ?thesis using B(6) x(2) T'_fv_\_pubval_\_eq_\_val[OF x True] unfolding B(9) subst_in_ik_if_subst_pubval_def by force next case False thus ?thesis using T_val_constr(8) T'_fv_\_pubval_\_eq_\_fresh_val[OF x] by force qed have "\ x \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \'" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ X'" "\ x = Fun (PubConst Value n) []" for x n proof (cases "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'") case True thus ?thesis using B'6''[OF _ x(2)] unfolding \'_def by auto next case False hence "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" using x(1) unfolding unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast thus ?thesis using B'6' x(2) unfolding \'_def by simp qed hence B'6: "subst_in_ik_if_subst_pubval (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ X') \ \'" unfolding subst_in_ik_if_subst_pubval_def by blast have B'7: "subst_eq_iff (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ X') \" proof (unfold subst_eq_iff_def; intro ballI) fix x y assume xy: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ X'" "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ X'" let ?Q = "\x y. \ x = \ y \ \ x = \ y" have *: "?Q x y" when xy: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" for x y using B(7) xy unfolding \_def subst_eq_iff_def by force have **: "?Q x y" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" and y: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" for x y proof - have xy_neq: "x \ y" using x y T'_fv_\_disj X'_disj by blast have x_eq: "\ x = \ x" using \_fv_\_eq_\ x by blast have x_eq_if_val: "\ x = \ x" when "\ x = Fun (Val n) []" for n using that x B(5) unfolding subst_eq_on_privvals_def by blast have x_neq_if_neq_val: "\ x \ \ x" when "\ x = Fun (PubConst Value n) []" for n by (metis that B(3) x UnI1 prot_fun.distinct(37) term.inject(2) valconsts_only_def) have y_eq_if_val: "\ y = \ y" when "\ y = Fun (Val n) []" for n using that y B'5 unfolding subst_eq_on_privvals_def by simp have y_eq: "\ y = Fun (Val (\ n)) []" when "\ y = Fun (PubConst Value n) []" "\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ z \ \ y" for n by (rule T'_fv_\_pubval_\_eq_\_fresh_val(1)[OF y that]) have y_eq': "\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \m. \ z = \ y \ \ y = \ z \ \ y = Fun (Val m) []" when "\ y = Fun (PubConst Value n) []" "\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ z = \ y" for n by (rule T'_fv_\_pubval_\_eq_\_val[OF y that]) have K_eq_if_I_eq: "\ x = \ y \ \ x = \ y" apply (cases "\n. \ x = Fun (PubConst Value n) []") subgoal using B(7,9) unfolding subst_eq_iff_def by (metis UnI1 x x_eq y_eq') subgoal by (metis x_eq x T'_fv_\_val_\_eq_\[OF y] 3(5) valconst_cases_def x_eq_if_val) done have K_neq_if_I_neq_val: "\ x \ \ y" when n: "\ y = Fun (Val n) []" and m: "\ x = Fun (PubConst Value m) []" for n m proof - have I_neq: "\ x \ \ y" using n m by simp note y_eq'' = y_eq_if_val[OF n] note x_neq = x_neq_if_neq_val[OF m] have x_ex: "\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ z = \ x" using x unfolding B(9) by blast have J1: "\ x \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using B(6) x m unfolding subst_in_ik_if_subst_pubval_def by fast have J2: "\ x \ \ x" by (metis m B(3) x UnI1 prot_fun.distinct(37) term.inject(2) valconsts_only_def) have J3: "\ y = \ y" using B(5) n y unfolding subst_eq_on_privvals_def by blast have K_x: "\ x \ \ x" using J2 x_eq by presburger have x_notin: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" using x T'_fv_\_disj X'_disj by blast have K_x': "\ x = \ x" using x_notin unfolding \_def by argo have K_y: "\ y = \ y" using y_eq'' J3 by argo have "\ x \ \ y" using I_neq x y B(7) unfolding subst_eq_iff_def by blast thus ?thesis using K_x' K_y by argo qed show ?thesis proof show "\ x = \ y \ \ x = \ y" by (rule K_eq_if_I_eq) next assume xy_eq: "\ x = \ y" show "\ x = \ y" proof (cases "\n. \ y = Fun (PubConst Value n) []") case True then obtain n where n: "\ y = Fun (PubConst Value n) []" by blast show ?thesis proof (cases "\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ z = \ y") case True thus ?thesis using B(7,9) unfolding subst_eq_iff_def by (metis xy_eq UnI1 x x_eq y_eq'[OF n]) next case False hence F: "\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ z \ \ y" by blast note y_eq'' = y_eq[OF n F] have n_in: "\ n \ T_val_fresh_vals" using \(2) x_eq xy_eq T_val_fresh_vals_notin_\ y n unfolding T'_pubvals_def by blast hence y_notin: "\(\ y \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using T_val_fresh_vals_notin_\ y_eq'' ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset[of "unlabel \"] by auto show ?thesis proof (cases "\m. \ x = Fun (PubConst Value m) []") case True thus ?thesis using y_notin B(6) x xy_eq x_eq unfolding B(9) subst_in_ik_if_subst_pubval_def by fastforce next case False then obtain m where m: "\ x = Fun (Val m) []" using 3(5) x unfolding valconst_cases_def by fast hence "\ x = \ x" using x B(5) unfolding subst_eq_on_privvals_def by blast hence "\ x = Fun (Val m) []" using m x_eq by argo moreover have "m \ T_val_fresh_vals" using m T_val_fresh_vals x unfolding \_vals_def by blast hence "m \ \ n" using n_in by blast ultimately have False using xy_eq y_eq'' by force thus ?thesis by simp qed qed next case False then obtain n where n: "\ y = Fun (Val n) []" using 3(3) y unfolding valconst_cases_def by fast note y_eq'' = y_eq_if_val[OF n] show ?thesis proof (cases "\m. \ x = Fun (Val m) []") case True thus ?thesis by (metis xy_eq x_eq y_eq'' x_eq_if_val) next case False then obtain m where m: "\ x = Fun (PubConst Value m) []" using 3(5) x unfolding valconst_cases_def by blast show ?thesis using K_neq_if_I_neq_val[OF n m] xy_eq by blast qed qed qed qed have ***: "?Q x y" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" and y: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" for x y proof assume xy_eq: "\ x = \ y" show "\ x = \ y" proof (cases "\n. \ x = Fun (PubConst Value n) []") case True thus ?thesis using xy_eq x y B(7) T'_fv_\_pubval_\_eq_\_fresh_val(1) T'_fv_\_pubval_\_eq_\_val unfolding B(9) subst_eq_iff_def by (metis (no_types) UnI1) qed (metis xy_eq x y T'_fv_\_val_\_eq_\) next assume xy_eq: "\ x = \ y" have case1: False when x': "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" and y': "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" and xy_eq': "\ x = \ y" and m: "\ x = Fun (PubConst Value m) []" and n: "\ y = Fun (Val n) []" for x y n m proof - have F: "\n. \ y = Fun (PubConst Value n) []" using n by auto note x_eq = T'_fv_\_pubval_\_eq_\_fresh_val[OF x' m] note y_eq = T'_fv_\_val_\_eq_\[OF y' F] have "\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ z = \ x" proof (rule ccontr) assume no_z: "\(\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ z = \ x)" hence "n \ \ m" using n y' x_eq(2) T_val_fresh_vals unfolding \_vals_def by blast thus False using xy_eq' x_eq y_eq(1) n no_z by auto qed then obtain z k where z: "z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" "\ z = \ x" "\ x = \ z" "\ x = Fun (Val k) []" using T'_fv_\_pubval_\_eq_\_val[OF x' m] by blast have "\ y = \ z" using z(2,3) y_eq xy_eq' by presburger hence "\ x = \ y" using z(1,2) ** B(9) \_\_fv_\_eq y' y_eq by metis thus False using n m by simp qed have case2: "m = n" when x': "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" and y': "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" and xy_eq': "\ x = \ y" and m: "\ x = Fun (PubConst Value m) []" and n: "\ y = Fun (PubConst Value n) []" for x y n m proof (cases "\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ z \ \ x") case True show ?thesis apply (cases "\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ z \ \ y") subgoal using xy_eq' m n T'_fv_\_pubval_\_eq_\_fresh_val[OF x' m True] T'_fv_\_pubval_\_eq_\_fresh_val[OF y' n] injD[OF \(1), of m n] by fastforce subgoal by (metis x' y' xy_eq' ** B(9) True) done qed (metis x' y' xy_eq' m n ** B(9) prot_fun.inject(6) term.inject(2)) have case3: "m = n" when x': "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" and y': "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" and xy_eq': "\ x = \ y" and m: "\ x = Fun (Val m) []" and n: "\ y = Fun (Val n) []" for x y n m using x' y' xy_eq' m n T'_fv_\_val_\_eq_\ by fastforce show "\ x = \ y" using x y xy_eq case1 case2 case3 3(3) unfolding valconst_cases_def by metis qed have ****: "?Q x y" when xy: "x \ X'" "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" for x y proof - have xy': "y \ X'" "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using xy X'_disj T'_fv_\_disj by (blast, blast, blast) have K_x: "\ x = \ x" using xy(2) unfolding \_def by argo have I_iff_J: "\ x = \ y \ \ x = \ y" using xy B(7) unfolding subst_eq_iff_def by fast show ?thesis using K_x I_iff_J K_x injD[OF \(1)] xy B(7,9) ** by (meson UnCI) qed show "?Q x y" using xy * **[of x y] **[of y x] ***[of x y] ****[of x y] ****[of y x] unfolding unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by (metis Un_iff) qed have B'8_9: "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \'" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \'" using B(8,9) T_val_constr(5) vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel T_val_constr"] unfolding \'_def unlabel_append vars\<^sub>s\<^sub>s\<^sub>t_append fv\<^sub>s\<^sub>s\<^sub>t_append by simp_all have I': "\n. \ x = Fun (PubConst Value n) []" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "\n. \ x = Fun (Val n) []" for x using x 3(3) unfolding valconst_cases_def by fast have B'10_11: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \'" "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \'" using B(10,11) unfolding N'_def \'_def unlabel_append trms\<^sub>s\<^sub>s\<^sub>t_append ik\<^sub>s\<^sub>s\<^sub>t_append by (blast, blast) have B'12: "\n \ N. \(Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \')" using B(12) \_vals_is_T'_vals \_vals_N_disj T_val_fresh_vals T_val_fresh_vals_is_T_val_constr_vals unfolding N'_def \'_def unfolding unlabel_append trms\<^sub>s\<^sub>s\<^sub>t_append ik\<^sub>s\<^sub>s\<^sub>t_append by fast have B'13: "db_eq (\@T') \' s_val T_upds" proof - let ?f = "filter is_Update \ unlabel" let ?g = "filter (\a. \l t ts. a = (l, insert\t,Fun (Set s_val) ts\))" have "?f (?g T_val_constr) = []" using T_val_constr(3,4,9) proof (induction T_val_constr) case (Cons a B) then obtain l b where a: "a = (l,b)" by (metis surj_pair) have IH: "?f (?g B) = []" using Cons.prems Cons.IH by auto show ?case proof (cases "T_upds = []") case True hence "is_Receive b" using a Cons.prems(1,2) by simp thus ?thesis using IH unfolding a Let_def by auto next case False hence "is_Insert b \ is_Receive b" using a Cons.prems(1,2) by simp hence "\t. a = (l2_val, insert\t,\s_val\\<^sub>s\)" when b: "\is_Receive b" using b Cons.prems(3) unfolding a by (metis list.set_intros(1)) thus ?thesis using IH unfolding a Let_def by auto qed qed simp hence "?f (?g (\@T')) = ?f (?g \)@?f (?g T')" "?f (?g (\@T_val_constr@T')) = ?f (?g \)@?f (?g T')" when "T_upds \ []" by simp_all moreover have "?f T_val_constr = []" when "T_upds = []" using T_val_constr_no_upds_if_no_T_upds[OF that] by force hence "?f (\@T') = ?f \@?f T'" "?f (\@T_val_constr@T') = ?f \@?f T'" when "T_upds = []" using that by auto ultimately show ?thesis using B(13) unfolding \'_def db_eq_def Let_def by presburger qed have B'14: "db_upds_consts_fresh (\@T') (fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ X') \" proof (unfold db_upds_consts_fresh_def; intro ballI allI impI; elim exE) fix x s n m assume x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ X'" and n: "insert\Fun (Val n) [],s\ \ set (unlabel (\@T')) \ delete\Fun (Val n) [],s\ \ set (unlabel (\@T'))" (is "?A (\@T')") and m: "\ x = Fun (PubConst Value m) []" have A_cases: "?A \ \ ?A T'" using n by force have n_in_case: "n \ \_vals" when A: "?A T'" proof - obtain t s' where t: "insert\t,s'\ \ set (unlabel (transaction_strand T)) \ delete\t,s'\ \ set (unlabel (transaction_strand T))" "Fun (Val n) [] = t \ \" using A dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(4,5) stateful_strand_step_mem_substD(4,5)[of _ _ _ \] subst_lsst_unlabel[of _ \] unfolding T'_def by (metis (no_types, opaque_lifting)) have "Fun (Val n) [] \ subst_range \" using t transaction_inserts_are_Value_vars(1)[OF T_wf(1,3), of t s'] transaction_deletes_are_Value_vars(1)[OF T_wf(1,3), of t s'] by force hence "Fun (Val n) [] \ subst_range \" using transaction_decl_fresh_renaming_substs_range'(4)[ OF step.hyps(3,4,5) _ \_empty] unfolding \_def by blast thus ?thesis unfolding \_vals_def by fast qed have in_A_case: "\ x \ Fun (Val n) []" when y: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" "\ x = \ y" "\ x = \ y" for y using A_cases proof assume "?A \" thus ?thesis using B(14) m y(1,3) unfolding db_upds_consts_fresh_def y(2) by auto next assume "?A T'" hence "n \ N'" using n_in_case unfolding N'_def by blast moreover have "\ y \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using B(6) y(1) m ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset unfolding y(2) subst_in_ik_if_subst_pubval_def by blast ultimately show ?thesis using B(12) y(3) by fastforce qed show "\ x \ Fun (Val n) []" proof (cases "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'") case True note 0 = T'_fv_\_pubval_\_eq_\_fresh_val[OF True m, unfolded B(9)[symmetric]] note 1 = T'_fv_\_pubval_\_eq_\_val[OF True m, unfolded B(9)[symmetric]] show ?thesis proof (cases "\y\ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'. \ y \ \ x") case True show ?thesis using A_cases 0[OF True] T_val_fresh_vals n_in_case unfolding \_vals_def by force next case False then obtain y where "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" "\ y = \ x" "\ x = \ y" using 1 by blast thus ?thesis using in_A_case by auto qed next case False hence x_in: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ X'" using x unfolding unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by fast hence x_eq: "\ x = \ x" using \_fv_\_eq_\ by blast show ?thesis using in_A_case[OF x_in _ x_eq] by blast qed qed have B'2: "?wt_model \ \'" proof (unfold welltyped_constraint_model_def; intro conjI) have "\ (\ x) = \\<^sub>v x" for x proof - have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using B(2) 3(1) unfolding welltyped_constraint_model_def by (blast,blast) hence *: "\y. \ (\ y) = \\<^sub>v y" "\y. \ (\ y) = \\<^sub>v y" unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by auto show ?thesis proof (cases "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'") case True note x = this show ?thesis proof (cases "\n. \ x = Fun (PubConst Value n) []") case True thus ?thesis using T'_fv_\_pubval_\_eq_val[OF x] T'_value_vars[OF x] by force next case False thus ?thesis using x * unfolding \_def by presburger qed next case False thus ?thesis using *(1) unfolding \_def by presburger qed qed thus "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" unfolding \_def wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by force show "constraint_model \ \'" proof (unfold constraint_model_def; intro conjI) have *: "strand_sem_stateful {} {} (unlabel \) \" "strand_sem_stateful {} {} (unlabel \) \" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using B(2) 3(1) unfolding welltyped_constraint_model_def constraint_model_def by fast+ show K0: "subst_domain \ = UNIV" proof - have "x \ subst_domain \" for x proof (cases "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'") case True thus ?thesis using T'_fv_\_pubval_\_eq_val[OF True] T'_fv_\_val_\_eq_\[OF True] *(3) unfolding subst_domain_def by (cases "\n. \ x = Fun (PubConst Value n) []") auto next case False thus ?thesis using *(4) unfolding \_def subst_domain_def by auto qed thus ?thesis by blast qed have "fv (\ x) = {}" for x using interpretation_grounds_all[OF *(3)] interpretation_grounds_all[OF *(4)] unfolding \_def by simp thus K1: "ground (subst_range \)" by simp have "wf\<^sub>t\<^sub>r\<^sub>m (Fun (Val n) [])" for n by fastforce moreover have "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" for x using *(5,6) by (fastforce,fastforce) ultimately have "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" for x unfolding \_def by auto thus K2: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by simp show "strand_sem_stateful {} {} (unlabel \') \" proof (unfold \'_def unlabel_append strand_sem_append_stateful Un_empty_left; intro conjI) let ?sem = "\M D A. strand_sem_stateful M D (unlabel A) \" let ?M1 = "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" let ?M2 = "?M1 \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr \\<^sub>s\<^sub>e\<^sub>t \)" let ?D1 = "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" let ?D2 = "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel T_val_constr) \ ?D1" show "?sem {} {} \" using \_\_fv_\_eq strand_sem_model_swap[OF _ *(2)] by blast show "?sem ?M1 ?D1 T_val_constr" using T_val_constr(3,4) strand_sem_stateful_if_no_send_or_check unfolding list_all_iff by blast have D2: "?D2 = ?D1 \ {(t \ \, s \ \) | t s. insert\t,s\ \ set (unlabel T_val_constr)}" using T_val_constr(3,4) dbupd\<^sub>s\<^sub>s\<^sub>t_no_deletes unfolding list_all_iff by blast have K3: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using K0 K1 by argo have rcv_\_is_\: "t \ \ = t \ \" when t: "(l,receive\ts\) \ set (transaction_receive T)" "t \ set ts" for l ts t proof - have "fv t \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" using t(2) stateful_strand_step_fv_subset_cases(2)[OF unlabel_in[OF t(1)]] by auto hence "t \ \ = t" using t \_dom \_ran admissible_transactionE(12,13)[OF T_adm] by blast thus ?thesis unfolding \_def \_empty by simp qed have eq_\_is_\: "t \ \ = t \ \" "s \ \ = s \ \" when t: "(l,\ac: t \ s\) \ set (transaction_checks T)" for l ac t s proof - have "fv t \ fv s \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using stateful_strand_step_fv_subset_cases(3)[OF unlabel_in[OF t]] by auto hence "t \ \ = t" "s \ \ = s" using t \_dom \_ran admissible_transactionE(12,13)[OF T_adm] by (blast, blast) thus "t \ \ = t \ \" "s \ \ = s \ \" unfolding \_def \_empty by simp_all qed have noteq_\_is_\: "t \ \ = t \ \" "s \ \ = s \ \" when t: "(l,\t != s\) \ set (transaction_checks T)" for l t s proof - have "fv t \ fv s \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using stateful_strand_step_fv_subset_cases(8)[OF unlabel_in[OF t]] by auto hence "t \ \ = t" "s \ \ = s" using t \_dom \_ran admissible_transactionE(12,13)[OF T_adm] by (blast, blast) thus "t \ \ = t \ \" "s \ \ = s \ \" unfolding \_def \_empty by simp_all qed have in_\_is_\: "t \ \ = t \ \" "s \ \ = s \ \" when t: "(l,\ac: t \ s\) \ set (transaction_checks T)" for l ac t s proof - have "fv t \ fv s \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using stateful_strand_step_fv_subset_cases(6)[OF unlabel_in[OF t]] by auto hence "t \ \ = t" "s \ \ = s" using t \_dom \_ran admissible_transactionE(12,13)[OF T_adm] by (blast, blast) thus "t \ \ = t \ \" "s \ \ = s \ \" unfolding \_def \_empty by simp_all qed have notin_\_is_\: "t \ \ = t \ \" "s \ \ = s \ \" when t: "(l,\t not in s\) \ set (transaction_checks T)" for l t s proof - have "fv t \ fv s \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using stateful_strand_step_fv_subset_cases(9)[OF unlabel_in[OF t]] by auto hence "t \ \ = t" "s \ \ = s" using t \_dom \_ran admissible_transactionE(12,13)[OF T_adm] by (blast, blast) thus "t \ \ = t \ \" "s \ \ = s \ \" unfolding \_def \_empty by simp_all qed have T'_trm_no_val: "\n. s = Fun (Val n) [] \ s = Fun (PubConst Value n) []" when t: "t \ trms_transaction T" "s \ t \ \" for t s proof - have ?thesis when "s \ t" using that t admissible_transactions_no_Value_consts'[OF T_adm] admissible_transactions_no_PubConsts[OF T_adm] by blast moreover have "Fun k [] \ u" when "Fun k [] \ u \ \" for k u using that proof (induction u) case (Var x) thus ?case using transaction_renaming_subst_is_renaming(2)[OF step.hyps(5), of x] by fastforce qed auto ultimately show ?thesis using t by blast qed define flt1 where "flt1 \ \A::('fun,'atom,'sets,'lbl) prot_constr. filter is_Update (unlabel A)" define flt2 where "flt2 \ \A::('fun,'atom,'sets,'lbl) prot_constr. filter (\a. \l t ts. a = (l, insert\t,\s_val\ts\\\<^sub>s\)) A" define flt3 where "flt3 \ \A::(('fun,'atom,'sets,'lbl) prot_fun, ('fun,'atom,'sets,'lbl) prot_var) stateful_strand. filter (\a. \t ts. a = insert\t,\s_val\ts\\\<^sub>s\) A" have flt2_subset: "set (unlabel (flt2 A)) \ set (unlabel A)" for A unfolding flt2_def unlabel_def by auto have flt2_unlabel: "unlabel (flt2 A) = flt3 (unlabel A)" for A unfolding flt2_def flt3_def by (induct A) auto have flt2_suffix: "suffix (filter (\a. \t ts. a = insert\t,\s_val\ts\\\<^sub>s\) A) (unlabel (flt2 B))" when "suffix A (unlabel B)" for A B using that unfolding flt2_def by (induct B arbitrary: A rule: List.rev_induct) auto have flt_AB: "flt1 (flt2 \) = flt1 (flt2 \)" proof - have *: "flt1 (flt2 \) = filter is_Update (flt3 (unlabel \))" "flt1 (flt2 \) = filter is_Update (flt3 (unlabel \))" using flt2_unlabel unfolding flt1_def by presburger+ have **: "filter is_Update (flt3 C) = flt3 (filter is_Update C)" for C proof (induction C) case Nil thus ?case unfolding flt3_def by force next case (Cons c C) thus ?case unfolding flt3_def by (cases c) auto qed show ?thesis proof (cases "T_upds = []") case True hence "filter is_Update (unlabel \) = filter is_Update (unlabel \)" using B(13) unfolding db_eq_def by fastforce thus ?thesis using ** unfolding * by presburger next case False thus ?thesis using B(13) unfolding flt1_def flt2_def db_eq_def Let_def by force qed qed have A_setops_Fun: "\t s. insert\t,s\ \ set (unlabel \) \ (\g ts. s = Fun g ts)" using reachable_constraints_setops_form[OF step.hyps(1) P] unfolding setops\<^sub>s\<^sub>s\<^sub>t_def by fastforce have A_insert_delete_not_subterm: "\ x = \ x \ (\(\ x \ t) \ \(\ x \ s) \ \(\ x \ t) \ \(\ x \ s))" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ fv t \ fv s" and x_neq: "\ x \ \ x" and ts: "insert\t,s\ \ set (unlabel \) \ delete\t,s\ \ set (unlabel \)" for x t s proof - have x_in: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" using ts x stateful_strand_step_fv_subset_cases(4,5) by blast note ts' = reachable_constraints_insert_delete_form[OF step.hyps(1) P ts] have *: "\ x = \ x" when n: "\ x = Fun (Val n) []" for n using n B'5 x_in unfolding subst_eq_on_privvals_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast have **: "\(\ x \ t)" "\(\ x \ s)" "\(\ x \ t)" "\(\ x \ s)" when n: "\ x = Fun (PubConst Value n) []" for n proof - show "\(\ x \ s)" using ts'(1) x_in 3(2,3) unfolding valconst_cases_def by fastforce show "\(\ x \ s)" using ts'(1) x_in B'3 unfolding valconsts_only_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by force show "\(\ x \ t)" using n ts'(3) by fastforce from ts'(3) have "\ x \ t" proof assume "\y. t = Var y" thus ?thesis using B'3 x_in unfolding valconsts_only_def by force next assume "\k. t = Fun (Val k) []" thus ?thesis using B'14 n x_in ts unfolding db_upds_consts_fresh_def by auto qed thus "\(\ x \ t)" using ts'(3) by auto qed show ?thesis using * ** 3(2,3) x_in unfolding valconst_cases_def by fast qed have flt2_insert_in_iff: "insert\u,v\ \ set (unlabel A) \ insert\u,v\ \ set (unlabel (flt2 A))" (is "?A A \ ?B A") when h: "s = \h\\<^sub>s" "h \ s_val" and t: "(t \ I,s \ I) = (u,v) \\<^sub>p I" for t s h u v A and I::"('fun,'atom,'sets,'lbl) prot_subst" proof show "?B A \ ?A A" using flt2_subset by fast show "?A A \ ?B A" proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) show ?case proof (cases "b = insert\u,v\") case True thus ?thesis using h t unfolding a flt2_def by force next case False thus ?thesis using Cons.prems Cons.IH unfolding a flt2_def by auto qed qed simp qed have flt2_inset_iff: "(t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (flt2 \)) \ {} \ (t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" (is "?A \ ?B") when h: "s = \h\\<^sub>s" "h \ s_val" for t s h proof let ?C1 = "\u v B C. suffix (delete\u,v\#B) (unlabel C)" let ?C2 = "\t s u v. (t,s) = (u,v) \\<^sub>p \" let ?C3 = "\t s C. \u v. ?C2 t s u v \ insert\u,v\ \ set C" let ?D = "\t s C. \u v B. ?C1 u v B C \ ?C2 t s u v \ ?C3 t s B" let ?db = "\C D. dbupd\<^sub>s\<^sub>s\<^sub>t C \ D" have "?C3 t s B" when "?D t s (flt2 \)" "?C1 u v B \" "?C2 t s u v" for u v B t s using that flt2_suffix flt2_subset by fastforce thus "?A \ ?B" using flt2_subset unfolding dbupd\<^sub>s\<^sub>s\<^sub>t_in_iff by blast show ?A when ?B using that proof (induction \ rule: List.rev_induct) case (snoc a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) have *: "?db (unlabel (A@[a])) {} = ?db [b] (?db (unlabel A) {})" "?db (unlabel (flt2 (A@[a]))) {} = ?db (unlabel (flt2 [a])) (?db (unlabel (flt2 A)) {})" using dbupd\<^sub>s\<^sub>s\<^sub>t_append[of _ _ \ "{}"] unfolding a flt2_def by auto show ?case proof (cases "\u v. b = insert\u,v\ \ (t \ \, s \ \) = (u,v) \\<^sub>p \") case True then obtain u v where "b = insert\u,v\" "(t \ \, s \ \) = (u, v) \\<^sub>p \" by force thus ?thesis using h *(2) unfolding a flt2_def by auto next case False hence IH: "(t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (flt2 A)) \ {}" using snoc.prems snoc.IH unfolding *(1) by (cases b) auto show ?thesis proof (cases "is_Delete b") case True then obtain u v where b: "b = delete\u,v\" by (cases b) auto have b': "unlabel (flt2 [a]) = [b]" "unlabel (flt2 (A@[a])) = unlabel (flt2 A)@[b]" unfolding a flt2_def b by (fastforce,fastforce) have "(t \ \, s \ \) \ (u,v) \\<^sub>p \" using *(1) snoc.prems unfolding b' b by simp thus ?thesis using *(2) IH unfolding b' b by simp next case False thus ?thesis using *(2) IH unfolding a flt2_def by (cases b) auto qed qed qed simp qed have inset_model_swap: "(t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {} \ (t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" (is "?in \ (unlabel \) \ ?in \ (unlabel \)") when h: "s = \h\\<^sub>s" "h \ s_val \ filter is_Update (unlabel \) = filter is_Update (unlabel \)" and t: "t = Var tx" and t_s_fv: "fv t \ fv s \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" and q: "\x \ fv t \ fv s. \ x = \ x \ (\(\ x \ t) \ \(\ x \ s) \ \(\ x \ t) \ \(\ x \ s))" "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv t \ fv s. \c. \ x = Fun c []" "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv t \ fv s. \c. \ x = Fun c []" "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv t \ fv s. \y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv t \ fv s. \ x = \ y \ \ x = \ y" for t s h tx proof - let ?upds = "\A. filter is_Update (unlabel A)" have flt2_fv: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (flt2 \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using fv\<^sub>s\<^sub>s\<^sub>t_mono[OF flt2_subset[of \]] by blast have upds_fv: "fv\<^sub>s\<^sub>s\<^sub>t (?upds \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" by auto have flt2_upds_fv: "fv\<^sub>s\<^sub>s\<^sub>t (?upds (flt2 \)) \ fv\<^sub>s\<^sub>s\<^sub>t (?upds \)" using flt2_subset[of \] by auto have h_neq: "Set h \ (Set s_val::('fun,'atom,'sets,'lbl) prot_fun)" when "h \ s_val" using that by simp have *: "\(fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r ` {}) = {}" "{} \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ = {}" "{} \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ = {}" by blast+ have "?in \ (?upds (flt2 \)) \ ?in \ (?upds (flt2 \))" proof let ?X = "fv\<^sub>s\<^sub>s\<^sub>t (?upds (flt2 \)) \ fv t \ fv s \ \(fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r ` ({}::(('fun,'atom,'sets,'lbl) prot_term \ ('fun,'atom,'sets,'lbl) prot_term) set))" let ?q0 = "\\ \. \x \ ?X. \ x = \ x \ (\(\ x \ t) \ \(\ x \ s) \\(\ x \ t) \ \(\ x \ s) \ (\(u,v) \ {}. \(\ x \ u) \ \(\ x \ v) \ \(\ x \ u) \ \(\ x \ v)) \ (\u v. insert\u,v\ \ set (?upds (flt2 \)) \ delete\u,v\ \ set (?upds (flt2 \)) \ \(\ x \ u) \ \(\ x \ v) \ \(\ x \ u) \ \(\ x \ v)))" let ?q1 = "\\. \x \ ?X. \c. \ x = Fun c []" let ?q2 = "\\ \. \x \ ?X. \y \ ?X. \ x = \ y \ \ x = \ y" have q0: "?q0 \ \" "?q0 \ \" proof - have upd_ex: "\u v. x \ fv u \ fv v \ (insert\u,v\ \ set (?upds A) \ delete\u,v\ \ set (?upds A))" when "x \ fv\<^sub>s\<^sub>s\<^sub>t (?upds A)" for x and A::"('fun,'atom,'sets,'lbl) prot_constr" using that proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) show ?case using Cons.IH Cons.prems unfolding a by (cases b) auto qed simp have "\(\ x \ t)" "\(\ x \ s)" "\(\ x \ t)" "\(\ x \ s)" when x: "x \ fv\<^sub>s\<^sub>s\<^sub>t (?upds (flt2 \)) \ fv t \ fv s" and x_neq: "\ x \ \ x" for x proof - have "\(\ x \ t) \ \(\ x \ s) \ \(\ x \ t) \ \(\ x \ s)" proof (cases "x \ fv t \ fv s") case True thus ?thesis using q(1) x_neq by blast next case False hence "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using x flt2_upds_fv upds_fv by blast hence "\n. \ x = Fun (Val n) []" "\n. \ x = Fun (Val n) [] \ \ x = Fun (PubConst Value n) []" using B'3 3(2) unfolding valconst_cases_def valconsts_only_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by (blast, blast) thus ?thesis unfolding t h(1) by auto qed thus "\(\ x \ t)" "\(\ x \ s)" "\(\ x \ t)" "\(\ x \ s)" by simp_all qed moreover have "\(\ x \ u)" "\(\ x \ v)" "\(\ x \ u)" "\(\ x \ v)" when x: "x \ fv\<^sub>s\<^sub>s\<^sub>t (?upds (flt2 \)) \ fv t \ fv s" and x_neq: "\ x \ \ x" and uv: "insert\u,v\ \ set (?upds (flt2 \)) \ delete\u,v\ \ set (?upds (flt2 \))" for x u v proof - have uv': "insert\u,v\ \ set (unlabel \) \ delete\u,v\ \ set (unlabel \)" using uv flt2_subset by auto have x_in: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ fv u \ fv v" using t_s_fv x flt2_upds_fv upds_fv by blast show "\(\ x \ u)" "\(\ x \ v)" "\(\ x \ u)" "\(\ x \ v)" using x_neq A_insert_delete_not_subterm[OF x_in x_neq uv'] by simp_all qed ultimately show "?q0 \ \" unfolding upd_ex unfolding * by (metis (no_types, lifting) empty_iff sup_bot_right) thus "?q0 \ \" by (metis (lifting) empty_iff) qed have q1: "?q1 \" "?q1 \" using q(2,3) flt2_upds_fv upds_fv by (blast,blast) have q2: "?q2 \ \" "?q2 \ \" using q(4) flt2_upds_fv upds_fv unfolding * by (blast,blast) show "?in \ (?upds (flt2 \)) \ ?in \ (?upds (flt2 \))" using dbupd\<^sub>s\<^sub>s\<^sub>t_subst_const_swap[OF _ q0(1) q1(1,2) q2(1)] by force show "?in \ (?upds (flt2 \)) \ ?in \ (?upds (flt2 \))" using dbupd\<^sub>s\<^sub>s\<^sub>t_subst_const_swap[OF _ q0(2) q1(2,1) q2(2)] by force qed hence flt2_subst_swap: "?in \ (unlabel (flt2 \)) \ ?in \ (unlabel (flt2 \))" using dbupd\<^sub>s\<^sub>s\<^sub>t_filter by blast (* TODO: merge with similar proof above? *) have "?in \ (?upds \) \ ?in \ (?upds \)" proof let ?X = "fv\<^sub>s\<^sub>s\<^sub>t (?upds \) \ fv t \ fv s \ \(fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r ` ({}::(('fun,'atom,'sets,'lbl) prot_term \ ('fun,'atom,'sets,'lbl) prot_term) set))" let ?q0 = "\\ \. \x \ ?X. \ x = \ x \ (\(\ x \ t) \ \(\ x \ s) \\(\ x \ t) \ \(\ x \ s) \ (\(u,v) \ {}. \(\ x \ u) \ \(\ x \ v) \ \(\ x \ u) \ \(\ x \ v)) \ (\u v. insert\u,v\ \ set (?upds \) \ delete\u,v\ \ set (?upds \) \ \(\ x \ u) \ \(\ x \ v) \ \(\ x \ u) \ \(\ x \ v)))" let ?q1 = "\\. \x \ ?X. \c. \ x = Fun c []" let ?q2 = "\\ \. \x \ ?X. \y \ ?X. \ x = \ y \ \ x = \ y" have q0: "?q0 \ \" "?q0 \ \" proof - have upd_ex: "\u v. x \ fv u \ fv v \ (insert\u,v\ \ set (?upds A) \ delete\u,v\ \ set (?upds A))" when "x \ fv\<^sub>s\<^sub>s\<^sub>t (?upds A)" for x and A::"('fun,'atom,'sets,'lbl) prot_constr" using that proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) show ?case using Cons.IH Cons.prems unfolding a by (cases b) auto qed simp have "\(\ x \ t)" "\(\ x \ s)" "\(\ x \ t)" "\(\ x \ s)" when x: "x \ fv\<^sub>s\<^sub>s\<^sub>t (?upds \) \ fv t \ fv s" and x_neq: "\ x \ \ x" for x proof - have "\(\ x \ t) \ \(\ x \ s) \ \(\ x \ t) \ \(\ x \ s)" proof (cases "x \ fv t \ fv s") case True thus ?thesis using q(1) x_neq by blast next case False hence "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using x flt2_upds_fv upds_fv by blast hence "\n. \ x = Fun (Val n) []" "\n. \ x = Fun (Val n) [] \ \ x = Fun (PubConst Value n) []" using B'3 3(2) unfolding valconst_cases_def valconsts_only_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by (blast, blast) thus ?thesis unfolding t h(1) by auto qed thus "\(\ x \ t)" "\(\ x \ s)" "\(\ x \ t)" "\(\ x \ s)" by simp_all qed moreover have "\(\ x \ u)" "\(\ x \ v)" "\(\ x \ u)" "\(\ x \ v)" when x: "x \ fv\<^sub>s\<^sub>s\<^sub>t (?upds \) \ fv t \ fv s" and x_neq: "\ x \ \ x" and uv: "insert\u,v\ \ set (?upds \) \ delete\u,v\ \ set (?upds \)" for x u v proof - have uv': "insert\u,v\ \ set (unlabel \) \ delete\u,v\ \ set (unlabel \)" using uv flt2_subset by auto have x_in: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ fv u \ fv v" using t_s_fv x flt2_upds_fv upds_fv by blast show "\(\ x \ u)" "\(\ x \ v)" "\(\ x \ u)" "\(\ x \ v)" using x_neq A_insert_delete_not_subterm[OF x_in x_neq uv'] by simp_all qed ultimately show "?q0 \ \" unfolding upd_ex unfolding * by (metis (no_types, lifting) empty_iff sup_bot_right) thus "?q0 \ \" by (metis (lifting) empty_iff) qed have q1: "?q1 \" "?q1 \" using q(2,3) flt2_upds_fv upds_fv by (blast,blast) have q2: "?q2 \ \" "?q2 \ \" using q(4) flt2_upds_fv upds_fv unfolding * by (blast,blast) show "?in \ (?upds \) \ ?in \ (?upds \)" using dbupd\<^sub>s\<^sub>s\<^sub>t_subst_const_swap[OF _ q0(1) q1(1,2) q2(1)] by force show "?in \ (?upds \) \ ?in \ (?upds \)" using dbupd\<^sub>s\<^sub>s\<^sub>t_subst_const_swap[OF _ q0(2) q1(2,1) q2(2)] by force qed hence db_subst_swap: "?in \ (unlabel \) \ ?in \ (unlabel \)" using dbupd\<^sub>s\<^sub>s\<^sub>t_filter by blast have "?in \ (unlabel \)" when A: "?in \ (unlabel \)" using h(2) proof assume h': "h \ s_val" have "?in \ (unlabel (flt2 \))" using A flt2_unlabel dbupd\<^sub>s\<^sub>s\<^sub>t_set_term_neq_in_iff[OF h_neq[OF h'] A_setops_Fun] unfolding h(1) flt3_def by simp hence "?in \ (unlabel (flt2 \))" using flt2_subst_swap by blast hence "?in \ (flt1 (flt2 \))" using dbupd\<^sub>s\<^sub>s\<^sub>t_filter unfolding flt1_def by blast hence "?in \ (flt1 (flt2 \))" using flt_AB by simp hence "?in \ (unlabel (flt2 \))" using dbupd\<^sub>s\<^sub>s\<^sub>t_filter unfolding flt1_def by blast thus ?thesis using flt2_inset_iff[OF h(1) h'] by fast next assume h': "filter is_Update (unlabel \) = filter is_Update (unlabel \)" have "?in \ (unlabel \)" using A db_subst_swap by blast hence "?in \ (flt1 \)" using dbupd\<^sub>s\<^sub>s\<^sub>t_filter unfolding flt1_def by blast hence "?in \ (flt1 \)" using h' unfolding flt1_def by simp thus ?thesis using dbupd\<^sub>s\<^sub>s\<^sub>t_filter unfolding flt1_def by blast qed moreover have "\?in \ (unlabel \)" when A: "\?in \ (unlabel \)" using h(2) proof assume h': "h \ s_val" have "\?in \ (unlabel (flt2 \))" using A flt2_unlabel dbupd\<^sub>s\<^sub>s\<^sub>t_set_term_neq_in_iff[OF h_neq[OF h'] A_setops_Fun] unfolding h(1) flt3_def by simp hence "\?in \ (unlabel (flt2 \))" using flt2_subst_swap by blast hence "\?in \ (flt1 (flt2 \))" using dbupd\<^sub>s\<^sub>s\<^sub>t_filter unfolding flt1_def by blast hence "\?in \ (flt1 (flt2 \))" using flt_AB by simp hence "\?in \ (unlabel (flt2 \))" using dbupd\<^sub>s\<^sub>s\<^sub>t_filter unfolding flt1_def by blast thus ?thesis using flt2_inset_iff[OF h(1) h'] by fast next assume h': "filter is_Update (unlabel \) = filter is_Update (unlabel \)" have "\?in \ (unlabel \)" using A db_subst_swap by blast hence "\?in \ (flt1 \)" using dbupd\<^sub>s\<^sub>s\<^sub>t_filter unfolding flt1_def by blast hence "\?in \ (flt1 \)" using h' unfolding flt1_def by simp thus ?thesis using dbupd\<^sub>s\<^sub>s\<^sub>t_filter unfolding flt1_def by blast qed ultimately show ?thesis by blast qed have "?M2 \ t \ \ \ \" when ts: "(l, receive\ts\) \ set (transaction_receive T)" "t \ set ts" for l t ts proof - have *: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \ \ \" using 5 ts by blast note t\\ = rcv_\_is_\[OF ts] have t_T'_trm: "t \ trms_transaction T" using trms\<^sub>s\<^sub>s\<^sub>t_memI(2)[OF unlabel_in[OF ts(1)] ts(2)] unfolding trms_transaction_unfold by blast have t_T'_trm': "t \ \ \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" using trms\<^sub>s\<^sub>s\<^sub>t_memI(2)[ OF stateful_strand_step_subst_inI(2)[ OF unlabel_in[OF ts(1)], unfolded unlabel_subst]] ts(2) unfolding T'_def trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq trms_transaction_subst_unfold by auto note t_no_val = T'_trm_no_val[OF t_T'_trm, unfolded t\\[symmetric]] have t_fv_T': "fv (t \ \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" using ts(2) stateful_strand_step_fv_subset_cases(2)[ OF stateful_strand_step_subst_inI(2)[OF unlabel_in[OF ts(1)], of \]] unfolding T'_def unlabel_subst fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq fv_transaction_subst_unfold by auto have ik_B_fv_subset: "fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" by (meson UnE fv_ik\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t subset_iff) let ?fresh_vals = "(\n. Fun (Val n) []) ` T_val_fresh_vals" have q0: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \ \ \" using * B(10) by (blast intro: ideduct_mono) have q1: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ fv (t \ \). valconst_cases \ x" using 3(2,3) t_fv_T' ik_B_fv_subset unfolding B(9) by blast have q2: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ fv (t \ \). \n. \ x = Fun (Val n) []" using B'3 t_fv_T' ik_B_fv_subset unfolding valconsts_only_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append B(9) by blast have T_val_constr_ik: "\M. ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr = M \ ?fresh_vals" "\M. ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr \\<^sub>s\<^sub>e\<^sub>t \ = (M \\<^sub>s\<^sub>e\<^sub>t \) \ ?fresh_vals" proof - obtain M where M: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr = M \ ?fresh_vals" using T_val_constr(8) by blast have "?fresh_vals \\<^sub>s\<^sub>e\<^sub>t \ = ?fresh_vals" by fastforce thus "\M. ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr = M \ ?fresh_vals" "\M. ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr \\<^sub>s\<^sub>e\<^sub>t \ = (M \\<^sub>s\<^sub>e\<^sub>t \) \ ?fresh_vals" using M by (fastforce, fastforce) qed have "\ x \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T_val_constr" when x: "x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ fv (t \ \)" "\ x = Fun (PubConst Value n) []" for x n using x(1) B'6'[OF _ x(2)] B'6''[OF _ x(2)] t_fv_T' ik_B_fv_subset unfolding B(9) unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast hence q3: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ fv (t \ \). (\n. \ x = Fun (PubConst Value n) []) \ \ x \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ ?fresh_vals" using T_val_constr_ik(1) T_val_constr(8) q2 unfolding B(9) \'_def unlabel_append ik\<^sub>s\<^sub>s\<^sub>t_append fv\<^sub>s\<^sub>s\<^sub>t_append by (metis (no_types, lifting) UnE UnI1 UnI2 image_iff mem_Collect_eq) have q4: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ fv (t \ \). (\n. \ x = Fun (Val n) []) \ \ x = \ x" using B'5 t_fv_T' ik_B_fv_subset unfolding subst_eq_on_privvals_def B(9) unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast have q5: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ fv (t \ \). \y \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ fv (t \ \). \ x = \ y \ \ x = \ y" using B'7 t_fv_T' ik_B_fv_subset unfolding subst_eq_iff_def B(9) unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast have q6: "\n. \(Fun (PubConst Value n) [] \\<^sub>s\<^sub>e\<^sub>t insert (t \ \) (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" proof - have "\n. s = Fun (PubConst Value n) []" when s: "s \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \'" for s proof - have "f \ PubConst Value n" when f: "f \ funs_term s" for f n using f s reachable_constraints_val_funs_private(1)[OF B'1 P, of f] unfolding is_PubConstValue_def is_PubConst_def the_PubConst_type_def by (metis (mono_tags, lifting) UN_I funs_term_subterms_eq(2) prot_fun.simps(85)) thus ?thesis by fastforce qed moreover have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \'" using ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset unfolding \'_def unlabel_append trms\<^sub>s\<^sub>s\<^sub>t_append by blast ultimately show ?thesis using t_no_val by blast qed show ?thesis using deduct_val_const_swap[OF q0 q1[unfolded valconst_cases_def] q2 q3 q4 q5 q6] T_val_constr_ik(2) by (blast intro: ideduct_mono) qed moreover have "t \ \ \ \ = s \ \ \ \" when ts: "(l, \ac: t \ s\) \ set (transaction_checks T)" for l ac t s proof - have q0: "t \ \ \ \ = s \ \ \ \" using 5 ts by blast have "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\ac: (t \ \) \ (s \ \)\) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using stateful_strand_step_fv_subset_cases(3)[ OF stateful_strand_step_subst_inI(3)[OF unlabel_in[OF ts], of \]] unfolding unlabel_subst by simp hence t_s_fv: "fv (t \ \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "fv (s \ \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" unfolding T'_def fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq fv_transaction_subst_unfold[of T \] by (fastforce, fastforce) have "t \ trms_transaction T" "s \ trms_transaction T" using trms\<^sub>s\<^sub>s\<^sub>t_memI(3,4)[OF unlabel_in[OF ts]] unfolding trms_transaction_unfold by (blast, blast) hence "\n. u = Fun (Val n) [] \ u = Fun (PubConst Value n) []" when u: "u \ t \ \ \ u \ s \ \" for u using u T'_trm_no_val unfolding eq_\_is_\[OF ts] by blast hence "\(\ x \ t \ \)" "\(\ x \ s \ \)" when x: "x \ fv (t \ \) \ fv (s \ \)" for x using x t_s_fv I' by (fast, fast) hence q1: "\x \ fv (t \ \) \ fv (s \ \). \ x = \ x \ (\(\ x \ t \ \) \ \(\ x \ s \ \))" by blast have q2: "\x \ fv (t \ \) \ fv (s \ \). \c. \ x = Fun c []" using t_s_fv 3(3) unfolding valconst_cases_def by blast have q3: "\x \ fv (t \ \) \ fv (s \ \). \c. \ x = Fun c []" using t_s_fv B'3 unfolding valconsts_only_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast have q4: "\x \ fv (t \ \) \ fv (s \ \). \y \ fv (t \ \) \ fv (s \ \). \ x = \ y \ \ x = \ y" using B'7 t_s_fv unfolding subst_eq_iff_def B(9) unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast show ?thesis by (rule subst_const_swap_eq'[OF q0 q1 q2 q3 q4]) qed moreover have "t \ \ \ \ \ s \ \ \ \" when ts: "(l, \t != s\) \ set (transaction_checks T)" for l t s proof - have q0: "t \ \ \ \ \ s \ \ \ \" using 5 ts by blast have "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\(t \ \) != (s \ \)\) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using stateful_strand_step_fv_subset_cases(8)[ OF stateful_strand_step_subst_inI(8)[OF unlabel_in[OF ts], of \]] unfolding unlabel_subst by simp hence t_s_fv: "fv (t \ \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "fv (s \ \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" unfolding T'_def fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq fv_transaction_subst_unfold[of T \] by (fastforce, fastforce) have "t \ trms_transaction T" "s \ trms_transaction T" using trms\<^sub>s\<^sub>s\<^sub>t_memI(9)[OF unlabel_in[OF ts]] unfolding trms_transaction_unfold by auto hence "\n. u = Fun (Val n) []" when u: "u \ t \ \ \ u \ s \ \" for u using u T'_trm_no_val unfolding noteq_\_is_\[OF ts] by blast hence "\(\ x \ t \ \)" "\(\ x \ s \ \)" when x: "x \ fv (t \ \) \ fv (s \ \)" for x using x t_s_fv B'3 unfolding valconsts_only_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by (fast, fast) hence q1: "\x \ fv (t \ \) \ fv (s \ \). \ x = \ x \ (\(\ x \ t \ \) \ \(\ x \ s \ \))" by blast have q2: "\x \ fv (t \ \) \ fv (s \ \). \c. \ x = Fun c []" using t_s_fv B'3 unfolding valconsts_only_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast have q3: "\x \ fv (t \ \) \ fv (s \ \). \c. \ x = Fun c []" using t_s_fv 3(3) unfolding valconst_cases_def by blast have q4: "\x \ fv (t \ \) \ fv (s \ \). \y \ fv (t \ \) \ fv (s \ \). \ x = \ y \ \ x = \ y" using B'7 t_s_fv unfolding subst_eq_iff_def B(9) unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast show ?thesis using q0 subst_const_swap_eq'[OF _ q1 q2 q3 q4] by fast qed moreover have "(t \ \ \ \, s \ \ \ \) \ ?D2" when ts: "(l, \ac: t \ s\) \ set (transaction_checks T)" for l ac t s proof - have s_neq_s_val: "s \ \s_val\\<^sub>s \ filter is_Update (unlabel \) = filter is_Update (unlabel \)" proof (cases "T_upds = []") case False thus ?thesis using step.hyps(2) ts x_val(7) unfolding transaction_strand_def by (cases ac) fastforce+ qed (use B(13)[unfolded db_eq_def] in simp) have ts': "\ac: t \ s\ \ set (unlabel (transaction_strand T))" using ts unlabel_in[OF ts] unfolding transaction_strand_def by fastforce have "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\ac: (t \ \) \ (s \ \)\) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using stateful_strand_step_fv_subset_cases(6)[ OF stateful_strand_step_subst_inI(6)[OF unlabel_in[OF ts], of \]] unfolding unlabel_subst by simp hence t_s_fv: "fv (t \ \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "fv (s \ \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" unfolding T'_def fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq fv_transaction_subst_unfold[of T \] by (fastforce, fastforce) have "t \ trms_transaction T" "s \ trms_transaction T" using ts' unfolding trms\<^sub>s\<^sub>s\<^sub>t_def by (force, force) hence "\n. u = Fun (Val n) [] \ u = Fun (PubConst Value n) []" when u: "u \ t \ \ \ u \ s \ \" for u using u T'_trm_no_val unfolding in_\_is_\[OF ts] by blast hence "\(\ x \ t \ \)" "\(\ x \ s \ \)" "\(\ x \ t \ \)" "\(\ x \ s \ \)" when x: "x \ fv (t \ \) \ fv (s \ \)" for x using x t_s_fv B'3 I' unfolding valconsts_only_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by (fast,fast,fast,fast) hence q1: "\x \ fv (t \ \) \ fv (s \ \). \ x = \ x \ (\(\ x \ t \ \) \ \(\ x \ s \ \) \ \(\ x \ t \ \) \ \(\ x \ s \ \))" by blast have q2: "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv (t \ \) \ fv (s \ \). \c. \ x = Fun c []" using t_s_fv 3(2,3) unfolding valconst_cases_def by blast have q3: "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv (t \ \) \ fv (s \ \). \c. \ x = Fun c []" using t_s_fv B'3 unfolding valconsts_only_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast have q4: "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv (t \ \) \ fv (s \ \). \y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv (t \ \) \ fv (s \ \). \ x = \ y \ \ x = \ y" using B'7 t_s_fv unfolding subst_eq_iff_def B(9) unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast obtain h tx where s: "s = \h\\<^sub>s" and tx: "t = Var tx" using ts' transaction_selects_are_Value_vars[OF T_wf(1,2), of t s] transaction_inset_checks_are_Value_vars[OF T_adm, of t s] by (cases ac) auto have h: "s \ \ = \h\\<^sub>s" "h \ s_val \ filter is_Update (unlabel \) = filter is_Update (unlabel \)" using s s_neq_s_val by (simp,blast) obtain ty where ty: "t \ \ = Var ty" using tx transaction_renaming_subst_is_renaming(2)[OF step.hyps(5), of tx] unfolding in_\_is_\[OF ts] by force have "(t \ \ \ \, s \ \ \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" using 5 ts by blast hence "(t \ \ \ \, s \ \ \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" using inset_model_swap[OF h ty _ q1 q2 q3 q4] t_s_fv by simp thus ?thesis unfolding D2 by blast qed moreover have "(t \ \ \ \, s \ \ \ \) \ ?D2" when ts: "(l, \t not in s\) \ set (transaction_checks T)" for l t s proof - have s_neq_s_val: "(T_upds \ [] \ s \ \s_val\\<^sub>s) \ (T_upds = [] \ filter is_Update (unlabel \) = filter is_Update (unlabel \))" proof (cases "T_upds = []") case False thus ?thesis using step.hyps(2) ts x_val(7) unfolding transaction_strand_def by force qed (use B(13)[unfolded db_eq_def] in simp) have ts': "\t not in s\ \ set (unlabel (transaction_strand T))" using ts unlabel_in[OF ts] unfolding transaction_strand_def by fastforce have "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\(t \ \) not in (s \ \)\) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using stateful_strand_step_fv_subset_cases(9)[ OF stateful_strand_step_subst_inI(9)[OF unlabel_in[OF ts], of \]] unfolding unlabel_subst by simp hence t_s_fv: "fv (t \ \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "fv (s \ \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" unfolding T'_def fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq fv_transaction_subst_unfold[of T \] by (fastforce, fastforce) have "t \ trms_transaction T" "s \ trms_transaction T" using ts' unfolding trms\<^sub>s\<^sub>s\<^sub>t_def by (force, force) hence "\n. u = Fun (Val n) [] \ u = Fun (PubConst Value n) []" when u: "u \ t \ \ \ u \ s \ \" for u using u T'_trm_no_val unfolding notin_\_is_\[OF ts] by blast hence "\(\ x \ t \ \)" "\(\ x \ s \ \)" "\(\ x \ t \ \)" "\(\ x \ s \ \)" when x: "x \ fv (t \ \) \ fv (s \ \)" for x using x t_s_fv B'3 I' unfolding valconsts_only_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by (fast,fast,fast,fast) hence q1: "\x \ fv (t \ \) \ fv (s \ \). \ x = \ x \ (\(\ x \ t \ \) \ \(\ x \ s \ \) \ \(\ x \ t \ \) \ \(\ x \ s \ \))" by blast have q2: "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv (t \ \) \ fv (s \ \). \c. \ x = Fun c []" using t_s_fv 3(2,3) unfolding valconst_cases_def by blast have q3: "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv (t \ \) \ fv (s \ \). \c. \ x = Fun c []" using t_s_fv B'3 unfolding valconsts_only_def unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast have q4: "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv (t \ \) \ fv (s \ \). \y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv (t \ \) \ fv (s \ \). \ x = \ y \ \ x = \ y" using B'7 t_s_fv unfolding subst_eq_iff_def B(9) unlabel_append fv\<^sub>s\<^sub>s\<^sub>t_append by blast obtain h tx where s: "s = \h\\<^sub>s" and tx: "t = Var tx" using transaction_notinset_checks_are_Value_vars(1,2)[OF T_adm ts', of t s] by auto have h: "s \ \ = \h\\<^sub>s" "h \ s_val \ filter is_Update (unlabel \) = filter is_Update (unlabel \)" "T_upds \ [] \ h \ s_val" using s s_neq_s_val by (simp,blast,blast) obtain ty where ty: "t \ \ = Var ty" using tx transaction_renaming_subst_is_renaming(2)[OF step.hyps(5), of tx] unfolding notin_\_is_\[OF ts] by force have *: "(t \ \ \ \, s \ \ \ \) \ (u \ \, v \ \)" when u: "insert\u,v\ \ set (unlabel T_val_constr)" and h': "h \ s_val" for u v proof - have "v = \s_val\\<^sub>s" using T_val_constr(9) unlabel_mem_has_label[OF u] by force thus ?thesis using h(1) h' by simp qed have "(t \ \ \ \, s \ \ \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" using 5 ts by blast hence **: "(t \ \ \ \, s \ \ \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" using inset_model_swap[OF h(1,2) ty _ q1 q2 q3 q4] t_s_fv by simp show ?thesis proof (cases "T_upds = []") case True have "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel T_val_constr) I D = D" for I D using T_val_constr_no_upds_if_no_T_upds[OF True] dbupd\<^sub>s\<^sub>s\<^sub>t_filter[of "unlabel T_val_constr"] by force thus ?thesis using ** by simp next case False thus ?thesis using ** * h(3) T_val_constr_no_upds_if_no_T_upds unfolding D2 by blast qed qed ultimately show "?sem ?M2 ?D2 T'" unfolding T'_def 4[OF T_adm K3 K2] by blast qed qed qed show ?case using B'1 B'2 B'3 B'4 B'5 B'6 B'7 B'8_9 B'10_11 B'12 B'13 B'14 unfolding \_def[symmetric] T'_def[symmetric] by blast qed obtain \ \ where B: "\ \ reachable_constraints P" "?wt_model \ \" "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \n. \ x = Fun (Val n) []" "?rcv_att \ n \ ?rcv_att \ n" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using lmm[OF finite.emptyI _ _ finite.emptyI] unfolding valconsts_only_def by auto show ?thesis using B(1,3) welltyped_constraint_model_attack_if_receive_attack[OF B(2)] B(4) 2 unfolding wt_attack_def B(5) by (meson list.set_intros(1)) qed private lemma add_occurs_msgs_soundness_aux2: assumes P: "\T \ set P. admissible_transaction T" and A: "\ \ reachable_constraints P" shows "\\ \ reachable_constraints (map add_occurs_msgs P). \ = rm_occurs_msgs_constr \" using A proof (induction rule: reachable_constraints.induct) case (step \ T \ \ \) define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" let ?A' = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" let ?B' = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand (add_occurs_msgs T) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" obtain A B C D E F where T: "T = Transaction A B C D E F" by (cases T) simp have P': "\T \ set P. admissible_transaction' T" "\T \ set P. admissible_transaction_no_occurs_msgs T" using P admissible_transactionE'(1,2) by (blast,blast) note T_adm' = bspec[OF P step.hyps(2)] note T_adm = bspec[OF P'(1) step.hyps(2)] note \_empty = admissible_transaction_decl_subst_empty[OF T_adm step.hyps(3)] note T_fresh_val = admissible_transactionE(2)[OF T_adm] note T_no_occ = admissible_transactionE'(2)[OF T_adm'] obtain \ where B: "\ \ reachable_constraints (map add_occurs_msgs P)" "\ = rm_occurs_msgs_constr \" using step.IH by blast note 0 = add_occurs_msgs_cases[OF T] note 1 = add_occurs_msgs_vars_eq[OF bspec[OF P'(1)]] note 2 = add_occurs_msgs_trms[of T] note 3 = add_occurs_msgs_transaction_strand_set[OF T] have 4: "add_occurs_msgs T \ set (map add_occurs_msgs P)" using step.hyps(2) by simp have 5: "transaction_decl_subst \ (add_occurs_msgs T)" using step.hyps(3) 0(4) unfolding transaction_decl_subst_def by argo have "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction (add_occurs_msgs T))" when t: "t \ subst_range \" for t proof - obtain c where c: "t = Fun (Val c) []" using t T_fresh_val transaction_fresh_subst_domain[OF step.hyps(4)] transaction_fresh_subst_sends_to_val[OF step.hyps(4), of _ thesis] by fastforce have *: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction T)" using t step.hyps(4) unfolding transaction_fresh_subst_def by (fast,fast) have "t \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ (\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. t \ occurs (Var x)) \ (\c. Fun c [] \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ t \ occurs (Fun c []))" when t: "t \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using t rm_occurs_msgs_constr_reachable_constraints_trms_cases[OF P' B(2,1)] by fast thus "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using *(1) unfolding c by fastforce show "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction (add_occurs_msgs T))" using *(2) unfolding 2 c by force qed hence 6: "transaction_fresh_subst \ (add_occurs_msgs T) (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using step.hyps(4) unfolding transaction_fresh_subst_def 0(5) 2 by fast have 7: "transaction_renaming_subst \ (map add_occurs_msgs P) (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using step.hyps(5) rm_occurs_msgs_constr_reachable_constraints_vars_eq[OF P' B(1)] B(2) 1(5) unfolding transaction_renaming_subst_def by simp have "?A' = rm_occurs_msgs_constr ?B'" using admissible_transaction_decl_fresh_renaming_subst_not_occurs[OF T_adm step.hyps(3,4,5)] rm_occurs_msgs_constr_transaction_strand''[OF T_adm T_no_occ] unfolding \_def[symmetric] by metis hence 8: "\@?A' = rm_occurs_msgs_constr (\@?B')" by (metis rm_occurs_msgs_constr_append B(2)) show ?case using reachable_constraints.step[OF B(1) 4 5 6 7] 8 unfolding \_def by blast qed (metis reachable_constraints.init rm_occurs_msgs_constr.simps(1)) private lemma add_occurs_msgs_soundness_aux3: assumes P: "\T \ set P. admissible_transaction T" and A: "\ \ reachable_constraints (map add_occurs_msgs P)" "welltyped_constraint_model \ (rm_occurs_msgs_constr \)" and I: "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \n. \ x = Fun (Val n) []" (is "?I \") shows "welltyped_constraint_model \ \" (is "?Q \ \") using A I proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) let ?f = rm_occurs_msgs_constr let ?sem = "\B. strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}) (unlabel B) \" define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" define \ where "\ \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" obtain T' where T': "T' \ set P" "T = add_occurs_msgs T'" using step.hyps(2) by fastforce then obtain A' B' C' D' E' F' where T'': "T' = Transaction A' B' C' D' E' F'" using prot_transaction.exhaust by blast have P': "\T \ set (map add_occurs_msgs P). admissible_transaction' T" "\T \ set (map add_occurs_msgs P). admissible_transaction_occurs_checks T" "\T \ set P. admissible_transaction' T" "\T \ set P. admissible_transaction_no_occurs_msgs T" using P admissible_transactionE' add_occurs_msgs_admissible_occurs_checks by (fastforce,fastforce,fastforce,fastforce) note T_adm = bspec[OF P'(1) step.hyps(2)] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] note T'_adm = bspec[OF P'(3) T'(1)] note T'_no_occ = bspec[OF P'(4) T'(1)] note T'_wf = admissible_transaction_is_wellformed_transaction(1)[OF T'_adm] note \_empty = admissible_transaction_decl_subst_empty[OF T_adm step.hyps(3)] note T_fresh_val = admissible_transactionE(2)[OF T_adm] have 0: "?Q \ (?f \)" "?I \" "?I \" by (metis step.prems(1) welltyped_constraint_model_prefix rm_occurs_msgs_constr_append, simp_all add: step.prems(2) \_def \_def) note IH = step.IH[OF 0(1,2)] have I': "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using step.prems(1) unfolding welltyped_constraint_model_def constraint_model_def by blast+ have 1: "\x \ fv_transaction T. \t. \ x = occurs t" "\x \ fv_transaction T. \ x \ Fun OccursSec []" using admissible_transaction_decl_fresh_renaming_subst_not_occurs[OF T_adm step.hyps(3,4,5)] unfolding \_def[symmetric] by simp_all have "(ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?f \) \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" using rm_occurs_msgs_constr_ik_subset by fast hence 2: "?sem (?f \)" using step.prems(1) strand_sem_append_stateful[of "{}" "{}" "unlabel (?f \)" "unlabel (?f \)"] rm_occurs_msgs_constr_dbupd\<^sub>s\<^sub>s\<^sub>t_eq[of \ \ "{}"] rm_occurs_msgs_constr_append[of \ \] strand_sem_ik_mono_stateful[of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?f \) \\<^sub>s\<^sub>e\<^sub>t \" _ _ _ "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \"] unfolding welltyped_constraint_model_def constraint_model_def \_def[symmetric] \_def[symmetric] by auto note 3 = rm_occurs_msgs_constr_transaction_strand''[ OF T'_adm T'_no_occ 1[unfolded T'(2)], unfolded \_def[symmetric] T'(2)[symmetric]] note 4 = add_occurs_msgs_cases[OF T'', unfolded T'(2)[symmetric]] define xs where "xs \ fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T'))" define flt where "flt \ filter (\x. x \ set (transaction_fresh T'))" define occs where "occs \ map (\x. occurs (Var x)::('fun,'atom,'sets,'lbl) prot_term)" note 6 = add_occurs_msgs_transaction_strand_cases(7,8,9)[ of T' \, unfolded xs_def[symmetric] flt_def[symmetric] occs_def[symmetric] T'(2)[symmetric]] have 7: "x \ fv_transaction T - set (transaction_fresh T)" when x: "x \ set (flt xs)" for x using that fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t add_occurs_msgs_vars_eq(3,9)[OF T'_adm] unfolding xs_def flt_def T'(2) by force have 9: "\y. \ x = Var y" when x: "x \ set (flt xs)" for x proof - have *: "x \ fst ` set (transaction_decl T ())" using admissible_transactionE(1)[OF T_adm] by simp have **: "x \ set (transaction_fresh T)" using 7[OF x] by simp show ?thesis using transaction_decl_fresh_renaming_substs_range(4)[OF step.hyps(3,4,5) * **] unfolding \_def by blast qed have 8: "\y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \ x = Var y" when x: "x \ set (flt xs)" for x proof - note * = 7[OF x] obtain y where y: "\ x = Var y" using 9[OF x] by blast have "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))" by (metis * Diff_iff fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq) have "\ x \ \ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)" using * by fast hence "fv (\ x) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv_transaction T)" by force hence "fv (\ x) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using fv\<^sub>s\<^sub>s\<^sub>t_subst_if_no_bvars[OF admissible_transactionE(4)[OF T_adm], of \] by (metis unlabel_subst) hence "fv (\ x) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" by (metis fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq \_def) thus ?thesis using y by simp qed have \_var_is_\_val: "\n. \ x = Fun (Val n) []" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" for x using step.prems(2) x unfolding \_def[symmetric] \_def[symmetric] by auto have T'_var_is_\\_val: "\n. \ x \ \ = Fun (Val n) []" when x: "x \ set (flt xs)" for x using 8[OF x] \_var_is_\_val by force (* TODO: extract lemma *) have poschecks_has_occ: "occurs (Fun (Val n) []) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" when x: "\ac: t \ s\ \ set (unlabel \)" and n: "t \ \ = Fun (Val n) []" for ac t s n proof - have *: "(t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" proof - obtain t' s' where t': "\ac: t' \ s'\ \ set (unlabel (transaction_checks T'))" "t = t' \ \" "s = s' \ \" using 4(8) x stateful_strand_step_mem_substD(6) wellformed_transaction_strand_unlabel_memberD(10)[OF T_wf(1)] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(6) unfolding \_def by (metis (no_types) unlabel_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst) have "(t' \ \ \ \, s' \ \ \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" using t'(1) 2 wellformed_transaction_unlabel_sem_iff[ OF T'_wf(1) I'(2,3), of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" \] unfolding 3 by blast thus ?thesis using t'(2,3) by simp qed have "t' \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" when "insert\t',s'\ \ set (unlabel \)" for t' s' using that by force have "t \ \ \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" proof - obtain t' s' where t': "insert\t',s'\ \ set (unlabel \)" "t \ \ = t' \ \" "s \ \ = s' \ \" using * db\<^sub>s\<^sub>s\<^sub>t_in_cases[of "t \ \" "s \ \" "unlabel \" \ "[]"] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of "unlabel \" \ "[]"] by auto have t'': "t' = t \ \ \ (\y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. t' = Var y \ \ y = t \ \)" using t'(1,2) stateful_strand_step_fv_subset_cases(4) unfolding n by (cases t') (force,force) thus ?thesis proof assume "t' = t \ \" thus ?thesis using t'(1) by force next assume "\y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. t' = Var y \ \ y = t \ \" then obtain y where y: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "\ y = t \ \" by blast have "\\<^sub>v y = TAtom Value" using y(2) wt_subst_trm''[OF I'(1), of "Var y"] unfolding n by simp hence "\B. prefix B \ \ t \ \ \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B" by (metis y constraint_model_Value_var_in_constr_prefix[OF step.hyps(1) IH P'(1,2)]) thus ?thesis unfolding prefix_def by auto qed qed thus ?thesis using reachable_constraints_occurs_fact_ik_case'[OF step.hyps(1) P'(1,2)] unfolding n by blast qed (* TODO: extract lemma *) have snds_has_occ: "occurs (Fun (Val n) []) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" when ts: "send\ts\ \ set (unlabel \)" and n: "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t set ts \\<^sub>s\<^sub>e\<^sub>t \" for ts n proof - have "receive\ts\ \ set (unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using ts dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(2) unfolding \_def by metis then obtain ts' where ts': "receive\ts'\ \ set (unlabel (transaction_strand T))" "ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" by (metis subst_lsst_memD(1) unlabel_in unlabel_mem_has_label) have "?sem (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using 2 strand_sem_append_stateful[of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}"] unfolding 3 transaction_dual_subst_unlabel_unfold by blast moreover have "list_all is_Receive (unlabel (transaction_receive T'))" using T'_wf unfolding wellformed_transaction_def by blast hence "list_all is_Send (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" by (metis subst_lsst_unlabel subst_sst_list_all(2) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(1)) hence "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = {}" using in_ik\<^sub>s\<^sub>s\<^sub>t_iff unfolding list_all_iff is_Send_def by fast ultimately have *: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" when "send\ts\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" "t \ set ts" for t ts using strand_sem_stateful_sends_deduct[OF _ that] by simp hence *: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \ \ \" when ts: "receive\ts\ \ set (unlabel (transaction_receive T'))" "t \ set ts" for t ts using ts(2) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(2)[of "ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" "transaction_receive T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] stateful_strand_step_subst_inI(2)[OF ts(1), of \, unfolded unlabel_subst] by auto have **: "set (flt xs) = fv_transaction T' - set (transaction_fresh T')" using fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t unfolding flt_def xs_def by fastforce have rcv_case: ?thesis when "ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" "Fun (Val n) [] \\<^sub>s\<^sub>e\<^sub>t set ts \\<^sub>s\<^sub>e\<^sub>t \" "receive\ts'\ \ set (unlabel (transaction_receive T'))" for ts ts' using that * reachable_constraints_occurs_fact_ik_case''[OF step.hyps(1) IH P'(1,2)] by auto have "receive\ts'\ \ set (unlabel (transaction_receive T))" using wellformed_transaction_strand_unlabel_memberD(1)[OF T_wf] ts'(1) by blast hence "(ts' = map (\x. occurs (Var x)) (flt xs) \ ts' \ []) \ receive\ts'\ \ set (unlabel (transaction_receive T'))" (is "?A \ ?B") using ** ts'(1) add_occurs_msgs_cases(13)[OF T''] unfolding T'(2)[symmetric] xs_def[symmetric] flt_def[symmetric] by force thus ?thesis proof assume ?A then obtain x where x: "x \ set (flt xs)" "Fun (Val n) [] \ \ x \ \" using ts' n by fastforce have x': "\ x \ \ = Fun (Val n) []" "x \ fv_transaction T" "x \ set (transaction_fresh T)" "x \ fv_transaction T'" "x \ set (transaction_fresh T')" using x(2) T'_var_is_\\_val[OF x(1)] 7[OF x(1)] ** x(1) by fastforce+ let ?snds = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?chks = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" have B_subsets: "set ?chks \ set (unlabel \)" unfolding \_def transaction_dual_subst_unlabel_unfold 4(8) by fastforce from admissible_transaction_fv_in_receives_or_selects_dual_subst[OF T'_adm x'(4,5), of \] show ?thesis proof assume "\ts. send\ts\ \ set ?snds \ \ x \\<^sub>s\<^sub>e\<^sub>t set ts" then obtain ss where ss: "send\ss\ \ set ?snds" "\ x \\<^sub>s\<^sub>e\<^sub>t set ss" by blast obtain ss' where ss': "ss = ss' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" "receive\ss'\ \ set (unlabel (transaction_receive T'))" by (metis ss(1) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(2) subst_lsst_memD(1) unlabel_in unlabel_mem_has_label) show ?thesis using rcv_case[OF ss'(1) _ ss'(2)] subst_subterms[OF ss(2), of \] x'(1) by argo qed (use B_subsets poschecks_has_occ[OF _ x'(1)] in blast) qed (metis rcv_case[OF ts'(2) n]) qed (* TODO: extract lemma *) have "occurs (\ x \ \) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" when x: "x \ set (flt xs)" for x proof - have "(\ac s. \ac: \ x \ s\ \ set (unlabel \)) \ (\ts. send\ts\ \ set (unlabel \) \ \ x \\<^sub>s\<^sub>e\<^sub>t set ts)" (is "(\ac s. ?A ac s) \ (\ts. ?B1 ts \ ?B2 ts)") using 7[OF x] admissible_transaction_fv_in_receives_or_selects_dual_subst[OF T_adm, of x \] unfolding \_def transaction_dual_subst_unlabel_unfold by auto thus ?thesis proof assume "\ac s. ?A ac s" then obtain ac s where s: "?A ac s" by blast show ?thesis using poschecks_has_occ[OF s] T'_var_is_\\_val[OF x] by force next assume "\ts. ?B1 ts \ ?B2 ts" then obtain ts where ts: "?B1 ts" "?B2 ts" by meson have ts': "\ x \ \ \\<^sub>s\<^sub>e\<^sub>t set ts \\<^sub>s\<^sub>e\<^sub>t \" by (metis ts(2) subst_subterms) show ?thesis using snds_has_occ[OF ts(1)] ts' T'_var_is_\\_val[OF x] by force qed qed hence "occurs (\ x \ \) \ \ \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" when "x \ set (flt xs)" for x using that by fast moreover have "occurs (\ x \ \) \ \ = occurs (\ x \ \)" for x using subst_ground_ident[OF interpretation_grounds[OF I'(2), of "\ x"], of \] by simp ultimately have "occurs (\ x \ \) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" when "x \ set (flt xs)" for x using that by auto hence "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" when "t \ set (occs (flt xs) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" for t using that unfolding occs_def by auto hence occs_sem: "?sem [\\, send\occs (flt xs) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\]" by auto (* TODO: extract lemma *) have "?sem \" proof - let ?IK = "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" let ?DB = "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ {}" let ?snds = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" let ?snds_occs = "(\\, send\occs (flt xs) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\)#?snds" let ?chks = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" let ?upds = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" let ?rcvs = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" note * = strand_sem_append_stateful[of _ _ _ _ \] note ** = transaction_dual_subst_unlabel_unfold have ***: "\M. M \ (ik\<^sub>s\<^sub>s\<^sub>t [] \\<^sub>s\<^sub>e\<^sub>t \) = M" "\D. dbupd\<^sub>s\<^sub>s\<^sub>t [] \ D = D" by simp_all have snds_sem: "?sem ?snds" "?sem ?snds_occs" using 2 occs_sem *[of ?IK ?DB] unfolding 3 ** by (blast, fastforce) have "list_all is_Receive (unlabel (transaction_receive T'))" using T'_wf unfolding wellformed_transaction_def by blast hence "list_all is_Send (unlabel ?snds)" "list_all is_Send (unlabel ?snds_occs)" using subst_sst_list_all(2) unlabel_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(1) by (metis, metis (no_types) list.pred_inject(2) stateful_strand_step.disc(1) unlabel_Cons(1)) hence "\a \ set (unlabel ?snds). \is_Receive a \ \is_Insert a \ \is_Delete a" "\a \ set (unlabel ?snds_occs). \is_Receive a \ \is_Insert a \ \is_Delete a" unfolding list_all_iff by (blast,blast) hence snds_no_upds: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?snds \\<^sub>s\<^sub>e\<^sub>t \ = {}" "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel ?snds) \ ?DB = ?DB" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?snds_occs) \\<^sub>s\<^sub>e\<^sub>t \ = {}" "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel ?snds_occs) \ ?DB = ?DB" by (metis ik\<^sub>s\<^sub>s\<^sub>t_snoc_no_receive_empty, metis dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd, metis ik\<^sub>s\<^sub>s\<^sub>t_snoc_no_receive_empty, metis dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd) have chks_sem: "?sem ?chks" using 2 snds_no_upds * unfolding 3 ** by auto have "list_all is_Check_or_Assignment (unlabel (transaction_checks T'))" using T'_wf unfolding wellformed_transaction_def by blast hence "list_all is_Check_or_Assignment (unlabel ?chks)" by (metis (no_types) subst_sst_list_all(11) unlabel_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(11)) hence "\a \ set (unlabel ?chks). \is_Receive a \ \is_Insert a \ \is_Delete a" unfolding list_all_iff by blast hence chks_no_upds: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?chks \\<^sub>s\<^sub>e\<^sub>t \ = {}" "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel ?chks) \ ?DB = ?DB" by (metis ik\<^sub>s\<^sub>s\<^sub>t_snoc_no_receive_empty, metis dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd) have upds_sem: "?sem ?upds" using 2 snds_no_upds chks_no_upds * unfolding 3 ** by auto have "list_all is_Send (unlabel (transaction_send T'))" using T'_wf unfolding wellformed_transaction_def by fast hence "list_all is_Send (unlabel (transaction_send T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" by (metis (no_types, opaque_lifting) subst_sst_list_all(1) unlabel_subst) hence rcvs_is_rcvs: "list_all is_Receive (unlabel ?rcvs)" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_list_all(2) by blast have rcvs_sem: "strand_sem_stateful M D (unlabel rcvs) \" when "list_all is_Receive (unlabel rcvs)" for M D and rcvs::"('fun, 'atom, 'sets, 'lbl) prot_strand" using rcvs_is_rcvs strand_sem_receive_prepend_stateful[of M D "[]" \, OF _ that] by auto have B_sem: "?sem (?snds@?chks@?upds@rcvs)" "?sem (?snds_occs@?chks@?upds@rcvs)" when "list_all is_Receive (unlabel rcvs)" for rcvs using strand_sem_append_stateful[of _ _ _ _ \] snds_sem snds_no_upds chks_sem chks_no_upds upds_sem rcvs_sem[OF that] by (force, force) show ?thesis proof (cases "transaction_fresh T' = []") case True show ?thesis using B_sem[OF rcvs_is_rcvs] unfolding \_def 6(1)[OF True] by force next case False note F = this show ?thesis proof (cases "\ts F'. transaction_send T' = \\, send\ts\\#F'") case True obtain ts F' rcvs' where F': "transaction_send T' = \\, send\ts\\#F'" "\ = (if flt xs = [] then ?snds else ?snds_occs)@?chks@?upds@rcvs'" "rcvs' = \\, receive\occs (transaction_fresh T')@ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (F' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using 6(3)[OF F True] unfolding \_def by blast have *: "list_all is_Receive (unlabel rcvs')" using rcvs_is_rcvs dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_Cons(1)[of \ ts "F' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] subst_lsst_cons[of "\\, send\ts\\" F' \] unfolding F'(1,3) list_all_iff by auto show ?thesis using B_sem[OF *] unfolding F'(2) by fastforce next case False have *: "list_all is_Receive (unlabel (\\, receive\occs (transaction_fresh T') \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\#?rcvs))" using rcvs_is_rcvs by auto show ?thesis using B_sem[OF *] unfolding \_def 6(2)[OF F False] by fastforce qed qed qed thus ?case using IH strand_sem_append_stateful[of "{}" "{}" "unlabel \" "unlabel \" \] unfolding welltyped_constraint_model_def constraint_model_def \_def[symmetric] \_def[symmetric] by simp qed simp theorem add_occurs_msgs_soundness: defines "wt_attack \ \\ \ l n. welltyped_constraint_model \ (\@[(l, send\[attack\n\]\)])" assumes P: "\T \ set P. admissible_transaction T" "has_initial_value_producing_transaction P" and A: "\ \ reachable_constraints P" "wt_attack \ \ l n" shows "\\ \ reachable_constraints (map add_occurs_msgs P). \\. wt_attack \ \ l n" proof - have P': "\T \ set (map add_occurs_msgs P). admissible_transaction' T" "\T \ set (map add_occurs_msgs P). admissible_transaction_occurs_checks T" "\T \ set P. admissible_transaction' T" "\T \ set P. admissible_transaction_no_occurs_msgs T" using P admissible_transactionE' add_occurs_msgs_admissible_occurs_checks by (fastforce,fastforce,fastforce,fastforce) obtain \' \ where A': "\' \ reachable_constraints P" "wt_attack \ \' l n" "\x\fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \'. \n. \ x = Fun (Val n) []" using add_occurs_msgs_soundness_aux1[OF P'(3) P(2) A[unfolded wt_attack_def]] unfolding wt_attack_def by blast have J: "welltyped_constraint_model \ \'" using A'(2) welltyped_constraint_model_prefix unfolding wt_attack_def by blast obtain \ where B: "\ \ reachable_constraints (map add_occurs_msgs P)" "\' = rm_occurs_msgs_constr \" using add_occurs_msgs_soundness_aux2[OF P(1) A'(1)] by blast have J': "welltyped_constraint_model \ \" using add_occurs_msgs_soundness_aux3[OF P(1) B(1) J[unfolded B(2)]] A'(3) rm_occurs_msgs_constr_reachable_constraints_fv_eq[OF P'(3,4) B(1)] unfolding wt_attack_def B(2) by blast obtain ts where ts: "receive\ts\ \ set (unlabel \)" "attack\n\ \ set ts" using reachable_constraints_receive_attack_if_attack''[OF P'(3) A'(1,2)[unfolded wt_attack_def]] rm_occurs_msgs_constr_receive_attack_iff[of n \] unfolding B(2)[symmetric] by auto have J'': "wt_attack \ \ l n" using welltyped_constraint_model_attack_if_receive_attack[OF J' ts] unfolding wt_attack_def by fast show ?thesis using B(1) J'' by blast qed end end subsection \Automatically Checking Protocol Security in a Typed Model\ context stateful_protocol_model begin definition abs_intruder_knowledge ("\\<^sub>i\<^sub>k") where "\\<^sub>i\<^sub>k S \ \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \)" definition abs_value_constants ("\\<^sub>v\<^sub>a\<^sub>l\<^sub>s") where "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s S \ \ {t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S) \\<^sub>s\<^sub>e\<^sub>t \. \n. t = Fun (Val n) []} \\<^sub>\\<^sub>s\<^sub>e\<^sub>t \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \)" definition abs_term_implications ("\\<^sub>t\<^sub>i") where "\\<^sub>t\<^sub>i \ T \ \ \ {(s,t) | s t x. s \ t \ x \ fv_transaction T \ x \ set (transaction_fresh T) \ Fun (Abs s) [] = \ x \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) \ Fun (Abs t) [] = \ x \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \)}" lemma abs_intruder_knowledge_append: "\\<^sub>i\<^sub>k (A@B) \ = (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B) \) \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B) \)" by (metis unlabel_append abs_set_union image_Un ik\<^sub>s\<^sub>s\<^sub>t_append abs_intruder_knowledge_def) lemma abs_value_constants_append: fixes A B::"('a,'b,'c,'d) prot_strand" shows "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s (A@B) \ = {t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t \. \n. t = Fun (Val n) []} \\<^sub>\\<^sub>s\<^sub>e\<^sub>t \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B) \) \ {t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B) \\<^sub>s\<^sub>e\<^sub>t \. \n. t = Fun (Val n) []} \\<^sub>\\<^sub>s\<^sub>e\<^sub>t \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B) \)" proof - define a0 where "a0 \ \\<^sub>0 (db\<^sub>s\<^sub>s\<^sub>t (unlabel (A@B)) \)" define M where "M \ \a::('a,'b,'c,'d) prot_strand. {t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t a) \\<^sub>s\<^sub>e\<^sub>t \. \n. t = Fun (Val n) []}" have "M (A@B) = M A \ M B" using image_Un[of "\x. x \ \" "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)"] unfolding M_def unlabel_append[of A B] trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel A" "unlabel B"] by blast hence "M (A@B) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0 = (M A \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0) \ (M B \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0)" by (simp add: abs_set_union) thus ?thesis unfolding abs_value_constants_def a0_def M_def by force qed lemma transaction_renaming_subst_has_no_pubconsts_abss: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_renaming_subst \ P A" shows "subst_range \ \ pubval_terms = {}" (is ?A) and "subst_range \ \ abs_terms = {}" (is ?B) proof - { fix t assume "t \ subst_range \" then obtain x where "t = Var x" using transaction_renaming_subst_is_renaming(1)[OF assms] by force hence "t \ pubval_terms" "t \ abs_terms" by simp_all } thus ?A ?B by auto qed lemma transaction_fresh_subst_has_no_pubconsts_abss: fixes \::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_fresh_subst \ T \" "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" shows "subst_range \ \ pubval_terms = {}" (is ?A) and "subst_range \ \ abs_terms = {}" (is ?B) proof - { fix t assume "t \ subst_range \" then obtain x where "x \ set (transaction_fresh T)" "\ x = t" using assms(1) unfolding transaction_fresh_subst_def by auto then obtain n where "t = Fun (Val n) []" using transaction_fresh_subst_sends_to_val[OF assms(1)] assms(2) by meson hence "t \ pubval_terms" "t \ abs_terms" unfolding is_PubConstValue_def by simp_all } thus ?A ?B by auto qed lemma reachable_constraints_GSMP_no_pubvals_abss: assumes "\ \ reachable_constraints P" and P: "\T \ set P. admissible_transaction' T" and \: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "\n. PubConst Value n \ \(funs_term ` (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "\n. Abs n \ \(funs_term ` (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" shows "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ GSMP (\T \ set P. trms_transaction T) - (pubval_terms \ abs_terms)" (is "?A \ ?B") using assms(1) \(4,5) proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \ \) define trms_P where "trms_P \ (\T \ set P. trms_transaction T)" define T' where "T' \ transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" have \_elim: "\ \\<^sub>s \ \\<^sub>s \ = \ \\<^sub>s \" using admissible_transaction_decl_subst_empty[OF bspec[OF P step.hyps(2)] step.hyps(3)] by simp note T_fresh = admissible_transactionE(2)[OF bspec[OF P step.hyps(2)]] note T_no_bvars = admissible_transactionE(4)[OF bspec[OF P step.hyps(2)]] have T_no_PubVal: "\T \ set P. \n. PubConst Value n \ \(funs_term ` trms_transaction T)" and T_no_Abs: "\T \ set P. \n. Abs n \ \(funs_term ` trms_transaction T)" using admissible_transactions_no_Value_consts''[OF bspec[OF P]] by metis+ have \': "\n. PubConst Value n \ \ (funs_term ` (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "\n. Abs n \ \ (funs_term ` (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using step.prems fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \"] unlabel_append[of \] by auto have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (rm_vars (set X) (\ \\<^sub>s \ \\<^sub>s \))" for X using wt_subst_rm_vars[of "\ \\<^sub>s \ \\<^sub>s \" "set X"] transaction_decl_fresh_renaming_substs_wt[OF step.hyps(3-5)] by metis hence wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t ((rm_vars (set X) (\ \\<^sub>s \ \\<^sub>s \)) \\<^sub>s \)" for X using \(2) wt_subst_compose by fast have wftrms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range ((rm_vars (set X) (\ \\<^sub>s \ \\<^sub>s \)) \\<^sub>s \))" for X using wf_trms_subst_compose[OF wf_trms_subst_rm_vars' \(3)] transaction_decl_fresh_renaming_substs_range_wf_trms[OF step.hyps(3-5)] by fast have "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T') \\<^sub>s\<^sub>e\<^sub>t \ \ ?B" proof fix t assume "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T') \\<^sub>s\<^sub>e\<^sub>t \" hence "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \\<^sub>s\<^sub>e\<^sub>t \" using trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq by blast then obtain s X where s: "s \ trms_transaction T" "t = s \ rm_vars (set X) (\ \\<^sub>s \ \\<^sub>s \) \\<^sub>s \" "set X \ bvars_transaction T" using trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subst'' unfolding T'_def by blast define \ where "\ \ rm_vars (set X) (\ \\<^sub>s \ \\<^sub>s \)" note 0 = pubval_terms_subst_range_comp[OF transaction_fresh_subst_has_no_pubconsts_abss(1)[OF step.hyps(4) T_fresh] transaction_renaming_subst_has_no_pubconsts_abss(1)[OF step.hyps(5)]] abs_terms_subst_range_comp[OF transaction_fresh_subst_has_no_pubconsts_abss(2)[OF step.hyps(4) T_fresh] transaction_renaming_subst_has_no_pubconsts_abss(2)[OF step.hyps(5)]] have 1: "s \ trms_P" using step.hyps(2) s(1) unfolding trms_P_def by auto have s_nin: "s \ pubval_terms" "s \ abs_terms" using 1 T_no_PubVal T_no_Abs funs_term_Fun_subterm unfolding trms_P_def is_PubConstValue_def is_Abs_def is_PubConst_def by (fastforce, blast) have 2: "(\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')) \ pubval_terms = {}" "(\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')) \ abs_terms = {}" "subst_range (\ \\<^sub>s \ \\<^sub>s \) \ pubval_terms = {}" "subst_range (\ \\<^sub>s \ \\<^sub>s \) \ abs_terms = {}" using 0 step.prems funs_term_Fun_subterm unfolding T'_def \_def \_elim by (fastforce simp add: is_PubConstValue_def is_PubConst_def, fastforce simp add: is_Abs_def, argo, argo) have "subst_range \ \ subst_range (\ \\<^sub>s \ \\<^sub>s \)" using rm_vars_img_subset unfolding \_def \_elim by blast hence 3: "subst_range \ \ pubval_terms = {}" "subst_range \ \ abs_terms = {}" using 2(3,4) step.prems funs_term_Fun_subterm unfolding T'_def \_def \_elim by (blast,blast) have "(\ ` fv (s \ \)) \ pubval_terms = {}" "(\ ` fv (s \ \)) \ abs_terms = {}" proof - have "\ = \ \\<^sub>s \ \\<^sub>s \" "bvars_transaction T = {}" "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" using s(3) T_no_bvars step.hyps(2) rm_vars_empty vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel T'"] bvars\<^sub>s\<^sub>s\<^sub>t_subst[of "unlabel (transaction_strand T)" "\ \\<^sub>s \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \ \\<^sub>s \"] unfolding \_def T'_def by simp_all hence "fv (s \ \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" using trms\<^sub>s\<^sub>s\<^sub>t_fv_subst_subset[OF s(1), of \] unlabel_subst[of "transaction_strand T" \] unfolding T'_def by auto moreover have "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" using fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')"] unlabel_append[of \ "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'"] fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of T'] by simp_all hence "\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ pubval_terms = {}" "\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ abs_terms = {}" using 2(1,2) by blast+ ultimately show "(\ ` fv (s \ \)) \ pubval_terms = {}" "(\ ` fv (s \ \)) \ abs_terms = {}" by blast+ qed hence \\\_disj: "((\ \\<^sub>s \) ` fv s) \ pubval_terms = {}" "((\ \\<^sub>s \) ` fv s) \ abs_terms = {}" using pubval_terms_subst_range_comp'[of \ "fv s" \] abs_terms_subst_range_comp'[of \ "fv s" \] pubval_terms_subst_range_disj[OF 3(1), of s] abs_terms_subst_range_disj[OF 3(2), of s] by (simp_all add: subst_apply_fv_unfold) have 4: "t \ pubval_terms" "t \ abs_terms" using s(2) s_nin \\\_disj pubval_terms_subst[of s "rm_vars (set X) (\ \\<^sub>s \ \\<^sub>s \) \\<^sub>s \"] pubval_terms_subst_range_disj[of "rm_vars (set X) (\ \\<^sub>s \ \\<^sub>s \) \\<^sub>s \" s] abs_terms_subst[of s "rm_vars (set X) (\ \\<^sub>s \ \\<^sub>s \) \\<^sub>s \"] abs_terms_subst_range_disj[of "rm_vars (set X) (\ \\<^sub>s \ \\<^sub>s \) \\<^sub>s \" s] unfolding \_def by blast+ have "t \ SMP trms_P" "fv t = {}" by (metis s(2) SMP.Substitution[OF SMP.MP[OF 1] wt wftrms, of X], metis s(2) subst_subst_compose[of s "rm_vars (set X) (\ \\<^sub>s \ \\<^sub>s \)" \] interpretation_grounds[OF \(1), of "s \ rm_vars (set X) (\ \\<^sub>s \ \\<^sub>s \)"]) hence 5: "t \ GSMP trms_P" unfolding GSMP_def by simp show "t \ ?B" using 4 5 by (auto simp add: trms_P_def) qed thus ?case using step.IH[OF \'] trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \"] unlabel_append[of \] image_Un[of "\x. x \ \" "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] by (simp add: T'_def) qed simp lemma \\<^sub>t\<^sub>i_covers_\\<^sub>0_aux: assumes \_reach: "\ \ reachable_constraints P" and T: "T \ set P" and \: "welltyped_constraint_model \ (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "t = Fun (Val n) [] \ t = Var x" and neq: "t \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) \ t \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \)" shows "\y \ fv_transaction T - set (transaction_fresh T). t \ \ = (\ \\<^sub>s \ \\<^sub>s \) y \ \ \ \\<^sub>v y = TAtom Value" proof - let ?\' = "\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" let ?\ = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))" let ?\' = "?\ \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" let ?\'' = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" have \_interp: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (metis \ welltyped_constraint_model_def constraint_model_def, metis \ welltyped_constraint_model_def, metis \ welltyped_constraint_model_def constraint_model_def) note T_adm = bspec[OF P(1) T] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] note T_adm_upds = admissible_transaction_is_wellformed_transaction(3)[OF T_adm] have T_fresh_vars_value_typed: "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" using T P(1) protocol_transaction_vars_TAtom_typed(3)[of T] P(1) by simp note \_empty = admissible_transaction_decl_subst_empty[OF T_adm \] note \\\_wt = transaction_decl_fresh_renaming_substs_wt[OF \ \ \] have \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (metis reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s admissible_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P(1) \_reach) hence t_wf: "wf\<^sub>t\<^sub>r\<^sub>m t" using t by auto have \_no_val_bvars: "\TAtom Value \ \\<^sub>v x" when "x \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" for x using P(1) reachable_constraints_no_bvars \_reach vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] that admissible_transactionE(4) by fast have x': "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" when "t = Var x" using that t by (simp add: var_subterm_trms\<^sub>s\<^sub>s\<^sub>t_is_vars\<^sub>s\<^sub>s\<^sub>t) have "\f \ funs_term (t \ \). is_Val f" using abs_eq_if_no_Val neq by metis hence "\n T. Fun (Val n) T \ t \ \" using funs_term_Fun_subterm unfolding is_Val_def by fast hence "TAtom Value \ \ (Var x)" when "t = Var x" using wt_subst_trm''[OF \_wt, of "Var x"] that subtermeq_imp_subtermtypeeq[of "t \ \"] wf_trm_subst[OF \_wf, of t] t_wf by fastforce hence x_val: "\\<^sub>v x = TAtom Value" when "t = Var x" using reachable_constraints_vars_TAtom_typed[OF \_reach P x'] that by fastforce hence x_fv: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" when "t = Var x" using x' using reachable_constraints_Value_vars_are_fv[OF \_reach P x'] that by blast then obtain m where m: "t \ \ = Fun (Val m) []" using constraint_model_Value_term_is_Val[ OF \_reach welltyped_constraint_model_prefix[OF \] P P_occ, of x] t(2) x_val by force hence 0: "\\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) m \ \\<^sub>0 (db\<^sub>s\<^sub>s\<^sub>t (unlabel \@?\'') \) m" using neq by (simp add: unlabel_def) have t_val: "\ t = TAtom Value" using x_val t by force obtain u s where s: "t \ \ = u \ \" "insert\u,s\ \ set ?\' \ delete\u,s\ \ set ?\'" using to_abs_neq_imp_db_update[OF 0] m by (metis (no_types, lifting) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst subst_lsst_unlabel) then obtain u' s' where s': "u = u' \ \ \\<^sub>s \ \\<^sub>s \" "s = s' \ \ \\<^sub>s \ \\<^sub>s \" "insert\u',s'\ \ set ?\ \ delete\u',s'\ \ set ?\" using stateful_strand_step_mem_substD(4,5) by blast hence s'': "insert\u',s'\ \ set (unlabel (transaction_strand T)) \ delete\u',s'\ \ set (unlabel (transaction_strand T))" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(4,5)[of u' s' "transaction_strand T"] by simp_all then obtain y where y: "y \ fv_transaction T" "u' = Var y" using transaction_inserts_are_Value_vars[OF T_wf T_adm_upds, of u' s'] transaction_deletes_are_Value_vars[OF T_wf T_adm_upds, of u' s'] stateful_strand_step_fv_subset_cases(4,5)[of u' s' "unlabel (transaction_strand T)"] by auto - hence 1: "t \ \ = (\ \\<^sub>s \ \\<^sub>s \) y \ \" using y s(1) s'(1) by (metis subst_apply_term.simps(1)) + hence 1: "t \ \ = (\ \\<^sub>s \ \\<^sub>s \) y \ \" using y s(1) s'(1) by (metis eval_term.simps(1)) have 2: "y \ set (transaction_fresh T)" when "(\ \\<^sub>s \ \\<^sub>s \) y \ \ \ \ y" using transaction_fresh_subst_grounds_domain[OF \, of y] subst_compose[of \ \ y] that \_empty by (auto simp add: subst_ground_ident) have 3: "y \ set (transaction_fresh T)" when "(\ \\<^sub>s \ \\<^sub>s \) y \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using 2 that \ unfolding transaction_fresh_subst_def by fastforce have 4: "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \\<^sub>v x = TAtom Value \ (\B. prefix B \ \ x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \ \ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B))" by (metis welltyped_constraint_model_prefix[OF \] constraint_model_Value_var_in_constr_prefix[OF \_reach _ P P_occ]) have 5: "\\<^sub>v y = TAtom Value" using 1 t_val wt_subst_trm''[OF \\\_wt, of "Var y"] wt_subst_trm''[OF \_wt, of t] wt_subst_trm''[OF \_wt, of "(\ \\<^sub>s \ \\<^sub>s \) y"] by (auto simp del: subst_subst_compose) have "y \ set (transaction_fresh T)" proof (cases "t = Var x") case True (* \ x occurs in \ but not in subst_range \, so y cannot be fresh *) hence *: "\ x = Fun (Val m) []" "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "\ x = (\ \\<^sub>s \ \\<^sub>s \) y \ \" using m t(1) 1 x_fv x' by (force, blast, force) obtain B where B: "prefix B \" "\ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" using *(2) 4 x_val[OF True] by fastforce hence "\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" using transaction_fresh_subst_range_fresh(1)[OF \] trms\<^sub>s\<^sub>s\<^sub>t_unlabel_prefix_subset(1)[of B] unfolding prefix_def by fast thus ?thesis using *(1,3) B(2) 2 by (metis subst_imgI term.distinct(1)) next case False hence "t \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using t by simp thus ?thesis using 1 3 by argo qed thus ?thesis using 1 5 y(1) by fast qed lemma \\<^sub>t\<^sub>i_covers_\\<^sub>0_Var: assumes \_reach: "\ \ reachable_constraints P" and T: "T \ set P" and \: "welltyped_constraint_model \ (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" shows "\ x \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \) \ timpl_closure_set {\ x \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)} (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \)" proof - define a0 where "a0 \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" define a0' where "a0' \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \)" define a3 where "a3 \ \\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \" have \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (metis reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s admissible_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P(1) \_reach) note T_adm = bspec[OF P(1) T] note \_empty = admissible_transaction_decl_subst_empty[OF T_adm \] note \\\_wt = transaction_decl_fresh_renaming_substs_wt[OF \ \ \] have \_interp: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (metis \ welltyped_constraint_model_def constraint_model_def, metis \ welltyped_constraint_model_def, metis \ welltyped_constraint_model_def constraint_model_def) have "\\<^sub>v x = Var Value \ (\a. \\<^sub>v x = Var (prot_atom.Atom a))" using reachable_constraints_vars_TAtom_typed[OF \_reach P, of x] x vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] by auto hence "\ x \\<^sub>\ a0' \ timpl_closure_set {\ x \\<^sub>\ a0} a3" proof assume x_val: "\\<^sub>v x = TAtom Value" show "\ x \\<^sub>\ a0' \ timpl_closure_set {\ x \\<^sub>\ a0} a3" proof (cases "\ x \\<^sub>\ a0 = \ x \\<^sub>\ a0'") case False hence "\y \ fv_transaction T - set (transaction_fresh T). \ x = (\ \\<^sub>s \ \\<^sub>s \) y \ \ \ \\<^sub>v y = TAtom Value" using \\<^sub>t\<^sub>i_covers_\\<^sub>0_aux[OF \_reach T \ \ \ \ P P_occ fv\<^sub>s\<^sub>s\<^sub>t_is_subterm_trms\<^sub>s\<^sub>s\<^sub>t[OF x], of _ x] unfolding a0_def a0'_def by fastforce then obtain y where y: "y \ fv_transaction T - set (transaction_fresh T)" "\ x = (\ \\<^sub>s \ \\<^sub>s \) y \ \" "\ x \\<^sub>\ a0 = (\ \\<^sub>s \ \\<^sub>s \) y \ \ \\<^sub>\ a0" "\ x \\<^sub>\ a0' = (\ \\<^sub>s \ \\<^sub>s \) y \ \ \\<^sub>\ a0'" "\\<^sub>v y = TAtom Value" by metis then obtain n where n: "(\ \\<^sub>s \ \\<^sub>s \) y \ \ = Fun (Val n) []" using \\<^sub>v_TAtom''(2)[of y] x x_val transaction_var_becomes_Val[ OF reachable_constraints.step[OF \_reach T \ \ \] \ \ \ \ P P_occ T, of y] by force have "a0 n \ a0' n" "y \ fv_transaction T" "y \ set (transaction_fresh T)" "absc (a0 n) = (\ \\<^sub>s \ \\<^sub>s \) y \ \ \\<^sub>\ a0" "absc (a0' n) = (\ \\<^sub>s \ \\<^sub>s \) y \ \ \\<^sub>\ a0'" using y n False by force+ hence 1: "(a0 n, a0' n) \ a3" unfolding a0_def a0'_def a3_def abs_term_implications_def by blast have 2: "\ x \\<^sub>\ a0' \ set \a0 n --\ a0' n\\\ x \\<^sub>\ a0\" using y n timpl_apply_const by auto show ?thesis using timpl_closure.TI[OF timpl_closure.FP 1] 2 term_variants_pred_iff_in_term_variants[ of "(\_. [])(Abs (a0 n) := [Abs (a0' n)])"] unfolding timpl_closure_set_def timpl_apply_term_def by auto qed (auto intro: timpl_closure_setI) next assume "\a. \\<^sub>v x = TAtom (Atom a)" then obtain a where x_atom: "\\<^sub>v x = TAtom (Atom a)" by moura obtain f T where fT: "\ x = Fun f T" using interpretation_grounds[OF \_interp, of "Var x"] by (cases "\ x") auto have fT_atom: "\ (Fun f T) = TAtom (Atom a)" using wt_subst_trm''[OF \_wt, of "Var x"] x_atom fT by simp have T: "T = []" using fT wf_trm_subst[OF \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s, of "Var x"] const_type_inv_wf[OF fT_atom] by fastforce have f: "\is_Val f" using fT_atom unfolding is_Val_def by auto have "\ x \\<^sub>\ b = \ x" for b using T fT abs_term_apply_const(2)[OF f] by auto thus "\ x \\<^sub>\ a0' \ timpl_closure_set {\ x \\<^sub>\ a0} a3" by (auto intro: timpl_closure_setI) qed thus ?thesis by (metis a0_def a0'_def a3_def) qed lemma \\<^sub>t\<^sub>i_covers_\\<^sub>0_Val: assumes \_reach: "\ \ reachable_constraints P" and T: "T \ set P" and \: "welltyped_constraint_model \ (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and n: "Fun (Val n) [] \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "Fun (Val n) [] \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \) \ timpl_closure_set {Fun (Val n) [] \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)} (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \)" proof - define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define a0 where "a0 \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" define a0' where "a0' \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \)" define a3 where "a3 \ \\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \" have \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (metis reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s admissible_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P(1) \_reach) note T_adm = bspec[OF P(1) T] have "Fun (Abs (a0' n)) [] \ timpl_closure_set {Fun (Abs (a0 n)) []} a3" proof (cases "a0 n = a0' n") case False then obtain x where x: "x \ fv_transaction T - set (transaction_fresh T)" "Fun (Val n) [] = (\ \\<^sub>s \ \\<^sub>s \) x \ \" using \\<^sub>t\<^sub>i_covers_\\<^sub>0_aux[OF \_reach T \ \ \ \ P P_occ n] by (fastforce simp add: a0_def a0'_def T'_def) hence "absc (a0 n) = (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0" "absc (a0' n) = (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0'" by simp_all hence 1: "(a0 n, a0' n) \ a3" using False x(1) unfolding a0_def a0'_def a3_def abs_term_implications_def T'_def by blast show ?thesis using timpl_apply_Abs[of "[]" "[]" "a0 n" "a0' n"] timpl_closure.TI[OF timpl_closure.FP[of "Fun (Abs (a0 n)) []" a3] 1] term_variants_pred_iff_in_term_variants[of "(\_. [])(Abs (a0 n) := [Abs (a0' n)])"] unfolding timpl_closure_set_def timpl_apply_term_def by force qed (auto intro: timpl_closure_setI) thus ?thesis by (simp add: a0_def a0'_def a3_def T'_def) qed lemma \\<^sub>t\<^sub>i_covers_\\<^sub>0_ik: assumes \_reach: "\ \ reachable_constraints P" and T: "T \ set P" and \: "welltyped_constraint_model \ (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and t: "t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" shows "t \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \) \ timpl_closure_set {t \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)} (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \)" proof - define a0 where "a0 \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" define a0' where "a0' \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \)" define a3 where "a3 \ \\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \" let ?U = "\T a. map (\s. s \ \ \\<^sub>\ a) T" have \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (metis reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s admissible_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P(1) \_reach) note T_adm = bspec[OF P(1) T] have "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "wf\<^sub>t\<^sub>r\<^sub>m t" using \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s t ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset by force+ hence "\t0 \ subterms t. t0 \ \ \\<^sub>\ a0' \ timpl_closure_set {t0 \ \ \\<^sub>\ a0} a3" proof (induction t) case (Var x) thus ?case using \\<^sub>t\<^sub>i_covers_\\<^sub>0_Var[OF \_reach T \ \ \ \ P P_occ, of x] ik\<^sub>s\<^sub>s\<^sub>t_var_is_fv[of x "unlabel \"] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] by (simp add: a0_def a0'_def a3_def) next case (Fun f S) have IH: "\t0 \ subterms t. t0 \ \ \\<^sub>\ a0' \ timpl_closure_set {t0 \ \ \\<^sub>\ a0} a3" when "t \ set S" for t using that Fun.prems(1) wf_trm_param[OF Fun.prems(2)] Fun.IH by (meson in_subterms_subset_Union params_subterms subsetCE) hence "t \\<^sub>\ a0' \ timpl_closure_set {t \\<^sub>\ a0} a3" when "t \ set (map (\s. s \ \) S)" for t using that by auto hence "t \\<^sub>\ a0' \ timpl_closure (t \\<^sub>\ a0) a3" when "t \ set (map (\s. s \ \) S)" for t using that timpl_closureton_is_timpl_closure by auto hence "(t \\<^sub>\ a0, t \\<^sub>\ a0') \ timpl_closure' a3" when "t \ set (map (\s. s \ \) S)" for t using that timpl_closure_is_timpl_closure' by auto hence IH': "((?U S a0) ! i, (?U S a0') ! i) \ timpl_closure' a3" when "i < length (map (\s. s \ \ \\<^sub>\ a0) S)" for i using that by auto show ?case proof (cases "\n. f = Val n") case True then obtain n where "Fun f S = Fun (Val n) []" using Fun.prems(2) unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by force moreover have "Fun f S \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset Fun.prems(1) by blast ultimately show ?thesis using \\<^sub>t\<^sub>i_covers_\\<^sub>0_Val[OF \_reach T \ \ \ \ P P_occ] by (simp add: a0_def a0'_def a3_def) next case False hence "Fun f S \ \ \\<^sub>\ a = Fun f (map (\t. t \ \ \\<^sub>\ a) S)" for a by (cases f) simp_all hence "(Fun f S \ \ \\<^sub>\ a0, Fun f S \ \ \\<^sub>\ a0') \ timpl_closure' a3" using timpl_closure_FunI[OF IH'] by simp hence "Fun f S \ \ \\<^sub>\ a0' \ timpl_closure_set {Fun f S \ \ \\<^sub>\ a0} a3" using timpl_closureton_is_timpl_closure timpl_closure_is_timpl_closure' by metis thus ?thesis using IH by simp qed qed thus ?thesis by (simp add: a0_def a0'_def a3_def) qed lemma transaction_prop1: assumes "\ \ abs_substs_fun ` set (transaction_check_comp msgcs (FP, OCC, TI) T)" and "x \ fv_transaction T" and "x \ set (transaction_fresh T)" and "\ x \ absdbupd (unlabel (transaction_updates T)) x (\ x)" and "transaction_check' msgcs (FP, OCC, TI) T" and TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" shows "(\ x, absdbupd (unlabel (transaction_updates T)) x (\ x)) \ (set TI)\<^sup>+" proof - let ?upd = "\x. absdbupd (unlabel (transaction_updates T)) x (\ x)" have 0: "fv_transaction T = set (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T)))" by (metis fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"]) have 1: "transaction_check_post (FP, OCC, TI) T \" using assms(1,5) unfolding transaction_check_def transaction_check'_def list_all_iff by blast have "(\ x, ?upd x) \ set TI \ (\ x, ?upd x) \ (set TI)\<^sup>+" using TI using assms(4) by blast thus ?thesis using assms(2,3,4) 0 1 in_trancl_closure_iff_in_trancl_fun[of _ _ TI] unfolding transaction_check_post_def List.member_def Let_def by blast qed lemma transaction_prop2: assumes \: "\ \ abs_substs_fun ` set (transaction_check_comp msgcs (FP, OCC, TI) T)" and x: "x \ fv_transaction T" "fst x = TAtom Value" and T_check: "transaction_check' msgcs (FP, OCC, TI) T" and T_adm: "admissible_transaction' T" and T_occ: "admissible_transaction_occurs_checks T" and FP: "analyzed (timpl_closure_set (set FP) (set TI))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" and OCC: "\t \ timpl_closure_set (set FP) (set TI). \f \ funs_term t. is_Abs f \ f \ Abs ` set OCC" "timpl_closure_set (absc ` set OCC) (set TI) \ absc ` set OCC" and TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" shows "x \ set (transaction_fresh T) \ \ x \ set OCC" (is "?A' \ ?A") and "absdbupd (unlabel (transaction_updates T)) x (\ x) \ set OCC" (is ?B) proof - let ?xs = "fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T))" let ?ys = "filter (\x. x \ set (transaction_fresh T) \ fst x = TAtom Value) ?xs" let ?C = "unlabel (transaction_checks T)" let ?poss = "transaction_poschecks_comp ?C" let ?negs = "transaction_negchecks_comp ?C" let ?\upd = "\y. absdbupd (unlabel (transaction_updates T)) y (\ y)" note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] have 0: "{x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value} = set ?ys" using fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] by force have 1: "transaction_check_pre (FP, OCC, TI) T \" using \ unfolding transaction_check_comp_def Let_def by fastforce have 2: "transaction_check_post (FP, OCC, TI) T \" using \ T_check unfolding transaction_check_def transaction_check'_def list_all_iff by auto have 3: "\ \ abs_substs_fun ` set (abs_substs_set ?ys OCC ?poss ?negs msgcs)" using \ unfolding transaction_check_comp_def Let_def by force show A: ?A when ?A' using that 0 3 x abs_substs_abss_bounded by blast have 4: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" when x': "x \ set (transaction_fresh T)" using x' admissible_transactionE(7)[OF T_adm] by blast have "intruder_synth_mod_timpls FP TI (occurs (absc (?\upd x)))" when x': "x \ set (transaction_fresh T)" proof - obtain l ts S where ts: "transaction_send T = (l,send\ts\)#S" "occurs (Var x) \ set ts" using admissible_transaction_occurs_checksE2[OF T_occ x'] by blast hence "occurs (Var x) \ set ts" "send\ts\ \ set (unlabel (transaction_send T))" using x' unfolding suffix_def by (fastforce, fastforce) thus ?thesis using 2 x unfolding transaction_check_post_def by fastforce qed hence "timpl_closure_set (set FP) (set TI) \\<^sub>c occurs (absc (?\upd x))" when x': "x \ set (transaction_fresh T)" using x' intruder_synth_mod_timpls_is_synth_timpl_closure_set[ OF TI, of FP "occurs (absc (?\upd x))"] by argo hence "Abs (?\upd x) \ \(funs_term ` timpl_closure_set (set FP) (set TI))" when x': "x \ set (transaction_fresh T)" using x' ideduct_synth_priv_fun_in_ik[ of "timpl_closure_set (set FP) (set TI)" "occurs (absc (?\upd x))"] by simp hence "\t \ timpl_closure_set (set FP) (set TI). Abs (?\upd x) \ funs_term t" when x': "x \ set (transaction_fresh T)" using x' by force hence 5: "?\upd x \ set OCC" when x': "x \ set (transaction_fresh T)" using x' OCC by fastforce have 6: "?\upd x \ set OCC" when x': "x \ set (transaction_fresh T)" proof (cases "\ x = ?\upd x") case False hence "(\ x, ?\upd x) \ (set TI)\<^sup>+" "\ x \ set OCC" using A 2 x' x TI unfolding transaction_check_post_def fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t Let_def in_trancl_closure_iff_in_trancl_fun[symmetric] List.member_def by blast+ thus ?thesis using timpl_closure_set_absc_subset_in[OF OCC(2)] by blast qed (simp add: A x' x(1)) show ?B by (metis 5 6) qed lemma transaction_prop3: assumes \_reach: "\ \ reachable_constraints P" and T: "T \ set P" and \: "welltyped_constraint_model \ (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and FP: "analyzed (timpl_closure_set (set FP) (set TI))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" "\t \ \\<^sub>i\<^sub>k \ \. timpl_closure_set (set FP) (set TI) \\<^sub>c t" and OCC: "\t \ timpl_closure_set (set FP) (set TI). \f \ funs_term t. is_Abs f \ f \ Abs ` set OCC" "timpl_closure_set (absc ` set OCC) (set TI) \ absc ` set OCC" "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \ \ absc ` set OCC" and TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" shows "\x \ set (transaction_fresh T). (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) = absc {}" (is ?A) and "\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T). intruder_synth_mod_timpls FP TI (t \ (\ \\<^sub>s \ \\<^sub>s \) \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \))" (is ?B) and "\x \ fv_transaction T - set (transaction_fresh T). \s. select\Var x,Fun (Set s) []\ \ set (unlabel (transaction_checks T)) \ (\ss. (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) = absc ss \ s \ ss)" (is ?C) and "\x \ fv_transaction T - set (transaction_fresh T). \s. \Var x in Fun (Set s) []\ \ set (unlabel (transaction_checks T)) \ (\ss. (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) = absc ss \ s \ ss)" (is ?D) and "\x \ fv_transaction T - set (transaction_fresh T). \s. \Var x not in Fun (Set s) []\ \ set (unlabel (transaction_checks T)) \ (\ss. (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) = absc ss \ s \ ss)" (is ?E) and "\x \ fv_transaction T - set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) \ absc ` set OCC" (is ?F) proof - let ?T' = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" let ?\ = "\ \\<^sub>s \ \\<^sub>s \" define a0 where "a0 \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" define a0' where "a0' \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@?T') \)" define fv_AT' where "fv_AT' \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@?T')" note T_adm = bspec[OF P(1) T] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] note T_adm' = admissible_transaction_is_wellformed_transaction(2-4)[OF T_adm] note \_empty = admissible_transaction_decl_subst_empty[OF T_adm \] hence \_elim: "?\ = \ \\<^sub>s \" by simp have \': "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "\n. PubConst Value n \ \(funs_term ` (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "\n. Abs n \ \(funs_term ` (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "\n. PubConst Value n \ \(funs_term ` (\ ` fv_AT'))" "\n. Abs n \ \(funs_term ` (\ ` fv_AT'))" using \ admissible_transaction_occurs_checks_prop'[ OF \_reach welltyped_constraint_model_prefix[OF \] P P_occ] admissible_transaction_occurs_checks_prop'[ OF reachable_constraints.step[OF \_reach T \ \ \] \ P P_occ] unfolding welltyped_constraint_model_def constraint_model_def is_Val_def is_Abs_def fv_AT'_def by (meson,meson,meson,metis,metis,metis,metis) have T_no_pubconsts: "\n. PubConst Value n \ \(funs_term ` trms_transaction T)" and T_no_abss: "\n. Abs n \ \(funs_term ` trms_transaction T)" and T_fresh_vars_value_typed: "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" and T_fv_const_typed: "\x \ fv_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" using protocol_transaction_vars_TAtom_typed protocol_transactions_no_pubconsts protocol_transactions_no_abss funs_term_Fun_subterm P T by (fast, fast, fast, fast) have wt_\\\: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \ \\<^sub>s \)" using \'(2) wt_subst_compose transaction_fresh_subst_wt[OF \] transaction_renaming_subst_wt[OF \] by blast have 1: "?\ y \ \ = \ y" when "y \ set (transaction_fresh T)" for y using transaction_fresh_subst_grounds_domain[OF \ that] subst_compose[of \ \ y] unfolding \_elim by (simp add: subst_ground_ident) have 2: "?\ y \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" when "y \ set (transaction_fresh T)" for y using 1[OF that] that \ unfolding transaction_fresh_subst_def by auto have 3: "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \\<^sub>v x = TAtom Value \ (\B. prefix B \ \ x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \ \ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B))" by (metis welltyped_constraint_model_prefix[OF \] constraint_model_Value_var_in_constr_prefix[OF \_reach _ P P_occ]) have 4: "\n. ?\ y \ \ = Fun (Val n) []" when "y \ fv_transaction T" "\\<^sub>v y = TAtom Value" for y using transaction_var_becomes_Val[ OF reachable_constraints.step[OF \_reach T \ \ \] \ \ \ \ P P_occ T] that T_fv_const_typed \\<^sub>v_TAtom''[of y] by metis have \_is_T_model: "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) (set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)) (unlabel ?T') \" using \ unlabel_append[of \ ?T'] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of "unlabel \" \ "[]"] strand_sem_append_stateful[of "{}" "{}" "unlabel \" "unlabel ?T'" \] by (simp add: welltyped_constraint_model_def constraint_model_def db\<^sub>s\<^sub>s\<^sub>t_def) have T_rcv_no_val_bvars: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ subst_domain ?\ = {}" using admissible_transaction_no_bvars[OF T_adm] bvars_transaction_unfold[of T] by blast show ?A proof fix y assume y: "y \ set (transaction_fresh T)" then obtain yn where yn: "(\ \\<^sub>s \ \\<^sub>s \) y \ \ = Fun (Val yn) []" "Fun (Val yn) [] \ subst_range \" by (metis \_elim T_fresh_vars_value_typed transaction_fresh_subst_sends_to_val'[OF \]) { \ \since \y\ is fresh \(\ \\<^sub>s \ \\<^sub>s \) y \ \\ cannot be part of the database state of \\ \\\ fix t' s assume t': "insert\t',s\ \ set (unlabel \)" "t' \ \ = Fun (Val yn) []" then obtain z where t'_z: "t' = Var z" using 2[OF y] yn(1) by (cases t') auto hence z: "z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "\ z = (\ \\<^sub>s \ \\<^sub>s \) y \ \" using t' yn(1) by force+ hence z': "\\<^sub>v z = TAtom Value" by (metis \.simps(1) \_consts_simps(2) t'(2) t'_z wt_subst_trm'' \'(2)) obtain B where B: "prefix B \" "\ z \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" using z z' 3 by fastforce hence "\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" using transaction_fresh_subst_range_fresh(1)[OF \] trms\<^sub>s\<^sub>s\<^sub>t_unlabel_prefix_subset(1)[of B] unfolding prefix_def by fast hence False using B(2) 1[OF y] z yn(1) by (metis subst_imgI term.distinct(1)) } hence "\s. (?\ y \ \, s) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" using db\<^sub>s\<^sub>s\<^sub>t_in_cases[of "?\ y \ \" _ "unlabel \" \ "[]"] yn(1) by (force simp add: db\<^sub>s\<^sub>s\<^sub>t_def) thus "?\ y \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) = absc {}" using to_abs_empty_iff_notin_db[of yn "db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ []"] yn(1) by (simp add: db\<^sub>s\<^sub>s\<^sub>t_def) qed show receives_covered: ?B proof fix t assume t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" hence t_in_T: "t \ trms_transaction T" using trms\<^sub>s\<^sub>s\<^sub>t_unlabel_prefix_subset(1)[of "transaction_receive T"] unfolding transaction_strand_def by fast obtain ts where ts: "t \ set ts" "receive\ts\ \ set (unlabel (transaction_receive T))" using t wellformed_transaction_send_receive_trm_cases(1)[OF T_wf] by blast have t_rcv: "receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t ?\\ \ set (unlabel (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?\))" using subst_lsst_unlabel_member[of "receive\ts\" "transaction_receive T" ?\] trms\<^sub>s\<^sub>s\<^sub>t_in[OF t] ts by fastforce have "list_all (\t. ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ t \ ?\ \ \) ts" using wellformed_transaction_sem_receives[OF T_wf \_is_T_model t_rcv] unfolding list_all_iff by fastforce hence *: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ t \ ?\ \ \" using ts(1) unfolding list_all_iff by fast have t_fv: "fv (t \ ?\) \ fv_AT'" using fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \"] unlabel_append[of \] fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?\"] ts(1) t_rcv fv_transaction_subst_unfold[of T ?\] unfolding fv_AT'_def by force have **: "\t \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0. timpl_closure_set (set FP) (set TI) \\<^sub>c t" using FP(3) by (auto simp add: a0_def abs_intruder_knowledge_def) note lms1 = pubval_terms_subst[OF _ pubval_terms_subst_range_disj[ OF transaction_fresh_subst_has_no_pubconsts_abss(1)[OF \], of t]] pubval_terms_subst[OF _ pubval_terms_subst_range_disj[ OF transaction_renaming_subst_has_no_pubconsts_abss(1)[OF \], of "t \ \"]] note lms2 = abs_terms_subst[OF _ abs_terms_subst_range_disj[ OF transaction_fresh_subst_has_no_pubconsts_abss(2)[OF \], of t]] abs_terms_subst[OF _ abs_terms_subst_range_disj[ OF transaction_renaming_subst_has_no_pubconsts_abss(2)[OF \], of "t \ \"]] have "t \ (\T\set P. trms_transaction T)" "fv (t \ \ \\<^sub>s \ \ \) = {}" using t_in_T T interpretation_grounds[OF \'(1)] by fast+ moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \ \\<^sub>s \))" using wf_trm_subst_rangeI[of \, OF transaction_fresh_subst_is_wf_trm[OF \]] wf_trm_subst_rangeI[of \, OF transaction_renaming_subst_is_wf_trm[OF \]] wf_trms_subst_compose[of \ \, THEN wf_trms_subst_compose[OF _ \'(3)]] by blast moreover have "t \ pubval_terms" using t_in_T T_no_pubconsts funs_term_Fun_subterm unfolding is_PubConstValue_def is_PubConst_def by fastforce hence "t \ ?\ \ pubval_terms" using lms1 T_fresh_vars_value_typed unfolding \_elim by auto hence "t \ ?\ \ \ \ pubval_terms" using \'(6) t_fv pubval_terms_subst'[of "t \ ?\" \] by auto moreover have "t \ abs_terms" using t_in_T T_no_abss funs_term_Fun_subterm unfolding is_Abs_def by force hence "t \ ?\ \ abs_terms" using lms2 T_fresh_vars_value_typed unfolding \_elim by auto hence "t \ ?\ \ \ \ abs_terms" using \'(7) t_fv abs_terms_subst'[of "t \ ?\" \] by auto ultimately have ***: "t \ \ \\<^sub>s \ \\<^sub>s \ \ \ \ GSMP (\T\set P. trms_transaction T) - (pubval_terms \ abs_terms)" using SMP.Substitution[OF SMP.MP[of t "\T\set P. trms_transaction T"], of "\ \\<^sub>s \ \\<^sub>s \"] subst_subst_compose[of t ?\ \] wt_\\\ unfolding GSMP_def \_elim by fastforce have ****: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ GSMP (\T\set P. trms_transaction T) - (pubval_terms \ abs_terms)" using reachable_constraints_GSMP_no_pubvals_abss[OF \_reach P \'(1-5)] ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset[of "unlabel \"] by blast show "intruder_synth_mod_timpls FP TI (t \ ?\ \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \))" using deduct_FP_if_deduct[OF **** ** * ***] deducts_eq_if_analyzed[OF FP(1)] intruder_synth_mod_timpls_is_synth_timpl_closure_set[OF TI, of FP] unfolding a0_def by force qed show ?C proof (intro ballI allI impI) fix y s assume y: "y \ fv_transaction T - set (transaction_fresh T)" and s: "select\Var y, Fun (Set s) []\ \ set (unlabel (transaction_checks T))" hence "select\Var y, Fun (Set s) []\ \ set (unlabel (transaction_strand T))" unfolding transaction_strand_def unlabel_def by auto hence y_val: "\\<^sub>v y = TAtom Value" using transaction_selects_are_Value_vars[OF T_wf T_adm'(1)] by fastforce have "select\?\ y, Fun (Set s) []\ \ set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?\))" using subst_lsst_unlabel_member[OF s] by fastforce hence "((\ \\<^sub>s \ \\<^sub>s \) y \ \, Fun (Set s) []) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" using wellformed_transaction_sem_pos_checks[ OF T_wf \_is_T_model, of Assign "(\ \\<^sub>s \ \\<^sub>s \) y" "Fun (Set s) []"] by simp thus "\ss. (\ \\<^sub>s \ \\<^sub>s \) y \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) = absc ss \ s \ ss" using to_abs_alt_def[of "db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \"] 4[of y] y y_val by auto qed show ?D proof (intro ballI allI impI) fix y s assume y: "y \ fv_transaction T - set (transaction_fresh T)" and s: "\Var y in Fun (Set s) []\ \ set (unlabel (transaction_checks T))" hence "\Var y in Fun (Set s) []\ \ set (unlabel (transaction_strand T))" unfolding transaction_strand_def unlabel_def by auto hence y_val: "\\<^sub>v y = TAtom Value" using transaction_inset_checks_are_Value_vars[OF T_adm] by fastforce have "\?\ y in Fun (Set s) []\ \ set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?\))" using subst_lsst_unlabel_member[OF s] by fastforce hence "(?\ y \ \, Fun (Set s) []) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" using wellformed_transaction_sem_pos_checks[ OF T_wf \_is_T_model, of Check "?\ y" "Fun (Set s) []"] by simp thus "\ss. ?\ y \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) = absc ss \ s \ ss" using to_abs_alt_def[of "db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \"] 4[of y] y y_val by auto qed show ?E proof (intro ballI allI impI) fix y s assume y: "y \ fv_transaction T - set (transaction_fresh T)" and s: "\Var y not in Fun (Set s) []\ \ set (unlabel (transaction_checks T))" hence "\Var y not in Fun (Set s) []\ \ set (unlabel (transaction_strand T))" unfolding transaction_strand_def unlabel_def by auto hence y_val: "\\<^sub>v y = TAtom Value" using transaction_notinset_checks_are_Value_vars(1)[ OF T_adm, of "[]" "[]" "[(Var y, Fun (Set s) [])]" "Var y" "Fun (Set s) []"] by force have "\?\ y not in Fun (Set s) []\ \ set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?\))" using subst_lsst_unlabel_member[OF s] by fastforce hence "(?\ y \ \, Fun (Set s) []) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" using wellformed_transaction_sem_neg_checks'(2)[ OF T_wf \_is_T_model, of "[]" "?\ y" "Fun (Set s) []"] by simp moreover have "list_all admissible_transaction_updates P" using Ball_set[of P "admissible_transaction"] P(1) Ball_set[of P admissible_transaction_updates] admissible_transaction_is_wellformed_transaction(3) by fast moreover have "list_all wellformed_transaction P" using P(1) Ball_set[of P "admissible_transaction"] Ball_set[of P wellformed_transaction] admissible_transaction_is_wellformed_transaction(1) by blast ultimately have "((\ \\<^sub>s \ \\<^sub>s \) y \ \, Fun (Set s) S) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" for S using reachable_constraints_db\<^sub>l\<^sub>s\<^sub>s\<^sub>t_set_args_empty[OF \_reach] unfolding admissible_transaction_updates_def by auto thus "\ss. (\ \\<^sub>s \ \\<^sub>s \) y \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) = absc ss \ s \ ss" using to_abs_alt_def[of "db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \"] 4[of y] y y_val by auto qed show ?F proof (intro ballI impI) fix y assume y: "y \ fv_transaction T - set (transaction_fresh T)" "\\<^sub>v y = TAtom Value" then obtain yn where yn: "?\ y \ \ = Fun (Val yn) []" using 4 by moura hence y_abs: "?\ y \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) = Fun (Abs (\\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) yn)) []" by simp have "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ (y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ (\t s. select\t,s\ \ set (unlabel (transaction_checks T)) \ y \ fv t \ fv s))" using admissible_transaction_fv_in_receives_or_selects[OF T_adm] y by blast thus "?\ y \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) \ absc ` set OCC" proof assume "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" then obtain ts where ts: "receive\ts\ \ set (unlabel (transaction_receive T))" "y \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" using wellformed_transaction_unlabel_cases(1)[OF T_wf] by (force simp add: unlabel_def) have *: "?\ y \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t ?\ \\<^sub>s \)" "list_all (\t. timpl_closure_set (set FP) (set TI) \\<^sub>c t \ ?\ \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)) ts" using ts fv_subterms_substI[of y _ "?\ \\<^sub>s \"] subst_compose[of ?\ \ y] subterms_subst_subset[of "?\ \\<^sub>s \"] receives_covered unfolding intruder_synth_mod_timpls_is_synth_timpl_closure_set[OF TI, symmetric] list_all_iff by fastforce+ have "Abs (\\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) yn) \ \(funs_term ` (timpl_closure_set (set FP) (set TI)))" using * y_abs abs_subterms_in[of "?\ y \ \" _ "\\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)"] ideduct_synth_priv_fun_in_ik[ OF _ funs_term_Fun_subterm'[of "Abs (\\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) yn)" "[]"]] unfolding list_all_iff by fastforce hence "?\ y \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) \ subterms\<^sub>s\<^sub>e\<^sub>t (timpl_closure_set (set FP) (set TI))" using y_abs wf_trms_subterms[OF timpl_closure_set_wf_trms[OF FP(2), of "set TI"]] funs_term_Fun_subterm[of "Abs (\\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \) yn)"] unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by fastforce hence "funs_term (?\ y \ \ \\<^sub>\ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)) \ (\t \ timpl_closure_set (set FP) (set TI). funs_term t)" using funs_term_subterms_eq(2)[of "timpl_closure_set (set FP) (set TI)"] by blast thus ?thesis using y_abs OCC(1) by fastforce next assume "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ (\t s. select\t,s\ \ set (unlabel (transaction_checks T)) \ y \ fv t \ fv s)" then obtain t t' where "select\t,t'\ \ set (unlabel (transaction_checks T))" "y \ fv t \ fv t'" by blast then obtain l s where "(l,select\Var y, Fun (Set s) []\) \ set (transaction_checks T)" using admissible_transaction_strand_step_cases(2)[OF T_adm] unfolding unlabel_def by fastforce then obtain U where U: "prefix (U@[(l,select\Var y, Fun (Set s) []\)]) (transaction_checks T)" using in_set_conv_decomp[of "(l, select\Var y,Fun (Set s) []\)" "transaction_checks T"] by (auto simp add: prefix_def) hence "select\Var y, Fun (Set s) []\ \ set (unlabel (transaction_checks T))" by (force simp add: prefix_def unlabel_def) hence "select\?\ y, Fun (Set s) []\ \ set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?\))" using subst_lsst_unlabel_member by fastforce hence "(Fun (Val yn) [], Fun (Set s) []) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" using yn wellformed_transaction_sem_pos_checks[ OF T_wf \_is_T_model, of Assign "?\ y" "Fun (Set s) []"] by fastforce hence "Fun (Val yn) [] \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \" using db\<^sub>s\<^sub>s\<^sub>t_in_cases[of "Fun (Val yn) []"] by (fastforce simp add: db\<^sub>s\<^sub>s\<^sub>t_def) thus ?thesis using OCC(3) yn abs_in[of "Fun (Val yn) []" _ "\\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)"] unfolding abs_value_constants_def by (metis (mono_tags, lifting) mem_Collect_eq subsetCE) qed qed qed lemma transaction_prop4: assumes \_reach: "\ \ reachable_constraints P" and T: "T \ set P" and \: "welltyped_constraint_model \ (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and x: "x \ set (transaction_fresh T)" and y: "y \ fv_transaction T - set (transaction_fresh T)" "\\<^sub>v y = TAtom Value" shows "(\ \\<^sub>s \ \\<^sub>s \) x \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" (is ?A) and "(\ \\<^sub>s \ \\<^sub>s \) y \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" (is ?B) proof - let ?T' = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" from \ have \': "welltyped_constraint_model \ \" using welltyped_constraint_model_prefix by auto note T_P_adm = bspec[OF P(1)] note T_adm = bspec[OF P(1) T] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] have be: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" using T_P_adm \_reach reachable_constraints_no_bvars admissible_transaction_no_bvars(2) by blast have T_no_bvars: "fv_transaction T = vars_transaction T" using admissible_transaction_no_bvars[OF T_adm] by simp note \_empty = admissible_transaction_decl_subst_empty[OF T_adm \] have \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" by (metis \ welltyped_constraint_model_def) have x_val: "fst x = TAtom Value" using x admissible_transactionE(14)[OF T_adm] unfolding list_all_iff by blast obtain xn where xn: "\ x = Fun (Val xn) []" using x_val transaction_fresh_subst_sends_to_val[OF \ x] \\<^sub>v_TAtom''(2)[of x] by meson hence xnxn: "(\ \\<^sub>s \ \\<^sub>s \) x = Fun (Val xn) []" unfolding subst_compose_def \_empty by auto from xn xnxn have a0: "(\ \\<^sub>s \ \\<^sub>s \) x \ \ = Fun (Val xn) []" by auto have b0: "\\<^sub>v x = TAtom Value" using P x T protocol_transaction_vars_TAtom_typed(3) by metis note 0 = a0 b0 have \_x_nin_A: "\ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" proof - have "\ x \ subst_range \" by (metis transaction_fresh_subst_sends_to_val[OF \ x b0]) moreover have "\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using \ transaction_fresh_subst_def[of \ T "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] by blast ultimately show ?thesis by auto qed have *: "y \ set (transaction_fresh T)" using assms by auto have **: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ (y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ (\t s. select\t,s\ \ set (unlabel (transaction_checks T)) \ y \ fv t \ fv s))" using * y admissible_transaction_fv_in_receives_or_selects[OF T_adm] by blast have y_fv: "y \ fv_transaction T" using y fv_transaction_unfold by blast have y_val: "fst y = TAtom Value" using y(2) \\<^sub>v_TAtom''(2) by blast have "\ x \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" proof (rule ccontr) assume "\\ x \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" then have a: "\ x \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" by auto then have \_x_I_in_A: "\ x \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \" using reachable_constraints_subterms_subst[OF \_reach \' P] by blast have "\u. u \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ \ u = \ x" proof - from \_x_I_in_A have "\tu. tu \ \ (subterms ` (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ tu \ \ = \ x \ \" by force then obtain tu where tu: "tu \ \ (subterms ` (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ tu \ \ = \ x \ \" by auto then have "tu \ \ x" using \_x_nin_A by auto moreover have "tu \ \ = \ x" using tu by (simp add: xn) ultimately have "\u. tu = Var u" unfolding xn by (cases tu) auto then obtain u where "tu = Var u" by auto have "u \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ \ u = \ x" proof - have "u \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using \tu = Var u\ tu var_subterm_trms\<^sub>s\<^sub>s\<^sub>t_is_vars\<^sub>s\<^sub>s\<^sub>t by fastforce then have "u \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using be vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] by blast moreover have "\ u = \ x" using \tu = Var u\ \tu \ \ = \ x\ by auto ultimately show ?thesis by auto qed then show "\u. u \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ \ u = \ x" by metis qed then obtain u where u: "u \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "\ u = \ x" by auto then have u_TA: "\\<^sub>v u = TAtom Value" using P(1) T x_val \\<^sub>v_TAtom''(2)[of x] wt_subst_trm''[OF \_wt, of "Var u"] wt_subst_trm''[of \ "Var x"] transaction_fresh_subst_wt[OF \] protocol_transaction_vars_TAtom_typed(3) by force have "\B. prefix B \ \ u \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \ \ u \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" using u u_TA by (metis welltyped_constraint_model_prefix[OF \] constraint_model_Value_var_in_constr_prefix[OF \_reach _ P P_occ]) then obtain B where "prefix B \ \ u \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \ \ u \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" by blast moreover have "\(subterms ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t xs) \ \(subterms ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ys)" when "prefix xs ys" for xs ys::"('fun,'atom,'sets,'lbl) prot_strand" using that subterms\<^sub>s\<^sub>e\<^sub>t_mono trms\<^sub>s\<^sub>s\<^sub>t_mono unlabel_mono set_mono_prefix by metis ultimately have "\ u \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by blast then have "\ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using u by auto then show "False" using \_x_nin_A by auto qed then show ?A - using subst_apply_term.simps(1)[of x \] + using eval_term.simps(1)[of _ x \] unfolding subst_compose_def xn \_empty by auto from ** show ?B proof define T' where "T' \ transaction_receive T" define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" assume y: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" hence "Var y \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" by (metis T'_def fv\<^sub>s\<^sub>s\<^sub>t_is_subterm_trms\<^sub>s\<^sub>s\<^sub>t) then obtain z where z: "z \ set (unlabel T')" "Var y \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p z)" by (induct T') auto have "is_Receive z" using Ball_set[of "unlabel T'" is_Receive] z(1) admissible_transaction_is_wellformed_transaction(1)[OF T_adm] unfolding wellformed_transaction_def T'_def by blast then obtain tys where "z = receive\tys\" by (cases z) auto hence tys: "receive\tys \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\ \ set (unlabel (T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "\ y \ subterms\<^sub>s\<^sub>e\<^sub>t (set tys \\<^sub>s\<^sub>e\<^sub>t \)" using z subst_mono unfolding subst_apply_labeled_stateful_strand_def unlabel_def by force+ hence y_deduct: "list_all (\t. ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \ \ \) tys" using transaction_receive_deduct[OF T_wf _ \ \ \] \ unfolding T'_def \_def welltyped_constraint_model_def list_all_iff by auto obtain ty where ty: "ty \ set tys" "\ y \ ty \ \" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ ty \ \ \ \" using tys y_deduct unfolding list_all_iff by blast obtain zn where zn: "(\ \\<^sub>s \ \\<^sub>s \) y \ \ = Fun (Val zn) []" using transaction_var_becomes_Val[ OF reachable_constraints.step[OF \_reach T \ \ \] \ \ \ \ P P_occ T y_fv y_val] by metis have "(\ \\<^sub>s \ \\<^sub>s \) y \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \)" using ty tys(2) y_deduct private_fun_deduct_in_ik[of _ _ "Val zn"] by (metis \_def zn subst_mono public.simps(3)) thus ?B using ik\<^sub>s\<^sub>s\<^sub>t_subst[of "unlabel \" \] unlabel_subst[of \ \] subterms\<^sub>s\<^sub>e\<^sub>t_mono[OF ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset[of "unlabel (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)"]] by fastforce next assume y': "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ (\t s. select\t,s\ \ set (unlabel (transaction_checks T)) \ y \ fv t \ fv s)" then obtain s where s: "select\Var y,s\ \ set (unlabel (transaction_checks T))" "fst y = TAtom Value" using admissible_transaction_strand_step_cases(1,2)[OF T_adm] by fastforce obtain z zn where zn: "(\ \\<^sub>s \ \\<^sub>s \) y = Var z" "\ z = Fun (Val zn) []" using transaction_var_becomes_Val[ OF reachable_constraints.step[OF \_reach T \ \ \] \ \ \ \ P P_occ T y_fv s(2)] transaction_decl_fresh_renaming_substs_range(4)[OF \ \ \ _ *] transaction_decl_subst_empty_inv[OF \[unfolded \_empty]] by auto have transaction_selects_db_here: "\n s. select\Var (TAtom Value, n), Fun (Set s) []\ \ set (unlabel (transaction_checks T)) \ (\ (TAtom Value, n) \ \, Fun (Set s) []) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" using transaction_selects_db[OF T_adm _ \ \ \] \ unfolding welltyped_constraint_model_def by auto have "\n. y = (Var Value, n)" using T \\<^sub>v_TAtom_inv(2) y_fv y(2) by blast moreover have "admissible_transaction_checks T" using admissible_transaction_is_wellformed_transaction(2)[OF T_adm] by blast then have "is_Fun_Set (the_set_term (select\Var y,s\))" using s unfolding admissible_transaction_checks_def by auto then have "\ss. s = Fun (Set ss) []" using is_Fun_Set_exi by auto ultimately obtain n ss where nss: "y = (TAtom Value, n)" "s = Fun (Set ss) []" by auto then have "select\Var (TAtom Value, n), Fun (Set ss) []\ \ set (unlabel (transaction_checks T))" using s by auto then have in_db: "(\ (TAtom Value, n) \ \, Fun (Set ss) []) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" using transaction_selects_db_here[of n ss] by auto have "(\ z, s) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" proof - have "(\ y \ \, s) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" using in_db nss by auto moreover have "\ y = Var z" using zn \_empty * \ by (metis (no_types, opaque_lifting) subst_compose_def subst_imgI subst_to_var_is_var term.distinct(1) transaction_fresh_subst_def var_comp(2)) then have "\ y \ \ = \ z" by auto ultimately show "(\ z, s) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" by auto qed then have "\t' s'. insert\t',s'\ \ set (unlabel \) \ \ z = t' \ \ \ s = s' \ \" using db\<^sub>s\<^sub>s\<^sub>t_in_cases[of "\ z" s "unlabel \" \ "[]"] unfolding db\<^sub>s\<^sub>s\<^sub>t_def by auto then obtain t' s' where t's': "insert\t',s'\ \ set (unlabel \) \ \ z = t' \ \ \ s = s' \ \" by auto then have "t' \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by force then have "t' \ \ \ (subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \\<^sub>s\<^sub>e\<^sub>t \" by auto then have "\ z \ (subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \\<^sub>s\<^sub>e\<^sub>t \" using t's' by auto then have "\ z \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using reachable_constraints_subterms_subst[ OF \_reach welltyped_constraint_model_prefix[OF \] P] by auto then show ?B using zn(1) by simp qed qed lemma transaction_prop5: fixes T \ \ \ \ \ T' a0 a0' \ defines "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" and "a0 \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" and "a0' \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \)" and "\ \ \\ x. if fst x = TAtom Value then (absc \ \) x else Var x" assumes \_reach: "\ \ reachable_constraints P" and T: "T \ set P" and \: "welltyped_constraint_model \ (\@T')" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and FP: "analyzed (timpl_closure_set (set FP) (set TI))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" "\t \ \\<^sub>i\<^sub>k \ \. timpl_closure_set (set FP) (set TI) \\<^sub>c t" and OCC: "\t \ timpl_closure_set (set FP) (set TI). \f \ funs_term t. is_Abs f \ f \ Abs ` set OCC" "timpl_closure_set (absc ` set OCC) (set TI) \ absc ` set OCC" "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \ \ absc ` set OCC" and TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and step: "list_all (transaction_check (FP, OCC, TI)) P" shows "\\ \ abs_substs_fun ` set (transaction_check_comp (\_ _. True) (FP, OCC, TI) T). \x \ fv_transaction T. \\<^sub>v x = TAtom Value \ (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0 = absc (\ x) \ (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0' = absc (absdbupd (unlabel (transaction_updates T)) x (\ x))" proof - define comp0 where "comp0 \ abs_substs_fun ` set (transaction_check_comp (\_ _. True) (FP, OCC, TI) T)" define check0 where "check0 \ transaction_check (FP, OCC, TI) T" define upd where "upd \ \\ x. absdbupd (unlabel (transaction_updates T)) x (\ x)" define b0 where "b0 \ \x. THE b. absc b = (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0" note all_defs = comp0_def check0_def a0_def a0'_def upd_def b0_def \_def T'_def have \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (metis reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s admissible_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P(1) \_reach) have \_interp: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (metis \ welltyped_constraint_model_def constraint_model_def, metis \ welltyped_constraint_model_def, metis \ welltyped_constraint_model_def constraint_model_def) have \_is_T_model: "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) (set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)) (unlabel T') \" using \ unlabel_append[of \ T'] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of "unlabel \" \ "[]"] strand_sem_append_stateful[of "{}" "{}" "unlabel \" "unlabel T'" \] by (simp add: welltyped_constraint_model_def constraint_model_def db\<^sub>s\<^sub>s\<^sub>t_def) note T_adm = bspec[OF P(1) T] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] have T_no_bvars: "fv_transaction T = vars_transaction T" "bvars_transaction T = {}" using admissible_transaction_no_bvars[OF T_adm] by simp_all have T_vars_const_typed: "\x \ fv_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" and T_fresh_vars_value_typed: "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" using T P protocol_transaction_vars_TAtom_typed(2,3)[of T] by simp_all note \_empty = admissible_transaction_decl_subst_empty[OF T_adm \] have wt_\\\: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \ \\<^sub>s \ \\<^sub>s \)" and wt_\\: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \ \\<^sub>s \)" using \_wt wt_subst_compose transaction_decl_fresh_renaming_substs_wt[OF \ \ \] by (blast, blast) have T_vars_vals: "\x \ fv_transaction T. \n. (\ \\<^sub>s \ \\<^sub>s \) x \ \ = Fun (Val n) []" proof fix x assume x: "x \ fv_transaction T" have "\n. (\ \\<^sub>s \) x \ \ = Fun (Val n) []" proof (cases "x \ subst_domain \") case True then obtain n where "\ x = Fun (Val n) []" using transaction_fresh_subst_sends_to_val[OF \] transaction_fresh_subst_domain[OF \] T_fresh_vars_value_typed by metis thus ?thesis by (simp add: subst_compose_def) next case False hence *: "(\ \\<^sub>s \) x = \ x" by (auto simp add: subst_compose_def) obtain y where y: "\\<^sub>v x = \\<^sub>v y" "\ x = Var y" using transaction_renaming_subst_wt[OF \] transaction_renaming_subst_is_renaming(1)[OF \] by (metis \.simps(1) prod.exhaust wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) hence "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" using x * T_no_bvars(2) unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \"] fv\<^sub>s\<^sub>s\<^sub>t_subst_fv_subset[of x "unlabel (transaction_strand T)" "\ \\<^sub>s \"] by auto hence "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \))" using fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \"] fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \"] unlabel_append[of \] by auto thus ?thesis using x y * T P (* T_vars_const_typed *) constraint_model_Value_term_is_Val[ OF reachable_constraints.step[OF \_reach T \ \ \] \[unfolded T'_def] P P_occ, of y] admissible_transaction_Value_vars[of T] \_empty by simp qed thus "\n. (\ \\<^sub>s \ \\<^sub>s \) x \ \ = Fun (Val n) []" using \_empty by simp qed have T_vars_absc: "\x \ fv_transaction T. \!n. (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0 = absc n" using T_vars_vals by fastforce hence "(absc \ b0) x = (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0" when "x \ fv_transaction T" for x using that unfolding b0_def by fastforce hence T_vars_absc': "t \ (absc \ b0) = t \ (\ \\<^sub>s \ \\<^sub>s \) \ \ \\<^sub>\ a0" when "fv t \ fv_transaction T" "\n T. Fun (Val n) T \ subterms t" for t using that(1) abs_term_subst_eq'[OF _ that(2), of "\ \\<^sub>s \ \\<^sub>s \ \\<^sub>s \" a0 "absc \ b0"] subst_compose[of "\ \\<^sub>s \ \\<^sub>s \" \] subst_subst_compose[of t "\ \\<^sub>s \ \\<^sub>s \" \] by fastforce have "\\ \ comp0. \x \ fv_transaction T. fst x = TAtom Value \ b0 x = \ x" proof - let ?C = "set (unlabel (transaction_checks T))" let ?xs = "fv_transaction T - set (transaction_fresh T)" note * = transaction_prop3[OF \_reach T \[unfolded T'_def] \ \ \ FP OCC TI P P_occ] have **: "\x \ set (transaction_fresh T). b0 x = {}" "\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T). intruder_synth_mod_timpls FP TI (t \ \ b0)" (is ?B) proof - show ?B proof (intro ballI impI) fix t assume t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" hence t': "fv t \ fv_transaction T" "\n T. Fun (Val n) T \ subterms t" using trms_transaction_unfold[of T] vars_transaction_unfold[of T] trms\<^sub>s\<^sub>s\<^sub>t_fv_vars\<^sub>s\<^sub>s\<^sub>t_subset[of t "unlabel (transaction_strand T)"] admissible_transactions_no_Value_consts'[OF T_adm] wellformed_transaction_send_receive_fv_subset(1)[OF T_wf t(1)] by blast+ have "intruder_synth_mod_timpls FP TI (t \ (absc \ b0))" using t(1) t' *(2) T_vars_absc' by (metis a0_def) moreover have "(absc \ b0) x = (\ b0) x" when "x \ fv t" for x using that T P admissible_transaction_Value_vars[of T] \fv t \ fv_transaction T\ \\<^sub>v_TAtom''(2)[of x] unfolding \_def by fastforce hence "t \ (absc \ b0) = t \ \ b0" using term_subst_eq[of t "absc \ b0" "\ b0"] by argo ultimately show "intruder_synth_mod_timpls FP TI (t \ \ b0)" using intruder_synth.simps[of "set FP"] by (cases "t \ \ b0") metis+ qed qed (simp add: *(1) a0_def b0_def) have ***: "\x \ ?xs. \s. select\Var x,Fun (Set s) []\ \ ?C \ s \ b0 x" "\x \ ?xs. \s. \Var x in Fun (Set s) []\ \ ?C \ s \ b0 x" "\x \ ?xs. \s. \Var x not in Fun (Set s) []\ \ ?C \ s \ b0 x" "\x \ ?xs. fst x = TAtom Value \ b0 x \ set OCC" unfolding a0_def b0_def using *(3,4) apply (force, force) using *(5) apply force using *(6) admissible_transaction_Value_vars[OF bspec[OF P T]] by force show ?thesis using transaction_check_comp_in[OF T_adm **[unfolded \_def] ***] unfolding comp0_def by metis qed hence 1: "\\ \ comp0. \x \ fv_transaction T. fst x = TAtom Value \ (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0 = absc (\ x)" using T_vars_absc unfolding b0_def a0_def by fastforce obtain \ where \: "\ \ comp0" "\x \ fv_transaction T. fst x = TAtom Value \ (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0 = absc (\ x)" using 1 by moura have 2: "\ x \ \ \\<^sub>\ \\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D) = absc (absdbupd (unlabel A) x d)" when "\ x \ \ \\<^sub>\ \\<^sub>0 D = absc d" and "\t u. insert\t,u\ \ set (unlabel A) \ (\y s. t = Var y \ u = Fun (Set s) [])" and "\t u. delete\t,u\ \ set (unlabel A) \ (\y s. t = Var y \ u = Fun (Set s) [])" and "\y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \ x \ \ = \ y \ \ \ x = y" and "\y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \n. \ y \ \ = Fun (Val n) []" and x: "\ x \ \ = Fun (Val n) []" and D: "\d \ set D. \s. snd d = Fun (Set s) []" for A::"('fun,'atom,'sets,'lbl) prot_strand" and x \ D n d using that(2,3,4,5) proof (induction A rule: List.rev_induct) case (snoc a A) then obtain l b where a: "a = (l,b)" by (metis surj_pair) have IH: "\\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D) n = absdbupd (unlabel A) x d" using snoc unlabel_append[of A "[a]"] a x by (simp del: unlabel_append) have b_prems: "\y \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p b. \ x \ \ = \ y \ \ \ x = y" "\y \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p b. \n. \ y \ \ = Fun (Val n) []" using snoc.prems(3,4) a by (simp_all add: unlabel_def) have *: "filter is_Update (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[a] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) = filter is_Update (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" "filter is_Update (unlabel (A@[a])) = filter is_Update (unlabel A)" when "\is_Update b" using that a by (cases b, simp_all add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def unlabel_def subst_apply_labeled_stateful_strand_def)+ note ** = IH a dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_append[of A "[a]" \] note *** = * absdbupd_filter[of "unlabel (A@[a])"] absdbupd_filter[of "unlabel A"] db\<^sub>s\<^sub>s\<^sub>t_filter[of "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[a] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))"] db\<^sub>s\<^sub>s\<^sub>t_filter[of "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))"] note **** = **(2,3) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_snoc[of A a \] unlabel_append[of "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "[dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]"] db\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "unlabel [dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]" \ D] have "\\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[a] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D) n = absdbupd (unlabel (A@[a])) x d" using ** *** proof (cases b) case (Insert t t') then obtain y s m where y: "t = Var y" "t' = Fun (Set s) []" "\ y \ \ = Fun (Val m) []" using snoc.prems(1) b_prems(2) a by (fastforce simp add: unlabel_def) hence a': "db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[a] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D = List.insert ((Fun (Val m) [], Fun (Set s) [])) (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ D)" "unlabel [dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \] = [insert\\ y, Fun (Set s) []\]" "unlabel [a] = [insert\Var y, Fun (Set s) []\]" using **** Insert by simp_all show ?thesis proof (cases "x = y") case True hence "\ x \ \ = \ y \ \" by simp hence "\\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[a] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D) n = insert s (\\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D) n)" by (metis (no_types, lifting) y(3) a'(1) x dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst to_abs_list_insert') thus ?thesis using True IH a'(3) absdbupd_append[of "unlabel A"] by (simp add: unlabel_def) next case False hence "\ x \ \ \ \ y \ \" using b_prems(1) y Insert by simp hence "\\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[a] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D) n = \\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D) n" by (metis (no_types, lifting) y(3) a'(1) x dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst to_abs_list_insert) thus ?thesis using False IH a'(3) absdbupd_append[of "unlabel A"] by (simp add: unlabel_def) qed next case (Delete t t') then obtain y s m where y: "t = Var y" "t' = Fun (Set s) []" "\ y \ \ = Fun (Val m) []" using snoc.prems(2) b_prems(2) a by (fastforce simp add: unlabel_def) hence a': "db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[a] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D = List.removeAll ((Fun (Val m) [], Fun (Set s) [])) (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ D)" "unlabel [dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \] = [delete\\ y, Fun (Set s) []\]" "unlabel [a] = [delete\Var y, Fun (Set s) []\]" using **** Delete by simp_all have "\s S. snd d = Fun (Set s) []" when "d \ set (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ D)" for d using snoc.prems(1,2) db\<^sub>l\<^sub>s\<^sub>s\<^sub>t_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_set_ex[OF that _ _ D] by (simp add: unlabel_def) moreover { fix t::"('fun,'atom,'sets,'lbl) prot_term" and D::"(('fun,'atom,'sets,'lbl) prot_term \ ('fun,'atom,'sets,'lbl) prot_term) list" assume "\d \ set D. \s. snd d = Fun (Set s) []" hence "removeAll (t, Fun (Set s) []) D = filter (\d. \S. d = (t, Fun (Set s) S)) D" by (induct D) auto } ultimately have a'': "List.removeAll ((Fun (Val m) [], Fun (Set s) [])) (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ D) = filter (\d. \S. d = (Fun (Val m) [], Fun (Set s) S)) (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ D)" by simp show ?thesis proof (cases "x = y") case True hence "\ x \ \ = \ y \ \" by simp hence "\\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[a] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D) n = (\\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D) n) - {s}" using y(3) a'' a'(1) x by (simp add: dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst to_abs_list_remove_all') thus ?thesis using True IH a'(3) absdbupd_append[of "unlabel A"] by (simp add: unlabel_def) next case False hence "\ x \ \ \ \ y \ \" using b_prems(1) y Delete by simp hence "\\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[a] \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D) n = \\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ D) n" by (metis (no_types, lifting) y(3) a'(1) x dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst to_abs_list_remove_all) thus ?thesis using False IH a'(3) absdbupd_append[of "unlabel A"] by (simp add: unlabel_def) qed qed simp_all thus ?case by (simp add: x) qed (simp add: that(1)) have 3: "x = y" when xy: "(\ \\<^sub>s \ \\<^sub>s \) x \ \ = (\ \\<^sub>s \ \\<^sub>s \) y \ \" "x \ fv_transaction T" "y \ fv_transaction T" for x y proof - have "x \ set (transaction_fresh T) \ y \ set (transaction_fresh T) \ ?thesis" using xy admissible_transaction_strand_sem_fv_ineq[OF T_adm \_is_T_model[unfolded T'_def]] by fast moreover { assume *: "x \ set (transaction_fresh T)" "y \ set (transaction_fresh T)" hence "\\<^sub>v x = TAtom Value" "\\<^sub>v y = TAtom Value" using T_fresh_vars_value_typed by (blast, blast) then obtain xn yn where "\ x = Fun (Val xn) []" "\ y = Fun (Val yn) []" using * transaction_fresh_subst_sends_to_val[OF \] by meson hence "\ x = \ y" using that(1) \_empty by (simp add: subst_compose) moreover have "inj_on \ (subst_domain \)" "x \ subst_domain \" "y \ subst_domain \" using * \ unfolding transaction_fresh_subst_def by auto ultimately have ?thesis unfolding inj_on_def by blast } moreover have False when "x \ set (transaction_fresh T)" "y \ set (transaction_fresh T)" using that(2) xy T_no_bvars admissible_transaction_Value_vars[OF bspec[OF P T], of y] transaction_prop4[OF \_reach T \[unfolded T'_def] \ \ \ P P_occ that(1), of y] by auto moreover have False when "x \ set (transaction_fresh T)" "y \ set (transaction_fresh T)" using that(1) xy T_no_bvars admissible_transaction_Value_vars[OF bspec[OF P T], of x] transaction_prop4[OF \_reach T \[unfolded T'_def] \ \ \ P P_occ that(2), of x] by fastforce ultimately show ?thesis by metis qed have 4: "\y s. t = Var y \ u = Fun (Set s) []" when "insert\t,u\ \ set (unlabel (transaction_strand T))" for t u using that admissible_transaction_strand_step_cases(3)[OF T_adm] T_wf by blast have 5: "\y s. t = Var y \ u = Fun (Set s) []" when "delete\t,u\ \ set (unlabel (transaction_strand T))" for t u using that admissible_transaction_strand_step_cases(3)[OF T_adm] T_wf by blast have 6: "\n. (\ \\<^sub>s \ \\<^sub>s \) y \ \ = Fun (Val n) []" when "y \ fv_transaction T" for y using that by (simp add: T_vars_vals) have "list_all wellformed_transaction P" "list_all admissible_transaction_updates P" using P(1) Ball_set[of P admissible_transaction'] Ball_set[of P wellformed_transaction] Ball_set[of P admissible_transaction_updates] admissible_transaction_is_wellformed_transaction(1,3) by fastforce+ hence 7: "\s. snd d = Fun (Set s) []" when "d \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" for d using that reachable_constraints_db\<^sub>l\<^sub>s\<^sub>s\<^sub>t_set_args_empty[OF \_reach] unfolding admissible_transaction_updates_def by (cases d) simp have "(\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0' = absc (upd \ x)" when x: "x \ fv_transaction T" "fst x = TAtom Value" for x proof - have "(\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ \\<^sub>0 (db'\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \ (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)) = absc (absdbupd (unlabel (transaction_strand T)) x (\ x))" using 2[of "\ \\<^sub>s \ \\<^sub>s \" x "db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \" "\ x" "transaction_strand T"] 3[OF _ x(1)] 4 5 6[OF that(1)] 6 7 x \(2) unfolding all_defs by blast thus ?thesis using x db\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \"] absdbupd_wellformed_transaction[OF T_wf] unfolding all_defs db\<^sub>s\<^sub>s\<^sub>t_def by force qed thus ?thesis using \ \\<^sub>v_TAtom''(2) unfolding all_defs by blast qed lemma transaction_prop6: fixes T \ \ \ \ \ T' a0 a0' defines "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" and "a0 \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" and "a0' \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \)" assumes \_reach: "\ \ reachable_constraints P" and T: "T \ set P" and \: "welltyped_constraint_model \ (\@T')" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and FP: "analyzed (timpl_closure_set (set FP) (set TI))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" "\t \ \\<^sub>i\<^sub>k \ \. timpl_closure_set (set FP) (set TI) \\<^sub>c t" and OCC: "\t \ timpl_closure_set (set FP) (set TI). \f \ funs_term t. is_Abs f \ f \ Abs ` set OCC" "timpl_closure_set (absc ` set OCC) (set TI) \ absc ` set OCC" "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \ \ absc ` set OCC" and TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and step: "list_all (transaction_check (FP, OCC, TI)) P" shows "\t \ timpl_closure_set (\\<^sub>i\<^sub>k \ \) (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \). timpl_closure_set (set FP) (set TI) \\<^sub>c t" (is ?A) and "timpl_closure_set (\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \) (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \) \ absc ` set OCC" (is ?B) and "\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T). is_Fun (t \ (\ \\<^sub>s \ \\<^sub>s \) \ \ \\<^sub>\ a0') \ timpl_closure_set (set FP) (set TI) \\<^sub>c t \ (\ \\<^sub>s \ \\<^sub>s \) \ \ \\<^sub>\ a0'" (is ?C) and "\x \ fv_transaction T. \\<^sub>v x = TAtom Value \ (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0' \ absc ` set OCC" (is ?D) proof - define comp0 where "comp0 \ abs_substs_fun ` set (transaction_check_comp (\_ _. True) (FP, OCC, TI) T)" define check0 where "check0 \ transaction_check (FP, OCC, TI) T" define upd where "upd \ \\ x. absdbupd (unlabel (transaction_updates T)) x (\ x)" define \ where "\ \ \\ x. if fst x = TAtom Value then (absc \ \) x else Var x" note T_adm = bspec[OF P T] note T_occ = bspec[OF P_occ T] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] have \_prop: "\ \ x = absc (\ x)" when "\\<^sub>v x = TAtom Value" for \ x using that \\<^sub>v_TAtom''(2)[of x] unfolding \_def by simp \ \The set-membership status of all value constants in T under \\\, \\\, \\\ are covered by the check\ have 0: "\\ \ comp0. \x \ fv_transaction T. \\<^sub>v x = TAtom Value \ (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0 = absc (\ x) \ (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0' = absc (upd \ x)" using transaction_prop5[OF \_reach T \[unfolded T'_def] \ \ \ FP OCC TI P P_occ step] unfolding a0_def a0'_def T'_def upd_def comp0_def by blast \ \All set-membership changes are covered by the term implication graph\ have 1: "(\ x, upd \ x) \ (set TI)\<^sup>+" when "\ \ comp0" "\ x \ upd \ x" "x \ fv_transaction T" "x \ set (transaction_fresh T)" for x \ using T that step Ball_set[of P "transaction_check (FP, OCC, TI)"] transaction_prop1[of \ "\_ _. True" FP OCC TI T x] TI unfolding upd_def comp0_def transaction_check_def by blast \ \All set-membership changes are covered by the fixed point\ have 2: (* "\ x \ set OCC" *) "upd \ x \ set OCC" when "\ \ comp0" "x \ fv_transaction T" "fst x = TAtom Value" for x \ using T that step Ball_set[of P "transaction_check (FP, OCC, TI)"] T_adm T_occ FP OCC TI transaction_prop2[of \ "\_ _. True" FP OCC TI T x] unfolding upd_def comp0_def transaction_check_def by blast obtain \ where \: "\ \ comp0" "\x \ fv_transaction T. \\<^sub>v x = TAtom Value \ (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0 = absc (\ x) \ (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0' = absc (upd \ x)" using 0 by moura have "\x. ab = (\ x, upd \ x) \ x \ fv_transaction T - set (transaction_fresh T) \ \ x \ upd \ x" when ab: "ab \ \\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \" for ab proof - obtain a b where ab': "ab = (a,b)" by (metis surj_pair) then obtain x where x: "a \ b" "x \ fv_transaction T" "x \ set (transaction_fresh T)" "absc a = (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0" "absc b = (\ \\<^sub>s \ \\<^sub>s \) x \ \ \\<^sub>\ a0'" using ab unfolding abs_term_implications_def a0_def a0'_def T'_def by blast hence "absc a = absc (\ x)" "absc b = absc (upd \ x)" using \(2) admissible_transaction_Value_vars[OF bspec[OF P T] x(2)] by metis+ thus ?thesis using x ab' by blast qed hence \\<^sub>t\<^sub>i_TI_subset: "\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \ \ {(a,b) \ (set TI)\<^sup>+. a \ b}" using 1[OF \(1)] by blast have "timpl_closure_set (timpl_closure_set (set FP) (set TI)) (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \) \\<^sub>c t" when t: "t \ timpl_closure_set (\\<^sub>i\<^sub>k \ \) (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \)" for t using timpl_closure_set_is_timpl_closure_union[of "\\<^sub>i\<^sub>k \ \" "\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \"] intruder_synth_timpl_closure_set FP(3) t by blast thus ?A using ideduct_synth_mono[OF _ timpl_closure_set_mono[OF subset_refl[of "timpl_closure_set (set FP) (set TI)"] \\<^sub>t\<^sub>i_TI_subset]] timpl_closure_set_timpls_trancl_eq'[of "timpl_closure_set (set FP) (set TI)" "set TI"] unfolding timpl_closure_set_idem by force have "timpl_closure_set (\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \) (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \) \ timpl_closure_set (absc ` set OCC) {(a,b) \ (set TI)\<^sup>+. a \ b}" using timpl_closure_set_mono[OF _ \\<^sub>t\<^sub>i_TI_subset] OCC(3) by blast thus ?B using OCC(2) timpl_closure_set_timpls_trancl_subset' by blast have "transaction_check_post (FP, OCC, TI) T \" using T \(1) step unfolding transaction_check_def transaction_check'_def comp0_def list_all_iff by fastforce hence 3: "timpl_closure_set (set FP) (set TI) \\<^sub>c t \ \ (upd \)" when "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "is_Fun (t \ \ (upd \))" for t using that unfolding transaction_check_post_def upd_def \_def intruder_synth_mod_timpls_is_synth_timpl_closure_set[OF TI, symmetric] by fastforce have 4: "\x \ fv t. (\ \\<^sub>s \ \\<^sub>s \ \\<^sub>s \) x \\<^sub>\ a0' = \ (upd \) x" when "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" for t using wellformed_transaction_send_receive_fv_subset(2)[OF T_wf that] \(2) subst_compose[of "\ \\<^sub>s \ \\<^sub>s \" \] \_prop admissible_transaction_Value_vars[OF bspec[OF P T]] by fastforce have 5: "\n T. Fun (Val n) T \ subterms t" when "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" for t using that admissible_transactions_no_Value_consts'[OF T_adm] trms_transaction_unfold[of T] by blast show ?D using 2[OF \(1)] \(2) \\<^sub>v_TAtom''(2) unfolding a0'_def T'_def by blast show ?C using 3 abs_term_subst_eq'[OF 4 5] by simp qed lemma reachable_constraints_covered_step: fixes \::"('fun,'atom,'sets,'lbl) prot_constr" assumes \_reach: "\ \ reachable_constraints P" and T: "T \ set P" and \: "welltyped_constraint_model \ (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \))" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and FP: "analyzed (timpl_closure_set (set FP) (set TI))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" "\t \ \\<^sub>i\<^sub>k \ \. timpl_closure_set (set FP) (set TI) \\<^sub>c t" "ground (set FP)" and OCC: "\t \ timpl_closure_set (set FP) (set TI). \f \ funs_term t. is_Abs f \ f \ Abs ` set OCC" "timpl_closure_set (absc ` set OCC) (set TI) \ absc ` set OCC" "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \ \ absc ` set OCC" and TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and transactions_covered: "list_all (transaction_check (FP, OCC, TI)) P" shows "\t \ \\<^sub>i\<^sub>k (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \. timpl_closure_set (set FP) (set TI) \\<^sub>c t" (is ?A) and "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) \ \ absc ` set OCC" (is ?B) proof - note step_props = transaction_prop6[OF \_reach T \ \ \ \ FP(1,2,3) OCC TI P P_occ transactions_covered] define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define a0 where "a0 \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" define a0' where "a0' \ \\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \)" define vals where "vals \ \S::('fun,'atom,'sets,'lbl) prot_constr. {t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S) \\<^sub>s\<^sub>e\<^sub>t \. \n. t = Fun (Val n) []}" define vals_sym where "vals_sym \ \S::('fun,'atom,'sets,'lbl) prot_constr. {t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S). (\n. t = Fun (Val n) []) \ (\m. t = Var (TAtom Value,m))}" have \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" by (metis \ welltyped_constraint_model_def) have \_grounds: "fv (t \ \) = {}" for t using \ interpretation_grounds[of \] unfolding welltyped_constraint_model_def constraint_model_def by auto have wt_\\\: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \ \\<^sub>s \ \\<^sub>s \)" and wt_\\: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \ \\<^sub>s \)" using \_wt wt_subst_compose transaction_decl_fresh_renaming_substs_wt[OF \ \ \] by (blast, blast) have "\T\set P. bvars_transaction T = {}" using P admissible_transactionE(4) by metis hence \_no_bvars: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" using reachable_constraints_no_bvars[OF \_reach] by metis have \_vals: "\n. \ (TAtom Value, m) = Fun (Val n) []" when "(TAtom Value, m) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" for m using constraint_model_Value_term_is_Val'[ OF \_reach welltyped_constraint_model_prefix[OF \] P P_occ] \_no_bvars vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] that by blast have vals_sym_vals: "t \ \ \ vals \" when t: "t \ vals_sym \" for t proof (cases t) case (Var x) then obtain m where *: "x = (TAtom Value,m)" using t unfolding vals_sym_def by blast moreover have "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using t unfolding vals_sym_def by blast hence "t \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \" "\n. \ (Var Value, m) = Fun (Val n) []" using Var * \_vals[of m] var_subterm_trms\<^sub>s\<^sub>s\<^sub>t_is_vars\<^sub>s\<^sub>s\<^sub>t[of x "unlabel \"] \\<^sub>v_TAtom[of Value m] reachable_constraints_Value_vars_are_fv[OF \_reach P(1), of x] by blast+ ultimately show ?thesis using Var unfolding vals_def by auto next case (Fun f T) then obtain n where "f = Val n" "T = []" using t unfolding vals_sym_def by blast moreover have "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using t unfolding vals_sym_def by blast hence "t \ \ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \" using Fun by blast ultimately show ?thesis using Fun unfolding vals_def by auto qed have vals_vals_sym: "\s. s \ vals_sym \ \ t = s \ \" when "t \ vals \" for t using that constraint_model_Val_is_Value_term[OF \] unfolding vals_def vals_sym_def by fast note T_adm = bspec[OF P T] note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] have 0: "\\<^sub>i\<^sub>k (\@T') \ = (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0' \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0'" "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s (\@T') \ = vals \ \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0' \ vals T' \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0'" by (metis abs_intruder_knowledge_append a0'_def, metis abs_value_constants_append[of \ T' \] a0'_def vals_def) have 1: "(ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0' = (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t (\ \\<^sub>s \ \\<^sub>s \) \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0'" by (metis T'_def dual_transaction_ik_is_transaction_send''[OF T_wf]) have 2: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) \ subst_domain \ = {}" "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) \ subst_domain \ = {}" "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) \ subst_domain \ = {}" using admissible_transactionE(4)[OF T_adm] by blast+ have "vals T' \ (\ \\<^sub>s \ \\<^sub>s \) ` fv_transaction T \\<^sub>s\<^sub>e\<^sub>t \" proof fix t assume "t \ vals T'" then obtain s n where s: "s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" "t = s \ \" "t = Fun (Val n) []" unfolding vals_def by fast then obtain u where u: "u \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))" "s = u \ (\ \\<^sub>s \ \\<^sub>s \)" using transaction_decl_fresh_renaming_substs_trms[OF \ \ \ 2] trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \"] unfolding T'_def by blast - have *: "t = u \ (\ \\<^sub>s \ \\<^sub>s \ \\<^sub>s \)" by (metis subst_subst_compose s(2) u(2)) + have *: "t = u \ (\ \\<^sub>s \ \\<^sub>s \ \\<^sub>s \)" using s(2) u(2) subst_subst_compose by simp then obtain x where x: "u = Var x" using s(3) admissible_transactions_no_Value_consts(1)[OF T_adm u(1)] by (cases u) force+ hence **: "x \ vars_transaction T" by (metis u(1) var_subterm_trms\<^sub>s\<^sub>s\<^sub>t_is_vars\<^sub>s\<^sub>s\<^sub>t) have "\\<^sub>v x = TAtom Value" using * x s(3) wt_subst_trm''[OF wt_\\\, of u] by simp thus "t \ (\ \\<^sub>s \ \\<^sub>s \) ` fv_transaction T \\<^sub>s\<^sub>e\<^sub>t \" using admissible_transaction_Value_vars_are_fv[OF T_adm **] x * - subst_apply_term.simps(1)[of x "\ \\<^sub>s \ \\<^sub>s \ \\<^sub>s \"] + eval_term.simps(1)[of _ x "\ \\<^sub>s \ \\<^sub>s \ \\<^sub>s \"] subst_comp_set_image[of "\ \\<^sub>s \ \\<^sub>s \" \ "fv_transaction T"] by blast qed hence 3: "vals T' \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0' \ ((\ \\<^sub>s \ \\<^sub>s \) ` fv_transaction T \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0'" by (simp add: abs_apply_terms_def image_mono) have "t \ \ \\<^sub>\ a0' \ timpl_closure_set (\\<^sub>i\<^sub>k \ \) (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \)" when "t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" for t using that abs_in[OF imageI[OF that]] \\<^sub>t\<^sub>i_covers_\\<^sub>0_ik[OF \_reach T \ \ \ \ P P_occ] timpl_closure_set_mono[of "{t \ \ \\<^sub>\ a0}" "\\<^sub>i\<^sub>k \ \" "\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \" "\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \"] unfolding a0_def a0'_def T'_def abs_intruder_knowledge_def by fast hence A: "\\<^sub>i\<^sub>k (\@T') \ \ timpl_closure_set (\\<^sub>i\<^sub>k \ \) (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \) \ (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t (\ \\<^sub>s \ \\<^sub>s \) \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0'" using 0(1) 1 by (auto simp add: abs_apply_terms_def) have "t \ \ \\<^sub>\ a0' \ timpl_closure_set {t \ \ \\<^sub>\ a0} (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \)" when t: "t \ vals_sym \" for t proof - have "(\n. t = Fun (Val n) [] \ t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ (\n. t = Var (TAtom Value,n) \ (TAtom Value,n) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" (is "?P \ ?Q") using t var_subterm_trms\<^sub>s\<^sub>s\<^sub>t_is_vars\<^sub>s\<^sub>s\<^sub>t[of _ "unlabel \"] \\<^sub>v_TAtom[of Value] reachable_constraints_Value_vars_are_fv[OF \_reach P(1)] unfolding vals_sym_def by fast thus ?thesis proof assume ?P then obtain n where n: "t = Fun (Val n) []" "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by moura thus ?thesis using \\<^sub>t\<^sub>i_covers_\\<^sub>0_Val[OF \_reach T \ \ \ \ P P_occ, of n] unfolding a0_def a0'_def T'_def by fastforce next assume ?Q thus ?thesis using \\<^sub>t\<^sub>i_covers_\\<^sub>0_Var[OF \_reach T \ \ \ \ P P_occ] unfolding a0_def a0'_def T'_def by fastforce qed qed moreover have "t \ \ \\<^sub>\ a0 \ \\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \" when "t \ vals_sym \" for t using that abs_in vals_sym_vals unfolding a0_def abs_value_constants_def vals_sym_def vals_def by (metis (mono_tags, lifting)) ultimately have "t \ \ \\<^sub>\ a0' \ timpl_closure_set (\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \) (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \)" when t: "t \ vals_sym \" for t using t timpl_closure_set_mono[of "{t \ \ \\<^sub>\ a0}" "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \" "\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \" "\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \"] by blast hence "t \\<^sub>\ a0' \ timpl_closure_set (\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \) (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \)" when t: "t \ vals \" for t using vals_vals_sym[OF t] by blast hence B: "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s (\@T') \ \ timpl_closure_set (\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \) (\\<^sub>t\<^sub>i \ T (\ \\<^sub>s \ \\<^sub>s \) \) \ ((\ \\<^sub>s \ \\<^sub>s \) ` fv_transaction T \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a0'" using 0(2) 3 by (simp add: abs_apply_terms_def image_subset_iff) have 4: "fv (t \ \ \\<^sub>s \ \\<^sub>s \ \ \ \\<^sub>\ a) = {}" for t a using \_grounds[of "t \ \ \\<^sub>s \ \\<^sub>s \"] abs_fv[of "t \ \ \\<^sub>s \ \\<^sub>s \ \ \" a] by argo have "is_Fun (t \ \ \\<^sub>s \ \\<^sub>s \ \ \ \\<^sub>\ a0')" for t using 4[of t a0'] by force thus ?A using A step_props(1,3) unfolding T'_def a0_def a0'_def abs_apply_terms_def by blast show ?B using B step_props(2,4) admissible_transaction_Value_vars[OF bspec[OF P T]] by (auto simp add: T'_def a0_def a0'_def abs_apply_terms_def) qed lemma reachable_constraints_covered: assumes \_reach: "\ \ reachable_constraints P" and \: "welltyped_constraint_model \ \" and FP: "analyzed (timpl_closure_set (set FP) (set TI))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" "ground (set FP)" and OCC: "\t \ timpl_closure_set (set FP) (set TI). \f \ funs_term t. is_Abs f \ f \ Abs ` set OCC" "timpl_closure_set (absc ` set OCC) (set TI) \ absc ` set OCC" and TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and transactions_covered: "list_all (transaction_check (FP, OCC, TI)) P" shows "\t \ \\<^sub>i\<^sub>k \ \. timpl_closure_set (set FP) (set TI) \\<^sub>c t" and "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \ \ absc ` set OCC" using \_reach \ proof (induction rule: reachable_constraints.induct) case init { case 1 show ?case by (simp add: abs_intruder_knowledge_def) } { case 2 show ?case by (simp add: abs_value_constants_def) } next case (step \ T \ \ \) { case 1 hence "welltyped_constraint_model \ \" by (metis welltyped_constraint_model_prefix) hence IH: "\t \ \\<^sub>i\<^sub>k \ \. timpl_closure_set (set FP) (set TI) \\<^sub>c t" "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \ \ absc ` set OCC" using step.IH by metis+ show ?case using reachable_constraints_covered_step[ OF step.hyps(1,2) "1.prems" step.hyps(3-5) FP(1,2) IH(1) FP(3) OCC IH(2) TI P P_occ transactions_covered] by metis } { case 2 hence "welltyped_constraint_model \ \" by (metis welltyped_constraint_model_prefix) hence IH: "\t \ \\<^sub>i\<^sub>k \ \. timpl_closure_set (set FP) (set TI) \\<^sub>c t" "\\<^sub>v\<^sub>a\<^sub>l\<^sub>s \ \ \ absc ` set OCC" using step.IH by metis+ show ?case using reachable_constraints_covered_step[ OF step.hyps(1,2) "2.prems" step.hyps(3-5) FP(1,2) IH(1) FP(3) OCC IH(2) TI P P_occ transactions_covered] by metis } qed lemma attack_in_fixpoint_if_attack_in_ik: fixes FP::"('fun,'atom,'sets,'lbl) prot_terms" assumes "\t \ IK \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a. FP \\<^sub>c t" and "attack\n\ \ IK" shows "attack\n\ \ FP" proof - have "attack\n\ \\<^sub>\ a \ IK \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a" by (rule abs_in[OF assms(2)]) hence "FP \\<^sub>c attack\n\ \\<^sub>\ a" using assms(1) by blast moreover have "attack\n\ \\<^sub>\ a = attack\n\" by simp ultimately have "FP \\<^sub>c attack\n\" by metis thus ?thesis using ideduct_synth_priv_const_in_ik[of FP "Attack n"] by simp qed lemma attack_in_fixpoint_if_attack_in_timpl_closure_set: fixes FP::"('fun,'atom,'sets,'lbl) prot_terms" assumes "attack\n\ \ timpl_closure_set FP TI" shows "attack\n\ \ FP" proof - have "\f \ funs_term (attack\n\). \is_Abs f" by auto thus ?thesis using timpl_closure_set_no_Abs_in_set[OF assms] by blast qed theorem prot_secure_if_fixpoint_covered_typed: assumes FP: "analyzed (timpl_closure_set (set FP) (set TI))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" "ground (set FP)" and OCC: "\t \ timpl_closure_set (set FP) (set TI). \f \ funs_term t. is_Abs f \ f \ Abs ` set OCC" "timpl_closure_set (absc ` set OCC) (set TI) \ absc ` set OCC" and TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and P: "\T \ set P. admissible_transaction T" "has_initial_value_producing_transaction P" and transactions_covered: "list_all (transaction_check (FP, OCC, TI)) (map add_occurs_msgs P)" and attack_notin_FP: "attack\n\ \ set FP" and \: "\ \ reachable_constraints P" shows "\\. welltyped_constraint_model \ (\@[(l, send\[attack\n\]\)])" (is "\\. ?Q \") proof assume "\\. ?Q \" then obtain \ \ where \: "\ \ reachable_constraints (map add_occurs_msgs P)" and \: "welltyped_constraint_model \ (\@[(l, send\[attack\n\]\)])" using add_occurs_msgs_soundness[OF P \] unfolding list_all_iff by blast have P': "\T \ set (map add_occurs_msgs P). admissible_transaction' T" "\T \ set (map add_occurs_msgs P). admissible_transaction_occurs_checks T" using P add_occurs_msgs_admissible_occurs_checks[OF admissible_transactionE'(1)] by auto have 0: "attack\n\ \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" using welltyped_constraint_model_prefix[OF \] reachable_constraints_covered(1)[OF \ _ FP OCC TI P' transactions_covered] attack_in_fixpoint_if_attack_in_ik[ of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" "\\<^sub>0 (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" "timpl_closure_set (set FP) (set TI)" n] attack_in_fixpoint_if_attack_in_timpl_closure_set attack_notin_FP unfolding abs_intruder_knowledge_def by blast have 1: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ attack\n\" using \ strand_sem_append_stateful[of "{}" "{}" "unlabel \" _ \] unfolding welltyped_constraint_model_def constraint_model_def by force show False using 0 private_const_deduct[OF _ 1] reachable_constraints_receive_attack_if_attack'(1)[ OF \ P'(1) welltyped_constraint_model_prefix[OF \] 1] by simp qed end subsection \Theorem: A Protocol is Secure if it is Covered by a Fixed-Point\ context stateful_protocol_model begin theorem prot_secure_if_fixpoint_covered: fixes P assumes FP: "analyzed (timpl_closure_set (set FP) (set TI))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" "ground (set FP)" and OCC: "\t \ timpl_closure_set (set FP) (set TI). \f \ funs_term t. is_Abs f \ f \ Abs ` set OCC" "timpl_closure_set (absc ` set OCC) (set TI) \ absc ` set OCC" and TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and M: "has_all_wt_instances_of \ (\T \ set P. trms_transaction T) N" "finite N" "tfr\<^sub>s\<^sub>e\<^sub>t N" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and P: "\T \ set P. admissible_transaction T" "\T \ set P. list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" "has_initial_value_producing_transaction P" and transactions_covered: "list_all (transaction_check (FP, OCC, TI)) (map add_occurs_msgs P)" and attack_notin_FP: "attack\n\ \ set FP" and A: "\ \ reachable_constraints P" shows "\\. constraint_model \ (\@[(l, send\[attack\n\]\)])" (is "\\. constraint_model \ ?A") proof assume "\\. constraint_model \ ?A" then obtain \ where "constraint_model \ ?A" by moura then obtain \\<^sub>\ where I: "welltyped_constraint_model \\<^sub>\ ?A" using reachable_constraints_typing_result[OF M P(1,2) A] by blast note a = FP OCC TI P(1,3) transactions_covered attack_notin_FP A show False using prot_secure_if_fixpoint_covered_typed[OF a] I by force qed end subsection \Alternative Protocol-Coverage Check\ context stateful_protocol_model begin context begin private lemma transaction_check_variant_soundness_aux0: assumes S: "S \ unlabel (transaction_strand T)" and xs: "xs \ filter (\x. x \ set (transaction_fresh T) \ fst x = TAtom Value) (fv_list\<^sub>s\<^sub>s\<^sub>t S)" and x: "fst x = Var Value" "x \ fv_transaction T" "x \ set (transaction_fresh T)" shows "x \ set xs" using x fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] unfolding xs S by auto private lemma transaction_check_variant_soundness_aux1: fixes T FP S C xs OCC negs poss as assumes C: "C \ unlabel (transaction_checks T)" and S: "S \ unlabel (transaction_strand T)" and xs: "xs \ filter (\x. x \ set (transaction_fresh T) \ fst x = TAtom Value) (fv_list\<^sub>s\<^sub>s\<^sub>t S)" and poss: "poss \ transaction_poschecks_comp C" and negs: "negs \ transaction_negchecks_comp C" and as: "as \ map (\x. (x, set (filter (\ab. poss x \ ab \ negs x \ ab = {}) OCC))) xs" and f: "f \ \x. case List.find (\p. fst p = x) as of Some p \ snd p | None \ {}" and x: "x \ set xs" shows "f x = set (filter (\ab. poss x \ ab \ negs x \ ab = {}) OCC)" proof - define g where "g \ \x. set (filter (\ab. poss x \ ab \ negs x \ ab = {}) OCC)" define gs where "gs \ map (\x. (x, g x)) xs" have 1: "(x, g x) \ set gs" using x unfolding gs_def by simp have 2: "distinct xs" unfolding xs fv_list\<^sub>s\<^sub>s\<^sub>t_def by auto have "\i < length xs. xs ! i = x \ (\j < i. xs ! j \ x)" when x: "x \ set xs" for x proof (rule ex1E[OF distinct_Ex1[OF 2 x]]) fix i assume i: "i < length xs \ xs ! i = x" and "\j. j < length xs \ xs ! j = x \ j = i" hence "\j < length xs. xs ! j = x \ j = i" by blast hence "\j < i. xs ! j = x \ j = i" using i by auto hence "\j < i. xs ! j \ x" by blast thus ?thesis using i by blast qed hence "\i < length gs. gs ! i = (x, g x) \ (\j < i. gs ! j \ (x, g x))" using 1 unfolding gs_def by fastforce hence "\i < length gs. fst (gs ! i) = x \ (x, g x) = gs ! i \ (\j < i. fst (gs ! j) \ x)" using nth_map[of _ xs "\x. (x, g x)"] length_map[of "\x. (x, g x)" xs] unfolding gs_def by (metis (no_types, lifting) fstI min.strict_order_iff min_less_iff_conj) hence "List.find (\p. fst p = x) gs = Some (x, g x)" using find_Some_iff[of "\p. fst p = x" gs "(x, g x)"] by blast thus ?thesis unfolding f as gs_def g_def by force qed private lemma transaction_check_variant_soundness_aux2: fixes T FP S C xs OCC negs poss as assumes C: "C \ unlabel (transaction_checks T)" and S: "S \ unlabel (transaction_strand T)" and xs: "xs \ filter (\x. x \ set (transaction_fresh T) \ fst x = TAtom Value) (fv_list\<^sub>s\<^sub>s\<^sub>t S)" and poss: "poss \ transaction_poschecks_comp C" and negs: "negs \ transaction_negchecks_comp C" and as: "as \ map (\x. (x, set (filter (\ab. poss x \ ab \ negs x \ ab = {}) OCC))) xs" and f: "f \ \x. case List.find (\p. fst p = x) as of Some p \ snd p | None \ {}" and x: "x \ set xs" shows "f x = {}" proof - define g where "g \ \x. set (filter (\ab. poss x \ ab \ negs x \ ab = {}) OCC)" define gs where "gs \ map (\x. (x, g x)) xs" have "(x, y) \ set gs" for y using x unfolding gs_def by force thus ?thesis using find_None_iff[of "\p. fst p = x" gs] unfolding f as gs_def g_def by fastforce qed private lemma synth_abs_substs_constrs_rel_if_synth_abs_substs_constrs: fixes T OCC negs poss defines "\ \ \\ x. if fst x = TAtom Value then (absc \ \) x else Var x" and "ts \ trms_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_receive T))" assumes ts_wf: "\t \ set ts. wf\<^sub>t\<^sub>r\<^sub>m t" and FP_ground: "ground (set FP)" and FP_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" shows "synth_abs_substs_constrs_rel FP OCC TI ts (synth_abs_substs_constrs (FP,OCC,TI) T)" proof - let ?R = "synth_abs_substs_constrs_rel FP OCC TI" let ?D = "synth_abs_substs_constrs_aux FP OCC TI" have *: "?R [t] (?D t)" when "wf\<^sub>t\<^sub>r\<^sub>m t" for t using that proof (induction t) case (Var x) thus ?case using synth_abs_substs_constrs_rel.SolveValueVar[of "?D (Var x)" OCC x TI FP] by fastforce next case (Fun f ss) let ?xs = "fv (Fun f ss)" let ?lst = "map (match_abss OCC TI (Fun f ss)) FP" define flt where "flt = (\\::(('fun,'atom,'sets,'lbl) prot_var \ 'sets set set) option. \ \ None)" define \ where "\ = map the (filter flt (map (match_abss OCC TI (Fun f ss)) FP))" define \1 where "\1 = fun_point_Union (set \)" define \2 where "\2 = fun_point_Inter (?D ` set ss)" have f: "arity f = length ss" using wf_trm_arity[OF Fun.prems] by simp have IH: "?R [s] (?D s)" when s: "s \ set ss" for s using Fun.IH[OF s wf_trm_subterm[OF Fun.prems Fun_param_is_subterm[OF s]]] s by force have \3: "\\. \ \ set \ \ (\s \ set FP. match_abss OCC TI (Fun f ss) s = Some \)" (is "\\. \ \ set \ \ ?P \") proof (intro allI iffI) fix \ assume "\ \ set \" then obtain u where "u \ set FP" "match_abss OCC TI (Fun f ss) u = Some \" unfolding \_def flt_def by fastforce thus "?P \" by blast next fix \ assume "?P \" then obtain u where u: "u \ set FP" "match_abss OCC TI (Fun f ss) u = Some \" by blast have "Some \ \ set ?lst" using u unfolding flt_def by force hence "Some \ \ set (filter flt ?lst)" unfolding flt_def by force moreover have "\\. d = Some \" when d: "d \ set (filter flt ?lst)" for d using d unfolding flt_def by simp ultimately have "\ \ set (map the (filter flt ?lst))" by force thus "\ \ set \" unfolding \_def by blast qed show ?case proof (cases "ss = []") case True note ss = this show ?thesis proof (cases "\public f \ Fun f ss \ set FP") case True thus ?thesis using synth_abs_substs_constrs_rel.SolvePrivConstNotin[of f FP OCC TI] unfolding ss by force next case False thus ?thesis using f synth_abs_substs_constrs_rel.SolvePubConst[of f FP OCC TI] synth_abs_substs_constrs_rel.SolvePrivConstIn[of f FP OCC TI] unfolding ss by auto qed next case False note ss = this hence f': "arity f > 0" using f by auto have IH': "?R ss (fun_point_Inter (?D ` set ss))" using IH synth_abs_substs_constrs_rel.SolveCons[OF ss, of FP OCC TI ?D] by blast have "?D (Fun f ss) = ( fun_point_union (fun_point_Union_list \) (fun_point_Inter_list (map ?D ss)))" using ss synth_abs_substs_constrs_aux.simps(2)[of FP OCC TI f ss] unfolding Let_def \_def flt_def by argo hence "?D (Fun f ss) = fun_point_union \1 \2" using fun_point_Inter_set_eq[of "map ?D ss"] fun_point_Union_set_eq[of \] unfolding \1_def \2_def by simp thus ?thesis using synth_abs_substs_constrs_rel.SolveComposed[ OF f' f[symmetric] \3 \1_def IH'] unfolding \2_def by argo qed qed note l0 = synth_abs_substs_constrs_rel.SolveNil[of FP OCC TI] note d0 = synth_abs_substs_constrs_def ts_def note l1 = * ts_wf synth_abs_substs_constrs_rel.SolveCons[of ts FP OCC TI ?D] note d1 = d0 Let_def fun_point_Inter_set_eq[symmetric] fun_point_Inter_def show ?thesis proof (cases "ts = []") case True thus ?thesis using l0 unfolding d0 by simp next case False thus ?thesis using l1 unfolding d1 by auto qed qed private function (sequential) match_abss'_timpls_transform ::"('c set \ 'c set) list \ ('a,'b,'c,'d) prot_subst \ ('a,'b,'c,'d) prot_term \ ('a,'b,'c,'d) prot_term \ (('a,'b,'c,'d) prot_var \ 'c set set) option" where "match_abss'_timpls_transform TI \ (Var x) (Fun (Abs a) _) = ( if \b ts. \ x = Fun (Abs b) ts \ (a = b \ (a,b) \ set TI) then Some ((\_. {})(x := {a})) else None)" | "match_abss'_timpls_transform TI \ (Fun f ts) (Fun g ss) = ( if f = g \ length ts = length ss then map_option fun_point_Union_list (those (map2 (match_abss'_timpls_transform TI \) ts ss)) else None)" | "match_abss'_timpls_transform _ _ _ _ = None" by pat_completeness auto termination proof - let ?m = "measures [size \ fst \ snd \ snd]" have 0: "wf ?m" by simp show ?thesis apply (standard, use 0 in fast) by (metis (no_types) comp_def fst_conv snd_conv measures_less Fun_zip_size_lt(1)) qed private lemma match_abss'_timpls_transform_Var_inv: assumes "match_abss'_timpls_transform TI \ (Var x) (Fun (Abs a) ts) = Some \" shows "\b ts. \ x = Fun (Abs b) ts \ (a = b \ (a, b) \ set TI)" and "\ = ((\_. {})(x := {a}))" using assms match_abss'_timpls_transform.simps(1)[of TI \ x a ts] by (metis option.distinct(1), metis option.distinct(1) option.inject) private lemma match_abss'_timpls_transform_Fun_inv: assumes "match_abss'_timpls_transform TI \ (Fun f ts) (Fun g ss) = Some \" shows "f = g" (is ?A) and "length ts = length ss" (is ?B) and "\\. Some \ = those (map2 (match_abss'_timpls_transform TI \) ts ss) \ \ = fun_point_Union_list \" (is ?C) and "\(t,s) \ set (zip ts ss). \\'. match_abss'_timpls_transform TI \ t s = Some \'" (is ?D) proof - note 0 = assms match_abss'_timpls_transform.simps(2)[of TI \ f ts g ss] option.distinct(1) show ?A by (metis 0) show ?B by (metis 0) show ?C by (metis (no_types, opaque_lifting) 0 map_option_eq_Some) thus ?D using map2_those_Some_case[of "match_abss'_timpls_transform TI \" ts ss] by fastforce qed private lemma match_abss'_timpl_transform_nonempty_is_fv: assumes "match_abss'_timpls_transform TI \ s t = Some \" and "\ x \ {}" shows "x \ fv s" using assms proof (induction s t arbitrary: TI \ \ rule: match_abss'_timpls_transform.induct) case (1 TI \ y a ts) show ?case using match_abss'_timpls_transform_Var_inv[OF "1.prems"(1)] "1.prems"(2) by fastforce next case (2 TI \ f ts g ss) note prems = "2.prems" note IH = "2.IH" obtain \ where \: "Some \ = those (map2 (match_abss'_timpls_transform TI \) ts ss)" "\ = fun_point_Union_list \" and fg: "f = g" "length ts = length ss" using match_abss'_timpls_transform_Fun_inv[OF prems(1)] by fast have "\\ \ set \. \ x \ {}" using fg(2) prems \ unfolding fun_point_Union_list_def by auto then obtain t' s' \ where ts': "(t',s') \ set (zip ts ss)" "match_abss'_timpls_transform TI \ t' s' = Some \" "\ x \ {}" using those_map2_SomeD[OF \(1)[symmetric]] by blast show ?case using ts'(3) IH[OF conjI[OF fg] ts'(1) _ ts'(2)] set_zip_leftD[OF ts'(1)] by force qed auto private lemma match_abss'_timpls_transformI: fixes s t::"('a,'b,'c,'d) prot_term" and \::"('a,'b,'c,'d) prot_subst" and \::"('a,'b,'c,'d) prot_var \ 'c set set" assumes TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and \: "timpls_transformable_to TI t (s \ \)" and \: "match_abss' s t = Some \" and t: "fv t = {}" and s: "\f \ funs_term s. \is_Abs f" "\x \ fv s. \a. \ x = \a\\<^sub>a\<^sub>b\<^sub>s" shows "match_abss'_timpls_transform TI \ s t = Some \" using \ \ t s proof (induction t arbitrary: s \ \) case (Fun f ts) note prems = Fun.prems note IH = Fun.IH show ?case proof (cases s) case (Var x) obtain a b where a: "f = Abs a" "\ = (\_. {})(x := {a})" and b: "\ x = \b\\<^sub>a\<^sub>b\<^sub>s" using match_abss'_Var_inv[OF prems(2)[unfolded Var]] prems(5)[unfolded Var] by auto thus ?thesis using prems(1) timpls_transformable_to_inv(3)[of TI f ts "Abs b" "[]"] unfolding Var by auto next case (Fun g ss) - note 0 = timpls_transformable_to_inv[OF prems(1)[unfolded Fun subst_apply_term.simps(2)]] + note 0 = timpls_transformable_to_inv[OF prems(1)[unfolded Fun eval_term.simps(2)]] note 1 = match_abss'_Fun_inv[OF prems(2)[unfolded Fun]] obtain \ where \: "those (map2 match_abss' ss ts) = Some \" "\ = fun_point_Union_list \" using 1(3) by force have "timpls_transformable_to TI t' (s' \ \)" "\\'. match_abss' s' t' = Some \'" when "(t',s') \ set (zip ts ss)" for s' t' by (metis 0(2) nth_map[of _ ss] zip_arg_index[OF that], use that 1(4) in_set_zip_swap[of t' s' ts ss] in fast) hence IH': "match_abss'_timpls_transform TI \ s' t' = Some \'" when "(t',s') \ set (zip ts ss)" "match_abss' s' t' = Some \'" for s' t' \' using that IH[of t' s' \ \'] prems(3-) unfolding Fun by (metis (no_types, lifting) set_zip_leftD set_zip_rightD subsetI subset_empty term.set_intros(2) term.set_intros(4)) have "those (map2 (match_abss'_timpls_transform TI \) ss ts) = Some \" using IH' \(1) 1(4) in_set_zip_swap[of _ _ ss ts] those_Some_iff[of "map2 match_abss' ss ts" \] those_Some_iff[of "map2 (match_abss'_timpls_transform TI \) ss ts" \] by auto thus ?thesis using \(2) 1(1,2) Fun by simp qed qed simp lemma timpls_transformable_to_match_abss'_nonempty_disj': fixes s t::"('a,'b,'c,'d) prot_term" and \::"('a,'b,'c,'d) prot_subst" and \::"('a,'b,'c,'d) prot_var \ 'c set set" assumes TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and \: "timpls_transformable_to TI t (s \ \)" and \: "match_abss' s t = Some \" and x: "x \ fv s" and t: "fv t = {}" and s: "\f \ funs_term s. \is_Abs f" "\x \ fv s. \a. \ x = \a\\<^sub>a\<^sub>b\<^sub>s" and a: "\ x = \a\\<^sub>a\<^sub>b\<^sub>s" shows "\b \ \ x. (b,a) \ (set TI)\<^sup>*" (is "?P \ x") proof - note 0 = match_abss'_subst_disj_nonempty[OF TI] have 1: "s \ \ \ timpl_closure t (set TI)" using timpls_transformable_to_iff_in_timpl_closure[OF TI] \ by blast have 2: "match_abss'_timpls_transform TI \ s t = Some \" using match_abss'_timpls_transformI[OF TI \ \ t s] by simp show "?P \ x" using 2 TI x t s a proof (induction TI \ s t arbitrary: \ rule: match_abss'_timpls_transform.induct) case (1 TI \ y c ts) thus ?case using match_abss'_timpls_transform_Var_inv[OF "1.prems"(1)] by auto next case (2 TI \ f ts g ss) obtain \ where fg: "f = g" "length ts = length ss" and \: "Some \ = those (map2 (match_abss'_timpls_transform TI \) ts ss)" "\ = fun_point_Union_list \" "\(t, s)\set (zip ts ss). \\'. match_abss'_timpls_transform TI \ t s = Some \'" using match_abss'_timpls_transform_Fun_inv[OF "2.prems"(1)] by blast have "(b,a) \ (set TI)\<^sup>*" when \': "\' \ set \" "b \ \' x" for \' b proof - obtain t' s' where t': "(t',s') \ set (zip ts ss)" "match_abss'_timpls_transform TI \ t' s' = Some \'" using those_map2_SomeD[OF \(1)[symmetric] \'(1)] by blast have *: "fv s' = {}" "\f \ funs_term t'. \is_Abs f" "\x \ fv t'. \a. \ x = \a\\<^sub>a\<^sub>b\<^sub>s" using "2.prems"(4-6) set_zip_leftD[OF t'(1)] set_zip_rightD[OF t'(1)] by (fastforce, fastforce, fastforce) have **: "x \ fv t'" using \'(2) match_abss'_timpl_transform_nonempty_is_fv[OF t'(2)] by blast show ?thesis using \'(2) "2.IH"[OF conjI[OF fg] t'(1) _ t'(2) "2.prems"(2) ** * "2.prems"(7)] by blast qed thus ?case using \(1) unfolding \(2) fun_point_Union_list_def by simp qed auto qed lemma timpls_transformable_to_match_abss'_nonempty_disj: fixes s t::"('a,'b,'c,'d) prot_term" and \::"('a,'b,'c,'d) prot_subst" and \::"('a,'b,'c,'d) prot_var \ 'c set set" assumes TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and \: "timpls_transformable_to TI t (s \ \)" and \: "match_abss' s t = Some \" and x: "x \ fv s" and t: "fv t = {}" and s: "\f \ funs_term s. \is_Abs f" "\x \ fv s. \a. \ x = \a\\<^sub>a\<^sub>b\<^sub>s" shows "\(ticl_abs TI ` \ x) \ {}" proof - have 0: "(a,b) \ (set TI)\<^sup>*" when y: "y \ fv s" "a \ \ y" "\ y = \b\\<^sub>a\<^sub>b\<^sub>s" for a b y using timpls_transformable_to_match_abss'_nonempty_disj'[OF TI \ \ y(1) t s y(3)] y(2) by blast obtain b where b: "\ x = \b\\<^sub>a\<^sub>b\<^sub>s" using x s(2) by blast have "b \ ticl_abs TI a" when a: "a \ \ x" for a using 0[OF x a b] unfolding ticl_abs_iff[OF TI] by blast thus ?thesis by blast qed lemma timpls_transformable_to_subst_subterm: fixes s t::"(('a,'b,'c,'d) prot_fun, 'v) term" and \ \::"(('a,'b,'c,'d) prot_fun, 'v) subst" assumes "timpls_transformable_to TI (t \ \) (t \ \)" and "s \ t" shows "timpls_transformable_to TI (s \ \) (s \ \)" using assms proof (induction "t \ \" "t \ \" arbitrary: t rule: timpls_transformable_to.induct) case (1 TI x y) thus ?case by (cases t) auto next case (2 TI f T g S) note prems = "2.prems" note hyps = "2.hyps"(2-) note IH = "2.hyps"(1) show ?case proof (cases "s = t") case False then obtain h U u where t: "t = Fun h U" "u \ set U" "s \ u" using prems(2) by (cases t) auto then obtain i where i: "i < length U" "U ! i = u" by (metis in_set_conv_nth) have "timpls_transformable_to TI (u \ \) (u \ \)" using t i prems(1) timpls_transformable_to_inv(2)[of TI h "U \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" h "U \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" i] by simp thus ?thesis using IH hyps t by auto qed (use prems in auto) qed simp_all lemma timpls_transformable_to_subst_match_case: assumes "timpls_transformable_to TI s (t \ \)" and "fv s = {}" and "\f \ funs_term t. \is_Abs f" and "distinct (fv_list t)" and "\x \ fv t. \a. \ x = \a\\<^sub>a\<^sub>b\<^sub>s" shows "\\. s = t \ \" using assms proof (induction s "t \ \" arbitrary: t rule: timpls_transformable_to.induct) case (2 TI f T g S) note prems = "2.prems" note hyps = "2.hyps"(2-) note IH = "2.hyps"(1) show ?case proof (cases t) case (Var x) then obtain a where a: "t \ \ = \a\\<^sub>a\<^sub>b\<^sub>s" using prems(5) by fastforce show ?thesis using hyps timpls_transformable_to_inv'[OF prems(1)[unfolded a]] unfolding Var by force next case (Fun h U) have g: "g = h" and S: "S = U \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using hyps unfolding Fun by simp_all note 0 = distinct_fv_list_Fun_param[OF prems(4)[unfolded Fun]] have 1: "\f \ funs_term u. \is_Abs f" when u: "u \ set U" for u using prems(3) u unfolding Fun by fastforce have 2: "fv t' = {}" when t': "t' \ set T" for t' using t' prems(2) by simp have 3: "\x \ fv u. \a. \ x = \a\\<^sub>a\<^sub>b\<^sub>s" when u: "u \ set U" for u using u prems(5) unfolding Fun by simp have "\is_Abs f" using prems(3) timpls_transformable_to_inv(3)[OF prems(1)[unfolded hyps[symmetric] S g]] unfolding Fun by auto hence f: "f = h" and T: "length T = length U" using prems(1) timpls_transformable_to_inv(1,3)[of TI f T h "U \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \"] unfolding Fun by fastforce+ define \ where "\ \ \i. if i < length T then SOME \. T ! i = U ! i \ \ else undefined" have "timpls_transformable_to TI (T ! i) (U ! i \ \)" when i: "i < length T" for i using prems(1)[unfolded Fun] T i timpls_transformable_to_inv(2)[of TI f T h "U \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" i] by auto hence "\\. T ! i = U ! i \ \" when i: "i < length T" for i using i T IH[OF _ _ _ 2 1 0 3, of "T ! i" "U ! i"] unfolding Fun g S by simp hence \: "T ! i = U ! i \ \ i" when i: "i < length T" for i using i someI2[of "\\. T ! i = U ! i \ \" _ "\\. T ! i = U ! i \ \"] unfolding \_def by fastforce define \ where "\ \ \x. if \i < length T. x \ fv (U ! i) then \ (SOME i. i < length T \ x \ fv (U ! i)) x else undefined" have "T ! i = U ! i \ \" when i: "i < length T" for i proof - have "j = i" when x: "x \ fv (U ! i)" and j: "j < length T" "x \ fv (U ! j)" for j x using x i j T distinct_fv_list_idx_fv_disjoint[OF prems(4)[unfolded Fun], of h U] by (metis (no_types, lifting) disjoint_iff_not_equal neqE term.dual_order.refl) hence "\ x = \ i x" when x: "x \ fv (U ! i)" for x using x i some_equality[of "\i. i < length T \ x \ fv (U ! i)" i] unfolding \_def by (metis (no_types, lifting)) thus ?thesis by (metis \ i term_subst_eq) qed hence "T = U \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" by (metis (no_types, lifting) T length_map nth_equalityI nth_map) hence "Fun f T = Fun f U \ \" by simp thus ?thesis using Fun f by fast qed qed simp_all lemma timpls_transformable_to_match_abss'_case: assumes "timpls_transformable_to TI s (t \ \)" and "fv s = {}" and "\f \ funs_term t. \is_Abs f" and "\x \ fv t. \a. \ x = \a\\<^sub>a\<^sub>b\<^sub>s" shows "\\. match_abss' t s = Some \" using assms proof (induction s "t \ \" arbitrary: t rule: timpls_transformable_to.induct) case (2 TI f T g S) note prems = "2.prems" note hyps = "2.hyps"(2-) note IH = "2.hyps"(1) show ?case proof (cases t) case (Var x) then obtain a where a: "t \ \ = \a\\<^sub>a\<^sub>b\<^sub>s" using prems(4) by fastforce thus ?thesis using timpls_transformable_to_inv'(4)[OF prems(1)[unfolded a]] by (metis (no_types) Var is_Abs_def term.sel(2) match_abss'.simps(1)) next case (Fun h U) have g: "g = h" and S: "S = U \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using hyps unfolding Fun by simp_all have 1: "\f \ funs_term u. \is_Abs f" when u: "u \ set U" for u using prems(3) u unfolding Fun by fastforce have 2: "fv t' = {}" when t': "t' \ set T" for t' using t' prems(2) by simp have 3: "\x \ fv u. \a. \ x = \a\\<^sub>a\<^sub>b\<^sub>s" when u: "u \ set U" for u using u prems(4) unfolding Fun by simp have "\is_Abs f" using prems(3) timpls_transformable_to_inv(3)[OF prems(1)[unfolded hyps[symmetric] S g]] unfolding Fun by auto hence f: "f = h" and T: "length T = length U" using prems(1) timpls_transformable_to_inv(1,3)[of TI f T h "U \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \"] unfolding Fun by fastforce+ define \ where "\ \ \i. if i < length T then SOME \. match_abss' (U ! i) (T ! i) = Some \ else undefined" have "timpls_transformable_to TI (T ! i) (U ! i \ \)" when i: "i < length T" for i using prems(1)[unfolded Fun] T i timpls_transformable_to_inv(2)[of TI f T h "U \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" i] by auto hence "\\. match_abss' (U ! i) (T ! i) = Some \" when i: "i < length T" for i using i T IH[OF _ _ _ 2 1 3, of "T ! i" "U ! i"] unfolding Fun g S by simp hence "match_abss' (U ! i) (T ! i) = Some (\ i)" when i: "i < length T" for i using i someI2[of "\\. match_abss' (U ! i) (T ! i) = Some \" _ "\\. match_abss' (U ! i) (T ! i) = Some \"] unfolding \_def by fastforce thus ?thesis using match_abss'_FunI[OF _ T] unfolding Fun f by auto qed qed simp_all lemma timpls_transformable_to_match_abss_case: assumes TI: "set TI = {(a,b) \ (set TI)\<^sup>+. a \ b}" and "timpls_transformable_to TI s (t \ \)" and "fv s = {}" and "\f \ funs_term t. \is_Abs f" and "\x \ fv t. \a. \ x = \a\\<^sub>a\<^sub>b\<^sub>s" shows "\\. match_abss OCC TI t s = Some \" proof - obtain \ where \: "match_abss' t s = Some \" using timpls_transformable_to_match_abss'_case[OF assms(2-)] by blast show ?thesis using \ timpls_transformable_to_match_abss'_nonempty_disj[OF assms(1,2) \ _ assms(3-5)] unfolding match_abss_def by simp qed private lemma transaction_check_variant_soundness_aux3: fixes T FP S C xs OCC negs poss as defines "\ \ \\ x. if fst x = TAtom Value then (absc \ \) x else Var x" and "C \ unlabel (transaction_checks T)" and "S \ unlabel (transaction_strand T)" and "ts \ trms_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_receive T))" and "xs \ filter (\x. x \ set (transaction_fresh T) \ fst x = TAtom Value) (fv_list\<^sub>s\<^sub>s\<^sub>t S)" assumes TI0: "\(a,b) \ set TI. \(c,d) \ set TI. b = c \ a \ d \ (a,d) \ set TI" "\(a,b) \ set TI. a \ b" and OCC: "\t \ set FP. \a. Abs a \ funs_term t \ a \ set OCC" and FP_ground: "ground (set FP)" and x: "x \ set xs" and xs: "\x. x \ set xs \ \ x \ set OCC" "\x. x \ set xs \ poss x \ \ x" "\x. x \ set xs \ \ x \ negs x = {}" "\x. x \ set xs \ \ x = {}" and ts: "\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T). intruder_synth_mod_timpls FP TI (t \ \ \)" "\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T). \f \ funs_term t. \is_Abs f" "\x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)). fst x = TAtom Value" and C: "\a x s. \a: Var x \ Fun (Set s) []\ \ set C \ s \ \ x" "\x s. \Var x not in Fun (Set s) []\ \ set C \ s \ \ x" and \: "synth_abs_substs_constrs_rel FP OCC TI ts \" shows "\ x \ \ x" proof - note defs = assms(1-5) note TI = trancl_eqI'[OF TI0] have \x0: "\ x \ set OCC" "poss x \ \ x" "\ x \ negs x = {}" using x xs by (blast,blast,blast) have ts0: "\t \ set ts. intruder_synth_mod_timpls FP TI (t \ \ \)" using ts(1) trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t unfolding ts_def by blast have ts1: "\Fun (Abs n) S \\<^sub>s\<^sub>e\<^sub>t set ts" for n S using ts(2) funs_term_Fun_subterm' unfolding ts_def trms_transaction_unfold trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t[symmetric] is_Abs_def by fastforce have ts2: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (set ts). fst x = TAtom Value" using ts(3) unfolding ts_def trms_transaction_unfold trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t[symmetric] by fastforce show ?thesis using \ ts0 ts1 ts2 proof (induction rule: synth_abs_substs_constrs_rel.induct) case (SolvePrivConstNotin c) hence "intruder_synth_mod_timpls FP TI (Fun c [])" by force hence "list_ex (\t. timpls_transformable_to TI t (Fun c [])) FP" using SolvePrivConstNotin.hyps(1,2) by simp then obtain t where t: "t \ set FP" "timpls_transformable_to TI t (Fun c [])" unfolding list_ex_iff by blast have "\is_Abs c" using SolvePrivConstNotin.prems(2)[of _ "[]"] by (metis in_subterms_Union is_Abs_def list.set_intros(1)) hence "t = Fun c []" using t(2) timpls_transformable_to_inv[of TI] by (cases t) auto thus ?case using t(1) SolvePrivConstNotin.hyps(3) by fast next case (SolveValueVar \1 y) have "list_ex (\t. timpls_transformable_to TI t \\ y\\<^sub>a\<^sub>b\<^sub>s) FP" using SolveValueVar.prems(1-3) unfolding \_def by simp then obtain t where t: "t \ set FP" "timpls_transformable_to TI t \\ y\\<^sub>a\<^sub>b\<^sub>s" unfolding list_ex_iff by blast obtain a where a: "t = \a\\<^sub>a\<^sub>b\<^sub>s" "a = \ y \ (a, \ y) \ set TI" proof - obtain ft tst where ft: "t = Fun ft tst" using t(2) timpls_transformable_to_inv_Var(1)[of TI _ "\\ y\\<^sub>a\<^sub>b\<^sub>s"] by (cases t) auto have "tst = []" "is_Abs ft" "the_Abs ft = \ y \ (the_Abs ft, \ y) \ set TI" using timpls_transformable_to_inv'(2,4,5)[OF t(2)[unfolded ft]] by (simp, force, force) thus thesis using that[of "the_Abs ft"] ft by force qed have "a \ set OCC" using t(1)[unfolded a(1)] OCC by auto thus ?case using \x0(1) t(1)[unfolded a(1)] a(2) unfolding SolveValueVar.hyps(1) ticl_abss_def ticl_abs_def by force next case (SolveComposed g us \ \1 \2) show ?case proof (cases "\t \ set us. intruder_synth_mod_timpls FP TI (t \ \ \)") case True hence "\ x \ \2 x" using SolveComposed.IH SolveComposed.prems(2,3) distinct_fv_list_Fun_param[of g us] by auto thus ?thesis unfolding fun_point_union_def by simp next case False hence "list_ex (\t. timpls_transformable_to TI t (Fun g us \ \ \)) FP" using SolveComposed.prems(1) intruder_synth_mod_timpls.simps(2)[of FP TI g "us \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \"] unfolding list_all_iff by auto then obtain t where t: "t \ set FP" "timpls_transformable_to TI t (Fun g us \ \ \)" unfolding list_ex_iff by blast have t_ground: "fv t = {}" using t(1) FP_ground by simp have g_no_abs: "\is_Abs f" when f: "f \ funs_term (Fun g us)" for f proof - obtain fts where "Fun f fts \ Fun g us" using funs_term_Fun_subterm[OF f] by blast thus ?thesis using SolveComposed.prems(2)[of _ fts] by (cases f) auto qed have g_\_abs: "\a. \ \ y = \a\\<^sub>a\<^sub>b\<^sub>s" when y: "y \ fv (Fun g us)" for y using y SolveComposed.prems(3) unfolding \_def by fastforce obtain \' where \': "match_abss OCC TI (Fun g us) t = Some \'" using g_no_abs g_\_abs timpls_transformable_to_match_abss_case[OF TI t(2) t_ground ] by blast let ?h1 = "\\ x. if x \ fv (Fun g us) then \ x else set OCC" let ?h2 = "\\ x. \(ticl_abs TI ` \ x)" obtain \'' where \'': "match_abss' (Fun g us) t = Some \''" "\' = ?h1 (?h2 \'')" "\x \ fv (Fun g us). \' x \ {} \ \' x \ {}" using match_abssD[OF \'] by blast have \'_\: "\' \ \" using t(1) \' SolveComposed.hyps(3) by metis have "\ x \ ticl_abs TI a" when a: "a \ \'' x" and x_in_g: "x \ fv (Fun g us)" for a proof - have "fst x = TAtom Value" using x_in_g SolveComposed.prems(3) by auto hence "\ \ x = \\ x\\<^sub>a\<^sub>b\<^sub>s" unfolding \_def by simp hence "(a, \ x) \ (set TI)\<^sup>*" using timpls_transformable_to_match_abss'_nonempty_disj'[ OF TI t(2) \''(1) x_in_g t_ground] g_no_abs g_\_abs a by fastforce thus "\ x \ ticl_abs TI a" unfolding ticl_abs_iff[OF TI] by force qed hence "\ x \ \' x" when x_in_g: "x \ fv (Fun g us)" using \''(2,3) x_in_g unfolding \''(1) by simp hence "\ x \ \' x" using match_abss_OCC_if_not_fv[OF \'] \x0(1) by blast hence "\ x \ \1 x" using \'_\ \x0(1) unfolding SolveComposed.hyps(4,5) fun_point_Union_def by auto thus ?thesis unfolding fun_point_union_def by simp qed qed (auto simp add: \x0 fun_point_Inter_def) qed private lemma transaction_check_variant_soundness_aux4: fixes T FP S C xs OCC negs poss as defines "\ \ \\ x. if fst x = TAtom Value then (absc \ \) x else Var x" and "C \ unlabel (transaction_checks T)" and "S \ unlabel (transaction_strand T)" and "xas \ (the_Abs \ the_Fun) ` set (filter (\t. is_Fun t \ is_Abs (the_Fun t)) FP)" and "ts \ trms_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_receive T))" and "xs \ filter (\x. x \ set (transaction_fresh T) \ fst x = TAtom Value) (fv_list\<^sub>s\<^sub>s\<^sub>t S)" and "poss \ transaction_poschecks_comp C" and "negs \ transaction_negchecks_comp C" and "as \ map (\x. (x, set (filter (\ab. poss x \ ab \ negs x \ ab = {}) OCC))) xs" and "f \ \x. case List.find (\p. fst p = x) as of Some p \ snd p | None \ {}" assumes T_adm: "admissible_transaction' T" and TI0: "\(a,b) \ set TI. \(c,d) \ set TI. b = c \ a \ d \ (a,d) \ set TI" "\(a,b) \ set TI. a \ b" and OCC: "\t \ set FP. \a. Abs a \ funs_term t \ a \ set OCC" and FP_ground: "ground (set FP)" and FP_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" and "x \ set xs" and "\x. x \ set xs \ \ x \ set OCC" and "\x. x \ set xs \ poss x \ \ x" and "\x. x \ set xs \ \ x \ negs x = {}" and "\x. x \ set xs \ \ x = {}" and "\t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T). intruder_synth_mod_timpls FP TI (t \ \ \)" and "\a x s. \a: Var x \ Fun (Set s) []\ \ set C \ s \ \ x" and "\x s. \Var x not in Fun (Set s) []\ \ set C \ s \ \ x" shows "\ x \ synth_abs_substs_constrs (FP,OCC,TI) T x" proof - let ?FPT = "(FP,OCC,TI)" let ?P = "\s u. let \ = mgu s u in \ \ None \ (\x \ fv s. is_Fun (the \ x) \ is_Abs (the_Fun (the \ x)))" define \0 where "\0 \ \x. if fst x = TAtom Value \ x \ fv_transaction T \ x \ set (transaction_fresh T) then {ab \ set OCC. poss x \ ab \ negs x \ ab = {}} else {}" define g where "g \ \x. set (filter (\ab. poss x \ ab \ negs x \ ab = {}) OCC)" define gs where "gs \ map (\x. (x, g x)) xs" note defs = assms(3-10) \0_def note assm = assms(11-)[unfolded defs] have ts0: "\t \ set ts. wf\<^sub>t\<^sub>r\<^sub>m t" using admissible_transaction_is_wellformed_transaction(4)[OF T_adm] unfolding admissible_transaction_terms_def wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[symmetric] ts_def trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t[symmetric] trms_transaction_unfold by fast have ts1: "\t \ set ts. \f \ funs_term t. \is_Abs f" using protocol_transactions_no_abss[OF T_adm] funs_term_Fun_subterm trms\<^sub>s\<^sub>s\<^sub>t_unlabel_prefix_subset(1) unfolding ts_def trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t[symmetric] is_Abs_def transaction_strand_def by (metis (no_types, opaque_lifting) in_subterms_Union in_subterms_subset_Union subset_iff) have ts2: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (set ts). fst x = TAtom Value" using admissible_transaction_Value_vars[OF T_adm] wellformed_transaction_send_receive_fv_subset(1)[ OF admissible_transaction_is_wellformed_transaction(1)[OF T_adm]] unfolding ts_def trms_transaction_unfold trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t[symmetric] \\<^sub>v_TAtom''(2) by fastforce have "f x = \0 x" for x proof (cases "fst x = Var Value \ x \ fv_transaction T \ x \ set (transaction_fresh T)") case True hence "\0 x = {ab \ set OCC. poss x \ ab \ negs x \ ab = {}}" unfolding \0_def by argo thus ?thesis using True transaction_check_variant_soundness_aux0[OF S_def xs_def, of x] transaction_check_variant_soundness_aux1[ OF C_def S_def xs_def poss_def negs_def as_def f_def, of x] by simp next case False hence 0: "\0 x = {}" unfolding \0_def by argo have "x \ set xs" using False fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] unfolding xs_def S_def by fastforce hence "List.find (\p. fst p = x) gs = None" using find_None_iff[of "\p. fst p = x" gs] unfolding gs_def by simp hence "f x = {}" unfolding f_def as_def gs_def g_def by force thus ?thesis using 0 by simp qed thus ?thesis using synth_abs_substs_constrs_rel_if_synth_abs_substs_constrs[ OF _ FP_ground FP_wf, of T, unfolded trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t ts_def[symmetric], OF ts0] transaction_check_variant_soundness_aux3[ OF TI0 OCC FP_ground assm(7-11), of "synth_abs_substs_constrs ?FPT T", unfolded trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t ts_def[symmetric], OF assm(12)[unfolded \_def trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t ts_def[symmetric]] ts1 ts2 assm(13-)[unfolded C_def]] unfolding defs synth_abs_substs_constrs_def Let_def by blast qed private lemma transaction_check_variant_soundness_aux5: fixes FP OCC TI T S C defines "msgcs \ \x a. a \ synth_abs_substs_constrs (FP,OCC,TI) T x" and "S \ unlabel (transaction_strand T)" and "C \ unlabel (transaction_checks T)" and "xs \ filter (\x. x \ set (transaction_fresh T) \ fst x = TAtom Value) (fv_list\<^sub>s\<^sub>s\<^sub>t S)" and "poss \ transaction_poschecks_comp C" and "negs \ transaction_negchecks_comp C" assumes T_adm: "admissible_transaction' T" and TI: "\(a,b) \ set TI. \(c,d) \ set TI. b = c \ a \ d \ (a,d) \ set TI" "\(a,b) \ set TI. a \ b" and OCC: "\t \ set FP. \a. Abs a \ funs_term t \ a \ set OCC" and FP: "ground (set FP)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" and \: "\ \ abs_substs_fun ` set (abs_substs_set xs OCC poss negs (\_ _. True))" "transaction_check_pre (FP,OCC,TI) T \" shows "\ \ abs_substs_fun ` set (abs_substs_set xs OCC poss negs msgcs)" proof - have 0: "\ x \ set OCC" "poss x \ \ x" "\ x \ negs x = {}" when x: "x \ set xs" for x using abs_substs_abss_bounded[OF \(1) x] by simp_all have 1: "\ x = {}" when x: "x \ set xs" for x by (rule abs_substs_abss_bounded'[OF \(1) x]) have 2: "msgcs x (\ x)" when x: "x \ set xs" for x using 0 1 x transaction_check_variant_soundness_aux4[OF T_adm TI OCC FP, of x \] transaction_check_pre_ReceiveE[OF \(2)] transaction_check_pre_InSetE[OF \(2)] transaction_check_pre_NotInSetE[OF \(2)] unfolding msgcs_def xs_def C_def S_def negs_def poss_def by fast show ?thesis using abs_substs_has_abs[of xs \ OCC poss negs msgcs] 0 1 2 by blast qed theorem transaction_check_variant_soundness: assumes P_adm: "\T \ set P. admissible_transaction' T" and TI: "\(a,b) \ set TI. \(c,d) \ set TI. b = c \ a \ d \ (a,d) \ set TI" "\(a,b) \ set TI. a \ b" and OCC: "\t \ set FP. \a. Abs a \ funs_term t \ a \ set OCC" and FP: "ground (set FP)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" and T_in: "T \ set P" and T_check: "transaction_check_coverage_rcv (FP,OCC,TI) T" shows "transaction_check (FP,OCC,TI) T" proof - have 0: "admissible_transaction' T" using P_adm T_in by blast show ?thesis using T_check transaction_check_variant_soundness_aux5[OF 0 TI OCC FP] unfolding transaction_check_def transaction_check'_def transaction_check_coverage_rcv_def transaction_check_comp_def list_all_iff Let_def by force qed end end subsection \Automatic Fixed-Point Computation\ context stateful_protocol_model begin fun reduce_fixpoint' where "reduce_fixpoint' FP _ [] = FP" | "reduce_fixpoint' FP TI (t#M) = ( let FP' = List.removeAll t FP in if intruder_synth_mod_timpls FP' TI t then FP' else reduce_fixpoint' FP TI M)" definition reduce_fixpoint where "reduce_fixpoint FP TI \ let f = \FP. reduce_fixpoint' FP TI FP in while (\M. set (f M) \ set M) f FP" definition compute_fixpoint_fun' where "compute_fixpoint_fun' P (n::nat option) enable_traces \ S0 \ let P' = map add_occurs_msgs P; sy = intruder_synth_mod_timpls; FP' = \S. fst (fst S); TI' = \S. snd (fst S); OCC' = \S. remdups ( (map (\t. the_Abs (the_Fun (args t ! 1))) (filter (\t. is_Fun t \ the_Fun t = OccursFact) (FP' S)))@ (map snd (TI' S))); equal_states = \S S'. set (FP' S) = set (FP' S') \ set (TI' S) = set (TI' S'); trace' = \S. snd S; close = \M f. let g = remdups \ f in while (\A. set (g A) \ set A) g M; close' = \M f. let g = remdups \ f in while (\A. set (g A) \ set A) g M; trancl_minus_refl = \TI. let aux = \ts p. map (\q. (fst p,snd q)) (filter ((=) (snd p) \ fst) ts) in filter (\p. fst p \ snd p) (close' TI (\ts. concat (map (aux ts) ts)@ts)); snd_Ana = \N M TI. let N' = filter (\t. \k \ set (fst (Ana t)). sy M TI k) N in filter (\t. \sy M TI t) (concat (map (\t. filter (\s. s \ set (snd (Ana t))) (args t)) N')); Ana_cl = \FP TI. close FP (\M. (M@snd_Ana M M TI)); TI_cl = \FP TI. close FP (\M. (M@filter (\t. \sy M TI t) (concat (map (\m. concat (map (\(a,b). \a --\ b\\m\) TI)) M)))); Ana_cl' = \FP TI. let K = \t. set (fst (Ana t)); flt = \M t. (\k \ K t. \sy M TI k) \ (\k \ K t. \f \ funs_term k. is_Abs f); N = \M. comp_timpl_closure_list (filter (flt M) M) TI in close FP (\M. M@snd_Ana (N M) M TI); \' = \S. \ (FP' S, OCC' S, TI' S); result = \S T \. let not_fresh = \x. x \ set (transaction_fresh T); xs = filter not_fresh (fv_list\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T))); u = \\ x. absdbupd (unlabel (transaction_strand T)) x (\ x) in (remdups (filter (\t. \sy (FP' S) (TI' S) t) (concat (map (\ts. the_msgs ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (absc \ u \)) (filter is_Send (unlabel (transaction_send T)))))), remdups (filter (\s. fst s \ snd s) (map (\x. (\ x, u \ x)) xs))); result_tuple = \S T \. (result S T (abs_substs_fun \), if enable_traces then \ else []); update_state = \S. if list_ex (\t. is_Fun t \ is_Attack (the_Fun t)) (FP' S) then S else let results = map (\T. map (result_tuple S T) (\' S T)) P'; newtrace_flt = (\n. let x = map fst (results ! n); y = map fst x; z = map snd x in set (concat y) - set (FP' S) \ {} \ set (concat z) - set (TI' S) \ {}); trace = if enable_traces then trace' S@[concat (map (\i. map (\a. (i,snd a)) (results ! i)) (filter newtrace_flt [0..x. fst x \ snd x) (concat (map snd U)@TI' S))), trace); W = ((Ana_cl (TI_cl (FP' V) (TI' V)) (TI' V), trancl_minus_refl (TI' V)), trace' V) in if \equal_states W S then W else ((Ana_cl' (FP' W) (TI' W), TI' W), trace' W); S = ((\h. case n of None \ while (\S. \equal_states S (h S)) h | Some m \ h ^^ m) update_state S0) in ((reduce_fixpoint (FP' S) (TI' S), OCC' S, TI' S), trace' S)" definition compute_fixpoint_fun where "compute_fixpoint_fun P \ let P' = remdups (filter (\T. transaction_updates T \ [] \ transaction_send T \ []) P); f = (\FPT T. let msgcs = synth_abs_substs_constrs FPT T in transaction_check_comp (\x a. a \ msgcs x) FPT T) in fst (compute_fixpoint_fun' P' None False f (([],[]),[]))" definition compute_fixpoint_with_trace where "compute_fixpoint_with_trace P \ compute_fixpoint_fun' P None True (transaction_check_comp (\_ _. True)) (([],[]),[])" definition compute_fixpoint_from_trace where "compute_fixpoint_from_trace P trace \ let P' = map add_occurs_msgs P; \ = \FPT T. let pre_check = transaction_check_pre FPT T; \s = map snd (filter (\(i,as). P' ! i = T) (concat trace)) in filter (\\. pre_check (abs_substs_fun \)) \s; f = compute_fixpoint_fun' \ map (nth P); g = \L FPT. fst ((f L (Some 1) False \ ((fst FPT,snd (snd FPT)),[]))) in fold g (map (map fst) trace) ([],[],[])" definition compute_reduced_attack_trace where "compute_reduced_attack_trace P trace \ let attack_in_fixpoint = list_ex (\t. \f \ funs_term t. is_Attack f) \ fst; is_attack_trace = attack_in_fixpoint \ compute_fixpoint_from_trace P; trace' = let is_attack_transaction = list_ex is_Fun_Attack \ concat \ map the_msgs \ filter is_Send \ unlabel \ transaction_send; trace' = if trace = [] then [] else butlast trace@[filter (is_attack_transaction \ nth P \ fst) (last trace)] in trace'; iter = \trace_prev trace_rest elem (prev,rest). let next = if is_attack_trace (trace_prev@(prev@rest)#trace_rest) then prev else prev@[elem] in (next, tl rest); iter' = \trace_part (trace_prev,trace_rest). let updated = foldr (iter trace_prev (tl trace_rest)) trace_part ([],tl (rev trace_part)) in (trace_prev@[rev (fst updated)], tl trace_rest); reduced_trace = fst (fold iter' trace' ([],trace')) in concat reduced_trace" end subsection \Locales for Protocols Proven Secure through Fixed-Point Coverage\ type_synonym ('f,'a,'s,'l) fixpoint_triple = "('f,'a,'s,'l) prot_term list \ 's set list \ ('s set \ 's set) list" context stateful_protocol_model begin definition "attack_notin_fixpoint (FPT::('fun,'atom,'sets,'lbl) fixpoint_triple) \ list_all (\t. \f \ funs_term t. \is_Attack f) (fst FPT)" definition "protocol_covered_by_fixpoint (FPT::('fun,'atom,'sets,'lbl) fixpoint_triple) P \ list_all (transaction_check FPT) (filter (\T. transaction_updates T \ [] \ transaction_send T \ []) (map add_occurs_msgs P))" definition "protocol_covered_by_fixpoint_coverage_rcv (FPT::('fun,'atom,'sets,'lbl) fixpoint_triple) P \ list_all (transaction_check_coverage_rcv FPT) (filter (\T. transaction_updates T \ [] \ transaction_send T \ []) (map add_occurs_msgs P))" definition "analyzed_fixpoint (FPT::('fun,'atom,'sets,'lbl) fixpoint_triple) \ let (FP, _, TI) = FPT in analyzed_closed_mod_timpls FP TI" definition "wellformed_protocol_SMP_set (P::('fun,'atom,'sets,'lbl) prot) N \ has_all_wt_instances_of \ (\T \ set P. trms_transaction T) (set N) \ comp_tfr\<^sub>s\<^sub>e\<^sub>t arity Ana \ (set N) \ list_all (\T. list_all (comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ Pair) (unlabel (transaction_strand T))) P" (* TODO: try to avoid checking the "list_all is_*" conditions here... *) definition "wellformed_protocol'' (P::('fun,'atom,'sets,'lbl) prot) N \ let f = \T. transaction_fresh T = [] \ transaction_updates T \ [] \ transaction_send T \ [] in list_all (\T. list_all is_Receive (unlabel (transaction_receive T)) \ list_all is_Check_or_Assignment (unlabel (transaction_checks T)) \ list_all is_Update (unlabel (transaction_updates T)) \ list_all is_Send (unlabel (transaction_send T))) P \ list_all admissible_transaction (filter f P) \ wellformed_protocol_SMP_set P N" definition "wellformed_protocol' (P::('fun,'atom,'sets,'lbl) prot) N \ wellformed_protocol'' P N \ has_initial_value_producing_transaction P" definition "wellformed_protocol (P::('fun,'atom,'sets,'lbl) prot) \ let f = \M. remdups (concat (map subterms_list M@map (fst \ Ana) M)); N0 = remdups (concat (map (trms_list\<^sub>s\<^sub>s\<^sub>t \ unlabel \ transaction_strand) P)); N = while (\A. set (f A) \ set A) f N0 in wellformed_protocol' P N" definition "wellformed_fixpoint' (FPT::('fun,'atom,'sets,'lbl) fixpoint_triple) \ let (FP, OCC, TI) = FPT; OCC' = set OCC in list_all (\t. wf\<^sub>t\<^sub>r\<^sub>m' arity t \ fv t = {}) FP \ list_all (\a. a \ OCC') (map snd TI) \ list_all (\t. \f \ funs_term t. is_Abs f \ the_Abs f \ OCC') FP" definition "wellformed_term_implication_graph (FPT::('fun,'atom,'sets,'lbl) fixpoint_triple) \ let (_, _, TI) = FPT in list_all (\(a,b). list_all (\(c,d). b = c \ a \ d \ List.member TI (a,d)) TI) TI \ list_all (\p. fst p \ snd p) TI" definition "wellformed_fixpoint (FPT::('fun,'atom,'sets,'lbl) fixpoint_triple) \ wellformed_fixpoint' FPT \ wellformed_term_implication_graph FPT" lemma wellformed_protocol_SMP_set_mono: assumes "wellformed_protocol_SMP_set P S" and "set P' \ set P" shows "wellformed_protocol_SMP_set P' S" using assms unfolding wellformed_protocol_SMP_set_def comp_tfr\<^sub>s\<^sub>e\<^sub>t_def has_all_wt_instances_of_def wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s'_def list_all_iff by fast lemma wellformed_protocol''_mono: assumes "wellformed_protocol'' P S" and "set P' \ set P" shows "wellformed_protocol'' P' S" using assms wellformed_protocol_SMP_set_mono[of P S P'] unfolding wellformed_protocol''_def list_all_iff by auto lemma wellformed_protocol'_mono: assumes "wellformed_protocol' P S" and "set P' \ set P" and "has_initial_value_producing_transaction P'" shows "wellformed_protocol' P' S" using assms wellformed_protocol_SMP_set_mono[of P S P'] wellformed_protocol''_mono unfolding wellformed_protocol'_def by blast lemma protocol_covered_by_fixpoint_if_protocol_covered_by_fixpoint_coverage_rcv: assumes P: "wellformed_protocol'' P P_SMP" and FPT: "wellformed_fixpoint FPT" and covered: "protocol_covered_by_fixpoint_coverage_rcv FPT P" shows "protocol_covered_by_fixpoint FPT P" proof - obtain FP OCC TI where FPT': "FPT = (FP,OCC,TI)" by (metis surj_pair) note defs = FPT' wellformed_protocol''_def wellformed_fixpoint_def wellformed_fixpoint'_def wellformed_term_implication_graph_def Let_def wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[symmetric] wf\<^sub>t\<^sub>r\<^sub>m_code[symmetric] member_def case_prod_unfold list_all_iff let ?f = "\T. transaction_fresh T = [] \ transaction_updates T \ [] \ transaction_send T \ []" have TI: "\(a,b) \ set TI. \(c,d) \ set TI. b = c \ a \ d \ (a,d) \ set TI" "\(a,b) \ set TI. a \ b" and OCC: "\t \ set FP. \a. Abs a \ funs_term t \ a \ set OCC" and FP: "ground (set FP)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" using FPT unfolding defs by (simp, simp, fastforce, simp, simp) have P_adm: "\T \ set (filter ?f (map add_occurs_msgs P)). admissible_transaction' T" using P add_occurs_msgs_admissible_occurs_checks(1)[OF admissible_transactionE'(1)] unfolding defs add_occurs_msgs_updates_send_filter_iff'[of P, symmetric] by auto show ?thesis using covered transaction_check_variant_soundness[OF P_adm TI OCC FP] unfolding protocol_covered_by_fixpoint_def protocol_covered_by_fixpoint_coverage_rcv_def FPT' list_all_iff by fastforce qed lemma protocol_covered_by_fixpoint_if_protocol_covered_by_fixpoint_coverage_rcv': assumes P: "wellformed_protocol'' P P_SMP" and P': "set P' \ set P" and FPT: "wellformed_fixpoint FPT" and covered: "protocol_covered_by_fixpoint_coverage_rcv FPT P'" shows "protocol_covered_by_fixpoint FPT P'" using protocol_covered_by_fixpoint_if_protocol_covered_by_fixpoint_coverage_rcv[OF _ FPT covered] wellformed_protocol''_mono[OF P P'] by simp lemma protocol_covered_by_fixpoint_trivial_case: assumes "list_all (\T. transaction_updates T = [] \ transaction_send T = []) (map add_occurs_msgs P)" shows "protocol_covered_by_fixpoint FPT P" using assms by (simp add: list_all_iff transaction_check_trivial_case protocol_covered_by_fixpoint_def) lemma protocol_covered_by_fixpoint_empty[simp]: "protocol_covered_by_fixpoint FPT []" by (simp add: protocol_covered_by_fixpoint_def) lemma protocol_covered_by_fixpoint_Cons[simp]: "protocol_covered_by_fixpoint FPT (T#P) \ transaction_check FPT (add_occurs_msgs T) \ protocol_covered_by_fixpoint FPT P" using transaction_check_trivial_case[of "add_occurs_msgs T"] unfolding protocol_covered_by_fixpoint_def case_prod_unfold by simp lemma protocol_covered_by_fixpoint_append[simp]: "protocol_covered_by_fixpoint FPT (P1@P2) \ protocol_covered_by_fixpoint FPT P1 \ protocol_covered_by_fixpoint FPT P2" by (simp add: protocol_covered_by_fixpoint_def case_prod_unfold) lemma protocol_covered_by_fixpoint_I1[intro]: assumes "list_all (protocol_covered_by_fixpoint FPT) P" shows "protocol_covered_by_fixpoint FPT (concat P)" using assms by (auto simp add: protocol_covered_by_fixpoint_def list_all_iff) lemma protocol_covered_by_fixpoint_I2[intro]: assumes "protocol_covered_by_fixpoint FPT P1" and "protocol_covered_by_fixpoint FPT P2" shows "protocol_covered_by_fixpoint FPT (P1@P2)" using assms by (auto simp add: protocol_covered_by_fixpoint_def) lemma protocol_covered_by_fixpoint_I3: assumes "\T \ set P. \\::('fun,'atom,'sets,'lbl) prot_var \ 'sets set. transaction_check_pre FPT (add_occurs_msgs T) \ \ transaction_check_post FPT (add_occurs_msgs T) \" shows "protocol_covered_by_fixpoint FPT P" using assms unfolding protocol_covered_by_fixpoint_def transaction_check_def transaction_check'_def transaction_check_comp_def list_all_iff Let_def case_prod_unfold Product_Type.fst_conv Product_Type.snd_conv by fastforce lemmas protocol_covered_by_fixpoint_intros = protocol_covered_by_fixpoint_I1 protocol_covered_by_fixpoint_I2 protocol_covered_by_fixpoint_I3 lemma prot_secure_if_prot_checks: fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple" assumes attack_notin_fixpoint: "attack_notin_fixpoint FP_OCC_TI" and transactions_covered: "protocol_covered_by_fixpoint FP_OCC_TI P" and analyzed_fixpoint: "analyzed_fixpoint FP_OCC_TI" and wellformed_protocol: "wellformed_protocol' P N" and wellformed_fixpoint: "wellformed_fixpoint FP_OCC_TI" shows "\\ \ reachable_constraints P. \\. constraint_model \ (\@[(l, send\[attack\n\]\)])" (is "?secure P") proof - define FP where "FP \ let (FP,_,_) = FP_OCC_TI in FP" define OCC where "OCC \ let (_,OCC,_) = FP_OCC_TI in OCC" define TI where "TI \ let (_,_,TI) = FP_OCC_TI in TI" define f where "f \ \T::('fun,'atom,'sets,'lbl) prot_transaction. transaction_fresh T = [] \ transaction_updates T \ [] \ transaction_send T \ []" define g where "g \ \T::('fun,'atom,'sets,'lbl) prot_transaction. transaction_fresh T = [] \ list_ex (\a. is_Update (snd a) \ is_Send (snd a)) (transaction_strand T)" note wellformed_prot_defs = wellformed_protocol'_def wellformed_protocol''_def wellformed_protocol_SMP_set_def have attack_notin_FP: "attack\n\ \ set FP" using attack_notin_fixpoint[unfolded attack_notin_fixpoint_def] unfolding list_all_iff FP_def by force have 1: "\(a,b) \ set TI. \(c,d) \ set TI. b = c \ a \ d \ (a,d) \ set TI" using wellformed_fixpoint unfolding wellformed_fixpoint_def wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[symmetric] Let_def TI_def list_all_iff member_def case_prod_unfold wellformed_term_implication_graph_def by auto have 0: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set FP)" and 2: "\(a,b) \ set TI. a \ b" and 3: "snd ` set TI \ set OCC" and 4: "\t \ set FP. \f \ funs_term t. is_Abs f \ f \ Abs ` set OCC" and 5: "ground (set FP)" using wellformed_fixpoint unfolding wellformed_fixpoint_def wf\<^sub>t\<^sub>r\<^sub>m_code[symmetric] is_Abs_def the_Abs_def list_all_iff Let_def case_prod_unfold set_map FP_def OCC_def TI_def wellformed_fixpoint'_def wellformed_term_implication_graph_def by (fast, fast, blast, fastforce, simp) have 8: "finite (set N)" and 9: "has_all_wt_instances_of \ (\T \ set (filter g P). trms_transaction T) (set N)" and 10: "tfr\<^sub>s\<^sub>e\<^sub>t (set N)" and 11: "\T \ set (filter f P). list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" and 12: "\T \ set (filter f P). admissible_transaction T" using wellformed_protocol[unfolded wellformed_prot_defs] tfr\<^sub>s\<^sub>e\<^sub>t_if_comp_tfr\<^sub>s\<^sub>e\<^sub>t[of "set N"] unfolding Let_def list_all_iff wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[symmetric] tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p[symmetric] has_all_wt_instances_of_def f_def[symmetric] by (fast, fastforce, fast, fastforce, fast) have 13: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set N)" using wellformed_protocol[unfolded wellformed_prot_defs] finite_SMP_representationD unfolding wf\<^sub>t\<^sub>r\<^sub>m_code[symmetric] wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s'_def comp_tfr\<^sub>s\<^sub>e\<^sub>t_def list_all_iff Let_def by fast have 14: "has_initial_value_producing_transaction (filter g P)" using wellformed_protocol has_initial_value_producing_transaction_update_send_ex_filter unfolding wellformed_protocol'_def Let_def g_def by blast note TI0 = trancl_eqI'[OF 1 2] have "analyzed (timpl_closure_set (set FP) (set TI))" using analyzed_fixpoint[unfolded analyzed_fixpoint_def] analyzed_closed_mod_timpls_is_analyzed_timpl_closure_set[OF TI0 0] unfolding FP_def TI_def by force note FP0 = this 0 5 note OCC0 = funs_term_OCC_TI_subset(1)[OF 4 3] timpl_closure_set_supset'[OF funs_term_OCC_TI_subset(2)[OF 4 3]] note M0 = 9 8 10 13 have "f T \ g T" when T: "T \ set P" for T proof - have *: "list_all stateful_strand_step.is_Receive (unlabel (transaction_receive T))" "list_all is_Check_or_Assignment (unlabel (transaction_checks T))" "list_all is_Update (unlabel (transaction_updates T))" "list_all stateful_strand_step.is_Send (unlabel (transaction_send T))" using T wellformed_protocol unfolding wellformed_protocol_def wellformed_prot_defs Let_def list_all_iff by (fast, fast, fast, fast) show ?thesis using transaction_updates_send_ex_iff[OF *] unfolding f_def g_def by (metis (no_types, lifting) list_ex_cong) qed hence 15: "\T \ set (filter g P). list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" and 16: "\T \ set (filter g P). admissible_transaction T" using 11 12 by auto have "list_all (transaction_check (FP, OCC, TI)) (map add_occurs_msgs (filter g P))" using transactions_covered[unfolded protocol_covered_by_fixpoint_def] transaction_check_trivial_case[of _ FP_OCC_TI] unfolding FP_def OCC_def TI_def list_all_iff Let_def case_prod_unfold by auto note P0 = 16 15 14 this attack_notin_FP show ?thesis using prot_secure_if_fixpoint_covered[OF FP0 OCC0 TI0 M0 P0] reachable_constraints_secure_if_filter_secure_case[unfolded g_def[symmetric]] by fast qed lemma prot_secure_if_prot_checks_coverage_rcv: fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple" assumes attack_notin_fixpoint: "attack_notin_fixpoint FP_OCC_TI" and transactions_covered: "protocol_covered_by_fixpoint_coverage_rcv FP_OCC_TI P" and analyzed_fixpoint: "analyzed_fixpoint FP_OCC_TI" and wellformed_protocol: "wellformed_protocol' P N" and wellformed_fixpoint: "wellformed_fixpoint FP_OCC_TI" shows "\\ \ reachable_constraints P. \\. constraint_model \ (\@[(l, send\[attack\n\]\)])" using prot_secure_if_prot_checks[ OF attack_notin_fixpoint _ analyzed_fixpoint wellformed_protocol wellformed_fixpoint] protocol_covered_by_fixpoint_if_protocol_covered_by_fixpoint_coverage_rcv[ OF _ wellformed_fixpoint transactions_covered] wellformed_protocol[unfolded wellformed_protocol'_def] by blast end locale secure_stateful_protocol = pm: stateful_protocol_model arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 for arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" + fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple" and P_SMP::"('fun,'atom,'sets,'lbl) prot_term list" assumes attack_notin_fixpoint: "pm.attack_notin_fixpoint FP_OCC_TI" and transactions_covered: "pm.protocol_covered_by_fixpoint FP_OCC_TI P" and analyzed_fixpoint: "pm.analyzed_fixpoint FP_OCC_TI" and wellformed_protocol: "pm.wellformed_protocol' P P_SMP" and wellformed_fixpoint: "pm.wellformed_fixpoint FP_OCC_TI" begin theorem protocol_secure: "\\ \ pm.reachable_constraints P. \\. pm.constraint_model \ (\@[(l, send\[attack\n\]\)])" by (rule pm.prot_secure_if_prot_checks[OF attack_notin_fixpoint transactions_covered analyzed_fixpoint wellformed_protocol wellformed_fixpoint]) corollary protocol_welltyped_secure: "\\ \ pm.reachable_constraints P. \\. pm.welltyped_constraint_model \ (\@[(l, send\[attack\n\]\)])" using protocol_secure unfolding pm.welltyped_constraint_model_def by fast end locale secure_stateful_protocol' = pm: stateful_protocol_model arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 for arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" + fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple" assumes attack_notin_fixpoint': "pm.attack_notin_fixpoint FP_OCC_TI" and transactions_covered': "pm.protocol_covered_by_fixpoint FP_OCC_TI P" and analyzed_fixpoint': "pm.analyzed_fixpoint FP_OCC_TI" and wellformed_protocol': "pm.wellformed_protocol P" and wellformed_fixpoint': "pm.wellformed_fixpoint FP_OCC_TI" begin sublocale secure_stateful_protocol arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 P FP_OCC_TI "let f = \M. remdups (concat (map subterms_list M@map (fst \ pm.Ana) M)); N0 = remdups (concat (map (trms_list\<^sub>s\<^sub>s\<^sub>t \ unlabel \ transaction_strand) P)) in while (\A. set (f A) \ set A) f N0" apply unfold_locales using attack_notin_fixpoint' transactions_covered' analyzed_fixpoint' wellformed_protocol'[unfolded pm.wellformed_protocol_def Let_def] wellformed_fixpoint' unfolding Let_def by blast+ end locale secure_stateful_protocol'' = pm: stateful_protocol_model arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 for arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" + fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" assumes checks: "let FPT = pm.compute_fixpoint_fun P in pm.attack_notin_fixpoint FPT \ pm.protocol_covered_by_fixpoint FPT P \ pm.analyzed_fixpoint FPT \ pm.wellformed_protocol P \ pm.wellformed_fixpoint FPT" begin sublocale secure_stateful_protocol' arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 P "pm.compute_fixpoint_fun P" using checks[unfolded Let_def case_prod_unfold] by unfold_locales meson+ end locale secure_stateful_protocol''' = pm: stateful_protocol_model arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 for arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" + fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple" and P_SMP::"('fun,'atom,'sets,'lbl) prot_term list" assumes checks': "let P' = P; FPT = FP_OCC_TI; P'_SMP = P_SMP in pm.attack_notin_fixpoint FPT \ pm.protocol_covered_by_fixpoint FPT P' \ pm.analyzed_fixpoint FPT \ pm.wellformed_protocol' P' P'_SMP \ pm.wellformed_fixpoint FPT" begin sublocale secure_stateful_protocol arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 P FP_OCC_TI P_SMP using checks'[unfolded Let_def case_prod_unfold] by unfold_locales meson+ end locale secure_stateful_protocol'''' = pm: stateful_protocol_model arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 for arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" + fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple" assumes checks'': "let P' = P; FPT = FP_OCC_TI in pm.attack_notin_fixpoint FPT \ pm.protocol_covered_by_fixpoint FPT P' \ pm.analyzed_fixpoint FPT \ pm.wellformed_protocol P' \ pm.wellformed_fixpoint FPT" begin sublocale secure_stateful_protocol' arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 P FP_OCC_TI using checks''[unfolded Let_def case_prod_unfold] by unfold_locales meson+ end locale secure_stateful_protocol_coverage_rcv = pm: stateful_protocol_model arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 for arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" + fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple" and P_SMP::"('fun,'atom,'sets,'lbl) prot_term list" assumes attack_notin_fixpoint_coverage_rcv: "pm.attack_notin_fixpoint FP_OCC_TI" and transactions_covered_coverage_rcv: "pm.protocol_covered_by_fixpoint_coverage_rcv FP_OCC_TI P" and analyzed_fixpoint_coverage_rcv: "pm.analyzed_fixpoint FP_OCC_TI" and wellformed_protocol_coverage_rcv: "pm.wellformed_protocol' P P_SMP" and wellformed_fixpoint_coverage_rcv: "pm.wellformed_fixpoint FP_OCC_TI" begin sublocale secure_stateful_protocol arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 P FP_OCC_TI P_SMP using pm.protocol_covered_by_fixpoint_if_protocol_covered_by_fixpoint_coverage_rcv[ OF _ wellformed_fixpoint_coverage_rcv transactions_covered_coverage_rcv] attack_notin_fixpoint_coverage_rcv analyzed_fixpoint_coverage_rcv wellformed_protocol_coverage_rcv wellformed_fixpoint_coverage_rcv wellformed_protocol_coverage_rcv[unfolded pm.wellformed_protocol'_def] by unfold_locales meson+ end locale secure_stateful_protocol_coverage_rcv' = pm: stateful_protocol_model arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 for arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" + fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple" assumes attack_notin_fixpoint_coverage_rcv': "pm.attack_notin_fixpoint FP_OCC_TI" and transactions_covered_coverage_rcv': "pm.protocol_covered_by_fixpoint_coverage_rcv FP_OCC_TI P" and analyzed_fixpoint_coverage_rcv': "pm.analyzed_fixpoint FP_OCC_TI" and wellformed_protocol_coverage_rcv': "pm.wellformed_protocol P" and wellformed_fixpoint_coverage_rcv': "pm.wellformed_fixpoint FP_OCC_TI" begin sublocale secure_stateful_protocol_coverage_rcv arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 P FP_OCC_TI "let f = \M. remdups (concat (map subterms_list M@map (fst \ pm.Ana) M)); N0 = remdups (concat (map (trms_list\<^sub>s\<^sub>s\<^sub>t \ unlabel \ transaction_strand) P)) in while (\A. set (f A) \ set A) f N0" apply unfold_locales using attack_notin_fixpoint_coverage_rcv' transactions_covered_coverage_rcv' analyzed_fixpoint_coverage_rcv' wellformed_protocol_coverage_rcv'[unfolded pm.wellformed_protocol_def Let_def] wellformed_fixpoint_coverage_rcv' unfolding Let_def by blast+ end locale secure_stateful_protocol_coverage_rcv'' = pm: stateful_protocol_model arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 for arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" + fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" assumes checks_coverage_rcv: "let FPT = pm.compute_fixpoint_fun P in pm.attack_notin_fixpoint FPT \ pm.protocol_covered_by_fixpoint_coverage_rcv FPT P \ pm.analyzed_fixpoint FPT \ pm.wellformed_protocol P \ pm.wellformed_fixpoint FPT" begin sublocale secure_stateful_protocol_coverage_rcv' arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 P "pm.compute_fixpoint_fun P" using checks_coverage_rcv[unfolded Let_def case_prod_unfold] by unfold_locales meson+ end locale secure_stateful_protocol_coverage_rcv''' = pm: stateful_protocol_model arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 for arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" + fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple" and P_SMP::"('fun,'atom,'sets,'lbl) prot_term list" assumes checks_coverage_rcv': "let P' = P; FPT = FP_OCC_TI; P'_SMP = P_SMP in pm.attack_notin_fixpoint FPT \ pm.protocol_covered_by_fixpoint_coverage_rcv FPT P' \ pm.analyzed_fixpoint FPT \ pm.wellformed_protocol' P' P'_SMP \ pm.wellformed_fixpoint FPT" begin sublocale secure_stateful_protocol_coverage_rcv arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 P FP_OCC_TI P_SMP using checks_coverage_rcv'[unfolded Let_def case_prod_unfold] by unfold_locales meson+ end locale secure_stateful_protocol_coverage_rcv'''' = pm: stateful_protocol_model arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 for arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" + fixes P::"('fun,'atom,'sets,'lbl) prot_transaction list" and FP_OCC_TI:: "('fun,'atom,'sets,'lbl) fixpoint_triple" assumes checks_coverage_rcv'': "let P' = P; FPT = FP_OCC_TI in pm.attack_notin_fixpoint FPT \ pm.protocol_covered_by_fixpoint_coverage_rcv FPT P' \ pm.analyzed_fixpoint FPT \ pm.wellformed_protocol P' \ pm.wellformed_fixpoint FPT" begin sublocale secure_stateful_protocol_coverage_rcv' arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 P FP_OCC_TI using checks_coverage_rcv''[unfolded Let_def case_prod_unfold] by unfold_locales meson+ end subsection \Automatic Protocol Composition\ context stateful_protocol_model begin definition welltyped_leakage_free_protocol where "welltyped_leakage_free_protocol S P \ let f = \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}; Sec = (f (set S)) - {m. {} \\<^sub>c m} in \\ \ reachable_constraints P. \\\<^sub>\ s. (\l ts. suffix [(l, receive\ts\)] \) \ s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>\ \ welltyped_constraint_model \\<^sub>\ (\@[\\, send\[s]\\])" definition wellformed_composable_protocols where "wellformed_composable_protocols (P::('fun,'atom,'sets,'lbl) prot list) N \ let Ts = concat P; steps = remdups (concat (map transaction_strand Ts)); MP0 = \T \ set Ts. trms_transaction T \ pair' Pair ` setops_transaction T in list_all (wf\<^sub>t\<^sub>r\<^sub>m' arity) N \ has_all_wt_instances_of \ MP0 (set N) \ comp_tfr\<^sub>s\<^sub>e\<^sub>t arity Ana \ (set N) \ list_all (comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ Pair \ snd) steps \ list_all admissible_transaction_terms Ts \ list_all (list_all (\x. \\<^sub>v x = TAtom Value \ (is_Var (\\<^sub>v x) \ is_Atom (the_Var (\\<^sub>v x)))) \ transaction_fresh) Ts \ list_all (\T. \x \ vars_transaction T. \TAtom AttackType \ \\<^sub>v x) Ts \ list_all (\T. \x \ vars_transaction T. \f \ funs_term (\\<^sub>v x). f \ Pair \ f \ OccursFact) Ts \ list_all (list_all (\s. is_Send (snd s) \ length (the_msgs (snd s)) = 1 \ is_Fun_Attack (hd (the_msgs (snd s))) \ the_Attack_label (the_Fun (hd (the_msgs (snd s)))) = fst s) \ transaction_strand) Ts \ list_all (\r. (\f \ \(funs_term ` (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd r))). f = OccursFact \ f = OccursSec) \ (is_Receive (snd r) \ is_Send (snd r)) \ fst r = \ \ (\t \ set (the_msgs (snd r)). (OccursFact \ funs_term t \ OccursSec \ funs_term t) \ is_Fun t \ length (args t) = 2 \ t = occurs (args t ! 1) \ is_Var (args t ! 1) \ (\ (args t ! 1) = TAtom Value))) steps" definition wellformed_composable_protocols' where "wellformed_composable_protocols' (P::('fun,'atom,'sets,'lbl) prot list) \ let Ts = concat P in list_all wellformed_transaction Ts \ list_all (list_all (\p. let (x,cs) = p in is_Var (\\<^sub>v x) \ is_Atom (the_Var (\\<^sub>v x)) \ (\c \ cs. \\<^sub>v x = \ (Fun (Fu c) []::('fun,'atom,'sets,'lbl) prot_term))) \ (\T. transaction_decl T ())) Ts" definition composable_protocols where "composable_protocols (P::('fun,'atom,'sets,'lbl) prot list) Ms S \ let steps = concat (map transaction_strand (concat P)); M_fun = (\l. case find ((=) l \ fst) Ms of Some M \ set (snd M) | None \ {}) in comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ Pair steps M_fun (set S)" lemma composable_protocols_par_comp_constr: fixes S f defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and "Sec \ (f (set S)) - {m. intruder_synth {} m}" assumes Ps_pc: "wellformed_composable_protocols Ps N" "wellformed_composable_protocols' Ps" "composable_protocols Ps Ms S" shows "\\ \ reachable_constraints (concat Ps). \\. constraint_model \ \ \ (\\\<^sub>\. welltyped_constraint_model \\<^sub>\ \ \ ((\n. welltyped_constraint_model \\<^sub>\ (proj n \)) \ (\\' l t. prefix \' \ \ suffix [(l, receive\t\)] \' \ strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' Sec \\<^sub>\)))" (is "\\ \ _. \_. _ \ ?Q \ \") proof (intro allI ballI impI) fix \ \ assume \: "\ \ reachable_constraints (concat Ps)" and \: "constraint_model \ \" let ?Ts = "concat Ps" let ?steps = "concat (map transaction_strand ?Ts)" let ?MP0 = "\T \ set ?Ts. trms_transaction T \ pair' Pair ` setops_transaction T" let ?M_fun = "\l. case find ((=) l \ fst) Ms of Some M \ set (snd M) | None \ {}" have M: "has_all_wt_instances_of \ ?MP0 (set N)" "finite (set N)" "tfr\<^sub>s\<^sub>e\<^sub>t (set N)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set N)" using Ps_pc tfr\<^sub>s\<^sub>e\<^sub>t_if_comp_tfr\<^sub>s\<^sub>e\<^sub>t[of "set N"] unfolding composable_protocols_def wellformed_composable_protocols_def Let_def list_all_iff wf\<^sub>t\<^sub>r\<^sub>m_code[symmetric] by fast+ have P: "\T \ set ?Ts. wellformed_transaction T" "\T \ set ?Ts. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms_transaction T)" "\T \ set ?Ts. list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" "comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ Pair ?steps ?M_fun (set S)" using Ps_pc tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p unfolding wellformed_composable_protocols_def wellformed_composable_protocols'_def composable_protocols_def Let_def list_all_iff unlabel_def wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[symmetric] admissible_transaction_terms_def by (meson, meson, fastforce, blast) show "?Q \ \" using reachable_constraints_par_comp_constr[OF M P \ \] unfolding Sec_def f_def by fast qed context begin private lemma reachable_constraints_no_leakage_alt_aux: fixes P lbls L defines "lbls \ \T. map (the_LabelN o fst) (filter (Not o has_LabelS) (transaction_strand T))" and "L \ set (remdups (concat (map lbls P)))" assumes l: "l \ L" shows "map (transaction_proj l) P = map transaction_star_proj P" proof - have 0: "\list_ex (has_LabelN l) (transaction_strand T)" when "T \ set P" for T using that l unfolding L_def lbls_def list_ex_iff by force have 1: "\list_ex (has_LabelN l) (transaction_strand T)" when T: "T \ set (map (transaction_proj l) P)" for T proof - obtain T' where T': "T' \ set P" "T = transaction_proj l T'" using T by auto show ?thesis using T'(2) 0[OF T'(1)] proj_set_subset[of l "transaction_strand T'"] transaction_strand_proj[of l T'] unfolding list_ex_iff by fastforce qed have "list_all has_LabelS (transaction_strand T)" when "T \ set (map (transaction_proj l) P)" for T using that 1[OF that] transaction_proj_idem[of l] transaction_strand_proj[of l "transaction_proj l T"] has_LabelS_proj_iff_not_has_LabelN[of l "transaction_strand (transaction_proj l T)"] by (metis (no_types) ex_map_conv) thus ?thesis using transaction_star_proj_ident_iff transaction_proj_member transaction_star_proj_negates_transaction_proj(1) by (metis (mono_tags, lifting) map_eq_conv) qed private lemma reachable_constraints_star_no_leakage: fixes Sec P lbls k defines "no_leakage \ \\. \\\<^sub>\ \' s. prefix \' \ \ (\l ts. suffix [(l, receive\ts\)] \') \ s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\ \ welltyped_constraint_model \\<^sub>\ (\'@[(k,send\[s]\)])" assumes Sec: "\s \ Sec. \{} \\<^sub>c s" "ground Sec" shows "\\ \ reachable_constraints (map transaction_star_proj P). no_leakage \" proof fix A assume A: "A \ reachable_constraints (map transaction_star_proj P)" have A': "\(l,a) \ set A. l = \" using reachable_constraints_preserves_labels[OF A] transaction_star_proj_has_star_labels unfolding list_all_iff by fastforce show "no_leakage A" using constr_sem_stateful_star_proj_no_leakage[OF Sec(2) A'] unlabel_append[of A] singleton_lst_proj(4)[of k] unfolding no_leakage_def welltyped_constraint_model_def constraint_model_def by fastforce qed private lemma reachable_constraints_no_leakage_alt: fixes Sec P lbls k defines "no_leakage \ \\. \\\<^sub>\ \' s. prefix \' \ \ (\l ts. suffix [(l, receive\ts\)] \') \ s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\ \ welltyped_constraint_model \\<^sub>\ (\'@[(k,send\[s]\)])" and "lbls \ \T. map (the_LabelN o fst) (filter (Not o has_LabelS) (transaction_strand T))" and "L \ set (remdups (concat (map lbls P)))" assumes Sec: "\s \ Sec. \{} \\<^sub>c s" "ground Sec" and lbl: "\l \ L. \\ \ reachable_constraints (map (transaction_proj l) P). no_leakage \" shows "\l. \\ \ reachable_constraints (map (transaction_proj l) P). \\\<^sub>\ \'. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\) \ prefix \' \ \ (\l' ts. suffix [(l', receive\ts\)] \') \ strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' Sec \\<^sub>\" proof (intro allI ballI) fix l \ assume \: "\ \ reachable_constraints (map (transaction_proj l) P)" let ?Q = "\\\<^sub>\ \'. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\) \ prefix \' \ \ (\l' t. suffix [(l', receive\t\)] \') \ strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' Sec \\<^sub>\" show "\\\<^sub>\ \'. ?Q \\<^sub>\ \'" proof assume "\\\<^sub>\ \'. ?Q \\<^sub>\ \'" then obtain \\<^sub>\ \' t n l' ts' where \\<^sub>\: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" and \': "prefix \' \" "suffix [(l', receive\ts'\)] \'" and t: "t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\" and n: "constr_sem_stateful \\<^sub>\ (proj_unl n \'@[send\[t]\])" unfolding strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by blast hence 0: "welltyped_constraint_model \\<^sub>\ (proj n \'@[(m,send\[t]\)])" for m unfolding welltyped_constraint_model_def constraint_model_def by fastforce have t_Sec: "\{} \\<^sub>c t" "t \ \\<^sub>\ = t" using t Sec subst_ground_ident[of t \\<^sub>\] by auto obtain B k' s where B: "constr_sem_stateful \\<^sub>\ (proj_unl n B@[send\[t]\])" "prefix (proj n B) (proj n \)" "suffix [(k', receive\s\)] (proj n B)" "t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n B) \\<^sub>\" using constr_sem_stateful_proj_priv_term_prefix_obtain[OF \'(1) n t t_Sec] by metis hence 1: "welltyped_constraint_model \\<^sub>\ (proj n B@[(m,send\[t]\)])" for m using 0 unfolding welltyped_constraint_model_def constraint_model_def by fastforce note 2 = reachable_constraints_transaction_proj_proj_eq note 3 = reachable_constraints_transaction_proj_star_proj note 4 = reachable_constraints_no_leakage_alt_aux note star_case = 0 t t_Sec(1) reachable_constraints_star_no_leakage[OF Sec] \'(2) 3[OF \] prefix_proj(1)[OF \'(1)] suffix_proj(1)[OF \'(2)] declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj_eq note lbl_case = 0 t(1) \ \' lbl 2(2)[OF \ \'(1)] show False proof (cases "l = n") case True thus ?thesis proof (cases "l \ L") case False hence "map (transaction_proj l) P = map transaction_star_proj P" using 4 unfolding L_def lbls_def by fast thus ?thesis using lbl_case(1-4,7) star_case(4,5) True by metis qed (metis lbl_case no_leakage_def) next case False hence "no_leakage (proj n \)" using star_case(4,6) unfolding no_leakage_def by fast thus ?thesis by (metis B(2-4) 1 no_leakage_def) qed qed qed private lemma reachable_constraints_no_leakage_alt'_aux1: fixes P::"('a,'b,'c,'d) prot_transaction list" defines "f \ list_all ((list_all (Not \ has_LabelS)) \ tl \ transaction_send)" assumes P: "f P" shows "f (map (transaction_proj l) P)" and "f (map transaction_star_proj P)" proof - let ?g = "\T. tl (transaction_send T)" have "set (?g (transaction_proj l T)) \ set (?g T)" (is "?A \ ?C") and "set (?g (transaction_star_proj T)) \ set (?g T)" (is "?B \ ?C") for T::"('a,'b,'c,'d) prot_transaction" proof - obtain T1 T2 T3 T4 T5 T6 where T: "T = Transaction T1 T2 T3 T4 T5 T6" by (cases T) simp have "transaction_send (transaction_proj l T) = proj l (transaction_send T)" "transaction_send (transaction_star_proj T) = filter has_LabelS (transaction_send T)" using transaction_proj.simps[of l T1 T2 T3 T4 T5 T6] transaction_star_proj.simps[of T1 T2 T3 T4 T5 T6] unfolding T proj_def Let_def by auto hence "set (?g (transaction_proj l T)) \ set (proj l (?g T))" "set (?g (transaction_star_proj T)) \ set (filter has_LabelS (?g T))" unfolding proj_def by (metis (no_types, lifting) filter.simps(2) list.collapse list.sel(2,3) list.set_sel(2) subsetI)+ thus "?A \ ?C" "?B \ ?C" using T unfolding proj_def by auto qed thus "f (map (transaction_proj l) P)" "f (map transaction_star_proj P)" using P unfolding f_def list_all_iff by fastforce+ qed private lemma reachable_constraints_no_leakage_alt'_aux2: fixes P defines "f \ \T. list_all is_Receive (unlabel (transaction_receive T)) \ list_all is_Check_or_Assignment (unlabel (transaction_checks T)) \ list_all is_Update (unlabel (transaction_updates T)) \ list_all is_Send (unlabel (transaction_send T))" assumes P: "list_all f P" shows "list_all f (map (transaction_proj l) P)" (is ?A) and "list_all f (map transaction_star_proj P)" (is ?B) proof - have "f (transaction_proj l T)" (is ?A') and "f (transaction_star_proj T)" (is ?B') when T_in: "T \ set P" for T proof - obtain T1 T2 T3 T4 T5 T6 where T: "T = Transaction T1 T2 T3 T4 T5 T6" by (cases T) have "f T" using P T_in unfolding list_all_iff by simp thus ?A' ?B' unfolding f_def T unlabel_def proj_def Let_def list_all_iff transaction_proj.simps[of l T1 T2 T3 T4 T5 T6] transaction_star_proj.simps[of T1 T2 T3 T4 T5 T6] by auto qed thus ?A ?B unfolding list_all_iff by auto qed private lemma reachable_constraints_no_leakage_alt': fixes Sec P lbls k defines "no_leakage \ \\. \\\<^sub>\ \' s. prefix \' \ \ (\l ts. suffix [(l, receive\ts\)] \') \ s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\ \ welltyped_constraint_model \\<^sub>\ (\'@[(k,send\[s]\)])" and "no_leakage' \ \\. \\\<^sub>\ s. (\l ts. suffix [(l, receive\ts\)] \) \ s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>\ \ welltyped_constraint_model \\<^sub>\ (\@[(k,send\[s]\)])" assumes P: "list_all wellformed_transaction P" "list_all ((list_all (Not \ has_LabelS)) \ tl \ transaction_send) P" and Sec: "\s \ Sec. \{} \\<^sub>c s" "ground Sec" and lbl: "\l \ L. \\ \ reachable_constraints (map (transaction_proj l) P). no_leakage' \" shows "\l \ L. \\ \ reachable_constraints (map (transaction_proj l) P). no_leakage \" (is ?A) and "\\ \ reachable_constraints (map transaction_star_proj P). no_leakage \" (is ?B) proof - define f where "f \ \T::('fun,'atom,'sets,'lbl) prot_transaction. list_all is_Receive (unlabel (transaction_receive T)) \ list_all is_Check_or_Assignment (unlabel (transaction_checks T)) \ list_all is_Update (unlabel (transaction_updates T)) \ list_all is_Send (unlabel (transaction_send T))" define g where "(g::('fun,'atom,'sets,'lbl) prot_transaction \ bool) \ list_all (Not \ has_LabelS) \ tl \ transaction_send" have P': "list_all f P" using P(1) unfolding wellformed_transaction_def f_def list_all_iff by fastforce note 0 = reachable_constraints_no_leakage_alt'_aux1[OF P(2), unfolded g_def[symmetric]] note 1 = reachable_constraints_no_leakage_alt'_aux2[ OF P'[unfolded f_def], unfolded f_def[symmetric]] note 2 = reachable_constraints_aligned_prefix_ex[unfolded f_def[symmetric] g_def[symmetric]] have 3: "\\ \ reachable_constraints (map transaction_star_proj P). no_leakage' \" using reachable_constraints_star_no_leakage[OF Sec] unfolding no_leakage'_def by blast show ?A proof (intro ballI) fix l \ assume l: "l \ L" and \: "\ \ reachable_constraints (map (transaction_proj l) P)" show "no_leakage \" proof (rule ccontr) assume "\no_leakage \" then obtain \\<^sub>\ \' s where \': "prefix \' \" "\l ts. suffix [(l, receive\ts\)] \'" "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\" "welltyped_constraint_model \\<^sub>\ (\'@[(k, send\[s]\)])" unfolding no_leakage_def by blast have s: "\{} \\<^sub>c s" "fv s = {}" using \'(3) Sec by auto have \\<^sub>\: "constr_sem_stateful \\<^sub>\ (unlabel \'@[send\[s]\])" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" using \'(4) unfolding welltyped_constraint_model_def constraint_model_def by auto show False using 2[OF 1(1) 0(1) s \ \'(1,2) \\<^sub>\(1)] l lbl \'(3) \\<^sub>\(2,3,4) singleton_lst_proj(4)[of k "send\[s]\"] unlabel_append[of _ "[(k, send\[s]\)]"] unfolding no_leakage'_def welltyped_constraint_model_def constraint_model_def by metis qed qed show ?B proof (intro ballI) fix \ assume \: "\ \ reachable_constraints (map transaction_star_proj P)" show "no_leakage \" proof (rule ccontr) assume "\no_leakage \" then obtain \\<^sub>\ \' s where \': "prefix \' \" "\l ts. suffix [(l, receive\ts\)] \'" "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\" "welltyped_constraint_model \\<^sub>\ (\'@[(k, send\[s]\)])" unfolding no_leakage_def by blast have s: "\{} \\<^sub>c s" "fv s = {}" using \'(3) Sec by auto have \\<^sub>\: "constr_sem_stateful \\<^sub>\ (unlabel \'@[send\[s]\])" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" using \'(4) unfolding welltyped_constraint_model_def constraint_model_def by auto show False using 2[OF 1(2) 0(2) s \ \'(1,2) \\<^sub>\(1)] 3 \'(3) \\<^sub>\(2,3,4) singleton_lst_proj(4)[of k "send\[s]\"] unlabel_append[of _ "[(k, send\[s]\)]"] unfolding no_leakage'_def welltyped_constraint_model_def constraint_model_def by metis qed qed qed lemma composable_protocols_par_comp_prot_alt: fixes S f Sec lbls Ps defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and "Sec \ (f (set S)) - {m. {} \\<^sub>c m}" and "lbls \ \T. map (the_LabelN o fst) (filter (Not o has_LabelS) (transaction_strand T))" and "L \ set (remdups (concat (map lbls (concat Ps))))" and "no_leakage \ \\. \\\<^sub>\ \' s. prefix \' \ \ (\l ts. suffix [(l, receive\ts\)] \') \ s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\ \ welltyped_constraint_model \\<^sub>\ (\'@[\\, send\[s]\\])" assumes Ps_pc: "wellformed_composable_protocols Ps N" "wellformed_composable_protocols' Ps" "composable_protocols Ps Ms S" and component_secure: "\\ \ reachable_constraints (map (transaction_proj l) (concat Ps)). \\. welltyped_constraint_model \ (\@[\l, send\[attack\ln l\]\\])" and no_leakage: "\l \ L. \\ \ reachable_constraints (map (transaction_proj l) (concat Ps)). no_leakage \" shows "\\ \ reachable_constraints (concat Ps). \\. constraint_model \ (\@[\l, send\[attack\ln l\]\\])" proof fix \ assume \: "\ \ reachable_constraints (concat Ps)" let ?att = "[\l, send\[attack\ln l\]\\]" define Q where "Q \ \\\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" define R where "R \ \\ \\<^sub>\. \\' l t. prefix \' \ \ suffix [(l, receive\t\)] \' \ strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' Sec \\<^sub>\" define M where "M \ \T\set (concat Ps). trms_transaction T \ pair' Pair ` setops_transaction T" have Sec: "\s \ Sec. \{} \\<^sub>c s" "ground Sec" unfolding Sec_def f_def by auto have par_comp': "\\ \ reachable_constraints (concat Ps). \\. constraint_model \ \ \ (\\\<^sub>\. welltyped_constraint_model \\<^sub>\ \ \ ((\n. welltyped_constraint_model \\<^sub>\ (proj n \)) \ R \ \\<^sub>\))" using \ composable_protocols_par_comp_constr[OF Ps_pc] unfolding Sec_def f_def R_def by fast have "\l. \\ \ reachable_constraints (map (transaction_proj l) (concat Ps)). \\\<^sub>\. Q \\<^sub>\ \ R \ \\<^sub>\" using reachable_constraints_no_leakage_alt[OF Sec no_leakage[unfolded no_leakage_def L_def lbls_def]] unfolding Q_def R_def by blast hence no_leakage': "\\ \ reachable_constraints (concat Ps). \\\<^sub>\. Q \\<^sub>\ \ R \ \\<^sub>\" using reachable_constraints_component_leaks_if_composed_leaks[OF Sec, of "concat Ps" "\\\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)"] unfolding Q_def R_def by blast have M: "has_all_wt_instances_of \ M (set N)" "finite (set N)" "tfr\<^sub>s\<^sub>e\<^sub>t (set N)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (set N)" and P: "\T \ set (concat Ps). wellformed_transaction T" "\T \ set (concat Ps). admissible_transaction_terms T" "\T \ set (concat Ps). \x \ vars_transaction T. \TAtom AttackType \ \\<^sub>v x" "\T \ set (concat Ps). \s \ set (transaction_strand T). is_Send (snd s) \ length (the_msgs (snd s)) = 1 \ is_Fun_Attack (hd (the_msgs (snd s))) \ the_Attack_label (the_Fun (hd (the_msgs (snd s)))) = fst s" "\T \ set (concat Ps). list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" using Ps_pc(1,2) tfr\<^sub>s\<^sub>e\<^sub>t_if_comp_tfr\<^sub>s\<^sub>e\<^sub>t tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p unfolding wellformed_composable_protocols_def wellformed_composable_protocols'_def list_all_iff Let_def M_def wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s'_def wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code unlabel_def \\<^sub>v_TAtom''(1,2) by (force, force, fast, fast, fast, fast, fast, simp, simp) have P_fresh: "\T \ set (concat Ps). \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" (is "\T \ ?P. \x \ ?frsh T. ?Q x") proof (intro ballI) fix T x assume T: "T \ ?P" "x \ ?frsh T" hence "\\<^sub>v x = TAtom Value \ (is_Var (\\<^sub>v x) \ is_Atom (the_Var (\\<^sub>v x)))" using Ps_pc(1) unfolding wellformed_composable_protocols_def list_all_iff Let_def by fastforce thus "?Q x" by (metis prot_atom.is_Atom_def term.collapse(1)) qed have P': "\T \ set (concat Ps). wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms_transaction T)" using P(2) admissible_transaction_terms_def by fast have "\welltyped_constraint_model \ (\@?att)" for \ proof assume "welltyped_constraint_model \ (\@?att)" hence \: "welltyped_constraint_model \ \" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ attack\ln l\" using strand_sem_append_stateful[of "{}" "{}" "unlabel \" "unlabel ?att"] unlabel_append[of \ ?att] unfolding welltyped_constraint_model_def constraint_model_def by auto obtain \\<^sub>\ where \\<^sub>\: "welltyped_constraint_model \\<^sub>\ \" "welltyped_constraint_model \\<^sub>\ (proj l \)" using \ \ no_leakage' par_comp' unfolding Q_def welltyped_constraint_model_def constraint_model_def by metis have "\l, receive\[attack\ln l\]\\ \ set \" using reachable_constraints_receive_attack_if_attack(3)[OF \ P(1-2) P_fresh P(3) \ P(4)] by auto hence "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l \) \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\ \ attack\ln l\" using in_proj_set[of l "receive\[attack\ln l\]\" \] in_ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t_iff[of "attack\ln l\" "proj l \"] intruder_deduct.Axiom[of "attack\ln l\" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l \) \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\"] by fastforce hence "welltyped_constraint_model \\<^sub>\ (proj l \@?att)" using \\<^sub>\ strand_sem_append_stateful[of "{}" "{}" "unlabel (proj l \)" "unlabel ?att" \\<^sub>\] unfolding welltyped_constraint_model_def constraint_model_def by auto thus False using component_secure reachable_constraints_transaction_proj[OF \, of l] by simp qed thus "\\. constraint_model \ (\@?att)" using reachable_constraints_typing_result'[OF M_def M P(1) P' P(5) \] by blast qed lemma composable_protocols_par_comp_prot: fixes S f Sec lbls Ps defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and "Sec \ (f (set S)) - {m. {} \\<^sub>c m}" and "lbls \ \T. map (the_LabelN o fst) (filter (Not o has_LabelS) (transaction_strand T))" and "L \ set (remdups (concat (map lbls (concat Ps))))" and "no_leakage \ \\. \\\<^sub>\ s. (\l ts. suffix [(l, receive\ts\)] \) \ s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>\ \ welltyped_constraint_model \\<^sub>\ (\@[\\, send\[s]\\])" assumes Ps_pc: "wellformed_composable_protocols Ps N" "wellformed_composable_protocols' Ps" "composable_protocols Ps Ms S" "list_all ((list_all (Not \ has_LabelS)) \ tl \ transaction_send) (concat Ps)" and component_secure: "\\ \ reachable_constraints (map (transaction_proj l) (concat Ps)). \\. welltyped_constraint_model \ (\@[\l, send\[attack\ln l\]\\])" and no_leakage: "\l \ L. \\ \ reachable_constraints (map (transaction_proj l) (concat Ps)). no_leakage \" shows "\\ \ reachable_constraints (concat Ps). \\. constraint_model \ (\@[\l, send\[attack\ln l\]\\])" proof - have P': "list_all wellformed_transaction (concat Ps)" using Ps_pc(2) unfolding wellformed_composable_protocols'_def by meson have Sec: "\s \ Sec. \{} \\<^sub>c s" "ground Sec" unfolding Sec_def f_def by auto note 0 = composable_protocols_par_comp_prot_alt[ OF Ps_pc(1-3) component_secure, unfolded lbls_def[symmetric] L_def[symmetric]] note 1 = reachable_constraints_no_leakage_alt'[ OF P' Ps_pc(4) Sec no_leakage[unfolded no_leakage_def]] show ?thesis using 0 1 unfolding f_def Sec_def by argo qed lemma composable_protocols_par_comp_prot': assumes P_defs: "Pc = concat Ps" "set Ps_with_stars = (\n. map (transaction_proj n) Pc) ` set (remdups (concat (map (\T. map (the_LabelN \ fst) (filter (Not \ has_LabelS) (transaction_strand T))) Pc)))" and Ps_wellformed: "list_all (list_all (Not \ has_LabelS) \ tl \ transaction_send) Pc" "wellformed_composable_protocols Ps N" "wellformed_composable_protocols' Ps" "composable_protocols Ps Ms S" and Ps_no_leakage: "list_all (welltyped_leakage_free_protocol S) Ps_with_stars" and P_def: "P = map (transaction_proj n) Pc" and P_wt_secure: "\\ \ reachable_constraints P. \\. welltyped_constraint_model \ (\@[\n, send\[attack\ln n\]\\])" shows "\\ \ reachable_constraints Pc. \\. constraint_model \ (\@[\n, send\[attack\ln n\]\\])" by (rule composable_protocols_par_comp_prot[ OF Ps_wellformed(2,3,4,1)[unfolded P_defs(1)] P_wt_secure[unfolded P_def[unfolded P_defs(1)]] transaction_proj_ball_subst[ OF P_defs(2)[unfolded P_defs(1)] Ps_no_leakage(1)[ unfolded list_all_iff welltyped_leakage_free_protocol_def Let_def]], unfolded P_defs(1)[symmetric]]) end context begin lemma welltyped_constraint_model_leakage_model_swap: fixes I \ \::"('fun,'atom,'sets,'lbl) prot_subst" and s assumes A: "welltyped_constraint_model I (A@[\\, send\[s \ \]\\])" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "subst_domain \ = fv s" "ground (subst_range \)" obtains J where "welltyped_constraint_model J (A@[\\, send\[s \ \]\\])" and "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J \ s \ \ \ J" proof note defs = welltyped_constraint_model_def constraint_model_def note \_s = subst_fv_dom_ground_if_ground_img[OF equalityD2[OF \(3)] \(4)] note \' = transaction_renaming_subst_is_renaming(2)[OF \] inj_on_subset[OF transaction_renaming_subst_is_injective[OF \] subset_UNIV[of "fv s"]] transaction_renaming_subst_var_obtain(2)[OF \, of _ s] transaction_renaming_subst_is_renaming(6)[OF \, of s] transaction_renaming_subst_vars_disj(8)[OF \] transaction_renaming_subst_wt[OF \] define \inv where "\inv \ subst_var_inv \ (fv s)" define \' where "\' \ rm_vars (UNIV - fv (s \ \)) (\inv \\<^sub>s \)" define J where "J \ \x. if x \ fv (s \ \) then \' x else I x" have \_invertible: "s = s \ \ \\<^sub>s \inv" using \'(1) inj_var_ran_subst_is_invertible'[of \ s] inj_on_subset[OF \'(2)] unfolding \inv_def by blast have \'_domain: "subst_domain \' = fv (s \ \)" proof - have "x \ subst_domain (\inv \\<^sub>s \)" when x: "x \ fv (s \ \)" for x proof - obtain y where y: "y \ fv s" "\ y = Var x" using \'(3)[OF x] by blast have "y \ subst_domain \" using y(1) \(3) by blast moreover have "(\inv \\<^sub>s \) x = \ y" using y \'(3)[OF x] \_invertible vars_term_subset_subst_eq[of "Var y" s "\ \\<^sub>s \inv" Var] unfolding \'_def \inv_def - by (metis (no_types, lifting) fv_subst_subset subst_apply_term.simps(1) + by (metis (no_types, lifting) fv_subst_subset eval_term.simps(1) subst_apply_term_empty subst_compose) ultimately show ?thesis using \(4) by fastforce qed thus ?thesis using rm_vars_dom[of "UNIV - fv (s \ \)" "\inv \\<^sub>s \"] unfolding \'_def by blast qed have \'_range: "fv t = {}" when t: "t \ (subst_range \')" for t proof - obtain x where "x \ fv (s \ \)" "\' x = t" using t \'_domain by auto thus ?thesis by (metis (no_types, lifting) \'_def subst_compose_def \(3,4) \_invertible fv_subst_subset subst_fv_dom_ground_if_ground_img subst_subst_compose Diff_iff) qed have J0: "x \ fv (s \ \) \ J x = \' x" "x \ fv (s \ \) \ J x = I x" for x unfolding J_def by (cases "x \ fv (s \ \)") (simp_all add: subst_compose) have J1: "subst_range J \ subst_range \' \ subst_range I" proof fix t assume "t \ subst_range J" then obtain x where x: "x \ subst_domain J" "J x = t" by auto hence "t = \' x \ x \ subst_domain \'" "t = I x \ x \ subst_domain I" by (metis subst_domI subst_dom_vars_in_subst)+ thus "t \ subst_range \' \ subst_range I" using x(2) J0[of x] by auto qed have "x \ fv (s \ \)" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[\\, send\[s \ \]\\])" for x using x \_s \'(4) \'(5) by auto hence "I x = J x" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[\\, send\[s \ \]\\])" for x using x unfolding J_def \'_def by auto hence "constr_sem_stateful J (unlabel (A@[\\, send\[s \ \]\\]))" using A strand_sem_model_swap[of "unlabel (A@[\\, send\[s \ \]\\])" I J "{}" "{}"] unfolding defs by blast moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t J" using A subst_var_inv_wt[OF \'(6), of "fv s"] wt_subst_trm''[OF \(1)] subst_compose[of "subst_var_inv \ (fv s)" \] unfolding defs J_def \'_def \inv_def wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by presburger moreover have "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t J" proof - have "fv t = {}" when t: "t \ (subst_range J)" for t using t A J1 \'_range unfolding defs by auto moreover have "x \ subst_domain J" for x proof (cases "x \ fv (s \ \)") case True thus ?thesis using J0(1)[of x] \'_domain unfolding subst_domain_def by auto next case False have "subst_domain I = UNIV" using A unfolding defs by fast thus ?thesis using J0(2)[OF False] unfolding subst_domain_def by auto qed ultimately show ?thesis by auto qed moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \')" using wf_trms_subst_compose[OF subst_var_inv_wf_trms[of \ "fv s"] \(2)] unfolding \'_def \inv_def by force hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range J)" using A J1 unfolding defs by fast ultimately show "welltyped_constraint_model J (A@[\\, send\[s \ \]\\])" unfolding defs by blast hence "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J \ s \ \" using \_s strand_sem_append_stateful[of "{}" "{}" "unlabel A" "[send\[s \ \]\]" J] unfolding defs by (simp add: subst_ground_ident) moreover have "s \ \ \ J = s \ \" proof - have "J x = \' x" when x: "x \ fv (s \ \)" for x using x unfolding J_def by argo hence "s \ \ \ J = s \ \ \ \'" using subst_agreement[of "s \ \" J \'] by force thus ?thesis - using \_invertible unfolding \'_def - by (metis rm_vars_subst_eq' subst_subst_compose) + using \_invertible unfolding \'_def rm_vars_subst_eq'[symmetric] by (metis subst_subst_compose) qed hence "s \ \ \ J = s \ \" by auto ultimately show "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J \ s \ \ \ J" by argo qed lemma welltyped_leakage_free_protocol_pointwise: "welltyped_leakage_free_protocol S P \ list_all (\s. welltyped_leakage_free_protocol [s] P) S" unfolding welltyped_leakage_free_protocol_def list_all_iff Let_def by fastforce lemma welltyped_leakage_free_no_deduct_constI: fixes c defines "s \ Fun c []::('fun,'atom,'sets,'lbl) prot_term" assumes s: "\\ \ reachable_constraints P. \\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[\\, send\[s]\\])" shows "welltyped_leakage_free_protocol [s] P" using s unfolding welltyped_leakage_free_protocol_def s_def by auto lemma welltyped_leakage_free_pub_termI: assumes s: "{} \\<^sub>c s" shows "welltyped_leakage_free_protocol [s] P" proof - define Q where "Q \ \M t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}" define f where "f \ \M. {t \ \ | t \. Q M t \}" define Sec where "Sec \ f (set [s]) - {m. {} \\<^sub>c m}" have 0: "fv s = {}" using s pgwt_ground pgwt_is_empty_synth by blast have 1: "s \ \ = s" for \ by (rule subst_ground_ident[OF 0]) have 2: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range Var)" using wt_subst_Var wf_trm_subst_range_Var by (blast,blast) have "f (set [s]) = {s}" proof show "f (set [s]) \ {s}" using 0 1 unfolding f_def Q_def by auto have "Q {s} s Var" using 0 2 unfolding Q_def by auto thus "{s} \ f (set [s])" using 1[of Var] unfolding f_def by force qed hence "Sec = {}" using s unfolding Sec_def by simp thus ?thesis unfolding welltyped_leakage_free_protocol_def Let_def Sec_def f_def Q_def by blast qed lemma welltyped_leakage_free_pub_constI: assumes c: "public\<^sub>f c" "arity\<^sub>f c = 0" shows "welltyped_leakage_free_protocol [\c\\<^sub>c] P" using c welltyped_leakage_free_pub_termI[OF intruder_synth.ComposeC[of "[]" "Fu c" "{}"]] by simp lemma welltyped_leakage_free_long_term_secretI: fixes n defines "Tatt \ \s'. Transaction (\(). []) [] [\n, receive\[s']\\] [] [] [\n, send\[attack\ln n\]\\]" assumes P_wt_secure: "\\ \ reachable_constraints P. \\. welltyped_constraint_model \ (\@[\n, send\[attack\ln n\]\\])" and s_long_term_secret: "\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ inj_on \ (fv s) \ \ ` fv s \ range Var \ Tatt (s \ \) \ set P" shows "welltyped_leakage_free_protocol [s] P" proof (rule ccontr) define Q where "Q \ \M t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}" define f where "f \ \M. {t \ \ | t \. Q M t \}" define Sec where "Sec \ f (set [s]) - {m. {} \\<^sub>c m}" note defs = Sec_def f_def Q_def note defs' = welltyped_constraint_model_def constraint_model_def assume "\welltyped_leakage_free_protocol [s] P" then obtain A I s' where A: "A \ reachable_constraints P" "s' \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I" "welltyped_constraint_model I (A@[\\, send\[s']\\])" unfolding welltyped_leakage_free_protocol_def defs by fastforce obtain \ where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "\ ` fv s \ range Var" "inj_on \ (fv s)" "Tatt (s \ \) \ set P" using s_long_term_secret by blast obtain \ where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "subst_domain \ = fv (s \ \)" "ground (subst_range \)" "s' = s \ \ \ \" proof - obtain \ where *: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "fv s' = {}" "s' = s \ \" using A(2) unfolding defs by auto define \ where "\ \ subst_var_inv \ (fv s) \\<^sub>s \" define \' where "\' \ rm_vars (UNIV - fv (s \ \)) \" have **: "s' = s \ \ \ \" using *(4) inj_var_ran_subst_is_invertible[OF \(3,2)] unfolding \_def by simp have "s' = s \ \ \ \'" using ** rm_vars_subst_eq'[of "s \ \" \] unfolding \'_def by simp moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using \(1) *(1) subst_var_inv_wt wt_subst_compose unfolding \_def by presburger hence "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \'" using wt_subst_rm_vars unfolding \'_def by blast moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using wf_trms_subst_compose[OF subst_var_inv_wf_trms *(2)] unfolding \_def by blast hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \')" using wf_trms_subst_rm_vars'[of \] unfolding \'_def by blast moreover have "fv (s \ \) \ subst_domain \" using *(3) ** ground_term_subst_domain_fv_subset unfolding \_def by blast hence "subst_domain \' = fv (s \ \)" using rm_vars_dom[of "UNIV - fv (s \ \)" \] unfolding \'_def by blast moreover have "ground (subst_range \')" proof - { fix t assume "t \ subst_range \'" then obtain x where "x \ fv (s \ \)" "\' x = t" using \subst_domain \' = fv (s \ \)\ by auto hence "t \ s \ \ \ \'" by (meson subst_mono_fv) hence "fv t = {}" using \s' = s \ \ \ \'\ *(3) ground_subterm by blast } thus ?thesis by force qed ultimately show thesis using that[of \'] by fast qed have \: "transaction_decl_subst Var (Tatt t)" and \: "transaction_fresh_subst Var (Tatt t) (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" for t unfolding transaction_decl_subst_def transaction_fresh_subst_def Tatt_def by simp_all obtain \::"('fun,'atom,'sets,'lbl) prot_subst" where \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" unfolding transaction_renaming_subst_def by blast obtain J where J: "welltyped_constraint_model J (A@[\\, send\[s \ \ \ \]\\])" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J \ s \ \ \ \ \ J" using welltyped_constraint_model_leakage_model_swap[OF A(3)[unfolded \(5)] \ \(1-4)] by blast define T where "T = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand (Tatt (s \ \)) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" define B where "B \ A@T" have "transaction_receive (Tatt t) = [\n, receive\[t]\\]" "transaction_checks (Tatt t) = []" "transaction_updates (Tatt t) = []" "transaction_send (Tatt t) = [\n, send\[attack\ln n\]\\]" for t unfolding Tatt_def by simp_all hence T_def': "T = [\n, send\[s \ \ \ \]\\, \n, receive\[attack\ln n\]\\]" using subst_lsst_append[of "transaction_receive (Tatt (s \ \))" _ \] subst_lsst_singleton[of "ln n" "receive\[s \ \]\" \] subst_lsst_singleton[of "ln n" "send\[attack\ln n\]\" \] unfolding transaction_strand_def T_def by fastforce have B0: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \\<^sub>s\<^sub>e\<^sub>t J \ attack\ln n\" using in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of "attack\ln n\" "unlabel T"] unfolding B_def T_def' by (force intro!: intruder_deduct.Axiom) have B1: "B \ reachable_constraints P" using reachable_constraints.step[OF A(1) \(4) \ \ \] unfolding B_def T_def by simp have "welltyped_constraint_model J B" using J strand_sem_append_stateful[of "{}" "{}" "unlabel A" _ J] unfolding defs' B_def T_def' by fastforce hence B2: "welltyped_constraint_model J (B@[\n, send\[attack\ln n\]\\])" using B0 strand_sem_append_stateful[of "{}" "{}" "unlabel B" "[send\[attack\ln n\]\]" J] unfolding defs' B_def by auto show False using P_wt_secure B1 B2 by blast qed lemma welltyped_leakage_free_value_constI: assumes P: "\T \ set P. wellformed_transaction T" "\T \ set P. admissible_transaction_terms T" "\T \ set P. transaction_decl T () = []" "\T \ set P. bvars_transaction T = {}" and P_fresh_declass: "\T \ set P. transaction_fresh T \ [] \ (transaction_send T \ [] \ (let (l,a) = hd (transaction_send T) in l = \ \ is_Send a \ Var ` set (transaction_fresh T) \ set (the_msgs a)))" shows "welltyped_leakage_free_protocol [\m: value\\<^sub>v] P" proof (rule ccontr) define Q where "Q \ \M t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}" define f where "f \ \M. {t \ \ | t \. Q M t \}" define Sec where "Sec \ f (set [\m: value\\<^sub>v]) - {m. {} \\<^sub>c m}" note defs = Sec_def f_def Q_def note defs' = welltyped_constraint_model_def constraint_model_def assume "\welltyped_leakage_free_protocol [\m: value\\<^sub>v] P" then obtain A I s where A: "A \ reachable_constraints P" "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I" "welltyped_constraint_model I (A@[\\, send\[s]\\])" unfolding welltyped_leakage_free_protocol_def defs by fastforce have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ s \ I" using welltyped_constraint_model_deduct_split[OF A(3)] by simp moreover have "s \ I = s" using A(2) unfolding defs by fast ultimately have s_deduct: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I) \ s" by (metis ik\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst) note I0 = welltyped_constraint_model_prefix[OF A(3)] have I1: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" using A(3) unfolding defs' by blast obtain f ts \ where f: "s = Fun f ts" "s = \m: value\\<^sub>v \ \" "\{} \\<^sub>c s" "s \ declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "fv s = {}" using A(2) unfolding defs by (cases s) auto have s1: "\ s = TAtom Value" by (metis \.simps(1) \\<^sub>v_TAtom f(2) wt_subst_trm''[OF \(1)]) have s2: "wf\<^sub>t\<^sub>r\<^sub>m s" using f(2) \(2) by force have s3: "ts = []" using f(1) s1 s2 const_type_inv_wf by blast obtain sn where sn: "s = Fun (Val sn) []" using s1 f(3) \_Fu_simps(4) \_Set_simps(3) unfolding f(1) s3 by (cases f) auto have "s \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" using private_fun_deduct_in_ik'[OF s_deduct[unfolded sn]] by (metis sn public.simps(3) ik\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst) hence s4: "s \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using constraint_model_Val_const_in_constr_prefix[OF A(1) I0 P(1,2)] unfolding sn by presburger obtain B T \ \ \ where B: "prefix (B@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) A" "B \ reachable_constraints P" "T \ set P" "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "s \ subst_range \" using constraint_model_Value_in_constr_prefix_fresh_action'[OF A(1) P(2-) s4[unfolded sn]] sn by blast obtain Tts Tsnds sx where T: "transaction_send T = \\, send\Tts\\#Tsnds" "Var ` set (transaction_fresh T) \ set Tts" and sx: "Var sx \ set Tts" "\ sx = s" using P_fresh_declass B(3,5,7) unfolding transaction_fresh_subst_def is_Send_def by (cases "transaction_send T") (fastforce,fastforce) have \_elim: "\ \\<^sub>s \ \\<^sub>s \ = \ \\<^sub>s \" using admissible_transaction_decl_subst_empty'[OF bspec[OF P(3) B(3)] B(4)] by simp have s5: "s \ set (Tts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \ \\<^sub>s I)" using sx unfolding \_elim sn by force have s6: "\\, receive\Tts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \ \\<^sub>s I\\ \ set (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)" proof - have "\\, send\Tts\\ \ set (transaction_send T)" using T(1) by simp hence "\\, send\Tts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\ \ set (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" for \ unfolding subst_apply_labeled_stateful_strand_def by force hence "\\, send\Tts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\ \ set (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" for \ using transaction_strand_subst_subsets(4)[of T \] by fast hence *: "\\, receive\Tts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\ \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" for \ using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(1)[of \ "Tts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \"] by blast have "\\, receive\Tts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \\\ \ set A" using B(1) *[of "\ \\<^sub>s \ \\<^sub>s \"] unfolding prefix_def by force thus ?thesis unfolding subst_apply_labeled_stateful_strand_def by force qed show False using s6 f(4) ideduct_mono[OF Axiom[OF s5], of "\{set ts|ts. \\,receive\ts\\ \ set (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)}"] unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by blast qed lemma welltyped_leakage_free_priv_constI: fixes c defines "s \ Fun c []::('fun,'atom,'sets,'lbl) prot_term" assumes c: "\public c" "arity c = 0" "\ s = TAtom ca" "ca \ Value" and P: "\T \ set P. \x \ vars_transaction T. is_Var (\\<^sub>v x)" "\T \ set P. \x \ vars_transaction T \ set (transaction_fresh T). \ s \ \\<^sub>v x" "\T \ set P. \t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). s \ set (snd (Ana t))" "\T \ set P. s \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = \a\\<^sub>\\<^sub>a)" "\T \ set P. wellformed_transaction T" shows "\\ \ reachable_constraints P. \\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[\\, send\[s]\\])" (is "\\ \ ?R. ?P \") and "welltyped_leakage_free_protocol [s] P" proof - show "\\ \ ?R. ?P \" proof fix A assume A: "A \ reachable_constraints P" define Q where "Q \ \M t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}" define f where "f \ \M. {t \ \ | t \. Q M t \}" define Sec where "Sec \ f (set [s]) - {m. {} \\<^sub>c m}" define f' where "f' \ \(T,\,\::('fun,'atom,'sets,'lbl) prot_subst,\). dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define g' where "g' \ concat \ map f'" let ?P_s_cases = "\M. s \ M \ (\m \ subterms\<^sub>s\<^sub>e\<^sub>t M. s \ set (snd (Ana m)))" let ?P_s_cases' = "\M \. s \ M \\<^sub>s\<^sub>e\<^sub>t \ \ (\m \ subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \. s \ set (snd (Ana m)))" note defs = Sec_def f_def Q_def note defs' = welltyped_constraint_model_def constraint_model_def show "?P A" proof (rule ccontr) assume "\?P A" then obtain I where I: "welltyped_constraint_model I (A@[\\, send\[s]\\])" by blast obtain Ts where Ts: "A = g' Ts" "\B. prefix B Ts \ g' B \ reachable_constraints P" "\B T \ \ \. prefix (B@[(T,\,\,\)]) Ts \ T \ set P \ transaction_decl_subst \ T \ transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g' B)) \ transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (g' B))" using reachable_constraints_as_transaction_lists[OF A(1)] unfolding g'_def f'_def by blast have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ s \ I" and I_s: "s \ I = s" using welltyped_constraint_model_deduct_split[OF I] unfolding s_def by simp_all hence s_deduct: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I) \ s" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ s" by (metis ik\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst, metis) have I_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" and I_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" and I_grounds: "ground (subst_range I)" and I_interp: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" using I unfolding defs' by (blast,blast,blast,blast) have Sec_unfold: "Sec = {s}" proof have "\{} \\<^sub>c s" using ideduct_synth_priv_const_in_ik[OF _ c(1)] unfolding s_def by blast thus "{s} \ Sec" unfolding defs s_def by fastforce qed (auto simp add: defs s_def) have s2: "wf\<^sub>t\<^sub>r\<^sub>m s" using c(1,2) unfolding s_def by fastforce have A_ik_fv: "\a. \\<^sub>v x = TAtom a \ a \ ca" when x: "x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" for x proof - obtain T where T: "T \ set P" "\\<^sub>v x \ \\<^sub>v ` fv_transaction T" using fv_ik\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[OF x] reachable_constraints_var_types_in_transactions(1)[OF A P(5)] by fast then obtain y where y: "y \ vars_transaction T" "\\<^sub>v y = \\<^sub>v x" using vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] by fastforce then obtain a where a: "\\<^sub>v y = TAtom a" using P(1) T(1) by blast hence "\\<^sub>v x = TAtom a" "\ s \ \\<^sub>v x" "\ s = TAtom ca" using y P(2) T(1) c(3) by auto thus ?thesis by force qed have I_s_x: "\s \ I x" when x: "x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" for x proof - obtain a where a: "\\<^sub>v x = TAtom a" "a \ ca" using A_ik_fv[OF x] by blast hence a': "\ (I x) = TAtom a" using wt_subst_trm''[OF I_wt, of "Var x"] by simp obtain f ts where f: "I x = Fun f ts" by (meson empty_fv_exists_fun interpretation_grounds_all[OF I_interp]) hence ts: "ts = []" using I_wf const_type_inv_wf[OF a'[unfolded f]] by fastforce have "c \ f" using f[unfolded ts] a a' c(3)[unfolded s_def] by force thus ?thesis using f ts unfolding s_def by simp qed have A_ik_I_const: "\f. arity f = 0 \ I x = Fun f []" when x: "x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" for x using x A_ik_fv I_wt empty_fv_exists_fun[OF interpretation_grounds_all[OF I_interp, of x]] wf_trm_subst_rangeD[OF I_wf, of x] const_type_inv const_type_inv_wf by (metis (no_types, lifting) \.simps(1) wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) hence A_ik_subst: "subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I) = subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t I" using subterms_subst''[of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" I] by blast have sublmm1: "s \ set (snd (Ana m))" when m: "m \\<^sub>s\<^sub>e\<^sub>t M" "s \ set (snd (Ana (m \ \)))" and M: "\y. y \ fv\<^sub>s\<^sub>e\<^sub>t M \ \s \ \ y" for m M \ proof - have m_fun: "is_Fun m" using m M Ana_subterm' vars_iff_subtermeq_set - unfolding s_def is_Var_def by (metis subst_apply_term.simps(1)) + unfolding s_def is_Var_def by (metis eval_term.simps(1)) obtain f K R ts i where f: "m = \f ts\\<^sub>t" "arity\<^sub>f f = length ts" "arity\<^sub>f f > 0" "Ana\<^sub>f f = (K, R)" and i: "i < length R" "s = ts ! (R ! i) \ \" and R_i: "\i < length R. map ((!) ts) R ! i = ts ! (R ! i) \ R ! i < length ts" proof - obtain f ts K R where f: "m \ \ = \f ts\\<^sub>t" "arity\<^sub>f f = length ts" "arity\<^sub>f f > 0" "Ana\<^sub>f f = (K, R)" "Ana (m \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) ts, map ((!) ts) R)" using m(2) Ana_nonempty_inv[of "m \ \"] by force obtain ts' where m': "m = \f ts'\\<^sub>t" "ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using f(1) m_fun by auto have R_i: "map ((!) ts) R ! i = ts ! (R ! i)" "R ! i < length ts" when i: "i < length R" for i using i Ana\<^sub>f_assm2_alt[OF f(4), of "R ! i"] f(2) by simp_all then obtain i where i: "s = ts ! (R ! i)" "i < length R" by (metis (no_types, lifting) m(2) f(5) in_set_conv_nth length_map snd_conv) have ts': "arity\<^sub>f f = length ts'" "length ts = length ts'" using m'(2) f(2) by simp_all have s': "s = ts' ! (R ! i) \ \" using R_i(2)[OF i(2)] i(1) unfolding ts'(2) m'(2) by simp show thesis using that f m' R_i ts' s' i by auto qed have "s = ts ! (R ! i)" proof (cases "ts ! (R ! i)") case (Var x) hence "Var x \ set ts" using R_i i nth_mem by fastforce hence "x \ fv\<^sub>s\<^sub>e\<^sub>t M" using m(1) f(1) fv_subterms_set by fastforce thus ?thesis using i M Var by fastforce qed (use i s_def in fastforce) thus "s \ set (snd (Ana m))" using f(1) Ana_Fu_intro[OF f(2-4)] i(1) by simp qed have "\s \ \ y" when m: "m \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "s \ set (snd (Ana (m \ \)))" and T: "T \ set P" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_ran: "\t. t \ subst_range \ \ (\c. t = Fun c [] \ arity c = 0) \ (\x. t = Var x)" and y: "y \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" for m T \ y proof assume "s \ \ y" hence "\\<^sub>v y = \ s" using wt_subst_trm''[OF \_wt, of "Var y"] \_ran[of "\ y"] by fastforce moreover have "y \ vars_transaction T" using y trms\<^sub>s\<^sub>s\<^sub>t_fv_vars\<^sub>s\<^sub>s\<^sub>t_subset unfolding vars_transaction_unfold[of T] by fastforce ultimately show False using P(2) T by force qed hence sublmm2: "s \ set (snd (Ana m))" when m: "m \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "s \ set (snd (Ana (m \ \)))" and T: "T \ set P" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_ran: "\t. t \ subst_range \ \ (\c. t = Fun c [] \ arity c = 0) \ (\x. t = Var x)" for m T \ using sublmm1[OF m] m T \_wt \_ran by blast have "s \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ (\m \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t I. s \ set (snd (Ana m)))" using private_const_deduct[OF c(1) s_deduct(2)[unfolded s_def]] I_s_x const_mem_subst_cases[of c] A_ik_subst unfolding s_def by blast hence "?P_s_cases (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" using sublmm1[of _ "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A"] I_s_x by blast then obtain T \ \ \ where T: "(T,\,\,\) \ set Ts" "?P_s_cases (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (f' (T,\,\,\)))" using ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t_concat[of "map f' Ts"] Ts(1)[unfolded g'_def] by fastforce obtain B where "prefix (B@[(T, \, \, \)]) Ts" by (metis T(1) prefix_snoc_in_iff) hence T_in_P: "T \ set P" and T_wf: "wellformed_transaction T" and \: "transaction_decl_subst \ T" and \: "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (concat (map f' B)))" and \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (concat (map f' B)))" using P(6) Ts(3)[unfolded g'_def] unfolding comp_def by (metis,metis,metis,metis,metis) note \\\_wt = transaction_decl_fresh_renaming_substs_wt[OF \ \ \] note \\\_ran = transaction_decl_fresh_renaming_substs_range'(1)[OF \ \ \] have "subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \) = subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" for M using \\\_ran subterms_subst''[of _ "\ \\<^sub>s \ \\<^sub>s \"] by (meson subst_imgI) hence s_cases: "?P_s_cases' (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) (\ \\<^sub>s \ \\<^sub>s \)" using T(2) dual_transaction_ik_is_transaction_send'[OF T_wf, of "\ \\<^sub>s \ \\<^sub>s \"] unfolding f'_def by auto from s_cases show False proof assume "s \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \" then obtain t where t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "s = t \ \ \\<^sub>s \ \\<^sub>s \" by moura have "s \ t" using P(4) T_in_P t(1) by blast then obtain x where x: "t = Var x" using t(2) unfolding s_def by (cases t) auto have "\\<^sub>v x = \ s" using x t(2) wt_subst_trm''[OF \\\_wt, of "Var x"] by simp moreover have "x \ vars_transaction T" using t(1) trms\<^sub>s\<^sub>s\<^sub>t_fv_vars\<^sub>s\<^sub>s\<^sub>t_subset unfolding x vars_transaction_unfold[of T] by fastforce ultimately show False using P(2) T_in_P by force qed (use sublmm2[OF _ _ T_in_P \\\_wt \\\_ran] P(3) T_in_P in blast) qed qed thus "welltyped_leakage_free_protocol [s] P" using welltyped_leakage_free_no_deduct_constI[of P c] unfolding s_def by blast qed lemma welltyped_leakage_free_priv_constI': assumes c: "\public\<^sub>f c" "arity\<^sub>f c = 0" "\\<^sub>f c = Some ca" and P: "\T \ set P. wellformed_transaction T" "\T \ set P. \x \ vars_transaction T \ set (transaction_fresh T). \ \c\\<^sub>c \ \\<^sub>v x" "\T \ set P. \x \ vars_transaction T. is_Var (\\<^sub>v x)" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = \a\\<^sub>\\<^sub>a)" "\T \ set P. \t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). \c\\<^sub>c \ set (snd (Ana t))" "\T \ set P. \c\\<^sub>c \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" shows "\\ \ reachable_constraints P. \\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[\\, send\[\c\\<^sub>c]\\])" and "welltyped_leakage_free_protocol [\c\\<^sub>c] P" using c welltyped_leakage_free_priv_constI[OF _ _ _ _ P(3,2,5,6,4,1), of "Atom ca"] by (force, force) lemma welltyped_leakage_free_set_constI: assumes P: "\T \ set P. wellformed_transaction T" "\T \ set P. \f \ \(funs_term ` (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))). \is_Set f" "\T \ set P. \x \ vars_transaction T \ set (transaction_fresh T). \\<^sub>v x \ TAtom SetType" "\T \ set P. \x \ vars_transaction T. is_Var (\\<^sub>v x)" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = \a\\<^sub>\\<^sub>a)" and c: "arity\<^sub>s c = 0" shows "\\ \ reachable_constraints P. \\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[\\, send\[\c\\<^sub>s]\\])" and "welltyped_leakage_free_protocol [\c\\<^sub>s] P" proof - have c'': "\c\\<^sub>s \ subterms t" when T: "T \ set P" and t: "t \\<^sub>s\<^sub>e\<^sub>t trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" for T t using t bspec[OF P(2) T] subtermeq_imp_funs_term_subset[of t] funs_term_Fun_subterm'[of "Set c" "[]::('fun,'atom,'sets,'lbl) prot_term list"] by fastforce have P': "\T \ set P. \t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). \c\\<^sub>s \ set (snd (Ana t))" "\T \ set P. \c\\<^sub>s \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" subgoal using Ana_subterm' c'' by fast subgoal using c'' by fast done have P'': "\T \ set P. \x \ vars_transaction T \ set (transaction_fresh T). \ \c\\<^sub>s \ \\<^sub>v x" using P(3) \_consts_simps(4)[OF c] by fastforce show "\\ \ reachable_constraints P. \\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[\\, send\[\c\\<^sub>s]\\])" "welltyped_leakage_free_protocol [\c\\<^sub>s] P" using c welltyped_leakage_free_priv_constI[OF _ _ _ _ P(4) P'' P' P(5,1), of SetType] by (force, force) qed lemma welltyped_leakage_free_occurssec_constI: defines "s \ Fun OccursSec []" assumes P: "\T \ set P. wellformed_transaction T" "\T \ set P. \x \ vars_transaction T \ set (transaction_fresh T). \\<^sub>v x \ TAtom OccursSecType" "\T \ set P. \t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). Fun OccursSec [] \ set (snd (Ana t))" "\T \ set P. Fun OccursSec [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "\T \ set P. \x \ vars_transaction T. is_Var (\\<^sub>v x)" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = \a\\<^sub>\\<^sub>a)" shows "\\ \ reachable_constraints P. \\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[\\, send\[s]\\])" and "welltyped_leakage_free_protocol [s] P" proof - have P': "\T \ set P. \x \ vars_transaction T \ set (transaction_fresh T). \ (Fun OccursSec []) \ \\<^sub>v x" using P(2) by auto show "\\ \ reachable_constraints P. \\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[\\, send\[s]\\])" "welltyped_leakage_free_protocol [s] P" using welltyped_leakage_free_priv_constI[OF _ _ _ _ P(5) P' P(3,4,6,1), of OccursSecType] unfolding s_def by auto qed lemma welltyped_leakage_free_occurs_factI: assumes P: "\T \ set P. admissible_transaction' T" and P_occ: "\T \ set P. admissible_transaction_occurs_checks T" and P_occ_star: "\T \ set P. \r \ set (transaction_send T). OccursFact \ \(funs_term ` (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd r))) \ fst r = \" shows "welltyped_leakage_free_protocol [occurs x] P" proof - define Q where "Q \ \M t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}" define f where "f \ \M. {t \ \ | t \. Q M t \}" define Sec where "Sec \ f (set [occurs x]) - {m. {} \\<^sub>c m}" define f' where "f' \ \(T,\,\::('fun,'atom,'sets,'lbl) prot_subst,\). dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define g' where "g' \ concat \ map f'" note defs = Sec_def f_def Q_def note defs' = welltyped_constraint_model_def constraint_model_def show ?thesis proof (rule ccontr) assume "\welltyped_leakage_free_protocol [occurs x] P" then obtain A I k where A: "A \ reachable_constraints P" "occurs k \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I" "welltyped_constraint_model I (A@[\\, send\[occurs k]\\])" unfolding welltyped_leakage_free_protocol_def defs by fastforce note A' = welltyped_constraint_model_prefix[OF A(3)] have occ_I: "occurs k \ I = occurs k" using A(2) unfolding defs by auto hence occ_in_ik: "occurs k \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "occurs k \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" using welltyped_constraint_model_deduct_split(2)[OF A(3)] reachable_constraints_occurs_fact_deduct_in_ik[OF A(1) A' P P_occ, of k] by (argo, argo) then obtain l ts where ts: "(l,receive\ts\) \ set A" "occurs k \ set ts" using in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of "occurs k" "unlabel A"] unfolding unlabel_def by moura obtain T a B \ \ \ where B: "prefix (B@f' (T,\,\,\)) A" and T: "T \ set P" "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" and a: "a \ set (transaction_strand T)" "fst (l,receive\ts\) = fst a" "(l,receive\ts\) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \ \\<^sub>s \" using reachable_constraints_transaction_action_obtain[OF A(1) ts(1), of thesis] unfolding f'_def by simp obtain ts' where ts': "a = (l,send\ts'\)" "ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \" using surj_pair[of a] a(2,3) by (cases "snd a") force+ obtain t where t: "t \ set ts'" "occurs k = t \ \ \\<^sub>s \ \\<^sub>s \" using ts(2) unfolding ts'(2) by force have occ_t: "OccursFact \ funs_term t" proof (cases t) case (Var y) thus ?thesis - using t(2) subst_apply_term.simps(1)[of y "\ \\<^sub>s \ \\<^sub>s \"] - transaction_decl_fresh_renaming_substs_range'(1)[OF T(2-), of "occurs k"] + using t(2) transaction_decl_fresh_renaming_substs_range'(1)[OF T(2-), of "occurs k"] by fastforce qed (use t(2) in simp) have P_wf: "\T \ set P. wellformed_transaction T" using P admissible_transaction_is_wellformed_transaction(1) by blast have l: "l = \" using wellformed_transaction_strand_memberD(8)[OF bspec[OF P_wf T(1)] a(1)[unfolded ts'(1)]] t(1) T(1) P_occ_star occ_t unfolding ts'(1) by fastforce have "occurs k \ \{set ts |ts. \\, receive\ts\\ \ set (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)}" using subst_lsst_memI[OF ts(1), of I] subst_set_map[OF ts(2), of I] unfolding occ_I l by auto thus False using A(2) unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by simp qed qed lemma welltyped_leakage_free_setop_pairI: assumes P: "\T \ set P. wellformed_transaction T" "\T \ set P. \x \ vars_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = \a\\<^sub>\\<^sub>a)" "\T \ set P. \f \ \(funs_term ` (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))). \is_Set f" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" "\T \ set P. transaction_decl T () = []" "\T \ set P. admissible_transaction_terms T" and c: "arity\<^sub>s c = 0" shows "welltyped_leakage_free_protocol [pair (x, \c\\<^sub>s)] P" proof - define Q where "Q \ \M t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}" define f where "f \ \M. {t \ \ | t \. Q M t \}" define Sec where "Sec \ f (set [pair (x, \c\\<^sub>s)]) - {m. {} \\<^sub>c m}" define f' where "f' \ \(T,\,\::('fun,'atom,'sets,'lbl) prot_subst,\). dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define g' where "g' \ concat \ map f'" note defs = Sec_def f_def Q_def note defs' = welltyped_constraint_model_def constraint_model_def have P': "\T \ set P. \x \ vars_transaction T \ set (transaction_fresh T). \\<^sub>v x \ TAtom SetType" "\T \ set P. \x \ vars_transaction T. is_Var (\\<^sub>v x)" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = \a\\<^sub>\\<^sub>a)" subgoal using P(2,4) by fastforce subgoal using P(2) by fastforce subgoal using P(4) by fast done note 0 = welltyped_leakage_free_set_constI(1)[OF P(1,3) P' c] show ?thesis proof (rule ccontr) assume "\welltyped_leakage_free_protocol [pair (x, \c\\<^sub>s)] P" then obtain A I k where A: "A \ reachable_constraints P" "pair (k, \c\\<^sub>s) \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I" "welltyped_constraint_model I (A@[\\, send\[pair (k, \c\\<^sub>s)]\\])" unfolding welltyped_leakage_free_protocol_def defs pair_def by fastforce note A' = welltyped_constraint_model_prefix[OF A(3)] have "pair (k, \c\\<^sub>s) \ I = pair (k, \c\\<^sub>s)" using A(2) unfolding defs by auto hence "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ pair (k, \c\\<^sub>s)" using welltyped_constraint_model_deduct_split(2)[OF A(3)] by argo then obtain n where n: "intruder_deduct_num (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I) n (pair (k, \c\\<^sub>s))" using deduct_num_if_deduct by fast have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using A(3) ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset unfolding defs' by simp_all hence "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" by blast hence "Pair \ \(funs_term ` (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I))" using reachable_constraints_no_Pair_fun'[OF A(1) P(4-6)] P by blast hence 1: "\pair (k, \c\\<^sub>s) \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" using funs_term_Fun_subterm'[of Pair] unfolding pair_def by auto have 2: "pair (k, \c\\<^sub>s) \ set (snd (Ana m))" when m: "m \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" for m using m 1 term.dual_order.trans Ana_subterm'[of "pair (k, \c\\<^sub>s)" m] by auto have "\ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ \c\\<^sub>s" using 0 A(1) A' welltyped_constraint_model_deduct_iff[of I A \ "\c\\<^sub>s"] by force moreover have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ \c\\<^sub>s" using 1 2 deduct_inv[OF n] deduct_if_deduct_num[of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" _ "\c\\<^sub>s"] unfolding pair_def by auto ultimately show False by blast qed qed lemma welltyped_leakage_free_short_term_secretI: fixes c x y f n d l l' defines "s \ \f [\c\\<^sub>c, \x: value\\<^sub>v]\\<^sub>t" and "Tatt \ Transaction (\(). []) [] [\\, receive\[occurs \y: value\\<^sub>v]\\, (l, receive\[\f [\c\\<^sub>c, \y: value\\<^sub>v]\\<^sub>t]\)] [(l', \\y: value\\<^sub>v not in \d\\<^sub>s\)] [] [\n, send\[attack\ln n\]\\]" assumes P: "\T \ set P. admissible_transaction' T" "\T \ set P. admissible_transaction_occurs_checks T" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = \a\\<^sub>\\<^sub>a)" and subterms_sec: "\\ \ reachable_constraints P. \\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[\\, send\[\c\\<^sub>c]\\])" and P_sec: "\\ \ reachable_constraints P. \\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[\n, send\[attack\ln n\]\\])" and P_Tatt: "Tatt \ set P" and P_d: "\T \ set P. \a \ set (transaction_updates T). is_Insert (snd a) \ the_set_term (snd a) = \d\\<^sub>s \ transaction_send T \ [] \ (let (l,b) = hd (transaction_send T) in l = \ \ is_Send b \ \f [\c\\<^sub>c, the_elem_term (snd a)]\\<^sub>t \ set (the_msgs b))" shows "welltyped_leakage_free_protocol [s] P" proof - define Q where "Q \ \M t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}" define Sec where "Sec \ {t \ \ | t \. Q (set [s]) t \} - {m. {} \\<^sub>c m}" define f' where "f' \ \(T,\,\::('fun,'atom,'sets,'lbl) prot_subst,\). dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define g' where "g' \ concat \ map f'" note defs = Sec_def Q_def note defs' = welltyped_constraint_model_def constraint_model_def show ?thesis proof (rule ccontr) assume "\welltyped_leakage_free_protocol [s] P" then obtain A I k where A: "A \ reachable_constraints P" "\f [\c\\<^sub>c, k]\\<^sub>t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I" "welltyped_constraint_model I (A@[\\, send\[\f [\c\\<^sub>c, k]\\<^sub>t]\\])" unfolding welltyped_leakage_free_protocol_def defs s_def by fastforce have I: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" using A(3) unfolding defs' by (blast,blast,blast) note A' = welltyped_constraint_model_prefix[OF A(3)] have "strand_sem_stateful {} {} (unlabel A) I" using A' unfolding defs' by simp hence A'': "strand_sem_stateful {} {} (unlabel A) (I(z := k))" when z: "z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" for z using z strand_sem_model_swap[of "unlabel A" I "I(z := k)" "{}" "{}"] by auto obtain \ where \: "\ (the_Var \x: value\\<^sub>v) = k" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "fv (\ (the_Var \x: value\\<^sub>v)) = {}" using A(2) unfolding defs s_def by auto have k: "\ k = TAtom Value" "fv k = {}" "wf\<^sub>t\<^sub>r\<^sub>m k" subgoal using \(1) wt_subst_trm''[OF \(2), of "\x: value\\<^sub>v"] by simp subgoal using \(1,4) by blast subgoal using \(1,3) by force done then obtain fk where fk: "k = Fun fk []" using const_type_inv_wf by (cases k) auto have fk': "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ \f [\c\\<^sub>c, k]\\<^sub>t" using fk welltyped_constraint_model_deduct_split(2)[OF A(3)] by auto have "\welltyped_constraint_model I (A@[\\, send\[\c\\<^sub>c]\\])" using subterms_sec(1) A(1) by blast hence "\ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ \c\\<^sub>c" using A' strand_sem_append_stateful[of "{}" "{}" "unlabel A" "unlabel [\\, send\[\c\\<^sub>c]\\]" I] unfolding defs' by auto hence "\f [\c\\<^sub>c, k]\\<^sub>t \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" using fk' deduct_inv'[OF fk'] by force moreover have "k \ \f [\c\\<^sub>c, k]\\<^sub>t" by simp ultimately have k_in_ik: "k \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" using in_subterms_subset_Union by blast hence "k \\<^sub>s\<^sub>e\<^sub>t ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ (\x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A). k \ I x)" using const_subterms_subst_cases[of fk I "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A"] unfolding fk by fast hence "fk \ \(funs_term ` ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ (\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. k \ I x)" unfolding fk by (meson UN_iff funs_term_Fun_subterm' fv_ik\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t ) hence "fk \ \(funs_term ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ (\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. k \ I x)" using ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset by blast moreover have "\\<^sub>v x = TAtom Value" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" for x using x reachable_constraints_var_types_in_transactions(1)[OF A(1) P(3)] P(1,2) admissible_transaction_Value_vars by force ultimately have "fk \ \(funs_term ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ (\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \\<^sub>v x = TAtom Value \ k \ I x)" by blast hence "fk \ \(funs_term ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ (\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \\<^sub>v x = TAtom Value \ I x = k)" using I(1,3) wf_trm_TComp_subterm wf_trm_subst_rangeD unfolding fk wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by (metis \.simps(1) term.distinct(1)) then obtain kn where kn: "fk = Val kn" using reachable_constraints_val_funs_private[OF A(1) P(1), of fk] constraint_model_Value_term_is_Val[OF A(1) A' P(1,2)] Fun_Value_type_inv[OF k(1)[unfolded fk]] unfolding fk is_PubConstValue_def by (cases fk) force+ obtain \::"('fun,'atom,'sets,'lbl) prot_subst" where \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" unfolding transaction_renaming_subst_def by blast obtain y' yn' where y': "\ (the_Var \y: value\\<^sub>v) = Var y'" "y \ yn'" "Var y' = \yn': value\\<^sub>v" using transaction_renaming_subst_is_renaming(1,2)[OF \] by force define B where "B \ A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand Tatt \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" define J where "J \ I(y' := k)" have "y' \ range_vars \" using y'(1) transaction_renaming_subst_is_renaming(3)[OF \] by (metis (no_types, lifting) in_mono subst_fv_imgI term.set_intros(3)) hence y'': "y' \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using transaction_renaming_subst_vars_disj(6)[OF \] by blast have 0: "(k,\d\\<^sub>s) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I)" proof assume a: "(k,\d\\<^sub>s) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I)" obtain l t t' where t: "(l,insert\t,t'\) \ set A" "t \ I = k" "t' \ I = \d\\<^sub>s" using db\<^sub>s\<^sub>s\<^sub>t_in_cases[OF a[unfolded db\<^sub>s\<^sub>s\<^sub>t_def]] unfolding unlabel_def by auto obtain T b B \ \ \ where T: "prefix (B@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) A" "T \ set P" "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "b \ set (transaction_strand T)" "(l, insert\t,t'\) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \ \\<^sub>s \" "fst (l, insert\t,t'\) = fst b" using reachable_constraints_transaction_action_obtain[OF A(1) t(1)] by metis define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" obtain b' where "b = (l,b')" using T(8) by (cases b) simp then obtain tb tb' where b': "b = (l,insert\tb,tb'\)" and tb: "t = tb \ \" and tb': "t' = tb' \ \" using T(7) unfolding \_def by (cases b') auto note T_adm = bspec[OF P(1) T(2)] note T_wf = admissible_transaction_is_wellformed_transaction(1,3)[OF T_adm] have b: "b \ set (transaction_updates T)" using transaction_strand_memberD[OF T(6)[unfolded b']] wellformed_transaction_cases[OF T_wf(1)] unfolding b' by blast have "\n. tb = \n: value\\<^sub>v" and *: "tb' = \d\\<^sub>s" using tb tb' T(6) t(3) transaction_inserts_are_Value_vars[OF T_wf, of tb tb'] unfolding b' unlabel_def by (force,force) have "is_Insert (snd b)" "the_set_term (snd b) = \d\\<^sub>s" "the_elem_term (snd b) = tb" unfolding b' * by simp_all hence "transaction_send T \ []" "let (l, a) = hd (transaction_send T) in l = \ \ is_Send a \ \f [\c\\<^sub>c, tb]\\<^sub>t \ set (the_msgs a)" using P_d T(2) b by (fast,fast) hence "\ts. \\,send\ts\\ \ set (transaction_send T) \ \f [\c\\<^sub>c, tb]\\<^sub>t \ set ts" unfolding is_Send_def by (cases "transaction_send T") auto then obtain ts where ts: "\\,send\ts\\ \ set (transaction_strand T)" "\f [\c\\<^sub>c, tb]\\<^sub>t \ set ts" unfolding transaction_strand_def by auto have "\\,receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\ \ set A" "\f [\c\\<^sub>c, t]\\<^sub>t \ set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using subst_lsst_memI[OF ts(1), of \] subst_set_map[OF ts(2), of \] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(1)[of \ "ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] set_mono_prefix[OF T(1)[unfolded \_def[symmetric]]] unfolding tb by auto hence "\f [\c\\<^sub>c, k]\\<^sub>t \ \{set ts |ts. \\, receive\ts\\ \ set (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)}" using t(2) subst_lsst_memI[of "\\,receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\" A I] by force thus False using A(2) unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto qed have "y' \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" using y'' fv_ik_subset_vars_sst'[of "unlabel A"] by blast hence 1: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J" unfolding J_def by (metis (no_types, lifting) fv_subset image_cong in_mono repl_invariance) have "(k,\d\\<^sub>s) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I {}" using 0 db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of "unlabel A" I "[]"] unfolding db\<^sub>s\<^sub>s\<^sub>t_def by force hence "(k,\d\\<^sub>s) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) J {}" using y'' vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel A"] db\<^sub>s\<^sub>s\<^sub>t_subst_swap[of "unlabel A" I J "[]"] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of "unlabel A" _ "[]"] unfolding db\<^sub>s\<^sub>s\<^sub>t_def J_def by (metis (no_types, lifting) Un_iff empty_set fun_upd_other) hence "((Var y' \ J, \d\\<^sub>s \ J) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) J {})" unfolding J_def fk by simp hence "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J) (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) J {}) (unlabel [\n, \Var y' not in \d\\<^sub>s\\]) J" using stateful_strand_sem_NegChecks_no_bvars(1)[ of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J" "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) J {}" "Var y'" "\d\\<^sub>s" J] by simp hence 2: "strand_sem_stateful {} {} (unlabel (A@[\n, \Var y' not in \d\\<^sub>s\\])) J" using A'' y' y'' vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel A"] strand_sem_append_stateful[ of "{}" "{}" "unlabel A" "unlabel [\n, \Var y' not in \d\\<^sub>s\\]" J] unfolding J_def by simp have B: "B \ reachable_constraints P" using reachable_constraints.step[OF A(1) P_Tatt _ _ \, of Var Var] unfolding B_def Tatt_def transaction_decl_subst_def transaction_fresh_subst_def by simp have Tatt': "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand Tatt \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = [\\, send\[occurs (Var y')]\\, (l, send\[\f [\c\\<^sub>c, Var y']\\<^sub>t]\), (l', \Var y' not in \d\\<^sub>s\), \n, receive\[attack\ln n\]\\]" using y' unfolding Tatt_def transaction_strand_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def subst_apply_labeled_stateful_strand_def by auto have J: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t J" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t J" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range J)" unfolding J_def subgoal using wt_subst_subst_upd[OF I(1)] k(1) y'(3) by simp subgoal by (metis I(2) k(2) fun_upd_apply interpretation_grounds_all interpretation_substI) subgoal using I(3) k(3) by fastforce done have 3: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J \ \f [\c\\<^sub>c, Var y']\\<^sub>t \ J" using 1 fk fk' unfolding J_def by auto have 4: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J \ occurs (Var y') \ J" using reachable_constraints_occurs_fact_ik_case'[ OF A(1) P(1,2) constraint_model_Val_const_in_constr_prefix'[ OF A(1) A' P(1) k_in_ik[unfolded fk kn]]] intruder_deduct.Axiom[of "occurs k" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J"] unfolding J_def fk kn by fastforce have "strand_sem_stateful {} {} (unlabel A) J" using 2 strand_sem_append_stateful by force hence "strand_sem_stateful {} {} (unlabel (A@[\\, send\[occurs (Var y')]\\, \n, send\[\f [\c\\<^sub>c, Var y']\\<^sub>t]\\, \n, \Var y' not in \d\\<^sub>s\\])) J" using 2 3 4 strand_sem_append_stateful[of "{}" "{}" _ _ J] unfolding unlabel_def ik\<^sub>s\<^sub>s\<^sub>t_def by force hence "strand_sem_stateful {} {} (unlabel (B@[\n, send\[attack\ln n\]\\])) J" using strand_sem_receive_send_append[of "{}" "{}" _ J "attack\ln n\"] strand_sem_append_stateful[of "{}" "{}" _ _ J] unfolding B_def Tatt' by simp hence "welltyped_constraint_model J (B@[\n, send\[attack\ln n\]\\])" using B J unfolding defs' by blast thus False using B(1) P_sec by blast qed qed lemma welltyped_leakage_free_short_term_secretI': fixes c x y f n d l l' \ defines "s \ \f [\c\\<^sub>c, Var (TAtom \,x)]\\<^sub>t" and "Tatt \ Transaction (\(). []) [] [(l, receive\[\f [\c\\<^sub>c, Var (TAtom \,y)]\\<^sub>t]\)] [(l', \Var (TAtom \,y) not in \d\\<^sub>s\)] [] [\n, send\[attack\ln n\]\\]" assumes P: "\T \ set P. wellformed_transaction T" "\T \ set P. \x \ set (unlabel (transaction_updates T)). is_Update x \ is_Fun (the_set_term x)" and subterms_sec: "\\ \ reachable_constraints P. \\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[\\, send\[\c\\<^sub>c]\\])" and P_sec: "\\ \ reachable_constraints P. \\\<^sub>\. welltyped_constraint_model \\<^sub>\ (\@[\n, send\[attack\ln n\]\\])" and P_Tatt: "Tatt \ set P" and P_d: "\T \ set P. \a \ set (transaction_updates T). is_Insert (snd a) \ the_set_term (snd a) = \d\\<^sub>s \ transaction_send T \ [] \ (let (l,b) = hd (transaction_send T) in l = \ \ is_Send b \ \f [\c\\<^sub>c, the_elem_term (snd a)]\\<^sub>t \ set (the_msgs b))" shows "welltyped_leakage_free_protocol [s] P" proof - define Q where "Q \ \M t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}" define Sec where "Sec \ {t \ \ | t \. Q (set [s]) t \} - {m. {} \\<^sub>c m}" define f' where "f' \ \(T,\,\::('fun,'atom,'sets,'lbl) prot_subst,\). dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" define g' where "g' \ concat \ map f'" note defs = Sec_def Q_def note defs' = welltyped_constraint_model_def constraint_model_def show ?thesis proof (rule ccontr) assume "\welltyped_leakage_free_protocol [s] P" then obtain A I k where A: "A \ reachable_constraints P" "\f [\c\\<^sub>c, k]\\<^sub>t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I" "welltyped_constraint_model I (A@[\\, send\[\f [\c\\<^sub>c, k]\\<^sub>t]\\])" unfolding welltyped_leakage_free_protocol_def defs s_def by fastforce have I: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" using A(3) unfolding defs' by (blast,blast,blast) note A' = welltyped_constraint_model_prefix[OF A(3)] have "strand_sem_stateful {} {} (unlabel A) I" using A' unfolding defs' by simp hence A'': "strand_sem_stateful {} {} (unlabel A) (I(z := k))" when z: "z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" for z using z strand_sem_model_swap[of "unlabel A" I "I(z := k)" "{}" "{}"] by auto obtain \ where \: "\ (TAtom \,x) = k" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "fv (\ (TAtom \,x)) = {}" using A(2) unfolding defs s_def by auto have k: "\ k = TAtom \" "fv k = {}" "wf\<^sub>t\<^sub>r\<^sub>m k" subgoal using \(1) wt_subst_trm''[OF \(2), of "Var (TAtom \,x)"] by simp subgoal using \(1,4) by blast subgoal using \(1,3) by force done then obtain fk where fk: "k = Fun fk []" using const_type_inv_wf by (cases k) auto have fk': "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ \f [\c\\<^sub>c, k]\\<^sub>t" using fk welltyped_constraint_model_deduct_split(2)[OF A(3)] by auto obtain \::"('fun,'atom,'sets,'lbl) prot_subst" where \: "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" unfolding transaction_renaming_subst_def by blast obtain y' yn' where y': "\ (TAtom \,y) = Var y'" "y \ yn'" "y' = (TAtom \,yn')" using transaction_renaming_subst_is_renaming(1,2)[OF \] by force define B where "B \ A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand Tatt \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" define J where "J \ I(y' := k)" have "y' \ range_vars \" using y'(1) transaction_renaming_subst_is_renaming(3)[OF \] by (metis (no_types, lifting) in_mono subst_fv_imgI term.set_intros(3)) hence y'': "y' \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using transaction_renaming_subst_vars_disj(6)[OF \] by blast have 0: "(k,\d\\<^sub>s) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I)" proof assume a: "(k,\d\\<^sub>s) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I)" obtain l t t' where t: "(l,insert\t,t'\) \ set A" "t \ I = k" "t' \ I = \d\\<^sub>s" using db\<^sub>s\<^sub>s\<^sub>t_in_cases[OF a[unfolded db\<^sub>s\<^sub>s\<^sub>t_def]] unfolding unlabel_def by auto obtain T b B \ \ \ where T: "prefix (B@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ \\<^sub>s \)) A" "T \ set P" "transaction_decl_subst \ T" "transaction_fresh_subst \ T (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "transaction_renaming_subst \ P (vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" "b \ set (transaction_strand T)" "(l, insert\t,t'\) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \ \\<^sub>s \" "fst (l, insert\t,t'\) = fst b" using reachable_constraints_transaction_action_obtain[OF A(1) t(1)] by metis define \ where "\ \ \ \\<^sub>s \ \\<^sub>s \" obtain b' where "b = (l,b')" using T(8) by (cases b) simp then obtain tb tb' where b': "b = (l,insert\tb,tb'\)" and tb: "t = tb \ \" and tb': "t' = tb' \ \" using T(7) unfolding \_def by (cases b') auto note T_wf = bspec[OF P(1) T(2)] bspec[OF P(2) T(2)] have b: "b \ set (transaction_updates T)" using transaction_strand_memberD[OF T(6)[unfolded b']] wellformed_transaction_cases[OF T_wf(1)] unfolding b' by blast have "is_Fun tb'" using bspec[OF P(2) T(2)] wellformed_transaction_strand_unlabel_memberD(6)[ OF T_wf(1) unlabel_in[OF T(6)[unfolded b']]] by fastforce hence *: "tb' = \d\\<^sub>s" using t(3) unfolding b' tb' by force have "is_Insert (snd b)" "the_set_term (snd b) = \d\\<^sub>s" "the_elem_term (snd b) = tb" unfolding b' * by simp_all hence "transaction_send T \ []" "let (l, a) = hd (transaction_send T) in l = \ \ is_Send a \ \f [\c\\<^sub>c, tb]\\<^sub>t \ set (the_msgs a)" using P_d T(2) b by (fast,fast) hence "\ts. \\,send\ts\\ \ set (transaction_send T) \ \f [\c\\<^sub>c, tb]\\<^sub>t \ set ts" unfolding is_Send_def by (cases "transaction_send T") auto then obtain ts where ts: "\\,send\ts\\ \ set (transaction_strand T)" "\f [\c\\<^sub>c, tb]\\<^sub>t \ set ts" unfolding transaction_strand_def by auto have "\\,receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\ \ set A" "\f [\c\\<^sub>c, t]\\<^sub>t \ set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using subst_lsst_memI[OF ts(1), of \] subst_set_map[OF ts(2), of \] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(1)[of \ "ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] set_mono_prefix[OF T(1)[unfolded \_def[symmetric]]] unfolding tb by auto hence "\f [\c\\<^sub>c, k]\\<^sub>t \ \{set ts |ts. \\, receive\ts\\ \ set (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)}" using t(2) subst_lsst_memI[of "\\,receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\" A I] by force thus False using A(2) unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto qed have "y' \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" using y'' fv_ik_subset_vars_sst'[of "unlabel A"] by blast hence 1: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J" unfolding J_def by (metis (no_types, lifting) fv_subset image_cong in_mono repl_invariance) have "(k,\d\\<^sub>s) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) I {}" using 0 db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of "unlabel A" I "[]"] unfolding db\<^sub>s\<^sub>s\<^sub>t_def by force hence "(k,\d\\<^sub>s) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) J {}" using y'' vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel A"] db\<^sub>s\<^sub>s\<^sub>t_subst_swap[of "unlabel A" I J "[]"] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of "unlabel A" _ "[]"] unfolding db\<^sub>s\<^sub>s\<^sub>t_def J_def by (metis (no_types, lifting) Un_iff empty_set fun_upd_other) hence "((Var y' \ J, \d\\<^sub>s \ J) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) J {})" unfolding J_def fk by simp hence "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J) (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) J {}) (unlabel [\n, \Var y' not in \d\\<^sub>s\\]) J" using stateful_strand_sem_NegChecks_no_bvars(1)[ of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J" "dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel A) J {}" "Var y'" "\d\\<^sub>s" J] by simp hence 2: "strand_sem_stateful {} {} (unlabel (A@[\n, \Var y' not in \d\\<^sub>s\\])) J" using A'' y' y'' vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel A"] strand_sem_append_stateful[ of "{}" "{}" "unlabel A" "unlabel [\n, \Var y' not in \d\\<^sub>s\\]" J] unfolding J_def by simp have B: "B \ reachable_constraints P" using reachable_constraints.step[OF A(1) P_Tatt _ _ \, of Var Var] unfolding B_def Tatt_def transaction_decl_subst_def transaction_fresh_subst_def by simp have Tatt': "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand Tatt \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = [(l, send\[\f [\c\\<^sub>c, Var y']\\<^sub>t]\), (l', \Var y' not in \d\\<^sub>s\), \n, receive\[attack\ln n\]\\]" using y' unfolding Tatt_def transaction_strand_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def subst_apply_labeled_stateful_strand_def by auto have J: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t J" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t J" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range J)" unfolding J_def subgoal using wt_subst_subst_upd[OF I(1)] k(1) y'(3) by simp subgoal by (metis I(2) k(2) fun_upd_apply interpretation_grounds_all interpretation_substI) subgoal using I(3) k(3) by fastforce done have 3: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t J \ \f [\c\\<^sub>c, Var y']\\<^sub>t \ J" using 1 fk fk' unfolding J_def by auto have "strand_sem_stateful {} {} (unlabel A) J" using 2 strand_sem_append_stateful by force hence "strand_sem_stateful {} {} (unlabel (A@[\n, send\[\f [\c\\<^sub>c, Var y']\\<^sub>t]\\, \n, \Var y' not in \d\\<^sub>s\\])) J" using 2 3 strand_sem_append_stateful[of "{}" "{}" _ _ J] unfolding unlabel_def ik\<^sub>s\<^sub>s\<^sub>t_def by force hence "strand_sem_stateful {} {} (unlabel (B@[\n, send\[attack\ln n\]\\])) J" using strand_sem_receive_send_append[of "{}" "{}" _ J "attack\ln n\"] strand_sem_append_stateful[of "{}" "{}" _ _ J] unfolding B_def Tatt' by simp hence "welltyped_constraint_model J (B@[\n, send\[attack\ln n\]\\])" using B J unfolding defs' by blast thus False using B(1) P_sec by blast qed qed definition welltyped_leakage_free_invkey_conditions' where "welltyped_leakage_free_invkey_conditions' invfun privfunsec declassifiedset S n P \ let f = \s. is_Var s \ fst (the_Var s) = TAtom Value; g = \s. is_Fun s \ args s = [] \ is_Set (the_Fun s) \ arity\<^sub>s (the_Set (the_Fun s)) = 0; h = \s. is_Fun s \ args s = [] \ is_Fu (the_Fun s) \ public\<^sub>f (the_Fu (the_Fun s)) \ arity\<^sub>f (the_Fu (the_Fun s)) = 0 in (\s\set S. f s \ (is_Fun s \ the_Fun s = Pair \ length (args s) = 2 \ g (args s ! 1)) \ g s \ h s \ s = \privfunsec\\<^sub>c \ s = Fun OccursSec [] \ (is_Fun s \ the_Fun s = OccursFact \ length (args s) = 2 \ args s ! 0 = Fun OccursSec []) \ (is_Fun s \ the_Fun s = Fu invfun \ length (args s) = 2 \ args s ! 0 = \privfunsec\\<^sub>c \ f (args s ! 1)) \ (is_Fun s \ is_Fu (the_Fun s) \ fv s = {} \ Transaction (\(). []) [] [\n, receive\[s]\\] [] [] [\n, send\[attack\ln n\]\\]\set P)) \ (\public\<^sub>f privfunsec \ arity\<^sub>f privfunsec = 0 \ \\<^sub>f privfunsec \ None) \ (\T\set P. transaction_fresh T \ [] \ transaction_send T \ [] \ (let (l, a) = hd (transaction_send T) in l = \ \ is_Send a \ Var ` set (transaction_fresh T) \ set (the_msgs a))) \ (\T\set P. \x\vars_transaction T. is_Var (\\<^sub>v x)) \ (\T\set P. \x\set (transaction_fresh T). \\<^sub>v x = Var Value \ (\a. \\<^sub>v x = \a\\<^sub>\\<^sub>a)) \ (\T\set P. \f\\(funs_term ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). \is_Set f) \ (\T\set P. \r\set (transaction_send T). OccursFact \ \(funs_term ` trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd r)) \ has_LabelS r) \ (\T\set P. \t\subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). \privfunsec\\<^sub>c \ set (snd (Ana t))) \ (\T\set P. \privfunsec\\<^sub>c \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \ (\T\set P. \a\set (transaction_updates T). is_Insert (snd a) \ the_set_term (snd a) = \declassifiedset\\<^sub>s \ transaction_send T \ [] \ (let (l, b) = hd (transaction_send T) in l = \ \ is_Send b \ \invfun [\privfunsec\\<^sub>c, the_elem_term (snd a)]\\<^sub>t \ set (the_msgs b)))" definition welltyped_leakage_free_invkey_conditions where "welltyped_leakage_free_invkey_conditions invfun privfunsec declassifiedset S n P \ let Tatt = \R. Transaction (\(). []) [] (R@[\n, receive\[\invfun [\privfunsec\\<^sub>c, \0: value\\<^sub>v]\\<^sub>t]\\]) [\\, \\0: value\\<^sub>v not in \declassifiedset\\<^sub>s\\] [] [\n, send\[attack\ln n\]\\] in welltyped_leakage_free_invkey_conditions' invfun privfunsec declassifiedset S n P \ has_initial_value_producing_transaction P \ (if Tatt [\\, receive\[occurs \0: value\\<^sub>v]\\] \ set P then \T\set P. admissible_transaction' T \ admissible_transaction_occurs_checks T else Tatt [] \ set P \ (\T\set P. wellformed_transaction T) \ (\T\set P. admissible_transaction_terms T) \ (\T\set P. bvars_transaction T = {}) \ (\T\set P. transaction_decl T () = []) \ (\T\set P. \x\set (transaction_fresh T). let \ = fst x in \ = TAtom Value \ \ \ \ \privfunsec\\<^sub>c) \ (\T\set P. \x\vars_transaction T. let \ = fst x in is_Var \ \ (the_Var \ = Value \ is_Atom (the_Var \)) \ \ \ \ \privfunsec\\<^sub>c) \ (\T\set P. \t\subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). Fun OccursSec [] \ set (snd (Ana t))) \ (\T\set P. Fun OccursSec [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \ (\T\set P. \x\set (unlabel (transaction_updates T)). is_Update x \ is_Fun (the_set_term x)) \ (\s\set S. is_Fun s \ the_Fun s \ OccursFact))" lemma welltyped_leakage_free_invkeyI: assumes P_wt_secure: "\\ \ reachable_constraints P. \\. welltyped_constraint_model \ (\@[\n, send\[attack\ln n\]\\])" and a: "welltyped_leakage_free_invkey_conditions invfun privsec declassset S n P" shows "welltyped_leakage_free_protocol S P" proof - let ?Tatt' = "\R C. Transaction (\(). []) [] R C [] [\n, send\[attack\ln n\]\\] ::('fun,'atom,'sets,'lbl) prot_transaction" let ?Tatt = "\R. ?Tatt' (R@[\n, receive\[\invfun [\privsec\\<^sub>c, \0: value\\<^sub>v]\\<^sub>t]\\]) [\\, \\0: value\\<^sub>v not in \declassset\\<^sub>s\\]" define Tatt1 where "Tatt1 \ ?Tatt [\\, receive\[occurs \0: value\\<^sub>v]\\]" define Tatt2 where "Tatt2 \ ?Tatt []" define Tatt_lts where "Tatt_lts \ \s. ?Tatt' [\n, receive\[s]\\] []" note defs = welltyped_leakage_free_invkey_conditions_def Let_def note defs' = defs welltyped_leakage_free_invkey_conditions'_def note Tatts = Tatt1_def Tatt2_def Tatt_lts_def obtain at where 0: "\public\<^sub>f privsec" "arity\<^sub>f privsec = 0" "\\<^sub>f privsec = Some at" using a unfolding defs' by fast have *: "\T\set P. admissible_transaction' T" "\T\set P. admissible_transaction_occurs_checks T" when "Tatt1 \ set P" using a that unfolding defs Tatts by (meson,meson) have **: "Tatt1 \ set P \ Tatt2 \ set P" using a unfolding defs Tatts by argo have ***: "\T\set P. \x\set (transaction_fresh T). \\<^sub>v x = TAtom Value \ \\<^sub>v x \ \ \privsec\\<^sub>c" "\T\set P. \x\vars_transaction T. \a. \\<^sub>v x = TAtom a \ (a = Value \ (\b. a = Atom b)) \ TAtom a \ \ \privsec\\<^sub>c" when "Tatt1 \ set P" subgoal using a that \\<^sub>v_TAtom''(2) unfolding defs Tatts by metis subgoal using a that \\<^sub>v_TAtom''(1,2) unfolding defs Tatts[symmetric] is_Atom_def is_Var_def by fastforce done have ****: "s \ occurs x" when "Tatt1 \ set P" "s \ set S" for s x using a that ** unfolding defs Tatts the_Fun_def by fastforce have 1: "\T\set P. transaction_fresh T \ [] \ transaction_send T \ [] \ (let (l, a) = hd (transaction_send T) in l = \ \ is_Send a \ Var ` set (transaction_fresh T) \ set (the_msgs a))" "\T\set P. \x\vars_transaction T. is_Var (\\<^sub>v x)" "\T\set P. \x\set (transaction_fresh T). \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = \a\\<^sub>\\<^sub>a)" "\T\set P. \f\\(funs_term ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). \is_Set f" "\T\set P. \r\set (transaction_send T). OccursFact \ \(funs_term ` trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd r)) \ has_LabelS r" "\T\set P. \t\subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). \privsec\\<^sub>c \ set (snd (Ana t))" "\T\set P. \privsec\\<^sub>c \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "\T\set P. \a\set (transaction_updates T). is_Insert (snd a) \ the_set_term (snd a) = \declassset\\<^sub>s \ transaction_send T \ [] \ (let (l, b) = hd (transaction_send T) in l = \ \ is_Send b \ \invfun [\privsec\\<^sub>c, the_elem_term (snd a)]\\<^sub>t \ set (the_msgs b))" using a unfolding defs' by (meson,meson,meson,meson,meson,meson,meson,meson) have 2: "\T\set P. wellformed_transaction T" "\T\set P. \x\vars_transaction T \ set (transaction_fresh T). \ \privsec\\<^sub>c \ \\<^sub>v x" "\T\set P. admissible_transaction_terms T" "\T\set P. \x\set (transaction_fresh T). \\<^sub>v x = TAtom Value" "\T\set P. transaction_decl T () = []" "\T\set P. \x\vars_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = \a\\<^sub>\\<^sub>a)" "\T\set P. \x\vars_transaction T \ set (transaction_fresh T). \\<^sub>v x \ TAtom SetType" "\T\set P. \x\vars_transaction T \ set (transaction_fresh T). \\<^sub>v x \ TAtom OccursSecType" "\T\set P. \t\subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). Fun OccursSec [] \ set (snd (Ana t))" "\T\set P. Fun OccursSec [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "\T\set P. bvars_transaction T = {}" "\T\set P. \x\set (unlabel (transaction_updates T)). is_Update x \ is_Fun (the_set_term x)" subgoal using a * unfolding defs by (metis admissible_transaction_is_wellformed_transaction(1)) subgoal apply (cases "Tatt1 \ set P") subgoal using a * admissible_transactionE(2,3) \_Fu_simps(4) unfolding defs Tatts by force subgoal using a \_Fu_simps(4) unfolding defs Tatts by fastforce done subgoal using a * admissible_transaction_is_wellformed_transaction(4) unfolding defs by metis subgoal using a * admissible_transactionE(2) unfolding defs Tatts[symmetric] by fastforce subgoal using a * admissible_transactionE(1) unfolding defs Tatts[symmetric] by metis subgoal using * *** admissible_transactionE(3) by fast subgoal using * *** admissible_transactionE(2,3) by (cases "Tatt1 \ set P") (force, fastforce) subgoal using * *** admissible_transactionE(2,3) by (cases "Tatt1 \ set P") (force, fastforce) subgoal using a * admissible_transaction_occurs_checksE6 unfolding defs by metis subgoal using a * admissible_transaction_occurs_checksE5 unfolding defs by metis subgoal using a * admissible_transaction_no_bvars(2) unfolding defs Tatts[symmetric] by fastforce subgoal using a * admissible_transaction_is_wellformed_transaction(3) unfolding defs Tatts[symmetric] admissible_transaction_updates_def by fastforce done have Tatt_lts_case: "\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ inj_on \ (fv s) \ \ ` fv s \ range Var \ ?Tatt' [\n, receive\[s \ \]\\] [] \ set P" when s: "fv s = {}" "Tatt_lts s \ set P" for s proof - have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" "inj_on Var (fv s)" "Var ` fv s \ range Var" "s \ Var = s" using s(1) by simp_all thus ?thesis using s(2) unfolding Tatts by metis qed have Tatt1_case: "?Tatt' [\\, receive\[occurs \0: value\\<^sub>v]\\, \n, receive\[\invfun [\privsec\\<^sub>c, \0: value\\<^sub>v]\\<^sub>t]\\] [\\, \\0: value\\<^sub>v not in \declassset\\<^sub>s\\] \ set P" when "Tatt1 \ set P" using that unfolding Tatts by auto have Tatt2_case: "?Tatt' [\n, receive\[\invfun [\privsec\\<^sub>c, \0: value\\<^sub>v]\\<^sub>t]\\] [\\, \\0: value\\<^sub>v not in \declassset\\<^sub>s\\] \ set P" when "Tatt2 \ set P" using that unfolding Tatts by auto note 3 = pair_def case_prod_conv note 4 = welltyped_leakage_free_priv_constI'[OF 0(1-3) 2(1,2) 1(2,3,6,7)] note 5 = welltyped_leakage_free_setop_pairI[OF 2(1,6) 1(4) 2(4,5,3), unfolded 3] welltyped_leakage_free_set_constI(2)[OF 2(1) 1(4) 2(7) 1(2,3), unfolded 3] welltyped_leakage_free_pub_constI 4(2) welltyped_leakage_free_occurssec_constI(2)[OF 2(1,8-10) 1(2,3)] welltyped_leakage_free_value_constI[OF 2(1,3,5,11) 1(1)] welltyped_leakage_free_short_term_secretI'[ OF 2(1,12) 4(1) P_wt_secure Tatt2_case 1(8)] welltyped_leakage_free_long_term_secretI[OF P_wt_secure Tatt_lts_case] welltyped_leakage_free_short_term_secretI[ OF * 1(3) 4(1) P_wt_secure Tatt1_case 1(8)] welltyped_leakage_free_occurs_factI[OF * 1(5)] ** **** have 6: "is_Fun s \ length (args s) = 2 \ (\f t u. s = Fun f [t, u])" for s::"('fun,'atom,'sets,'lbl) prot_term" by auto define pubconst_cond where "pubconst_cond \ \s::('fun,'atom,'sets,'lbl) prot_term. is_Fun s \ args s = [] \ is_Fu (the_Fun s) \ public\<^sub>f (the_Fu (the_Fun s)) \ arity\<^sub>f (the_Fu (the_Fun s)) = 0" define valuevar_cond where "valuevar_cond \ \s::('fun,'atom,'sets,'lbl) prot_term. is_Var s \ fst (the_Var s) = TAtom Value" define setconst_cond where "setconst_cond \ \s::('fun,'atom,'sets,'lbl) prot_term. is_Fun s \ args s = [] \ is_Set (the_Fun s) \ arity\<^sub>s (the_Set (the_Fun s)) = 0" define setop_pair_cond where "setop_pair_cond \ \s. is_Fun s \ the_Fun s = Pair \ length (args s) = 2 \ setconst_cond (args s ! 1)" define occursfact_cond where "occursfact_cond \ \s::('fun,'atom,'sets,'lbl) prot_term. is_Fun s \ the_Fun s = OccursFact \ length (args s) = 2 \ args s ! 0 = Fun OccursSec []" define invkey_cond where "invkey_cond \ \s. is_Fun s \ the_Fun s = Fu invfun \ length (args s) = 2 \ args s ! 0 = \privsec\\<^sub>c \ valuevar_cond (args s ! 1)" define ground_lts_cond where "ground_lts_cond \ \s. is_Fun s \ is_Fu (the_Fun s) \ fv s = {} \ Tatt_lts s \ set P" note cond_defs = pubconst_cond_def valuevar_cond_def setconst_cond_def setop_pair_cond_def occursfact_cond_def invkey_cond_def ground_lts_cond_def have "(\m. s = \m: value\\<^sub>v) \ valuevar_cond s" "(\x c. arity\<^sub>s c = 0 \ s = Fun Pair [x, \c\\<^sub>s]) \ setop_pair_cond s" "(\c. arity\<^sub>s c = 0 \ s = \c\\<^sub>s) \ setconst_cond s" "(\c. public\<^sub>f c \ arity\<^sub>f c = 0 \ s = \c\\<^sub>c) \ pubconst_cond s" "(\x. s = occurs x) \ occursfact_cond s" "(\x. s = \invfun [\privsec\\<^sub>c, \x: value\\<^sub>v]\\<^sub>t) \ invkey_cond s" "(\f ts. s = \f ts\\<^sub>t \ fv s = {} \ Tatt_lts s \ set P) \ ground_lts_cond s" for s::"('fun,'atom,'sets,'lbl) prot_term" unfolding is_Set_def the_Set_def is_Fu_def cond_defs by (fastforce,use 6[of s] in fastforce,fastforce,force,fastforce,fastforce,fastforce) moreover have "(\s \ set S. valuevar_cond s \ setop_pair_cond s \ setconst_cond s \ pubconst_cond s \ s = \privsec\\<^sub>c \ s = Fun OccursSec [] \ occursfact_cond s \ invkey_cond s \ ground_lts_cond s) \ (\public\<^sub>f privsec \ arity\<^sub>f privsec = 0 \ \\<^sub>f privsec \ None)" using a unfolding defs' cond_defs Tatts by meson ultimately have 7: "\s \ set S. (\x c. arity\<^sub>s c = 0 \ s = Fun Pair [x, \c\\<^sub>s]) \ (\c. arity\<^sub>s c = 0 \ s = \c\\<^sub>s) \ (\c. public\<^sub>f c \ arity\<^sub>f c = 0 \ s = \c\\<^sub>c) \ s = \privsec\\<^sub>c \ s = Fun OccursSec [] \ (\m. s = \m: value\\<^sub>v) \ (\x. s = occurs x) \ (\x. s = \invfun [\privsec\\<^sub>c, \x: value\\<^sub>v]\\<^sub>t) \ (\f ts. s = \f ts\\<^sub>t \ fv s = {} \ Tatt_lts s \ set P)" unfolding Let_def by fastforce show ?thesis by (rule iffD2[OF welltyped_leakage_free_protocol_pointwise]; unfold list_all_iff; intro ballI) (use bspec[OF 7] 5 in blast) qed end end locale composable_stateful_protocols = pm: stateful_protocol_model arity\<^sub>f arity\<^sub>s public\<^sub>f Ana\<^sub>f \\<^sub>f label_witness1 label_witness2 for arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets,nat) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"nat" and label_witness2::"nat" + fixes Pc::"('fun,'atom,'sets,nat) prot_transaction list" and Ps Ps_with_star_projs::"('fun,'atom,'sets,nat) prot_transaction list list" and Pc_SMP Sec_symbolic::"('fun,'atom,'sets,nat) prot_term list" and Ps_GSMPs::"(nat \ ('fun,'atom,'sets,nat) prot_term list) list" assumes Pc_def: "Pc = concat Ps" and Ps_with_star_projs_def: "let Pc' = Pc; L = [0..n. (map (transaction_proj n) Pc')) L \ set L = set (remdups (concat ( map (\T. map (the_LabelN \ fst) (filter (Not \ has_LabelS) (transaction_strand T))) Pc')))" and Pc_wellformed_composable: "list_all (list_all (Not \ has_LabelS) \ tl \ transaction_send) Pc" "pm.wellformed_composable_protocols Ps Pc_SMP" "pm.wellformed_composable_protocols' Ps" "pm.composable_protocols Ps Ps_GSMPs Sec_symbolic" begin theorem composed_protocol_preserves_component_goals: assumes components_leakage_free: "list_all (pm.welltyped_leakage_free_protocol Sec_symbolic) Ps_with_star_projs" and n_def: "n < length Ps_with_star_projs" and P_def: "P = Ps_with_star_projs ! n" and P_welltyped_secure: "\\ \ pm.reachable_constraints P. \\. pm.welltyped_constraint_model \ (\@[\n, send\[attack\ln n\]\\])" shows "\\ \ pm.reachable_constraints Pc. \\. pm.constraint_model \ (\@[\n, send\[attack\ln n\]\\])" proof - note 0 = Ps_with_star_projs_def[unfolded Let_def] have 1: "set Ps_with_star_projs = (\n. map (transaction_proj n) Pc) ` set (remdups (concat (map (\T. map (the_LabelN \ fst) (filter (Not \ has_LabelS) (transaction_strand T))) Pc)))" by (metis (mono_tags, lifting) 0 image_set) have 2: "Ps_with_star_projs ! n = map (transaction_proj n) Pc" using conjunct1[OF 0] n_def by fastforce show ?thesis by (rule pm.composable_protocols_par_comp_prot'[ OF Pc_def 1 Pc_wellformed_composable components_leakage_free 2 P_welltyped_secure[unfolded P_def]]) qed end end diff --git a/thys/First_Order_Terms/Subterm_and_Context.thy b/thys/First_Order_Terms/Subterm_and_Context.thy --- a/thys/First_Order_Terms/Subterm_and_Context.thy +++ b/thys/First_Order_Terms/Subterm_and_Context.thy @@ -1,603 +1,603 @@ section \Subterms and Contexts\ text \ We define the (proper) sub- and superterm relations on first order terms, as well as contexts (you can think of contexts as terms with exactly one hole, where we can plug-in another term). Moreover, we establish several connections between these concepts as well as to other concepts such as substitutions. \ theory Subterm_and_Context imports Term "Abstract-Rewriting.Abstract_Rewriting" begin subsection \Subterms\ text \The \<^emph>\superterm\ relation.\ inductive_set supteq :: "(('f, 'v) term \ ('f, 'v) term) set" where refl [simp, intro]: "(t, t) \ supteq" | subt [intro]: "u \ set ss \ (u, t) \ supteq \ (Fun f ss, t) \ supteq" text \The \<^emph>\proper superterm\ relation.\ inductive_set supt :: "(('f, 'v) term \ ('f, 'v) term) set" where arg [simp, intro]: "s \ set ss \ (Fun f ss, s) \ supt" | subt [intro]: "s \ set ss \ (s, t) \ supt \ (Fun f ss, t) \ supt" hide_const suptp supteqp hide_fact suptp.arg suptp.cases suptp.induct suptp.intros suptp.subt suptp_supt_eq hide_fact supteqp.cases supteqp.induct supteqp.intros supteqp.refl supteqp.subt supteqp_supteq_eq hide_fact (open) supt.arg supt.subt supteq.refl supteq.subt subsubsection \Syntactic Sugar\ text \Infix syntax.\ abbreviation "supt_pred s t \ (s, t) \ supt" abbreviation "supteq_pred s t \ (s, t) \ supteq" abbreviation (input) "subt_pred s t \ supt_pred t s" abbreviation (input) "subteq_pred s t \ supteq_pred t s" notation supt ("{\}") and supt_pred ("(_/ \ _)" [56, 56] 55) and subt_pred (infix "\" 55) and supteq ("{\}") and supteq_pred ("(_/ \ _)" [56, 56] 55) and subteq_pred (infix "\" 55) abbreviation subt ("{\}") where "{\} \ {\}\" abbreviation subteq ("{\}") where "{\} \ {\}\" text \Quantifier syntax.\ syntax "_All_supteq" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_Ex_supteq" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_All_supt" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_Ex_supt" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_All_subteq" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_Ex_subteq" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_All_subt" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_Ex_subt" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) (* for parsing *) translations "\x\y. P" \ "\x. x \ y \ P" "\x\y. P" \ "\x. x \ y \ P" "\x\y. P" \ "\x. x \ y \ P" "\x\y. P" \ "\x. x \ y \ P" "\x\y. P" \ "\x. x \ y \ P" "\x\y. P" \ "\x. x \ y \ P" "\x\y. P" \ "\x. x \ y \ P" "\x\y. P" \ "\x. x \ y \ P" (* for printing *) print_translation \ let val All_binder = Mixfix.binder_name @{const_syntax All}; val Ex_binder = Mixfix.binder_name @{const_syntax Ex}; val impl = @{const_syntax "implies"}; val conj = @{const_syntax "conj"}; val supt = @{const_syntax "supt_pred"}; val supteq = @{const_syntax "supteq_pred"}; val trans = [((All_binder, impl, supt), ("_All_supt", "_All_subt")), ((All_binder, impl, supteq), ("_All_supteq", "_All_subteq")), ((Ex_binder, conj, supt), ("_Ex_supt", "_Ex_subt")), ((Ex_binder, conj, supteq), ("_Ex_supteq", "_Ex_subteq"))]; fun matches_bound v t = case t of (Const ("_bound", _) $ Free (v', _)) => (v = v') | _ => false fun contains_var v = Term.exists_subterm (fn Free (x, _) => x = v | _ => false) fun mk x c n P = Syntax.const c $ Syntax_Trans.mark_bound_body x $ n $ P fun tr' q = (q, K (fn [Const ("_bound", _) $ Free (v, T), Const (c, _) $ (Const (d, _) $ t $ u) $ P] => (case AList.lookup (=) trans (q, c, d) of NONE => raise Match | SOME (l, g) => if matches_bound v t andalso not (contains_var v u) then mk (v, T) l u P else if matches_bound v u andalso not (contains_var v t) then mk (v, T) g t P else raise Match) | _ => raise Match)); in [tr' All_binder, tr' Ex_binder] end \ subsubsection \Transitivity Reasoning for Subterms\ lemma supt_trans [trans]: "s \ t \ t \ u \ s \ u" by (induct s t rule: supt.induct) auto lemma trans_supt: "trans {\}" by (auto simp: trans_def dest: supt_trans) lemma supteq_trans [trans]: "s \ t \ t \ u \ s \ u" by (induct s t rule: supteq.induct) (auto) text \Auxiliary lemmas about term size.\ lemma size_simp5: "s \ set ss \ s \ t \ size t < size s \ size t < Suc (size_list size ss)" by (induct ss) auto lemma size_simp6: "s \ set ss \ s \ t \ size t \ size s \ size t \ Suc (size_list size ss)" by (induct ss) auto lemma size_simp1: "t \ set ts \ size t < Suc (size_list size ts)" by (induct ts) auto lemma size_simp2: "t \ set ts \ size t < Suc (Suc (size s + size_list size ts))" by (induct ts) auto lemma size_simp3: assumes "(x, y) \ set (zip xs ys)" shows "size x < Suc (size_list size xs)" using set_zip_leftD [OF assms] size_simp1 by auto lemma size_simp4: assumes "(x, y) \ set (zip xs ys)" shows "size y < Suc (size_list size ys)" using set_zip_rightD [OF assms] using size_simp1 by auto lemmas size_simps = size_simp1 size_simp2 size_simp3 size_simp4 size_simp5 size_simp6 declare size_simps [termination_simp] lemma supt_size: "s \ t \ size s > size t" by (induct rule: supt.induct) (auto simp: size_simps) lemma supteq_size: "s \ t \ size s \ size t" by (induct rule: supteq.induct) (auto simp: size_simps) text \Reflexivity and Asymmetry.\ lemma reflcl_supteq [simp]: "supteq\<^sup>= = supteq" by auto lemma trancl_supteq [simp]: "supteq\<^sup>+ = supteq" by (rule trancl_id) (auto simp: trans_def intro: supteq_trans) lemma rtrancl_supteq [simp]: "supteq\<^sup>* = supteq" unfolding trancl_reflcl[symmetric] by auto lemma eq_supteq: "s = t \ s \ t" by auto lemma supt_neqD: "s \ t \ s \ t" using supt_size by auto lemma supteq_Var [simp]: "x \ vars_term t \ t \ Var x" proof (induct t) case (Var y) then show ?case by (cases "x = y") auto next case (Fun f ss) then show ?case by (auto) qed lemmas vars_term_supteq = supteq_Var lemma term_not_arg [iff]: "Fun f ss \ set ss" (is "?t \ set ss") proof assume "?t \ set ss" then have "?t \ ?t" by (auto) then have "?t \ ?t" by (auto dest: supt_neqD) then show False by simp qed lemma supt_Fun [simp]: assumes "s \ Fun f ss" (is "s \ ?t") and "s \ set ss" shows "False" proof - from \s \ set ss\ have "?t \ s" by (auto) then have "size ?t > size s" by (rule supt_size) from \s \ ?t\ have "size s > size ?t" by (rule supt_size) with \size ?t > size s\ show False by simp qed lemma supt_supteq_conv: "s \ t = (s \ t \ s \ t)" proof assume "s \ t" then show "s \ t \ s \ t" proof (induct rule: supt.induct) case (subt s ss t f) have "s \ s" .. from \s \ set ss\ have "Fun f ss \ s" by (auto) from \s \ t \ s \ t\ have "s \ t" .. with \Fun f ss \ s\ have first: "Fun f ss \ t" by (rule supteq_trans) from \s \ set ss\ and \s \ t\ have "Fun f ss \ t" .. then have second: "Fun f ss \ t" by (auto dest: supt_neqD) from first and second show "Fun f ss \ t \ Fun f ss \ t" by auto qed (auto simp: size_simps) next assume "s \ t \ s \ t" then have "s \ t" and "s \ t" by auto then show "s \ t" by (induct) (auto) qed lemma supt_not_sym: "s \ t \ \ (t \ s)" proof assume "s \ t" and "t \ s" then have "s \ s" by (rule supt_trans) then show False unfolding supt_supteq_conv by simp qed lemma supt_irrefl[iff]: "\ t \ t" using supt_not_sym[of t t] by auto lemma irrefl_subt: "irrefl {\}" by (auto simp: irrefl_def) lemma supt_imp_supteq: "s \ t \ s \ t" unfolding supt_supteq_conv by auto lemma supt_supteq_not_supteq: "s \ t = (s \ t \ \ (t \ s))" using supt_not_sym unfolding supt_supteq_conv by auto lemma supteq_supt_conv: "(s \ t) = (s \ t \ s = t)" by (auto simp: supt_supteq_conv) lemma supteq_antisym: assumes "s \ t" and "t \ s" shows "s = t" using assms unfolding supteq_supt_conv by (auto simp: supt_not_sym) text \The subterm relation is an order on terms.\ interpretation subterm: order "(\)" "(\)" by (unfold_locales) (rule supt_supteq_not_supteq, auto intro: supteq_trans supteq_antisym supt_supteq_not_supteq) text \More transitivity rules.\ lemma supt_supteq_trans[trans]: "s \ t \ t \ u \ s \ u" by (metis subterm.le_less_trans) lemma supteq_supt_trans[trans]: "s \ t \ t \ u \ s \ u" by (metis subterm.less_le_trans) declare subterm.le_less_trans[trans] declare subterm.less_le_trans[trans] lemma suptE [elim]: "s \ t \ (s \ t \ P) \ (s \ t \ P) \ P" by (auto simp: supt_supteq_conv) lemmas suptI [intro] = subterm.dual_order.not_eq_order_implies_strict lemma supt_supteq_set_conv: "{\} = {\} - Id" using supt_supteq_conv [to_set] by auto lemma supteq_supt_set_conv: "{\} = {\}\<^sup>=" by (auto simp: supt_supteq_conv) lemma supteq_imp_vars_term_subset: "s \ t \ vars_term t \ vars_term s" by (induct rule: supteq.induct) auto lemma set_supteq_into_supt [simp]: assumes "t \ set ts" and "t \ s" shows "Fun f ts \ s" proof - from \t \ s\ have "t = s \ t \ s" by auto then show ?thesis proof assume "t = s" with \t \ set ts\ show ?thesis by auto next assume "t \ s" from supt.subt[OF \t \ set ts\ this] show ?thesis . qed qed text \The superterm relation is strongly normalizing.\ lemma SN_supt: "SN {\}" unfolding SN_iff_wf by (rule wf_subset) (auto intro: supt_size) lemma supt_not_refl[elim!]: assumes "t \ t" shows False proof - from assms have "t \ t" by auto then show False by simp qed lemma supteq_not_supt [elim]: assumes "s \ t" and "\ (s \ t)" shows "s = t" using assms by (induct) auto lemma supteq_not_supt_conv [simp]: "{\} - {\} = Id" by auto lemma supteq_subst [simp, intro]: assumes "s \ t" shows "s \ \ \ t \ \" using assms proof (induct rule: supteq.induct) case (subt u ss t f) from \u \ set ss\ have "u \ \ \ set (map (\t. t \ \) ss)" "u \ \ \ u \ \" by auto - then have "Fun f ss \ \ \ u \ \" unfolding subst_apply_term.simps by blast + then have "Fun f ss \ \ \ u \ \" unfolding eval_term.simps by blast from this and \u \ \ \ t \ \\ show ?case by (rule supteq_trans) qed auto lemma supt_subst [simp, intro]: assumes "s \ t" shows "s \ \ \ t \ \" using assms proof (induct rule: supt.induct) case (arg s ss f) then have "s \ \ \ set (map (\t. t \ \) ss)" by simp - then show ?case unfolding subst_apply_term.simps by (rule supt.arg) + then show ?case unfolding eval_term.simps by (rule supt.arg) next case (subt u ss t f) from \u \ set ss\ have "u \ \ \ set (map (\t. t \ \) ss)" by simp - then have "Fun f ss \ \ \ u \ \" unfolding subst_apply_term.simps by (rule supt.arg) + then have "Fun f ss \ \ \ u \ \" unfolding eval_term.simps by (rule supt.arg) with \u \ \ \ t \ \\ show ?case by (metis supt_trans) qed lemma subterm_induct: assumes "\t. \s\t. P s \ P t" shows [case_names subterm]: "P t" by (rule wf_induct[OF wf_measure[of size], of P t], rule assms, insert supt_size, auto) subsection \Contexts\ text \A \<^emph>\context\ is a term containing exactly one \<^emph>\hole\.\ datatype (funs_ctxt: 'f, vars_ctxt: 'v) ctxt = Hole ("\") | More 'f "('f, 'v) term list" "('f, 'v) ctxt" "('f, 'v) term list" text \ We also say that we apply a context~@{term C} to a term~@{term t} when we replace the hole in a @{term C} by @{term t}. \ fun ctxt_apply_term :: "('f, 'v) ctxt \ ('f, 'v) term \ ('f, 'v) term" ("_\_\" [1000, 0] 1000) where "\\s\ = s" | "(More f ss1 C ss2)\s\ = Fun f (ss1 @ C\s\ # ss2)" lemma ctxt_eq [simp]: "(C\s\ = C\t\) = (s = t)" by (induct C) auto fun ctxt_compose :: "('f, 'v) ctxt \ ('f, 'v) ctxt \ ('f, 'v) ctxt" (infixl "\\<^sub>c" 75) where "\ \\<^sub>c D = D" | "(More f ss1 C ss2) \\<^sub>c D = More f ss1 (C \\<^sub>c D) ss2" interpretation ctxt_monoid_mult: monoid_mult "\" "(\\<^sub>c)" proof fix C D E :: "('f, 'v) ctxt" show "C \\<^sub>c D \\<^sub>c E = C \\<^sub>c (D \\<^sub>c E)" by (induct C) simp_all show "\ \\<^sub>c C = C" by simp show "C \\<^sub>c \ = C" by (induct C) simp_all qed instantiation ctxt :: (type, type) monoid_mult begin definition [simp]: "1 = \" definition [simp]: "(*) = (\\<^sub>c)" instance by (intro_classes) (simp_all add: ac_simps) end lemma ctxt_ctxt_compose [simp]: "(C \\<^sub>c D)\t\ = C\D\t\\" by (induct C) simp_all lemmas ctxt_ctxt = ctxt_ctxt_compose [symmetric] text \Applying substitutions to contexts.\ fun subst_apply_ctxt :: "('f, 'v) ctxt \ ('f, 'v, 'w) gsubst \ ('f, 'w) ctxt" (infixl "\\<^sub>c" 67) where "\ \\<^sub>c \ = \" | "(More f ss1 D ss2) \\<^sub>c \ = More f (map (\t. t \ \) ss1) (D \\<^sub>c \) (map (\t. t \ \) ss2)" lemma subst_apply_term_ctxt_apply_distrib [simp]: "C\t\ \ \ = (C \\<^sub>c \)\t \ \\" by (induct C) auto lemma subst_compose_ctxt_compose_distrib [simp]: "(C \\<^sub>c D) \\<^sub>c \ = (C \\<^sub>c \) \\<^sub>c (D \\<^sub>c \)" by (induct C) auto lemma ctxt_compose_subst_compose_distrib [simp]: "C \\<^sub>c (\ \\<^sub>s \) = (C \\<^sub>c \) \\<^sub>c \" by (induct C) (auto) subsection \The Connection between Contexts and the Superterm Relation\ lemma ctxt_imp_supteq [simp]: shows "C\t\ \ t" by (induct C) auto lemma supteq_ctxtE[elim]: assumes "s \ t" obtains C where "s = C\t\" using assms proof (induct arbitrary: thesis) case (refl s) have "s = \\s\" by simp from refl[OF this] show ?case . next case (subt u ss t f) then obtain C where "u = C\t\" by auto from split_list[OF \u \ set ss\] obtain ss1 and ss2 where "ss = ss1 @ u # ss2" by auto then have "Fun f ss = (More f ss1 C ss2)\t\" using \u = C\t\\ by simp with subt show ?case by best qed lemma ctxt_supteq[intro]: assumes "s = C\t\" shows "s \ t" proof (cases C) case Hole then show ?thesis using assms by auto next case (More f ss1 D ss2) with assms have s: "s = Fun f (ss1 @ D\t\ # ss2)" (is "_ = Fun _ ?ss") by simp have "D\t\ \ set ?ss" by simp moreover have "D\t\ \ t" by (induct D) auto ultimately show ?thesis unfolding s .. qed lemma supteq_ctxt_conv: "(s \ t) = (\C. s = C\t\)" by auto lemma supt_ctxtE[elim]: assumes "s \ t" obtains C where "C \ \" and "s = C\t\" using assms proof (induct arbitrary: thesis) case (arg s ss f) from split_list[OF \s \ set ss\] obtain ss1 and ss2 where ss: "ss = ss1 @ s # ss2" by auto let ?C = "More f ss1 \ ss2" have "?C \ \" by simp moreover have "Fun f ss = ?C\s\" by (simp add: ss) ultimately show ?case using arg by best next case (subt s ss t f) then obtain C where "C \ \" and "s = C\t\" by auto from split_list[OF \s \ set ss\] obtain ss1 and ss2 where ss: "ss = ss1 @ s # ss2" by auto have "More f ss1 C ss2 \ \" by simp moreover have "Fun f ss = (More f ss1 C ss2)\t\" using \s = C\t\\ by (simp add: ss) ultimately show ?case using subt(4) by best qed lemma ctxt_supt[intro 2]: assumes "C \ \" and "s = C\t\" shows "s \ t" proof (cases C) case Hole with assms show ?thesis by simp next case (More f ss1 D ss2) with assms have s: "s = Fun f (ss1 @ D\t\ # ss2)" by simp have "D\t\ \ set (ss1 @ D\t\ # ss2)" by simp then have "s \ D\t\" unfolding s .. also have "D\t\ \ t" by (induct D) auto finally show "s \ t" . qed lemma supt_ctxt_conv: "(s \ t) = (\C. C \ \ \ s = C\t\)" (is "_ = ?rhs") proof assume "s \ t" then have "s \ t" and "s \ t" by auto from \s \ t\ obtain C where "s = C\t\" by auto with \s \ t\ have "C \ \" by auto with \s = C\t\\ show "?rhs" by auto next assume "?rhs" then show "s \ t" by auto qed lemma nectxt_imp_supt_ctxt: "C \ \ \ C\t\ \ t" by auto lemma supt_var: "\ (Var x \ u)" proof assume "Var x \ u" then obtain C where "C \ \" and "Var x = C\u\" .. then show False by (cases C) auto qed lemma supt_const: "\ (Fun f [] \ u)" proof assume "Fun f [] \ u" then obtain C where "C \ \" and "Fun f [] = C\u\" .. then show False by (cases C) auto qed lemma supteq_var_imp_eq: "(Var x \ t) = (t = Var x)" (is "_ = (_ = ?x)") proof assume "t = Var x" then show "Var x \ t" by auto next assume st: "?x \ t" from st obtain C where "?x = C\t\" by best then show "t = ?x" by (cases C) auto qed lemma Var_supt [elim!]: assumes "Var x \ t" shows P using assms supt_var[of x t] by simp lemma Fun_supt [elim]: assumes "Fun f ts \ s" obtains t where "t \ set ts" and "t \ s" using assms by (cases) (auto simp: supt_supteq_conv) lemma inj_ctxt_apply_term: "inj (ctxt_apply_term C)" by (auto simp: inj_on_def) lemma ctxt_subst_eq: "(\x. x \ vars_ctxt C \ \ x = \ x) \ C \\<^sub>c \ = C \\<^sub>c \" proof (induct C) case (More f bef C aft) { fix t assume t:"t \ set bef" have "t \ \ = t \ \" using t More(2) by (auto intro: term_subst_eq) } moreover { fix t assume t:"t \ set aft" have "t \ \ = t \ \" using t More(2) by (auto intro: term_subst_eq) } moreover have "C \\<^sub>c \ = C \\<^sub>c \" using More by auto ultimately show ?case by auto qed auto text \ A \<^emph>\signature\ is a set of function symbol/arity pairs, where the arity of a function symbol, denotes the number of arguments it expects. \ type_synonym 'f sig = "('f \ nat) set" text \The set of all function symbol/ arity pairs occurring in a term.\ fun funas_term :: "('f, 'v) term \ 'f sig" where "funas_term (Var _) = {}" | "funas_term (Fun f ts) = {(f, length ts)} \ \(set (map funas_term ts))" lemma supt_imp_funas_term_subset: assumes "s \ t" shows "funas_term t \ funas_term s" using assms by induct auto lemma supteq_imp_funas_term_subset[simp]: assumes "s \ t" shows "funas_term t \ funas_term s" using assms by induct auto text \The set of all function symbol/arity pairs occurring in a context.\ fun funas_ctxt :: "('f, 'v) ctxt \ 'f sig" where "funas_ctxt Hole = {}" | "funas_ctxt (More f ss1 D ss2) = {(f, Suc (length (ss1 @ ss2)))} \ funas_ctxt D \ \(set (map funas_term (ss1 @ ss2)))" lemma funas_term_ctxt_apply [simp]: "funas_term (C\t\) = funas_ctxt C \ funas_term t" by (induct C, auto) lemma funas_term_subst: "funas_term (t \ \) = funas_term t \ \(funas_term ` \ ` vars_term t)" by (induct t) auto lemma funas_ctxt_compose [simp]: "funas_ctxt (C \\<^sub>c D) = funas_ctxt C \ funas_ctxt D" by (induct C) auto lemma funas_ctxt_subst [simp]: "funas_ctxt (C \\<^sub>c \) = funas_ctxt C \ \(funas_term ` \ ` vars_ctxt C)" by (induct C, auto simp: funas_term_subst) end diff --git a/thys/First_Order_Terms/Term.thy b/thys/First_Order_Terms/Term.thy --- a/thys/First_Order_Terms/Term.thy +++ b/thys/First_Order_Terms/Term.thy @@ -1,749 +1,771 @@ (* Author: Christian Sternagel Author: René Thiemann License: LGPL *) section \First-Order Terms\ theory Term imports Main "HOL-Library.Multiset" begin datatype (funs_term : 'f, vars_term : 'v) "term" = is_Var: Var (the_Var: 'v) | Fun 'f (args : "('f, 'v) term list") where "args (Var _) = []" abbreviation "is_Fun t \ \ is_Var t" lemma is_VarE [elim]: "is_Var t \ (\x. t = Var x \ P) \ P" by (cases t) auto lemma is_FunE [elim]: "is_Fun t \ (\f ts. t = Fun f ts \ P) \ P" by (cases t) auto lemma inj_on_Var[simp]: \<^marker>\contributor \Martin Desharnais\\ "inj_on Var A" by (rule inj_onI) simp lemma member_image_the_Var_image_subst: \<^marker>\contributor \Martin Desharnais\\ assumes is_var_\: "\x. is_Var (\ x)" shows "x \ the_Var ` \ ` V \ Var x \ \ ` V" using is_var_\ image_iff by (metis (no_types, opaque_lifting) term.collapse(1) term.sel(1)) lemma image_the_Var_image_subst_renaming_eq: \<^marker>\contributor \Martin Desharnais\\ assumes is_var_\: "\x. is_Var (\ x)" shows "the_Var ` \ ` V = (\x \ V. vars_term (\ x))" proof (rule Set.equalityI; rule Set.subsetI) from is_var_\ show "\x. x \ the_Var ` \ ` V \ x \ (\x\V. vars_term (\ x))" using term.set_sel(3) by force next from is_var_\ show "\x. x \ (\x\V. vars_term (\ x)) \ x \ the_Var ` \ ` V" by (smt (verit, best) Term.term.simps(17) UN_iff image_eqI singletonD term.collapse(1)) qed text \The variables of a term as multiset.\ fun vars_term_ms :: "('f, 'v) term \ 'v multiset" where "vars_term_ms (Var x) = {#x#}" | "vars_term_ms (Fun f ts) = \\<^sub># (mset (map vars_term_ms ts))" lemma set_mset_vars_term_ms [simp]: "set_mset (vars_term_ms t) = vars_term t" by (induct t) auto text \Reorient equations of the form @{term "Var x = t"} and @{term "Fun f ss = t"} to facilitate simplification.\ setup \ Reorient_Proc.add (fn Const (@{const_name Var}, _) $ _ => true | _ => false) #> Reorient_Proc.add (fn Const (@{const_name Fun}, _) $ _ $ _ => true | _ => false) \ simproc_setup reorient_Var ("Var x = t") = \K Reorient_Proc.proc\ simproc_setup reorient_Fun ("Fun f ss = t") = \K Reorient_Proc.proc\ text \The \emph{root symbol} of a term is defined by:\ fun root :: "('f, 'v) term \ ('f \ nat) option" where "root (Var x) = None" | "root (Fun f ts) = Some (f, length ts)" lemma finite_vars_term [simp]: "finite (vars_term t)" by (induct t) simp_all lemma finite_Union_vars_term: "finite (\t \ set ts. vars_term t)" by auto +text \We define the evaluation of terms, under interpretation of function symbols and assignment of + variables, as follows:\ + +fun eval_term ("_\(2_)\_" [999,1,100]100) where + "I\Var x\\ = \ x" +| "I\Fun f ss\\ = I f [I\s\\. s \ ss]" + +notation eval_term ("_\(2_)\" [999,1]100) + +lemma eval_same_vars: + assumes "\x \ vars_term s. \ x = \ x" + shows "I\s\\ = I\s\\" + by (insert assms, induct s, auto intro!:map_cong[OF refl] cong[of "I _"]) + +lemma eval_same_vars_cong: + assumes ref: "s = t" and v: "\x. x \ vars_term s \ \ x = \ x" + shows "I\s\\ = I\t\\" + by (fold ref, rule eval_same_vars, auto dest:v) + +lemma eval_with_fresh_var: "x \ vars_term s \ I\s\\(x:=a) = I\s\\" + by (auto intro: eval_same_vars) + +lemma eval_map_term: "I\map_term ff fv s\\ = (I \ ff)\s\(\ \ fv)" + by (induct s, auto intro: cong[of "I _"]) + + text \A substitution is a mapping \\\ from variables to terms. We call a substitution that alters the type of variables a generalized substitution, since it does not have all properties that are expected of (standard) substitutions (e.g., there is no empty substitution).\ type_synonym ('f, 'v, 'w) gsubst = "'v \ ('f, 'w) term" type_synonym ('f, 'v) subst = "('f, 'v, 'v) gsubst" -fun subst_apply_term :: "('f, 'v) term \ ('f, 'v, 'w) gsubst \ ('f, 'w) term" (infixl "\" 67) - where - "Var x \ \ = \ x" - | "Fun f ss \ \ = Fun f (map (\t. t \ \) ss)" +abbreviation subst_apply_term :: "('f, 'v) term \ ('f, 'v, 'w) gsubst \ ('f, 'w) term" (infixl "\" 67) + where "subst_apply_term \ eval_term Fun" definition subst_compose :: "('f, 'u, 'v) gsubst \ ('f, 'v, 'w) gsubst \ ('f, 'u, 'w) gsubst" (infixl "\\<^sub>s" 75) where "\ \\<^sub>s \ = (\x. (\ x) \ \)" lemma subst_subst_compose [simp]: "t \ (\ \\<^sub>s \) = t \ \ \ \" - by (induct t \ rule: subst_apply_term.induct) (simp_all add: subst_compose_def) + by (induct t) (simp_all add: subst_compose_def) lemma subst_compose_assoc: "\ \\<^sub>s \ \\<^sub>s \ = \ \\<^sub>s (\ \\<^sub>s \)" proof (rule ext) fix x show "(\ \\<^sub>s \ \\<^sub>s \) x = (\ \\<^sub>s (\ \\<^sub>s \)) x" proof - have "(\ \\<^sub>s \ \\<^sub>s \) x = \(x) \ \ \ \" by (simp add: subst_compose_def) also have "\ = \(x) \ (\ \\<^sub>s \)" by simp finally show ?thesis by (simp add: subst_compose_def) qed qed lemma subst_apply_term_empty [simp]: "t \ Var = t" proof (induct t) case (Fun f ts) from map_ext [rule_format, of ts _ id, OF Fun] show ?case by simp qed simp interpretation subst_monoid_mult: monoid_mult "Var" "(\\<^sub>s)" by (unfold_locales) (simp add: subst_compose_assoc, simp_all add: subst_compose_def) lemma term_subst_eq: assumes "\x. x \ vars_term t \ \ x = \ x" shows "t \ \ = t \ \" using assms by (induct t) (auto) lemma term_subst_eq_rev: "t \ \ = t \ \ \ \x \ vars_term t. \ x = \ x" by (induct t) simp_all lemma term_subst_eq_conv: "t \ \ = t \ \ \ (\x \ vars_term t. \ x = \ x)" - using term_subst_eq [of t \ \] and term_subst_eq_rev [of t \ \] by auto + by (auto intro!: term_subst_eq term_subst_eq_rev) lemma subst_term_eqI: assumes "(\t. t \ \ = t \ \)" shows "\ = \" using assms [of "Var x" for x] by (intro ext) simp definition subst_domain :: "('f, 'v) subst \ 'v set" where "subst_domain \ = {x. \ x \ Var x}" fun subst_range :: "('f, 'v) subst \ ('f, 'v) term set" where "subst_range \ = \ ` subst_domain \" lemma vars_term_ms_subst [simp]: "vars_term_ms (t \ \) = (\x\#vars_term_ms t. vars_term_ms (\ x))" (is "_ = ?V t") proof (induct t) case (Fun f ts) have IH: "map (\ t. vars_term_ms (t \ \)) ts = map (\ t. ?V t) ts" by (rule map_cong[OF refl Fun]) show ?case by (simp add: o_def IH, induct ts, auto) qed simp lemma vars_term_ms_subst_mono: assumes "vars_term_ms s \# vars_term_ms t" shows "vars_term_ms (s \ \) \# vars_term_ms (t \ \)" proof - from assms[unfolded mset_subset_eq_exists_conv] obtain u where t: "vars_term_ms t = vars_term_ms s + u" by auto show ?thesis unfolding vars_term_ms_subst unfolding t by auto qed text \The variables introduced by a substitution.\ definition range_vars :: "('f, 'v) subst \ 'v set" where "range_vars \ = \(vars_term ` subst_range \)" lemma mem_range_varsI: \<^marker>\contributor \Martin Desharnais\\ assumes "\ x = Var y" and "x \ y" shows "y \ range_vars \" unfolding range_vars_def UN_iff proof (rule bexI[of _ "Var y"]) show "y \ vars_term (Var y)" by simp next from assms show "Var y \ subst_range \" by (simp_all add: subst_domain_def) qed lemma subst_domain_Var [simp]: "subst_domain Var = {}" by (simp add: subst_domain_def) lemma subst_range_Var[simp]: \<^marker>\contributor \Martin Desharnais\\ "subst_range Var = {}" by simp lemma range_vars_Var[simp]: \<^marker>\contributor \Martin Desharnais\\ "range_vars Var = {}" by (simp add: range_vars_def) lemma subst_apply_term_ident: \<^marker>\contributor \Martin Desharnais\\ "vars_term t \ subst_domain \ = {} \ t \ \ = t" proof (induction t) case (Var x) thus ?case by (simp add: subst_domain_def) next case (Fun f ts) thus ?case by (auto intro: list.map_ident_strong) qed lemma vars_term_subst_apply_term: \<^marker>\contributor \Martin Desharnais\\ "vars_term (t \ \) = (\x \ vars_term t. vars_term (\ x))" by (induction t) (auto simp add: insert_Diff_if subst_domain_def) corollary vars_term_subst_apply_term_subset: \<^marker>\contributor \Martin Desharnais\\ "vars_term (t \ \) \ vars_term t - subst_domain \ \ range_vars \" unfolding vars_term_subst_apply_term proof (induction t) case (Var x) show ?case by (cases "\ x = Var x") (auto simp add: range_vars_def subst_domain_def) next case (Fun f xs) thus ?case by auto qed definition is_renaming :: "('f, 'v) subst \ bool" where "is_renaming \ \ (\x. is_Var (\ x)) \ inj_on \ (subst_domain \)" lemma inv_renaming_sound: \<^marker>\contributor \Martin Desharnais\\ assumes is_var_\: "\x. is_Var (\ x)" and "inj \" shows "\ \\<^sub>s (Var \ (inv (the_Var \ \))) = Var" proof - define \' where "\' = the_Var \ \" have \_def: "\ = Var \ \'" unfolding \'_def using is_var_\ by auto from is_var_\ \inj \\ have "inj \'" unfolding inj_def \_def comp_def by fast hence "inv \' \ \' = id" using inv_o_cancel[of \'] by simp hence "Var \ (inv \' \ \') = Var" by simp hence "\x. (Var \ (inv \' \ \')) x = Var x" by metis hence "\x. ((Var \ \') \\<^sub>s (Var \ (inv \'))) x = Var x" unfolding subst_compose_def by auto thus "\ \\<^sub>s (Var \ (inv \')) = Var" using \_def by auto qed lemma ex_inverse_of_renaming: \<^marker>\contributor \Martin Desharnais\\ assumes "\x. is_Var (\ x)" and "inj \" shows "\\. \ \\<^sub>s \ = Var" using inv_renaming_sound[OF assms] by blast lemma vars_term_subst: "vars_term (t \ \) = \(vars_term ` \ ` vars_term t)" by (induct t) simp_all lemma range_varsE [elim]: assumes "x \ range_vars \" and "\t. x \ vars_term t \ t \ subst_range \ \ P" shows "P" using assms by (auto simp: range_vars_def) lemma range_vars_subst_compose_subset: "range_vars (\ \\<^sub>s \) \ (range_vars \ - subst_domain \) \ range_vars \" (is "?L \ ?R") proof fix x assume "x \ ?L" then obtain y where "y \ subst_domain (\ \\<^sub>s \)" and "x \ vars_term ((\ \\<^sub>s \) y)" by (auto simp: range_vars_def) then show "x \ ?R" proof (cases) assume "y \ subst_domain \" and "x \ vars_term ((\ \\<^sub>s \) y)" moreover then obtain v where "v \ vars_term (\ y)" and "x \ vars_term (\ v)" by (auto simp: subst_compose_def vars_term_subst) ultimately show ?thesis by (cases "v \ subst_domain \") (auto simp: range_vars_def subst_domain_def) qed (auto simp: range_vars_def subst_compose_def subst_domain_def) qed definition "subst x t = Var (x := t)" lemma subst_simps [simp]: "subst x t x = t" "subst x (Var x) = Var" by (auto simp: subst_def) lemma subst_subst_domain [simp]: "subst_domain (subst x t) = (if t = Var x then {} else {x})" proof - { fix y have "y \ {y. subst x t y \ Var y} \ y \ (if t = Var x then {} else {x})" by (cases "x = y", auto simp: subst_def) } then show ?thesis by (simp add: subst_domain_def) qed lemma subst_subst_range [simp]: "subst_range (subst x t) = (if t = Var x then {} else {t})" by (cases "t = Var x") (auto simp: subst_domain_def subst_def) lemma subst_apply_left_idemp [simp]: assumes "\ x = t \ \" shows "s \ subst x t \ \ = s \ \" using assms by (induct s) (auto simp: subst_def) lemma subst_compose_left_idemp [simp]: assumes "\ x = t \ \" shows "subst x t \\<^sub>s \ = \" by (rule subst_term_eqI) (simp add: assms) lemma subst_ident [simp]: assumes "x \ vars_term t" shows "t \ subst x u = t" proof - have "t \ subst x u = t \ Var" by (rule term_subst_eq) (auto simp: assms subst_def) then show ?thesis by simp qed lemma subst_self_idemp [simp]: "x \ vars_term t \ subst x t \\<^sub>s subst x t = subst x t" by (metis subst_simps(1) subst_compose_left_idemp subst_ident) type_synonym ('f, 'v) terms = "('f, 'v) term set" text \Applying a substitution to every term of a given set.\ abbreviation subst_apply_set :: "('f, 'v) terms \ ('f, 'v, 'w) gsubst \ ('f, 'w) terms" (infixl "\\<^sub>s\<^sub>e\<^sub>t" 60) where "T \\<^sub>s\<^sub>e\<^sub>t \ \ (\t. t \ \) ` T" text \Composition of substitutions\ lemma subst_compose: "(\ \\<^sub>s \) x = \ x \ \" by (auto simp: subst_compose_def) lemmas subst_subst = subst_subst_compose [symmetric] lemma subst_apply_eq_Var: assumes "s \ \ = Var x" obtains y where "s = Var y" and "\ y = Var x" using assms by (induct s) auto lemma subst_domain_subst_compose: "subst_domain (\ \\<^sub>s \) = (subst_domain \ - {x. \y. \ x = Var y \ \ y = Var x}) \ (subst_domain \ - subst_domain \)" by (auto simp: subst_domain_def subst_compose_def elim: subst_apply_eq_Var) text \A substitution is idempotent iff the variables in its range are disjoint from its domain. (See also "Term Rewriting and All That" \<^cite>\\Lemma 4.5.7\ in "AllThat"\.)\ lemma subst_idemp_iff: "\ \\<^sub>s \ = \ \ subst_domain \ \ range_vars \ = {}" proof assume "\ \\<^sub>s \ = \" then have "\x. \ x \ \ = \ x \ Var" by simp (metis subst_compose_def) then have *: "\x. \y\vars_term (\ x). \ y = Var y" unfolding term_subst_eq_conv by simp { fix x y assume "\ x \ Var x" and "x \ vars_term (\ y)" with * [of y] have False by simp } then show "subst_domain \ \ range_vars \ = {}" by (auto simp: subst_domain_def range_vars_def) next assume "subst_domain \ \ range_vars \ = {}" then have *: "\x y. \ x = Var x \ \ y = Var y \ x \ vars_term (\ y)" by (auto simp: subst_domain_def range_vars_def) have "\x. \y\vars_term (\ x). \ y = Var y" proof fix x y assume "y \ vars_term (\ x)" with * [of y x] show "\ y = Var y" by auto qed then show "\ \\<^sub>s \ = \" by (simp add: subst_compose_def term_subst_eq_conv [symmetric]) qed lemma subst_compose_apply_eq_apply_lhs: \<^marker>\contributor \Martin Desharnais\\ assumes "range_vars \ \ subst_domain \ = {}" "x \ subst_domain \" shows "(\ \\<^sub>s \) x = \ x" proof (cases "\ x") case (Var y) show ?thesis proof (cases "x = y") case True with Var have \\ x = Var x\ by simp moreover from \x \ subst_domain \\ have "\ x = Var x" by (simp add: disjoint_iff subst_domain_def) ultimately show ?thesis by (simp add: subst_compose_def) next case False have "y \ range_vars \" unfolding range_vars_def UN_iff proof (rule bexI) show "y \ vars_term (Var y)" by simp next from Var False show "Var y \ subst_range \" by (simp_all add: subst_domain_def) qed hence "y \ subst_domain \" using \range_vars \ \ subst_domain \ = {}\ by (simp add: disjoint_iff) with Var show ?thesis unfolding subst_compose_def by (simp add: subst_domain_def) qed next case (Fun f ys) hence "Fun f ys \ subst_range \ \ (\y\set ys. y \ subst_range \)" using subst_domain_def by fastforce hence "\x \ vars_term (Fun f ys). x \ range_vars \" by (metis UN_I range_vars_def term.distinct(1) term.sel(4) term.set_cases(2)) hence "Fun f ys \ \ = Fun f ys \ Var" unfolding term_subst_eq_conv using \range_vars \ \ subst_domain \ = {}\ by (simp add: disjoint_iff subst_domain_def) - hence "Fun f ys \ \ = Fun f ys" - by simp - with Fun show ?thesis + from this[unfolded subst_apply_term_empty] Fun show ?thesis by (simp add: subst_compose_def) qed lemma subst_apply_term_subst_apply_term_eq_subst_apply_term_lhs: \<^marker>\contributor \Martin Desharnais\\ assumes "range_vars \ \ subst_domain \ = {}" and "vars_term t \ subst_domain \ = {}" shows "t \ \ \ \ = t \ \" proof - from assms have "\x. x \ vars_term t \ (\ \\<^sub>s \) x = \ x" using subst_compose_apply_eq_apply_lhs by fastforce hence "t \ \ \\<^sub>s \ = t \ \" - using term_subst_eq_conv[of t "\ \\<^sub>s \" \] by metis + using term_subst_eq_conv by metis thus ?thesis by simp qed fun num_funs :: "('f, 'v) term \ nat" where "num_funs (Var x) = 0" | "num_funs (Fun f ts) = Suc (sum_list (map num_funs ts))" lemma num_funs_0: assumes "num_funs t = 0" obtains x where "t = Var x" using assms by (induct t) auto lemma num_funs_subst: "num_funs (t \ \) \ num_funs t" by (induct t) (simp_all, metis comp_apply sum_list_mono) lemma sum_list_map_num_funs_subst: assumes "sum_list (map (num_funs \ (\t. t \ \)) ts) = sum_list (map num_funs ts)" shows "\i < length ts. num_funs (ts ! i \ \) = num_funs (ts ! i)" using assms proof (induct ts) case (Cons t ts) then have "num_funs (t \ \) + sum_list (map (num_funs \ (\t. t \ \)) ts) = num_funs t + sum_list (map num_funs ts)" by (simp add: o_def) moreover have "num_funs (t \ \) \ num_funs t" by (metis num_funs_subst) moreover have "sum_list (map (num_funs \ (\t. t \ \)) ts) \ sum_list (map num_funs ts)" using num_funs_subst [of _ \] by (induct ts) (auto intro: add_mono) ultimately show ?case using Cons by (auto) (case_tac i, auto) qed simp lemma is_Fun_num_funs_less: assumes "x \ vars_term t" and "is_Fun t" shows "num_funs (\ x) < num_funs (t \ \)" using assms proof (induct t) case (Fun f ts) then obtain u where u: "u \ set ts" "x \ vars_term u" by auto then have "num_funs (u \ \) \ sum_list (map (num_funs \ (\t. t \ \)) ts)" by (intro member_le_sum_list) simp moreover have "num_funs (\ x) \ num_funs (u \ \)" using Fun.hyps [OF u] and u by (cases u; simp) ultimately show ?case by simp qed simp lemma finite_subst_domain_subst: "finite (subst_domain (subst x y))" by simp lemma subst_domain_compose: "subst_domain (\ \\<^sub>s \) \ subst_domain \ \ subst_domain \" by (auto simp: subst_domain_def subst_compose_def) lemma vars_term_disjoint_imp_unifier: fixes \ :: "('f, 'v, 'w) gsubst" assumes "vars_term s \ vars_term t = {}" and "s \ \ = t \ \" shows "\\ :: ('f, 'v, 'w) gsubst. s \ \ = t \ \" proof - let ?\ = "\x. if x \ vars_term s then \ x else \ x" have "s \ \ = s \ ?\" unfolding term_subst_eq_conv by (induct s) (simp_all) moreover have "t \ \ = t \ ?\" using assms(1) unfolding term_subst_eq_conv by (induct s arbitrary: t) (auto) ultimately have "s \ ?\ = t \ ?\" using assms(2) by simp then show ?thesis by blast qed lemma vars_term_subset_subst_eq: assumes "vars_term t \ vars_term s" and "s \ \ = s \ \" shows "t \ \ = t \ \" using assms by (induct t) (induct s, auto) subsection \Restrict the Domain of a Substitution\ definition restrict_subst_domain where \<^marker>\contributor \Martin Desharnais\\ "restrict_subst_domain V \ x \ (if x \ V then \ x else Var x)" lemma restrict_subst_domain_empty[simp]: \<^marker>\contributor \Martin Desharnais\\ "restrict_subst_domain {} \ = Var" unfolding restrict_subst_domain_def by auto lemma restrict_subst_domain_Var[simp]: \<^marker>\contributor \Martin Desharnais\\ "restrict_subst_domain V Var = Var" unfolding restrict_subst_domain_def by auto lemma subst_domain_restrict_subst_domain[simp]: \<^marker>\contributor \Martin Desharnais\\ "subst_domain (restrict_subst_domain V \) = V \ subst_domain \" unfolding restrict_subst_domain_def subst_domain_def by auto lemma subst_apply_term_restrict_subst_domain: \<^marker>\contributor \Martin Desharnais\\ "vars_term t \ V \ t \ restrict_subst_domain V \ = t \ \" by (rule term_subst_eq) (simp add: restrict_subst_domain_def subsetD) subsection \Rename the Domain of a Substitution\ definition rename_subst_domain where \<^marker>\contributor \Martin Desharnais\\ "rename_subst_domain \ \ x = (if Var x \ \ ` subst_domain \ then \ (the_inv \ (Var x)) else Var x)" lemma rename_subst_domain_Var_lhs[simp]: \<^marker>\contributor \Martin Desharnais\\ "rename_subst_domain Var \ = \" by (rule ext) (simp add: rename_subst_domain_def inj_image_mem_iff the_inv_f_f subst_domain_def) lemma rename_subst_domain_Var_rhs[simp]: \<^marker>\contributor \Martin Desharnais\\ "rename_subst_domain \ Var = Var" by (rule ext) (simp add: rename_subst_domain_def) lemma subst_domain_rename_subst_domain_subset: \<^marker>\contributor \Martin Desharnais\\ assumes is_var_\: "\x. is_Var (\ x)" shows "subst_domain (rename_subst_domain \ \) \ the_Var ` \ ` subst_domain \" by (auto simp add: subst_domain_def rename_subst_domain_def member_image_the_Var_image_subst[OF is_var_\]) lemma subst_range_rename_subst_domain_subset: \<^marker>\contributor \Martin Desharnais\\ assumes "inj \" shows "subst_range (rename_subst_domain \ \) \ subst_range \" proof (intro Set.equalityI Set.subsetI) fix t assume "t \ subst_range (rename_subst_domain \ \)" then obtain x where t_def: "t = rename_subst_domain \ \ x" and "rename_subst_domain \ \ x \ Var x" by (auto simp: image_iff subst_domain_def) show "t \ subst_range \" proof (cases \Var x \ \ ` subst_domain \\) case True then obtain x' where "\ x' = Var x" and "x' \ subst_domain \" by auto then show ?thesis using the_inv_f_f[OF \inj \\, of x'] by (simp add: t_def rename_subst_domain_def) next case False hence False using \rename_subst_domain \ \ x \ Var x\ by (simp add: t_def rename_subst_domain_def) thus ?thesis .. qed qed lemma range_vars_rename_subst_domain_subset: \<^marker>\contributor \Martin Desharnais\\ assumes "inj \" shows "range_vars (rename_subst_domain \ \) \ range_vars \" unfolding range_vars_def using subst_range_rename_subst_domain_subset[OF \inj \\] by (metis Union_mono image_mono) lemma renaming_cancels_rename_subst_domain: \<^marker>\contributor \Martin Desharnais\\ assumes is_var_\: "\x. is_Var (\ x)" and "inj \" and vars_t: "vars_term t \ subst_domain \" shows "t \ \ \ rename_subst_domain \ \ = t \ \" unfolding subst_subst proof (intro term_subst_eq ballI) fix x assume "x \ vars_term t" with vars_t have x_in: "x \ subst_domain \" by blast obtain x' where \_x: "\ x = Var x'" using is_var_\ by (meson is_Var_def) with x_in have x'_in: "Var x' \ \ ` subst_domain \" by (metis image_eqI) have "(\ \\<^sub>s rename_subst_domain \ \) x = \ x \ rename_subst_domain \ \" by (simp add: subst_compose_def) also have "\ = rename_subst_domain \ \ x'" using \_x by simp also have "\ = \ (the_inv \ (Var x'))" by (simp add: rename_subst_domain_def if_P[OF x'_in]) also have "\ = \ (the_inv \ (\ x))" by (simp add: \_x) also have "\ = \ x" using \inj \\ by (simp add: the_inv_f_f) finally show "(\ \\<^sub>s rename_subst_domain \ \) x = \ x" by simp qed subsection \Rename the Domain and Range of a Substitution\ definition rename_subst_domain_range where \<^marker>\contributor \Martin Desharnais\\ "rename_subst_domain_range \ \ x = (if Var x \ \ ` subst_domain \ then ((Var o the_inv \) \\<^sub>s \ \\<^sub>s \) (Var x) else Var x)" lemma rename_subst_domain_range_Var_lhs[simp]: \<^marker>\contributor \Martin Desharnais\\ "rename_subst_domain_range Var \ = \" by (rule ext) (simp add: rename_subst_domain_range_def inj_image_mem_iff the_inv_f_f subst_domain_def subst_compose_def) lemma rename_subst_domain_range_Var_rhs[simp]: \<^marker>\contributor \Martin Desharnais\\ "rename_subst_domain_range \ Var = Var" by (rule ext) (simp add: rename_subst_domain_range_def) lemma subst_compose_renaming_rename_subst_domain_range: \<^marker>\contributor \Martin Desharnais\\ fixes \ \ :: "('f, 'v) subst" assumes is_var_\: "\x. is_Var (\ x)" and "inj \" shows "\ \\<^sub>s rename_subst_domain_range \ \ = \ \\<^sub>s \" proof (rule ext) fix x from is_var_\ obtain x' where "\ x = Var x'" by (meson is_Var_def is_renaming_def) with \inj \\ have inv_\_x': "the_inv \ (Var x') = x" by (metis the_inv_f_f) show "(\ \\<^sub>s rename_subst_domain_range \ \) x = (\ \\<^sub>s \) x" proof (cases "x \ subst_domain \") case True hence "Var x' \ \ ` subst_domain \" using \\ x = Var x'\ by (metis imageI) thus ?thesis by (simp add: \\ x = Var x'\ rename_subst_domain_range_def subst_compose_def inv_\_x') next case False hence "Var x' \ \ ` subst_domain \" proof (rule contrapos_nn) assume "Var x' \ \ ` subst_domain \" hence "\ x \ \ ` subst_domain \" unfolding \\ x = Var x'\ . thus "x \ subst_domain \" unfolding inj_image_mem_iff[OF \inj \\] . qed with False \\ x = Var x'\ show ?thesis by (simp add: subst_compose_def subst_domain_def rename_subst_domain_range_def) qed qed corollary subst_apply_term_renaming_rename_subst_domain_range: \<^marker>\contributor \Martin Desharnais\\ \ \This might be easier to find with @{command find_theorems}.\ fixes t :: "('f, 'v) term" and \ \ :: "('f, 'v) subst" assumes is_var_\: "\x. is_Var (\ x)" and "inj \" shows "t \ \ \ rename_subst_domain_range \ \ = t \ \ \ \" unfolding subst_subst unfolding subst_compose_renaming_rename_subst_domain_range[OF assms] by (rule refl) text \A term is called \<^emph>\ground\ if it does not contain any variables.\ fun ground :: "('f, 'v) term \ bool" where "ground (Var x) \ False" | "ground (Fun f ts) \ (\t \ set ts. ground t)" lemma ground_vars_term_empty: "ground t \ vars_term t = {}" by (induct t) simp_all lemma ground_subst [simp]: "ground (t \ \) \ (\x \ vars_term t. ground (\ x))" by (induct t) simp_all lemma ground_subst_apply: assumes "ground t" shows "t \ \ = t" proof - have "t = t \ Var" by simp also have "\ = t \ \" by (rule term_subst_eq, insert assms[unfolded ground_vars_term_empty], auto) finally show ?thesis by simp qed text \Just changing the variables in a term\ abbreviation "map_vars_term f \ term.map_term (\x. x) f" lemma map_vars_term_as_subst: "map_vars_term f t = t \ (\ x. Var (f x))" by (induct t) simp_all lemma map_vars_term_eq: "map_vars_term f s = s \ (Var \ f)" by (induct s) auto lemma ground_map_vars_term [simp]: "ground (map_vars_term f t) = ground t" by (induct t) simp_all lemma map_vars_term_subst [simp]: "map_vars_term f (t \ \) = t \ (\ x. map_vars_term f (\ x))" by (induct t) simp_all lemma map_vars_term_compose: "map_vars_term m1 (map_vars_term m2 t) = map_vars_term (m1 o m2) t" by (induct t) simp_all lemma map_vars_term_id [simp]: "map_vars_term id t = t" by (induct t) (auto intro: map_idI) lemma apply_subst_map_vars_term: "map_vars_term m t \ \ = t \ (\ \ m)" by (induct t) (auto) end diff --git a/thys/First_Order_Terms/Unification.thy b/thys/First_Order_Terms/Unification.thy --- a/thys/First_Order_Terms/Unification.thy +++ b/thys/First_Order_Terms/Unification.thy @@ -1,888 +1,888 @@ (* Author: Christian Sternagel Author: René Thiemann License: LGPL *) subsection \A Concrete Unification Algorithm\ theory Unification imports Abstract_Unification Option_Monad Renaming2 begin definition "decompose s t = (case (s, t) of (Fun f ss, Fun g ts) \ if f = g then zip_option ss ts else None | _ \ None)" lemma decompose_same_Fun[simp]: \<^marker>\contributor \Martin Desharnais\\ "decompose (Fun f ss) (Fun f ss) = Some (zip ss ss)" by (simp add: decompose_def) lemma decompose_Some [dest]: "decompose (Fun f ss) (Fun g ts) = Some E \ f = g \ length ss = length ts \ E = zip ss ts" by (cases "f = g") (auto simp: decompose_def) lemma decompose_None [dest]: "decompose (Fun f ss) (Fun g ts) = None \ f \ g \ length ss \ length ts" by (cases "f = g") (auto simp: decompose_def) text \Applying a substitution to a list of equations.\ definition subst_list :: "('f, 'v) subst \ ('f, 'v) equation list \ ('f, 'v) equation list" where "subst_list \ ys = map (\p. (fst p \ \, snd p \ \)) ys" lemma mset_subst_list [simp]: "mset (subst_list (subst x t) ys) = subst_mset (subst x t) (mset ys)" by (auto simp: subst_mset_def subst_list_def) lemma subst_list_append: "subst_list \ (xs @ ys) = subst_list \ xs @ subst_list \ ys" by (auto simp: subst_list_def) function (sequential) unify :: "('f, 'v) equation list \ ('v \ ('f, 'v) term) list \ ('v \ ('f, 'v) term) list option" where "unify [] bs = Some bs" | "unify ((Fun f ss, Fun g ts) # E) bs = (case decompose (Fun f ss) (Fun g ts) of None \ None | Some us \ unify (us @ E) bs)" | "unify ((Var x, t) # E) bs = (if t = Var x then unify E bs else if x \ vars_term t then None else unify (subst_list (subst x t) E) ((x, t) # bs))" | "unify ((t, Var x) # E) bs = (if x \ vars_term t then None else unify (subst_list (subst x t) E) ((x, t) # bs))" by pat_completeness auto termination by (standard, rule wf_inv_image [of "unif\" "mset \ fst", OF wf_converse_unif]) (force intro: UNIF1.intros simp: unif_def union_commute)+ lemma unify_append_prefix_same: \<^marker>\contributor \Martin Desharnais\\ "(\e \ set es1. fst e = snd e) \ unify (es1 @ es2) bs = unify es2 bs" proof (induction "es1 @ es2" bs arbitrary: es1 es2 bs rule: unify.induct) case (1 bs) thus ?case by simp next case (2 f ss g ts E bs) show ?case proof (cases es1) case Nil thus ?thesis by simp next case (Cons e es1') hence e_def: "e = (Fun f ss, Fun g ts)" and E_def: "E = es1' @ es2" using "2" by simp_all hence "f = g" and "ss = ts" using "2.prems" local.Cons by auto hence "unify (es1 @ es2) bs = unify ((zip ts ts @ es1') @ es2) bs" by (simp add: Cons e_def) also have "\ = unify es2 bs" proof (rule "2.hyps"(1)) show "decompose (Fun f ss) (Fun g ts) = Some (zip ts ts)" by (simp add: \f = g\ \ss = ts\) next show "zip ts ts @ E = (zip ts ts @ es1') @ es2" by (simp add: E_def) next show "\e\set (zip ts ts @ es1'). fst e = snd e" using "2.prems" by (auto simp: Cons zip_same) qed finally show ?thesis . qed next case (3 x t E bs) show ?case proof (cases es1) case Nil thus ?thesis by simp next case (Cons e es1') hence e_def: "e = (Var x, t)" and E_def: "E = es1' @ es2" using 3 by simp_all show ?thesis proof (cases "t = Var x") case True show ?thesis using 3(1)[OF True E_def] using "3.hyps"(3) "3.prems" local.Cons by fastforce next case False thus ?thesis using "3.prems" e_def local.Cons by force qed qed next case (4 v va x E bs) then show ?case proof (cases es1) case Nil thus ?thesis by simp next case (Cons e es1') hence e_def: "e = (Fun v va, Var x)" and E_def: "E = es1' @ es2" using 4 by simp_all thus ?thesis using "4.prems" local.Cons by fastforce qed qed corollary unify_Cons_same: \<^marker>\contributor \Martin Desharnais\\ "fst e = snd e \ unify (e # es) bs = unify es bs" by (rule unify_append_prefix_same[of "[_]", simplified]) corollary unify_same: \<^marker>\contributor \Martin Desharnais\\ "(\e \ set es. fst e = snd e) \ unify es bs = Some bs" by (rule unify_append_prefix_same[of _ "[]", simplified]) definition subst_of :: "('v \ ('f, 'v) term) list \ ('f, 'v) subst" where "subst_of ss = List.foldr (\(x, t) \. \ \\<^sub>s subst x t) ss Var" text \Computing the mgu of two terms.\ definition mgu :: "('f, 'v) term \ ('f, 'v) term \ ('f, 'v) subst option" where "mgu s t = (case unify [(s, t)] [] of None \ None | Some res \ Some (subst_of res))" lemma subst_of_simps [simp]: "subst_of [] = Var" "subst_of ((x, Var x) # ss) = subst_of ss" "subst_of (b # ss) = subst_of ss \\<^sub>s subst (fst b) (snd b)" by (simp_all add: subst_of_def split: prod.splits) lemma subst_of_append [simp]: "subst_of (ss @ ts) = subst_of ts \\<^sub>s subst_of ss" by (induct ss) (auto simp: ac_simps) text \The concrete algorithm \unify\ can be simulated by the inference rules of \UNIF\.\ lemma unify_Some_UNIF: assumes "unify E bs = Some cs" shows "\ds ss. cs = ds @ bs \ subst_of ds = compose ss \ UNIF ss (mset E) {#}" using assms proof (induction E bs arbitrary: cs rule: unify.induct) case (2 f ss g ts E bs) then obtain us where "decompose (Fun f ss) (Fun g ts) = Some us" and [simp]: "f = g" "length ss = length ts" "us = zip ss ts" and "unify (us @ E) bs = Some cs" by (auto split: option.splits) from "2.IH" [OF this(1, 5)] obtain xs ys where "cs = xs @ bs" and [simp]: "subst_of xs = compose ys" and *: "UNIF ys (mset (us @ E)) {#}" by auto then have "UNIF (Var # ys) (mset ((Fun f ss, Fun g ts) # E)) {#}" by (force intro: UNIF1.decomp simp: ac_simps) moreover have "cs = xs @ bs" by fact moreover have "subst_of xs = compose (Var # ys)" by simp ultimately show ?case by blast next case (3 x t E bs) show ?case proof (cases "t = Var x") assume "t = Var x" then show ?case using 3 by auto (metis UNIF.step compose_simps(2) UNIF1.trivial) next assume "t \ Var x" with 3 obtain xs ys where [simp]: "cs = (ys @ [(x, t)]) @ bs" and [simp]: "subst_of ys = compose xs" and "x \ vars_term t" and "UNIF xs (mset (subst_list (subst x t) E)) {#}" by (cases "x \ vars_term t") force+ then have "UNIF (subst x t # xs) (mset ((Var x, t) # E)) {#}" by (force intro: UNIF1.Var_left simp: ac_simps) moreover have "cs = (ys @ [(x, t)]) @ bs" by simp moreover have "subst_of (ys @ [(x, t)]) = compose (subst x t # xs)" by simp ultimately show ?case by blast qed next case (4 f ss x E bs) with 4 obtain xs ys where [simp]: "cs = (ys @ [(x, Fun f ss)]) @ bs" and [simp]: "subst_of ys = compose xs" and "x \ vars_term (Fun f ss)" and "UNIF xs (mset (subst_list (subst x (Fun f ss)) E)) {#}" by (cases "x \ vars_term (Fun f ss)") force+ then have "UNIF (subst x (Fun f ss) # xs) (mset ((Fun f ss, Var x) # E)) {#}" by (force intro: UNIF1.Var_right simp: ac_simps) moreover have "cs = (ys @ [(x, Fun f ss)]) @ bs" by simp moreover have "subst_of (ys @ [(x, Fun f ss)]) = compose (subst x (Fun f ss) # xs)" by simp ultimately show ?case by blast qed force lemma unify_sound: assumes "unify E [] = Some cs" shows "is_imgu (subst_of cs) (set E)" proof - from unify_Some_UNIF [OF assms] obtain ss where "subst_of cs = compose ss" and "UNIF ss (mset E) {#}" by auto with UNIF_empty_imp_is_mgu_compose [OF this(2)] and UNIF_idemp [OF this(2)] show ?thesis by (auto simp add: is_imgu_def is_mgu_def) (metis subst_compose_assoc) qed lemma mgu_sound: assumes "mgu s t = Some \" shows "is_imgu \ {(s, t)}" proof - obtain ss where "unify [(s, t)] [] = Some ss" and "\ = subst_of ss" using assms by (auto simp: mgu_def split: option.splits) then have "is_imgu \ (set [(s, t)])" by (metis unify_sound) then show ?thesis by simp qed text \If \unify\ gives up, then the given set of equations cannot be reduced to the empty set by \UNIF\.\ lemma unify_None: assumes "unify E ss = None" shows "\E'. E' \ {#} \ (mset E, E') \ unif\<^sup>!" using assms proof (induction E ss rule: unify.induct) case (1 bs) then show ?case by simp next case (2 f ss g ts E bs) moreover { assume *: "decompose (Fun f ss) (Fun g ts) = None" have ?case proof (cases "unifiable (set E)") case True then have "(mset E, {#}) \ unif\<^sup>*" by (simp add: unifiable_imp_empty) from unif_rtrancl_mono [OF this, of "{#(Fun f ss, Fun g ts)#}"] obtain \ where "(mset E + {#(Fun f ss, Fun g ts)#}, {#(Fun f ss \ \, Fun g ts \ \)#}) \ unif\<^sup>*" by (auto simp: subst_mset_def) moreover have "{#(Fun f ss \ \, Fun g ts \ \)#} \ NF unif" using decompose_None [OF *] by (auto simp: single_is_union NF_def unif_def elim!: UNIF1.cases) (metis length_map) ultimately show ?thesis by auto (metis normalizability_I add_mset_not_empty) next case False moreover have "\ unifiable {(Fun f ss, Fun g ts)}" using * by (auto simp: unifiable_def) ultimately have "\ unifiable (set ((Fun f ss, Fun g ts) # E))" by (auto simp: unifiable_def unifiers_def) then show ?thesis by (simp add: not_unifiable_imp_not_empty_NF) qed } moreover { fix us assume *: "decompose (Fun f ss) (Fun g ts) = Some us" and "unify (us @ E) bs = None" from "2.IH" [OF this] obtain E' where "E' \ {#}" and "(mset (us @ E), E') \ unif\<^sup>!" by blast moreover have "(mset ((Fun f ss, Fun g ts) # E), mset (us @ E)) \ unif" proof - have "g = f" and "length ss = length ts" and "us = zip ss ts" using * by auto then show ?thesis by (auto intro: UNIF1.decomp simp: unif_def ac_simps) qed ultimately have ?case by auto } ultimately show ?case by (auto split: option.splits) next case (3 x t E bs) { assume [simp]: "t = Var x" obtain E' where "E' \ {#}" and "(mset E, E') \ unif\<^sup>!" using 3 by auto moreover have "(mset ((Var x, t) # E), mset E) \ unif" by (auto intro: UNIF1.trivial simp: unif_def) ultimately have ?case by auto } moreover { assume *: "t \ Var x" "x \ vars_term t" then obtain E' where "E' \ {#}" and "(mset (subst_list (subst x t) E), E') \ unif\<^sup>!" using 3 by auto moreover have "(mset ((Var x, t) # E), mset (subst_list (subst x t) E)) \ unif" using * by (auto intro: UNIF1.Var_left simp: unif_def) ultimately have ?case by auto } moreover { assume *: "t \ Var x" "x \ vars_term t" then have "x \ vars_term t" "is_Fun t" by auto then have "\ unifiable {(Var x, t)}" by (rule in_vars_is_Fun_not_unifiable) then have **: "\ unifiable {(Var x \ \, t \ \)}" for \ :: "('b, 'a) subst" using subst_set_reflects_unifiable [of \ "{(Var x, t)}"] by (auto simp: subst_set_def) have ?case proof (cases "unifiable (set E)") case True then have "(mset E, {#}) \ unif\<^sup>*" by (simp add: unifiable_imp_empty) from unif_rtrancl_mono [OF this, of "{#(Var x, t)#}"] obtain \ where "(mset E + {#(Var x, t)#}, {#(Var x \ \, t \ \)#}) \ unif\<^sup>*" by (auto simp: subst_mset_def) moreover obtain E' where "E' \ {#}" and "({#(Var x \ \, t \ \)#}, E') \ unif\<^sup>!" using not_unifiable_imp_not_empty_NF and ** by (metis set_mset_single) ultimately show ?thesis by auto next case False moreover have "\ unifiable {(Var x, t)}" using * by (force simp: unifiable_def) ultimately have "\ unifiable (set ((Var x, t) # E))" by (auto simp: unifiable_def unifiers_def) then show ?thesis by (simp add: not_unifiable_imp_not_empty_NF) qed } ultimately show ?case by blast next case (4 f ss x E bs) define t where "t = Fun f ss" { assume *: "x \ vars_term t" then obtain E' where "E' \ {#}" and "(mset (subst_list (subst x t) E), E') \ unif\<^sup>!" using 4 by (auto simp: t_def) moreover have "(mset ((t, Var x) # E), mset (subst_list (subst x t) E)) \ unif" using * by (auto intro: UNIF1.Var_right simp: unif_def) ultimately have ?case by (auto simp: t_def) } moreover { assume "x \ vars_term t" then have *: "x \ vars_term t" "t \ Var x" by (auto simp: t_def) then have "x \ vars_term t" "is_Fun t" by auto then have "\ unifiable {(Var x, t)}" by (rule in_vars_is_Fun_not_unifiable) then have **: "\ unifiable {(Var x \ \, t \ \)}" for \ :: "('b, 'a) subst" using subst_set_reflects_unifiable [of \ "{(Var x, t)}"] by (auto simp: subst_set_def) have ?case proof (cases "unifiable (set E)") case True then have "(mset E, {#}) \ unif\<^sup>*" by (simp add: unifiable_imp_empty) from unif_rtrancl_mono [OF this, of "{#(t, Var x)#}"] obtain \ where "(mset E + {#(t, Var x)#}, {#(t \ \, Var x \ \)#}) \ unif\<^sup>*" by (auto simp: subst_mset_def) moreover obtain E' where "E' \ {#}" and "({#(t \ \, Var x \ \)#}, E') \ unif\<^sup>!" using not_unifiable_imp_not_empty_NF and ** by (metis unifiable_insert_swap set_mset_single) ultimately show ?thesis by (auto simp: t_def) next case False moreover have "\ unifiable {(t, Var x)}" using * by (simp add: unifiable_def) ultimately have "\ unifiable (set ((t, Var x) # E))" by (auto simp: unifiable_def unifiers_def) then show ?thesis by (simp add: not_unifiable_imp_not_empty_NF t_def) qed } ultimately show ?case by blast qed lemma unify_complete: assumes "unify E bs = None" shows "unifiers (set E) = {}" proof - from unify_None [OF assms] obtain E' where "E' \ {#}" and "(mset E, E') \ unif\<^sup>!" by blast then have "\ unifiable (set E)" using irreducible_reachable_imp_not_unifiable by force then show ?thesis by (auto simp: unifiable_def) qed corollary ex_unify_if_unifiers_not_empty: \<^marker>\contributor \Martin Desharnais\\ "unifiers es \ {} \ set xs = es \ \ys. unify xs [] = Some ys" using unify_complete by auto lemma mgu_complete: "mgu s t = None \ unifiers {(s, t)} = {}" proof - assume "mgu s t = None" then have "unify [(s, t)] [] = None" by (cases "unify [(s, t)] []", auto simp: mgu_def) then have "unifiers (set [(s, t)]) = {}" by (rule unify_complete) then show ?thesis by simp qed corollary ex_mgu_if_unifiers_not_empty: \<^marker>\contributor \Martin Desharnais\\ "unifiers {(t,u)} \ {} \ \\. mgu t u = Some \" using mgu_complete by auto corollary ex_mgu_if_subst_apply_term_eq_subst_apply_term: \<^marker>\contributor \Martin Desharnais\\ fixes t u :: "('f, 'v) Term.term" and \ :: "('f, 'v) subst" assumes t_eq_u: "t \ \ = u \ \" shows "\\ :: ('f, 'v) subst. Unification.mgu t u = Some \" proof - from t_eq_u have "unifiers {(t, u)} \ {}" unfolding unifiers_def by auto thus ?thesis by (rule ex_mgu_if_unifiers_not_empty) qed lemma finite_subst_domain_subst_of: "finite (subst_domain (subst_of xs))" proof (induct xs) case (Cons x xs) moreover have "finite (subst_domain (subst (fst x) (snd x)))" by (metis finite_subst_domain_subst) ultimately show ?case using subst_domain_compose [of "subst_of xs" "subst (fst x) (snd x)"] by (simp del: subst_subst_domain) (metis finite_subset infinite_Un) qed simp lemma unify_subst_domain: \<^marker>\contributor \Martin Desharnais\\ assumes unif: "unify E [] = Some xs" shows "subst_domain (subst_of xs) \ (\e \ set E. vars_term (fst e) \ vars_term (snd e))" proof - from unify_Some_UNIF[OF unif] obtain xs' where "subst_of xs = compose xs'" and "UNIF xs' (mset E) {#}" by auto thus ?thesis using UNIF_subst_domain_subset by (metis (mono_tags, lifting) multiset.set_map set_mset_mset vars_mset_def) qed lemma mgu_subst_domain: assumes "mgu s t = Some \" shows "subst_domain \ \ vars_term s \ vars_term t" proof - obtain xs where "unify [(s, t)] [] = Some xs" and "\ = subst_of xs" using assms by (simp add: mgu_def split: option.splits) thus ?thesis using unify_subst_domain by fastforce qed lemma mgu_finite_subst_domain: "mgu s t = Some \ \ finite (subst_domain \)" by (drule mgu_subst_domain) (simp add: finite_subset) lemma unify_range_vars: \<^marker>\contributor \Martin Desharnais\\ assumes unif: "unify E [] = Some xs" shows "range_vars (subst_of xs) \ (\e \ set E. vars_term (fst e) \ vars_term (snd e))" proof - from unify_Some_UNIF[OF unif] obtain xs' where "subst_of xs = compose xs'" and "UNIF xs' (mset E) {#}" by auto thus ?thesis using UNIF_range_vars_subset by (metis (mono_tags, lifting) multiset.set_map set_mset_mset vars_mset_def) qed lemma mgu_range_vars: \<^marker>\contributor \Martin Desharnais\\ assumes "mgu s t = Some \" shows "range_vars \ \ vars_term s \ vars_term t" proof - obtain xs where "unify [(s, t)] [] = Some xs" and "\ = subst_of xs" using assms by (simp add: mgu_def split: option.splits) thus ?thesis using unify_range_vars by fastforce qed lemma unify_subst_domain_range_vars_disjoint: \<^marker>\contributor \Martin Desharnais\\ assumes unif: "unify E [] = Some xs" shows "subst_domain (subst_of xs) \ range_vars (subst_of xs) = {}" proof - from unify_Some_UNIF[OF unif] obtain xs' where "subst_of xs = compose xs'" and "UNIF xs' (mset E) {#}" by auto thus ?thesis using UNIF_subst_domain_range_vars_Int by metis qed lemma mgu_subst_domain_range_vars_disjoint: \<^marker>\contributor \Martin Desharnais\\ assumes "mgu s t = Some \" shows "subst_domain \ \ range_vars \ = {}" proof - obtain xs where "unify [(s, t)] [] = Some xs" and "\ = subst_of xs" using assms by (simp add: mgu_def split: option.splits) thus ?thesis using unify_subst_domain_range_vars_disjoint by metis qed corollary subst_apply_term_eq_subst_apply_term_if_mgu: \<^marker>\contributor \Martin Desharnais\\ assumes mgu_t_u: "mgu t u = Some \" shows "t \ \ = u \ \" using mgu_sound[OF mgu_t_u] by (simp add: is_imgu_def unifiers_def) lemma mgu_same: "mgu t t = Some Var" \<^marker>\contributor \Martin Desharnais\\ by (simp add: mgu_def unify_same) lemma mgu_is_Var_if_not_in_equations: \<^marker>\contributor \Martin Desharnais\\ fixes \ :: "('f, 'v) subst" and E :: "('f, 'v) equations" and x :: 'v assumes mgu_\: "is_mgu \ E" and x_not_in: "x \ (\e\E. vars_term (fst e) \ vars_term (snd e))" shows "is_Var (\ x)" proof - from mgu_\ have unif_\: "\ \ unifiers E" and minimal_\: "\\ \ unifiers E. \\. \ = \ \\<^sub>s \" by (simp_all add: is_mgu_def) define \ :: "('f, 'v) subst" where "\ = (\x. if x \ (\e \ E. vars_term (fst e) \ vars_term (snd e)) then \ x else Var x)" have \\ \ unifiers E\ unfolding unifiers_def mem_Collect_eq proof (rule ballI) fix e assume "e \ E" with unif_\ have "fst e \ \ = snd e \ \" by blast moreover from \e \ E\ have "fst e \ \ = fst e \ \" and "snd e \ \ = snd e \ \" unfolding term_subst_eq_conv by (auto simp: \_def) ultimately show "fst e \ \ = snd e \ \" by simp qed with minimal_\ obtain \ where "\ \\<^sub>s \ = \" by auto with x_not_in have "(\ \\<^sub>s \) x = Var x" by (simp add: \_def) thus "is_Var (\ x)" by (metis subst_apply_eq_Var subst_compose term.disc(1)) qed corollary mgu_ball_is_Var: \<^marker>\contributor \Martin Desharnais\\ "is_mgu \ E \ \x \ - (\e\E. vars_term (fst e) \ vars_term (snd e)). is_Var (\ x)" by (rule ballI) (rule mgu_is_Var_if_not_in_equations[folded Compl_iff]) lemma mgu_inj_on: \<^marker>\contributor \Martin Desharnais\\ fixes \ :: "('f, 'v) subst" and E :: "('f, 'v) equations" assumes mgu_\: "is_mgu \ E" shows "inj_on \ (- (\e \ E. vars_term (fst e) \ vars_term (snd e)))" proof (rule inj_onI) fix x y assume x_in: "x \ - (\e\E. vars_term (fst e) \ vars_term (snd e))" and y_in: "y \ - (\e\E. vars_term (fst e) \ vars_term (snd e))" and "\ x = \ y" from mgu_\ have unif_\: "\ \ unifiers E" and minimal_\: "\\ \ unifiers E. \\. \ = \ \\<^sub>s \" by (simp_all add: is_mgu_def) define \ :: "('f, 'v) subst" where "\ = (\x. if x \ (\e \ E. vars_term (fst e) \ vars_term (snd e)) then \ x else Var x)" have \\ \ unifiers E\ unfolding unifiers_def mem_Collect_eq proof (rule ballI) fix e assume "e \ E" with unif_\ have "fst e \ \ = snd e \ \" by blast moreover from \e \ E\ have "fst e \ \ = fst e \ \" and "snd e \ \ = snd e \ \" unfolding term_subst_eq_conv by (auto simp: \_def) ultimately show "fst e \ \ = snd e \ \" by simp qed with minimal_\ obtain \ where "\ \\<^sub>s \ = \" by auto hence "(\ \\<^sub>s \) x = Var x" and "(\ \\<^sub>s \) y = Var y" using ComplD[OF x_in] ComplD[OF y_in] by (simp_all add: \_def) with \\ x = \ y\ show "x = y" by (simp add: subst_compose_def) qed lemma imgu_subst_domain_subset: \<^marker>\contributor \Martin Desharnais\\ fixes \ :: "('f, 'v) subst" and E :: "('f, 'v) equations" and Evars :: "'v set" assumes imgu_\: "is_imgu \ E" and fin_E: "finite E" defines "Evars \ (\e \ E. vars_term (fst e) \ vars_term (snd e))" shows "subst_domain \ \ Evars" proof (intro Set.subsetI) from imgu_\ have unif_\: "\ \ unifiers E" and minimal_\: "\\ \ unifiers E. \ \\<^sub>s \ = \" by (simp_all add: is_imgu_def) from fin_E obtain es :: "('f, 'v) equation list" where "set es = E" using finite_list by auto then obtain xs :: "('v \ ('f, 'v) Term.term) list" where unify_es: "unify es [] = Some xs" using unif_\ ex_unify_if_unifiers_not_empty by blast define \ :: "('f, 'v) subst" where "\ = subst_of xs" have dom_\: "subst_domain \ \ Evars" using unify_subst_domain[OF unify_es, unfolded \set es = E\, folded Evars_def \_def] . have range_vars_\: "range_vars \ \ Evars" using unify_range_vars[OF unify_es, unfolded \set es = E\, folded Evars_def \_def] . have "\ \ unifiers E" using \set es = E\ unify_es \_def is_imgu_def unify_sound by blast with minimal_\ have \_comp_\: "\x. (\ \\<^sub>s \) x = \ x" by auto fix x :: 'v assume "x \ subst_domain \" hence "\ x \ Var x" by (simp add: subst_domain_def) show "x \ Evars" proof (cases "x \ subst_domain \") case True thus ?thesis using dom_\ by auto next case False hence "\ x = Var x" by (simp add: subst_domain_def) hence "\ x \ \ = Var x" using \_comp_\[of x] by (simp add: subst_compose) thus ?thesis - proof (rule subst_apply_eq_Var[of "\ x" \ x]) + proof (rule subst_apply_eq_Var) show "\y. \ x = Var y \ \ y = Var x \ ?thesis" using \\ x \ Var x\ range_vars_\ mem_range_varsI[of \ _ x] by auto qed qed qed lemma imgu_range_vars_of_equations_vars_subset: \<^marker>\contributor \Martin Desharnais\\ fixes \ :: "('f, 'v) subst" and E :: "('f, 'v) equations" and Evars :: "'v set" assumes imgu_\: "is_imgu \ E" and fin_E: "finite E" defines "Evars \ (\e \ E. vars_term (fst e) \ vars_term (snd e))" shows "(\x \ Evars. vars_term (\ x)) \ Evars" proof (rule Set.subsetI) from imgu_\ have unif_\: "\ \ unifiers E" and minimal_\: "\\ \ unifiers E. \ \\<^sub>s \ = \" by (simp_all add: is_imgu_def) from fin_E obtain es :: "('f, 'v) equation list" where "set es = E" using finite_list by auto then obtain xs :: "('v \ ('f, 'v) Term.term) list" where unify_es: "unify es [] = Some xs" using unif_\ ex_unify_if_unifiers_not_empty by blast define \ :: "('f, 'v) subst" where "\ = subst_of xs" have dom_\: "subst_domain \ \ Evars" using unify_subst_domain[OF unify_es, unfolded \set es = E\, folded Evars_def \_def] . have range_vars_\: "range_vars \ \ Evars" using unify_range_vars[OF unify_es, unfolded \set es = E\, folded Evars_def \_def] . hence ball_vars_apply_\_subset: "\x \ subst_domain \. vars_term (\ x) \ Evars" unfolding range_vars_def by (simp add: SUP_le_iff) have "\ \ unifiers E" using \set es = E\ unify_es \_def is_imgu_def unify_sound by blast with minimal_\ have \_comp_\: "\x. (\ \\<^sub>s \) x = \ x" by auto fix y :: 'v assume "y \ (\x \ Evars. vars_term (\ x))" then obtain x :: 'v where x_in: "x \ Evars" and y_in: "y \ vars_term (\ x)" by (auto simp: subst_domain_def) have vars_\_x: "vars_term (\ x) \ Evars" using ball_vars_apply_\_subset subst_domain_def x_in by fastforce show "y \ Evars" proof (rule ccontr) assume "y \ Evars" hence "y \ vars_term (\ x)" using vars_\_x by blast moreover have "y \ vars_term ((\ \\<^sub>s \) x)" proof - have "\ y = Var y" using \y \ Evars\ dom_\ by (auto simp add: subst_domain_def) thus ?thesis unfolding subst_compose_def vars_term_subst_apply_term UN_iff using y_in by force qed ultimately show False using \_comp_\[of x] by simp qed qed lemma imgu_range_vars_subset: \<^marker>\contributor \Martin Desharnais\\ fixes \ :: "('f, 'v) subst" and E :: "('f, 'v) equations" assumes imgu_\: "is_imgu \ E" and fin_E: "finite E" shows "range_vars \ \ (\e \ E. vars_term (fst e) \ vars_term (snd e))" proof - have "range_vars \ = (\x \ subst_domain \. vars_term (\ x))" by (simp add: range_vars_def) also have "\ \ (\x \ (\e \ E. vars_term (fst e) \ vars_term (snd e)). vars_term (\ x))" using imgu_subst_domain_subset[OF imgu_\ fin_E] by fast also have "\ \ (\e \ E. vars_term (fst e) \ vars_term (snd e))" using imgu_range_vars_of_equations_vars_subset[OF imgu_\ fin_E] by metis finally show ?thesis . qed definition the_mgu :: "('f, 'v) term \ ('f, 'v) term \ ('f ,'v) subst" where "the_mgu s t = (case mgu s t of None \ Var | Some \ \ \)" lemma the_mgu_is_imgu: fixes \ :: "('f, 'v) subst" assumes "s \ \ = t \ \" shows "is_imgu (the_mgu s t) {(s, t)}" proof - from assms have "unifiers {(s, t)} \ {}" by (force simp: unifiers_def) then obtain \ where "mgu s t = Some \" and "the_mgu s t = \" using mgu_complete by (auto simp: the_mgu_def) with mgu_sound show ?thesis by blast qed lemma the_mgu: fixes \ :: "('f, 'v) subst" assumes "s \ \ = t \ \" shows "s \ the_mgu s t = t \ the_mgu s t \ \ = the_mgu s t \\<^sub>s \" proof - have *: "\ \ unifiers {(s, t)}" by (force simp: assms unifiers_def) show ?thesis proof (cases "mgu s t") assume "mgu s t = None" then have "unifiers {(s, t)} = {}" by (rule mgu_complete) with * show ?thesis by simp next fix \ assume "mgu s t = Some \" moreover then have "is_imgu \ {(s, t)}" by (rule mgu_sound) ultimately have "is_imgu (the_mgu s t) {(s, t)}" by (unfold the_mgu_def, simp) with * show ?thesis by (auto simp: is_imgu_def unifiers_def) qed qed subsubsection \Unification of two terms where variables should be considered disjoint\ definition mgu_var_disjoint_generic :: "('v \ 'u) \ ('w \ 'u) \ ('f, 'v) term \ ('f, 'w) term \ (('f, 'v, 'u) gsubst \ ('f, 'w, 'u) gsubst) option" where "mgu_var_disjoint_generic vu wu s t = (case mgu (map_vars_term vu s) (map_vars_term wu t) of None \ None | Some \ \ Some (\ \ vu, \ \ wu))" lemma mgu_var_disjoint_generic_sound: assumes unif: "mgu_var_disjoint_generic vu wu s t = Some (\1, \2)" shows "s \ \1 = t \ \2" proof - from unif[unfolded mgu_var_disjoint_generic_def] obtain \ where unif2: "mgu (map_vars_term vu s) (map_vars_term wu t) = Some \" by (cases "mgu (map_vars_term vu s) (map_vars_term wu t)", auto) from mgu_sound[OF unif2[unfolded mgu_var_disjoint_generic_def], unfolded is_imgu_def unifiers_def] have "map_vars_term vu s \ \ = map_vars_term wu t \ \" by auto from this[unfolded apply_subst_map_vars_term] unif[unfolded mgu_var_disjoint_generic_def unif2] show ?thesis by simp qed (* if terms s and t can become identical via two substitutions \ and \ then mgu_var_disjoint yields two more general substitutions \1 \2 *) lemma mgu_var_disjoint_generic_complete: fixes \ :: "('f, 'v, 'u) gsubst" and \ :: "('f, 'w, 'u) gsubst" and vu :: "'v \ 'u" and wu:: "'w \ 'u" assumes inj: "inj vu" "inj wu" and vwu: "range vu \ range wu = {}" and unif_disj: "s \ \ = t \ \" shows "\\1 \2 \. mgu_var_disjoint_generic vu wu s t = Some (\1, \2) \ \ = \1 \\<^sub>s \ \ \ = \2 \\<^sub>s \ \ s \ \1 = t \ \2" proof - note inv1[simp] = the_inv_f_f[OF inj(1)] note inv2[simp] = the_inv_f_f[OF inj(2)] obtain \ :: "('f,'u)subst" where gamma: "\ = (\ x. if x \ range vu then \ (the_inv vu x) else \ (the_inv wu x))" by auto have ids: "s \ \ = map_vars_term vu s \ \" unfolding gamma by (induct s, auto) have idt: "t \ \ = map_vars_term wu t \ \" unfolding gamma by (induct t, insert vwu, auto) from unif_disj ids idt have unif: "map_vars_term vu s \ \ = map_vars_term wu t \ \" (is "?s \ \ = ?t \ \") by auto from the_mgu[OF unif] have unif2: "?s \ the_mgu ?s ?t = ?t \ the_mgu ?s ?t" and inst: "\ = the_mgu ?s ?t \\<^sub>s \" by auto have "mgu ?s ?t = Some (the_mgu ?s ?t)" unfolding the_mgu_def using mgu_complete[unfolded unifiers_def] unif by (cases "mgu ?s ?t", auto) with inst obtain \ where mu: "mgu ?s ?t = Some \" and gamma_mu: "\ = \ \\<^sub>s \" by auto let ?tau1 = "\ \ vu" let ?tau2 = "\ \ wu" show ?thesis unfolding mgu_var_disjoint_generic_def mu option.simps proof (intro exI conjI, rule refl) show "\ = ?tau1 \\<^sub>s \" proof (rule ext) fix x have "(?tau1 \\<^sub>s \) x = \ (vu x)" using fun_cong[OF gamma_mu, of "vu x"] by (simp add: subst_compose_def) also have "... = \ x" unfolding gamma by simp finally show "\ x = (?tau1 \\<^sub>s \) x" by simp qed next show "\ = ?tau2 \\<^sub>s \" proof (rule ext) fix x have "(?tau2 \\<^sub>s \) x = \ (wu x)" using fun_cong[OF gamma_mu, of "wu x"] by (simp add: subst_compose_def) also have "... = \ x" unfolding gamma using vwu by auto finally show "\ x = (?tau2 \\<^sub>s \) x" by simp qed next have "s \ ?tau1 = map_vars_term vu s \ \" unfolding apply_subst_map_vars_term .. also have "... = map_vars_term wu t \ \" unfolding unif2[unfolded the_mgu_def mu option.simps] .. also have "... = t \ ?tau2" unfolding apply_subst_map_vars_term .. finally show "s \ ?tau1 = t \ ?tau2" . qed qed abbreviation "mgu_var_disjoint_sum \ mgu_var_disjoint_generic Inl Inr" lemma mgu_var_disjoint_sum_sound: "mgu_var_disjoint_sum s t = Some (\1, \2) \ s \ \1 = t \ \2" by (rule mgu_var_disjoint_generic_sound) lemma mgu_var_disjoint_sum_complete: fixes \ :: "('f, 'v, 'v + 'w) gsubst" and \ :: "('f, 'w, 'v + 'w) gsubst" assumes unif_disj: "s \ \ = t \ \" shows "\\1 \2 \. mgu_var_disjoint_sum s t = Some (\1, \2) \ \ = \1 \\<^sub>s \ \ \ = \2 \\<^sub>s \ \ s \ \1 = t \ \2" by (rule mgu_var_disjoint_generic_complete[OF _ _ _ unif_disj], auto simp: inj_on_def) lemma mgu_var_disjoint_sum_instance: fixes \ :: "('f, 'v) subst" and \ :: "('f, 'v) subst" assumes unif_disj: "s \ \ = t \ \" shows "\\1 \2 \. mgu_var_disjoint_sum s t = Some (\1, \2) \ \ = \1 \\<^sub>s \ \ \ = \2 \\<^sub>s \ \ s \ \1 = t \ \2" proof - let ?map = "\ m \ v. map_vars_term m (\ v)" let ?m = "?map (Inl :: ('v \ 'v + 'v))" let ?m' = "?map (case_sum (\ x. x) (\ x. x))" from unif_disj have id: "map_vars_term Inl (s \ \) = map_vars_term Inl (t \ \)" by simp from mgu_var_disjoint_sum_complete[OF id[unfolded map_vars_term_subst]] obtain \1 \2 \ where mgu: "mgu_var_disjoint_sum s t = Some (\1,\2)" and \: "?m \ = \1 \\<^sub>s \" and \: "?m \ = \2 \\<^sub>s \" and unif: "s \ \1 = t \ \2" by blast { fix \ :: "('f, 'v) subst" have "?m' (?m \) = \" by (simp add: map_vars_term_compose o_def term.map_ident) } note id = this { fix \ :: "('f,'v,'v+'v)gsubst" and \ :: "('f,'v + 'v)subst" have "?m' (\ \\<^sub>s \) = \ \\<^sub>s ?m' \" by (rule ext, unfold subst_compose_def, simp add: map_vars_term_subst) } note id' = this from arg_cong[OF \, of ?m', unfolded id id'] have \: "\ = \1 \\<^sub>s ?m' \" . from arg_cong[OF \, of ?m', unfolded id id'] have \: "\ = \2 \\<^sub>s ?m' \" . show ?thesis by (intro exI conjI, rule mgu, rule \, rule \, rule unif) qed subsubsection \A variable disjoint unification algorithm without changing the type\ text \We pass the renaming function as additional argument\ definition mgu_vd :: "'v :: infinite renaming2 \ _ \ _" where "mgu_vd r = mgu_var_disjoint_generic (rename_1 r) (rename_2 r)" lemma mgu_vd_sound: "mgu_vd r s t = Some (\1, \2) \ s \ \1 = t \ \2" unfolding mgu_vd_def by (rule mgu_var_disjoint_generic_sound) lemma mgu_vd_complete: fixes \ :: "('f, 'v :: infinite) subst" and \ :: "('f, 'v) subst" assumes unif_disj: "s \ \ = t \ \" shows "\\1 \2 \. mgu_vd r s t = Some (\1, \2) \ \ = \1 \\<^sub>s \ \ \ = \2 \\<^sub>s \ \ s \ \1 = t \ \2" unfolding mgu_vd_def by (rule mgu_var_disjoint_generic_complete[OF rename_12 unif_disj]) end diff --git a/thys/Rewrite_Properties_Reduction/Rewriting/Replace_Constant.thy b/thys/Rewrite_Properties_Reduction/Rewriting/Replace_Constant.thy --- a/thys/Rewrite_Properties_Reduction/Rewriting/Replace_Constant.thy +++ b/thys/Rewrite_Properties_Reduction/Rewriting/Replace_Constant.thy @@ -1,370 +1,370 @@ theory Replace_Constant imports Rewriting begin subsection \Removing/Replacing constants in a rewrite sequence that do not appear in the rewrite system\ lemma funas_term_const_subst_conv: "(c, 0) \ funas_term l \ \ (l \ constT c)" proof (induct l) case (Fun f ts) then show ?case by auto (metis Fun_supt supteq_supt_conv term.inject(2))+ qed (auto simp add: supteq_var_imp_eq) lemma fresh_const_single_step_replace: assumes lin: "linear_sys \" and fresh: "(c, 0) \ funas_rel \" and occ: "p \ poss_of_term (constT c) s" and step: "(s, t) \ rstep \" shows "(s[p \ u], t) \ rstep \ \ (\ q. q \ poss_of_term (constT c) t \ (s[p \ u], t[q \ u]) \ rstep \)" proof - from occ have const: "p \ poss s \ s |_ p = constT c" by auto from step obtain C l r \ where t [simp]: "s = C\l \ \\" "t = C\r \ \\" and rule: "(l, r) \ \" by blast from rule lin have lin: "linear_term l" "linear_term r" by fastforce+ from fresh rule have nt_lhs: "(c, 0) \ funas_term l" by (auto simp: funas_rel_def) consider (par) "p \ (hole_pos C)" | (below) "hole_pos C \\<^sub>p p" using occ by (auto dest: poss_of_term_const_ctxt_apply) then show ?thesis proof cases case par then have possc: "p \ possc C" using const t possc_def by blast then have "p \ poss_of_term (constT c) t" "(s[p \ u], t[p \ u]) \ rstep \" using const par_hole_pos_replace_term_context_at[OF par] using possc_subt_at_ctxt_apply[OF possc par, of "r \ \" "l \ \"] rule by auto (metis par par_pos_replace_pres replace_at_hole_pos) then show ?thesis by blast next case below then obtain q where [simp]:"p = hole_pos C @ q" and poss: "q \ poss (l \ \)" using const position_less_eq_def by (metis (full_types) ctxt_at_pos_hole_pos ctxt_at_pos_subt_at_pos poss_append_poss t(1)) have const: "l \ \ |_ q = constT c" using const by auto from nt_lhs have "\ r. r \ varposs l \ r \\<^sub>p q" using const poss proof (induct l arbitrary: q) case (Var x) then show ?case by auto next case (Fun f ts) from Fun(1)[OF nth_mem, of "hd q" "tl q"] Fun(2-) obtain r where "r \ varposs (ts ! hd q) \ r \\<^sub>p tl q" by (cases q) auto then show ?case using Fun(2- 4) by (intro exI[of _ "hd q # r"]) auto qed then obtain x v where varposs: "v \ varposs l" "v \\<^sub>p q" "l |_ v = Var x" unfolding varposs_def by blast let ?\ = "\x. if Var x = l |_ v then (\ x)[q -\<^sub>p v \ u] else \ x" show ?thesis proof (cases "x \ vars_term r") case True then obtain q' where varposs_r: "q' \ varposs r" "r |_ q' = Var x" by (metis vars_term_varposs_iff) have "(s[p \ u], t[(hole_pos C) @ q' @ (q -\<^sub>p v) \ u]) \ rstep \" using lin varposs rule varposs_r by (auto simp: linear_term_varposs_subst_replace_term intro!: rstep_ctxtI) (smt (verit, ccfv_SIG) pos_diff_append_itself rrstep.intros rrstep_rstep_mono subset_eq term_subst_eq) moreover have "(hole_pos C) @ q' @ q -\<^sub>p v \ poss_of_term (constT c) t" using varposs_r varposs poss const poss_pos_diffI[OF varposs(2) poss] using subt_at_append_dist[of q' "q -\<^sub>p v" "r \ \"] by (auto simp: poss_append_poss varposs_imp_poss[THEN subst_subt_at_dist] varposs_imp_poss[THEN subsetD[OF subst_poss_mono]]) - (metis pos_les_eq_append_diff subst_apply_term.simps(1) subst_subt_at_dist subt_at_append_dist varposs_imp_poss) + (metis pos_les_eq_append_diff eval_term.simps(1) subst_subt_at_dist subt_at_append_dist varposs_imp_poss) ultimately show ?thesis by auto next case False then have [simp]: "r \ \ = r \ ?\" using varposs by (auto simp add: term_subst_eq_conv) have "(s[p \ u], t) \ rstep \" using rule varposs lin by (auto simp: linear_term_varposs_subst_replace_term) then show ?thesis by auto qed qed qed lemma fresh_const_steps_replace: assumes lin: "linear_sys \" and fresh: "(c, 0) \ funas_rel \" and occ: "p \ poss_of_term (constT c) s" and steps: "(s, t) \ (rstep \)\<^sup>+" shows "(s[p \ u], t) \ (rstep \)\<^sup>+ \ (\ q. q \ poss_of_term (constT c) t \ (s[p \ u], t[q \ u]) \ (rstep \)\<^sup>+)" using steps occ proof (induct arbitrary: p rule: converse_trancl_induct) case (base s) from fresh_const_single_step_replace[OF lin fresh base(2, 1)] show ?case by (meson r_into_trancl') next case (step s t) from fresh_const_single_step_replace[OF lin fresh step(4, 1)] consider (a) "(s[p \ u], t) \ rstep \" | (b) "\q. q \ poss_of_term (constT c) t \ (s[p \ u], t[q \ u]) \ rstep \" by blast then show ?case proof cases case a then show ?thesis using step(2) by auto next case b then obtain q where "q \ poss_of_term (constT c) t" "(s[p \ u], t[q \ u]) \ rstep \" by blast from step(3)[OF this(1)] this(2) show ?thesis by (metis trancl_into_trancl2) qed qed lemma remove_const_lhs_steps: assumes lin: "linear_sys \" and fresh: "(c, 0) \ funas_rel \" and const: "(c, 0) \ funas_term t" and pos: "p \ poss_of_term (constT c) s" and steps: "(s, t) \ (rstep \)\<^sup>+" shows "(s[p \ u], t) \ (rstep \)\<^sup>+" using steps pos const fresh_const_steps_replace by (metis fresh funas_term_const_subst_conv lin poss_of_termE subt_at_imp_supteq) text \Now we can show that we may remove a constant substitution\ definition const_replace_closed where "const_replace_closed c U = (\ s t u p. p \ poss_of_term (constT c) s \ (s, t) \ U \ (\ q. q \ poss_of_term (constT c) t \ (s[p \ u], t[q \ u]) \ U) \ (s[p \ u], t) \ U)" lemma const_replace_closedD: assumes "const_replace_closed c U" "p \ poss_of_term (constT c) s" "(s, t) \ U" shows "(s[p \ u], t) \ U \ (\ q. q \ poss_of_term (constT c) t \ (s[p \ u], t[q \ u]) \ U)" using assms unfolding const_replace_closed_def by blast lemma const_replace_closedI: assumes "\ s t u p. p \ poss_of_term (constT c) s \ (s, t) \ U \ (\ q. q \ poss_of_term (constT c) t \ (s[p \ u], t[q \ u]) \ U) \ (s[p \ u], t) \ U" shows "const_replace_closed c U" using assms unfolding const_replace_closed_def by auto abbreviation const_subst :: "'f \ 'v \ ('f, 'v) Term.term" where "const_subst c \ (\ x. Fun c [])" lemma lin_fresh_rstep_const_replace_closed: "linear_sys \ \ (c, 0) \ funas_rel \ \ const_replace_closed c (rstep \)" using fresh_const_single_step_replace[of \ c] by (intro const_replace_closedI) (auto simp: constT_nfunas_term_poss_of_term_empty, blast) lemma const_replace_closed_symcl: "const_replace_closed c U \ const_replace_closed c (U\<^sup>=)" unfolding const_replace_closed_def by (metis Un_iff pair_in_Id_conv) lemma const_replace_closed_trancl: "const_replace_closed c U \ const_replace_closed c (U\<^sup>+)" proof (intro const_replace_closedI) fix s t u p assume const: "const_replace_closed c U" and wit: "p \ poss_of_term (constT c) s" and steps :"(s, t) \ U\<^sup>+" show "(\q. q \ poss_of_term (constT c) t \ (s[p \ u], t[q \ u]) \ U\<^sup>+) \ (s[p \ u], t) \ U\<^sup>+" using steps wit proof (induct arbitrary: p rule: converse_trancl_induct) case (base s) show ?case using const_replace_closedD[OF const base(2, 1)] by blast next case (step s v) from const_replace_closedD[OF const step(4, 1)] consider (a) "(s[p \ u], v) \ U" | (b) "\ q. q \ poss_of_term (constT c) v \ (s[p \ u], v[q \ u]) \ U" by auto then show ?case proof cases case a then show ?thesis using step(2) by (meson trancl_into_trancl2) next case b then show ?thesis using step(3, 4) by (meson trancl_into_trancl2) qed qed qed lemma const_replace_closed_rtrancl: "const_replace_closed c U \ const_replace_closed c (U\<^sup>*)" proof (intro const_replace_closedI) fix s t u p assume const: "const_replace_closed c U" and wit: "p \ poss_of_term (constT c) s" and steps :"(s, t) \ U\<^sup>*" show "(\q. q \ poss_of_term (constT c) t \ (s[p \ u], t[q \ u]) \ U\<^sup>*) \ (s[p \ u], t) \ U\<^sup>*" using const_replace_closed_trancl[OF const] wit steps by (metis const_replace_closedD rtrancl_eq_or_trancl) qed lemma const_replace_closed_relcomp: "const_replace_closed c U \ const_replace_closed c V \ const_replace_closed c (U O V)" proof (intro const_replace_closedI) fix s t u p assume const: "const_replace_closed c U" "const_replace_closed c V" and wit: "p \ poss_of_term (constT c) s" and step: "(s, t) \ U O V" from step obtain w where w: "(s, w) \ U" "(w, t) \ V" by auto from const_replace_closedD[OF const(1) wit this(1)] consider (a) "(s[p \ u], w) \ U" | (b) "(\q. q \ poss_of_term (constT c) w \ (s[p \ u], w[q \ u]) \ U)" by auto then show "(\q. q \ poss_of_term (constT c) t \ (s[p \ u], t[q \ u]) \ U O V) \ (s[p \ u], t) \ U O V" proof cases case a then show ?thesis using w(2) by auto next case b then show ?thesis using const_replace_closedD[OF const(2) _ w(2)] by (meson relcomp.simps) qed qed text \@{const const_replace_closed} allow the removal of a fresh constant substitution\ lemma const_replace_closed_remove_subst_lhs: assumes repcl: "const_replace_closed c U" and const: "(c, 0) \ funas_term t" and steps: "(s \ const_subst c, t) \ U" shows "(s, t) \ U" using steps proof (induct "card (varposs s)" arbitrary: s) case (Suc n) obtain p ps where vl: "varposs s = insert p ps" "p \ ps" using Suc(2) by (metis card_le_Suc_iff dual_order.refl) let ?s = "s[p \ Fun c []]" have vp: "p \ varposs s" using vl by auto then have [simp]: "?s \ const_subst c = s \ const_subst c" by (induct s arbitrary: p) (auto simp: nth_list_update map_update intro!: nth_equalityI) have "varposs ?s = ps" using vl varposs_ground_replace_at[of p s "constT c"] by auto then have "n = card (varposs ?s)" using vl Suc(2) by (auto simp: card_insert_if finite_varposs) from Suc(1)[OF this] have IH: "(s[p \ constT c], t) \ U" "p \ poss_of_term (constT c) s[p \ constT c]" using Suc(2, 3) vl poss_of_term_replace_term_at varposs_imp_poss vp using \s[p \ constT c] \ const_subst c = s \ const_subst c\ by fastforce+ show ?case using const_replace_closedD[OF repcl] const IH(2, 1) by (metis constT_nfunas_term_poss_of_term_empty empty_iff replace_term_at_same_pos replace_term_at_subt_at_id) qed (auto simp: ground_subst_apply card_eq_0_iff finite_varposs varposs_empty_gound) subsubsection \Removal lemma applied to various rewrite relations\ lemma remove_const_subst_step_lhs: assumes lin: "linear_sys \" and fresh: "(c, 0) \ funas_rel \" and const: "(c, 0) \ funas_term t" and step: "(s \ const_subst c, t) \ (rstep \)" shows "(s, t) \ (rstep \)" using lin_fresh_rstep_const_replace_closed[OF lin fresh, THEN const_replace_closed_remove_subst_lhs] const step by blast lemma remove_const_subst_steps_lhs: assumes lin: "linear_sys \" and fresh: "(c, 0) \ funas_rel \" and const: "(c, 0) \ funas_term t" and steps: "(s \ const_subst c, t) \ (rstep \)\<^sup>+" shows "(s, t) \ (rstep \)\<^sup>+" using lin_fresh_rstep_const_replace_closed[THEN const_replace_closed_trancl, OF lin fresh, THEN const_replace_closed_remove_subst_lhs] using const steps by blast lemma remove_const_subst_steps_eq_lhs: assumes lin: "linear_sys \" and fresh: "(c, 0) \ funas_rel \" and const: "(c, 0) \ funas_term t" and steps: "(s \ const_subst c, t) \ (rstep \)\<^sup>*" shows "(s, t) \ (rstep \)\<^sup>*" using steps const by (cases "s = t") (auto simp: rtrancl_eq_or_trancl funas_term_subst ground_subst_apply vars_term_empty_ground dest: remove_const_subst_steps_lhs[OF lin fresh const] split: if_splits) lemma remove_const_subst_steps_rhs: assumes lin: "linear_sys \" and fresh: "(c, 0) \ funas_rel \" and const: "(c, 0) \ funas_term s" and steps: "(s, t \ const_subst c) \ (rstep \)\<^sup>+" shows "(s, t) \ (rstep \)\<^sup>+" proof - from steps have revs: "(t \ const_subst c, s) \ (rstep (\\))\<^sup>+" unfolding rew_converse_outwards by auto have "(t, s) \ (rstep (\\))\<^sup>+" using assms by (intro remove_const_subst_steps_lhs[OF _ _ _ revs]) (auto simp: funas_rel_def) then show ?thesis unfolding rew_converse_outwards by auto qed lemma remove_const_subst_steps_eq_rhs: assumes lin: "linear_sys \" and fresh: "(c, 0) \ funas_rel \" and const: "(c, 0) \ funas_term s" and steps: "(s, t \ const_subst c) \ (rstep \)\<^sup>*" shows "(s, t) \ (rstep \)\<^sup>*" using steps const by (cases "s = t") (auto simp: rtrancl_eq_or_trancl funas_term_subst ground_subst_apply vars_term_empty_ground dest!: remove_const_subst_steps_rhs[OF lin fresh const] split: if_splits) text \Main lemmas\ lemma const_subst_eq_ground_eq: assumes "s \ const_subst c = t \ const_subst d" "c \ d" and "(c, 0) \ funas_term t" "(d, 0) \ funas_term s" shows "s = t" using assms proof (induct s arbitrary: t) case (Var x) then show ?case by (cases t) auto next case (Fun f ts) from Fun(2-) obtain g us where [simp]: "t = Fun g us" by (cases t) auto have [simp]: "g = f" and l: "length ts = length us" using Fun(2) by (auto intro: map_eq_imp_length_eq) have "i < length ts \ ts ! i = us ! i" for i using Fun(1)[OF nth_mem, of i "us ! i" for i] Fun(2-) l by (auto simp: map_eq_conv') then show ?case using l by (auto intro: nth_equalityI) qed lemma remove_const_subst_steps: assumes "linear_sys \" and "(c, 0) \ funas_rel \" and "(d, 0) \ funas_rel \" and "c \ d" "(c, 0) \ funas_term t" "(d, 0) \ funas_term s" and "(s \ const_subst c, t \ const_subst d) \ (rstep \)\<^sup>*" shows "(s, t) \ (rstep \)\<^sup>*" proof (cases "s \ const_subst c = t \ const_subst d") case True from const_subst_eq_ground_eq[OF this] assms(4 - 6) show ?thesis by auto next case False then have step: "(s \ const_subst c, t \ const_subst d) \ (rstep \)\<^sup>+" using assms(7) by (auto simp: rtrancl_eq_or_trancl) then have "(s, t \ const_subst d) \ (rstep \)\<^sup>+" using assms by (intro remove_const_subst_steps_lhs[OF _ _ _ step]) (auto simp: funas_term_subst) from remove_const_subst_steps_rhs[OF _ _ _ this] show ?thesis using assms by auto qed lemma remove_const_subst_relcomp_lhs: assumes sys: "linear_sys \" "linear_sys \" and fr: "(c, 0) \ funas_rel \" and fs:"(c, 0) \ funas_rel \" and funas: "(c, 0) \ funas_term t" and seq: "(s \ const_subst c, t) \ (rstep \)\<^sup>* O (rstep \)\<^sup>*" shows "(s, t) \ (rstep \)\<^sup>* O (rstep \)\<^sup>*" using seq using lin_fresh_rstep_const_replace_closed[OF sys(1) fr, THEN const_replace_closed_rtrancl] using lin_fresh_rstep_const_replace_closed[OF sys(2) fs, THEN const_replace_closed_rtrancl] using const_replace_closed_relcomp by (intro const_replace_closed_remove_subst_lhs[OF _ funas seq]) force lemma remove_const_subst_relcomp_rhs: assumes sys: "linear_sys \" "linear_sys \" and fr: "(c, 0) \ funas_rel \" and fs:"(c, 0) \ funas_rel \" and funas: "(c, 0) \ funas_term s" and seq: "(s, t \ const_subst c) \ (rstep \)\<^sup>* O (rstep \)\<^sup>*" shows "(s, t) \ (rstep \)\<^sup>* O (rstep \)\<^sup>*" proof - from seq have "(t \ const_subst c,s) \ ((rstep \)\<^sup>* O (rstep \)\<^sup>*)\" by auto then have "(t \ const_subst c,s) \ ((rstep \)\<^sup>*)\ O ((rstep \)\<^sup>*)\" using converse_relcomp by blast note seq = this[unfolded rtrancl_converse[symmetric] rew_converse_inwards] from sys fr fs have "linear_sys (\\)" "linear_sys (\\)" "(c, 0) \ funas_rel (\\)" "(c, 0) \ funas_rel (\\)" by (auto simp: funas_rel_def) from remove_const_subst_relcomp_lhs[OF this funas seq] have "(t, s) \ (rstep (\\))\<^sup>* O (rstep (\\))\<^sup>*" by simp then show ?thesis unfolding rew_converse_outwards converse_relcomp[symmetric] by simp qed lemma remove_const_subst_relcomp: assumes sys: "linear_sys \" "linear_sys \" and fr: "(c, 0) \ funas_rel \" "(d, 0) \ funas_rel \" and fs:"(c, 0) \ funas_rel \" "(d, 0) \ funas_rel \" and diff: "c \ d" and funas: "(c, 0) \ funas_term t" "(d, 0) \ funas_term s" and seq: "(s \ const_subst c, t \ const_subst d) \ (rstep \)\<^sup>* O (rstep \)\<^sup>*" shows "(s, t) \ (rstep \)\<^sup>* O (rstep \)\<^sup>*" proof - from diff funas(1) have *: "(c, 0) \ funas_term (t \ const_subst d)" by (auto simp: funas_term_subst) show ?thesis using remove_const_subst_relcomp_rhs[OF sys fr(2) fs(2) funas(2) remove_const_subst_relcomp_lhs[OF sys fr(1) fs(1) * seq]] by blast qed end \ No newline at end of file diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Intruder_Deduction.thy b/thys/Stateful_Protocol_Composition_and_Typing/Intruder_Deduction.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/Intruder_Deduction.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/Intruder_Deduction.thy @@ -1,1193 +1,1191 @@ (* Title: Intruder_Deduction.thy Author: Andreas Viktor Hess, DTU SPDX-License-Identifier: BSD-3-Clause *) section \Dolev-Yao Intruder Model\ theory Intruder_Deduction imports Messages More_Unification begin subsection \Syntax for the Intruder Deduction Relations\ consts INTRUDER_SYNTH::"('f,'v) terms \ ('f,'v) term \ bool" (infix "\\<^sub>c" 50) consts INTRUDER_DEDUCT::"('f,'v) terms \ ('f,'v) term \ bool" (infix "\" 50) subsection \Intruder Model Locale\ text \ The intruder model is parameterized over arbitrary function symbols (e.g, cryptographic operators) and variables. It requires three functions: - \arity\ that assigns an arity to each function symbol. - \public\ that partitions the function symbols into those that will be available to the intruder and those that will not. - \Ana\, the analysis interface, that defines how messages can be decomposed (e.g., decryption). \ locale intruder_model = fixes arity :: "'fun \ nat" and public :: "'fun \ bool" and Ana :: "('fun,'var) term \ (('fun,'var) term list \ ('fun,'var) term list)" assumes Ana_keys_fv: "\t K R. Ana t = (K,R) \ fv\<^sub>s\<^sub>e\<^sub>t (set K) \ fv t" and Ana_keys_wf: "\t k K R f T. Ana t = (K,R) \ (\g S. Fun g S \ t \ length S = arity g) \ k \ set K \ Fun f T \ k \ length T = arity f" and Ana_var[simp]: "\x. Ana (Var x) = ([],[])" and Ana_fun_subterm: "\f T K R. Ana (Fun f T) = (K,R) \ set R \ set T" and Ana_subst: "\t \ K R. \Ana t = (K,R); K \ [] \ R \ []\ \ Ana (t \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \,R \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" begin lemma Ana_subterm: assumes "Ana t = (K,T)" shows "set T \ subterms t" using assms by (cases t) (simp add: psubsetI, metis Ana_fun_subterm Fun_gt_params UN_I term.order_refl params_subterms psubsetI subset_antisym subset_trans) lemma Ana_subterm': "s \ set (snd (Ana t)) \ s \ t" using Ana_subterm by (cases "Ana t") auto lemma Ana_vars: assumes "Ana t = (K,M)" shows "fv\<^sub>s\<^sub>e\<^sub>t (set K) \ fv t" "fv\<^sub>s\<^sub>e\<^sub>t (set M) \ fv t" by (rule Ana_keys_fv[OF assms]) (use Ana_subterm[OF assms] subtermeq_vars_subset in auto) abbreviation \ where "\ \ UNIV::'var set" abbreviation \n ("\\<^sup>_") where "\\<^sup>n \ {f::'fun. arity f = n}" abbreviation \npub ("\\<^sub>p\<^sub>u\<^sub>b\<^sup>_") where "\\<^sub>p\<^sub>u\<^sub>b\<^sup>n \ {f. public f} \ \\<^sup>n" abbreviation \npriv ("\\<^sub>p\<^sub>r\<^sub>i\<^sub>v\<^sup>_") where "\\<^sub>p\<^sub>r\<^sub>i\<^sub>v\<^sup>n \ {f. \public f} \ \\<^sup>n" abbreviation \\<^sub>p\<^sub>u\<^sub>b where "\\<^sub>p\<^sub>u\<^sub>b \ (\n. \\<^sub>p\<^sub>u\<^sub>b\<^sup>n)" abbreviation \\<^sub>p\<^sub>r\<^sub>i\<^sub>v where "\\<^sub>p\<^sub>r\<^sub>i\<^sub>v \ (\n. \\<^sub>p\<^sub>r\<^sub>i\<^sub>v\<^sup>n)" abbreviation \ where "\ \ (\n. \\<^sup>n)" abbreviation \ where "\ \ \\<^sup>0" abbreviation \\<^sub>p\<^sub>u\<^sub>b where "\\<^sub>p\<^sub>u\<^sub>b \ {f. public f} \ \" abbreviation \\<^sub>p\<^sub>r\<^sub>i\<^sub>v where "\\<^sub>p\<^sub>r\<^sub>i\<^sub>v \ {f. \public f} \ \" abbreviation \\<^sub>f where "\\<^sub>f \ \ - \" abbreviation \\<^sub>f\<^sub>p\<^sub>u\<^sub>b where "\\<^sub>f\<^sub>p\<^sub>u\<^sub>b \ \\<^sub>f \ \\<^sub>p\<^sub>u\<^sub>b" abbreviation \\<^sub>f\<^sub>p\<^sub>r\<^sub>i\<^sub>v where "\\<^sub>f\<^sub>p\<^sub>r\<^sub>i\<^sub>v \ \\<^sub>f \ \\<^sub>p\<^sub>r\<^sub>i\<^sub>v" lemma disjoint_fun_syms: "\\<^sub>f \ \ = {}" by auto lemma id_union_univ: "\\<^sub>f \ \ = UNIV" "\ = UNIV" by auto lemma const_arity_eq_zero[dest]: "c \ \ \ arity c = 0" by simp lemma const_pub_arity_eq_zero[dest]: "c \ \\<^sub>p\<^sub>u\<^sub>b \ arity c = 0 \ public c" by simp lemma const_priv_arity_eq_zero[dest]: "c \ \\<^sub>p\<^sub>r\<^sub>i\<^sub>v \ arity c = 0 \ \public c" by simp lemma fun_arity_gt_zero[dest]: "f \ \\<^sub>f \ arity f > 0" by fastforce lemma pub_fun_public[dest]: "f \ \\<^sub>f\<^sub>p\<^sub>u\<^sub>b \ public f" by fastforce lemma pub_fun_arity_gt_zero[dest]: "f \ \\<^sub>f\<^sub>p\<^sub>u\<^sub>b \ arity f > 0" by fastforce lemma \\<^sub>f_unfold: "\\<^sub>f = {f::'fun. arity f > 0}" by auto lemma \_unfold: "\ = {f::'fun. arity f = 0}" by auto lemma \pub_unfold: "\\<^sub>p\<^sub>u\<^sub>b = {f::'fun. arity f = 0 \ public f}" by auto lemma \priv_unfold: "\\<^sub>p\<^sub>r\<^sub>i\<^sub>v = {f::'fun. arity f = 0 \ \public f}" by auto lemma \npub_unfold: "(\\<^sub>p\<^sub>u\<^sub>b\<^sup>n) = {f::'fun. arity f = n \ public f}" by auto lemma \npriv_unfold: "(\\<^sub>p\<^sub>r\<^sub>i\<^sub>v\<^sup>n) = {f::'fun. arity f = n \ \public f}" by auto lemma \fpub_unfold: "\\<^sub>f\<^sub>p\<^sub>u\<^sub>b = {f::'fun. arity f > 0 \ public f}" by auto lemma \fpriv_unfold: "\\<^sub>f\<^sub>p\<^sub>r\<^sub>i\<^sub>v = {f::'fun. arity f > 0 \ \public f}" by auto lemma \n_m_eq: "\(\\<^sup>n) \ {}; (\\<^sup>n) = (\\<^sup>m)\ \ n = m" by auto subsection \Term Well-formedness\ definition "wf\<^sub>t\<^sub>r\<^sub>m t \ \f T. Fun f T \ t \ length T = arity f" abbreviation "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s T \ \t \ T. wf\<^sub>t\<^sub>r\<^sub>m t" lemma Ana_keys_wf': "Ana t = (K,T) \ wf\<^sub>t\<^sub>r\<^sub>m t \ k \ set K \ wf\<^sub>t\<^sub>r\<^sub>m k" using Ana_keys_wf unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by metis lemma wf_trm_Var[simp]: "wf\<^sub>t\<^sub>r\<^sub>m (Var x)" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by simp lemma wf_trm_subst_range_Var[simp]: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range Var)" by simp lemma wf_trm_subst_range_iff: "(\x. wf\<^sub>t\<^sub>r\<^sub>m (\ x)) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by force lemma wf_trm_subst_rangeD: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ wf\<^sub>t\<^sub>r\<^sub>m (\ x)" by (metis wf_trm_subst_range_iff) lemma wf_trm_subst_rangeI[intro]: "(\x. wf\<^sub>t\<^sub>r\<^sub>m (\ x)) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (metis wf_trm_subst_range_iff) lemma wf_trmI[intro]: assumes "\t. t \ set T \ wf\<^sub>t\<^sub>r\<^sub>m t" "length T = arity f" shows "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" using assms unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto lemma wf_trm_subterm: "\wf\<^sub>t\<^sub>r\<^sub>m t; s \ t\ \ wf\<^sub>t\<^sub>r\<^sub>m s" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by (induct t) auto lemma wf_trm_subtermeq: assumes "wf\<^sub>t\<^sub>r\<^sub>m t" "s \ t" shows "wf\<^sub>t\<^sub>r\<^sub>m s" proof (cases "s = t") case False thus "wf\<^sub>t\<^sub>r\<^sub>m s" using assms(2) wf_trm_subterm[OF assms(1)] by simp qed (metis assms(1)) lemma wf_trm_param: assumes "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" "t \ set T" shows "wf\<^sub>t\<^sub>r\<^sub>m t" by (meson assms subtermeqI'' wf_trm_subtermeq) lemma wf_trm_param_idx: assumes "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" and "i < length T" shows "wf\<^sub>t\<^sub>r\<^sub>m (T ! i)" using wf_trm_param[OF assms(1), of "T ! i"] assms(2) by fastforce lemma wf_trm_subst: assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "wf\<^sub>t\<^sub>r\<^sub>m t = wf\<^sub>t\<^sub>r\<^sub>m (t \ \)" proof show "wf\<^sub>t\<^sub>r\<^sub>m t \ wf\<^sub>t\<^sub>r\<^sub>m (t \ \)" proof (induction t) case (Fun f T) hence "\t. t \ set T \ wf\<^sub>t\<^sub>r\<^sub>m t" by (meson wf\<^sub>t\<^sub>r\<^sub>m_def Fun_param_is_subterm term.order_trans) hence "\t. t \ set T \ wf\<^sub>t\<^sub>r\<^sub>m (t \ \)" using Fun.IH by auto moreover have "length (map (\t. t \ \) T) = arity f" using Fun.prems unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto ultimately show ?case by fastforce qed (simp add: wf_trm_subst_rangeD[OF assms]) show "wf\<^sub>t\<^sub>r\<^sub>m (t \ \) \ wf\<^sub>t\<^sub>r\<^sub>m t" proof (induction t) case (Fun f T) hence "wf\<^sub>t\<^sub>r\<^sub>m t" when "t \ set (map (\s. s \ \) T)" for t - by (metis that wf\<^sub>t\<^sub>r\<^sub>m_def Fun_param_is_subterm term.order_trans subst_apply_term.simps(2)) + by (metis that wf\<^sub>t\<^sub>r\<^sub>m_def Fun_param_is_subterm term.order_trans eval_term.simps(2)) hence "wf\<^sub>t\<^sub>r\<^sub>m t" when "t \ set T" for t using that Fun.IH by auto moreover have "length (map (\t. t \ \) T) = arity f" using Fun.prems unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto ultimately show ?case by fastforce qed (simp add: assms) qed lemma wf_trm_subst_singleton: assumes "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m t'" shows "wf\<^sub>t\<^sub>r\<^sub>m (t \ Var(v := t'))" proof - have "wf\<^sub>t\<^sub>r\<^sub>m ((Var(v := t')) w)" for w using assms(2) unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by simp thus ?thesis using assms(1) wf_trm_subst[of "Var(v := t')" t, OF wf_trm_subst_rangeI] by simp qed lemma wf_trm_subst_rm_vars: assumes "wf\<^sub>t\<^sub>r\<^sub>m (t \ \)" shows "wf\<^sub>t\<^sub>r\<^sub>m (t \ rm_vars X \)" using assms proof (induction t) case (Fun f T) have "wf\<^sub>t\<^sub>r\<^sub>m (t \ \)" when "t \ set T" for t using that wf_trm_param[of f "map (\t. t \ \) T"] Fun.prems by auto hence "wf\<^sub>t\<^sub>r\<^sub>m (t \ rm_vars X \)" when "t \ set T" for t using that Fun.IH by simp moreover have "length T = arity f" using Fun.prems unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto ultimately show ?case unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto qed simp lemma wf_trm_subst_rm_vars': "wf\<^sub>t\<^sub>r\<^sub>m (\ v) \ wf\<^sub>t\<^sub>r\<^sub>m (rm_vars X \ v)" by auto lemma wf_trms_subst: assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (M \\<^sub>s\<^sub>e\<^sub>t \)" by (metis (no_types, lifting) assms imageE wf_trm_subst) lemma wf_trms_subst_rm_vars: assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (M \\<^sub>s\<^sub>e\<^sub>t \)" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (M \\<^sub>s\<^sub>e\<^sub>t rm_vars X \)" using assms wf_trm_subst_rm_vars by blast lemma wf_trms_subst_rm_vars': assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (rm_vars X \))" using assms by force lemma wf_trms_subst_compose: assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using assms subst_img_comp_subset' wf_trm_subst by blast lemma wf_trm_subst_compose: fixes \::"('fun, 'v) subst" assumes "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" "\x. wf\<^sub>t\<^sub>r\<^sub>m (\ x)" shows "wf\<^sub>t\<^sub>r\<^sub>m ((\ \\<^sub>s \) x)" using wf_trm_subst[of \ "\ x", OF wf_trm_subst_rangeI[OF assms(2)]] assms(1) subst_subst_compose[of "Var x" \ \] - subst_apply_term.simps(1)[of x \] - subst_apply_term.simps(1)[of x "\ \\<^sub>s \"] -by argo +by auto lemma wf_trms_Var_range: assumes "subst_range \ \ range Var" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using assms by fastforce lemma wf_trms_subst_compose_Var_range: assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and "subst_range \ \ range Var" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using assms wf_trms_subst_compose wf_trms_Var_range by metis+ lemma wf_trm_subst_inv: "wf\<^sub>t\<^sub>r\<^sub>m (t \ \) \ wf\<^sub>t\<^sub>r\<^sub>m t" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by (induct t) auto lemma wf_trms_subst_inv: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (M \\<^sub>s\<^sub>e\<^sub>t \) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" using wf_trm_subst_inv by fast lemma wf_trm_subterms: "wf\<^sub>t\<^sub>r\<^sub>m t \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subterms t)" using wf_trm_subterm by blast lemma wf_trms_subterms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subterms\<^sub>s\<^sub>e\<^sub>t M)" using wf_trm_subterms by blast lemma wf_trm_arity: "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T) \ length T = arity f" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by blast lemma wf_trm_subterm_arity: "wf\<^sub>t\<^sub>r\<^sub>m t \ Fun f T \ t \ length T = arity f" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by blast lemma unify_list_wf_trm: assumes "Unification.unify E B = Some U" "\(s,t) \ set E. wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t" and "\(v,t) \ set B. wf\<^sub>t\<^sub>r\<^sub>m t" shows "\(v,t) \ set U. wf\<^sub>t\<^sub>r\<^sub>m t" using assms proof (induction E B arbitrary: U rule: Unification.unify.induct) case (1 B U) thus ?case by auto next case (2 f T g S E B U) have wf_fun: "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" "wf\<^sub>t\<^sub>r\<^sub>m (Fun g S)" using "2.prems"(2) by auto from "2.prems"(1) obtain E' where *: "decompose (Fun f T) (Fun g S) = Some E'" and [simp]: "f = g" "length T = length S" "E' = zip T S" and **: "Unification.unify (E'@E) B = Some U" by (auto split: option.splits) hence "t \ Fun f T" "t' \ Fun g S" when "(t,t') \ set E'" for t t' using that by (metis zip_arg_subterm(1), metis zip_arg_subterm(2)) hence "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m t'" when "(t,t') \ set E'" for t t' using wf_trm_subterm wf_fun \f = g\ that by blast+ thus ?case using "2.IH"[OF * ** _ "2.prems"(3)] "2.prems"(2) by fastforce next case (3 v t E B) hence *: "\(w,x) \ set ((v, t) # B). wf\<^sub>t\<^sub>r\<^sub>m x" and **: "\(s,t) \ set E. wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m t" by auto show ?case proof (cases "t = Var v") case True thus ?thesis using "3.prems" "3.IH"(1) by auto next case False hence "v \ fv t" using "3.prems"(1) by auto hence "Unification.unify (subst_list (subst v t) E) ((v, t)#B) = Some U" using \t \ Var v\ "3.prems"(1) by auto moreover have "\(s, t) \ set (subst_list (subst v t) E). wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t" using wf_trm_subst_singleton[OF _ \wf\<^sub>t\<^sub>r\<^sub>m t\] "3.prems"(2) unfolding subst_list_def subst_def by auto ultimately show ?thesis using "3.IH"(2)[OF \t \ Var v\ \v \ fv t\ _ _ *] by metis qed next case (4 f T v E B U) hence *: "\(w,x) \ set ((v, Fun f T) # B). wf\<^sub>t\<^sub>r\<^sub>m x" and **: "\(s,t) \ set E. wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" by auto have "v \ fv (Fun f T)" using "4.prems"(1) by force hence "Unification.unify (subst_list (subst v (Fun f T)) E) ((v, Fun f T)#B) = Some U" using "4.prems"(1) by auto moreover have "\(s, t) \ set (subst_list (subst v (Fun f T)) E). wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t" using wf_trm_subst_singleton[OF _ \wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)\] "4.prems"(2) unfolding subst_list_def subst_def by auto ultimately show ?case using "4.IH"[OF \v \ fv (Fun f T)\ _ _ *] by metis qed lemma mgu_wf_trm: assumes "mgu s t = Some \" "wf\<^sub>t\<^sub>r\<^sub>m s" "wf\<^sub>t\<^sub>r\<^sub>m t" shows "wf\<^sub>t\<^sub>r\<^sub>m (\ v)" proof - from assms obtain \' where "subst_of \' = \" "\(v,t) \ set \'. wf\<^sub>t\<^sub>r\<^sub>m t" using unify_list_wf_trm[of "[(s,t)]" "[]"] by (auto simp: mgu_def split: option.splits) thus ?thesis proof (induction \' arbitrary: \ v rule: List.rev_induct) case (snoc x \' \ v) define \ where "\ = subst_of \'" hence "wf\<^sub>t\<^sub>r\<^sub>m (\ v)" for v using snoc.prems(2) snoc.IH[of \] by fastforce moreover obtain w t where x: "x = (w,t)" by (metis surj_pair) hence \: "\ = Var(w := t) \\<^sub>s \" using snoc.prems(1) by (simp add: subst_def \_def) moreover have "wf\<^sub>t\<^sub>r\<^sub>m t" using snoc.prems(2) x by auto ultimately show ?case using wf_trm_subst[of _ t] unfolding subst_compose_def by auto qed (simp add: wf\<^sub>t\<^sub>r\<^sub>m_def) qed lemma mgu_wf_trms: assumes "mgu s t = Some \" "wf\<^sub>t\<^sub>r\<^sub>m s" "wf\<^sub>t\<^sub>r\<^sub>m t" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_wf_trm[OF assms] by simp subsection \Definitions: Intruder Deduction Relations\ text \ A standard Dolev-Yao intruder. \ inductive intruder_deduct::"('fun,'var) terms \ ('fun,'var) term \ bool" where Axiom[simp]: "t \ M \ intruder_deduct M t" | Compose[simp]: "\length T = arity f; public f; \t. t \ set T \ intruder_deduct M t\ \ intruder_deduct M (Fun f T)" | Decompose: "\intruder_deduct M t; Ana t = (K, T); \k. k \ set K \ intruder_deduct M k; t\<^sub>i \ set T\ \ intruder_deduct M t\<^sub>i" text \ A variant of the intruder relation which limits the intruder to composition only. \ inductive intruder_synth::"('fun,'var) terms \ ('fun,'var) term \ bool" where AxiomC[simp]: "t \ M \ intruder_synth M t" | ComposeC[simp]: "\length T = arity f; public f; \t. t \ set T \ intruder_synth M t\ \ intruder_synth M (Fun f T)" adhoc_overloading INTRUDER_DEDUCT intruder_deduct adhoc_overloading INTRUDER_SYNTH intruder_synth lemma intruder_deduct_induct[consumes 1, case_names Axiom Compose Decompose]: assumes "M \ t" "\t. t \ M \ P M t" "\T f. \length T = arity f; public f; \t. t \ set T \ M \ t; \t. t \ set T \ P M t\ \ P M (Fun f T)" "\t K T t\<^sub>i. \M \ t; P M t; Ana t = (K, T); \k. k \ set K \ M \ k; \k. k \ set K \ P M k; t\<^sub>i \ set T\ \ P M t\<^sub>i" shows "P M t" using assms by (induct rule: intruder_deduct.induct) blast+ lemma intruder_synth_induct[consumes 1, case_names AxiomC ComposeC]: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" assumes "M \\<^sub>c t" "\t. t \ M \ P M t" "\T f. \length T = arity f; public f; \t. t \ set T \ M \\<^sub>c t; \t. t \ set T \ P M t\ \ P M (Fun f T)" shows "P M t" using assms by (induct rule: intruder_synth.induct) auto subsection \Definitions: Analyzed Knowledge and Public Ground Well-formed Terms (PGWTs)\ definition analyzed::"('fun,'var) terms \ bool" where "analyzed M \ \t. M \ t \ M \\<^sub>c t" definition analyzed_in where "analyzed_in t M \ \K R. (Ana t = (K,R) \ (\k \ set K. M \\<^sub>c k)) \ (\r \ set R. M \\<^sub>c r)" definition decomp_closure::"('fun,'var) terms \ ('fun,'var) terms \ bool" where "decomp_closure M M' \ \t. M \ t \ (\t' \ M. t \ t') \ t \ M'" inductive public_ground_wf_term::"('fun,'var) term \ bool" where PGWT[simp]: "\public f; arity f = length T; \t. t \ set T \ public_ground_wf_term t\ \ public_ground_wf_term (Fun f T)" abbreviation "public_ground_wf_terms \ {t. public_ground_wf_term t}" lemma public_const_deduct: assumes "c \ \\<^sub>p\<^sub>u\<^sub>b" shows "M \ Fun c []" "M \\<^sub>c Fun c []" proof - have "arity c = 0" "public c" using const_arity_eq_zero \c \ \\<^sub>p\<^sub>u\<^sub>b\ by auto thus "M \ Fun c []" "M \\<^sub>c Fun c []" using intruder_synth.ComposeC[OF _ \public c\, of "[]"] intruder_deduct.Compose[OF _ \public c\, of "[]"] by auto qed lemma public_const_deduct'[simp]: assumes "arity c = 0" "public c" shows "M \ Fun c []" "M \\<^sub>c Fun c []" using intruder_deduct.Compose[of "[]" c] intruder_synth.ComposeC[of "[]" c] assms by simp_all lemma private_fun_deduct_in_ik: assumes t: "M \ t" "Fun f T \ subterms t" and f: "\public f" shows "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t M" using t proof (induction t rule: intruder_deduct.induct) case Decompose thus ?case by (meson Ana_subterm psubsetD term.order_trans) qed (auto simp add: f in_subterms_Union) lemma private_fun_deduct_in_ik': assumes t: "M \ Fun f T" and f: "\public f" shows "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t M" by (rule private_fun_deduct_in_ik[OF t term.order_refl f]) lemma pgwt_public: "\public_ground_wf_term t; Fun f T \ t\ \ public f" by (induct t rule: public_ground_wf_term.induct) auto lemma pgwt_ground: "public_ground_wf_term t \ fv t = {}" by (induct t rule: public_ground_wf_term.induct) auto lemma pgwt_fun: "public_ground_wf_term t \ \f T. t = Fun f T" using pgwt_ground[of t] by (cases t) auto lemma pgwt_arity: "\public_ground_wf_term t; Fun f T \ t\ \ arity f = length T" by (induct t rule: public_ground_wf_term.induct) auto lemma pgwt_wellformed: "public_ground_wf_term t \ wf\<^sub>t\<^sub>r\<^sub>m t" by (induct t rule: public_ground_wf_term.induct) auto lemma pgwt_deducible: "public_ground_wf_term t \ M \\<^sub>c t" by (induct t rule: public_ground_wf_term.induct) auto lemma pgwt_is_empty_synth: "public_ground_wf_term t \ {} \\<^sub>c t" proof - { fix M::"('fun,'var) term set" assume "M \\<^sub>c t" "M = {}" hence "public_ground_wf_term t" by (induct t rule: intruder_synth.induct) auto } thus ?thesis using pgwt_deducible by auto qed lemma ideduct_synth_subst_apply: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" assumes "{} \\<^sub>c t" "\v. M \\<^sub>c \ v" shows "M \\<^sub>c t \ \" proof - { fix M'::"('fun,'var) term set" assume "M' \\<^sub>c t" "M' = {}" hence "M \\<^sub>c t \ \" proof (induction t rule: intruder_synth.induct) case (ComposeC T f M') hence "length (map (\t. t \ \) T) = arity f" "\x. x \ set (map (\t. t \ \) T) \ M \\<^sub>c x" by auto thus ?case using intruder_synth.ComposeC[of "map (\t. t \ \) T" f M] \public f\ by fastforce qed simp } thus ?thesis using assms by metis qed subsection \Lemmata: Monotonicity, Deduction of Private Constants, etc.\ context begin lemma ideduct_mono: "\M \ t; M \ M'\ \ M' \ t" proof (induction rule: intruder_deduct.induct) case (Decompose M t K T t\<^sub>i) have "\k. k \ set K \ M' \ k" using Decompose.IH \M \ M'\ by simp moreover have "M' \ t" using Decompose.IH \M \ M'\ by simp ultimately show ?case using Decompose.hyps intruder_deduct.Decompose by blast qed auto lemma ideduct_synth_mono: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" shows "\M \\<^sub>c t; M \ M'\ \ M' \\<^sub>c t" by (induct rule: intruder_synth.induct) auto context begin \ \Used by \inductive_set\\ private lemma ideduct_mono_set[mono_set]: "M \ N \ M \ t \ N \ t" "M \ N \ M \\<^sub>c t \ N \\<^sub>c t" using ideduct_mono ideduct_synth_mono by (blast, blast) end lemma ideduct_reduce: "\M \ M' \ t; \t'. t' \ M' \ M \ t'\ \ M \ t" proof (induction rule: intruder_deduct_induct) case Decompose thus ?case using intruder_deduct.Decompose by blast qed auto lemma ideduct_synth_reduce: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" shows "\M \ M' \\<^sub>c t; \t'. t' \ M' \ M \\<^sub>c t'\ \ M \\<^sub>c t" by (induct rule: intruder_synth_induct) auto lemma ideduct_mono_eq: assumes "\t. M \ t \ M' \ t" shows "M \ N \ t \ M' \ N \ t" proof show "M \ N \ t \ M' \ N \ t" proof (induction t rule: intruder_deduct_induct) case (Axiom t) thus ?case proof (cases "t \ M") case True hence "M \ t" using intruder_deduct.Axiom by metis thus ?thesis using assms ideduct_mono[of M' t "M' \ N"] by simp qed auto next case (Compose T f) thus ?case using intruder_deduct.Compose by auto next case (Decompose t K T t\<^sub>i) thus ?case using intruder_deduct.Decompose[of "M' \ N" t K T] by auto qed show "M' \ N \ t \ M \ N \ t" proof (induction t rule: intruder_deduct_induct) case (Axiom t) thus ?case proof (cases "t \ M'") case True hence "M' \ t" using intruder_deduct.Axiom by metis thus ?thesis using assms ideduct_mono[of M t "M \ N"] by simp qed auto next case (Compose T f) thus ?case using intruder_deduct.Compose by auto next case (Decompose t K T t\<^sub>i) thus ?case using intruder_deduct.Decompose[of "M \ N" t K T] by auto qed qed lemma deduct_synth_subterm: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" assumes "M \\<^sub>c t" "s \ subterms t" "\m \ M. \s \ subterms m. M \\<^sub>c s" shows "M \\<^sub>c s" using assms by (induct t rule: intruder_synth.induct) auto lemma deduct_if_synth[intro, dest]: "M \\<^sub>c t \ M \ t" by (induct rule: intruder_synth.induct) auto private lemma ideduct_ik_eq: assumes "\t \ M. M' \ t" shows "M' \ t \ M' \ M \ t" by (meson assms ideduct_mono ideduct_reduce sup_ge1) private lemma synth_if_deduct_empty: "{} \ t \ {} \\<^sub>c t" proof (induction t rule: intruder_deduct_induct) case (Decompose t K M m) then obtain f T where "t = Fun f T" "m \ set T" using Ana_fun_subterm Ana_var by (cases t) fastforce+ with Decompose.IH(1) show ?case by (induction rule: intruder_synth_induct) auto qed auto private lemma ideduct_deduct_synth_mono_eq: assumes "\t. M \ t \ M' \\<^sub>c t" "M \ M'" and "\t. M' \ N \ t \ M' \ N \ D \\<^sub>c t" shows "M \ N \ t \ M' \ N \ D \\<^sub>c t" proof - have "\m \ M'. M \ m" using assms(1) by auto hence "\t. M \ t \ M' \ t" by (metis assms(1,2) deduct_if_synth ideduct_reduce sup.absorb2) hence "\t. M' \ N \ t \ M \ N \ t" by (meson ideduct_mono_eq) thus ?thesis by (meson assms(3)) qed lemma ideduct_subst: "M \ t \ M \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" proof (induction t rule: intruder_deduct_induct) case (Compose T f) hence "length (map (\t. t \ \) T) = arity f" "\t. t \ set T \ M \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" by auto thus ?case using intruder_deduct.Compose[OF _ Compose.hyps(2), of "map (\t. t \ \) T"] by auto next case (Decompose t K M' m') hence "Ana (t \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" "\k. k \ set (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) \ M \\<^sub>s\<^sub>e\<^sub>t \ \ k" "m' \ \ \ set (M' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Ana_subst[OF Decompose.hyps(2)] by fastforce+ thus ?case using intruder_deduct.Decompose[OF Decompose.IH(1)] by metis qed simp lemma ideduct_synth_subst: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" and \::"('fun,'var) subst" shows "M \\<^sub>c t \ M \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" proof (induction t rule: intruder_synth_induct) case (ComposeC T f) hence "length (map (\t. t \ \) T) = arity f" "\t. t \ set T \ M \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" by auto thus ?case using intruder_synth.ComposeC[OF _ ComposeC.hyps(2), of "map (\t. t \ \) T"] by auto qed simp lemma ideduct_vars: assumes "M \ t" shows "fv t \ fv\<^sub>s\<^sub>e\<^sub>t M" using assms proof (induction t rule: intruder_deduct_induct) case (Decompose t K T t\<^sub>i) thus ?case using Ana_vars(2) fv_subset by blast qed auto lemma ideduct_synth_vars: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" assumes "M \\<^sub>c t" shows "fv t \ fv\<^sub>s\<^sub>e\<^sub>t M" using assms by (induct t rule: intruder_synth_induct) auto lemma ideduct_synth_priv_fun_in_ik: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" assumes "M \\<^sub>c t" "f \ funs_term t" "\public f" shows "f \ \(funs_term ` M)" using assms by (induct t rule: intruder_synth_induct) auto lemma ideduct_synth_priv_const_in_ik: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" assumes "M \\<^sub>c Fun c []" "\public c" shows "Fun c [] \ M" using intruder_synth.cases[OF assms(1)] assms(2) by fast lemma ideduct_synth_ik_replace: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" assumes "\t \ M. N \\<^sub>c t" and "M \\<^sub>c t" shows "N \\<^sub>c t" using assms(2,1) by (induct t rule: intruder_synth.induct) auto end subsection \Lemmata: Analyzed Intruder Knowledge Closure\ lemma deducts_eq_if_analyzed: "analyzed M \ M \ t \ M \\<^sub>c t" unfolding analyzed_def by auto lemma closure_is_superset: "decomp_closure M M' \ M \ M'" unfolding decomp_closure_def by force lemma deduct_if_closure_deduct: "\M' \ t; decomp_closure M M'\ \ M \ t" proof (induction t rule: intruder_deduct.induct) case (Decompose M' t K T t\<^sub>i) thus ?case using intruder_deduct.Decompose[OF _ \Ana t = (K,T)\ _ \t\<^sub>i \ set T\] by simp qed (auto simp add: decomp_closure_def) lemma deduct_if_closure_synth: "\decomp_closure M M'; M' \\<^sub>c t\ \ M \ t" using deduct_if_closure_deduct by blast lemma decomp_closure_subterms_composable: assumes "decomp_closure M M'" and "M' \\<^sub>c t'" "M' \ t" "t \ t'" shows "M' \\<^sub>c t" using \M' \\<^sub>c t'\ assms proof (induction t' rule: intruder_synth.induct) case (AxiomC t' M') have "M \ t" using \M' \ t\ deduct_if_closure_deduct AxiomC.prems(1) by blast moreover { have "\s \ M. t' \ s" using \t' \ M'\ AxiomC.prems(1) unfolding decomp_closure_def by blast hence "\s \ M. t \ s" using \t \ t'\ term.order_trans by auto } ultimately have "t \ M'" using AxiomC.prems(1) unfolding decomp_closure_def by blast thus ?case by simp next case (ComposeC T f M') let ?t' = "Fun f T" { assume "t = ?t'" have "M' \\<^sub>c t" using \M' \\<^sub>c ?t'\ \t = ?t'\ by simp } moreover { assume "t \ ?t'" have "\x \ set T. t \ x" using \t \ ?t'\ \t \ ?t'\ by simp hence "M' \\<^sub>c t" using ComposeC.IH ComposeC.prems(1,3) ComposeC.hyps(3) by blast } ultimately show ?case using cases_simp[of "t = ?t'" "M' \\<^sub>c t"] by simp qed lemma decomp_closure_analyzed: assumes "decomp_closure M M'" shows "analyzed M'" proof - { fix t assume "M' \ t" have "M' \\<^sub>c t" using \M' \ t\ assms proof (induction t rule: intruder_deduct.induct) case (Decompose M' t K T t\<^sub>i) hence "M' \ t\<^sub>i" using Decompose.hyps intruder_deduct.Decompose by blast moreover have "t\<^sub>i \ t" using Decompose.hyps(4) Ana_subterm[OF Decompose.hyps(2)] by blast moreover have "M' \\<^sub>c t" using Decompose.IH(1) Decompose.prems by blast ultimately show "M' \\<^sub>c t\<^sub>i" using decomp_closure_subterms_composable Decompose.prems by blast qed auto } moreover have "\t. M \\<^sub>c t \ M \ t" by auto ultimately show ?thesis by (auto simp add: decomp_closure_def analyzed_def) qed lemma analyzed_if_all_analyzed_in: assumes M: "\t \ M. analyzed_in t M" shows "analyzed M" proof (unfold analyzed_def, intro allI iffI) fix t assume t: "M \ t" thus "M \\<^sub>c t" proof (induction t rule: intruder_deduct_induct) case (Decompose t K T t\<^sub>i) { assume "t \ M" hence ?case using M Decompose.IH(2) Decompose.hyps(2,4) unfolding analyzed_in_def by fastforce } moreover { fix f S assume "t = Fun f S" "\s. s \ set S \ M \\<^sub>c s" hence ?case using Ana_fun_subterm[of f S] Decompose.hyps(2,4) by blast } ultimately show ?case using intruder_synth.cases[OF Decompose.IH(1), of ?case] by blast qed simp_all qed auto lemma analyzed_is_all_analyzed_in: "(\t \ M. analyzed_in t M) \ analyzed M" proof show "analyzed M \ \t \ M. analyzed_in t M" unfolding analyzed_in_def analyzed_def by (auto intro: intruder_deduct.Decompose[OF intruder_deduct.Axiom]) qed (rule analyzed_if_all_analyzed_in) lemma ik_has_synth_ik_closure: fixes M :: "('fun,'var) terms" shows "\M'. (\t. M \ t \ M' \\<^sub>c t) \ decomp_closure M M' \ (finite M \ finite M')" proof - let ?M' = "{t. M \ t \ (\t' \ M. t \ t')}" have M'_closes: "decomp_closure M ?M'" unfolding decomp_closure_def by simp hence "M \ ?M'" using closure_is_superset by simp have "\t. ?M' \\<^sub>c t \ M \ t" using deduct_if_closure_synth[OF M'_closes] by blast moreover have "\t. M \ t \ ?M' \ t" using ideduct_mono[OF _ \M \ ?M'\] by simp moreover have "analyzed ?M'" using decomp_closure_analyzed[OF M'_closes] . ultimately have "\t. M \ t \ ?M' \\<^sub>c t" unfolding analyzed_def by blast moreover have "finite M \ finite ?M'" by auto ultimately show ?thesis using M'_closes by blast qed lemma deducts_eq_if_empty_ik: "{} \ t \ {} \\<^sub>c t" using analyzed_is_all_analyzed_in[of "{}"] deducts_eq_if_analyzed[of "{}" t] by blast subsection \Intruder Variants: Numbered and Composition-Restricted Intruder Deduction Relations\ text \ A variant of the intruder relation which restricts composition to only those terms that satisfy a given predicate Q. \ inductive intruder_deduct_restricted:: "('fun,'var) terms \ (('fun,'var) term \ bool) \ ('fun,'var) term \ bool" ("\_;_\ \\<^sub>r _" 50) where AxiomR[simp]: "t \ M \ \M; Q\ \\<^sub>r t" | ComposeR[simp]: "\length T = arity f; public f; \t. t \ set T \ \M; Q\ \\<^sub>r t; Q (Fun f T)\ \ \M; Q\ \\<^sub>r Fun f T" | DecomposeR: "\\M; Q\ \\<^sub>r t; Ana t = (K, T); \k. k \ set K \ \M; Q\ \\<^sub>r k; t\<^sub>i \ set T\ \ \M; Q\ \\<^sub>r t\<^sub>i" text \ A variant of the intruder relation equipped with a number representing the height of the derivation tree (i.e., \\M; k\ \\<^sub>n t\ iff k is the maximum number of applications of the compose and decompose rules in any path of the derivation tree for \M \ t\). \ inductive intruder_deduct_num:: "('fun,'var) terms \ nat \ ('fun,'var) term \ bool" ("\_; _\ \\<^sub>n _" 50) where AxiomN[simp]: "t \ M \ \M; 0\ \\<^sub>n t" | ComposeN[simp]: "\length T = arity f; public f; \t. t \ set T \ \M; steps t\ \\<^sub>n t\ \ \M; Suc (Max (insert 0 (steps ` set T)))\ \\<^sub>n Fun f T" | DecomposeN: "\\M; n\ \\<^sub>n t; Ana t = (K, T); \k. k \ set K \ \M; steps k\ \\<^sub>n k; t\<^sub>i \ set T\ \ \M; Suc (Max (insert n (steps ` set K)))\ \\<^sub>n t\<^sub>i" lemma intruder_deduct_restricted_induct[consumes 1, case_names AxiomR ComposeR DecomposeR]: assumes "\M; Q\ \\<^sub>r t" "\t. t \ M \ P M Q t" "\T f. \length T = arity f; public f; \t. t \ set T \ \M; Q\ \\<^sub>r t; \t. t \ set T \ P M Q t; Q (Fun f T) \ \ P M Q (Fun f T)" "\t K T t\<^sub>i. \\M; Q\ \\<^sub>r t; P M Q t; Ana t = (K, T); \k. k \ set K \ \M; Q\ \\<^sub>r k; \k. k \ set K \ P M Q k; t\<^sub>i \ set T\ \ P M Q t\<^sub>i" shows "P M Q t" using assms by (induct t rule: intruder_deduct_restricted.induct) blast+ lemma intruder_deduct_num_induct[consumes 1, case_names AxiomN ComposeN DecomposeN]: assumes "\M; n\ \\<^sub>n t" "\t. t \ M \ P M 0 t" "\T f steps. \length T = arity f; public f; \t. t \ set T \ \M; steps t\ \\<^sub>n t; \t. t \ set T \ P M (steps t) t\ \ P M (Suc (Max (insert 0 (steps ` set T)))) (Fun f T)" "\t K T t\<^sub>i steps n. \\M; n\ \\<^sub>n t; P M n t; Ana t = (K, T); \k. k \ set K \ \M; steps k\ \\<^sub>n k; t\<^sub>i \ set T; \k. k \ set K \ P M (steps k) k\ \ P M (Suc (Max (insert n (steps ` set K)))) t\<^sub>i" shows "P M n t" using assms by (induct rule: intruder_deduct_num.induct) blast+ lemma ideduct_restricted_mono: "\\M; P\ \\<^sub>r t; M \ M'\ \ \M'; P\ \\<^sub>r t" proof (induction rule: intruder_deduct_restricted_induct) case (DecomposeR t K T t\<^sub>i) have "\k. k \ set K \ \M'; P\ \\<^sub>r k" using DecomposeR.IH \M \ M'\ by simp moreover have "\M'; P\ \\<^sub>r t" using DecomposeR.IH \M \ M'\ by simp ultimately show ?case using DecomposeR intruder_deduct_restricted.DecomposeR[OF _ DecomposeR.hyps(2) _ DecomposeR.hyps(4)] by blast qed auto subsection \Lemmata: Intruder Deduction Equivalences\ lemma deduct_if_restricted_deduct: "\M;P\ \\<^sub>r m \ M \ m" proof (induction m rule: intruder_deduct_restricted_induct) case (DecomposeR t K T t\<^sub>i) thus ?case using intruder_deduct.Decompose by blast qed simp_all lemma restricted_deduct_if_restricted_ik: assumes "\M;P\ \\<^sub>r m" "\m \ M. P m" and P: "\t t'. P t \ t' \ t \ P t'" shows "P m" using assms(1) proof (induction m rule: intruder_deduct_restricted_induct) case (DecomposeR t K T t\<^sub>i) obtain f S where "t = Fun f S" using Ana_var \t\<^sub>i \ set T\ \Ana t = (K, T)\ by (cases t) auto thus ?case using DecomposeR assms(2) P Ana_subterm by blast qed (simp_all add: assms(2)) lemma deduct_restricted_if_synth: assumes P: "P m" "\t t'. P t \ t' \ t \ P t'" and m: "M \\<^sub>c m" shows "\M; P\ \\<^sub>r m" using m P(1) proof (induction m rule: intruder_synth_induct) case (ComposeC T f) hence "\M; P\ \\<^sub>r t" when t: "t \ set T" for t using t P(2) subtermeqI''[of _ T f] by fastforce thus ?case using intruder_deduct_restricted.ComposeR[OF ComposeC.hyps(1,2)] ComposeC.prems(1) by metis qed simp lemma deduct_zero_in_ik: assumes "\M; 0\ \\<^sub>n t" shows "t \ M" proof - { fix k assume "\M; k\ \\<^sub>n t" hence "k > 0 \ t \ M" by (induct t) auto } thus ?thesis using assms by auto qed lemma deduct_if_deduct_num: "\M; k\ \\<^sub>n t \ M \ t" by (induct t rule: intruder_deduct_num.induct) (metis intruder_deduct.Axiom, metis intruder_deduct.Compose, metis intruder_deduct.Decompose) lemma deduct_num_if_deduct: "M \ t \ \k. \M; k\ \\<^sub>n t" proof (induction t rule: intruder_deduct_induct) case (Compose T f) then obtain steps where *: "\t \ set T. \M; steps t\ \\<^sub>n t" by moura then obtain n where "\t \ set T. steps t \ n" using finite_nat_set_iff_bounded_le[of "steps ` set T"] by auto thus ?case using ComposeN[OF Compose.hyps(1,2), of M steps] * by force next case (Decompose t K T t\<^sub>i) hence "\u. u \ insert t (set K) \ \k. \M; k\ \\<^sub>n u" by auto then obtain steps where *: "\M; steps t\ \\<^sub>n t" "\t \ set K. \M; steps t\ \\<^sub>n t" by moura then obtain n where "steps t \ n" "\t \ set K. steps t \ n" using finite_nat_set_iff_bounded_le[of "steps ` insert t (set K)"] by auto thus ?case using DecomposeN[OF _ Decompose.hyps(2) _ Decompose.hyps(4), of M _ steps] * by force qed (metis AxiomN) lemma deduct_normalize: assumes M: "\m \ M. \f T. Fun f T \ m \ P f T" and t: "\M; k\ \\<^sub>n t" "Fun f T \ t" "\P f T" shows "\l \ k. (\M; l\ \\<^sub>n Fun f T) \ (\t \ set T. \j < l. \M; j\ \\<^sub>n t)" using t proof (induction t rule: intruder_deduct_num_induct) case (AxiomN t) thus ?case using M by auto next case (ComposeN T' f' steps) thus ?case proof (cases "Fun f' T' = Fun f T") case True hence "\M; Suc (Max (insert 0 (steps ` set T')))\ \\<^sub>n Fun f T" "T = T'" using intruder_deduct_num.ComposeN[OF ComposeN.hyps] by auto moreover have "\t. t \ set T \ \M; steps t\ \\<^sub>n t" using True ComposeN.hyps(3) by auto moreover have "\t. t \ set T \ steps t < Suc (Max (insert 0 (steps ` set T)))" using Max_less_iff[of "insert 0 (steps ` set T)" "Suc (Max (insert 0 (steps ` set T)))"] by auto ultimately show ?thesis by auto next case False then obtain t' where t': "t' \ set T'" "Fun f T \ t'" using ComposeN by auto hence "\l \ steps t'. (\M; l\ \\<^sub>n Fun f T) \ (\t \ set T. \j < l. \M; j\ \\<^sub>n t)" using ComposeN.IH[OF _ _ ComposeN.prems(2)] by auto moreover have "steps t' < Suc (Max (insert 0 (steps ` set T')))" using Max_less_iff[of "insert 0 (steps ` set T')" "Suc (Max (insert 0 (steps ` set T')))"] using t'(1) by auto ultimately show ?thesis using ComposeN.hyps(3)[OF t'(1)] by (meson Suc_le_eq le_Suc_eq le_trans) qed next case (DecomposeN t K T' t\<^sub>i steps n) hence *: "Fun f T \ t" using term.order_trans[of "Fun f T" t\<^sub>i t] Ana_subterm[of t K T'] by blast have "\l \ n. (\M; l\ \\<^sub>n Fun f T) \ (\t' \ set T. \j < l. \M; j\ \\<^sub>n t')" using DecomposeN.IH(1)[OF * DecomposeN.prems(2)] by auto moreover have "n < Suc (Max (insert n (steps ` set K)))" using Max_less_iff[of "insert n (steps ` set K)" "Suc (Max (insert n (steps ` set K)))"] by auto ultimately show ?case using DecomposeN.hyps(4) by (meson Suc_le_eq le_Suc_eq le_trans) qed lemma deduct_inv: assumes "\M; n\ \\<^sub>n t" shows "t \ M \ (\f T. t = Fun f T \ public f \ length T = arity f \ (\t \ set T. \l < n. \M; l\ \\<^sub>n t)) \ (\m \ subterms\<^sub>s\<^sub>e\<^sub>t M. (\l < n. \M; l\ \\<^sub>n m) \ (\k \ set (fst (Ana m)). \l < n. \M; l\ \\<^sub>n k) \ t \ set (snd (Ana m)))" (is "?P t n \ ?Q t n \ ?R t n") using assms proof (induction n arbitrary: t rule: nat_less_induct) case (1 n t) thus ?case proof (cases n) case 0 hence "t \ M" using deduct_zero_in_ik "1.prems"(1) by metis thus ?thesis by auto next case (Suc n') hence "\M; Suc n'\ \\<^sub>n t" "\m < Suc n'. \x. (\M; m\ \\<^sub>n x) \ ?P x m \ ?Q x m \ ?R x m" using "1.prems" "1.IH" by blast+ hence "?P t (Suc n') \ ?Q t (Suc n') \ ?R t (Suc n')" proof (induction t rule: intruder_deduct_num_induct) case (AxiomN t) thus ?case by simp next case (ComposeN T f steps) have "\t. t \ set T \ steps t < Suc (Max (insert 0 (steps ` set T)))" using Max_less_iff[of "insert 0 (steps ` set T)" "Suc (Max (insert 0 (steps ` set T)))"] by auto thus ?case using ComposeN.hyps by metis next case (DecomposeN t K T t\<^sub>i steps n) have 0: "n < Suc (Max (insert n (steps ` set K)))" "\k. k \ set K \ steps k < Suc (Max (insert n (steps ` set K)))" using Max_less_iff[of "insert n (steps ` set K)" "Suc (Max (insert n (steps ` set K)))"] by auto have IH1: "?P t j \ ?Q t j \ ?R t j" when jt: "j < n" "\M; j\ \\<^sub>n t" for j t using jt DecomposeN.prems(1) 0(1) by simp have IH2: "?P t n \ ?Q t n \ ?R t n" using DecomposeN.IH(1) IH1 by simp have 1: "\k \ set (fst (Ana t)). \l < Suc (Max (insert n (steps ` set K))). \M; l\ \\<^sub>n k" using DecomposeN.hyps(1,2,3) 0(2) by auto have 2: "t\<^sub>i \ set (snd (Ana t))" using DecomposeN.hyps(2,4) by fastforce have 3: "t \ subterms\<^sub>s\<^sub>e\<^sub>t M" when "t \ set (snd (Ana m))" "m \\<^sub>s\<^sub>e\<^sub>t M" for m using that(1) Ana_subterm[of m _ "snd (Ana m)"] in_subterms_subset_Union[OF that(2)] by (metis (no_types, lifting) prod.collapse psubsetD subsetCE subsetD) have 4: "?R t\<^sub>i (Suc (Max (insert n (steps ` set K))))" when "?R t n" using that 0(1) 1 2 3 DecomposeN.hyps(1) by (metis (no_types, lifting)) have 5: "?R t\<^sub>i (Suc (Max (insert n (steps ` set K))))" when "?P t n" using that 0(1) 1 2 DecomposeN.hyps(1) by blast have 6: ?case when *: "?Q t n" proof - obtain g S where g: "t = Fun g S" "public g" "length S = arity g" "\t \ set S. \l < n. \M; l\ \\<^sub>n t" using * by moura then obtain l where l: "l < n" "\M; l\ \\<^sub>n t\<^sub>i" using 0(1) DecomposeN.hyps(2,4) Ana_fun_subterm[of g S K T] by blast have **: "l < Suc (Max (insert n (steps ` set K)))" using l(1) 0(1) by simp show ?thesis using IH1[OF l] less_trans[OF _ **] by fastforce qed show ?case using IH2 4 5 6 by argo qed thus ?thesis using Suc by fast qed qed lemma deduct_inv': assumes "M \ Fun f ts" shows "Fun f ts \\<^sub>s\<^sub>e\<^sub>t M \ (\t \ set ts. M \ t)" proof - obtain k where k: "intruder_deduct_num M k (Fun f ts)" using deduct_num_if_deduct[OF assms] by fast have "Fun f ts \\<^sub>s\<^sub>e\<^sub>t M \ (\t \ set ts. \l. intruder_deduct_num M l t)" using deduct_inv[OF k] Ana_subterm'[of "Fun f ts"] in_subterms_subset_Union by blast thus ?thesis using deduct_if_deduct_num by blast qed lemma restricted_deduct_if_deduct: assumes M: "\m \ M. \f T. Fun f T \ m \ P (Fun f T)" and P_subterm: "\f T t. M \ Fun f T \ P (Fun f T) \ t \ set T \ P t" and P_Ana_key: "\t K T k. M \ t \ P t \ Ana t = (K, T) \ M \ k \ k \ set K \ P k" and m: "M \ m" "P m" shows "\M; P\ \\<^sub>r m" proof - { fix k assume "\M; k\ \\<^sub>n m" hence ?thesis using m(2) proof (induction k arbitrary: m rule: nat_less_induct) case (1 n m) thus ?case proof (cases n) case 0 hence "m \ M" using deduct_zero_in_ik "1.prems"(1) by metis thus ?thesis by auto next case (Suc n') hence "\M; Suc n'\ \\<^sub>n m" "\m < Suc n'. \x. (\M; m\ \\<^sub>n x) \ P x \ \M;P\ \\<^sub>r x" using "1.prems" "1.IH" by blast+ thus ?thesis using "1.prems"(2) proof (induction m rule: intruder_deduct_num_induct) case (ComposeN T f steps) have *: "steps t < Suc (Max (insert 0 (steps ` set T)))" when "t \ set T" for t using Max_less_iff[of "insert 0 (steps ` set T)"] that by blast have **: "P t" when "t \ set T" for t using P_subterm ComposeN.prems(2) that Fun_param_is_subterm[OF that] intruder_deduct.Compose[OF ComposeN.hyps(1,2)] deduct_if_deduct_num[OF ComposeN.hyps(3)] by blast have "\M; P\ \\<^sub>r t" when "t \ set T" for t using ComposeN.prems(1) ComposeN.hyps(3)[OF that] *[OF that] **[OF that] by blast thus ?case by (metis intruder_deduct_restricted.ComposeR[OF ComposeN.hyps(1,2)] ComposeN.prems(2)) next case (DecomposeN t K T t\<^sub>i steps l) show ?case proof (cases "P t") case True hence "\k. k \ set K \ P k" using P_Ana_key DecomposeN.hyps(1,2,3) deduct_if_deduct_num by blast moreover have "\k m x. k \ set K \ m < steps k \ \M; m\ \\<^sub>n x \ P x \ \M;P\ \\<^sub>r x" proof - fix k m x assume *: "k \ set K" "m < steps k" "\M; m\ \\<^sub>n x" "P x" have "steps k \ insert l (steps ` set K)" using *(1) by simp hence "m < Suc (Max (insert l (steps ` set K)))" using less_trans[OF *(2), of "Suc (Max (insert l (steps ` set K)))"] Max_less_iff[of "insert l (steps ` set K)" "Suc (Max (insert l (steps ` set K)))"] by auto thus "\M;P\ \\<^sub>r x" using DecomposeN.prems(1) *(3,4) by simp qed ultimately have "\k. k \ set K \ \M; P\ \\<^sub>r k" using DecomposeN.IH(2) by auto moreover have "\M; P\ \\<^sub>r t" using True DecomposeN.prems(1) DecomposeN.hyps(1) le_imp_less_Suc Max_less_iff[of "insert l (steps ` set K)" "Suc (Max (insert l (steps ` set K)))"] by blast ultimately show ?thesis using intruder_deduct_restricted.DecomposeR[OF _ DecomposeN.hyps(2) _ DecomposeN.hyps(4)] by metis next case False obtain g S where gS: "t = Fun g S" using DecomposeN.hyps(2,4) by (cases t) moura+ hence *: "Fun g S \ t" "\P (Fun g S)" using False by force+ have "\jM; j\ \\<^sub>n t\<^sub>i" using gS DecomposeN.hyps(2,4) Ana_fun_subterm[of g S K T] deduct_normalize[of M "\f T. P (Fun f T)", OF M DecomposeN.hyps(1) *] by force hence "\jM; j\ \\<^sub>n t\<^sub>i" using Max_less_iff[of "insert l (steps ` set K)" "Suc (Max (insert l (steps ` set K)))"] less_trans[of _ l "Suc (Max (insert l (steps ` set K)))"] by blast thus ?thesis using DecomposeN.prems(1,2) by meson qed qed auto qed qed } thus ?thesis using deduct_num_if_deduct m(1) by metis qed lemma restricted_deduct_if_deduct': assumes "\m \ M. P m" and "\t t'. P t \ t' \ t \ P t'" and "\t K T k. P t \ Ana t = (K, T) \ k \ set K \ P k" and "M \ m" "P m" shows "\M; P\ \\<^sub>r m" using restricted_deduct_if_deduct[of M P m] assms by blast lemma private_const_deduct: assumes c: "\public c" "M \ (Fun c []::('fun,'var) term)" shows "Fun c [] \ M \ (\m \ subterms\<^sub>s\<^sub>e\<^sub>t M. M \ m \ (\k \ set (fst (Ana m)). M \ m) \ Fun c [] \ set (snd (Ana m)))" proof - obtain n where "\M; n\ \\<^sub>n Fun c []" using c(2) deduct_num_if_deduct by moura hence "Fun c [] \ M \ (\m \ subterms\<^sub>s\<^sub>e\<^sub>t M. (\l < n. \M; l\ \\<^sub>n m) \ (\k \ set (fst (Ana m)). \l < n. \M; l\ \\<^sub>n k) \ Fun c [] \ set (snd (Ana m)))" using deduct_inv[of M n "Fun c []"] c(1) by fast thus ?thesis using deduct_if_deduct_num[of M] by blast qed lemma private_fun_deduct_in_ik'': assumes t: "M \ Fun f T" "Fun c [] \ set T" "\m \ subterms\<^sub>s\<^sub>e\<^sub>t M. Fun f T \ set (snd (Ana m))" and c: "\public c" "Fun c [] \ M" "\m \ subterms\<^sub>s\<^sub>e\<^sub>t M. Fun c [] \ set (snd (Ana m))" shows "Fun f T \ M" proof - have *: "\n. \M; n\ \\<^sub>n Fun c []" using private_const_deduct[OF c(1)] c(2,3) deduct_if_deduct_num by blast obtain n where n: "\M; n\ \\<^sub>n Fun f T" using t(1) deduct_num_if_deduct by blast show ?thesis using deduct_inv[OF n] t(2,3) * by blast qed end subsection \Executable Definitions for Code Generation\ fun intruder_synth' where "intruder_synth' pu ar M (Var x) = (Var x \ M)" | "intruder_synth' pu ar M (Fun f T) = ( Fun f T \ M \ (pu f \ length T = ar f \ list_all (intruder_synth' pu ar M) T))" definition "wf\<^sub>t\<^sub>r\<^sub>m' ar t \ (\s \ subterms t. is_Fun s \ ar (the_Fun s) = length (args s))" definition "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' ar M \ (\t \ M. wf\<^sub>t\<^sub>r\<^sub>m' ar t)" definition "analyzed_in' An pu ar t M \ (case An t of (K,T) \ (\k \ set K. intruder_synth' pu ar M k) \ (\s \ set T. intruder_synth' pu ar M s))" lemma (in intruder_model) intruder_synth'_induct[consumes 1, case_names Var Fun]: assumes "intruder_synth' public arity M t" "\x. intruder_synth' public arity M (Var x) \ P (Var x)" "\f T. (\z. z \ set T \ intruder_synth' public arity M z \ P z) \ intruder_synth' public arity M (Fun f T) \ P (Fun f T) " shows "P t" using assms by (induct public arity M t rule: intruder_synth'.induct) auto lemma (in intruder_model) wf\<^sub>t\<^sub>r\<^sub>m_code[code_unfold]: "wf\<^sub>t\<^sub>r\<^sub>m t = wf\<^sub>t\<^sub>r\<^sub>m' arity t" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def wf\<^sub>t\<^sub>r\<^sub>m'_def by auto lemma (in intruder_model) wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[code_unfold]: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M = wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity M" using wf\<^sub>t\<^sub>r\<^sub>m_code unfolding wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s'_def by auto lemma (in intruder_model) intruder_synth_code[code_unfold]: "intruder_synth M t = intruder_synth' public arity M t" (is "?A \ ?B") proof show "?A \ ?B" proof (induction t rule: intruder_synth_induct) case (AxiomC t) thus ?case by (cases t) auto qed (fastforce simp add: list_all_iff) show "?B \ ?A" proof (induction t rule: intruder_synth'_induct) case (Fun f T) thus ?case proof (cases "Fun f T \ M") case False hence "public f" "length T = arity f" "list_all (intruder_synth' public arity M) T" using Fun.hyps by fastforce+ thus ?thesis using Fun.IH intruder_synth.ComposeC[of T f M] Ball_set[of T] by blast qed simp qed simp qed lemma (in intruder_model) analyzed_in_code[code_unfold]: "analyzed_in t M = analyzed_in' Ana public arity t M" using intruder_synth_code[of M] unfolding analyzed_in_def analyzed_in'_def by fastforce end diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Lazy_Intruder.thy b/thys/Stateful_Protocol_Composition_and_Typing/Lazy_Intruder.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/Lazy_Intruder.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/Lazy_Intruder.thy @@ -1,996 +1,994 @@ (* Title: Lazy_Intruder.thy Author: Andreas Viktor Hess, DTU SPDX-License-Identifier: BSD-3-Clause *) section \The Lazy Intruder\ theory Lazy_Intruder imports Strands_and_Constraints Intruder_Deduction begin context intruder_model begin subsection \Definition of the Lazy Intruder\ text \The lazy intruder constraint reduction system, defined as a relation on constraint states\ inductive_set LI_rel:: "((('fun,'var) strand \ (('fun,'var) subst)) \ ('fun,'var) strand \ (('fun,'var) subst)) set" and LI_rel' (infix "\" 50) and LI_rel_trancl (infix "\\<^sup>+" 50) and LI_rel_rtrancl (infix "\\<^sup>*" 50) where "A \ B \ (A,B) \ LI_rel" | "A \\<^sup>+ B \ (A,B) \ LI_rel\<^sup>+" | "A \\<^sup>* B \ (A,B) \ LI_rel\<^sup>*" | Compose: "\simple S; length T = arity f; public f\ \ (S@Send [Fun f T]#S',\) \ (S@(map Send1 T)@S',\)" | Unify: "\simple S; Fun f T' \ ik\<^sub>s\<^sub>t S; Some \ = mgu (Fun f T) (Fun f T')\ \ (S@Send [Fun f T]#S',\) \ ((S@S') \\<^sub>s\<^sub>t \,\ \\<^sub>s \)" | Equality: "\simple S; Some \ = mgu t t'\ \ (S@Equality _ t t'#S',\) \ ((S@S') \\<^sub>s\<^sub>t \,\ \\<^sub>s \)" text \A "pre-processing step" to be applied before constraint reduction. It transforms constraints such that exactly one message is transmitted in each message transmission step. It is sound and complete and preserves the various well-formedness properties required by the lazy intruder.\ fun LI_preproc where "LI_preproc [] = []" | "LI_preproc (Send ts#S) = map Send1 ts@LI_preproc S" | "LI_preproc (Receive ts#S) = map Receive1 ts@LI_preproc S" | "LI_preproc (x#S) = x#LI_preproc S" definition LI_preproc_prop where "LI_preproc_prop S \ \ts. Send ts \ set S \ Receive ts \ set S \ (\t. ts = [t])" subsection \Lemmata: Preprocessing \ lemma LI_preproc_preproc_prop: "LI_preproc_prop (LI_preproc S)" by (induct S rule: LI_preproc.induct) (auto simp add: LI_preproc_prop_def) lemma LI_preproc_sem_eq: "\M; S\\<^sub>c \ \ \M; LI_preproc S\\<^sub>c \" (is "?A \ ?B") proof show "?A \ ?B" proof (induction S rule: strand_sem_induct) case (ConsSnd M ts S) hence "\M; LI_preproc S\\<^sub>c \" "\M; map Send1 ts\\<^sub>c \" using strand_sem_Send_map(5) by auto moreover have "ik\<^sub>s\<^sub>t (map Send1 ts) \\<^sub>s\<^sub>e\<^sub>t \ = {}" unfolding ik\<^sub>s\<^sub>t_is_rcv_set by fastforce ultimately show ?case using strand_sem_append(1) by simp next case (ConsRcv M ts S) hence "\(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; LI_preproc S\\<^sub>c \" "\M; map Receive1 ts\\<^sub>c \" using strand_sem_Receive_map(3) by auto moreover have "ik\<^sub>s\<^sub>t (map Receive1 ts) \\<^sub>s\<^sub>e\<^sub>t \ = set ts \\<^sub>s\<^sub>e\<^sub>t \" unfolding ik\<^sub>s\<^sub>t_is_rcv_set by force ultimately show ?case using strand_sem_append(1) by (simp add: Un_commute) qed simp_all show "?B \ ?A" proof (induction S arbitrary: M rule: LI_preproc.induct) case (2 ts S) have "ik\<^sub>s\<^sub>t (map Send1 ts) \\<^sub>s\<^sub>e\<^sub>t \ = {}" unfolding ik\<^sub>s\<^sub>t_is_rcv_set by fastforce hence "\M; S\\<^sub>c \" "\M; map Send1 ts\\<^sub>c \" using 2 strand_sem_append(1) by auto thus ?case using strand_sem_Send_map(5) by simp next case (3 ts S) have "ik\<^sub>s\<^sub>t (map Receive1 ts) \\<^sub>s\<^sub>e\<^sub>t \ = set ts \\<^sub>s\<^sub>e\<^sub>t \" unfolding ik\<^sub>s\<^sub>t_is_rcv_set by force hence "\M \ (set ts \\<^sub>s\<^sub>e\<^sub>t \); S\\<^sub>c \" "\M; map Receive1 ts\\<^sub>c \" using 3 strand_sem_append(1) by auto thus ?case using strand_sem_Receive_map(3) by (simp add: Un_commute) qed simp_all qed lemma LI_preproc_sem_eq': "(\ \\<^sub>c \S, \\) \ (\ \\<^sub>c \LI_preproc S, \\)" using LI_preproc_sem_eq unfolding constr_sem_c_def by simp lemma LI_preproc_vars_eq: "fv\<^sub>s\<^sub>t (LI_preproc S) = fv\<^sub>s\<^sub>t S" "bvars\<^sub>s\<^sub>t (LI_preproc S) = bvars\<^sub>s\<^sub>t S" "vars\<^sub>s\<^sub>t (LI_preproc S) = vars\<^sub>s\<^sub>t S" by (induct S rule: LI_preproc.induct) auto lemma LI_preproc_trms_eq: "trms\<^sub>s\<^sub>t (LI_preproc S) = trms\<^sub>s\<^sub>t S" by (induct S rule: LI_preproc.induct) auto lemma LI_preproc_wf\<^sub>s\<^sub>t: assumes "wf\<^sub>s\<^sub>t X S" shows "wf\<^sub>s\<^sub>t X (LI_preproc S)" using assms proof (induction S arbitrary: X rule: wf\<^sub>s\<^sub>t_induct) case (ConsRcv X ts S) hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ X" "wf\<^sub>s\<^sub>t X (LI_preproc S)" by auto thus ?case using wf_Receive1_prefix by simp next case (ConsSnd X ts S) hence "wf\<^sub>s\<^sub>t (X \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) (LI_preproc S)" by simp thus ?case using wf_Send1_prefix by simp qed simp_all lemma LI_preproc_preserves_wellformedness: assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" shows "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r (LI_preproc S) \" using assms LI_preproc_vars_eq[of S] LI_preproc_wf\<^sub>s\<^sub>t[of "{}" S] unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by argo lemma LI_preproc_prop_SendE: assumes "LI_preproc_prop S" and "Send ts \ set S" shows "(\x. ts = [Var x]) \ (\f T. ts = [Fun f T])" proof - obtain t where "ts = [t]" using assms unfolding LI_preproc_prop_def by auto thus ?thesis by (cases t) auto qed lemma LI_preproc_prop_split: "LI_preproc_prop (S@S') \ LI_preproc_prop S \ LI_preproc_prop S'" (is "?A \ ?B") proof show "?A \ ?B" proof (induction S) case (Cons x S) thus ?case unfolding LI_preproc_prop_def by (cases x) auto qed (simp add: LI_preproc_prop_def) show "?B \ ?A" proof (induction S) case (Cons x S) thus ?case unfolding LI_preproc_prop_def by (cases x) auto qed (simp add: LI_preproc_prop_def) qed subsection \Lemma: The Lazy Intruder is Well-founded\ context begin private lemma LI_compose_measure_lt: "((S@(map Send1 T)@S',\\<^sub>1), (S@Send [Fun f T]#S',\\<^sub>2)) \ measure\<^sub>s\<^sub>t" using strand_fv_card_map_fun_eq[of S f T S'] strand_size_map_fun_lt[of T f] by (simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) private lemma LI_unify_measure_lt: assumes "Some \ = mgu (Fun f T) t" "fv t \ fv\<^sub>s\<^sub>t S" shows "(((S@S') \\<^sub>s\<^sub>t \,\\<^sub>1), (S@Send [Fun f T]#S',\\<^sub>2)) \ measure\<^sub>s\<^sub>t" proof (cases "\ = Var") assume "\ = Var" hence "(S@S') \\<^sub>s\<^sub>t \ = S@S'" by blast thus ?thesis using strand_fv_card_rm_fun_le[of S S' f T] by (auto simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) next assume "\ \ Var" then obtain v where "v \ fv (Fun f T) \ fv t" "subst_elim \ v" using mgu_eliminates[OF assms(1)[symmetric]] by metis hence v_in: "v \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using assms(2) by (auto simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) have "range_vars \ \ fv (Fun f T) \ fv\<^sub>s\<^sub>t S" using assms(2) mgu_vars_bounded[OF assms(1)[symmetric]] by auto hence img_bound: "range_vars \ \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" by auto have finite_fv: "finite (fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S'))" by auto have "v \ fv\<^sub>s\<^sub>t ((S@Send [Fun f T]#S') \\<^sub>s\<^sub>t \)" using strand_fv_subst_subset_if_subst_elim[OF \subst_elim \ v\] v_in by metis hence v_not_in: "v \ fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" by auto have "fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using strand_subst_fv_bounded_if_img_bounded[OF img_bound] by simp hence "fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using v_in v_not_in by blast hence "card (fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)) < card (fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S'))" using psubset_card_mono[OF finite_fv] by simp thus ?thesis by (auto simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) qed private lemma LI_equality_measure_lt: assumes "Some \ = mgu t t'" shows "(((S@S') \\<^sub>s\<^sub>t \,\\<^sub>1), (S@Equality a t t'#S',\\<^sub>2)) \ measure\<^sub>s\<^sub>t" proof (cases "\ = Var") assume "\ = Var" hence "(S@S') \\<^sub>s\<^sub>t \ = S@S'" by blast thus ?thesis using strand_fv_card_rm_eq_le[of S S' a t t'] by (auto simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) next assume "\ \ Var" then obtain v where "v \ fv t \ fv t'" "subst_elim \ v" using mgu_eliminates[OF assms(1)[symmetric]] by metis hence v_in: "v \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using assms by auto have "range_vars \ \ fv t \ fv t' \ fv\<^sub>s\<^sub>t S" using assms mgu_vars_bounded[OF assms(1)[symmetric]] by auto hence img_bound: "range_vars \ \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" by auto have finite_fv: "finite (fv\<^sub>s\<^sub>t (S@Equality a t t'#S'))" by auto have "v \ fv\<^sub>s\<^sub>t ((S@Equality a t t'#S') \\<^sub>s\<^sub>t \)" using strand_fv_subst_subset_if_subst_elim[OF \subst_elim \ v\] v_in by metis hence v_not_in: "v \ fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" by auto have "fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using strand_subst_fv_bounded_if_img_bounded[OF img_bound] by simp hence "fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using v_in v_not_in by blast hence "card (fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)) < card (fv\<^sub>s\<^sub>t (S@Equality a t t'#S'))" using psubset_card_mono[OF finite_fv] by simp thus ?thesis by (auto simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) qed private lemma LI_in_measure: "(S\<^sub>1,\\<^sub>1) \ (S\<^sub>2,\\<^sub>2) \ ((S\<^sub>2,\\<^sub>2),(S\<^sub>1,\\<^sub>1)) \ measure\<^sub>s\<^sub>t" proof (induction rule: LI_rel.induct) case (Compose S T f S' \) thus ?case using LI_compose_measure_lt[of S T S'] by metis next case (Unify S f U \ T S' \) hence "fv (Fun f U) \ fv\<^sub>s\<^sub>t S" using fv_snd_rcv_strand_subset(2)[of S] by force thus ?case using LI_unify_measure_lt[OF Unify.hyps(3), of S S'] by metis qed (metis LI_equality_measure_lt) private lemma LI_in_measure_trans: "(S\<^sub>1,\\<^sub>1) \\<^sup>+ (S\<^sub>2,\\<^sub>2) \ ((S\<^sub>2,\\<^sub>2),(S\<^sub>1,\\<^sub>1)) \ measure\<^sub>s\<^sub>t" by (induction rule: trancl.induct, metis surjective_pairing LI_in_measure) (metis (no_types, lifting) surjective_pairing LI_in_measure measure\<^sub>s\<^sub>t_trans trans_def) private lemma LI_converse_wellfounded_trans: "wf ((LI_rel\<^sup>+)\)" proof - have "(LI_rel\<^sup>+)\ \ measure\<^sub>s\<^sub>t" using LI_in_measure_trans by auto thus ?thesis using measure\<^sub>s\<^sub>t_wellfounded wf_subset by metis qed private lemma LI_acyclic_trans: "acyclic (LI_rel\<^sup>+)" using wf_acyclic[OF LI_converse_wellfounded_trans] acyclic_converse by metis private lemma LI_acyclic: "acyclic LI_rel" using LI_acyclic_trans acyclic_subset by (simp add: acyclic_def) lemma LI_no_infinite_chain: "\(\f. \i. f i \\<^sup>+ f (Suc i))" proof - have "\(\f. \i. (f (Suc i), f i) \ (LI_rel\<^sup>+)\)" using wf_iff_no_infinite_down_chain LI_converse_wellfounded_trans by metis thus ?thesis by simp qed private lemma LI_unify_finite: assumes "finite M" shows "finite {((S@Send [Fun f T]#S',\), ((S@S') \\<^sub>s\<^sub>t \,\ \\<^sub>s \)) | \ T'. simple S \ Fun f T' \ M \ Some \ = mgu (Fun f T) (Fun f T')}" using assms proof (induction M rule: finite_induct) case (insert m M) thus ?case proof (cases m) case (Fun g U) let ?a = "\\. ((S@Send [Fun f T]#S',\), ((S@S') \\<^sub>s\<^sub>t \,\ \\<^sub>s \))" let ?A = "\B. {?a \ | \ T'. simple S \ Fun f T' \ B \ Some \ = mgu (Fun f T) (Fun f T')}" have "?A (insert m M) = (?A M) \ (?A {m})" by auto moreover have "finite (?A {m})" proof (cases "\\. Some \ = mgu (Fun f T) (Fun g U)") case True then obtain \ where \: "Some \ = mgu (Fun f T) (Fun g U)" by blast have A_m_eq: "\\'. ?a \' \ ?A {m} \ ?a \ = ?a \'" proof - fix \' assume "?a \' \ ?A {m}" hence "\\. Some \ = mgu (Fun f T) (Fun g U) \ ?a \ = ?a \'" using \m = Fun g U\ by auto thus "?a \ = ?a \'" by (metis \ option.inject) qed have "?A {m} = {} \ ?A {m} = {?a \}" proof (cases "simple S \ ?A {m} \ {}") case True hence "simple S" "?A {m} \ {}" by meson+ hence "?A {m} = {?a \ | \. Some \ = mgu (Fun f T) (Fun g U)}" using \m = Fun g U\ by auto hence "?a \ \ ?A {m}" using \ by auto show ?thesis proof (rule ccontr) assume "\(?A {m} = {} \ ?A {m} = {?a \})" then obtain B where B: "?A {m} = insert (?a \) B" "?a \ \ B" "B \ {}" using \?A {m} \ {}\ \?a \ \ ?A {m}\ by (metis (no_types, lifting) Set.set_insert) then obtain b where b: "?a \ \ b" "b \ B" by (metis (no_types, lifting) ex_in_conv) then obtain \' where \': "b = ?a \'" using B(1) by blast moreover have "?a \' \ ?A {m}" using B(1) b(2) \' by auto hence "?a \ = ?a \'" by (blast dest!: A_m_eq) ultimately show False using b(1) by simp qed qed auto thus ?thesis by (metis (no_types, lifting) finite.emptyI finite_insert) next case False hence "?A {m} = {}" using \m = Fun g U\ by blast thus ?thesis by (metis finite.emptyI) qed ultimately show ?thesis using insert.IH by auto qed simp qed fastforce end subsection \Lemma: The Lazy Intruder Preserves Well-formedness\ context begin private lemma LI_preserves_subst_wf_single: assumes "(S\<^sub>1,\\<^sub>1) \ (S\<^sub>2,\\<^sub>2)" "fv\<^sub>s\<^sub>t S\<^sub>1 \ bvars\<^sub>s\<^sub>t S\<^sub>1 = {}" "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>1" and "subst_domain \\<^sub>1 \ vars\<^sub>s\<^sub>t S\<^sub>1 = {}" "range_vars \\<^sub>1 \ bvars\<^sub>s\<^sub>t S\<^sub>1 = {}" shows "fv\<^sub>s\<^sub>t S\<^sub>2 \ bvars\<^sub>s\<^sub>t S\<^sub>2 = {}" "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>2" and "subst_domain \\<^sub>2 \ vars\<^sub>s\<^sub>t S\<^sub>2 = {}" "range_vars \\<^sub>2 \ bvars\<^sub>s\<^sub>t S\<^sub>2 = {}" using assms proof (induction rule: LI_rel.induct) case (Compose S X f S' \) { case 1 thus ?case using vars_st_snd_map by auto } { case 2 thus ?case using vars_st_snd_map by auto } { case 3 thus ?case using vars_st_snd_map by force } { case 4 thus ?case using vars_st_snd_map by auto } next case (Unify S f U \ T S' \) hence "fv (Fun f U) \ fv\<^sub>s\<^sub>t S" using fv_subset_if_in_strand_ik' by blast hence *: "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using mgu_vars_bounded[OF Unify.hyps(3)[symmetric]] unfolding range_vars_alt_def by (fastforce simp del: subst_range.simps) have "fv\<^sub>s\<^sub>t (S@S') \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" "vars\<^sub>s\<^sub>t (S@S') \ vars\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" by auto hence **: "fv\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" "vars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) \ vars\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using subst_sends_strand_fv_to_img[of "S@S'" \] strand_subst_vars_union_bound[of "S@S'" \] * by blast+ have "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" by (fact mgu_gives_wellformed_subst[OF Unify.hyps(3)[symmetric]]) { case 1 have "bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = bvars\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using bvars_subst_ident[of "S@S'" \] by auto thus ?case using 1 ** by blast } { case 2 hence "subst_domain \ \ subst_domain \ = {}" "subst_domain \ \ range_vars \ = {}" using * by blast+ thus ?case by (metis wf_subst_compose[OF \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\]) } { case 3 hence "subst_domain \ \ vars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using ** by blast moreover have "v \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" when "v \ subst_domain \" for v using * that by blast hence "subst_domain \ \ fv\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using mgu_eliminates_dom[OF Unify.hyps(3)[symmetric], THEN strand_fv_subst_subset_if_subst_elim, of _ "S@Send [Fun f T]#S'"] unfolding subst_elim_def by auto moreover have "bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = bvars\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using bvars_subst_ident[of "S@S'" \] by auto hence "subst_domain \ \ bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using 3(1) * by blast ultimately show ?case using ** * subst_domain_compose[of \ \] vars\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>t[of "S@S' \\<^sub>s\<^sub>t \"] by blast } { case 4 have ***: "bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = bvars\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using bvars_subst_ident[of "S@S'" \] by auto hence "range_vars \ \ bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using 4(1) * by blast thus ?case using subst_img_comp_subset[of \ \] 4(4) *** by blast } next case (Equality S \ t t' a S' \) hence *: "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using mgu_vars_bounded[OF Equality.hyps(2)[symmetric]] unfolding range_vars_alt_def by fastforce have "fv\<^sub>s\<^sub>t (S@S') \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" "vars\<^sub>s\<^sub>t (S@S') \ vars\<^sub>s\<^sub>t (S@Equality a t t'#S')" by auto hence **: "fv\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" "vars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) \ vars\<^sub>s\<^sub>t (S@Equality a t t'#S')" using subst_sends_strand_fv_to_img[of "S@S'" \] strand_subst_vars_union_bound[of "S@S'" \] * by blast+ have "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" by (fact mgu_gives_wellformed_subst[OF Equality.hyps(2)[symmetric]]) { case 1 have "bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = bvars\<^sub>s\<^sub>t (S@Equality a t t'#S')" using bvars_subst_ident[of "S@S'" \] by auto thus ?case using 1 ** by blast } { case 2 hence "subst_domain \ \ subst_domain \ = {}" "subst_domain \ \ range_vars \ = {}" using * by blast+ thus ?case by (metis wf_subst_compose[OF \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\]) } { case 3 hence "subst_domain \ \ vars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using ** by blast moreover have "v \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" when "v \ subst_domain \" for v using * that by blast hence "subst_domain \ \ fv\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using mgu_eliminates_dom[OF Equality.hyps(2)[symmetric], THEN strand_fv_subst_subset_if_subst_elim, of _ "S@Equality a t t'#S'"] unfolding subst_elim_def by auto moreover have "bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = bvars\<^sub>s\<^sub>t (S@Equality a t t'#S')" using bvars_subst_ident[of "S@S'" \] by auto hence "subst_domain \ \ bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using 3(1) * by blast ultimately show ?case using ** * subst_domain_compose[of \ \] vars\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>t[of "S@S' \\<^sub>s\<^sub>t \"] by blast } { case 4 have ***: "bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = bvars\<^sub>s\<^sub>t (S@Equality a t t'#S')" using bvars_subst_ident[of "S@S'" \] by auto hence "range_vars \ \ bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using 4(1) * by blast thus ?case using subst_img_comp_subset[of \ \] 4(4) *** by blast } qed private lemma LI_preserves_subst_wf: assumes "(S\<^sub>1,\\<^sub>1) \\<^sup>* (S\<^sub>2,\\<^sub>2)" "fv\<^sub>s\<^sub>t S\<^sub>1 \ bvars\<^sub>s\<^sub>t S\<^sub>1 = {}" "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>1" and "subst_domain \\<^sub>1 \ vars\<^sub>s\<^sub>t S\<^sub>1 = {}" "range_vars \\<^sub>1 \ bvars\<^sub>s\<^sub>t S\<^sub>1 = {}" shows "fv\<^sub>s\<^sub>t S\<^sub>2 \ bvars\<^sub>s\<^sub>t S\<^sub>2 = {}" "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>2" and "subst_domain \\<^sub>2 \ vars\<^sub>s\<^sub>t S\<^sub>2 = {}" "range_vars \\<^sub>2 \ bvars\<^sub>s\<^sub>t S\<^sub>2 = {}" using assms proof (induction S\<^sub>2 \\<^sub>2 rule: rtrancl_induct2) case (step S\<^sub>i \\<^sub>i S\<^sub>j \\<^sub>j) { case 1 thus ?case using LI_preserves_subst_wf_single[OF \(S\<^sub>i,\\<^sub>i) \ (S\<^sub>j,\\<^sub>j)\] step.IH by metis } { case 2 thus ?case using LI_preserves_subst_wf_single[OF \(S\<^sub>i,\\<^sub>i) \ (S\<^sub>j,\\<^sub>j)\] step.IH by metis } { case 3 thus ?case using LI_preserves_subst_wf_single[OF \(S\<^sub>i,\\<^sub>i) \ (S\<^sub>j,\\<^sub>j)\] step.IH by metis } { case 4 thus ?case using LI_preserves_subst_wf_single[OF \(S\<^sub>i,\\<^sub>i) \ (S\<^sub>j,\\<^sub>j)\] step.IH by metis } qed metis lemma LI_preserves_wellformedness: assumes "(S\<^sub>1,\\<^sub>1) \\<^sup>* (S\<^sub>2,\\<^sub>2)" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1" shows "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>2 \\<^sub>2" proof - have *: "wf\<^sub>s\<^sub>t {} S\<^sub>j" when "(S\<^sub>i, \\<^sub>i) \ (S\<^sub>j, \\<^sub>j)" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>i \\<^sub>i" for S\<^sub>i \\<^sub>i S\<^sub>j \\<^sub>j using that proof (induction rule: LI_rel.induct) case (Compose S T f S' \) thus ?case by (metis wf_send_compose wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) next case (Unify S f U \ T S' \) have "fv (Fun f T) \ fv (Fun f U) \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using Unify.hyps(2) by force hence "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using mgu_vars_bounded[OF Unify.hyps(3)[symmetric]] by (metis subset_trans) hence "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@Send [Fun f T]#S') = {}" using Unify.prems unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by blast thus ?case using wf_unify[OF _ Unify.hyps(2) MGU_is_Unifier[OF mgu_gives_MGU], of "{}", OF _ Unify.hyps(3)[symmetric], of S'] Unify.prems(1) by (auto simp add: wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) next case (Equality S \ t t' a S' \) have "fv t \ fv t' \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using Equality.hyps(2) by force hence "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using mgu_vars_bounded[OF Equality.hyps(2)[symmetric]] by (metis subset_trans) hence "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@Equality a t t'#S') = {}" using Equality.prems unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by blast thus ?case using wf_equality[OF _ Equality.hyps(2)[symmetric], of "{}" S a S'] Equality.prems(1) by (auto simp add: wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) qed show ?thesis using assms proof (induction rule: rtrancl_induct2) case (step S\<^sub>i \\<^sub>i S\<^sub>j \\<^sub>j) thus ?case using LI_preserves_subst_wf_single[OF \(S\<^sub>i,\\<^sub>i) \ (S\<^sub>j,\\<^sub>j)\] *[OF \(S\<^sub>i,\\<^sub>i) \ (S\<^sub>j,\\<^sub>j)\] by (metis wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) qed simp qed lemma LI_preserves_trm_wf: assumes "(S,\) \\<^sup>* (S',\')" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S')" proof - { fix S \ S' \' assume "(S,\) \ (S',\')" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S')" proof (induction rule: LI_rel.induct) case (Compose S T f S' \) hence "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" and *: "t \ set S \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p t)" "t \ set S' \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p t)" for t by auto hence "wf\<^sub>t\<^sub>r\<^sub>m t" when "t \ set T" for t using that unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p t)" when "t \ set (map Send1 T)" for t using that unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto thus ?case using * by force next case (Unify S f U \ T S' \) have "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" "wf\<^sub>t\<^sub>r\<^sub>m (Fun f U)" using Unify.prems(1) Unify.hyps(2) wf_trm_subterm[of _ "Fun f U"] by (simp, force) hence range_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_wf_trm[OF Unify.hyps(3)[symmetric]] by simp { fix s assume "s \ set (S@S' \\<^sub>s\<^sub>t \)" hence "\s' \ set (S@S'). s = s' \\<^sub>s\<^sub>t\<^sub>p \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')" using Unify.prems(1) by (auto simp add: subst_apply_strand_def) moreover { fix s' assume s': "s = s' \\<^sub>s\<^sub>t\<^sub>p \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')" "s' \ set (S@S')" from s'(2) have "trms\<^sub>s\<^sub>t\<^sub>p (s' \\<^sub>s\<^sub>t\<^sub>p \) = trms\<^sub>s\<^sub>t\<^sub>p s' \\<^sub>s\<^sub>e\<^sub>t (rm_vars (set (bvars\<^sub>s\<^sub>t\<^sub>p s')) \)" proof (induction s') case (Inequality X F) thus ?case by (induct F) (auto simp add: subst_apply_pairs_def) qed auto hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s)" using wf_trm_subst[OF wf_trms_subst_rm_vars'[OF range_wf]] \wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')\ s'(1) by simp } ultimately have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s)" by auto } thus ?case by auto next case (Equality S \ t t' a S' \) hence "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m t'" by simp_all hence range_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_wf_trm[OF Equality.hyps(2)[symmetric]] by simp { fix s assume "s \ set (S@S' \\<^sub>s\<^sub>t \)" hence "\s' \ set (S@S'). s = s' \\<^sub>s\<^sub>t\<^sub>p \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')" using Equality.prems(1) by (auto simp add: subst_apply_strand_def) moreover { fix s' assume s': "s = s' \\<^sub>s\<^sub>t\<^sub>p \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')" "s' \ set (S@S')" from s'(2) have "trms\<^sub>s\<^sub>t\<^sub>p (s' \\<^sub>s\<^sub>t\<^sub>p \) = trms\<^sub>s\<^sub>t\<^sub>p s' \\<^sub>s\<^sub>e\<^sub>t (rm_vars (set (bvars\<^sub>s\<^sub>t\<^sub>p s')) \)" proof (induction s') case (Inequality X F) thus ?case by (induct F) (auto simp add: subst_apply_pairs_def) qed auto hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s)" using wf_trm_subst[OF wf_trms_subst_rm_vars'[OF range_wf]] \wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')\ s'(1) by simp } ultimately have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s)" by auto } thus ?case by auto qed } with assms show ?thesis by (induction rule: rtrancl_induct2) metis+ qed lemma LI_preproc_prop_subst: "LI_preproc_prop S \ LI_preproc_prop (S \\<^sub>s\<^sub>t \)" proof (induction S) case (Cons x S) thus ?case unfolding LI_preproc_prop_def by (cases x) auto qed (simp add: LI_preproc_prop_def) lemma LI_preserves_LI_preproc_prop: assumes "(S\<^sub>1,\\<^sub>1) \\<^sup>* (S\<^sub>2,\\<^sub>2)" "LI_preproc_prop S\<^sub>1" shows "LI_preproc_prop S\<^sub>2" using assms proof (induction rule: rtrancl_induct2) case (step S\<^sub>i \\<^sub>i S\<^sub>j \\<^sub>j) hence "LI_preproc_prop S\<^sub>i" by metis with step.hyps(2) show ?case proof (induction rule: LI_rel.induct) case (Unify S f T' \ T S' \) thus ?case using LI_preproc_prop_subst LI_preproc_prop_split by (metis append.left_neutral append_Cons) next case (Equality S \ t t' uu S' \) thus ?case using LI_preproc_prop_subst LI_preproc_prop_split by (metis append.left_neutral append_Cons) qed (auto simp add: LI_preproc_prop_def) qed simp end subsection \Theorem: Soundness of the Lazy Intruder\ context begin private lemma LI_soundness_single: assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1" "(S\<^sub>1,\\<^sub>1) \ (S\<^sub>2,\\<^sub>2)" "\ \\<^sub>c \S\<^sub>2,\\<^sub>2\" shows "\ \\<^sub>c \S\<^sub>1,\\<^sub>1\" using assms(2,1,3) proof (induction rule: LI_rel.induct) case (Compose S T f S' \) have "ik\<^sub>s\<^sub>t (map Send1 T) \\<^sub>s\<^sub>e\<^sub>t \ = {}" by fastforce hence *: "\{}; S\\<^sub>c \" "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; map Send1 T\\<^sub>c \" "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \" using Compose unfolding constr_sem_c_def by (force, force, fastforce) have "ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c Fun f T \ \" using *(2) Compose.hyps(2) ComposeC[OF _ Compose.hyps(3), of "map (\x. x \ \) T"] unfolding subst_compose_def by force thus "\ \\<^sub>c \S@Send [Fun f T]#S',\\" using *(1,3) \\ \\<^sub>c \S@map Send1 T@S',\\\ by (auto simp add: constr_sem_c_def) next case (Unify S f U \ T S' \) have "(\ \\<^sub>s \) supports \" "\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \" using Unify.prems(2) unfolding constr_sem_c_def by metis+ then obtain \ where \: "\ \\<^sub>s \ \\<^sub>s \ = \" unfolding subst_compose_def by auto have \fun_id: "Fun f U \ \ = Fun f U" "Fun f T \ \ = Fun f T" - using Unify.prems(1) trm_subst_ident[of "Fun f U" \] + using Unify.prems(1) subst_apply_term_ident[of "Fun f U" \] fv_subset_if_in_strand_ik[of "Fun f U" S] Unify.hyps(2) fv_snd_rcv_strand_subset(2)[of S] strand_vars_split(1)[of S "Send [Fun f T]#S'"] unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def apply blast - using Unify.prems(1) trm_subst_ident[of "Fun f T" \] + using Unify.prems(1) subst_apply_term_ident[of "Fun f T" \] unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by fastforce hence \\_disj: "subst_domain \ \ subst_domain \ = {}" "subst_domain \ \ range_vars \ = {}" "subst_domain \ \ range_vars \ = {}" using trm_subst_disj mgu_vars_bounded[OF Unify.hyps(3)[symmetric]] apply (blast,blast) using Unify.prems(1) unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by blast hence \\_support: "\ supports \" "\ supports \" by (simp_all add: subst_support_comp_split[OF \(\ \\<^sub>s \) supports \\]) have "fv (Fun f T) \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" "fv (Fun f U) \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using Unify.hyps(2) by force+ hence \_vars_bound: "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t (S@Send [Fun f T]#S')" using mgu_vars_bounded[OF Unify.hyps(3)[symmetric]] by blast have "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; [Send [Fun f T]]\\<^sub>c \" proof - from Unify.hyps(2) have "Fun f U \ \ \ ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \" by blast hence "Fun f U \ \ \ ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \" by blast moreover have "Unifier \ (Fun f T) (Fun f U)" by (fact MGU_is_Unifier[OF mgu_gives_MGU[OF Unify.hyps(3)[symmetric]]]) ultimately have "Fun f T \ \ \ ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \" using \ by (metis \fun_id subst_subst_compose) thus ?thesis by simp qed have "\{}; S\\<^sub>c \" "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \" proof - have "(S@S' \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \ = S@S' \\<^sub>s\<^sub>t \" "(S@S') \\<^sub>s\<^sub>t \ = S@S'" proof - have "subst_domain \ \ vars\<^sub>s\<^sub>t (S@S') = {}" using Unify.prems(1) by (auto simp add: wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) hence "subst_domain \ \ vars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using \\_disj(2) strand_subst_vars_union_bound[of "S@S'" \] by blast thus "(S@S' \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \ = S@S' \\<^sub>s\<^sub>t \" "(S@S') \\<^sub>s\<^sub>t \ = S@S'" using strand_subst_comp \subst_domain \ \ vars\<^sub>s\<^sub>t (S@S') = {}\ by (blast,blast) qed moreover have "subst_idem \" by (fact mgu_gives_subst_idem[OF Unify.hyps(3)[symmetric]]) moreover have "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@S') = {}" "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@S') = {}" using wf_constr_bvars_disj[OF Unify.prems(1)] wf_constr_bvars_disj'[OF Unify.prems(1) \_vars_bound] by auto ultimately have "\{}; S@S'\\<^sub>c \" using \\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \\ \ strand_sem_subst(1)[of \ "S@S' \\<^sub>s\<^sub>t \" "{}" "\ \\<^sub>s \"] strand_sem_subst(2)[of \ "S@S'" "{}" "\ \\<^sub>s \"] strand_sem_subst_subst_idem[of \ "S@S'" "{}" \] unfolding constr_sem_c_def by (metis subst_compose_assoc) thus "\{}; S\\<^sub>c \" "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \" by auto qed show "\ \\<^sub>c \S@Send [Fun f T]#S',\\" using \\_support(1) \\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; [Send [Fun f T]]\\<^sub>c \\ \\{}; S\\<^sub>c \\ \\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \\ by (auto simp add: constr_sem_c_def) next case (Equality S \ t t' a S' \) have "(\ \\<^sub>s \) supports \" "\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \" using Equality.prems(2) unfolding constr_sem_c_def by metis+ then obtain \ where \: "\ \\<^sub>s \ \\<^sub>s \ = \" unfolding subst_compose_def by auto have "fv t \ vars\<^sub>s\<^sub>t (S@Equality a t t'#S')" "fv t' \ vars\<^sub>s\<^sub>t (S@Equality a t t'#S')" by auto moreover have "subst_domain \ \ vars\<^sub>s\<^sub>t (S@Equality a t t'#S') = {}" using Equality.prems(1) unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by auto - ultimately have \fun_id: "t \ \ = t" "t' \ \ = t'" - using trm_subst_ident[of t \] trm_subst_ident[of t' \] - by auto + ultimately have \fun_id: "t \ \ = t" "t' \ \ = t'" by auto hence \\_disj: "subst_domain \ \ subst_domain \ = {}" "subst_domain \ \ range_vars \ = {}" "subst_domain \ \ range_vars \ = {}" using trm_subst_disj mgu_vars_bounded[OF Equality.hyps(2)[symmetric]] apply (blast,blast) using Equality.prems(1) unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by blast hence \\_support: "\ supports \" "\ supports \" by (simp_all add: subst_support_comp_split[OF \(\ \\<^sub>s \) supports \\]) have "fv t \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" "fv t' \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" by auto hence \_vars_bound: "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using mgu_vars_bounded[OF Equality.hyps(2)[symmetric]] by blast have "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; [Equality a t t']\\<^sub>c \" proof - have "t \ \ = t' \ \" using MGU_is_Unifier[OF mgu_gives_MGU[OF Equality.hyps(2)[symmetric]]] by metis hence "t \ (\ \\<^sub>s \) = t' \ (\ \\<^sub>s \)" by (metis \fun_id subst_subst_compose) hence "t \ \ = t' \ \" by (metis \ subst_subst_compose) thus ?thesis by simp qed have "\{}; S\\<^sub>c \" "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \" proof - have "(S@S' \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \ = S@S' \\<^sub>s\<^sub>t \" "(S@S') \\<^sub>s\<^sub>t \ = S@S'" proof - have "subst_domain \ \ vars\<^sub>s\<^sub>t (S@S') = {}" using Equality.prems(1) by (fastforce simp add: wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def simp del: subst_range.simps) hence "subst_domain \ \ fv\<^sub>s\<^sub>t (S@S') = {}" by blast hence "subst_domain \ \ fv\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using \\_disj(2) subst_sends_strand_fv_to_img[of "S@S'" \] by blast thus "(S@S' \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \ = S@S' \\<^sub>s\<^sub>t \" "(S@S') \\<^sub>s\<^sub>t \ = S@S'" using strand_subst_comp \subst_domain \ \ vars\<^sub>s\<^sub>t (S@S') = {}\ by (blast,blast) qed moreover have "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@S') = {}" "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@S') = {}" using wf_constr_bvars_disj[OF Equality.prems(1)] wf_constr_bvars_disj'[OF Equality.prems(1) \_vars_bound] by auto ultimately have "\{}; S@S'\\<^sub>c \" using \\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \\ \ strand_sem_subst(1)[of \ "S@S' \\<^sub>s\<^sub>t \" "{}" "\ \\<^sub>s \"] strand_sem_subst(2)[of \ "S@S'" "{}" "\ \\<^sub>s \"] strand_sem_subst_subst_idem[of \ "S@S'" "{}" \] mgu_gives_subst_idem[OF Equality.hyps(2)[symmetric]] unfolding constr_sem_c_def by (metis subst_compose_assoc) thus "\{}; S\\<^sub>c \" "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \" by auto qed show "\ \\<^sub>c \S@Equality a t t'#S',\\" using \\_support(1) \\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; [Equality a t t']\\<^sub>c \\ \\{}; S\\<^sub>c \\ \\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \\ by (auto simp add: constr_sem_c_def) qed theorem LI_soundness: assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1" "(LI_preproc S\<^sub>1,\\<^sub>1) \\<^sup>* (S\<^sub>2,\\<^sub>2)" "\ \\<^sub>c \S\<^sub>2, \\<^sub>2\" shows "\ \\<^sub>c \S\<^sub>1, \\<^sub>1\" using assms(2,1,3) proof (induction S\<^sub>2 \\<^sub>2 rule: rtrancl_induct2) case (step S\<^sub>i \\<^sub>i S\<^sub>j \\<^sub>j) thus ?case using LI_preproc_preserves_wellformedness[OF \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1\] LI_preserves_wellformedness[OF \(LI_preproc S\<^sub>1, \\<^sub>1) \\<^sup>* (S\<^sub>i, \\<^sub>i)\] LI_soundness_single[OF _ \(S\<^sub>i, \\<^sub>i) \ (S\<^sub>j, \\<^sub>j)\ \\ \\<^sub>c \S\<^sub>j, \\<^sub>j\\] by metis qed (metis LI_preproc_sem_eq') end subsection \Theorem: Completeness of the Lazy Intruder\ context begin private lemma LI_completeness_single: assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1" "\ \\<^sub>c \S\<^sub>1, \\<^sub>1\" "\simple S\<^sub>1" "LI_preproc_prop S\<^sub>1" shows "\S\<^sub>2 \\<^sub>2. (S\<^sub>1,\\<^sub>1) \ (S\<^sub>2,\\<^sub>2) \ (\ \\<^sub>c \S\<^sub>2, \\<^sub>2\)" using not_simple_elim[OF \\simple S\<^sub>1\] proof - { \ \In this case \S\<^sub>1\ isn't simple because it contains an equality constraint, so we can simply proceed with the reduction by computing the MGU for the equation\ assume "\S' S'' a t t'. S\<^sub>1 = S'@Equality a t t'#S'' \ simple S'" then obtain S a t t' S' where S\<^sub>1: "S\<^sub>1 = S@Equality a t t'#S'" "simple S" by moura hence *: "wf\<^sub>s\<^sub>t {} S" "\ \\<^sub>c \S, \\<^sub>1\" "\\<^sub>1 supports \" "t \ \ = t' \ \" using \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1\ wf_eq_fv[of "{}" S t t' S'] fv_snd_rcv_strand_subset(5)[of S] by (auto simp add: constr_sem_c_def wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) from * have "Unifier \ t t'" by simp then obtain \ where \: "Some \ = mgu t t'" "subst_idem \" "subst_domain \ \ range_vars \ \ fv t \ fv t'" using mgu_always_unifies mgu_gives_subst_idem mgu_vars_bounded by metis+ have "\ \\<^sub>\ \" using mgu_gives_MGU[OF \(1)[symmetric]] by (metis \Unifier \ t t'\) hence "\ supports \" using subst_support_if_mgt_subst_idem[OF _ \(2)] by metis hence "(\\<^sub>1 \\<^sub>s \) supports \" using subst_support_comp \\\<^sub>1 supports \\ by metis have "\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \" proof - have "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t S\<^sub>1" using \(3) S\<^sub>1(1) by auto hence "\{}; S\<^sub>1 \\<^sub>s\<^sub>t \\\<^sub>c \" using \subst_idem \\ \\ \\<^sub>\ \\ \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ strand_sem_subst wf_constr_bvars_disj'(1)[OF assms(1)] unfolding subst_idem_def constr_sem_c_def by (metis (no_types) subst_compose_assoc) thus "\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \" using S\<^sub>1(1) by force qed moreover have "(S@Equality a t t'#S', \\<^sub>1) \ (S@S' \\<^sub>s\<^sub>t \, \\<^sub>1 \\<^sub>s \)" using LI_rel.Equality[OF \simple S\ \(1)] S\<^sub>1 by metis ultimately have ?thesis using S\<^sub>1(1) \(\\<^sub>1 \\<^sub>s \) supports \\ by (auto simp add: constr_sem_c_def) } moreover { \ \In this case \S\<^sub>1\ isn't simple because it contains a deduction constraint for a composed term, so we must look at how this composed term is derived under the interpretation \\\\ assume "\S' S'' ts. S\<^sub>1 = S'@Send ts#S'' \ (\x. ts = [Var x]) \ simple S'" hence "\S' S'' f T. S\<^sub>1 = S'@Send [Fun f T]#S'' \ simple S'" using LI_preproc_prop_SendE[OF \LI_preproc_prop S\<^sub>1\] by fastforce with assms obtain S f T S' where S\<^sub>1: "S\<^sub>1 = S@Send [Fun f T]#S'" "simple S" by moura hence "wf\<^sub>s\<^sub>t {} S" "\ \\<^sub>c \S, \\<^sub>1\" "\\<^sub>1 supports \" using \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1\ by (auto simp add: constr_sem_c_def wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) \ \Lemma for a common subcase\ have fun_sat: "\ \\<^sub>c \S@(map Send1 T)@S', \\<^sub>1\" when T: "\t. t \ set T \ ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" proof - have "\t. t \ set T \ \ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; [Send1 t]\\<^sub>c \" using T by simp hence "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; map Send1 T\\<^sub>c \" using \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ strand_sem_Send_map by blast moreover have "ik\<^sub>s\<^sub>t (S@[Send1 (Fun f T)]) \\<^sub>s\<^sub>e\<^sub>t \ = ik\<^sub>s\<^sub>t (S@(map Send1 T)) \\<^sub>s\<^sub>e\<^sub>t \" by auto hence "\ik\<^sub>s\<^sub>t (S@(map Send1 T)) \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \" using \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ unfolding S\<^sub>1(1) constr_sem_c_def by force ultimately show ?thesis using \\ \\<^sub>c \S, \\<^sub>1\\ strand_sem_append(1)[of "{}" S \ "map Send1 T"] strand_sem_append(1)[of "{}" "S@map Send1 T" \ S'] unfolding constr_sem_c_def by simp qed from S\<^sub>1 \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ have "ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c Fun f T \ \" by (auto simp add: constr_sem_c_def) hence ?thesis proof cases \ \Case 1: \\(f(T))\ has been derived using the \AxiomC\ rule.\ case AxiomC hence ex_t: "\t. t \ ik\<^sub>s\<^sub>t S \ Fun f T \ \ = t \ \" by auto show ?thesis proof (cases "\T'. Fun f T' \ ik\<^sub>s\<^sub>t S \ Fun f T \ \ \ Fun f T' \ \") \ \Case 1.1: \f(T)\ is equal to a variable in the intruder knowledge under \\\. Hence there must exists a deduction constraint in the simple prefix of the constraint in which this variable occurs/"is sent" for the first time. Since this variable itself cannot have been derived from the \AxiomC\ rule (because it must be equal under the interpretation to \f(T)\, which is by assumption not in the intruder knowledge under \\\) it must be the case that we can derive it using the \ComposeC\ rule. Hence we can apply the \Compose\ rule of the lazy intruder to \f(T)\.\ case True have "\v. Var v \ ik\<^sub>s\<^sub>t S \ Fun f T \ \ = \ v" proof - obtain t where "t \ ik\<^sub>s\<^sub>t S" "Fun f T \ \ = t \ \" using ex_t by moura thus ?thesis using \\T'. Fun f T' \ ik\<^sub>s\<^sub>t S \ Fun f T \ \ \ Fun f T' \ \\ by (cases t) auto qed hence "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. Fun f T \ \ = \ v" using vars_subset_if_in_strand_ik2[of _ S] by fastforce then obtain v S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where S: "S = S\<^sub>p\<^sub>r\<^sub>e@Send [Var v]#S\<^sub>s\<^sub>u\<^sub>f" "Fun f T \ \ = \ v" "\(\w \ wfrestrictedvars\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e. Fun f T \ \ = \ w)" using \wf\<^sub>s\<^sub>t {} S\ wf_simple_strand_first_Send_var_split[OF _ \simple S\, of "Fun f T" \] by auto hence "\w. Var w \ ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \ \ v \ Var w \ \" by force moreover have "\T'. Fun f T' \ ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \ Fun f T \ \ \ Fun f T' \ \" using \\T'. Fun f T' \ ik\<^sub>s\<^sub>t S \ Fun f T \ \ \ Fun f T' \ \\ S(1) by (meson contra_subsetD ik_append_subset(1)) hence "\g T'. Fun g T' \ ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \ \ v \ Fun g T' \ \" using S(2) by simp ultimately have "\t \ ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e. \ v \ t \ \" by (metis term.exhaust) hence "\ v \ (ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e) \\<^sub>s\<^sub>e\<^sub>t \" by auto have "ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c \ v" using S\<^sub>1(1) S(1) \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ by (auto simp add: constr_sem_c_def) hence "ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c Fun f T \ \" using \Fun f T \ \ = \ v\ by metis hence "length T = arity f" "public f" "\t. t \ set T \ ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" using \Fun f T \ \ = \ v\ \\ v \ ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t \\ intruder_synth.simps[of "ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t \" "\ v"] by auto hence *: "\t. t \ set T \ ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" using S(1) by (auto intro: ideduct_synth_mono) hence "\ \\<^sub>c \S@(map Send1 T)@S', \\<^sub>1\" by (metis fun_sat) moreover have "(S@Send [Fun f T]#S', \\<^sub>1) \ (S@map Send1 T@S', \\<^sub>1)" by (metis LI_rel.Compose[OF \simple S\ \length T = arity f\ \public f\]) ultimately show ?thesis using S\<^sub>1 by auto next \ \Case 1.2: \\(f(T))\ can be derived from an interpreted composed term in the intruder knowledge. Use the \Unify\ rule on this composed term to further reduce the constraint.\ case False then obtain T' where t: "Fun f T' \ ik\<^sub>s\<^sub>t S" "Fun f T \ \ = Fun f T' \ \" by auto hence "fv (Fun f T') \ fv\<^sub>s\<^sub>t S\<^sub>1" using S\<^sub>1(1) fv_subset_if_in_strand_ik'[OF t(1)] fv_snd_rcv_strand_subset(2)[of S] by auto from t have "Unifier \ (Fun f T) (Fun f T')" by simp then obtain \ where \: "Some \ = mgu (Fun f T) (Fun f T')" "subst_idem \" "subst_domain \ \ range_vars \ \ fv (Fun f T) \ fv (Fun f T')" using mgu_always_unifies mgu_gives_subst_idem mgu_vars_bounded by metis+ have "\ \\<^sub>\ \" using mgu_gives_MGU[OF \(1)[symmetric]] by (metis \Unifier \ (Fun f T) (Fun f T')\) hence "\ supports \" using subst_support_if_mgt_subst_idem[OF _ \(2)] by metis hence "(\\<^sub>1 \\<^sub>s \) supports \" using subst_support_comp \\\<^sub>1 supports \\ by metis have "\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \" proof - have "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t S\<^sub>1" using \(3) S\<^sub>1(1) \fv (Fun f T') \ fv\<^sub>s\<^sub>t S\<^sub>1\ unfolding range_vars_alt_def by (fastforce simp del: subst_range.simps) hence "\{}; S\<^sub>1 \\<^sub>s\<^sub>t \\\<^sub>c \" using \subst_idem \\ \\ \\<^sub>\ \\ \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ strand_sem_subst wf_constr_bvars_disj'(1)[OF assms(1)] unfolding subst_idem_def constr_sem_c_def by (metis (no_types) subst_compose_assoc) thus "\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \" using S\<^sub>1(1) by force qed moreover have "(S@Send [Fun f T]#S', \\<^sub>1) \ (S@S' \\<^sub>s\<^sub>t \, \\<^sub>1 \\<^sub>s \)" using LI_rel.Unify[OF \simple S\ t(1) \(1)] S\<^sub>1 by metis ultimately show ?thesis using S\<^sub>1(1) \(\\<^sub>1 \\<^sub>s \) supports \\ by (auto simp add: constr_sem_c_def) qed next \ \Case 2: \\(f(T))\ has been derived using the \ComposeC\ rule. Simply use the \Compose\ rule of the lazy intruder to proceed with the reduction.\ case (ComposeC T' g) hence "f = g" "length T = arity f" "public f" and "\x. x \ set T \ ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c x \ \" by auto hence "\ \\<^sub>c \S@(map Send1 T)@S', \\<^sub>1\" using fun_sat by metis moreover have "(S\<^sub>1, \\<^sub>1) \ (S@(map Send1 T)@S', \\<^sub>1)" using S\<^sub>1 LI_rel.Compose[OF \simple S\ \length T = arity f\ \public f\] by metis ultimately show ?thesis by metis qed } moreover have "\A B X F. S\<^sub>1 = A@Inequality X F#B \ ineq_model \ X F" using assms(2) by (auto simp add: constr_sem_c_def) ultimately show ?thesis using not_simple_elim[OF \\simple S\<^sub>1\] by metis qed theorem LI_completeness: assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1" "\ \\<^sub>c \S\<^sub>1, \\<^sub>1\" shows "\S\<^sub>2 \\<^sub>2. (LI_preproc S\<^sub>1,\\<^sub>1) \\<^sup>* (S\<^sub>2,\\<^sub>2) \ simple S\<^sub>2 \ (\ \\<^sub>c \S\<^sub>2, \\<^sub>2\)" proof (cases "simple (LI_preproc S\<^sub>1)") case False let ?Stuck = "\S\<^sub>2 \\<^sub>2. \(\S\<^sub>3 \\<^sub>3. (S\<^sub>2,\\<^sub>2) \ (S\<^sub>3,\\<^sub>3) \ (\ \\<^sub>c \S\<^sub>3, \\<^sub>3\))" let ?Sats = "{((S,\),(S',\')). (S,\) \ (S',\') \ (\ \\<^sub>c \S, \\) \ (\ \\<^sub>c \S', \'\)}" have simple_if_stuck: "\S\<^sub>2 \\<^sub>2. \(LI_preproc S\<^sub>1,\\<^sub>1) \\<^sup>+ (S\<^sub>2,\\<^sub>2); \ \\<^sub>c \S\<^sub>2, \\<^sub>2\; ?Stuck S\<^sub>2 \\<^sub>2\ \ simple S\<^sub>2" using LI_preproc_preserves_wellformedness[OF \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1\] LI_preserves_LI_preproc_prop[OF _ LI_preproc_preproc_prop] LI_completeness_single[OF LI_preserves_wellformedness] trancl_into_rtrancl by metis have base: "\b. ((LI_preproc S\<^sub>1,\\<^sub>1),b) \ ?Sats" using LI_preproc_preserves_wellformedness[OF \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1\] LI_completeness_single[OF _ _ False LI_preproc_preproc_prop] LI_preproc_sem_eq' \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ by auto have *: "\S \ S' \'. ((S,\),(S',\')) \ ?Sats\<^sup>+ \ (S,\) \\<^sup>+ (S',\') \ (\ \\<^sub>c \S', \'\)" proof - fix S \ S' \' assume "((S,\),(S',\')) \ ?Sats\<^sup>+" thus "(S,\) \\<^sup>+ (S',\') \ (\ \\<^sub>c \S', \'\)" by (induct rule: trancl_induct2) auto qed have "\S\<^sub>2 \\<^sub>2. ((LI_preproc S\<^sub>1,\\<^sub>1),(S\<^sub>2,\\<^sub>2)) \ ?Sats\<^sup>+ \ ?Stuck S\<^sub>2 \\<^sub>2" proof (rule ccontr) assume "\(\S\<^sub>2 \\<^sub>2. ((LI_preproc S\<^sub>1,\\<^sub>1),(S\<^sub>2,\\<^sub>2)) \ ?Sats\<^sup>+ \ ?Stuck S\<^sub>2 \\<^sub>2)" hence sat_not_stuck: "\S\<^sub>2 \\<^sub>2. ((LI_preproc S\<^sub>1,\\<^sub>1),(S\<^sub>2,\\<^sub>2)) \ ?Sats\<^sup>+ \ \?Stuck S\<^sub>2 \\<^sub>2" by blast have "\S \. ((LI_preproc S\<^sub>1,\\<^sub>1),(S,\)) \ ?Sats\<^sup>+ \ (\b. ((S,\),b) \ ?Sats)" proof (intro allI impI) fix S \ assume a: "((LI_preproc S\<^sub>1,\\<^sub>1),(S,\)) \ ?Sats\<^sup>+" have "\b. ((LI_preproc S\<^sub>1,\\<^sub>1),b) \ ?Sats\<^sup>+ \ \c. b \ c \ ((LI_preproc S\<^sub>1,\\<^sub>1),c) \ ?Sats\<^sup>+" proof - fix b assume in_sat: "((LI_preproc S\<^sub>1,\\<^sub>1),b) \ ?Sats\<^sup>+" hence "\c. (b,c) \ ?Sats" using * sat_not_stuck by (cases b) blast thus "\c. b \ c \ ((LI_preproc S\<^sub>1,\\<^sub>1),c) \ ?Sats\<^sup>+" using trancl_into_trancl[OF in_sat] by blast qed hence "\S' \'. (S,\) \ (S',\') \ ((LI_preproc S\<^sub>1,\\<^sub>1),(S',\')) \ ?Sats\<^sup>+" using a by auto then obtain S' \' where S'\': "(S,\) \ (S',\')" "((LI_preproc S\<^sub>1,\\<^sub>1),(S',\')) \ ?Sats\<^sup>+" by auto hence "\ \\<^sub>c \S', \'\" using * by blast moreover have "(LI_preproc S\<^sub>1, \\<^sub>1) \\<^sup>+ (S,\)" using a trancl_mono by blast ultimately have "((S,\),(S',\')) \ ?Sats" using S'\'(1) * a by blast thus "\b. ((S,\),b) \ ?Sats" using S'\'(2) by blast qed hence "\f. \i::nat. (f i, f (Suc i)) \ ?Sats" using infinite_chain_intro'[OF base] by blast moreover have "?Sats \ LI_rel\<^sup>+" by auto hence "\(\f. \i::nat. (f i, f (Suc i)) \ ?Sats)" using LI_no_infinite_chain infinite_chain_mono by blast ultimately show False by auto qed hence "\S\<^sub>2 \\<^sub>2. (LI_preproc S\<^sub>1, \\<^sub>1) \\<^sup>+ (S\<^sub>2, \\<^sub>2) \ simple S\<^sub>2 \ (\ \\<^sub>c \S\<^sub>2, \\<^sub>2\)" using simple_if_stuck * by blast thus ?thesis by (meson trancl_into_rtrancl) qed (use LI_preproc_sem_eq' \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ in blast) end subsection \Corollary: Soundness and Completeness as a Single Theorem\ corollary LI_soundness_and_completeness: assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1" shows "\ \\<^sub>c \S\<^sub>1, \\<^sub>1\ \ (\S\<^sub>2 \\<^sub>2. (LI_preproc S\<^sub>1,\\<^sub>1) \\<^sup>* (S\<^sub>2,\\<^sub>2) \ simple S\<^sub>2 \ (\ \\<^sub>c \S\<^sub>2, \\<^sub>2\))" by (metis LI_soundness[OF assms] LI_completeness[OF assms]) end end diff --git a/thys/Stateful_Protocol_Composition_and_Typing/More_Unification.thy b/thys/Stateful_Protocol_Composition_and_Typing/More_Unification.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/More_Unification.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/More_Unification.thy @@ -1,3480 +1,3460 @@ (* Based on src/HOL/ex/Unification.thy packaged with Isabelle/HOL 2015 having the following license: ISABELLE COPYRIGHT NOTICE, LICENCE AND DISCLAIMER. Copyright (c) 1986-2015, University of Cambridge, Technische Universitaet Muenchen, and contributors. All rights reserved. Redistribution and use in source and binary forms, with or without modification, are permitted provided that the following conditions are met: * Redistributions of source code must retain the above copyright notice, this list of conditions and the following disclaimer. * Redistributions in binary form must reproduce the above copyright notice, this list of conditions and the following disclaimer in the documentation and/or other materials provided with the distribution. * Neither the name of the University of Cambridge or the Technische Universitaet Muenchen nor the names of their contributors may be used to endorse or promote products derived from this software without specific prior written permission. THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. *) (* Title: More_Unification.thy Author: Andreas Viktor Hess, DTU SPDX-License-Identifier: BSD-3-Clause Originally based on src/HOL/ex/Unification.thy (Isabelle/HOL 2015) by: Author: Martin Coen, Cambridge University Computer Laboratory Author: Konrad Slind, TUM & Cambridge University Computer Laboratory Author: Alexander Krauss, TUM *) section \Definitions and Properties Related to Substitutions and Unification\ theory More_Unification imports Messages "First_Order_Terms.Unification" begin subsection \Substitutions\ abbreviation subst_apply_list (infix "\\<^sub>l\<^sub>i\<^sub>s\<^sub>t" 51) where "T \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \ map (\t. t \ \) T" abbreviation subst_apply_pair (infixl "\\<^sub>p" 60) where "d \\<^sub>p \ \ (case d of (t,t') \ (t \ \, t' \ \))" abbreviation subst_apply_pair_set (infixl "\\<^sub>p\<^sub>s\<^sub>e\<^sub>t" 60) where "M \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \ (\d. d \\<^sub>p \) ` M" definition subst_apply_pairs (infix "\\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s" 51) where "F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \ \ map (\f. f \\<^sub>p \) F" abbreviation subst_more_general_than (infixl "\\<^sub>\" 50) where "\ \\<^sub>\ \ \ \\. \ = \ \\<^sub>s \" abbreviation subst_support (infix "supports" 50) where "\ supports \ \ (\x. \ x \ \ = \ x)" abbreviation rm_var where "rm_var v s \ s(v := Var v)" abbreviation rm_vars where "rm_vars vs \ \ (\v. if v \ vs then Var v else \ v)" definition subst_elim where "subst_elim \ v \ \t. v \ fv (t \ \)" definition subst_idem where "subst_idem s \ s \\<^sub>s s = s" lemma subst_support_def: "\ supports \ \ \ = \ \\<^sub>s \" unfolding subst_compose_def by metis lemma subst_supportD: "\ supports \ \ \ \\<^sub>\ \" using subst_support_def by auto lemma rm_vars_empty[simp]: "rm_vars {} s = s" "rm_vars (set []) s = s" by simp_all lemma rm_vars_singleton: "rm_vars {v} s = rm_var v s" by auto lemma subst_apply_terms_empty: "M \\<^sub>s\<^sub>e\<^sub>t Var = M" by simp lemma subst_agreement: "(t \ r = t \ s) \ (\v \ fv t. Var v \ r = Var v \ s)" by (induct t) auto lemma repl_invariance[dest?]: "v \ fv t \ t \ s(v := u) = t \ s" by (simp add: subst_agreement) lemma subst_idx_map: assumes "\i \ set I. i < length T" shows "(map ((!) T) I) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = map ((!) (map (\t. t \ \) T)) I" using assms by auto lemma subst_idx_map': assumes "\i \ fv\<^sub>s\<^sub>e\<^sub>t (set K). i < length T" shows "(K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t ((!) (map (\t. t \ \) T))" (is "?A = ?B") proof - have "T ! i \ \ = (map (\t. t \ \) T) ! i" when "i < length T" for i using that by auto hence "T ! i \ \ = (map (\t. t \ \) T) ! i" when "i \ fv\<^sub>s\<^sub>e\<^sub>t (set K)" for i using that assms by auto hence "k \ (!) T \ \ = k \ (!) (map (\t. t \ \) T)" when "fv k \ fv\<^sub>s\<^sub>e\<^sub>t (set K)" for k using that by (induction k) force+ thus ?thesis by auto qed lemma subst_remove_var: "v \ fv s \ v \ fv (t \ Var(v := s))" by (induct t) simp_all lemma subst_set_map: "x \ set X \ x \ s \ set (map (\x. x \ s) X)" by simp lemma subst_set_idx_map: assumes "\i \ I. i < length T" shows "(!) T ` I \\<^sub>s\<^sub>e\<^sub>t \ = (!) (map (\t. t \ \) T) ` I" (is "?A = ?B") proof have *: "T ! i \ \ = (map (\t. t \ \) T) ! i" when "i < length T" for i using that by auto show "?A \ ?B" using * assms by blast show "?B \ ?A" using * assms by auto qed lemma subst_set_idx_map': assumes "\i \ fv\<^sub>s\<^sub>e\<^sub>t K. i < length T" shows "K \\<^sub>s\<^sub>e\<^sub>t (!) T \\<^sub>s\<^sub>e\<^sub>t \ = K \\<^sub>s\<^sub>e\<^sub>t (!) (map (\t. t \ \) T)" (is "?A = ?B") proof have "T ! i \ \ = (map (\t. t \ \) T) ! i" when "i < length T" for i using that by auto hence "T ! i \ \ = (map (\t. t \ \) T) ! i" when "i \ fv\<^sub>s\<^sub>e\<^sub>t K" for i using that assms by auto hence *: "k \ (!) T \ \ = k \ (!) (map (\t. t \ \) T)" when "fv k \ fv\<^sub>s\<^sub>e\<^sub>t K" for k using that by (induction k) force+ show "?A \ ?B" using * by auto show "?B \ ?A" using * by force qed lemma subst_term_list_obtain: assumes "\i < length T. \s. P (T ! i) s \ S ! i = s \ \" and "length T = length S" shows "\U. length T = length U \ (\i < length T. P (T ! i) (U ! i)) \ S = map (\u. u \ \) U" using assms proof (induction T arbitrary: S) case (Cons t T S') then obtain s S where S': "S' = s#S" by (cases S') auto have "\i < length T. \s. P (T ! i) s \ S ! i = s \ \" "length T = length S" using Cons.prems S' by force+ then obtain U where U: "length T = length U" "\i < length T. P (T ! i) (U ! i)" "S = map (\u. u \ \) U" using Cons.IH by moura obtain u where u: "P t u" "s = u \ \" using Cons.prems(1) S' by auto have 1: "length (t#T) = length (u#U)" using Cons.prems(2) U(1) by fastforce have 2: "\i < length (t#T). P ((t#T) ! i) ((u#U) ! i)" using u(1) U(2) by (simp add: nth_Cons') have 3: "S' = map (\u. u \ \) (u#U)" using U u S' by simp show ?case using 1 2 3 by blast qed simp lemma subst_mono: "t \ u \ t \ s \ u \ s" by (induct u) auto lemma subst_mono_fv: "x \ fv t \ s x \ t \ s" by (induct t) auto lemma subst_mono_neq: assumes "t \ u" shows "t \ s \ u \ s" proof (cases u) case (Var v) hence False using \t \ u\ by simp thus ?thesis .. next case (Fun f X) then obtain x where "x \ set X" "t \ x" using \t \ u\ by auto hence "t \ s \ x \ s" using subst_mono by metis obtain Y where "Fun f X \ s = Fun f Y" by auto hence "x \ s \ set Y" using \x \ set X\ by auto hence "x \ s \ Fun f X \ s" using \Fun f X \ s = Fun f Y\ Fun_param_is_subterm by simp hence "t \ s \ Fun f X \ s" using \t \ s \ x \ s\ by (metis term.dual_order.trans term.order.eq_iff) thus ?thesis using \u = Fun f X\ \t \ u\ by metis qed lemma subst_no_occs[dest]: "\Var v \ t \ t \ Var(v := s) = t" by (induct t) (simp_all add: map_idI) lemma var_comp[simp]: "\ \\<^sub>s Var = \" "Var \\<^sub>s \ = \" unfolding subst_compose_def by simp_all lemma subst_comp_all: "M \\<^sub>s\<^sub>e\<^sub>t (\ \\<^sub>s \) = (M \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \" -using subst_subst_compose[of _ \ \] by auto +unfolding subst_subst_compose by auto lemma subst_all_mono: "M \ M' \ M \\<^sub>s\<^sub>e\<^sub>t s \ M' \\<^sub>s\<^sub>e\<^sub>t s" by auto lemma subst_comp_set_image: "(\ \\<^sub>s \) ` X = \ ` X \\<^sub>s\<^sub>e\<^sub>t \" using subst_compose by fastforce lemma subst_ground_ident[dest?]: "fv t = {} \ t \ s = t" by (induct t, simp, metis subst_agreement empty_iff subst_apply_term_empty) lemma subst_ground_ident_compose: "fv (\ x) = {} \ (\ \\<^sub>s \) x = \ x" "fv (t \ \) = {} \ t \ (\ \\<^sub>s \) = t \ \" -using subst_subst_compose[of t \ \] +unfolding subst_subst_compose by (simp_all add: subst_compose_def subst_ground_ident) lemma subst_all_ground_ident[dest?]: "ground M \ M \\<^sub>s\<^sub>e\<^sub>t s = M" proof - assume "ground M" hence "\t. t \ M \ fv t = {}" by auto hence "\t. t \ M \ t \ s = t" by (metis subst_ground_ident) moreover have "\t. t \ M \ t \ s \ M \\<^sub>s\<^sub>e\<^sub>t s" by (metis imageI) ultimately show "M \\<^sub>s\<^sub>e\<^sub>t s = M" by (simp add: image_cong) qed -lemma subst_eqI[intro]: "(\t. t \ \ = t \ \) \ \ = \" -proof - - assume "\t. t \ \ = t \ \" - hence "\v. Var v \ \ = Var v \ \" by auto - thus "\ = \" by auto -qed - lemma subst_cong: "\\ = \'; \ = \'\ \ (\ \\<^sub>s \) = (\' \\<^sub>s \')" by auto lemma subst_mgt_bot[simp]: "Var \\<^sub>\ \" by simp lemma subst_mgt_refl[simp]: "\ \\<^sub>\ \" by (metis var_comp(1)) lemma subst_mgt_trans: "\\ \\<^sub>\ \; \ \\<^sub>\ \\ \ \ \\<^sub>\ \" by (metis subst_compose_assoc) lemma subst_mgt_comp: "\ \\<^sub>\ \ \\<^sub>s \" by auto lemma subst_mgt_comp': "\ \\<^sub>s \ \\<^sub>\ \ \ \ \\<^sub>\ \" by (metis subst_compose_assoc) lemma var_self: "(\w. if w = v then Var v else Var w) = Var" using subst_agreement by auto lemma var_same[simp]: "Var(v := t) = Var \ t = Var v" by (intro iffI, metis fun_upd_same, simp add: var_self) lemma subst_eq_if_eq_vars: "(\v. (Var v) \ \ = (Var v) \ \) \ \ = \" by (auto simp add: subst_agreement) lemma subst_all_empty[simp]: "{} \\<^sub>s\<^sub>e\<^sub>t \ = {}" by simp lemma subst_all_insert:"(insert t M) \\<^sub>s\<^sub>e\<^sub>t \ = insert (t \ \) (M \\<^sub>s\<^sub>e\<^sub>t \)" by auto lemma subst_apply_fv_subset: "fv t \ V \ fv (t \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" by (induct t) auto lemma subst_apply_fv_empty: assumes "fv t = {}" shows "fv (t \ \) = {}" using assms subst_apply_fv_subset[of t "{}" \] by auto lemma subst_compose_fv: assumes "fv (\ x) = {}" shows "fv ((\ \\<^sub>s \) x) = {}" using assms subst_apply_fv_empty unfolding subst_compose_def by fast lemma subst_compose_fv': fixes \ \::"('a,'b) subst" assumes "y \ fv ((\ \\<^sub>s \) x)" shows "\z. z \ fv (\ x)" using assms subst_compose_fv by fast lemma subst_apply_fv_unfold: "fv (t \ \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv t)" by (induct t) auto lemma subst_apply_fv_unfold_set: "fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>s\<^sub>e\<^sub>t (set ts)) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" by (simp add: subst_apply_fv_unfold) lemma subst_apply_fv_unfold': "fv (t \ \) = (\v \ fv t. fv (\ v))" using subst_apply_fv_unfold by simp lemma subst_apply_fv_union: "fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv (t \ \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ fv t))" proof - have "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ fv t)) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv t)" by auto thus ?thesis using subst_apply_fv_unfold by metis qed lemma fv\<^sub>s\<^sub>e\<^sub>t_subst: "fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>s\<^sub>e\<^sub>t M)" -using subst_apply_fv_unfold[of _ \] by auto +by (simp add: subst_apply_fv_unfold) lemma subst_list_set_fv: "fv\<^sub>s\<^sub>e\<^sub>t (set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>s\<^sub>e\<^sub>t (set ts))" using subst_apply_fv_unfold_set[of \ ts] by simp lemma subst_elimI[intro]: "(\t. v \ fv (t \ \)) \ subst_elim \ v" by (auto simp add: subst_elim_def) lemma subst_elimI'[intro]: "(\w. v \ fv (Var w \ \)) \ subst_elim \ v" by (simp add: subst_elim_def subst_apply_fv_unfold') lemma subst_elimD[dest]: "subst_elim \ v \ v \ fv (t \ \)" by (auto simp add: subst_elim_def) lemma subst_elimD'[dest]: "subst_elim \ v \ \ v \ Var v" -by (metis subst_elim_def subst_apply_term.simps(1) term.set_intros(3)) +by (metis subst_elim_def eval_term.simps(1) term.set_intros(3)) lemma subst_elimD''[dest]: "subst_elim \ v \ v \ fv (\ w)" -by (metis subst_elim_def subst_apply_term.simps(1)) +by (metis subst_elim_def eval_term.simps(1)) lemma subst_elim_rm_vars_dest[dest]: "subst_elim (\::('a,'b) subst) v \ v \ vs \ subst_elim (rm_vars vs \) v" proof - assume assms: "subst_elim \ v" "v \ vs" obtain f::"('a, 'b) subst \ 'b \ 'b" where "\\ v. (\w. v \ fv (Var w \ \)) = (v \ fv (Var (f \ v) \ \))" by moura hence *: "\a \. a \ fv (Var (f \ a) \ \) \ subst_elim \ a" by blast have "Var (f (rm_vars vs \) v) \ \ \ Var (f (rm_vars vs \) v) \ rm_vars vs \ \ v \ fv (Var (f (rm_vars vs \) v) \ rm_vars vs \)" using assms(1) by fastforce moreover { assume "Var (f (rm_vars vs \) v) \ \ \ Var (f (rm_vars vs \) v) \ rm_vars vs \" hence "rm_vars vs \ (f (rm_vars vs \) v) \ \ (f (rm_vars vs \) v)" by auto hence "f (rm_vars vs \) v \ vs" by meson hence ?thesis using * assms(2) by force } ultimately show ?thesis using * by blast qed lemma occs_subst_elim: "\Var v \ t \ subst_elim (Var(v := t)) v \ (Var(v := t)) = Var" proof (cases "Var v = t") assume "Var v \ t" "\Var v \ t" hence "v \ fv t" by (simp add: vars_iff_subterm_or_eq) thus ?thesis by (auto simp add: subst_remove_var) qed auto lemma occs_subst_elim': "\Var v \ t \ subst_elim (Var(v := t)) v" proof - assume "\Var v \ t" hence "v \ fv t" by (auto simp add: vars_iff_subterm_or_eq) thus "subst_elim (Var(v := t)) v" by (simp add: subst_elim_def subst_remove_var) qed lemma subst_elim_comp: "subst_elim \ v \ subst_elim (\ \\<^sub>s \) v" by (auto simp add: subst_elim_def) lemma var_subst_idem: "subst_idem Var" by (simp add: subst_idem_def) lemma var_upd_subst_idem: assumes "\Var v \ t" shows "subst_idem (Var(v := t))" -unfolding subst_idem_def -proof - let ?\ = "Var(v := t)" - from assms have t_\_id: "t \ ?\ = t" by blast - fix s show "s \ (?\ \\<^sub>s ?\) = s \ ?\" - unfolding subst_compose_def - by (induction s, metis t_\_id fun_upd_def subst_apply_term.simps(1), simp) -qed + using subst_no_occs[OF assms] by (simp add: subst_idem_def subst_def[symmetric]) lemma zip_map_subst: "zip xs (xs \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) = map (\t. (t, t \ \)) xs" by (induction xs) auto lemma map2_map_subst: "map2 f xs (xs \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) = map (\t. f t (t \ \)) xs" by (induction xs) auto subsection \Lemmata: Domain and Range of Substitutions\ lemma range_vars_alt_def: "range_vars s \ fv\<^sub>s\<^sub>e\<^sub>t (subst_range s)" unfolding range_vars_def by simp lemma subst_dom_var_finite[simp]: "finite (subst_domain Var)" by simp lemma subst_range_Var[simp]: "subst_range Var = {}" by simp lemma range_vars_Var[simp]: "range_vars Var = {}" by fastforce lemma finite_subst_img_if_finite_dom: "finite (subst_domain \) \ finite (range_vars \)" unfolding range_vars_alt_def by auto lemma finite_subst_img_if_finite_dom': "finite (subst_domain \) \ finite (subst_range \)" by auto lemma subst_img_alt_def: "subst_range s = {t. \v. s v = t \ t \ Var v}" by (auto simp add: subst_domain_def) lemma subst_fv_img_alt_def: "range_vars s = (\t \ {t. \v. s v = t \ t \ Var v}. fv t)" unfolding range_vars_alt_def by (auto simp add: subst_domain_def) lemma subst_domI[intro]: "\ v \ Var v \ v \ subst_domain \" by (simp add: subst_domain_def) lemma subst_imgI[intro]: "\ v \ Var v \ \ v \ subst_range \" by (simp add: subst_domain_def) lemma subst_fv_imgI[intro]: "\ v \ Var v \ fv (\ v) \ range_vars \" unfolding range_vars_alt_def by auto lemma subst_eqI': assumes "t \ \ = t \ \" "subst_domain \ = subst_domain \" "subst_domain \ \ fv t" shows "\ = \" by (metis assms(2,3) term_subst_eq_rev[OF assms(1)] in_mono ext subst_domI) lemma subst_domain_subst_Fun_single[simp]: "subst_domain (Var(x := Fun f T)) = {x}" (is "?A = ?B") unfolding subst_domain_def by simp lemma subst_range_subst_Fun_single[simp]: "subst_range (Var(x := Fun f T)) = {Fun f T}" (is "?A = ?B") by simp lemma range_vars_subst_Fun_single[simp]: "range_vars (Var(x := Fun f T)) = fv (Fun f T)" unfolding range_vars_alt_def by force lemma var_renaming_is_Fun_iff: assumes "subst_range \ \ range Var" shows "is_Fun t = is_Fun (t \ \)" proof (cases t) case (Var x) hence "\y. \ x = Var y" using assms by auto thus ?thesis using Var by auto qed simp lemma subst_fv_dom_img_subset: "fv t \ subst_domain \ \ fv (t \ \) \ range_vars \" unfolding range_vars_alt_def by (induct t) auto lemma subst_fv_dom_img_subset_set: "fv\<^sub>s\<^sub>e\<^sub>t M \ subst_domain \ \ fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) \ range_vars \" proof - assume assms: "fv\<^sub>s\<^sub>e\<^sub>t M \ subst_domain \" obtain f::"'a set \ (('b, 'a) term \ 'a set) \ ('b, 'a) terms \ ('b, 'a) term" where "\x y z. (\v. v \ z \ \ y v \ x) \ (f x y z \ z \ \ y (f x y z) \ x)" by moura hence *: "\T g A. (\ \ (g ` T) \ A \ (\t. t \ T \ g t \ A)) \ (\ (g ` T) \ A \ f A g T \ T \ \ g (f A g T) \ A)" by (metis (no_types) SUP_le_iff) hence **: "\t. t \ M \ fv t \ subst_domain \" by (metis (no_types) assms fv\<^sub>s\<^sub>e\<^sub>t.simps) have "\t::('b, 'a) term. \f T. t \ f ` T \ (\t'::('b, 'a) term. t = f t' \ t' \ T)" by blast hence "f (range_vars \) fv (M \\<^sub>s\<^sub>e\<^sub>t \) \ M \\<^sub>s\<^sub>e\<^sub>t \ \ fv (f (range_vars \) fv (M \\<^sub>s\<^sub>e\<^sub>t \)) \ range_vars \" by (metis (full_types) ** subst_fv_dom_img_subset) thus ?thesis by (metis (no_types) * fv\<^sub>s\<^sub>e\<^sub>t.simps) qed lemma subst_fv_dom_ground_if_ground_img: assumes "fv t \ subst_domain s" "ground (subst_range s)" shows "fv (t \ s) = {}" using subst_fv_dom_img_subset[OF assms(1)] assms(2) by force lemma subst_fv_dom_ground_if_ground_img': assumes "fv t \ subst_domain s" "\x. x \ subst_domain s \ fv (s x) = {}" shows "fv (t \ s) = {}" using subst_fv_dom_ground_if_ground_img[OF assms(1)] assms(2) by auto lemma subst_fv_unfold: "fv (t \ s) = (fv t - subst_domain s) \ fv\<^sub>s\<^sub>e\<^sub>t (s ` (fv t \ subst_domain s))" proof (induction t) case (Var v) thus ?case proof (cases "v \ subst_domain s") case True thus ?thesis by auto next case False hence "fv (Var v \ s) = {v}" "fv (Var v) \ subst_domain s = {}" by auto thus ?thesis by auto qed next case Fun thus ?case by auto qed lemma subst_fv_unfold_ground_img: "range_vars s = {} \ fv (t \ s) = fv t - subst_domain s" -using subst_fv_unfold[of t s] unfolding range_vars_alt_def by auto +by (auto simp: range_vars_alt_def subst_fv_unfold) lemma subst_img_update: "\\ v = Var v; t \ Var v\ \ range_vars (\(v := t)) = range_vars \ \ fv t" proof - assume "\ v = Var v" "t \ Var v" hence "(\s \ {s. \w. (\(v := t)) w = s \ s \ Var w}. fv s) = fv t \ range_vars \" unfolding range_vars_alt_def by (auto simp add: subst_domain_def) thus "range_vars (\(v := t)) = range_vars \ \ fv t" by (metis Un_commute subst_fv_img_alt_def) qed lemma subst_dom_update1: "v \ subst_domain \ \ subst_domain (\(v := Var v)) = subst_domain \" by (auto simp add: subst_domain_def) lemma subst_dom_update2: "t \ Var v \ subst_domain (\(v := t)) = insert v (subst_domain \)" by (auto simp add: subst_domain_def) lemma subst_dom_update3: "t = Var v \ subst_domain (\(v := t)) = subst_domain \ - {v}" by (auto simp add: subst_domain_def) lemma var_not_in_subst_dom[elim]: "v \ subst_domain s \ s v = Var v" by (simp add: subst_domain_def) lemma subst_dom_vars_in_subst[elim]: "v \ subst_domain s \ s v \ Var v" by (simp add: subst_domain_def) lemma subst_not_dom_fixed: "\v \ fv t; v \ subst_domain s\ \ v \ fv (t \ s)" by (induct t) auto lemma subst_not_img_fixed: "\v \ fv (t \ s); v \ range_vars s\ \ v \ fv t" unfolding range_vars_alt_def by (induct t) force+ lemma ground_range_vars[intro]: "ground (subst_range s) \ range_vars s = {}" unfolding range_vars_alt_def by metis lemma ground_subst_no_var[intro]: "ground (subst_range s) \ x \ range_vars s" using ground_range_vars[of s] by blast lemma ground_img_obtain_fun: assumes "ground (subst_range s)" "x \ subst_domain s" obtains f T where "s x = Fun f T" "Fun f T \ subst_range s" "fv (Fun f T) = {}" proof - from assms(2) obtain t where t: "s x = t" "t \ subst_range s" by moura hence "fv t = {}" using assms(1) by auto thus ?thesis using t that by (cases t) simp_all qed lemma ground_term_subst_domain_fv_subset: "fv (t \ \) = {} \ fv t \ subst_domain \" by (induct t) auto lemma ground_subst_range_empty_fv: "ground (subst_range \) \ x \ subst_domain \ \ fv (\ x) = {}" by simp lemma subst_Var_notin_img: "x \ range_vars s \ t \ s = Var x \ t = Var x" using subst_not_img_fixed[of x t s] by (induct t) auto lemma fv_in_subst_img: "\s v = t; t \ Var v\ \ fv t \ range_vars s" unfolding range_vars_alt_def by auto lemma empty_dom_iff_empty_subst: "subst_domain \ = {} \ \ = Var" by auto lemma subst_dom_cong: "(\v t. \ v = t \ \ v = t) \ subst_domain \ \ subst_domain \" by (auto simp add: subst_domain_def) lemma subst_img_cong: "(\v t. \ v = t \ \ v = t) \ range_vars \ \ range_vars \" unfolding range_vars_alt_def by (auto simp add: subst_domain_def) lemma subst_dom_elim: "subst_domain s \ range_vars s = {} \ fv (t \ s) \ subst_domain s = {}" proof (induction t) case (Var v) thus ?case using fv_in_subst_img[of s] by (cases "s v = Var v") (auto simp add: subst_domain_def) next case Fun thus ?case by auto qed lemma subst_dom_insert_finite: "finite (subst_domain s) = finite (subst_domain (s(v := t)))" proof assume "finite (subst_domain s)" have "subst_domain (s(v := t)) \ insert v (subst_domain s)" by (auto simp add: subst_domain_def) thus "finite (subst_domain (s(v := t)))" by (meson \finite (subst_domain s)\ finite_insert rev_finite_subset) next assume *: "finite (subst_domain (s(v := t)))" hence "finite (insert v (subst_domain s))" proof (cases "t = Var v") case True hence "finite (subst_domain s - {v})" by (metis * subst_dom_update3) thus ?thesis by simp qed (metis * subst_dom_update2[of t v s]) thus "finite (subst_domain s)" by simp qed lemma trm_subst_disj: "t \ \ = t \ fv t \ subst_domain \ = {}" proof (induction t) case (Fun f X) hence "map (\x. x \ \) X = X" by simp hence "\x. x \ set X \ x \ \ = x" using map_eq_conv by fastforce thus ?case using Fun.IH by auto qed (simp add: subst_domain_def) -lemma trm_subst_ident[intro]: "fv t \ subst_domain \ = {} \ t \ \ = t" -proof - - assume "fv t \ subst_domain \ = {}" - hence "\v \ fv t. \w \ subst_domain \. v \ w" by auto - thus ?thesis - by (metis subst_agreement subst_apply_term.simps(1) subst_apply_term_empty subst_domI) -qed +declare subst_apply_term_ident[intro] lemma trm_subst_ident'[intro]: "v \ subst_domain \ \ (Var v) \ \ = Var v" -using trm_subst_ident by (simp add: subst_domain_def) +using subst_apply_term_ident by (simp add: subst_domain_def) lemma trm_subst_ident''[intro]: "(\x. x \ fv t \ \ x = Var x) \ t \ \ = t" proof - assume "\x. x \ fv t \ \ x = Var x" hence "fv t \ subst_domain \ = {}" by (auto simp add: subst_domain_def) - thus ?thesis using trm_subst_ident by auto + thus ?thesis using subst_apply_term_ident by auto qed lemma set_subst_ident: "fv\<^sub>s\<^sub>e\<^sub>t M \ subst_domain \ = {} \ M \\<^sub>s\<^sub>e\<^sub>t \ = M" proof - assume "fv\<^sub>s\<^sub>e\<^sub>t M \ subst_domain \ = {}" hence "\t \ M. t \ \ = t" by auto thus ?thesis by force qed lemma trm_subst_ident_subterms[intro]: "fv t \ subst_domain \ = {} \ subterms t \\<^sub>s\<^sub>e\<^sub>t \ = subterms t" using set_subst_ident[of "subterms t" \] fv_subterms[of t] by simp lemma trm_subst_ident_subterms'[intro]: "v \ fv t \ subterms t \\<^sub>s\<^sub>e\<^sub>t Var(v := s) = subterms t" using trm_subst_ident_subterms[of t "Var(v := s)"] by (meson subst_no_occs trm_subst_disj vars_iff_subtermeq) lemma const_mem_subst_cases: assumes "Fun c [] \ M \\<^sub>s\<^sub>e\<^sub>t \" shows "Fun c [] \ M \ Fun c [] \ \ ` fv\<^sub>s\<^sub>e\<^sub>t M" proof - obtain m where m: "m \ M" "m \ \ = Fun c []" using assms by auto thus ?thesis by (cases m) force+ qed lemma const_mem_subst_cases': assumes "Fun c [] \ M \\<^sub>s\<^sub>e\<^sub>t \" shows "Fun c [] \ M \ Fun c [] \ subst_range \" using const_mem_subst_cases[OF assms] by force lemma fv_subterms_substI[intro]: "y \ fv t \ \ y \ subterms t \\<^sub>s\<^sub>e\<^sub>t \" using image_iff vars_iff_subtermeq by fastforce lemma fv_subterms_subst_eq[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (subterms (t \ \)) = fv\<^sub>s\<^sub>e\<^sub>t (subterms t \\<^sub>s\<^sub>e\<^sub>t \)" using fv_subterms by (induct t) force+ lemma fv_subterms_set_subst: "fv\<^sub>s\<^sub>e\<^sub>t (subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \) = fv\<^sub>s\<^sub>e\<^sub>t (subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \))" using fv_subterms_subst_eq[of _ \] by auto lemma fv_subterms_set_subst': "fv\<^sub>s\<^sub>e\<^sub>t (subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \) = fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" using fv_subterms_set[of "M \\<^sub>s\<^sub>e\<^sub>t \"] fv_subterms_set_subst[of \ M] by simp lemma fv_subst_subset: "x \ fv t \ fv (\ x) \ fv (t \ \)" by (metis fv_subset image_eqI subst_apply_fv_unfold) lemma fv_subst_subset': "fv s \ fv t \ fv (s \ \) \ fv (t \ \)" using fv_subst_subset by (induct s) force+ lemma fv_subst_obtain_var: fixes \::"('a,'b) subst" assumes "x \ fv (t \ \)" shows "\y \ fv t. x \ fv (\ y)" using assms by (induct t) force+ lemma set_subst_all_ident: "fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) \ subst_domain \ = {} \ M \\<^sub>s\<^sub>e\<^sub>t (\ \\<^sub>s \) = M \\<^sub>s\<^sub>e\<^sub>t \" by (metis set_subst_ident subst_comp_all) lemma subterms_subst: "subterms (t \ d) = (subterms t \\<^sub>s\<^sub>e\<^sub>t d) \ subterms\<^sub>s\<^sub>e\<^sub>t (d ` (fv t \ subst_domain d))" by (induct t) (auto simp add: subst_domain_def) lemma subterms_subst': fixes \::"('a,'b) subst" assumes "\x \ fv t. (\f. \ x = Fun f []) \ (\y. \ x = Var y)" shows "subterms (t \ \) = subterms t \\<^sub>s\<^sub>e\<^sub>t \" using assms proof (induction t) case (Var x) thus ?case proof (cases "x \ subst_domain \") case True hence "(\f. \ x = Fun f []) \ (\y. \ x = Var y)" using Var by simp hence "subterms (\ x) = {\ x}" by auto thus ?thesis by simp qed auto qed auto lemma subterms_subst'': fixes \::"('a,'b) subst" assumes "\x \ fv\<^sub>s\<^sub>e\<^sub>t M. (\f. \ x = Fun f []) \ (\y. \ x = Var y)" shows "subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) = subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" using subterms_subst'[of _ \] assms by auto lemma subterms_subst_subterm: fixes \::"('a,'b) subst" assumes "\x \ fv a. (\f. \ x = Fun f []) \ (\y. \ x = Var y)" and "b \ subterms (a \ \)" shows "\c \ subterms a. c \ \ = b" using subterms_subst'[OF assms(1)] assms(2) by auto lemma subterms_subst_subset: "subterms t \\<^sub>s\<^sub>e\<^sub>t \ \ subterms (t \ \)" by (induct t) auto lemma subterms_subst_subset': "subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \ \ subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" using subterms_subst_subset by fast lemma subterms\<^sub>s\<^sub>e\<^sub>t_subst: fixes \::"('a,'b) subst" assumes "t \ subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" shows "t \ subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \ \ (\x \ fv\<^sub>s\<^sub>e\<^sub>t M. t \ subterms (\ x))" using assms subterms_subst[of _ \] by auto lemma rm_vars_dom: "subst_domain (rm_vars V s) = subst_domain s - V" by (auto simp add: subst_domain_def) lemma rm_vars_dom_subset: "subst_domain (rm_vars V s) \ subst_domain s" by (auto simp add: subst_domain_def) lemma rm_vars_dom_eq': "subst_domain (rm_vars (UNIV - V) s) = subst_domain s \ V" using rm_vars_dom[of "UNIV - V" s] by blast lemma rm_vars_dom_eqI: assumes "t \ \ = t \ \" shows "subst_domain (rm_vars (UNIV - fv t) \) = subst_domain (rm_vars (UNIV - fv t) \)" by (meson assms Diff_iff UNIV_I term_subst_eq_rev) lemma rm_vars_img: "subst_range (rm_vars V s) = s ` subst_domain (rm_vars V s)" by (auto simp add: subst_domain_def) lemma rm_vars_img_subset: "subst_range (rm_vars V s) \ subst_range s" by (auto simp add: subst_domain_def) lemma rm_vars_img_fv_subset: "range_vars (rm_vars V s) \ range_vars s" unfolding range_vars_alt_def by (auto simp add: subst_domain_def) lemma rm_vars_fv_obtain: assumes "x \ fv (t \ rm_vars X \) - X" shows "\y \ fv t - X. x \ fv (rm_vars X \ y)" using assms by (induct t) (fastforce, force) lemma rm_vars_apply: "v \ subst_domain (rm_vars V s) \ (rm_vars V s) v = s v" by (auto simp add: subst_domain_def) lemma rm_vars_apply': "subst_domain \ \ vs = {} \ rm_vars vs \ = \" by force lemma rm_vars_ident: "fv t \ vs = {} \ t \ (rm_vars vs \) = t \ \" by (induct t) auto lemma rm_vars_fv_subset: "fv (t \ rm_vars X \) \ fv t \ fv (t \ \)" by (induct t) auto lemma rm_vars_fv_disj: assumes "fv t \ X = {}" "fv (t \ \) \ X = {}" shows "fv (t \ rm_vars X \) \ X = {}" using rm_vars_ident[OF assms(1)] assms(2) by auto lemma rm_vars_ground_supports: assumes "ground (subst_range \)" shows "rm_vars X \ supports \" proof fix x have *: "ground (subst_range (rm_vars X \))" using rm_vars_img_subset[of X \] assms by (auto simp add: subst_domain_def) show "rm_vars X \ x \ \ = \ x " proof (cases "x \ subst_domain (rm_vars X \)") case True hence "fv (rm_vars X \ x) = {}" using * by auto thus ?thesis using True by auto qed (simp add: subst_domain_def) qed lemma rm_vars_split: assumes "ground (subst_range \)" shows "\ = rm_vars X \ \\<^sub>s rm_vars (subst_domain \ - X) \" proof - let ?s1 = "rm_vars X \" let ?s2 = "rm_vars (subst_domain \ - X) \" have doms: "subst_domain ?s1 \ subst_domain \" "subst_domain ?s2 \ subst_domain \" by (auto simp add: subst_domain_def) { fix x assume "x \ subst_domain \" hence "\ x = Var x" "?s1 x = Var x" "?s2 x = Var x" using doms by auto hence "\ x = (?s1 \\<^sub>s ?s2) x" by (simp add: subst_compose_def) } moreover { fix x assume "x \ subst_domain \" "x \ X" hence "?s1 x = Var x" "?s2 x = \ x" using doms by auto hence "\ x = (?s1 \\<^sub>s ?s2) x" by (simp add: subst_compose_def) } moreover { fix x assume "x \ subst_domain \" "x \ X" hence "?s1 x = \ x" "fv (\ x) = {}" using assms doms by auto hence "\ x = (?s1 \\<^sub>s ?s2) x" by (simp add: subst_compose subst_ground_ident) } ultimately show ?thesis by blast qed lemma rm_vars_fv_img_disj: assumes "fv t \ X = {}" "X \ range_vars \ = {}" shows "fv (t \ rm_vars X \) \ X = {}" using assms proof (induction t) case (Var x) hence *: "(rm_vars X \) x = \ x" by auto show ?case proof (cases "x \ subst_domain \") case True hence "\ x \ subst_range \" by auto hence "fv (\ x) \ X = {}" using Var.prems(2) unfolding range_vars_alt_def by fastforce thus ?thesis using * by auto next case False thus ?thesis using Var.prems(1) by auto qed next case Fun thus ?case by auto qed lemma subst_apply_dom_ident: "t \ \ = t \ subst_domain \ \ subst_domain \ \ t \ \ = t" proof (induction t) case (Fun f T) thus ?case by (induct T) auto qed (auto simp add: subst_domain_def) lemma rm_vars_subst_apply_ident: assumes "t \ \ = t" shows "t \ (rm_vars vs \) = t" using rm_vars_dom[of vs \] subst_apply_dom_ident[OF assms, of "rm_vars vs \"] by auto lemma rm_vars_subst_eq: "t \ \ = t \ rm_vars (subst_domain \ - subst_domain \ \ fv t) \" by (auto intro: term_subst_eq) lemma rm_vars_subst_eq': "t \ \ = t \ rm_vars (UNIV - fv t) \" by (auto intro: term_subst_eq) lemma rm_vars_comp: assumes "range_vars \ \ vs = {}" shows "t \ rm_vars vs (\ \\<^sub>s \) = t \ (rm_vars vs \ \\<^sub>s rm_vars vs \)" using assms proof (induction t) case (Var x) thus ?case proof (cases "x \ vs") - case True thus ?thesis using Var by auto + case True thus ?thesis using Var + by (simp add: subst_compose_def) next case False have "subst_domain (rm_vars vs \) \ vs = {}" by (auto simp add: subst_domain_def) moreover have "fv (\ x) \ vs = {}" using Var False unfolding range_vars_alt_def by force ultimately have "\ x \ (rm_vars vs \) = \ x \ \" using rm_vars_ident by (simp add: subst_domain_def) moreover have "(rm_vars vs (\ \\<^sub>s \)) x = (\ \\<^sub>s \) x" by (metis False) - ultimately show ?thesis using subst_compose by auto + ultimately show ?thesis by (auto simp: subst_compose) qed next case Fun thus ?case by auto qed lemma rm_vars_fv\<^sub>s\<^sub>e\<^sub>t_subst: assumes "x \ fv\<^sub>s\<^sub>e\<^sub>t (rm_vars X \ ` Y)" shows "x \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` Y) \ x \ X" using assms by auto lemma disj_dom_img_var_notin: assumes "subst_domain \ \ range_vars \ = {}" "\ v = t" "t \ Var v" shows "v \ fv t" "\v \ fv (t \ \). v \ subst_domain \" proof - have "v \ subst_domain \" "fv t \ range_vars \" using fv_in_subst_img[of \ v t, OF assms(2)] assms(2,3) by (auto simp add: subst_domain_def) thus "v \ fv t" using assms(1) by auto have *: "fv t \ subst_domain \ = {}" using assms(1) \fv t \ range_vars \\ by auto hence "t \ \ = t" by blast thus "\v \ fv (t \ \). v \ subst_domain \" using * by auto qed lemma subst_sends_dom_to_img: "v \ subst_domain \ \ fv (Var v \ \) \ range_vars \" unfolding range_vars_alt_def by auto lemma subst_sends_fv_to_img: "fv (t \ s) \ fv t \ range_vars s" proof (induction t) case (Var v) thus ?case proof (cases "Var v \ s = Var v") case True thus ?thesis by simp next case False hence "v \ subst_domain s" by (meson trm_subst_ident') hence "fv (Var v \ s) \ range_vars s" using subst_sends_dom_to_img by simp thus ?thesis by auto qed next case Fun thus ?case by auto qed lemma ident_comp_subst_trm_if_disj: assumes "subst_domain \ \ range_vars \ = {}" "v \ subst_domain \" shows "(\ \\<^sub>s \) v = \ v" proof - from assms have " subst_domain \ \ fv (\ v) = {}" using fv_in_subst_img unfolding range_vars_alt_def by auto thus "(\ \\<^sub>s \) v = \ v" unfolding subst_compose_def by blast qed lemma ident_comp_subst_trm_if_disj': "fv (\ v) \ subst_domain \ = {} \ (\ \\<^sub>s \) v = \ v" unfolding subst_compose_def by blast lemma subst_idemI[intro]: "subst_domain \ \ range_vars \ = {} \ subst_idem \" using ident_comp_subst_trm_if_disj[of \ \] var_not_in_subst_dom[of _ \] subst_eq_if_eq_vars[of \] by (metis subst_idem_def subst_compose_def var_comp(2)) lemma subst_idemI'[intro]: "ground (subst_range \) \ subst_idem \" proof (intro subst_idemI) assume "ground (subst_range \)" hence "range_vars \ = {}" by (metis ground_range_vars) thus "subst_domain \ \ range_vars \ = {}" by blast qed lemma subst_idemE: "subst_idem \ \ subst_domain \ \ range_vars \ = {}" proof - assume "subst_idem \" hence "\v. fv (\ v) \ subst_domain \ = {}" unfolding subst_idem_def subst_compose_def by (metis trm_subst_disj) thus ?thesis unfolding range_vars_alt_def by auto qed lemma subst_idem_rm_vars: "subst_idem \ \ subst_idem (rm_vars X \)" proof - assume "subst_idem \" hence "subst_domain \ \ range_vars \ = {}" by (metis subst_idemE) moreover have "subst_domain (rm_vars X \) \ subst_domain \" "range_vars (rm_vars X \) \ range_vars \" unfolding range_vars_alt_def by (auto simp add: subst_domain_def) ultimately show ?thesis by blast qed lemma subst_fv_bounded_if_img_bounded: "range_vars \ \ fv t \ V \ fv (t \ \) \ fv t \ V" proof (induction t) case (Var v) thus ?case unfolding range_vars_alt_def by (cases "\ v = Var v") auto qed (metis (no_types, lifting) Un_assoc Un_commute subst_sends_fv_to_img sup.absorb_iff2) lemma subst_fv_bound_singleton: "fv (t \ Var(v := t')) \ fv t \ fv t'" using subst_fv_bounded_if_img_bounded[of "Var(v := t')" t "fv t'"] unfolding range_vars_alt_def by (auto simp add: subst_domain_def) lemma subst_fv_bounded_if_img_bounded': assumes "range_vars \ \ fv\<^sub>s\<^sub>e\<^sub>t M" shows "fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t M" proof fix v assume *: "v \ fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" obtain t where t: "t \ M" "t \ \ \ M \\<^sub>s\<^sub>e\<^sub>t \" "v \ fv (t \ \)" proof - assume **: "\t. \t \ M; t \ \ \ M \\<^sub>s\<^sub>e\<^sub>t \; v \ fv (t \ \)\ \ thesis" have "v \ \ (fv ` ((\t. t \ \) ` M))" using * by (metis fv\<^sub>s\<^sub>e\<^sub>t.simps) hence "\t. t \ M \ v \ fv (t \ \)" by blast thus ?thesis using ** imageI by blast qed from \t \ M\ obtain M' where "t \ M'" "M = insert t M'" by (meson Set.set_insert) hence "fv\<^sub>s\<^sub>e\<^sub>t M = fv t \ fv\<^sub>s\<^sub>e\<^sub>t M'" by simp hence "fv (t \ \) \ fv\<^sub>s\<^sub>e\<^sub>t M" using subst_fv_bounded_if_img_bounded assms by simp thus "v \ fv\<^sub>s\<^sub>e\<^sub>t M" using assms \v \ fv (t \ \)\ by auto qed lemma ground_img_if_ground_subst: "(\v t. s v = t \ fv t = {}) \ range_vars s = {}" unfolding range_vars_alt_def by auto lemma ground_subst_fv_subset: "ground (subst_range \) \ fv (t \ \) \ fv t" using subst_fv_bounded_if_img_bounded[of \] unfolding range_vars_alt_def by force lemma ground_subst_fv_subset': "ground (subst_range \) \ fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t M" using subst_fv_bounded_if_img_bounded'[of \ M] unfolding range_vars_alt_def by auto lemma subst_to_var_is_var[elim]: "t \ s = Var v \ \w. t = Var w" -using subst_apply_term.elims by blast +by (auto elim!: eval_term.elims) lemma subst_dom_comp_inI: assumes "y \ subst_domain \" and "y \ subst_domain \" shows "y \ subst_domain (\ \\<^sub>s \)" using assms subst_domain_subst_compose[of \ \] by blast lemma subst_comp_notin_dom_eq: "x \ subst_domain \1 \ (\1 \\<^sub>s \2) x = \2 x" unfolding subst_compose_def by fastforce lemma subst_dom_comp_eq: assumes "subst_domain \ \ range_vars \ = {}" shows "subst_domain (\ \\<^sub>s \) = subst_domain \ \ subst_domain \" proof (rule ccontr) assume "subst_domain (\ \\<^sub>s \) \ subst_domain \ \ subst_domain \" hence "subst_domain (\ \\<^sub>s \) \ subst_domain \ \ subst_domain \" using subst_domain_compose[of \ \] by (simp add: subst_domain_def) then obtain v where "v \ subst_domain (\ \\<^sub>s \)" "v \ subst_domain \ \ subst_domain \" by auto hence v_in_some_subst: "\ v \ Var v \ \ v \ Var v" and "\ v \ \ = Var v" unfolding subst_compose_def by (auto simp add: subst_domain_def) then obtain w where "\ v = Var w" using subst_to_var_is_var by fastforce show False proof (cases "v = w") case True hence "\ v = Var v" using \\ v = Var w\ by simp hence "\ v \ Var v" using v_in_some_subst by simp thus False using \\ v = Var v\ \\ v \ \ = Var v\ by simp next case False hence "v \ subst_domain \" using v_in_some_subst \\ v \ \ = Var v\ by auto hence "v \ range_vars \" using assms by auto moreover have "\ w = Var v" using \\ v \ \ = Var v\ \\ v = Var w\ by simp hence "v \ range_vars \" using \v \ w\ subst_fv_imgI[of \ w] by simp ultimately show False .. qed qed lemma subst_img_comp_subset[simp]: "range_vars (\1 \\<^sub>s \2) \ range_vars \1 \ range_vars \2" proof let ?img = "range_vars" fix x assume "x \ ?img (\1 \\<^sub>s \2)" then obtain v t where vt: "x \ fv t" "t = (\1 \\<^sub>s \2) v" "t \ Var v" unfolding range_vars_alt_def subst_compose_def by (auto simp add: subst_domain_def) { assume "x \ ?img \1" hence "x \ ?img \2" by (metis (no_types, opaque_lifting) fv_in_subst_img Un_iff subst_compose_def - vt subsetCE subst_apply_term.simps(1) subst_sends_fv_to_img) + vt subsetCE eval_term.simps(1) subst_sends_fv_to_img) } thus "x \ ?img \1 \ ?img \2" by auto qed lemma subst_img_comp_subset': assumes "t \ subst_range (\1 \\<^sub>s \2)" shows "t \ subst_range \2 \ (\t' \ subst_range \1. t = t' \ \2)" proof - obtain x where x: "x \ subst_domain (\1 \\<^sub>s \2)" "(\1 \\<^sub>s \2) x = t" "t \ Var x" using assms by (auto simp add: subst_domain_def) { assume "x \ subst_domain \1" hence "(\1 \\<^sub>s \2) x = \2 x" unfolding subst_compose_def by auto hence ?thesis using x by auto } moreover { assume "x \ subst_domain \1" hence ?thesis using subst_compose x(2) by fastforce } ultimately show ?thesis by metis qed lemma subst_img_comp_subset'': "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range (\1 \\<^sub>s \2)) \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \2) \ ((subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \1)) \\<^sub>s\<^sub>e\<^sub>t \2)" proof fix t assume "t \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range (\1 \\<^sub>s \2))" then obtain x where x: "x \ subst_domain (\1 \\<^sub>s \2)" "t \ subterms ((\1 \\<^sub>s \2) x)" by auto show "t \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \2) \ (subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \1) \\<^sub>s\<^sub>e\<^sub>t \2)" proof (cases "x \ subst_domain \1") case True thus ?thesis using subst_compose[of \1 \2] x(2) subterms_subst by fastforce next case False hence "(\1 \\<^sub>s \2) x = \2 x" unfolding subst_compose_def by auto thus ?thesis using x by (auto simp add: subst_domain_def) qed qed lemma subst_img_comp_subset''': "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range (\1 \\<^sub>s \2)) - range Var \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \2) - range Var \ ((subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \1) - range Var) \\<^sub>s\<^sub>e\<^sub>t \2)" proof fix t assume t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range (\1 \\<^sub>s \2)) - range Var" then obtain f T where fT: "t = Fun f T" by (cases t) simp_all then obtain x where x: "x \ subst_domain (\1 \\<^sub>s \2)" "Fun f T \ subterms ((\1 \\<^sub>s \2) x)" using t by auto have "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \2) \ (subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \1) - range Var \\<^sub>s\<^sub>e\<^sub>t \2)" proof (cases "x \ subst_domain \1") case True hence "Fun f T \ (subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \2)) \ (subterms (\1 x) \\<^sub>s\<^sub>e\<^sub>t \2)" - using x(2) subterms_subst[of "\1 x" \2] - unfolding subst_compose[of \1 \2 x] by auto + using x(2) + by (auto simp: subst_compose subterms_subst) moreover have ?thesis when *: "Fun f T \ subterms (\1 x) \\<^sub>s\<^sub>e\<^sub>t \2" proof - obtain s where s: "s \ subterms (\1 x)" "Fun f T = s \ \2" using * by moura show ?thesis proof (cases s) case (Var y) hence "Fun f T \ subst_range \2" using s by force thus ?thesis by blast next case (Fun g S) hence "Fun f T \ (subterms (\1 x) - range Var) \\<^sub>s\<^sub>e\<^sub>t \2" using s by blast thus ?thesis using True by auto qed qed ultimately show ?thesis by blast next case False hence "(\1 \\<^sub>s \2) x = \2 x" unfolding subst_compose_def by auto thus ?thesis using x by (auto simp add: subst_domain_def) qed thus "t \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \2) - range Var \ (subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \1) - range Var \\<^sub>s\<^sub>e\<^sub>t \2)" using fT by auto qed lemma subst_img_comp_subset_const: assumes "Fun c [] \ subst_range (\1 \\<^sub>s \2)" shows "Fun c [] \ subst_range \2 \ Fun c [] \ subst_range \1 \ (\x. Var x \ subst_range \1 \ \2 x = Fun c [])" proof (cases "Fun c [] \ subst_range \2") case False then obtain t where t: "t \ subst_range \1" "Fun c [] = t \ \2" using subst_img_comp_subset'[OF assms] by auto thus ?thesis by (cases t) auto qed (simp add: subst_img_comp_subset'[OF assms]) lemma subst_img_comp_subset_const': fixes \ \::"('f,'v) subst" assumes "(\ \\<^sub>s \) x = Fun c []" shows "\ x = Fun c [] \ (\z. \ x = Var z \ \ z = Fun c [])" proof (cases "\ x = Fun c []") case False then obtain t where "\ x = t" "t \ \ = Fun c []" using assms unfolding subst_compose_def by auto thus ?thesis by (cases t) auto qed simp lemma subst_img_comp_subset_ground: assumes "ground (subst_range \1)" shows "subst_range (\1 \\<^sub>s \2) \ subst_range \1 \ subst_range \2" proof fix t assume t: "t \ subst_range (\1 \\<^sub>s \2)" then obtain x where x: "x \ subst_domain (\1 \\<^sub>s \2)" "t = (\1 \\<^sub>s \2) x" by auto show "t \ subst_range \1 \ subst_range \2" proof (cases "x \ subst_domain \1") case True hence "fv (\1 x) = {}" using assms ground_subst_range_empty_fv by fast hence "t = \1 x" using x(2) unfolding subst_compose_def by blast thus ?thesis using True by simp next case False hence "t = \2 x" "x \ subst_domain \2" using x subst_domain_compose[of \1 \2] by (metis subst_comp_notin_dom_eq, blast) thus ?thesis using x by simp qed qed lemma subst_fv_dom_img_single: assumes "v \ fv t" "\ v = t" "\w. v \ w \ \ w = Var w" shows "subst_domain \ = {v}" "range_vars \ = fv t" proof - show "subst_domain \ = {v}" using assms by (fastforce simp add: subst_domain_def) have "fv t \ range_vars \" by (metis fv_in_subst_img assms(1,2) vars_iff_subterm_or_eq) moreover have "\v. \ v \ Var v \ \ v = t" using assms by fastforce ultimately show "range_vars \ = fv t" unfolding range_vars_alt_def by (auto simp add: subst_domain_def) qed lemma subst_comp_upd1: "\(v := t) \\<^sub>s \ = (\ \\<^sub>s \)(v := t \ \)" unfolding subst_compose_def by auto lemma subst_comp_upd2: assumes "v \ subst_domain s" "v \ range_vars s" shows "s(v := t) = s \\<^sub>s (Var(v := t))" unfolding subst_compose_def proof - { fix w have "(s(v := t)) w = s w \ Var(v := t)" proof (cases "w = v") case True hence "s w = Var w" using \v \ subst_domain s\ by (simp add: subst_domain_def) thus ?thesis using \w = v\ by simp next case False hence "(s(v := t)) w = s w" by simp moreover have "s w \ Var(v := t) = s w" using \w \ v\ \v \ range_vars s\ by (metis fv_in_subst_img fun_upd_apply insert_absorb insert_subset - repl_invariance subst_apply_term.simps(1) subst_apply_term_empty) + repl_invariance eval_term.simps(1) subst_apply_term_empty) ultimately show ?thesis .. qed } thus "s(v := t) = (\w. s w \ Var(v := t))" by auto qed lemma ground_subst_dom_iff_img: "ground (subst_range \) \ x \ subst_domain \ \ \ x \ subst_range \" by (auto simp add: subst_domain_def) lemma finite_dom_subst_exists: "finite S \ \\::('f,'v) subst. subst_domain \ = S" proof (induction S rule: finite.induct) case (insertI A a) then obtain \::"('f,'v) subst" where "subst_domain \ = A" by blast fix f::'f have "subst_domain (\(a := Fun f [])) = insert a A" using \subst_domain \ = A\ by (auto simp add: subst_domain_def) thus ?case by metis qed (auto simp add: subst_domain_def) lemma subst_inj_is_bij_betw_dom_img_if_ground_img: assumes "ground (subst_range \)" shows "inj \ \ bij_betw \ (subst_domain \) (subst_range \)" (is "?A \ ?B") proof show "?A \ ?B" by (metis bij_betw_def injD inj_onI subst_range.simps) next assume ?B hence "inj_on \ (subst_domain \)" unfolding bij_betw_def by auto moreover have "\x. x \ UNIV - subst_domain \ \ \ x = Var x" by auto hence "inj_on \ (UNIV - subst_domain \)" using inj_onI[of "UNIV - subst_domain \"] by (metis term.inject(1)) moreover have "\x y. x \ subst_domain \ \ y \ subst_domain \ \ \ x \ \ y" using assms by (auto simp add: subst_domain_def) ultimately show ?A by (metis injI inj_onD subst_domI term.inject(1)) qed lemma bij_finite_ground_subst_exists: assumes "finite (S::'v set)" "infinite (U::('f,'v) term set)" "ground U" shows "\\::('f,'v) subst. subst_domain \ = S \ bij_betw \ (subst_domain \) (subst_range \) \ subst_range \ \ U" proof - obtain T' where "T' \ U" "card T' = card S" "finite T'" by (meson assms(2) finite_Diff2 infinite_arbitrarily_large) then obtain f::"'v \ ('f,'v) term" where f_bij: "bij_betw f S T'" using finite_same_card_bij[OF assms(1)] by metis hence *: "\v. v \ S \ f v \ Var v" using \ground U\ \T' \ U\ bij_betwE by fastforce let ?\ = "\v. if v \ S then f v else Var v" have "subst_domain ?\ = S" proof show "subst_domain ?\ \ S" by (auto simp add: subst_domain_def) { fix v assume "v \ S" "v \ subst_domain ?\" hence "f v = Var v" by (simp add: subst_domain_def) hence False using *[OF \v \ S\] by metis } thus "S \ subst_domain ?\" by blast qed hence "\v w. \v \ subst_domain ?\; w \ subst_domain ?\\ \ ?\ w \ ?\ v" using \ground U\ bij_betwE[OF f_bij] set_rev_mp[OF _ \T' \ U\] by (metis (no_types, lifting) UN_iff empty_iff vars_iff_subterm_or_eq fv\<^sub>s\<^sub>e\<^sub>t.simps) hence "inj_on ?\ (subst_domain ?\)" using f_bij \subst_domain ?\ = S\ unfolding bij_betw_def inj_on_def by metis hence "bij_betw ?\ (subst_domain ?\) (subst_range ?\)" using inj_on_imp_bij_betw[of ?\] by simp moreover have "subst_range ?\ = T'" using \bij_betw f S T'\ \subst_domain ?\ = S\ unfolding bij_betw_def by auto hence "subst_range ?\ \ U" using \T' \ U\ by auto ultimately show ?thesis using \subst_domain ?\ = S\ by (metis (lifting)) qed lemma bij_finite_const_subst_exists: assumes "finite (S::'v set)" "finite (T::'f set)" "infinite (U::'f set)" shows "\\::('f,'v) subst. subst_domain \ = S \ bij_betw \ (subst_domain \) (subst_range \) \ subst_range \ \ (\c. Fun c []) ` (U - T)" proof - obtain T' where "T' \ U - T" "card T' = card S" "finite T'" by (meson assms(2,3) finite_Diff2 infinite_arbitrarily_large) then obtain f::"'v \ 'f" where f_bij: "bij_betw f S T'" using finite_same_card_bij[OF assms(1)] by metis let ?\ = "\v. if v \ S then Fun (f v) [] else Var v" have "subst_domain ?\ = S" by (simp add: subst_domain_def) moreover have "\v w. \v \ subst_domain ?\; w \ subst_domain ?\\ \ ?\ w \ ?\ v" by auto hence "inj_on ?\ (subst_domain ?\)" using f_bij unfolding bij_betw_def inj_on_def by (metis \subst_domain ?\ = S\ term.inject(2)) hence "bij_betw ?\ (subst_domain ?\) (subst_range ?\)" using inj_on_imp_bij_betw[of ?\] by simp moreover have "subst_range ?\ = ((\c. Fun c []) ` T')" using \bij_betw f S T'\ unfolding bij_betw_def inj_on_def by (auto simp add: subst_domain_def) hence "subst_range ?\ \ ((\c. Fun c []) ` (U - T))" using \T' \ U - T\ by auto ultimately show ?thesis by (metis (lifting)) qed lemma bij_finite_const_subst_exists': assumes "finite (S::'v set)" "finite (T::('f,'v) terms)" "infinite (U::'f set)" shows "\\::('f,'v) subst. subst_domain \ = S \ bij_betw \ (subst_domain \) (subst_range \) \ subst_range \ \ ((\c. Fun c []) ` U) - T" proof - have "finite (\(funs_term ` T))" using assms(2) by auto then obtain \ where \: "subst_domain \ = S" "bij_betw \ (subst_domain \) (subst_range \)" "subst_range \ \ (\c. Fun c []) ` (U - (\(funs_term ` T)))" using bij_finite_const_subst_exists[OF assms(1) _ assms(3)] by blast moreover have "(\c. Fun c []) ` (U - (\(funs_term ` T))) \ ((\c. Fun c []) ` U) - T" by auto ultimately show ?thesis by blast qed lemma bij_betw_iteI: assumes "bij_betw f A B" "bij_betw g C D" "A \ C = {}" "B \ D = {}" shows "bij_betw (\x. if x \ A then f x else g x) (A \ C) (B \ D)" proof - have "bij_betw (\x. if x \ A then f x else g x) A B" by (metis bij_betw_cong[of A f "\x. if x \ A then f x else g x" B] assms(1)) moreover have "bij_betw (\x. if x \ A then f x else g x) C D" using bij_betw_cong[of C g "\x. if x \ A then f x else g x" D] assms(2,3) by force ultimately show ?thesis using bij_betw_combine[OF _ _ assms(4)] by metis qed lemma subst_comp_split: assumes "subst_domain \ \ range_vars \ = {}" shows "\ = (rm_vars (subst_domain \ - V) \) \\<^sub>s (rm_vars V \)" (is ?P) and "\ = (rm_vars V \) \\<^sub>s (rm_vars (subst_domain \ - V) \)" (is ?Q) proof - let ?rm1 = "rm_vars (subst_domain \ - V) \" and ?rm2 = "rm_vars V \" have "subst_domain ?rm2 \ range_vars ?rm1 = {}" "subst_domain ?rm1 \ range_vars ?rm2 = {}" using assms unfolding range_vars_alt_def by (force simp add: subst_domain_def)+ hence *: "\v. v \ subst_domain ?rm1 \ (?rm1 \\<^sub>s ?rm2) v = \ v" "\v. v \ subst_domain ?rm2 \ (?rm2 \\<^sub>s ?rm1) v = \ v" using ident_comp_subst_trm_if_disj[of ?rm2 ?rm1] ident_comp_subst_trm_if_disj[of ?rm1 ?rm2] by (auto simp add: subst_domain_def) hence "\v. v \ subst_domain ?rm1 \ (?rm1 \\<^sub>s ?rm2) v = \ v" "\v. v \ subst_domain ?rm2 \ (?rm2 \\<^sub>s ?rm1) v = \ v" unfolding subst_compose_def by (auto simp add: subst_domain_def) hence "\v. (?rm1 \\<^sub>s ?rm2) v = \ v" "\v. (?rm2 \\<^sub>s ?rm1) v = \ v" using * by blast+ thus ?P ?Q by auto qed lemma subst_comp_eq_if_disjoint_vars: assumes "(subst_domain \ \ range_vars \) \ (subst_domain \ \ range_vars \) = {}" shows "\ \\<^sub>s \ = \ \\<^sub>s \" proof - { fix x assume "x \ subst_domain \" hence "(\ \\<^sub>s \) x = \ x" "(\ \\<^sub>s \) x = \ x" using assms unfolding range_vars_alt_def by (force simp add: subst_compose)+ hence "(\ \\<^sub>s \) x = (\ \\<^sub>s \) x" by metis } moreover { fix x assume "x \ subst_domain \" hence "(\ \\<^sub>s \) x = \ x" "(\ \\<^sub>s \) x = \ x" using assms unfolding range_vars_alt_def by (auto simp add: subst_compose subst_domain_def) hence "(\ \\<^sub>s \) x = (\ \\<^sub>s \) x" by metis } moreover { fix x assume "x \ subst_domain \" "x \ subst_domain \" hence "(\ \\<^sub>s \) x = (\ \\<^sub>s \) x" by (simp add: subst_compose subst_domain_def) } ultimately show ?thesis by auto qed lemma subst_eq_if_disjoint_vars_ground: fixes \ \::"('f,'v) subst" assumes "subst_domain \ \ subst_domain \ = {}" "ground (subst_range \)" "ground (subst_range \)" shows "t \ \ \ \ = t \ \ \ \" by (metis assms subst_comp_eq_if_disjoint_vars range_vars_alt_def subst_subst_compose sup_bot.right_neutral) lemma subst_img_bound: "subst_domain \ \ range_vars \ \ fv t \ range_vars \ \ fv (t \ \)" proof - assume "subst_domain \ \ range_vars \ \ fv t" hence "subst_domain \ \ fv t" by blast thus ?thesis by (metis (no_types) range_vars_alt_def le_iff_sup subst_apply_fv_unfold subst_apply_fv_union subst_range.simps) qed lemma subst_all_fv_subset: "fv t \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv (t \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" proof - assume *: "fv t \ fv\<^sub>s\<^sub>e\<^sub>t M" { fix v assume "v \ fv t" hence "v \ fv\<^sub>s\<^sub>e\<^sub>t M" using * by auto then obtain t' where "t' \ M" "v \ fv t'" by auto hence "fv (\ v) \ fv (t' \ \)" - by (metis subst_apply_term.simps(1) subst_apply_fv_subset subst_apply_fv_unfold + by (metis eval_term.simps(1) subst_apply_fv_subset subst_apply_fv_unfold subtermeq_vars_subset vars_iff_subtermeq) hence "fv (\ v) \ fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" using \t' \ M\ by auto } - thus ?thesis using subst_apply_fv_unfold[of t \] by auto + thus ?thesis by (auto simp: subst_apply_fv_unfold) qed lemma subst_support_if_mgt_subst_idem: assumes "\ \\<^sub>\ \" "subst_idem \" shows "\ supports \" proof - from \\ \\<^sub>\ \\ obtain \ where \: "\ = \ \\<^sub>s \" by blast - hence "\v. \ v \ \ = Var v \ (\ \\<^sub>s \ \\<^sub>s \)" by simp + hence "\v. \ v \ \ = Var v \ (\ \\<^sub>s \ \\<^sub>s \)" by (simp add: subst_compose) hence "\v. \ v \ \ = Var v \ (\ \\<^sub>s \)" using \subst_idem \ \ unfolding subst_idem_def by simp hence "\v. \ v \ \ = Var v \ \" using \ by simp thus "\ supports \" by simp qed lemma subst_support_iff_mgt_if_subst_idem: assumes "subst_idem \" shows "\ \\<^sub>\ \ \ \ supports \" proof show "\ \\<^sub>\ \ \ \ supports \" by (fact subst_support_if_mgt_subst_idem[OF _ \subst_idem \\]) show "\ supports \ \ \ \\<^sub>\ \" by (fact subst_supportD) qed lemma subst_support_comp: fixes \ \ \::"('a,'b) subst" assumes "\ supports \" "\ supports \" shows "(\ \\<^sub>s \) supports \" -by (metis (no_types) assms subst_agreement subst_apply_term.simps(1) subst_subst_compose) +by (metis (no_types) assms subst_agreement eval_term.simps(1) subst_subst_compose) lemma subst_support_comp': fixes \ \ \::"('a,'b) subst" assumes "\ supports \" shows "\ supports (\ \\<^sub>s \)" "\ supports \ \ \ supports (\ \\<^sub>s \)" using assms unfolding subst_support_def by (metis subst_compose_assoc, metis) lemma subst_support_comp_split: fixes \ \ \::"('a,'b) subst" assumes "(\ \\<^sub>s \) supports \" shows "subst_domain \ \ range_vars \ = {} \ \ supports \" and "subst_domain \ \ subst_domain \ = {} \ \ supports \" proof - assume "subst_domain \ \ range_vars \ = {}" hence "subst_idem \" by (metis subst_idemI) have "\ \\<^sub>\ \" using assms subst_compose_assoc[of \ \ \] unfolding subst_compose_def by metis show "\ supports \" using subst_support_if_mgt_subst_idem[OF \\ \\<^sub>\ \\ \subst_idem \\] by auto next assume "subst_domain \ \ subst_domain \ = {}" moreover have "\v \ subst_domain (\ \\<^sub>s \). (\ \\<^sub>s \) v \ \ = \ v" using assms by metis ultimately have "\v \ subst_domain \. \ v \ \ = \ v" using var_not_in_subst_dom unfolding subst_compose_def - by (metis IntI empty_iff subst_apply_term.simps(1)) + by (metis IntI empty_iff eval_term.simps(1)) thus "\ supports \" by force qed lemma subst_idem_support: "subst_idem \ \ \ supports \ \\<^sub>s \" unfolding subst_idem_def by (metis subst_support_def subst_compose_assoc) lemma subst_idem_iff_self_support: "subst_idem \ \ \ supports \" using subst_support_def[of \ \] unfolding subst_idem_def by auto lemma subterm_subst_neq: "t \ t' \ t \ s \ t' \ s" by (metis subst_mono_neq) lemma fv_Fun_subst_neq: "x \ fv (Fun f T) \ \ x \ Fun f T \ \" using subterm_subst_neq[of "Var x" "Fun f T"] vars_iff_subterm_or_eq[of x "Fun f T"] by auto lemma subterm_subst_unfold: assumes "t \ s \ \" shows "(\s'. s' \ s \ t = s' \ \) \ (\x \ fv s. t \ \ x)" using assms proof (induction s) case (Fun f T) thus ?case proof (cases "t = Fun f T \ \") case True thus ?thesis using Fun by auto next case False then obtain s' where s': "s' \ set T" "t \ s' \ \" using Fun by auto hence "(\s''. s'' \ s' \ t = s'' \ \) \ (\x \ fv s'. t \ \ x)" by (metis Fun.IH) thus ?thesis using s'(1) by auto qed qed simp lemma subterm_subst_img_subterm: assumes "t \ s \ \" "\s'. s' \ s \ t \ s' \ \" shows "\w \ fv s. t \ \ w" using subterm_subst_unfold[OF assms(1)] assms(2) by force lemma subterm_subst_not_img_subterm: assumes "t \ s \ \" "\(\w \ fv s. t \ \ w)" shows "\f T. Fun f T \ s \ t = Fun f T \ \" proof (rule ccontr) assume "\(\f T. Fun f T \ s \ t = Fun f T \ \)" hence "\f T. Fun f T \ s \ t \ Fun f T \ \" by simp moreover have "\x. Var x \ s \ t \ Var x \ \" using assms(2) vars_iff_subtermeq by force ultimately have "\s'. s' \ s \ t \ s' \ \" by (metis "term.exhaust") thus False using assms subterm_subst_img_subterm by blast qed lemma subst_apply_img_var: assumes "v \ fv (t \ \)" "v \ fv t" obtains w where "w \ fv t" "v \ fv (\ w)" using assms by (induct t) auto lemma subst_apply_img_var': assumes "x \ fv (t \ \)" "x \ fv t" shows "\y \ fv t. x \ fv (\ y)" by (metis assms subst_apply_img_var) lemma nth_map_subst: fixes \::"('f,'v) subst" and T::"('f,'v) term list" and i::nat shows "i < length T \ (map (\t. t \ \) T) ! i = (T ! i) \ \" by (fact nth_map) lemma subst_subterm: assumes "Fun f T \ t \ \" shows "(\S. Fun f S \ t \ Fun f S \ \ = Fun f T) \ (\s \ subst_range \. Fun f T \ s)" using assms subterm_subst_not_img_subterm by (cases "\s \ subst_range \. Fun f T \ s") fastforce+ lemma subst_subterm': assumes "Fun f T \ t \ \" shows "\S. length S = length T \ (Fun f S \ t \ (\s \ subst_range \. Fun f S \ s))" using subst_subterm[OF assms] by auto lemma subst_subterm'': assumes "s \ subterms (t \ \)" shows "(\u \ subterms t. s = u \ \) \ s \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \)" proof (cases s) case (Var x) thus ?thesis using assms subterm_subst_not_img_subterm vars_iff_subtermeq by (cases "s = t \ \") fastforce+ next case (Fun f T) thus ?thesis - using subst_subterm[of f T t \] assms + using subst_subterm assms by fastforce qed lemma fv_ground_subst_compose: assumes "subst_domain \ = subst_domain \" and "range_vars \ = {}" "range_vars \ = {}" shows "fv (t \ \ \\<^sub>s \) = fv (t \ \ \\<^sub>s \)" proof (induction t) case (Var x) show ?case proof (cases "x \ subst_domain \") case True thus ?thesis - using assms unfolding range_vars_alt_def by (auto simp add: subst_apply_fv_empty) + using assms unfolding range_vars_alt_def by (auto simp: subst_compose subst_apply_fv_empty) next case False hence "\ x = Var x" "\ x = Var x" using assms(1) by (blast,blast) - thus ?thesis by simp + thus ?thesis by (simp add: subst_compose) qed qed simp subsection \More Small Lemmata\ lemma funs_term_subst: "funs_term (t \ \) = funs_term t \ (\x \ fv t. funs_term (\ x))" by (induct t) auto lemma fv\<^sub>s\<^sub>e\<^sub>t_subst_img_eq: assumes "X \ (subst_domain \ \ range_vars \) = {}" shows "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (Y - X)) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` Y) - X" using assms unfolding range_vars_alt_def by force lemma subst_Fun_index_eq: assumes "i < length T" "Fun f T \ \ = Fun g T' \ \" shows "T ! i \ \ = T' ! i \ \" proof - have "map (\x. x \ \) T = map (\x. x \ \) T'" using assms by simp thus ?thesis by (metis assms(1) length_map nth_map) qed lemma fv_exists_if_unifiable_and_neq: fixes t t'::"('a,'b) term" and \ \::"('a,'b) subst" assumes "t \ t'" "t \ \ = t' \ \" shows "fv t \ fv t' \ {}" proof assume "fv t \ fv t' = {}" hence "fv t = {}" "fv t' = {}" by auto hence "t \ \ = t" "t' \ \ = t'" by auto hence "t = t'" using assms(2) by metis thus False using assms(1) by auto qed lemma const_subterm_subst: "Fun c [] \ t \ Fun c [] \ t \ \" by (induct t) auto lemma const_subterm_subst_var_obtain: assumes "Fun c [] \ t \ \" "\Fun c [] \ t" obtains x where "x \ fv t" "Fun c [] \ \ x" using assms by (induct t) auto lemma const_subterm_subst_cases: assumes "Fun c [] \ t \ \" shows "Fun c [] \ t \ (\x \ fv t. x \ subst_domain \ \ Fun c [] \ \ x)" proof (cases "Fun c [] \ t") case False then obtain x where "x \ fv t" "Fun c [] \ \ x" using const_subterm_subst_var_obtain[OF assms] by moura thus ?thesis by (cases "x \ subst_domain \") auto qed simp lemma const_subterms_subst_cases: assumes "Fun c [] \\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" shows "Fun c [] \\<^sub>s\<^sub>e\<^sub>t M \ (\x \ fv\<^sub>s\<^sub>e\<^sub>t M. x \ subst_domain \ \ Fun c [] \ \ x)" using assms const_subterm_subst_cases[of c _ \] by auto lemma const_subterms_subst_cases': assumes "Fun c [] \\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" shows "Fun c [] \\<^sub>s\<^sub>e\<^sub>t M \ Fun c [] \\<^sub>s\<^sub>e\<^sub>t subst_range \" using const_subterms_subst_cases[OF assms] by auto lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_fv_subset: assumes "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" shows "fv (\ x) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using assms proof (induction F) case (Cons f F) then obtain t t' where f: "f = (t,t')" by (metis surj_pair) show ?case proof (cases "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F") case True thus ?thesis using Cons.IH unfolding subst_apply_pairs_def by auto next case False hence "x \ fv t \ fv t'" using Cons.prems f by simp hence "fv (\ x) \ fv (t \ \) \ fv (t' \ \)" using fv_subst_subset[of x] by force thus ?thesis using f unfolding subst_apply_pairs_def by auto qed qed simp lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_step_subst: "fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" proof (induction F) case (Cons f F) obtain t t' where "f = (t,t')" by moura thus ?case using Cons by (simp add: subst_apply_pairs_def subst_apply_fv_unfold) qed (simp_all add: subst_apply_pairs_def) lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_obtain_var: fixes \::"('a,'b) subst" assumes "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" shows "\y \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F. x \ fv (\ y)" using assms proof (induction F) case (Cons f F) then obtain t s where f: "f = (t,s)" by (metis surj_pair) from Cons.IH show ?case proof (cases "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)") case False hence "x \ fv (t \ \) \ x \ fv (s \ \)" using f Cons.prems by (simp add: subst_apply_pairs_def) hence "(\y \ fv t. x \ fv (\ y)) \ (\y \ fv s. x \ fv (\ y))" by (metis fv_subst_obtain_var) thus ?thesis using f by (auto simp add: subst_apply_pairs_def) qed (auto simp add: Cons.IH) qed (simp add: subst_apply_pairs_def) lemma pair_subst_ident[intro]: "(fv t \ fv t') \ subst_domain \ = {} \ (t,t') \\<^sub>p \ = (t,t')" by auto lemma pairs_substI[intro]: assumes "subst_domain \ \ (\(s,t) \ M. fv s \ fv t) = {}" shows "M \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ = M" proof - { fix m assume M: "m \ M" then obtain s t where m: "m = (s,t)" by (metis surj_pair) hence "(fv s \ fv t) \ subst_domain \ = {}" using assms M by auto hence "m \\<^sub>p \ = m" using m by auto } thus ?thesis by (simp add: image_cong) qed lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F))" proof (induction F) case (Cons g G) obtain t t' where "g = (t,t')" by (metis surj_pair) thus ?case using Cons.IH by (simp add: subst_apply_pairs_def subst_apply_fv_unfold) qed (simp add: subst_apply_pairs_def) lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_subset: assumes "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \ subst_domain \" shows "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ subst_domain \ \ subst_domain \" using assms proof (induction F) case (Cons g G) hence IH: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G \ subst_domain \ \ subst_domain \" by (simp add: subst_apply_pairs_def) obtain t t' where g: "g = (t,t')" by (metis surj_pair) hence "fv (t \ \) \ subst_domain \" "fv (t' \ \) \ subst_domain \" using Cons.prems by (simp_all add: subst_apply_pairs_def) hence "fv t \ subst_domain \ \ subst_domain \" "fv t' \ subst_domain \ \ subst_domain \" - using subst_apply_fv_unfold[of _ \] by force+ + unfolding subst_apply_fv_unfold by force+ thus ?case using IH g by (simp add: subst_apply_pairs_def) qed (simp add: subst_apply_pairs_def) lemma pairs_subst_comp: "F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \ \\<^sub>s \ = ((F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" by (induct F) (auto simp add: subst_apply_pairs_def) lemma pairs_substI'[intro]: "subst_domain \ \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F = {} \ F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \ = F" by (induct F) (force simp add: subst_apply_pairs_def)+ lemma subst_pair_compose[simp]: "d \\<^sub>p (\ \\<^sub>s \) = d \\<^sub>p \ \\<^sub>p \" proof - obtain t s where "d = (t,s)" by moura thus ?thesis by auto qed lemma subst_pairs_compose[simp]: "D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t (\ \\<^sub>s \) = D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" by auto lemma subst_apply_pair_pair: "(t, s) \\<^sub>p \ = (t \ \, s \ \)" by (rule prod.case) lemma subst_apply_pairs_nil[simp]: "[] \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \ = []" unfolding subst_apply_pairs_def by simp lemma subst_apply_pairs_singleton[simp]: "[(t,s)] \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \ = [(t \ \,s \ \)]" unfolding subst_apply_pairs_def by simp lemma subst_apply_pairs_Var[iff]: "F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s Var = F" by (simp add: subst_apply_pairs_def) lemma subst_apply_pairs_pset_subst: "set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) = set F \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" unfolding subst_apply_pairs_def by force lemma subst_subterms: "t \\<^sub>s\<^sub>e\<^sub>t M \ t \ \ \\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" using subst_mono_neq by fastforce lemma subst_subterms_fv: "x \ fv\<^sub>s\<^sub>e\<^sub>t M \ \ x \ subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" using fv_subterms_substI by fastforce lemma subst_subterms_Var: "Var x \\<^sub>s\<^sub>e\<^sub>t M \ \ x \ subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" using subst_subterms_fv[of x M \] by force lemma fv_subset_subterms_subset: "\ ` fv\<^sub>s\<^sub>e\<^sub>t M \ subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" using subst_subterms_fv by fast lemma subst_const_swap_eq: fixes \ \::"('a,'b) subst" assumes t: "t \ \ = s \ \" and \: "\x \ fv t \ fv s. \k. \ x = Fun k []" "\x \ fv t. \(\ x \ s)" "\x \ fv s. \(\ x \ t)" and \_def: "\ \ \x. p (\ x)" shows "t \ \ = s \ \" using t \ proof (induction t arbitrary: s) case (Var x) thus ?case unfolding \_def by (cases s) auto next case (Fun f ts) note prems = Fun.prems obtain ss where s: "s = Fun f ss" and ss: "ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = ss \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using prems by (cases s) auto have "ts ! i \ \ = ss ! i \ \" when i: "i < length ts" for i proof - have *: "ts ! i \ set ts" using i by simp have **: "ts ! i \ \ = ss ! i \ \" using i prems(1) unfolding s by (metis subst_Fun_index_eq) have ***: "ss ! i \ set ss" using i ss by (metis length_map nth_mem) show ?thesis using Fun.IH[OF * **] prems(2,3,4) * *** unfolding s by auto qed hence IH: "ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = ss \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using ss by (metis (mono_tags, lifting) length_map nth_equalityI nth_map) show ?case using IH unfolding s by auto qed lemma term_subst_set_eq: assumes "\x. x \ fv\<^sub>s\<^sub>e\<^sub>t M \ \ x = \ x" shows "M \\<^sub>s\<^sub>e\<^sub>t \ = M \\<^sub>s\<^sub>e\<^sub>t \" proof - have "t \ \ = t \ \" when "t \ M" for t using that assms term_subst_eq[of _ \ \] by fastforce thus ?thesis by simp qed lemma subst_const_swap_eq': assumes "t \ \ = s \ \" and "\x \ fv t \ fv s. \ x = \ x \ \(\ x \ t) \ \(\ x \ s)" (is "?A t s") and "\x \ fv t \ fv s. \c. \ x = Fun c []" (is "?B t s") and "\x \ fv t \ fv s. \c. \ x = Fun c []" (is "?C t s") and "\x \ fv t \ fv s. \y \ fv t \ fv s. \ x = \ y \ \ x = \ y" (is "?D t s") shows "t \ \ = s \ \" using assms proof (induction t arbitrary: s) case (Var x) note prems = Var.prems have "(\y. s = Var y) \ (\c. s = Fun c [])" using prems(1,3) by (cases s) auto thus ?case proof assume "\y. s = Var y" then obtain y where y: "s = Var y" by blast hence "\ x = \ y" using prems(1) by simp hence "\ x = \ y" using prems(5) y by fastforce thus ?thesis using y by force next assume "\c. s = Fun c []" then obtain c where c: "s = Fun c []" by blast have "\ x = \ x \ \(\ x \ Fun c [])" using prems(2) c by auto thus ?thesis using prems(1) c by simp qed next case (Fun f ts) note prems = Fun.prems note IH = Fun.IH show ?case proof (cases s) case (Var x) note s = this hence ts: "ts = []" using prems(1,3) by auto show ?thesis using prems unfolding s ts by auto next case (Fun g ss) note s = this hence g: "f = g" using prems(1) by fastforce have ss: "ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = ss \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using prems(1) unfolding s by (cases s) auto have len: "length ts = length ss" using ss by (metis length_map) have "ts ! i \ \ = ss ! i \ \" when i: "i < length ts" for i proof - have 0: "ts ! i \ set ts" using i by simp have 1: "ts ! i \ \ = ss ! i \ \" using i prems(1) unfolding s by (metis subst_Fun_index_eq) have 2: "ss ! i \ set ss" using i by (metis len nth_mem) have 3: "fv (ts ! i) \ fv (Fun f ts)" "fv (ss ! i) \ fv (Fun g ss)" "subterms (ts ! i) \ subterms (Fun f ts)" "subterms (ss ! i) \ subterms (Fun g ss)" subgoal by (meson index_Fun_fv_subset i) subgoal by (metis index_Fun_fv_subset i len) subgoal using ss i by fastforce subgoal using ss i len by fastforce done have 4: "?A (ts ! i) (ss ! i)" "?B (ts ! i) (ss ! i)" "?C (ts ! i) (ss ! i)" "?D (ts ! i) (ss ! i)" subgoal using 3 prems(2) unfolding s by blast subgoal using 3(1,2) prems(3) unfolding s by blast subgoal using 3(1,2) prems(4) unfolding s by blast subgoal using 3(1,2) prems(5) unfolding s by blast done thus ?thesis using IH[OF 0 1 4] prems(2-) 0 2 unfolding s by blast qed hence "ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = ss \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" by (metis (mono_tags, lifting) ss length_map nth_equalityI nth_map) thus ?thesis unfolding s g by auto qed qed lemma subst_const_swap_eq_mem: assumes "t \ \ \ M \\<^sub>s\<^sub>e\<^sub>t \" and "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. \ x = \ x \ \(\ x \\<^sub>s\<^sub>e\<^sub>t insert t M)" and "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. \c. \ x = Fun c []" (is "?B (fv\<^sub>s\<^sub>e\<^sub>t M \ fv t)") and "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. \c. \ x = Fun c []" (is "?C (fv\<^sub>s\<^sub>e\<^sub>t M \ fv t)") and "\x \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. \y \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv t. \ x = \ y \ \ x = \ y" (is "?D (fv\<^sub>s\<^sub>e\<^sub>t M \ fv t)") shows "t \ \ \ M \\<^sub>s\<^sub>e\<^sub>t \" proof - let ?A = "\t s. \x \ fv t \ fv s. \ x = \ x \ \(\ x \ t) \ \(\ x \ s)" obtain s where s: "s \ M" "s \ \ = t \ \" using assms(1) by fastforce have 0: "fv s \ fv\<^sub>s\<^sub>e\<^sub>t M" "subterms s \ subterms\<^sub>s\<^sub>e\<^sub>t (insert t M)" "subterms t \ subterms\<^sub>s\<^sub>e\<^sub>t (insert t M)" using s(1) by auto have 1: "?A s t" "?B (fv s \ fv t)" "?C (fv s \ fv t)" "?D (fv s \ fv t)" subgoal using assms(2) 0 by fast subgoal using assms(3) 0 by blast subgoal using assms(4) 0 by blast subgoal using assms(5) 0 by blast done have "s \ \ = t \ \" by (rule subst_const_swap_eq'[OF s(2) 1]) thus ?thesis by (metis s(1) imageI) qed subsection \Finite Substitutions\ inductive_set fsubst::"('a,'b) subst set" where fvar: "Var \ fsubst" | FUpdate: "\\ \ fsubst; v \ subst_domain \; t \ Var v\ \ \(v := t) \ fsubst" lemma finite_dom_iff_fsubst: "finite (subst_domain \) \ \ \ fsubst" proof assume "finite (subst_domain \)" thus "\ \ fsubst" proof (induction "subst_domain \" arbitrary: \ rule: finite.induct) case emptyI hence "\ = Var" using empty_dom_iff_empty_subst by metis thus ?case using fvar by simp next case (insertI \'\<^sub>d\<^sub>o\<^sub>m v) thus ?case proof (cases "v \ \'\<^sub>d\<^sub>o\<^sub>m") case True hence "\'\<^sub>d\<^sub>o\<^sub>m = subst_domain \" using \insert v \'\<^sub>d\<^sub>o\<^sub>m = subst_domain \\ by auto thus ?thesis using insertI.hyps(2) by metis next case False let ?\' = "\w. if w \ \'\<^sub>d\<^sub>o\<^sub>m then \ w else Var w" have "subst_domain ?\' = \'\<^sub>d\<^sub>o\<^sub>m" using \v \ \'\<^sub>d\<^sub>o\<^sub>m\ \insert v \'\<^sub>d\<^sub>o\<^sub>m = subst_domain \\ by (auto simp add: subst_domain_def) hence "?\' \ fsubst" using insertI.hyps(2) by simp moreover have "?\'(v := \ v) = (\w. if w \ insert v \'\<^sub>d\<^sub>o\<^sub>m then \ w else Var w)" by auto hence "?\'(v := \ v) = \" using \insert v \'\<^sub>d\<^sub>o\<^sub>m = subst_domain \\ by (auto simp add: subst_domain_def) ultimately show ?thesis using FUpdate[of ?\' v "\ v"] False insertI.hyps(3) by (auto simp add: subst_domain_def) qed qed next assume "\ \ fsubst" thus "finite (subst_domain \)" by (induct \, simp, metis subst_dom_insert_finite) qed lemma fsubst_induct[case_names fvar FUpdate, induct set: finite]: assumes "finite (subst_domain \)" "P Var" and "\\ v t. \finite (subst_domain \); v \ subst_domain \; t \ Var v; P \\ \ P (\(v := t))" shows "P \" using assms finite_dom_iff_fsubst fsubst.induct by metis lemma fun_upd_fsubst: "s(v := t) \ fsubst \ s \ fsubst" using subst_dom_insert_finite[of s] finite_dom_iff_fsubst by blast lemma finite_img_if_fsubst: "s \ fsubst \ finite (subst_range s)" using finite_dom_iff_fsubst finite_subst_img_if_finite_dom' by blast subsection \Unifiers and Most General Unifiers (MGUs)\ abbreviation Unifier::"('f,'v) subst \ ('f,'v) term \ ('f,'v) term \ bool" where "Unifier \ t u \ (t \ \ = u \ \)" abbreviation MGU::"('f,'v) subst \ ('f,'v) term \ ('f,'v) term \ bool" where "MGU \ t u \ Unifier \ t u \ (\\. Unifier \ t u \ \ \\<^sub>\ \)" lemma MGUI[intro]: shows "\t \ \ = u \ \; \\::('f,'v) subst. t \ \ = u \ \ \ \ \\<^sub>\ \\ \ MGU \ t u" by auto lemma UnifierD[dest]: fixes \::"('f,'v) subst" and f g::'f and X Y::"('f,'v) term list" assumes "Unifier \ (Fun f X) (Fun g Y)" shows "f = g" "length X = length Y" proof - from assms show "f = g" by auto from assms have "Fun f X \ \ = Fun g Y \ \" by auto hence "length (map (\x. x \ \) X) = length (map (\x. x \ \) Y)" by auto thus "length X = length Y" by auto qed lemma MGUD[dest]: fixes \::"('f,'v) subst" and f g::'f and X Y::"('f,'v) term list" assumes "MGU \ (Fun f X) (Fun g Y)" shows "f = g" "length X = length Y" -using assms by (auto intro!: UnifierD[of f X \ g Y]) +using assms by (auto dest: map_eq_imp_length_eq) lemma MGU_sym[sym]: "MGU \ s t \ MGU \ t s" by auto lemma Unifier_sym[sym]: "Unifier \ s t \ Unifier \ t s" by auto lemma MGU_nil: "MGU Var s t \ s = t" by fastforce lemma Unifier_comp: "Unifier (\ \\<^sub>s \) t u \ Unifier \ (t \ \) (u \ \)" by simp lemma Unifier_comp': "Unifier \ (t \ \) (u \ \) \ Unifier (\ \\<^sub>s \) t u" by simp lemma Unifier_excludes_subterm: assumes \: "Unifier \ t u" shows "\t \ u" proof assume "t \ u" hence "t \ \ \ u \ \" using subst_mono_neq by metis hence "t \ \ \ u \ \" by simp moreover from \ have "t \ \ = u \ \" by auto ultimately show False .. qed lemma MGU_is_Unifier: "MGU \ t u \ Unifier \ t u" by (rule conjunct1) lemma MGU_Var1: assumes "\Var v \ t" shows "MGU (Var(v := t)) (Var v) t" proof (intro MGUI exI) show "Var v \ (Var(v := t)) = t \ (Var(v := t))" using assms subst_no_occs by fastforce next fix \::"('a,'b) subst" assume th: "Var v \ \ = t \ \" show "\ = (Var(v := t)) \\<^sub>s \" - proof - fix s show "s \ \ = s \ ((Var(v := t)) \\<^sub>s \)" using th by (induct s) auto - qed + using th by (auto simp: subst_compose_def) qed lemma MGU_Var2: "v \ fv t \ MGU (Var(v := t)) (Var v) t" by (metis (no_types) MGU_Var1 vars_iff_subterm_or_eq) lemma MGU_Var3: "MGU Var (Var v) (Var w) \ v = w" by fastforce lemma MGU_Const1: "MGU Var (Fun c []) (Fun d []) \ c = d" by fastforce lemma MGU_Const2: "MGU \ (Fun c []) (Fun d []) \ c = d" by auto lemma MGU_Fun: assumes "MGU \ (Fun f X) (Fun g Y)" shows "f = g" "length X = length Y" proof - let ?F = "\\ X. map (\x. x \ \) X" from assms have "\f = g; ?F \ X = ?F \ Y; \\'. f = g \ ?F \' X = ?F \' Y \ \ \\<^sub>\ \'\ \ length X = length Y" using map_eq_imp_length_eq by auto thus "f = g" "length X = length Y" using assms by auto qed lemma Unifier_Fun: assumes "Unifier \ (Fun f (x#X)) (Fun g (y#Y))" shows "Unifier \ x y" "Unifier \ (Fun f X) (Fun g Y)" using assms by simp_all lemma Unifier_subst_idem_subst: "subst_idem r \ Unifier s (t \ r) (u \ r) \ Unifier (r \\<^sub>s s) (t \ r) (u \ r)" by (metis (no_types, lifting) subst_idem_def subst_subst_compose) lemma subst_idem_comp: "subst_idem r \ Unifier s (t \ r) (u \ r) \ (\q. Unifier q (t \ r) (u \ r) \ s \\<^sub>s q = q) \ subst_idem (r \\<^sub>s s)" by (frule Unifier_subst_idem_subst, blast, metis subst_idem_def subst_compose_assoc) lemma Unifier_mgt: "\Unifier \ t u; \ \\<^sub>\ \\ \ Unifier \ t u" by auto lemma Unifier_support: "\Unifier \ t u; \ supports \\ \ Unifier \ t u" using subst_supportD Unifier_mgt by metis lemma MGU_mgt: "\MGU \ t u; MGU \ t u\ \ \ \\<^sub>\ \" by auto lemma Unifier_trm_fv_bound: "\Unifier s t u; v \ fv t\ \ v \ subst_domain s \ range_vars s \ fv u" proof (induction t arbitrary: s u) case (Fun f X) hence "v \ fv (u \ s) \ v \ subst_domain s" by (metis subst_not_dom_fixed) thus ?case by (metis (no_types) Un_iff contra_subsetD subst_sends_fv_to_img) qed (metis (no_types) UnI1 UnI2 subsetCE no_var_subterm subst_sends_dom_to_img subst_to_var_is_var trm_subst_ident' vars_iff_subterm_or_eq) lemma Unifier_rm_var: "\Unifier \ s t; v \ fv s \ fv t\ \ Unifier (rm_var v \) s t" by (auto simp add: repl_invariance) lemma Unifier_ground_rm_vars: assumes "ground (subst_range s)" "Unifier (rm_vars X s) t t'" shows "Unifier s t t'" by (rule Unifier_support[OF assms(2) rm_vars_ground_supports[OF assms(1)]]) lemma Unifier_dom_restrict: assumes "Unifier s t t'" "fv t \ fv t' \ S" shows "Unifier (rm_vars (UNIV - S) s) t t'" proof - let ?s = "rm_vars (UNIV - S) s" show ?thesis using term_subst_eq_conv[of t s ?s] term_subst_eq_conv[of t' s ?s] assms by auto qed subsection \Well-formedness of Substitutions and Unifiers\ inductive_set wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set::"('a,'b) subst set" where Empty[simp]: "Var \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set" | Insert[simp]: "\\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set; v \ subst_domain \; v \ range_vars \; fv t \ (insert v (subst_domain \)) = {}\ \ \(v := t) \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set" definition wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t::"('a,'b) subst \ bool" where "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ subst_domain \ \ range_vars \ = {} \ finite (subst_domain \)" definition wf\<^sub>M\<^sub>G\<^sub>U::"('a,'b) subst \ ('a,'b) term \ ('a,'b) term \ bool" where "wf\<^sub>M\<^sub>G\<^sub>U \ s t \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ MGU \ s t \ subst_domain \ \ range_vars \ \ fv s \ fv t" lemma wf_subst_subst_idem: "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ subst_idem \" using subst_idemI[of \] unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by fast lemma wf_subst_properties: "\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set = wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" proof show "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ \ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set" unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def proof - assume "subst_domain \ \ range_vars \ = {} \ finite (subst_domain \)" hence "finite (subst_domain \)" "subst_domain \ \ range_vars \ = {}" by auto thus "\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set" proof (induction \ rule: fsubst_induct) case fvar thus ?case by simp next case (FUpdate \ v t) have "subst_domain \ \ subst_domain (\(v := t))" "range_vars \ \ range_vars (\(v := t))" using FUpdate.hyps(2,3) subst_img_update unfolding range_vars_alt_def by (fastforce simp add: subst_domain_def)+ hence "subst_domain \ \ range_vars \ = {}" using FUpdate.prems(1) by blast hence "\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set" using FUpdate.IH by metis have *: "range_vars (\(v := t)) = range_vars \ \ fv t" using FUpdate.hyps(2) subst_img_update[OF _ FUpdate.hyps(3)] by fastforce hence "fv t \ insert v (subst_domain \) = {}" using FUpdate.prems subst_dom_update2[OF FUpdate.hyps(3)] by blast moreover have "subst_domain (\(v := t)) = insert v (subst_domain \)" by (meson FUpdate.hyps(3) subst_dom_update2) hence "v \ range_vars \" using FUpdate.prems * by blast ultimately show ?case using Insert[OF \\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set\ \v \ subst_domain \\] by metis qed qed show "\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def proof (induction \ rule: wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set.induct) case Empty thus ?case by simp next case (Insert \ v t) hence 1: "subst_domain \ \ range_vars \ = {}" by simp hence 2: "subst_domain (\(v := t)) \ range_vars \ = {}" using Insert.hyps(3) by (auto simp add: subst_domain_def) have 3: "fv t \ subst_domain (\(v := t)) = {}" using Insert.hyps(4) by (auto simp add: subst_domain_def) have 4: "\ v = Var v" using \v \ subst_domain \\ by (simp add: subst_domain_def) from Insert.IH have "finite (subst_domain \)" by simp hence 5: "finite (subst_domain (\(v := t)))" using subst_dom_insert_finite[of \] by simp have "subst_domain (\(v := t)) \ range_vars (\(v := t)) = {}" proof (cases "t = Var v") case True hence "range_vars (\(v := t)) = range_vars \" using 4 fun_upd_triv term.inject(1) unfolding range_vars_alt_def by (auto simp add: subst_domain_def) thus "subst_domain (\(v := t)) \ range_vars (\(v := t)) = {}" using 1 2 3 by auto next case False hence "range_vars (\(v := t)) = fv t \ (range_vars \)" using 4 subst_img_update[of \ v] by auto thus "subst_domain (\(v := t)) \ range_vars (\(v := t)) = {}" using 1 2 3 by blast qed thus ?case using 5 by blast qed qed lemma wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_induct[consumes 1, case_names Empty Insert]: assumes "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "P Var" and "\\ v t. \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \; P \; v \ subst_domain \; v \ range_vars \; fv t \ insert v (subst_domain \) = {}\ \ P (\(v := t))" shows "P \" proof - from assms(1,3) wf_subst_properties have "\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set" "\\ v t. \\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set; P \; v \ subst_domain \; v \ range_vars \; fv t \ insert v (subst_domain \) = {}\ \ P (\(v := t))" by blast+ thus "P \" using wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set.induct assms(2) by blast qed lemma wf_subst_fsubst: "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ \ \ fsubst" unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def using finite_dom_iff_fsubst by blast lemma wf_subst_nil: "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by simp lemma wf_MGU_nil: "MGU Var s t \ wf\<^sub>M\<^sub>G\<^sub>U Var s t" using wf_subst_nil subst_domain_Var range_vars_Var unfolding wf\<^sub>M\<^sub>G\<^sub>U_def by fast lemma wf_MGU_dom_bound: "wf\<^sub>M\<^sub>G\<^sub>U \ s t \ subst_domain \ \ fv s \ fv t" unfolding wf\<^sub>M\<^sub>G\<^sub>U_def by blast lemma wf_subst_single: assumes "v \ fv t" "\ v = t" "\w. v \ w \ \ w = Var w" shows "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" proof - have *: "subst_domain \ = {v}" by (metis subst_fv_dom_img_single(1)[OF assms]) have "subst_domain \ \ range_vars \ = {}" using * assms subst_fv_dom_img_single(2) by (metis inf_bot_left insert_disjoint(1)) moreover have "finite (subst_domain \)" using * by simp ultimately show ?thesis by (metis wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) qed lemma wf_subst_reduction: "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t s \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (rm_var v s)" proof - assume "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t s" moreover have "subst_domain (rm_var v s) \ subst_domain s" by (auto simp add: subst_domain_def) moreover have "range_vars (rm_var v s) \ range_vars s" unfolding range_vars_alt_def by (auto simp add: subst_domain_def) ultimately have "subst_domain (rm_var v s) \ range_vars (rm_var v s) = {}" by (meson compl_le_compl_iff disjoint_eq_subset_Compl subset_trans wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) moreover have "finite (subst_domain (rm_var v s))" using \subst_domain (rm_var v s) \ subst_domain s\ \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t s\ rev_finite_subset unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by blast ultimately show "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (rm_var v s)" by (metis wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) qed lemma wf_subst_compose: assumes "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \1" "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \2" and "subst_domain \1 \ subst_domain \2 = {}" and "subst_domain \1 \ range_vars \2 = {}" shows "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\1 \\<^sub>s \2)" using assms proof (induction \1 rule: wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_induct) case Empty thus ?case unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by simp next case (Insert \1 v t) have "t \ Var v" using Insert.hyps(4) by auto hence dom1v_unfold: "subst_domain (\1(v := t)) = insert v (subst_domain \1)" using subst_dom_update2 by metis hence doms_disj: "subst_domain \1 \ subst_domain \2 = {}" using Insert.prems(2) disjoint_insert(1) by blast moreover have dom_img_disj: "subst_domain \1 \ range_vars \2 = {}" using Insert.hyps(2) Insert.prems(3) by (fastforce simp add: subst_domain_def) ultimately have "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\1 \\<^sub>s \2)" using Insert.IH[OF \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \2\] by metis have dom_comp_is_union: "subst_domain (\1 \\<^sub>s \2) = subst_domain \1 \ subst_domain \2" using subst_dom_comp_eq[OF dom_img_disj] . have "v \ subst_domain \2" using Insert.prems(2) \t \ Var v\ by (fastforce simp add: subst_domain_def) hence "\2 v = Var v" "\1 v = Var v" using Insert.hyps(2) by (simp_all add: subst_domain_def) hence "(\1 \\<^sub>s \2) v = Var v" "(\1(v := t) \\<^sub>s \2) v = t \ \2" "((\1 \\<^sub>s \2)(v := t)) v = t" unfolding subst_compose_def by simp_all have fv_t2_bound: "fv (t \ \2) \ fv t \ range_vars \2" by (meson subst_sends_fv_to_img) have 1: "v \ subst_domain (\1 \\<^sub>s \2)" using \(\1 \\<^sub>s \2) v = Var v\ by (auto simp add: subst_domain_def) have "insert v (subst_domain \1) \ range_vars \2 = {}" using Insert.prems(3) dom1v_unfold by blast hence "v \ range_vars \1 \ range_vars \2" using Insert.hyps(3) by blast hence 2: "v \ range_vars (\1 \\<^sub>s \2)" by (meson set_rev_mp subst_img_comp_subset) have "subst_domain \2 \ range_vars \2 = {}" using \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \2\ unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by simp hence "fv (t \ \2) \ subst_domain \2 = {}" using subst_dom_elim unfolding range_vars_alt_def by simp moreover have "v \ range_vars \2" using Insert.prems(3) dom1v_unfold by blast hence "v \ fv t \ range_vars \2" using Insert.hyps(4) by blast hence "v \ fv (t \ \2)" using \fv (t \ \2) \ fv t \ range_vars \2\ by blast moreover have "fv (t \ \2) \ subst_domain \1 = {}" using dom_img_disj fv_t2_bound \fv t \ insert v (subst_domain \1) = {}\ by blast ultimately have 3: "fv (t \ \2) \ insert v (subst_domain (\1 \\<^sub>s \2)) = {}" using dom_comp_is_union by blast have "\1(v := t) \\<^sub>s \2 = (\1 \\<^sub>s \2)(v := t \ \2)" using subst_comp_upd1[of \1 v t \2] . moreover have "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t ((\1 \\<^sub>s \2)(v := t \ \2))" using "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set.Insert"[OF _ 1 2 3] \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\1 \\<^sub>s \2)\ wf_subst_properties by metis ultimately show ?case by presburger qed lemma wf_subst_append: fixes \1 \2::"('f,'v) subst" assumes "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \1" "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \2" and "subst_domain \1 \ subst_domain \2 = {}" and "subst_domain \1 \ range_vars \2 = {}" and "range_vars \1 \ subst_domain \2 = {}" shows "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\v. if \1 v = Var v then \2 v else \1 v)" using assms proof (induction \1 rule: wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_induct) case Empty thus ?case unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by simp next case (Insert \1 v t) let ?if = "\w. if \1 w = Var w then \2 w else \1 w" let ?if_upd = "\w. if (\1(v := t)) w = Var w then \2 w else (\1(v := t)) w" from Insert.hyps(4) have "?if_upd = ?if(v := t)" by fastforce have dom_insert: "subst_domain (\1(v := t)) = insert v (subst_domain \1)" using Insert.hyps(4) by (auto simp add: subst_domain_def) have "\1 v = Var v" "t \ Var v" using Insert.hyps(2,4) by auto hence img_insert: "range_vars (\1(v := t)) = range_vars \1 \ fv t" using subst_img_update by metis from Insert.prems(2) dom_insert have "subst_domain \1 \ subst_domain \2 = {}" by (auto simp add: subst_domain_def) moreover have "subst_domain \1 \ range_vars \2 = {}" using Insert.prems(3) dom_insert by (simp add: subst_domain_def) moreover have "range_vars \1 \ subst_domain \2 = {}" using Insert.prems(4) img_insert by blast ultimately have "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t ?if" using Insert.IH[OF Insert.prems(1)] by metis have dom_union: "subst_domain ?if = subst_domain \1 \ subst_domain \2" by (auto simp add: subst_domain_def) hence "v \ subst_domain ?if" using Insert.hyps(2) Insert.prems(2) dom_insert by (auto simp add: subst_domain_def) moreover have "v \ range_vars ?if" using Insert.prems(3) Insert.hyps(3) dom_insert unfolding range_vars_alt_def by (auto simp add: subst_domain_def) moreover have "fv t \ insert v (subst_domain ?if) = {}" using Insert.hyps(4) Insert.prems(4) img_insert unfolding range_vars_alt_def by (fastforce simp add: subst_domain_def) ultimately show ?case using wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set.Insert \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t ?if\ \?if_upd = ?if(v := t)\ wf_subst_properties by (metis (no_types, lifting)) qed lemma wf_subst_elim_append: assumes "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "subst_elim \ v" "v \ fv t" shows "subst_elim (\(w := t)) v" using assms proof (induction \ rule: wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_induct) case (Insert \ v' t') hence "\q. v \ fv (Var q \ \(v' := t'))" using subst_elimD by blast hence "\q. v \ fv (Var q \ \(v' := t', w := t))" using \v \ fv t\ by simp - thus ?case by (metis subst_elimI' subst_apply_term.simps(1)) + thus ?case by (metis subst_elimI' eval_term.simps(1)) qed (simp add: subst_elim_def) lemma wf_subst_elim_dom: assumes "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\v \ subst_domain \. subst_elim \ v" using assms proof (induction \ rule: wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_induct) case (Insert \ w t) have dom_insert: "subst_domain (\(w := t)) \ insert w (subst_domain \)" by (auto simp add: subst_domain_def) hence "\v \ subst_domain \. subst_elim (\(w := t)) v" using Insert.IH Insert.hyps(2,4) by (metis Insert.hyps(1) IntI disjoint_insert(2) empty_iff wf_subst_elim_append) moreover have "w \ fv t" using Insert.hyps(4) by simp hence "\q. w \ fv (Var q \ \(w := t))" by (metis fv_simps(1) fv_in_subst_img Insert.hyps(3) contra_subsetD - fun_upd_def singletonD subst_apply_term.simps(1)) + fun_upd_def singletonD eval_term.simps(1)) hence "subst_elim (\(w := t)) w" by (metis subst_elimI') ultimately show ?case using dom_insert by blast qed simp lemma wf_subst_support_iff_mgt: "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ \ supports \ \ \ \\<^sub>\ \" using subst_support_def subst_support_if_mgt_subst_idem wf_subst_subst_idem by blast subsection \Interpretations\ abbreviation interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t::"('a,'b) subst \ bool" where "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ subst_domain \ = UNIV \ ground (subst_range \)" lemma interpretation_substI: "(\v. fv (\ v) = {}) \ interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" proof - assume "\v. fv (\ v) = {}" moreover { fix v assume "fv (\ v) = {}" hence "v \ subst_domain \" by auto } ultimately show ?thesis by auto qed lemma interpretation_grounds[simp]: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ fv (t \ \) = {}" using subst_fv_dom_ground_if_ground_img[of t \] by blast lemma interpretation_grounds_all: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ (\v. fv (\ v) = {})" by (metis range_vars_alt_def UNIV_I fv_in_subst_img subset_empty subst_dom_vars_in_subst) lemma interpretation_grounds_all': "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ ground (M \\<^sub>s\<^sub>e\<^sub>t \)" using subst_fv_dom_ground_if_ground_img[of _ \] by simp lemma interpretation_comp: assumes "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" proof - have \_fv: "fv (\ v) = {}" for v using interpretation_grounds_all[OF assms] by simp hence \_fv': "fv (t \ \) = {}" for t - by (metis all_not_in_conv subst_elimD subst_elimI' subst_apply_term.simps(1)) + by (metis all_not_in_conv subst_elimD subst_elimI' eval_term.simps(1)) from assms have "(\ \\<^sub>s \) v \ Var v" for v unfolding subst_compose_def by (metis fv_simps(1) \_fv' insert_not_empty) hence "subst_domain (\ \\<^sub>s \) = UNIV" by (simp add: subst_domain_def) moreover have "fv ((\ \\<^sub>s \) v) = {}" for v unfolding subst_compose_def using \_fv' by simp hence "ground (subst_range (\ \\<^sub>s \))" by simp ultimately show "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" .. from assms have "(\ \\<^sub>s \) v \ Var v" for v unfolding subst_compose_def by (metis fv_simps(1) \_fv insert_not_empty subst_to_var_is_var) hence "subst_domain (\ \\<^sub>s \) = UNIV" by (simp add: subst_domain_def) moreover have "fv ((\ \\<^sub>s \) v) = {}" for v - unfolding subst_compose_def by (simp add: \_fv trm_subst_ident) + unfolding subst_compose_def by (simp add: \_fv subst_apply_term_ident) hence "ground (subst_range (\ \\<^sub>s \))" by simp ultimately show "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" .. qed lemma interpretation_subst_exists: "\\::('f,'v) subst. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" proof - obtain c::"'f" where "c \ UNIV" by simp then obtain \::"('f,'v) subst" where "\v. \ v = Fun c []" by simp hence "subst_domain \ = UNIV" "ground (subst_range \)" by (simp_all add: subst_domain_def) thus ?thesis by auto qed lemma interpretation_subst_exists': "\\::('f,'v) subst. subst_domain \ = X \ ground (subst_range \)" proof - obtain \::"('f,'v) subst" where \: "subst_domain \ = UNIV" "ground (subst_range \)" using interpretation_subst_exists by moura let ?\ = "rm_vars (UNIV - X) \" have 1: "subst_domain ?\ = X" using \ by (auto simp add: subst_domain_def) hence 2: "ground (subst_range ?\)" using \ by force show ?thesis using 1 2 by blast qed lemma interpretation_subst_idem: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ subst_idem \" unfolding subst_idem_def -using interpretation_grounds_all[of \] trm_subst_ident subst_eq_if_eq_vars -by fastforce + using interpretation_grounds_all[of \] subst_apply_term_ident subst_eq_if_eq_vars +by (fastforce simp: subst_compose) lemma subst_idem_comp_upd_eq: assumes "v \ subst_domain \" "subst_idem \" shows "\ \\<^sub>s \ = \(v := \ v) \\<^sub>s \" proof - from assms(1) have "(\ \\<^sub>s \) v = \ v" unfolding subst_compose_def by auto moreover have "\w. w \ v \ (\ \\<^sub>s \) w = (\(v := \ v) \\<^sub>s \) w" unfolding subst_compose_def by auto moreover have "(\(v := \ v) \\<^sub>s \) v = \ v" using assms(2) unfolding subst_idem_def subst_compose_def by (metis fun_upd_same) ultimately show ?thesis by (metis fun_upd_same fun_upd_triv subst_comp_upd1) qed lemma interpretation_dom_img_disjoint: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ subst_domain \ \ range_vars \ = {}" unfolding range_vars_alt_def by auto subsection \Basic Properties of MGUs\ lemma MGU_is_mgu_singleton: "MGU \ t u = is_mgu \ {(t,u)}" unfolding is_mgu_def unifiers_def by auto lemma Unifier_in_unifiers_singleton: "Unifier \ s t \ \ \ unifiers {(s,t)}" unfolding unifiers_def by auto lemma subst_list_singleton_fv_subset: "(\x \ set (subst_list (subst v t) E). fv (fst x) \ fv (snd x)) \ fv t \ (\x \ set E. fv (fst x) \ fv (snd x))" proof (induction E) case (Cons x E) let ?fvs = "\L. \x \ set L. fv (fst x) \ fv (snd x)" let ?fvx = "fv (fst x) \ fv (snd x)" let ?fvxsubst = "fv (fst x \ Var(v := t)) \ fv (snd x \ Var(v := t))" have "?fvs (subst_list (subst v t) (x#E)) = ?fvxsubst \ ?fvs (subst_list (subst v t) E)" unfolding subst_list_def subst_def by auto hence "?fvs (subst_list (subst v t) (x#E)) \ ?fvxsubst \ fv t \ ?fvs E" using Cons.IH by blast moreover have "?fvs (x#E) = ?fvx \ ?fvs E" by auto moreover have "?fvxsubst \ ?fvx \ fv t" using subst_fv_bound_singleton[of _ v t] by blast ultimately show ?case unfolding range_vars_alt_def by auto qed (simp add: subst_list_def) lemma subst_of_dom_subset: "subst_domain (subst_of L) \ set (map fst L)" proof (induction L rule: List.rev_induct) case (snoc x L) then obtain v t where x: "x = (v,t)" by (metis surj_pair) hence "subst_of (L@[x]) = Var(v := t) \\<^sub>s subst_of L" - unfolding subst_of_def subst_def by (induct L) auto + unfolding subst_of_def subst_def by (induct L) (auto simp: subst_compose) hence "subst_domain (subst_of (L@[x])) \ insert v (subst_domain (subst_of L))" using x subst_domain_compose[of "Var(v := t)" "subst_of L"] by (auto simp add: subst_domain_def) thus ?case using snoc.IH x by auto qed simp lemma wf_MGU_is_imgu_singleton: "wf\<^sub>M\<^sub>G\<^sub>U \ s t \ is_imgu \ {(s,t)}" proof - assume 1: "wf\<^sub>M\<^sub>G\<^sub>U \ s t" have 2: "subst_idem \" by (metis wf_subst_subst_idem 1 wf\<^sub>M\<^sub>G\<^sub>U_def) have 3: "\\' \ unifiers {(s,t)}. \ \\<^sub>\ \'" "\ \ unifiers {(s,t)}" by (metis 1 Unifier_in_unifiers_singleton wf\<^sub>M\<^sub>G\<^sub>U_def)+ have "\\ \ unifiers {(s,t)}. \ = \ \\<^sub>s \" by (metis 2 3 subst_idem_def subst_compose_assoc) thus "is_imgu \ {(s,t)}" by (metis is_imgu_def \\ \ unifiers {(s,t)}\) qed lemmas mgu_subst_range_vars = mgu_range_vars lemmas mgu_same_empty = mgu_same lemma mgu_var: assumes "x \ fv t" shows "mgu (Var x) t = Some (Var(x := t))" proof - have "unify [(Var x,t)] [] = Some [(x,t)]" using assms by (auto simp add: subst_list_def) moreover have "subst_of [(x,t)] = Var(x := t)" unfolding subst_of_def subst_def by simp ultimately show ?thesis by (simp add: mgu_def) qed lemma mgu_gives_wellformed_subst: assumes "mgu s t = Some \" shows "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using mgu_finite_subst_domain[OF assms] mgu_subst_domain_range_vars_disjoint[OF assms] unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by auto lemma mgu_gives_wellformed_MGU: assumes "mgu s t = Some \" shows "wf\<^sub>M\<^sub>G\<^sub>U \ s t" using mgu_subst_domain[OF assms] mgu_sound[OF assms] mgu_subst_range_vars [OF assms] MGU_is_mgu_singleton[of s \ t] is_imgu_imp_is_mgu[of \ "{(s,t)}"] mgu_gives_wellformed_subst[OF assms] unfolding wf\<^sub>M\<^sub>G\<^sub>U_def by blast lemma mgu_gives_subst_idem: "mgu s t = Some \ \ subst_idem \" using mgu_sound[of s t \] unfolding is_imgu_def subst_idem_def by auto lemma mgu_always_unifies: "Unifier \ M N \ \\. mgu M N = Some \" using mgu_complete Unifier_in_unifiers_singleton by blast lemma mgu_gives_MGU: "mgu s t = Some \ \ MGU \ s t" using mgu_sound[of s t \, THEN is_imgu_imp_is_mgu] MGU_is_mgu_singleton by metis lemma mgu_vars_bounded[dest?]: "mgu M N = Some \ \ subst_domain \ \ range_vars \ \ fv M \ fv N" using mgu_gives_wellformed_MGU unfolding wf\<^sub>M\<^sub>G\<^sub>U_def by blast lemma mgu_vars_bounded': assumes \: "mgu M N = Some \" and MN: "fv M = {} \ fv N = {}" shows "subst_domain \ = fv M \ fv N" (is ?A) and "range_vars \ = {}" (is ?B) proof - let ?C = "\t. subst_domain \ = fv t" have 0: "fv N = {} \ subst_domain \ \ fv M" "fv N = {} \ range_vars \ \ fv M" "fv M = {} \ subst_domain \ \ fv N" "fv M = {} \ range_vars \ \ fv N" using mgu_vars_bounded[OF \] by simp_all note 1 = mgu_gives_MGU[OF \] mgu_subst_domain_range_vars_disjoint[OF \] note 2 = subst_fv_imgI[of \] subst_dom_vars_in_subst[of _ \] note 3 = ground_term_subst_domain_fv_subset[of _ \] note 4 = subst_apply_fv_empty[of _ \] have "fv (\ x) = {}" when x: "x \ fv M" and N: "fv N = {}" for x using x N 0(1,2) 1 2[of x] 3[of M] 4[of N] by auto hence "?C M" ?B when N: "fv N = {}" using 0(1,2)[OF N] N by (fastforce, fastforce) moreover have "fv (\ x) = {}" when x: "x \ fv N" and M: "fv M = {}" for x using x M 0(3,4) 1 2[of x] 3[of N] 4[of M] by auto hence "?C N" ?B when M: "fv M = {}" using 0(3,4)[OF M] M by (fastforce, fastforce) ultimately show ?A ?B using MN by auto qed lemma mgu_eliminates[dest?]: assumes "mgu M N = Some \" shows "(\v \ fv M \ fv N. subst_elim \ v) \ \ = Var" (is "?P M N \") proof (cases "\ = Var") case False then obtain v where v: "v \ subst_domain \" by auto hence "v \ fv M \ fv N" using mgu_vars_bounded[OF assms] by blast thus ?thesis using wf_subst_elim_dom[OF mgu_gives_wellformed_subst[OF assms]] v by blast qed simp lemma mgu_eliminates_dom: assumes "mgu x y = Some \" "v \ subst_domain \" shows "subst_elim \ v" using mgu_gives_wellformed_subst[OF assms(1)] unfolding wf\<^sub>M\<^sub>G\<^sub>U_def wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def subst_elim_def by (metis disjoint_iff_not_equal subst_dom_elim assms(2)) lemma unify_list_distinct: assumes "Unification.unify E B = Some U" "distinct (map fst B)" and "(\x \ set E. fv (fst x) \ fv (snd x)) \ set (map fst B) = {}" shows "distinct (map fst U)" using assms proof (induction E B arbitrary: U rule: Unification.unify.induct) case 1 thus ?case by simp next case (2 f X g Y E B U) let ?fvs = "\L. \x \ set L. fv (fst x) \ fv (snd x)" from "2.prems"(1) obtain E' where *: "decompose (Fun f X) (Fun g Y) = Some E'" and [simp]: "f = g" "length X = length Y" "E' = zip X Y" and **: "Unification.unify (E'@E) B = Some U" by (auto split: option.splits) hence "\t t'. (t,t') \ set E' \ fv t \ fv (Fun f X) \ fv t' \ fv (Fun g Y)" by (metis zip_arg_subterm subtermeq_vars_subset) hence "?fvs E' \ fv (Fun f X) \ fv (Fun g Y)" by fastforce moreover have "fv (Fun f X) \ set (map fst B) = {}" "fv (Fun g Y) \ set (map fst B) = {}" using "2.prems"(3) by auto ultimately have "?fvs E' \ set (map fst B) = {}" by blast moreover have "?fvs E \ set (map fst B) = {}" using "2.prems"(3) by auto ultimately have "?fvs (E'@E) \ set (map fst B) = {}" by auto thus ?case using "2.IH"[OF * ** "2.prems"(2)] by metis next case (3 v t E B) let ?fvs = "\L. \x \ set L. fv (fst x) \ fv (snd x)" let ?E' = "subst_list (subst v t) E" from "3.prems"(3) have "v \ set (map fst B)" "fv t \ set (map fst B) = {}" by force+ hence *: "distinct (map fst ((v, t)#B))" using "3.prems"(2) by auto show ?case proof (cases "t = Var v") case True thus ?thesis using "3.prems" "3.IH"(1) by auto next case False hence "v \ fv t" using "3.prems"(1) by auto hence "Unification.unify (subst_list (subst v t) E) ((v, t)#B) = Some U" using \t \ Var v\ "3.prems"(1) by auto moreover have "?fvs ?E' \ set (map fst ((v, t)#B)) = {}" proof - have "v \ ?fvs ?E'" unfolding subst_list_def subst_def by (simp add: \v \ fv t\ subst_remove_var) moreover have "?fvs ?E' \ fv t \ ?fvs E" by (metis subst_list_singleton_fv_subset) hence "?fvs ?E' \ set (map fst B) = {}" using "3.prems"(3) by auto ultimately show ?thesis by auto qed ultimately show ?thesis using "3.IH"(2)[OF \t \ Var v\ \v \ fv t\ _ *] by metis qed next case (4 f X v E B U) let ?fvs = "\L. \x \ set L. fv (fst x) \ fv (snd x)" let ?E' = "subst_list (subst v (Fun f X)) E" have *: "?fvs E \ set (map fst B) = {}" using "4.prems"(3) by auto from "4.prems"(1) have "v \ fv (Fun f X)" by force from "4.prems"(3) have **: "v \ set (map fst B)" "fv (Fun f X) \ set (map fst B) = {}" by force+ hence ***: "distinct (map fst ((v, Fun f X)#B))" using "4.prems"(2) by auto from "4.prems"(3) have ****: "?fvs ?E' \ set (map fst ((v, Fun f X)#B)) = {}" proof - have "v \ ?fvs ?E'" unfolding subst_list_def subst_def using \v \ fv (Fun f X)\ subst_remove_var[of v "Fun f X"] by simp moreover have "?fvs ?E' \ fv (Fun f X) \ ?fvs E" by (metis subst_list_singleton_fv_subset) hence "?fvs ?E' \ set (map fst B) = {}" using * ** by blast ultimately show ?thesis by auto qed have "Unification.unify (subst_list (subst v (Fun f X)) E) ((v, Fun f X) # B) = Some U" using \v \ fv (Fun f X)\ "4.prems"(1) by auto thus ?case using "4.IH"[OF \v \ fv (Fun f X)\ _ *** ****] by metis qed lemma mgu_None_is_subst_neq: fixes s t::"('a,'b) term" and \::"('a,'b) subst" assumes "mgu s t = None" shows "s \ \ \ t \ \" using assms mgu_always_unifies by force lemma mgu_None_if_neq_ground: assumes "t \ t'" "fv t = {}" "fv t' = {}" shows "mgu t t' = None" proof (rule ccontr) assume "mgu t t' \ None" then obtain \ where \: "mgu t t' = Some \" by auto hence "t \ \ = t" "t' \ \ = t'" using assms subst_ground_ident by auto thus False using assms(1) MGU_is_Unifier[OF mgu_gives_MGU[OF \]] by auto qed lemma mgu_None_commutes: "mgu s t = None \ mgu t s = None" + thm mgu_complete[of s t] Unifier_in_unifiers_singleton[of ] using mgu_complete[of s t] Unifier_in_unifiers_singleton[of s _ t] Unifier_sym[of t _ s] Unifier_in_unifiers_singleton[of t _ s] mgu_sound[of t s] unfolding is_imgu_def by fastforce lemma mgu_img_subterm_subst: fixes \::"('f,'v) subst" and s t u::"('f,'v) term" assumes "mgu s t = Some \" "u \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) - range Var" shows "u \ ((subterms s \ subterms t) - range Var) \\<^sub>s\<^sub>e\<^sub>t \" proof - define subterms_tuples::"('f,'v) equation list \ ('f,'v) terms" where subtt_def: "subterms_tuples \ \E. subterms\<^sub>s\<^sub>e\<^sub>t (fst ` set E) \ subterms\<^sub>s\<^sub>e\<^sub>t (snd ` set E)" define subterms_img::"('f,'v) subst \ ('f,'v) terms" where subti_def: "subterms_img \ \d. subterms\<^sub>s\<^sub>e\<^sub>t (subst_range d)" define d where "d \ \v t. subst v t::('f,'v) subst" define V where "V \ range Var::('f,'v) terms" define R where "R \ \d::('f,'v) subst. ((subterms s \ subterms t) - V) \\<^sub>s\<^sub>e\<^sub>t d" define M where "M \ \E d. subterms_tuples E \ subterms_img d" define Q where "Q \ (\E d. M E d - V \ R d - V)" define Q' where "Q' \ (\E d d'. (M E d - V) \\<^sub>s\<^sub>e\<^sub>t d' \ (R d - V) \\<^sub>s\<^sub>e\<^sub>t (d'::('f,'v) subst))" have Q_subst: "Q (subst_list (subst v t') E) (subst_of ((v, t')#B))" when v_fv: "v \ fv t'" and Q_assm: "Q ((Var v, t')#E) (subst_of B)" for v t' E B proof - define E' where "E' \ subst_list (subst v t') E" define B' where "B' \ subst_of ((v, t')#B)" have E': "E' = subst_list (d v t') E" and B': "B' = subst_of B \\<^sub>s d v t'" using subst_of_simps(3)[of "(v, t')"] unfolding subst_def E'_def B'_def d_def by simp_all have vt_img_subt: "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range (d v t')) = subterms t'" and vt_dom: "subst_domain (d v t') = {v}" using v_fv by (auto simp add: subst_domain_def d_def subst_def) have *: "subterms u1 \ subterms\<^sub>s\<^sub>e\<^sub>t (fst ` set E)" "subterms u2 \ subterms\<^sub>s\<^sub>e\<^sub>t (snd ` set E)" when "(u1,u2) \ set E" for u1 u2 using that by auto have **: "subterms\<^sub>s\<^sub>e\<^sub>t (d v t' ` (fv u \ subst_domain (d v t'))) \ subterms t'" for u::"('f,'v) term" using vt_dom unfolding d_def by force have 1: "subterms_tuples E' - V \ (subterms t' - V) \ (subterms_tuples E - V \\<^sub>s\<^sub>e\<^sub>t d v t')" (is "?A \ ?B") proof fix u assume "u \ ?A" then obtain u1 u2 where u12: "(u1,u2) \ set E" "u \ (subterms (u1 \ (d v t')) - V) \ (subterms (u2 \ (d v t')) - V)" unfolding subtt_def subst_list_def E'_def d_def by moura hence "u \ (subterms t' - V) \ (((subterms_tuples E) \\<^sub>s\<^sub>e\<^sub>t d v t') - V)" using subterms_subst[of u1 "d v t'"] subterms_subst[of u2 "d v t'"] *[OF u12(1)] **[of u1] **[of u2] unfolding subtt_def subst_list_def by auto moreover have "(subterms_tuples E \\<^sub>s\<^sub>e\<^sub>t d v t') - V \ (subterms_tuples E - V \\<^sub>s\<^sub>e\<^sub>t d v t') \ {t'}" unfolding subst_def subtt_def V_def d_def by force ultimately show "u \ ?B" using u12 v_fv by auto qed have 2: "subterms_img B' - V \ (subterms t' - V) \ (subterms_img (subst_of B) - V \\<^sub>s\<^sub>e\<^sub>t d v t')" using B' vt_img_subt subst_img_comp_subset'''[of "subst_of B" "d v t'"] unfolding subti_def subst_def V_def by argo have 3: "subterms_tuples ((Var v, t')#E) - V = (subterms t' - V) \ (subterms_tuples E - V)" by (auto simp add: subst_def subtt_def V_def) have "fv\<^sub>s\<^sub>e\<^sub>t (subterms t' - V) \ subst_domain (d v t') = {}" using v_fv vt_dom fv_subterms[of t'] by fastforce hence 4: "subterms t' - V \\<^sub>s\<^sub>e\<^sub>t d v t' = subterms t' - V" using set_subst_ident[of "subterms t' - range Var" "d v t'"] by (simp add: V_def) have "M E' B' - V \ M ((Var v, t')#E) (subst_of B) - V \\<^sub>s\<^sub>e\<^sub>t d v t'" using 1 2 3 4 unfolding M_def by blast moreover have "Q' ((Var v, t')#E) (subst_of B) (d v t')" using Q_assm unfolding Q_def Q'_def by auto moreover have "R (subst_of B) \\<^sub>s\<^sub>e\<^sub>t d v t' = R (subst_of ((v,t')#B))" unfolding R_def d_def by auto ultimately have "M (subst_list (d v t') E) (subst_of ((v, t')#B)) - V \ R (subst_of ((v, t')#B)) - V" unfolding Q'_def E'_def B'_def d_def by blast thus ?thesis unfolding Q_def M_def R_def d_def by blast qed have "u \ subterms s \ subterms t - V \\<^sub>s\<^sub>e\<^sub>t subst_of U" when assms': "unify E B = Some U" "u \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range (subst_of U)) - V" "Q E (subst_of B)" for E B U and T::"('f,'v) term list" using assms' proof (induction E B arbitrary: U rule: Unification.unify.induct) case (1 B) thus ?case by (auto simp add: Q_def M_def R_def subti_def) next case (2 g X h Y E B U) from "2.prems"(1) obtain E' where E': "decompose (Fun g X) (Fun h Y) = Some E'" "g = h" "length X = length Y" "E' = zip X Y" "Unification.unify (E'@E) B = Some U" by (auto split: option.splits) moreover have "subterms_tuples (E'@E) \ subterms_tuples ((Fun g X, Fun h Y)#E)" proof fix u assume "u \ subterms_tuples (E'@E)" then obtain u1 u2 where u12: "(u1,u2) \ set (E'@E)" "u \ subterms u1 \ subterms u2" unfolding subtt_def by fastforce thus "u \ subterms_tuples ((Fun g X, Fun h Y)#E)" proof (cases "(u1,u2) \ set E'") case True hence "subterms u1 \ subterms (Fun g X)" "subterms u2 \ subterms (Fun h Y)" using E'(4) subterms_subset params_subterms subsetCE by (metis set_zip_leftD, metis set_zip_rightD) thus ?thesis using u12 unfolding subtt_def by auto next case False thus ?thesis using u12 unfolding subtt_def by fastforce qed qed hence "Q (E'@E) (subst_of B)" using "2.prems"(3) unfolding Q_def M_def by blast ultimately show ?case using "2.IH"[of E' U] "2.prems" by meson next case (3 v t' E B) show ?case proof (cases "t' = Var v") case True thus ?thesis using "3.prems" "3.IH"(1) unfolding Q_def M_def V_def subtt_def by auto next case False hence 1: "v \ fv t'" using "3.prems"(1) by auto hence "unify (subst_list (subst v t') E) ((v, t')#B) = Some U" using False "3.prems"(1) by auto thus ?thesis using Q_subst[OF 1 "3.prems"(3)] "3.IH"(2)[OF False 1 _ "3.prems"(2)] by metis qed next case (4 g X v E B U) have 1: "v \ fv (Fun g X)" using "4.prems"(1) not_None_eq by fastforce hence 2: "unify (subst_list (subst v (Fun g X)) E) ((v, Fun g X)#B) = Some U" using "4.prems"(1) by auto have 3: "Q ((Var v, Fun g X)#E) (subst_of B)" using "4.prems"(3) unfolding Q_def M_def subtt_def by auto show ?case using Q_subst[OF 1 3] "4.IH"[OF 1 2 "4.prems"(2)] by metis qed moreover obtain D where "unify [(s, t)] [] = Some D" "\ = subst_of D" using assms(1) by (auto simp: mgu_def split: option.splits) moreover have "Q [(s,t)] (subst_of [])" unfolding Q_def M_def R_def subtt_def subti_def by force ultimately show ?thesis using assms(2) unfolding V_def by auto qed lemma mgu_img_consts: fixes \::"('f,'v) subst" and s t::"('f,'v) term" and c::'f and z::'v assumes "mgu s t = Some \" "Fun c [] \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \)" shows "Fun c [] \ subterms s \ subterms t" proof - obtain u where "u \ (subterms s \ subterms t) - range Var" "u \ \ = Fun c []" using mgu_img_subterm_subst[OF assms(1), of "Fun c []"] assms(2) by force thus ?thesis by (cases u) auto qed lemma mgu_img_consts': fixes \::"('f,'v) subst" and s t::"('f,'v) term" and c::'f and z::'v assumes "mgu s t = Some \" "\ z = Fun c []" shows "Fun c [] \ s \ Fun c [] \ t" using mgu_img_consts[OF assms(1)] assms(2) by (metis Un_iff in_subterms_Union subst_imgI term.distinct(1)) lemma mgu_img_composed_var_term: fixes \::"('f,'v) subst" and s t::"('f,'v) term" and f::'f and Z::"'v list" assumes "mgu s t = Some \" "Fun f (map Var Z) \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \)" shows "\Z'. map \ Z' = map Var Z \ Fun f (map Var Z') \ subterms s \ subterms t" proof - obtain u where u: "u \ (subterms s \ subterms t) - range Var" "u \ \ = Fun f (map Var Z)" using mgu_img_subterm_subst[OF assms(1), of "Fun f (map Var Z)"] assms(2) by fastforce then obtain T where T: "u = Fun f T" "map (\t. t \ \) T = map Var Z" by (cases u) auto have "\t \ set T. \x. t = Var x" using T(2) by (induct T arbitrary: Z) auto then obtain Z' where Z': "map Var Z' = T" by (metis ex_map_conv) hence "map \ Z' = map Var Z" using T(2) by (induct Z' arbitrary: T Z) auto thus ?thesis using u(1) T(1) Z' by auto qed lemma mgu_ground_instance_case: assumes t: "fv (t \ \) = {}" shows "mgu t (t \ \) = Some (rm_vars (UNIV - fv t) \)" (is ?A) and mgu_ground_commutes: "mgu t (t \ \) = mgu (t \ \) t" (is ?B) proof - define \ where "\ \ rm_vars (UNIV - fv t) \" have \: "t \ \ = t \ \" using rm_vars_subst_eq'[of t \] unfolding \_def by metis have 0: "Unifier \ t (t \ \)" using subst_ground_ident[OF t, of \] term_subst_eq[of t \ \] unfolding \_def by (metis Diff_iff) obtain \ where \: "mgu t (t \ \) = Some \" "MGU \ t (t \ \)" using mgu_always_unifies[OF 0] mgu_gives_MGU[of t "t \ \"] unfolding \ by blast have 1: "subst_domain \ = fv t" using t MGU_is_Unifier[OF \(2)] subset_antisym[OF mgu_subst_domain[OF \(1)]] ground_term_subst_domain_fv_subset[of t \] subst_apply_fv_empty[OF t, of \] unfolding \ by auto have 2: "subst_domain \ = fv t" using 0 rm_vars_dom[of "UNIV - fv t" \] ground_term_subst_domain_fv_subset[of t \] subst_apply_fv_empty[OF t, of \] unfolding \_def by auto have "\ x = \ x" for x using 1 2 MGU_is_Unifier[OF \(2)] term_subst_eq_conv[of t \ \] subst_ground_ident[OF t[unfolded \], of \] subst_domI[of _ x] by metis hence "\ = \" by presburger thus A: ?A using \(1) unfolding \ \_def by blast have "Unifier \ (t \ \) t" using 0 by simp then obtain \' where \': "mgu (t \ \) t = Some \'" "MGU \' (t \ \) t" using mgu_always_unifies mgu_gives_MGU[of "t \ \" t] unfolding \ by fastforce have 3: "subst_domain \' = fv t" using t MGU_is_Unifier[OF \'(2)] subset_antisym[OF mgu_subst_domain[OF \'(1)]] ground_term_subst_domain_fv_subset[of t \'] subst_apply_fv_empty[OF t, of \'] unfolding \ by auto have "\' x = \ x" for x using 2 3 MGU_is_Unifier[OF \'(2)] term_subst_eq_conv[of t \' \] subst_ground_ident[OF t[unfolded \], of \'] subst_domI[of _ x] by metis hence "\' = \" by presburger thus ?B using A \'(1) unfolding \ \_def by argo qed subsection \Lemmata: The "Inequality Lemmata"\ text \Subterm injectivity (a stronger injectivity property)\ definition subterm_inj_on where "subterm_inj_on f A \ \x\A. \y\A. (\v. v \ f x \ v \ f y) \ x = y" lemma subterm_inj_on_imp_inj_on: "subterm_inj_on f A \ inj_on f A" unfolding subterm_inj_on_def inj_on_def by fastforce lemma subst_inj_on_is_bij_betw: "inj_on \ (subst_domain \) = bij_betw \ (subst_domain \) (subst_range \)" unfolding inj_on_def bij_betw_def by auto lemma subterm_inj_on_alt_def: "subterm_inj_on f A \ (inj_on f A \ (\s \ f`A. \u \ f`A. (\v. v \ s \ v \ u) \ s = u))" (is "?A \ ?B") unfolding subterm_inj_on_def inj_on_def by fastforce lemma subterm_inj_on_alt_def': "subterm_inj_on \ (subst_domain \) \ (inj_on \ (subst_domain \) \ (\s \ subst_range \. \u \ subst_range \. (\v. v \ s \ v \ u) \ s = u))" (is "?A \ ?B") by (metis subterm_inj_on_alt_def subst_range.simps) lemma subterm_inj_on_subset: assumes "subterm_inj_on f A" and "B \ A" shows "subterm_inj_on f B" proof - have "inj_on f A" "\s\f ` A. \u\f ` A. (\v. v \ s \ v \ u) \ s = u" using subterm_inj_on_alt_def[of f A] assms(1) by auto moreover have "f ` B \ f ` A" using assms(2) by auto ultimately have "inj_on f B" "\s\f ` B. \u\f ` B. (\v. v \ s \ v \ u) \ s = u" using inj_on_subset[of f A] assms(2) by blast+ thus ?thesis by (metis subterm_inj_on_alt_def) qed lemma inj_subst_unif_consts: fixes \ \ \::"('f,'v) subst" and s t::"('f,'v) term" assumes \: "subterm_inj_on \ (subst_domain \)" "\x \ (fv s \ fv t) - X. \c. \ x = Fun c []" "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ (subterms s \ subterms t) = {}" "ground (subst_range \)" "subst_domain \ \ X = {}" and \: "ground (subst_range \)" "subst_domain \ = subst_domain \" and unif: "Unifier \ (s \ \) (t \ \)" shows "\\. Unifier \ (s \ \) (t \ \)" proof - let ?xs = "subst_domain \" let ?ys = "(fv s \ fv t) - ?xs" have "\\::('f,'v) subst. s \ \ = t \ \" by (metis subst_subst_compose unif) then obtain \::"('f,'v) subst" where \: "mgu s t = Some \" using mgu_always_unifies by moura have 1: "\\::('f,'v) subst. s \ \ \ \ = t \ \ \ \" by (metis unif) have 2: "\\::('f,'v) subst. s \ \ \ \ = t \ \ \ \ \ \ \\<^sub>\ \ \\<^sub>s \" using mgu_gives_MGU[OF \] by simp have 3: "\(z::'v) (c::'f). \ z = Fun c [] \ Fun c [] \ s \ Fun c [] \ t" by (rule mgu_img_consts'[OF \]) have 4: "subst_domain \ \ range_vars \ = {}" by (metis mgu_gives_wellformed_subst[OF \] wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) have 5: "subst_domain \ \ range_vars \ \ fv s \ fv t" by (metis mgu_gives_wellformed_MGU[OF \] wf\<^sub>M\<^sub>G\<^sub>U_def) { fix x and \::"('f,'v) subst" assume "x \ subst_domain \" hence "(\ \\<^sub>s \) x = \ x" using \(4) ident_comp_subst_trm_if_disj[of \ \] unfolding range_vars_alt_def by fast } then obtain \::"('f,'v) subst" where \: "\x \ subst_domain \. \ x = (\ \\<^sub>s \) x" using 1 2 by moura have *: "\x. x \ subst_domain \ \ subst_domain \ \ \y \ ?ys. \ x = Var y" proof - fix x assume "x \ subst_domain \ \ ?xs" hence x: "x \ subst_domain \" "x \ subst_domain \" by auto then obtain c where c: "\ x = Fun c []" using \(2,5) 5 by moura hence *: "(\ \\<^sub>s \) x = Fun c []" using \ x by fastforce hence **: "x \ subst_domain (\ \\<^sub>s \)" "Fun c [] \ subst_range (\ \\<^sub>s \)" by (auto simp add: subst_domain_def) have "\ x = Fun c [] \ (\z. \ x = Var z \ \ z = Fun c [])" by (rule subst_img_comp_subset_const'[OF *]) moreover have "\ x \ Fun c []" proof (rule ccontr) assume "\\ x \ Fun c []" hence "Fun c [] \ s \ Fun c [] \ t" using 3 by metis moreover have "\u \ subst_range \. u \ subterms s \ subterms t" using \(3) by force hence "Fun c [] \ subterms s \ subterms t" by (metis c \ground (subst_range \)\x(2) ground_subst_dom_iff_img) ultimately show False by auto qed moreover have "\x' \ subst_domain \. \ x \ Var x'" proof (rule ccontr) assume "\(\x' \ subst_domain \. \ x \ Var x')" then obtain x' where x': "x' \ subst_domain \" "\ x = Var x'" by moura hence "\ x' = Fun c []" "(\ \\<^sub>s \) x = Fun c []" using * unfolding subst_compose_def by auto moreover have "x \ x'" using x(1) x'(2) 4 by (auto simp add: subst_domain_def) moreover have "x' \ subst_domain \" using x'(2) mgu_eliminates_dom[OF \] - by (metis (no_types) subst_elim_def subst_apply_term.simps(1) vars_iff_subterm_or_eq) + by (metis (no_types) subst_elim_def eval_term.simps(1) vars_iff_subterm_or_eq) moreover have "(\ \\<^sub>s \) x = \ x" "(\ \\<^sub>s \) x' = \ x'" using \ x(2) x'(1) by auto ultimately show False using subterm_inj_on_imp_inj_on[OF \(1)] * by (simp add: inj_on_def subst_compose_def x'(2) subst_domain_def) qed ultimately show "\y \ ?ys. \ x = Var y" by (metis 5 x(2) subtermeqI' vars_iff_subtermeq DiffI Un_iff subst_fv_imgI sup.orderE) qed have **: "inj_on \ (subst_domain \ \ ?xs)" proof (intro inj_onI) fix x y assume *: "x \ subst_domain \ \ subst_domain \" "y \ subst_domain \ \ subst_domain \" "\ x = \ y" hence "(\ \\<^sub>s \) x = (\ \\<^sub>s \) y" unfolding subst_compose_def by auto hence "\ x = \ y" using \ * by auto thus "x = y" using inj_onD[OF subterm_inj_on_imp_inj_on[OF \(1)]] *(1,2) by simp qed define \ where "\ = (\y'. if Var y' \ \ ` (subst_domain \ \ ?xs) then Var ((inv_into (subst_domain \ \ ?xs) \) (Var y')) else Var y'::('f,'v) term)" have a1: "Unifier (\ \\<^sub>s \) s t" using mgu_gives_MGU[OF \] by auto define \' where "\' = \ \\<^sub>s \" have d1: "subst_domain \' \ ?ys" proof fix z assume z: "z \ subst_domain \'" have "z \ ?xs \ z \ subst_domain \'" proof (cases "z \ subst_domain \") case True moreover assume "z \ ?xs" ultimately have z_in: "z \ subst_domain \ \ ?xs" by simp then obtain y where y: "\ z = Var y" "y \ ?ys" using * by moura hence "\ y = Var ((inv_into (subst_domain \ \ ?xs) \) (Var y))" using \_def z_in by simp hence "\ y = Var z" by (metis y(1) z_in ** inv_into_f_eq) hence "\' z = Var z" using \'_def y(1) subst_compose_def[of \ \] by simp thus ?thesis by (simp add: subst_domain_def) next case False hence "\ z = Var z" by (simp add: subst_domain_def) moreover assume "z \ ?xs" hence "\ z = Var z" using \_def * by force ultimately show ?thesis using \'_def subst_compose_def[of \ \] by (simp add: subst_domain_def) qed moreover have "subst_domain \ \ range_vars \" unfolding \'_def \_def range_vars_alt_def by (auto simp add: subst_domain_def) hence "subst_domain \' \ subst_domain \ \ range_vars \" using subst_domain_compose[of \ \] unfolding \'_def by blast ultimately show "z \ ?ys" using 5 z by auto qed have d2: "Unifier (\' \\<^sub>s \) s t" using a1 \'_def by auto have d3: "\ \\<^sub>s \' \\<^sub>s \ = \' \\<^sub>s \" proof - { fix z::'v assume z: "z \ ?xs" then obtain u where u: "\ z = u" "fv u = {}" using \ by auto hence "(\ \\<^sub>s \' \\<^sub>s \) z = u" by (simp add: subst_compose subst_ground_ident) moreover have "z \ subst_domain \'" using d1 z by auto hence "\' z = Var z" by (simp add: subst_domain_def) hence "(\' \\<^sub>s \) z = u" using u(1) by (simp add: subst_compose) ultimately have "(\ \\<^sub>s \' \\<^sub>s \) z = (\' \\<^sub>s \) z" by metis } moreover { fix z::'v assume "z \ ?ys" hence "z \ subst_domain \" using \(2) by auto hence "(\ \\<^sub>s \' \\<^sub>s \) z = (\' \\<^sub>s \) z" by (simp add: subst_compose subst_domain_def) } moreover { fix z::'v assume "z \ ?xs" "z \ ?ys" hence "\ z = Var z" "\' z = Var z" using \(2) d1 by blast+ hence "(\ \\<^sub>s \' \\<^sub>s \) z = (\' \\<^sub>s \) z" by (simp add: subst_compose) } ultimately show ?thesis by auto qed from d2 d3 have "Unifier (\' \\<^sub>s \) (s \ \) (t \ \)" by (metis subst_subst_compose) thus ?thesis by metis qed lemma inj_subst_unif_comp_terms: fixes \ \ \::"('f,'v) subst" and s t::"('f,'v) term" assumes \: "subterm_inj_on \ (subst_domain \)" "ground (subst_range \)" "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ (subterms s \ subterms t) = {}" "(fv s \ fv t) - subst_domain \ \ X" and tfr: "\f U. Fun f U \ subterms s \ subterms t \ U = [] \ (\u \ set U. u \ Var ` X)" and \: "ground (subst_range \)" "subst_domain \ = subst_domain \" and unif: "Unifier \ (s \ \) (t \ \)" shows "\\. Unifier \ (s \ \) (t \ \)" proof - let ?xs = "subst_domain \" let ?ys = "(fv s \ fv t) - ?xs" have "ground (subst_range \)" using \(2) by auto have "\\::('f,'v) subst. s \ \ = t \ \" by (metis subst_subst_compose unif) then obtain \::"('f,'v) subst" where \: "mgu s t = Some \" using mgu_always_unifies by moura have 1: "\\::('f,'v) subst. s \ \ \ \ = t \ \ \ \" by (metis unif) have 2: "\\::('f,'v) subst. s \ \ \ \ = t \ \ \ \ \ \ \\<^sub>\ \ \\<^sub>s \" using mgu_gives_MGU[OF \] by simp have 3: "\(z::'v) (c::'f). Fun c [] \ \ z \ Fun c [] \ s \ Fun c [] \ t" using mgu_img_consts[OF \] by force have 4: "subst_domain \ \ range_vars \ = {}" using mgu_gives_wellformed_subst[OF \] by (metis wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) have 5: "subst_domain \ \ range_vars \ \ fv s \ fv t" using mgu_gives_wellformed_MGU[OF \] by (metis wf\<^sub>M\<^sub>G\<^sub>U_def) { fix x and \::"('f,'v) subst" assume "x \ subst_domain \" hence "(\ \\<^sub>s \) x = \ x" using \ground (subst_range \)\ ident_comp_subst_trm_if_disj[of \ \ x] unfolding range_vars_alt_def by blast } then obtain \::"('f,'v) subst" where \: "\x \ subst_domain \. \ x = (\ \\<^sub>s \) x" using 1 2 by moura have ***: "\x. x \ subst_domain \ \ subst_domain \ \ fv (\ x) \ ?ys" proof - fix x assume "x \ subst_domain \ \ ?xs" hence x: "x \ subst_domain \" "x \ subst_domain \" by auto moreover have "\(\x' \ ?xs. x' \ fv (\ x))" proof (rule ccontr) assume "\\(\x' \ ?xs. x' \ fv (\ x))" then obtain x' where x': "x' \ fv (\ x)" "x' \ ?xs" by metis have "x \ x'" "x' \ subst_domain \" "\ x' = Var x'" using 4 x(1) x'(1) unfolding range_vars_alt_def by auto hence "(\ \\<^sub>s \) x' \ (\ \\<^sub>s \) x" "\ x' = (\ \\<^sub>s \) x'" using \ x(2) x'(2) by (metis subst_compose subst_mono vars_iff_subtermeq x'(1), - metis subst_apply_term.simps(1) subst_compose_def) + metis eval_term.simps(1) subst_compose_def) hence "\ x' \ \ x" using \ x(2) x'(2) by auto thus False using \(1) x'(2) x(2) \x \ x'\ unfolding subterm_inj_on_def by (meson subtermeqI') qed ultimately show "fv (\ x) \ ?ys" using 5 subst_dom_vars_in_subst[of x \] subst_fv_imgI[of \ x] by blast qed have **: "inj_on \ (subst_domain \ \ ?xs)" proof (intro inj_onI) fix x y assume *: "x \ subst_domain \ \ subst_domain \" "y \ subst_domain \ \ subst_domain \" "\ x = \ y" hence "(\ \\<^sub>s \) x = (\ \\<^sub>s \) y" unfolding subst_compose_def by auto hence "\ x = \ y" using \ * by auto thus "x = y" using inj_onD[OF subterm_inj_on_imp_inj_on[OF \(1)]] *(1,2) by simp qed have *: "\x. x \ subst_domain \ \ subst_domain \ \ \y \ ?ys. \ x = Var y" proof (rule ccontr) fix xi assume xi_assms: "xi \ subst_domain \ \ subst_domain \" "\(\y \ ?ys. \ xi = Var y)" hence xi_\: "xi \ subst_domain \" and \_xi_comp: "\(\y. \ xi = Var y)" using ***[of xi] 5 by auto then obtain f T where f: "\ xi = Fun f T" by (cases "\ xi") moura have "\g Y'. Y' \ [] \ Fun g (map Var Y') \ \ xi \ set Y' \ ?ys" proof - have "\c. Fun c [] \ \ xi \ Fun c [] \ \ xi" using \ xi_\ by (metis const_subterm_subst subst_compose) hence 1: "\c. \(Fun c [] \ \ xi)" using 3[of _ xi] xi_\ \(3) by auto have "\(\x. \ xi = Var x)" using f by auto hence "\g S. Fun g S \ \ xi \ (\s \ set S. (\c. s = Fun c []) \ (\x. s = Var x))" using nonvar_term_has_composed_shallow_term[of "\ xi"] by auto then obtain g S where gS: "Fun g S \ \ xi" "\s \ set S. (\c. s = Fun c []) \ (\x. s = Var x)" by moura have "\s \ set S. \x. s = Var x" using 1 term.order_trans gS by (metis (no_types, lifting) UN_I term.order_refl subsetCE subterms.simps(2) sup_ge2) then obtain S' where 2: "map Var S' = S" by (metis ex_map_conv) have "S \ []" using 1 term.order_trans[OF _ gS(1)] by fastforce hence 3: "S' \ []" "Fun g (map Var S') \ \ xi" using gS(1) 2 by auto have "set S' \ fv (Fun g (map Var S'))" by simp hence 4: "set S' \ fv (\ xi)" using 3(2) fv_subterms by force show ?thesis using ***[OF xi_assms(1)] 2 3 4 by auto qed then obtain g Y' where g: "Y' \ []" "Fun g (map Var Y') \ \ xi" "set Y' \ ?ys" by moura then obtain X where X: "map \ X = map Var Y'" "Fun g (map Var X) \ subterms s \ subterms t" using mgu_img_composed_var_term[OF \, of g Y'] by force hence "\(u::('f,'v) term) \ set (map Var X). u \ Var ` ?ys" using \(4) tfr g(1) by fastforce then obtain j where j: "j < length X" "X ! j \ ?ys" by (metis image_iff[of _ Var "fv s \ fv t - subst_domain \"] nth_map[of _ X Var] in_set_conv_nth[of _ "map Var X"] length_map[of Var X]) define yj' where yj': "yj' \ Y' ! j" define xj where xj: "xj \ X ! j" have "xj \ fv s \ fv t" using j X(1) g(3) 5 xj yj' by (metis length_map nth_map term.simps(1) in_set_conv_nth le_supE subsetCE subst_domI) hence xj_\: "xj \ subst_domain \" using j unfolding xj by simp have len: "length X = length Y'" by (rule map_eq_imp_length_eq[OF X(1)]) have "Var yj' \ \ xi" using term.order_trans[OF _ g(2)] j(1) len unfolding yj' by auto hence "\ yj' \ \ xi" - using \ xi_\ by (metis subst_apply_term.simps(1) subst_compose_def subst_mono) + using \ xi_\ by (metis eval_term.simps(1) subst_compose_def subst_mono) moreover have \_xj_var: "Var yj' = \ xj" using X(1) len j(1) nth_map unfolding xj yj' by metis - hence "\ yj' = \ xj" using \ xj_\ by (metis subst_apply_term.simps(1) subst_compose_def) + hence "\ yj' = \ xj" using \ xj_\ by (metis eval_term.simps(1) subst_compose_def) moreover have "xi \ xj" using \_xi_comp \_xj_var by auto ultimately show False using \(1) xi_\ xj_\ unfolding subterm_inj_on_def by blast qed define \ where "\ = (\y'. if Var y' \ \ ` (subst_domain \ \ ?xs) then Var ((inv_into (subst_domain \ \ ?xs) \) (Var y')) else Var y'::('f,'v) term)" have a1: "Unifier (\ \\<^sub>s \) s t" using mgu_gives_MGU[OF \] by auto define \' where "\' = \ \\<^sub>s \" have d1: "subst_domain \' \ ?ys" proof fix z assume z: "z \ subst_domain \'" have "z \ ?xs \ z \ subst_domain \'" proof (cases "z \ subst_domain \") case True moreover assume "z \ ?xs" ultimately have z_in: "z \ subst_domain \ \ ?xs" by simp then obtain y where y: "\ z = Var y" "y \ ?ys" using * by moura hence "\ y = Var ((inv_into (subst_domain \ \ ?xs) \) (Var y))" using \_def z_in by simp hence "\ y = Var z" by (metis y(1) z_in ** inv_into_f_eq) hence "\' z = Var z" using \'_def y(1) subst_compose_def[of \ \] by simp thus ?thesis by (simp add: subst_domain_def) next case False hence "\ z = Var z" by (simp add: subst_domain_def) moreover assume "z \ ?xs" hence "\ z = Var z" using \_def * by force ultimately show ?thesis using \'_def subst_compose_def[of \ \] by (simp add: subst_domain_def) qed moreover have "subst_domain \ \ range_vars \" unfolding \'_def \_def range_vars_alt_def subst_domain_def by auto hence "subst_domain \' \ subst_domain \ \ range_vars \" using subst_domain_compose[of \ \] unfolding \'_def by blast ultimately show "z \ ?ys" using 5 z by blast qed have d2: "Unifier (\' \\<^sub>s \) s t" using a1 \'_def by auto have d3: "\ \\<^sub>s \' \\<^sub>s \ = \' \\<^sub>s \" proof - { fix z::'v assume z: "z \ ?xs" then obtain u where u: "\ z = u" "fv u = {}" using \ by auto hence "(\ \\<^sub>s \' \\<^sub>s \) z = u" by (simp add: subst_compose subst_ground_ident) moreover have "z \ subst_domain \'" using d1 z by auto hence "\' z = Var z" by (simp add: subst_domain_def) hence "(\' \\<^sub>s \) z = u" using u(1) by (simp add: subst_compose) ultimately have "(\ \\<^sub>s \' \\<^sub>s \) z = (\' \\<^sub>s \) z" by metis } moreover { fix z::'v assume "z \ ?ys" hence "z \ subst_domain \" using \(2) by auto hence "(\ \\<^sub>s \' \\<^sub>s \) z = (\' \\<^sub>s \) z" by (simp add: subst_compose subst_domain_def) } moreover { fix z::'v assume "z \ ?xs" "z \ ?ys" hence "\ z = Var z" "\' z = Var z" using \(2) d1 by blast+ hence "(\ \\<^sub>s \' \\<^sub>s \) z = (\' \\<^sub>s \) z" by (simp add: subst_compose) } ultimately show ?thesis by auto qed from d2 d3 have "Unifier (\' \\<^sub>s \) (s \ \) (t \ \)" by (metis subst_subst_compose) thus ?thesis by metis qed context begin private lemma sat_ineq_subterm_inj_subst_aux: fixes \::"('f,'v) subst" assumes "Unifier \ (s \ \) (t \ \)" "ground (subst_range \)" "(fv s \ fv t) - X \ subst_domain \" "subst_domain \ \ X = {}" shows "\\::('f,'v) subst. subst_domain \ = X \ ground (subst_range \) \ s \ \ \ \ = t \ \ \ \" proof - have "\\. Unifier \ (s \ \) (t \ \) \ interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" proof - obtain \'::"('f,'v) subst" where *: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \'" using interpretation_subst_exists by metis hence "Unifier (\ \\<^sub>s \') (s \ \) (t \ \)" using assms(1) by simp thus ?thesis using * interpretation_comp by blast qed then obtain \' where \': "Unifier \' (s \ \) (t \ \)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \'" by moura define \'' where "\'' = rm_vars (UNIV - X) \'" have *: "fv (s \ \) \ X" "fv (t \ \) \ X" using assms(2,3) subst_fv_unfold_ground_img[of \] unfolding range_vars_alt_def by (simp_all add: Diff_subset_conv Un_commute) hence **: "subst_domain \'' = X" "ground (subst_range \'')" using rm_vars_img_subset[of "UNIV - X" \'] rm_vars_dom[of "UNIV - X" \'] \'(2) unfolding \''_def by auto hence "\t. t \ \ \ \'' = t \ \'' \ \" using subst_eq_if_disjoint_vars_ground[OF _ _ assms(2)] assms(4) by blast moreover have "Unifier \'' (s \ \) (t \ \)" using Unifier_dom_restrict[OF \'(1)] \''_def * by blast ultimately show ?thesis using ** by auto qed text \ The "inequality lemma": This lemma gives sufficient syntactic conditions for finding substitutions \\\ under which terms \s\ and \t\ are not unifiable. This is useful later when establishing the typing results since we there want to find well-typed solutions to inequality constraints / "negative checks" constraints, and this lemma gives conditions for protocols under which such constraints are well-typed satisfiable if satisfiable. \ lemma sat_ineq_subterm_inj_subst: fixes \ \ \::"('f,'v) subst" assumes \: "subterm_inj_on \ (subst_domain \)" "ground (subst_range \)" "subst_domain \ \ X = {}" "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ (subterms s \ subterms t) = {}" "(fv s \ fv t) - subst_domain \ \ X" and tfr: "(\x \ (fv s \ fv t) - X. \c. \ x = Fun c []) \ (\f U. Fun f U \ subterms s \ subterms t \ U = [] \ (\u \ set U. u \ Var ` X))" and \: "\\::('f,'v) subst. subst_domain \ = X \ ground (subst_range \) \ s \ \ \ \ \ t \ \ \ \" "(fv s \ fv t) - X \ subst_domain \" "subst_domain \ \ X = {}" "ground (subst_range \)" "subst_domain \ = subst_domain \" and \: "subst_domain \ = X" "ground (subst_range \)" shows "s \ \ \ \ \ t \ \ \ \" proof - have "\\. \Unifier \ (s \ \) (t \ \)" by (metis \(1) sat_ineq_subterm_inj_subst_aux[OF _ \(4,2,3)]) hence "\Unifier \ (s \ \) (t \ \)" using inj_subst_unif_consts[OF \(1) _ \(4,2,3) \(4,5)] inj_subst_unif_comp_terms[OF \(1,2,4,5) _ \(4,5)] tfr by metis moreover have "subst_domain \ \ subst_domain \ = {}" using \(2,3) \(1) by auto ultimately show ?thesis using \ subst_eq_if_disjoint_vars_ground[OF _ \(2) \(2)] by metis qed end lemma ineq_subterm_inj_cond_subst: assumes "X \ range_vars \ = {}" and "\f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t S \ T = [] \ (\u \ set T. u \ Var`X)" shows "\f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (S \\<^sub>s\<^sub>e\<^sub>t \) \ T = [] \ (\u \ set T. u \ Var`X)" proof (intro allI impI) let ?M = "\S. subterms\<^sub>s\<^sub>e\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \" let ?N = "\S. subterms\<^sub>s\<^sub>e\<^sub>t (\ ` (fv\<^sub>s\<^sub>e\<^sub>t S \ subst_domain \))" fix f T assume "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (S \\<^sub>s\<^sub>e\<^sub>t \)" hence 1: "Fun f T \ ?M S \ Fun f T \ ?N S" using subterms_subst[of _ \] by auto have 2: "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ \u \ set T. u \ Var`X" using fv_subset_subterms[of "Fun f T" "subst_range \"] assms(1) unfolding range_vars_alt_def by force have 3: "\x \ subst_domain \. \ x \ Var`X" proof fix x assume "x \ subst_domain \" hence "fv (\ x) \ range_vars \" using subst_dom_vars_in_subst subst_fv_imgI unfolding range_vars_alt_def by auto thus "\ x \ Var`X" using assms(1) by auto qed show "T = [] \ (\s \ set T. s \ Var`X)" using 1 proof assume "Fun f T \ ?M S" then obtain u where u: "u \ subterms\<^sub>s\<^sub>e\<^sub>t S" "u \ \ = Fun f T" by fastforce show ?thesis proof (cases u) case (Var x) hence "Fun f T \ subst_range \" using u(2) by (simp add: subst_domain_def) hence "\u \ set T. u \ Var`X" using 2 by force thus ?thesis by auto next case (Fun g S) hence "S = [] \ (\u \ set S. u \ Var`X)" using assms(2) u(1) by metis thus ?thesis proof assume "S = []" thus ?thesis using u(2) Fun by simp next assume "\u \ set S. u \ Var`X" then obtain u' where u': "u' \ set S" "u' \ Var`X" by moura hence "u' \ \ \ set T" using u(2) Fun by auto thus ?thesis using u'(2) 3 by (cases u') force+ qed qed next assume "Fun f T \ ?N S" thus ?thesis using 2 by force qed qed subsection \Lemmata: Sufficient Conditions for Term Matching\ definition subst_var_inv::"('a,'b) subst \ 'b set \ ('a,'b) subst" where "subst_var_inv \ X \ (\x. if Var x \ \ ` X then Var ((inv_into X \) (Var x)) else Var x)" lemma subst_var_inv_subst_domain: assumes "x \ subst_domain (subst_var_inv \ X)" shows "Var x \ \ ` X" by (meson assms subst_dom_vars_in_subst subst_var_inv_def) lemma subst_var_inv_subst_domain': assumes "X \ subst_domain \" shows "x \ subst_domain (subst_var_inv \ X) \ Var x \ \ ` X" proof show "Var x \ \ ` X \ x \ subst_domain (subst_var_inv \ X)" by (metis (no_types, lifting) assms f_inv_into_f in_mono inv_into_into subst_domI subst_dom_vars_in_subst subst_var_inv_def term.inject(1)) qed (rule subst_var_inv_subst_domain) lemma subst_var_inv_Var_range: "subst_range (subst_var_inv \ X) \ range Var" unfolding subst_var_inv_def by auto text \Injective substitutions from variables to variables are invertible\ lemma inj_var_ran_subst_is_invertible: assumes \_inj_on_X: "inj_on \ X" and \_var_on_X: "\ ` X \ range Var" and fv_t: "fv t \ X" shows "t = t \ \ \\<^sub>s subst_var_inv \ X" proof - have "\ x \ subst_var_inv \ X = Var x" when x: "x \ X" for x proof - obtain y where y: "\ x = Var y" using x \_var_on_X fv_t by auto hence "Var y \ \ ` X" using x by simp thus ?thesis using y inv_into_f_eq[OF \_inj_on_X x y] unfolding subst_var_inv_def by simp qed thus ?thesis using fv_t by (simp add: subst_compose_def trm_subst_ident'' subset_eq) qed lemma inj_var_ran_subst_is_invertible': assumes \_inj_on_t: "inj_on \ (fv t)" and \_var_on_t: "\ ` fv t \ range Var" shows "t = t \ \ \\<^sub>s subst_var_inv \ (fv t)" using assms inj_var_ran_subst_is_invertible by fast text \Sufficient conditions for matching unifiable terms\ lemma inj_var_ran_unifiable_has_subst_match: assumes "t \ \ = s \ \" "inj_on \ (fv t)" "\ ` fv t \ range Var" shows "t = s \ \ \\<^sub>s subst_var_inv \ (fv t)" using assms inj_var_ran_subst_is_invertible by fastforce end diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Compositionality.thy b/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Compositionality.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Compositionality.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Compositionality.thy @@ -1,3795 +1,3795 @@ (* Title: Stateful_Compositionality.thy Author: Andreas Viktor Hess, DTU SPDX-License-Identifier: BSD-3-Clause *) section \Stateful Protocol Compositionality\ theory Stateful_Compositionality imports Stateful_Typing Parallel_Compositionality Labeled_Stateful_Strands begin text\\label{sec:Stateful-Compositionality}\ subsection \Small Lemmata\ lemma (in typed_model) wt_subst_sstp_vars_type_subset: fixes a::"('fun,'var) stateful_strand_step" assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "\t \ subst_range \. fv t = {} \ (\x. t = Var x)" shows "\ ` Var ` fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ \ ` Var ` fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" (is ?A) and "\ ` Var ` set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)) = \ ` Var ` set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a)" (is ?B) and "\ ` Var ` vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ \ ` Var ` vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" (is ?C) proof - show ?A proof fix \ assume \: "\ \ \ ` Var ` fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" then obtain x where x: "x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "\ (Var x) = \" by moura show "\ \ \ ` Var ` fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" proof (cases "x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a") case False hence "\y \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a. \ y = Var x" proof (cases a) case (NegChecks X F G) hence *: "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" "x \ set X" using fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck(1)[of X "F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \" "G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \"] fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck(1)[of X F G] False x(1) by fastforce+ obtain y where y: "y \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" "x \ fv (rm_vars (set X) \ y)" using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_obtain_var[of _ _ "rm_vars (set X) \"] fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_obtain_var[of _ _ "rm_vars (set X) \"] *(1) by blast have "fv (rm_vars (set X) \ z) = {} \ (\u. rm_vars (set X) \ z = Var u)" for z using assms(2) rm_vars_img_subset[of "set X" \] by blast hence "rm_vars (set X) \ y = Var x" using y(2) by fastforce hence "\y \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a. rm_vars (set X) \ y = Var x" using y fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck(1)[of X F G] NegChecks *(2) by fastforce thus ?thesis by (metis (full_types) *(2) term.inject(1)) qed (use assms(2) x(1) subst_apply_img_var'[of x _ \] in fastforce)+ then obtain y where y: "y \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" "\ y = Var x" by moura hence "\ (Var y) = \" using x(2) assms(1) by (simp add: wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) thus ?thesis using y(1) by auto qed (use x in auto) qed show ?B by (metis bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst) show ?C proof fix \ assume \: "\ \ \ ` Var ` vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" then obtain x where x: "x \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "\ (Var x) = \" by moura show "\ \ \ ` Var ` vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" proof (cases "x \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a") case False hence "\y \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a. \ y = Var x" proof (cases a) case (NegChecks X F G) hence *: "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" "x \ set X" using vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck[of X "F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \" "G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \"] vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck[of X F G] False x(1) by (fastforce, blast) obtain y where y: "y \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" "x \ fv (rm_vars (set X) \ y)" using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_obtain_var[of _ _ "rm_vars (set X) \"] fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_obtain_var[of _ _ "rm_vars (set X) \"] *(1) by blast have "fv (rm_vars (set X) \ z) = {} \ (\u. rm_vars (set X) \ z = Var u)" for z using assms(2) rm_vars_img_subset[of "set X" \] by blast hence "rm_vars (set X) \ y = Var x" using y(2) by fastforce hence "\y \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a. rm_vars (set X) \ y = Var x" using y vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck[of X F G] NegChecks by blast thus ?thesis by (metis (full_types) *(2) term.inject(1)) qed (use assms(2) x(1) subst_apply_img_var'[of x _ \] in fastforce)+ then obtain y where y: "y \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" "\ y = Var x" by moura hence "\ (Var y) = \" using x(2) assms(1) by (simp add: wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) thus ?thesis using y(1) by auto qed (use x in auto) qed qed lemma (in typed_model) wt_subst_lsst_vars_type_subset: fixes A::"('fun,'var,'a) labeled_stateful_strand" assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "\t \ subst_range \. fv t = {} \ (\x. t = Var x)" shows "\ ` Var ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ \ ` Var ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" (is ?A) and "\ ` Var ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = \ ` Var ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" (is ?B) and "\ ` Var ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ \ ` Var ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" (is ?C) proof - have "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) = fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) = set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" when "a = (l,b)" for a l b and A::"('fun,'var,'a) labeled_stateful_strand" using that unlabel_Cons(1)[of l b A] unlabel_subst[of "a#A" \] subst_lsst_cons[of a A \] subst_sst_cons[of b "unlabel A" \] subst_apply_labeled_stateful_strand_step.simps(1)[of l b \] vars\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons[of l b A] vars\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons[of l "b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" "A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] fv\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons[of l b A] fv\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons[of l "b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" "A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons[of l b A] bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons[of l "b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" "A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] by simp_all hence *: "\ ` Var ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = \ ` Var ` vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ \ ` Var ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "\ ` Var ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) = \ ` Var ` vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \ \ ` Var ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "\ ` Var ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = \ ` Var ` fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ \ ` Var ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "\ ` Var ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) = \ ` Var ` fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \ \ ` Var ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "\ ` Var ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = \ ` Var ` set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)) \ \ ` Var ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "\ ` Var ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) = \ ` Var ` set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b) \ \ ` Var ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" when "a = (l,b)" for a l b and A::"('fun,'var,'a) labeled_stateful_strand" using that by fast+ have "?A \ ?B \ ?C" proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) show ?case using Cons.IH wt_subst_sstp_vars_type_subset[OF assms, of b] *[OF a, of A] by (metis Un_mono) qed simp thus ?A ?B ?C by metis+ qed lemma (in stateful_typed_model) fv_pair_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subset: assumes "d \ set D" shows "fv (pair (snd d)) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" using assms unfolding pair_def by (induct D) (auto simp add: unlabel_def) lemma (in stateful_typed_model) labeled_sat_ineq_lift: assumes "\M; map (\d. \X\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\dbproj i D. d \ set Di]\\<^sub>d \" (is "?R1 D") and "\(j,p) \ {(i,t,s)} \ set D \ set Di. \(k,q) \ {(i,t,s)} \ set D \ set Di. (\\. Unifier \ (pair p) (pair q)) \ j = k" (is "?R2 D") shows "\M; map (\d. \X\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\D. d \ set Di]\\<^sub>d \" using assms proof (induction D) case (Cons dl D) obtain d l where dl: "dl = (l,d)" by (metis surj_pair) have 1: "?R1 D" proof (cases "i = l") case True thus ?thesis using Cons.prems(1) dl by (cases "dl \ set Di") (auto simp add: dbproj_def) next case False thus ?thesis using Cons.prems(1) dl by (auto simp add: dbproj_def) qed have "set D \ set (dl#D)" by auto hence 2: "?R2 D" using Cons.prems(2) by blast have "i \ l \ dl \ set Di \ \M; [\X\\\: [(pair (t,s), pair (snd dl))]\\<^sub>s\<^sub>t]\\<^sub>d \" using Cons.prems(1) dl by (auto simp add: ineq_model_def dbproj_def) moreover have "\\. Unifier \ (pair (t,s)) (pair d) \ i = l" using Cons.prems(2) dl by force ultimately have 3: "dl \ set Di \ \M; [\X\\\: [(pair (t,s), pair (snd dl))]\\<^sub>s\<^sub>t]\\<^sub>d \" using strand_sem_not_unif_is_sat_ineq[of "pair (t,s)" "pair d"] dl by fastforce show ?case using Cons.IH[OF 1 2] 3 dl by auto qed simp lemma (in stateful_typed_model) labeled_sat_ineq_dbproj: assumes "\M; map (\d. \X\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\D. d \ set Di]\\<^sub>d \" (is "?P D") shows "\M; map (\d. \X\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\dbproj i D. d \ set Di]\\<^sub>d \" (is "?Q D") using assms proof (induction D) case (Cons di D) obtain d j where di: "di = (j,d)" by (metis surj_pair) have "?P D" using Cons.prems by (cases "di \ set Di") auto hence IH: "?Q D" by (metis Cons.IH) show ?case using di IH proof (cases "i = j \ di \ set Di") case True have 1: "\M; [\X\\\: [(pair (t,s), pair (snd di))]\\<^sub>s\<^sub>t]\\<^sub>d \" using Cons.prems True by auto have 2: "dbproj i (di#D) = di#dbproj i D" using True dbproj_Cons(1) di by auto show ?thesis using 1 2 IH by auto qed (auto simp add: dbproj_def) qed (simp add: dbproj_def) lemma (in stateful_typed_model) labeled_sat_ineq_dbproj_sem_equiv: assumes "\(j,p) \ ((\(t, s). (i, t, s)) ` set F') \ set D. \(k,q) \ ((\(t, s). (i, t, s)) ` set F') \ set D. (\\. Unifier \ (pair p) (pair q)) \ j = k" and "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (map snd D) \ set X = {}" shows "\M; map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd D))\\<^sub>d \ \ \M; map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))\\<^sub>d \" proof - let ?A = "set (map snd D) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" let ?B = "set (map snd (dbproj i D)) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" let ?C = "set (map snd D) - set (map snd (dbproj i D))" let ?F = "(\(t, s). (i, t, s)) ` set F'" let ?P = "\\. subst_domain \ = set X \ ground (subst_range \)" have 1: "\(t, t') \ set (map snd D). (fv t \ fv t') \ set X = {}" "\(t, t') \ set (map snd (dbproj i D)). (fv t \ fv t') \ set X = {}" using assms(2) dbproj_subset[of i D] unfolding unlabel_def by force+ have 2: "?B \ ?A" unfolding dbproj_def by auto have 3: "\Unifier \ (pair f) (pair d)" when f: "f \ set F'" and d: "d \ set (map snd D) - set (map snd (dbproj i D))" for f d and \::"('fun,'var) subst" proof - obtain k where k: "(k,d) \ set D - set (dbproj i D)" using d by force have "(i,f) \ ((\(t, s). (i, t, s)) ` set F') \ set D" "(k,d) \ ((\(t, s). (i, t, s)) ` set F') \ set D" using f k by auto hence "i = k" when "Unifier \ (pair f) (pair d)" for \ using assms(1) that by blast moreover have "k \ i" using k d unfolding dbproj_def by simp ultimately show ?thesis by metis qed have "f \\<^sub>p \ \ d \\<^sub>p \" when "f \ set F'" "d \ ?C" for f d and \::"('fun,'var) subst" by (metis fun_pair_eq_subst 3[OF that]) hence "f \\<^sub>p (\ \\<^sub>s \) \ ?C \\<^sub>p\<^sub>s\<^sub>e\<^sub>t (\ \\<^sub>s \)" when "f \ set F'" for f and \::"('fun,'var) subst" using that by blast moreover have "?C \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ = ?C \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" when "?P \" for \ using assms(2) that pairs_substI[of \ "(set (map snd D) - set (map snd (dbproj i D)))"] by blast ultimately have 4: "f \\<^sub>p (\ \\<^sub>s \) \ ?C \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" when "f \ set F'" "?P \" for f and \::"('fun,'var) subst" by (metis that subst_pairs_compose) { fix f and \::"('fun,'var) subst" assume "f \ set F'" "?P \" hence "f \\<^sub>p (\ \\<^sub>s \) \ ?C \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" by (metis 4) hence "f \\<^sub>p (\ \\<^sub>s \) \ ?A - ?B" by force } hence 5: "\f\set F'. \\. ?P \ \ f \\<^sub>p (\ \\<^sub>s \) \ ?A - ?B" by metis show ?thesis using negchecks_model_db_subset[OF 2] negchecks_model_db_supset[OF 2 5] tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_sem_equiv[OF 1(1)] tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_sem_equiv[OF 1(2)] tr_NegChecks_constr_iff(1) strand_sem_eq_defs(2) by (metis (no_types, lifting)) qed lemma (in stateful_typed_model) labeled_sat_eqs_list_all: assumes "\(j, p) \ {(i,t,s)} \ set D. \(k,q) \ {(i,t,s)} \ set D. (\\. Unifier \ (pair p) (pair q)) \ j = k" (is "?P D") and "\M; map (\d. \ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t) D\\<^sub>d \" (is "?Q D") shows "list_all (\d. fst d = i) D" using assms proof (induction D rule: List.rev_induct) case (snoc di D) obtain d j where di: "di = (j,d)" by (metis surj_pair) have "pair (t,s) \ \ = pair d \ \" using di snoc.prems(2) by auto hence "\\. Unifier \ (pair (t,s)) (pair d)" by auto hence 1: "i = j" using snoc.prems(1) di by fastforce have "set D \ set (D@[di])" by auto hence 2: "?P D" using snoc.prems(1) by blast have 3: "?Q D" using snoc.prems(2) by auto show ?case using di 1 snoc.IH[OF 2 3] by simp qed simp lemma (in stateful_typed_model) labeled_sat_eqs_subseqs: assumes "Di \ set (subseqs D)" and "\(j, p) \ {(i,t,s)} \ set D. \(k, q) \ {(i,t,s)} \ set D. (\\. Unifier \ (pair p) (pair q)) \ j = k" (is "?P D") and "\M; map (\d. \ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t) Di\\<^sub>d \" shows "Di \ set (subseqs (dbproj i D))" proof - have "set Di \ set D" by (rule subseqs_subset[OF assms(1)]) hence "?P Di" using assms(2) by blast thus ?thesis using labeled_sat_eqs_list_all[OF _ assms(3)] subseqs_mem_dbproj[OF assms(1)] by simp qed lemma (in stateful_typing_result) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p: assumes "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel S)" shows "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S))" using assms proof (induction S) case (Cons a S) have prems: "tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd a)" "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel S)" using Cons.prems unlabel_Cons(2)[of a S] by simp_all hence IH: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S))" by (metis Cons.IH) obtain l b where a: "a = (l,b)" by (metis surj_pair) with Cons show ?case proof (cases b) case (Equality c t t') hence "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#S) = a#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S" by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_Cons(3) a) thus ?thesis using a IH prems by fastforce next case (NegChecks X F G) hence "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#S) = a#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S" by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_Cons(7) a) thus ?thesis using a IH prems by fastforce qed auto qed simp lemma (in stateful_typed_model) setops\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq: "setops\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)) = setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) thus ?case using Cons.IH by (cases b) (simp_all add: setops\<^sub>s\<^sub>s\<^sub>t_def) qed simp subsection \Locale Setup and Definitions\ locale labeled_stateful_typed_model = stateful_typed_model arity public Ana \ Pair + labeled_typed_model arity public Ana \ label_witness1 label_witness2 for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,'var) term \ (('fun,'var) term list \ ('fun,'var) term list)" and \::"('fun,'var) term \ ('fun,'atom::finite) term_type" and Pair::"'fun" and label_witness1::"'lbl" and label_witness2::"'lbl" begin definition lpair where "lpair lp \ case lp of (i,p) \ (i,pair p)" lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_pair_image[simp]: "lpair ` (setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,send\ts\)) = {}" "lpair ` (setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,receive\ts\)) = {}" "lpair ` (setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,\ac: t \ t'\)) = {}" "lpair ` (setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,insert\t,s\)) = {(i, pair (t,s))}" "lpair ` (setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,delete\t,s\)) = {(i, pair (t,s))}" "lpair ` (setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,\ac: t \ s\)) = {(i, pair (t,s))}" "lpair ` (setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,\X\\\: F \\: F'\)) = ((\(t,s). (i, pair (t,s))) ` set F')" unfolding lpair_def by force+ definition par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\::('fun,'var,'lbl) labeled_stateful_strand) (Secrets::('fun,'var) terms) \ (\l1 l2. l1 \ l2 \ GSMP_disjoint (trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \)) (trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \)) Secrets) \ (\s \ Secrets. \{} \\<^sub>c s) \ ground Secrets \ (\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. (\\. Unifier \ (pair p) (pair q)) \ i = j)" definition declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ \ {s. \{set ts | ts. \\, receive\ts\\ \ set (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)} \ s}" definition strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t ("_ leaks _ under _") where "(\::('fun,'var,'lbl) labeled_stateful_strand) leaks Secrets under \ \ (\t \ Secrets - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \. \n. \ \\<^sub>s (proj_unl n \@[send\[t]\]))" type_synonym ('a,'b,'c) labeleddbstate = "('c strand_label \ (('a,'b) term \ ('a,'b) term)) set" type_synonym ('a,'b,'c) labeleddbstatelist = "('c strand_label \ (('a,'b) term \ ('a,'b) term)) list" definition typing_cond\<^sub>s\<^sub>s\<^sub>t where "typing_cond\<^sub>s\<^sub>s\<^sub>t \ \ wf\<^sub>s\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t \) \ tfr\<^sub>s\<^sub>s\<^sub>t \" text \ For proving the compositionality theorem for stateful constraints the idea is to first define a variant of the reduction technique that was used to establish the stateful typing result. This variant performs database-state projections, and it allows us to reduce the compositionality problem for stateful constraints to ordinary constraints. \ fun tr\<^sub>p\<^sub>c:: "('fun,'var,'lbl) labeled_stateful_strand \ ('fun,'var,'lbl) labeleddbstatelist \ ('fun,'var,'lbl) labeled_strand list" where "tr\<^sub>p\<^sub>c [] D = [[]]" | "tr\<^sub>p\<^sub>c ((i,send\ts\)#A) D = map ((#) (i,send\ts\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>c A D)" | "tr\<^sub>p\<^sub>c ((i,receive\ts\)#A) D = map ((#) (i,receive\ts\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>c A D)" | "tr\<^sub>p\<^sub>c ((i,\ac: t \ t'\)#A) D = map ((#) (i,\ac: t \ t'\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>c A D)" | "tr\<^sub>p\<^sub>c ((i,insert\t,s\)#A) D = tr\<^sub>p\<^sub>c A (List.insert (i,(t,s)) D)" | "tr\<^sub>p\<^sub>c ((i,delete\t,s\)#A) D = ( concat (map (\Di. map (\B. (map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\dbproj i D. d \ set Di])@B) (tr\<^sub>p\<^sub>c A [d\D. d \ set Di])) (subseqs (dbproj i D))))" | "tr\<^sub>p\<^sub>c ((i,\ac: t \ s\)#A) D = concat (map (\B. map (\d. (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#B) (dbproj i D)) (tr\<^sub>p\<^sub>c A D))" | "tr\<^sub>p\<^sub>c ((i,\X\\\: F \\: F' \)#A) D = map ((@) (map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D))))) (tr\<^sub>p\<^sub>c A D)" end locale labeled_stateful_typing = labeled_stateful_typed_model arity public Ana \ Pair label_witness1 label_witness2 + stateful_typing_result arity public Ana \ Pair for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,'var) term \ (('fun,'var) term list \ ('fun,'var) term list)" and \::"('fun,'var) term \ ('fun,'atom::finite) term_type" and Pair::"'fun" and label_witness1::"'lbl" and label_witness2::"'lbl" begin sublocale labeled_typing by unfold_locales end subsection \Small Lemmata\ context labeled_stateful_typed_model begin lemma declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_alt_def: "declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ = {s. \{set ts | ts. \\, receive\ts\\ \ set \} \\<^sub>s\<^sub>e\<^sub>t \ \ s}" proof - have 0: "(l, receive\ts\) \ set (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = (\ts'. (l, receive\ts'\) \ set \ \ ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" (is "?A \ = ?B \") for ts l proof show "?A \ \ ?B \" proof (induction \) case (Cons a \) obtain k b where a: "a = (k,b)" by (metis surj_pair) show ?case proof (cases "?A \") case False hence "(l,receive\ts\) = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" using Cons.prems subst_lsst_cons[of a \ \] by auto thus ?thesis unfolding a by (cases b) auto qed (use Cons.IH in auto) qed simp show "?B \ \ ?A \" proof (induction \) case (Cons a \) obtain k b where a: "a = (k,b)" by (metis surj_pair) show ?case proof (cases "?B \") case False hence "\ts'. a = (l, receive\ts'\) \ ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using Cons.prems by auto thus ?thesis using subst_lsst_cons[of a \ \] unfolding a by (cases b) auto qed (use Cons.IH subst_lsst_cons[of a \ \] in auto) qed simp qed let ?M = "\A. \{set ts |ts. \\, receive\ts\\ \ set A}" have 1: "?M (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = ?M \ \\<^sub>s\<^sub>e\<^sub>t \" (is "?A = ?B") proof show "?A \ ?B" proof fix t assume t: "t \ ?A" then obtain ts where ts: "t \ set ts" "\\, receive\ts\\ \ set (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by blast thus "t \ ?B" using 0[of \ ts] by fastforce qed show "?B \ ?A" proof fix t assume t: "t \ ?B" then obtain ts where ts: "t \ set ts \\<^sub>s\<^sub>e\<^sub>t \" "\\, receive\ts\\ \ set \" by blast hence "\\, receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\ \ set (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using 0[of \ "ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \"] by blast thus "t \ ?A" using ts(1) by force qed qed show ?thesis using 1 unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by argo qed lemma declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_prefix_subset: assumes AB: "prefix A B" shows "declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I \ declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t B I" proof fix t assume t: "t \ declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I" obtain C where C: "B = A@C" using prefixE[OF AB] by metis show "t \ declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t B I" using t ideduct_mono[of "\{set ts |ts. (\, receive\ts\) \ set A} \\<^sub>s\<^sub>e\<^sub>t I" t "\{set ts |ts. (\, receive\ts\) \ set B} \\<^sub>s\<^sub>e\<^sub>t I"] unfolding C declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_alt_def by auto qed lemma declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_star_receive_supset: "{t | t ts. \\, receive\ts\\ \ set \ \ t \ set ts} \\<^sub>s\<^sub>e\<^sub>t \ \ declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \" unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_alt_def by (fastforce intro: intruder_deduct.Axiom) lemma declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj_eq: "declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A) I" using proj_mem_iff(2)[of _ A] unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_alt_def by simp lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_nil: assumes "ground Sec" "\s \ Sec. \s'\subterms s. {} \\<^sub>c s' \ s' \ Sec" "\s \ Sec. \{} \\<^sub>c s" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] Sec" using assms unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by simp lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subset: assumes A: "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A Sec" and BA: "set B \ set A" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t B Sec" proof - let ?L = "\n A. trms\<^sub>s\<^sub>s\<^sub>t (proj_unl n A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl n A)" have "?L n B \ ?L n A" for n using trms\<^sub>s\<^sub>s\<^sub>t_mono[OF proj_set_mono(2)[OF BA]] setops\<^sub>s\<^sub>s\<^sub>t_mono[OF proj_set_mono(2)[OF BA]] by blast hence "GSMP_disjoint (?L m B) (?L n B) Sec" when nm: "m \ n" for n m::'lbl using GSMP_disjoint_subset[of "?L m A" "?L n A" Sec "?L m B" "?L n B"] A nm unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by simp thus "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t B Sec" using A setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_mono[OF BA] unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by blast qed lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_split: assumes "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B) Sec" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A Sec" "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t B Sec" using par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subset[OF assms] by simp_all lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj: assumes "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A Sec" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A) Sec" using par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subset[OF assms] by simp lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t: assumes A: "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A S" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) S" proof (unfold par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def case_prod_unfold; intro conjI) show "ground S" "\s \ S. \{} \\<^sub>c s" using A unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by fast+ let ?M = "\l B. (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l B) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l B))" let ?P = "\B. \l1 l2. l1 \ l2 \ GSMP_disjoint (?M l1 B) (?M l2 B) S" let ?Q = "\B. \p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B. \q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B. (\\. Unifier \ (pair (snd p)) (pair (snd q))) \ fst p = fst q" have "?P A" "?Q A" using A unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def case_prod_unfold by blast+ thus "?P (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "?Q (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" by (metis setops\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq proj_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t, metis setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq) qed lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst: assumes A: "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A S" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "subst_domain \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = {}" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) S" proof (unfold par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def case_prod_unfold; intro conjI) show "ground S" "\s \ S. \{} \\<^sub>c s" using A unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by fast+ let ?N = "\l B. trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l B) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l B)" define M where "M \ \l (B::('fun,'var,'lbl) labeled_stateful_strand). ?N l B" let ?P = "\p q. \\. Unifier \ (pair (snd p)) (pair (snd q))" let ?Q = "\B. \p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B. \q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B. ?P p q \ fst p = fst q" let ?R = "\B. \l1 l2. l1 \ l2 \ GSMP_disjoint (?N l1 B) (?N l2 B) S" have d: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l A) \ subst_domain \ = {}" for l using \(3) unfolding proj_def bvars\<^sub>s\<^sub>s\<^sub>t_def unlabel_def by auto have "GSMP_disjoint (M l1 A) (M l2 A) S" when l: "l1 \ l2" for l1 l2 using l A unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def M_def by presburger moreover have "M l (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = (M l A) \\<^sub>s\<^sub>e\<^sub>t \" for l using fun_pair_subst_set[of \ "setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l A)", symmetric] trms\<^sub>s\<^sub>s\<^sub>t_subst[OF d[of l]] setops\<^sub>s\<^sub>s\<^sub>t_subst[OF d[of l]] proj_subst[of l A \] unfolding M_def unlabel_subst by auto ultimately have "GSMP_disjoint (M l1 (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) (M l2 (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) S" when l: "l1 \ l2" for l1 l2 using l GSMP_wt_subst_subset[OF _ \(1,2), of _ "M l1 A"] GSMP_wt_subst_subset[OF _ \(1,2), of _ "M l2 A"] unfolding GSMP_disjoint_def by fastforce thus "?R (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" unfolding M_def by blast have "?Q A" using A unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by force thus "?Q (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using \(3) proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) have 0: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) = set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd a)) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" unfolding bvars\<^sub>s\<^sub>s\<^sub>t_def unlabel_def by simp have "?Q A" "subst_domain \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = {}" using Cons.prems 0 unfolding setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto hence IH: "?Q (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using Cons.IH unfolding setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by blast have 1: "fst p = fst q" when p: "p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" and q: "q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" and pq: "?P p q" for p q using a p q pq by (cases b) auto have 2: "fst p = fst q" when p: "p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and q: "q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" and pq: "?P p q" for p q proof - obtain p' X where p': "p' \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "fst p = fst p'" "X \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A)" "snd p = snd p' \\<^sub>p rm_vars X \" using setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_subst[OF p] 0 by blast obtain q' Y where q': "q' \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" "fst q = fst q'" "Y \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A)" "snd q = snd q' \\<^sub>p rm_vars Y \" using setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_in_subst[OF q] 0 by blast have "pair (snd p) = pair (snd p') \ \" "pair (snd q) = pair (snd q') \ \" using fun_pair_subst[of "snd p'" "rm_vars X \"] fun_pair_subst[of "snd q'" "rm_vars Y \"] p'(3,4) q'(3,4) Cons.prems(2) rm_vars_apply'[of \ X] rm_vars_apply'[of \ Y] by fastforce+ hence "\\. Unifier \ (pair (snd p')) (pair (snd q'))" using pq Unifier_comp' by metis thus ?thesis using Cons.prems p'(1,2) q'(1,2) by simp qed show ?case by (metis 1 2 IH Un_iff setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_cons subst_lsst_cons) qed simp qed lemma wf_pair_negchecks_map': assumes "wf\<^sub>s\<^sub>t X (unlabel A)" shows "wf\<^sub>s\<^sub>t X (unlabel (map (\G. (i,\Y\\\: (F@G)\\<^sub>s\<^sub>t)) M@A))" using assms by (induct M) auto lemma wf_pair_eqs_ineqs_map': fixes A::"('fun,'var,'lbl) labeled_strand" assumes "wf\<^sub>s\<^sub>t X (unlabel A)" "Di \ set (subseqs (dbproj i D))" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X" shows "wf\<^sub>s\<^sub>t X (unlabel ( (map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\dbproj i D. d \ set Di])@A))" proof - let ?f = "[d\dbproj i D. d \ set Di]" define c1 where c1: "c1 = map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di" define c2 where c2: "c2 = map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) ?f" define c3 where c3: "c3 = map (\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t) (unlabel Di)" define c4 where c4: "c4 = map (\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) (unlabel ?f)" have ci_eqs: "c3 = unlabel c1" "c4 = unlabel c2" unfolding c1 c2 c3 c4 unlabel_def by auto have 1: "wf\<^sub>s\<^sub>t X (unlabel (c2@A))" using wf_fun_pair_ineqs_map[OF assms(1)] ci_eqs(2) unlabel_append[of c2 A] c4 by metis have 2: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel Di) \ X" using assms(3) subseqs_set_subset(1)[OF assms(2)] unfolding unlabel_def dbproj_def by fastforce { fix B::"('fun,'var) strand" assume "wf\<^sub>s\<^sub>t X B" hence "wf\<^sub>s\<^sub>t X (unlabel c1@B)" using 2 unfolding c1 unlabel_def by (induct Di) auto } thus ?thesis using 1 unfolding c1 c2 unlabel_def by simp qed lemma trms\<^sub>s\<^sub>s\<^sub>t_setops\<^sub>s\<^sub>s\<^sub>t_wt_instance_ex: defines "M \ \A. trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" assumes B: "\b \ set B. \a \ set A. \\. b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "\t \ M B. \s \ M A. \\. t = s \ \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" proof let ?P = "\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" fix t assume "t \ M B" then obtain b where b: "b \ set B" "t \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd b) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd b)" unfolding M_def unfolding unlabel_def trms\<^sub>s\<^sub>s\<^sub>t_def setops\<^sub>s\<^sub>s\<^sub>t_def by auto then obtain a \ where a: "a \ set A" "b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using B by meson note \' = wt_subst_rm_vars[OF \(1)] wf_trms_subst_rm_vars'[OF \(2)] have "t \ M (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using b(2) a unfolding M_def subst_apply_labeled_stateful_strand_def unlabel_def trms\<^sub>s\<^sub>s\<^sub>t_def setops\<^sub>s\<^sub>s\<^sub>t_def by auto moreover have "\s \ M A. \\. t = s \ \ \ ?P \" when "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subst'[OF that] \' unfolding M_def by blast moreover have "\s \ M A. \\. t = s \ \ \ ?P \" when t: "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A \\<^sub>s\<^sub>s\<^sub>t \)" proof - obtain p where p: "p \ setops\<^sub>s\<^sub>s\<^sub>t (unlabel A \\<^sub>s\<^sub>s\<^sub>t \)" "t = pair p" using t by blast then obtain q X where q: "q \ setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "p = q \\<^sub>p rm_vars (set X) \" using setops\<^sub>s\<^sub>s\<^sub>t_subst'[OF p(1)] by blast hence "t = pair q \ rm_vars (set X) \" using fun_pair_subst[of q "rm_vars (set X) \"] p(2) by presburger thus ?thesis using \'[of "set X"] q(1) unfolding M_def by blast qed ultimately show "\s \ M A. \\. t = s \ \ \ ?P \" unfolding M_def unlabel_subst by fast qed lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_wt_instance_ex: assumes B: "\b \ set B. \a \ set A. \\. b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "\p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B. \q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \\. fst p = fst q \ snd p = snd q \\<^sub>p \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" proof let ?P = "\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" fix p assume "p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B" then obtain b where b: "b \ set B" "p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p b" unfolding setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by blast then obtain a \ where a: "a \ set A" "b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using B by meson hence p: "p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using b(2) unfolding setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def subst_apply_labeled_stateful_strand_def by auto obtain X q where q: "q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "fst p = fst q" "snd p = snd q \\<^sub>p rm_vars X \" using setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_subst[OF p] by blast show "\q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \\. fst p = fst q \ snd p = snd q \\<^sub>p \ \ ?P \" using q wt_subst_rm_vars[OF \(1)] wf_trms_subst_rm_vars'[OF \(2)] by blast qed lemma deduct_proj_priv_term_prefix_ex_stateful: assumes A: "ik\<^sub>s\<^sub>s\<^sub>t (proj_unl l A) \\<^sub>s\<^sub>e\<^sub>t I \ t" and t: "\{} \\<^sub>c t" shows "\B k s. (k = \ \ k = ln l) \ prefix (B@[(k,receive\s\)]) A \ declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((B@[(k,receive\s\)])) I = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I \ ik\<^sub>s\<^sub>s\<^sub>t (proj_unl l (B@[(k,receive\s\)])) = ik\<^sub>s\<^sub>s\<^sub>t (proj_unl l A)" using A proof (induction A rule: List.rev_induct) case Nil have "ik\<^sub>s\<^sub>s\<^sub>t (proj_unl l []) \\<^sub>s\<^sub>e\<^sub>t I = {}" by auto thus ?case using Nil t deducts_eq_if_empty_ik[of t] by argo next case (snoc a A) obtain k b where a: "a = (k,b)" by (metis surj_pair) let ?P = "k = \ \ k = (ln l)" let ?Q = "\s. b = receive\s\" have 0: "ik\<^sub>s\<^sub>s\<^sub>t (proj_unl l (A@[a])) = ik\<^sub>s\<^sub>s\<^sub>t (proj_unl l A)" when "?P \ \?Q" using that ik\<^sub>s\<^sub>s\<^sub>t_snoc_no_receive_eq[OF that, of I "proj_unl l A"] unfolding ik\<^sub>s\<^sub>s\<^sub>t_def a by (cases "k = \ \ k = (ln l)") auto have 1: "declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@[a]) I = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I" when "?P \ \?Q" using that snoc.prems unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_alt_def a by (metis (no_types, lifting) UnCI UnE empty_iff insert_iff list.set prod.inject set_append) note 2 = snoc.prems snoc.IH 0 1 show ?case proof (cases ?P) case True note T = this thus ?thesis proof (cases ?Q) case True thus ?thesis using T unfolding a by blast qed (use 2 in auto) qed (use 2 in auto) qed lemma constr_sem_stateful_proj_priv_term_prefix_obtain: assumes \': "prefix \' \" "constr_sem_stateful \\<^sub>\ (proj_unl n \'@[send\[t]\])" and t: "t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\" "\{} \\<^sub>c t" "t \ \\<^sub>\ = t" obtains B k' s where "k' = \ \ k' = ln n" "prefix B \'" "suffix [(k', receive\s\)] B" "declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \\<^sub>\ = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n B) = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n \')" "constr_sem_stateful \\<^sub>\ (proj_unl n B@[send\[t]\])" "prefix (proj n B) (proj n \)" "suffix [(k', receive\s\)] (proj n B)" "t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n B) \\<^sub>\" proof - have "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n \') \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\ \ t" using \'(2) t(3) strand_sem_append_stateful[of "{}" "{}" "proj_unl n \'" "[send\[t]\]" \\<^sub>\] by simp then obtain B k' s where B: "k' = \ \ k' = ln n" "prefix B \'" "suffix [(k', receive\s\)] B" "declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \\<^sub>\ = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\" "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n B) = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n \')" using deduct_proj_priv_term_prefix_ex_stateful[OF _ t(2), of \\<^sub>\ n \'] unfolding suffix_def by blast have B': "constr_sem_stateful \\<^sub>\ (proj_unl n B@[send\[t]\])" using B(5) \'(2) strand_sem_append_stateful[of "{}" "{}" "proj_unl n \'" "[send\[t]\]" \\<^sub>\] strand_sem_append_stateful[of "{}" "{}" "proj_unl n B" _ \\<^sub>\] prefix_proj(2)[OF B(2), of n] by (metis (no_types, lifting) append_Nil2 prefix_def strand_sem_stateful.simps(2)) have B'': "prefix (proj n B) (proj n \)" "suffix [(k', receive\s\)] (proj n B)" "t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n B) \\<^sub>\" using \' t B(1-4) declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj_eq[of B \\<^sub>\ n] unfolding suffix_def prefix_def proj_def by auto show ?thesis by (rule that[OF B B' B'']) qed lemma constr_sem_stateful_star_proj_no_leakage: fixes Sec P lbls k defines "no_leakage \ \\. \\\<^sub>\ \ s. prefix \ \ \ s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>\ \ \\<^sub>\ \\<^sub>s (unlabel \@[send\[s]\])" assumes Sec: "ground Sec" and \: "\(l,a) \ set \. l = \" shows "no_leakage \" proof (rule ccontr) assume "\no_leakage \" then obtain I B s where B: "prefix B \" "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t B I" "I \\<^sub>s (unlabel B@[send\[s]\])" unfolding no_leakage_def by blast have 1: "\(\{set ts | ts. \\, receive\ts\\ \ set (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)} \ s)" using B(2) unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by fast have 2: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I) \ s" using B(2,3) Sec strand_sem_append_stateful[of "{}" "{}" "unlabel B" "[send\[s]\]" I] - trm_subst_ident[of s I] unlabel_subst[of B] ik\<^sub>s\<^sub>s\<^sub>t_subst[of "unlabel B"] + subst_apply_term_ident[of s I] unlabel_subst[of B] ik\<^sub>s\<^sub>s\<^sub>t_subst[of "unlabel B"] by force have "l = \" when "(l,c) \ set B" for l c using that \ B(1) set_mono_prefix by blast hence "l = \" when "(l,c) \ set (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)" for l c using that unfolding subst_apply_labeled_stateful_strand_def by auto hence 3: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I) = (\{set ts | ts. \\, receive\ts\\ \ set (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I)})" using in_ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t_iff[of _ "B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t I"] unfolding ik\<^sub>s\<^sub>s\<^sub>t_def unlabel_def by auto show False using 1 2 3 by force qed end subsection \Lemmata: Properties of the Constraint Translation Function\ context labeled_stateful_typed_model begin lemma tr_par_labeled_rcv_iff: "B \ set (tr\<^sub>p\<^sub>c A D) \ (i, receive\t\\<^sub>s\<^sub>t) \ set B \ (i, receive\t\) \ set A" by (induct A D arbitrary: B rule: tr\<^sub>p\<^sub>c.induct) auto lemma tr_par_declassified_eq: "B \ set (tr\<^sub>p\<^sub>c A D) \ declassified\<^sub>l\<^sub>s\<^sub>t B I = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I" using tr_par_labeled_rcv_iff unfolding declassified\<^sub>l\<^sub>s\<^sub>t_alt_def declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_alt_def by simp lemma tr_par_ik_eq: assumes "B \ set (tr\<^sub>p\<^sub>c A D)" shows "ik\<^sub>s\<^sub>t (unlabel B) = ik\<^sub>s\<^sub>s\<^sub>t (unlabel A)" proof - have "{t. \i. (i, receive\t\\<^sub>s\<^sub>t) \ set B} = {t. \i. (i, receive\t\) \ set A}" using tr_par_labeled_rcv_iff[OF assms] by simp moreover have "\C. {t. \i. (i, receive\t\\<^sub>s\<^sub>t) \ set C} = {t. receive\t\\<^sub>s\<^sub>t \ set (unlabel C)}" "\C. {t. \i. (i, receive\t\) \ set C} = {t. receive\t\ \ set (unlabel C)}" unfolding unlabel_def by force+ ultimately show ?thesis unfolding ik\<^sub>s\<^sub>s\<^sub>t_def ik\<^sub>s\<^sub>t_is_rcv_set by fast qed lemma tr_par_deduct_iff: assumes "B \ set (tr\<^sub>p\<^sub>c A D)" shows "ik\<^sub>s\<^sub>t (unlabel B) \\<^sub>s\<^sub>e\<^sub>t I \ t \ ik\<^sub>s\<^sub>s\<^sub>t (unlabel A) \\<^sub>s\<^sub>e\<^sub>t I \ t" using tr_par_ik_eq[OF assms] by metis lemma tr_par_vars_subset: assumes "A' \ set (tr\<^sub>p\<^sub>c A D)" shows "fv\<^sub>l\<^sub>s\<^sub>t A' \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" (is ?P) and "bvars\<^sub>l\<^sub>s\<^sub>t A' \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A)" (is ?Q) proof - show ?P using assms proof (induction "unlabel A" arbitrary: A A' D rule: strand_sem_stateful_induct) case (ConsIn A' D ac t s AA A A') then obtain i B where iB: "A = (i,\ac: t \ s\)#B" "AA = unlabel B" unfolding unlabel_def by moura then obtain A'' d where *: "d \ set (dbproj i D)" "A' = (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c B D)" using ConsIn.prems(1) by moura hence "fv\<^sub>l\<^sub>s\<^sub>t A'' \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel B) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" "fv (pair (snd d)) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" apply (metis ConsIn.hyps(1)[OF iB(2)]) using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_mono[OF dbproj_subset[of i D]] fv_pair_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subset[OF *(1)] by blast thus ?case using * iB unfolding pair_def by auto next case (ConsDel A' D t s AA A A') then obtain i B where iB: "A = (i,delete\t,s\)#B" "AA = unlabel B" unfolding unlabel_def by moura define fltD1 where "fltD1 = (\Di. filter (\d. d \ set Di) D)" define fltD2 where "fltD2 = (\Di. filter (\d. d \ set Di) (dbproj i D))" define constr where "constr = (\Di. (map (\d. (i, \check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i, \[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) (fltD2 Di)))" from iB obtain A'' Di where *: "Di \ set (subseqs (dbproj i D))" "A' = (constr Di)@A''" "A'' \ set (tr\<^sub>p\<^sub>c B (fltD1 Di))" using ConsDel.prems(1) unfolding constr_def fltD1_def fltD2_def by moura hence "fv\<^sub>l\<^sub>s\<^sub>t A'' \ fv\<^sub>s\<^sub>s\<^sub>t AA \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel (fltD1 Di))" unfolding constr_def fltD1_def by (metis ConsDel.hyps(1) iB(2)) hence 1: "fv\<^sub>l\<^sub>s\<^sub>t A'' \ fv\<^sub>s\<^sub>s\<^sub>t AA \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_mono[of "unlabel (fltD1 Di)" "unlabel D"] unfolding unlabel_def fltD1_def by force have 2: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel Di) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel (fltD1 Di)) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" using subseqs_set_subset(1)[OF *(1)] unfolding fltD1_def unlabel_def dbproj_def by auto have 5: "fv\<^sub>l\<^sub>s\<^sub>t A' = fv\<^sub>l\<^sub>s\<^sub>t (constr Di) \ fv\<^sub>l\<^sub>s\<^sub>t A''" using * unfolding unlabel_def by force have "fv\<^sub>l\<^sub>s\<^sub>t (constr Di) \ fv t \ fv s \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel Di) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel (fltD1 Di))" unfolding unlabel_def constr_def fltD1_def fltD2_def pair_def dbproj_def by auto hence 3: "fv\<^sub>l\<^sub>s\<^sub>t (constr Di) \ fv t \ fv s \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" using 2 by blast have 4: "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) = fv t \ fv s \ fv\<^sub>s\<^sub>s\<^sub>t AA" using iB by auto have "fv\<^sub>s\<^sub>t (unlabel A') \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" using 1 3 4 5 by blast thus ?case by metis next case (ConsNegChecks A' D X F F' AA A A') then obtain i B where iB: "A = (i,NegChecks X F F')#B" "AA = unlabel B" unfolding unlabel_def by moura define D' where "D' \ \(fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ` set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (unlabel (dbproj i D))))" define constr where "constr = map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" from iB obtain A'' where *: "A'' \ set (tr\<^sub>p\<^sub>c B D)" "A' = constr@A''" using ConsNegChecks.prems(1) unfolding constr_def by moura hence "fv\<^sub>l\<^sub>s\<^sub>t A'' \ fv\<^sub>s\<^sub>s\<^sub>t AA \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" by (metis ConsNegChecks.hyps(1) iB(2)) hence **: "fv\<^sub>l\<^sub>s\<^sub>t A'' \ fv\<^sub>s\<^sub>s\<^sub>t AA \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" by auto have 1: "fv\<^sub>l\<^sub>s\<^sub>t constr \ (D' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) - set X" unfolding D'_def constr_def unlabel_def by auto have "set (unlabel (dbproj i D)) \ set (unlabel D)" unfolding unlabel_def dbproj_def by auto hence 2: "D' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" using tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_vars_subset'[of F' "unlabel (dbproj i D)"] fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_mono unfolding D'_def by blast have 3: "fv\<^sub>l\<^sub>s\<^sub>t A' \ ((fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) - set X) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ fv\<^sub>l\<^sub>s\<^sub>t A''" using 1 2 *(2) unfolding unlabel_def by fastforce have 4: "fv\<^sub>s\<^sub>s\<^sub>t AA \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel A)" by (metis ConsNegChecks.hyps(2) fv\<^sub>s\<^sub>s\<^sub>t_cons_subset) have 5: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel A)" using ConsNegChecks.hyps(2) unfolding unlabel_def by force show ?case using ** 3 4 5 by blast qed (fastforce simp add: unlabel_def)+ show ?Q using assms apply (induct "unlabel A" arbitrary: A A' D rule: strand_sem_stateful_induct) by (fastforce simp add: unlabel_def)+ qed lemma tr_par_vars_disj: assumes "A' \ set (tr\<^sub>p\<^sub>c A D)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" and "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" shows "fv\<^sub>l\<^sub>s\<^sub>t A' \ bvars\<^sub>l\<^sub>s\<^sub>t A' = {}" using assms tr_par_vars_subset by fast lemma tr_par_trms_subset: assumes "A' \ set (tr\<^sub>p\<^sub>c A D)" shows "trms\<^sub>l\<^sub>s\<^sub>t A' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D" using assms proof (induction A D arbitrary: A' rule: tr\<^sub>p\<^sub>c.induct) case 1 thus ?case by simp next case (2 i t A D) then obtain A'' where A'': "A' = (i,send\t\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" by moura hence "trms\<^sub>l\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D" by (metis "2.IH") thus ?case using A'' by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) next case (3 i t A D) then obtain A'' where A'': "A' = (i,receive\t\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" by moura hence "trms\<^sub>l\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D" by (metis "3.IH") thus ?case using A'' by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) next case (4 i ac t t' A D) then obtain A'' where A'': "A' = (i,\ac: t \ t'\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" by moura hence "trms\<^sub>l\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D" by (metis "4.IH") thus ?case using A'' by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) next case (5 i t s A D) hence "A' \ set (tr\<^sub>p\<^sub>c A (List.insert (i,t,s) D))" by simp hence "trms\<^sub>l\<^sub>s\<^sub>t A' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set (List.insert (i,t,s) D)" by (metis "5.IH") thus ?case by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) next case (6 i t s A D) from 6 obtain Di A'' B C where A'': "Di \ set (subseqs (dbproj i D))" "A'' \ set (tr\<^sub>p\<^sub>c A [d\D. d \ set Di])" "A' = (B@C)@A''" "B = map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di" "C = map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\dbproj i D. d \ set Di]" by moura hence "trms\<^sub>l\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set [d\D. d \ set Di]" by (metis "6.IH") moreover have "set [d\D. d \ set Di] \ set D" using set_filter by auto ultimately have "trms\<^sub>l\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D" by blast hence "trms\<^sub>l\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel ((i,delete\t,s\)#A)) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel ((i,delete\t,s\)#A)) \ pair ` snd ` set D" using setops\<^sub>s\<^sub>s\<^sub>t_cons_subset trms\<^sub>s\<^sub>s\<^sub>t_cons by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) moreover have "set Di \ set D" "set [d\dbproj i D . d \ set Di] \ set D" using subseqs_set_subset[OF A''(1)] unfolding dbproj_def by auto hence "trms\<^sub>s\<^sub>t (unlabel B) \ insert (pair (t, s)) (pair ` snd ` set D)" "trms\<^sub>s\<^sub>t (unlabel C) \ insert (pair (t, s)) (pair ` snd ` set D)" using A''(4,5) unfolding unlabel_def by auto hence "trms\<^sub>s\<^sub>t (unlabel (B@C)) \ insert (pair (t,s)) (pair ` snd ` set D)" using unlabel_append[of B C] by auto moreover have "pair (t,s) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (delete\t,s\#unlabel A)" by (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) ultimately show ?case using A''(3) trms\<^sub>s\<^sub>t_append[of "unlabel (B@C)" "unlabel A'"] unlabel_append[of "B@C" A''] by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) next case (7 i ac t s A D) from 7 obtain d A'' where A'': "d \ set (dbproj i D)" "A'' \ set (tr\<^sub>p\<^sub>c A D)" "A' = (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#A''" by moura hence "trms\<^sub>l\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D" by (metis "7.IH") moreover have "trms\<^sub>s\<^sub>t (unlabel A') = {pair (t,s), pair (snd d)} \ trms\<^sub>s\<^sub>t (unlabel A'')" using A''(1,3) by auto ultimately show ?case using A''(1) unfolding dbproj_def by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) next case (8 i X F F' A D) define constr where "constr = map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" define B where "B \ \(trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ` set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D))))" from 8 obtain A'' where A'': "A'' \ set (tr\<^sub>p\<^sub>c A D)" "A' = constr@A''" unfolding constr_def by moura have "trms\<^sub>s\<^sub>t (unlabel A'') \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair`snd`set D" by (metis A''(1) "8.IH") moreover have "trms\<^sub>s\<^sub>t (unlabel constr) \ B \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair ` snd ` set D" unfolding unlabel_def constr_def B_def by auto ultimately have "trms\<^sub>s\<^sub>t (unlabel A') \ B \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D" using A'' unlabel_append[of constr A''] by auto moreover have "set (dbproj i D) \ set D" unfolding dbproj_def by auto hence "B \ pair ` set F' \ pair ` snd ` set D" using tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_trms_subset'[of F' "map snd (dbproj i D)"] unfolding B_def by force moreover have "pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel ((i, \X\\\: F \\: F'\)#A)) = pair ` set F' \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" by auto ultimately show ?case by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) qed lemma tr_par_wf_trms: assumes "A' \ set (tr\<^sub>p\<^sub>c A [])" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t (unlabel A))" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>t A')" using tr_par_trms_subset[OF assms(1)] setops\<^sub>s\<^sub>s\<^sub>t_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s(2)[OF assms(2)] by auto lemma tr_par_wf': assumes "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" and "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X" and "wf'\<^sub>s\<^sub>s\<^sub>t X (unlabel A)" "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" and "A' \ set (tr\<^sub>p\<^sub>c A D)" shows "wf\<^sub>l\<^sub>s\<^sub>t X A'" proof - define P where "P = (\(D::('fun,'var,'lbl) labeleddbstatelist) (A::('fun,'var,'lbl) labeled_stateful_strand). (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}) \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {})" have "P D A" using assms(1,4) by (simp add: P_def) with assms(5,3,2) show ?thesis proof (induction A arbitrary: X A' D) case Nil thus ?case by simp next case (Cons a A) obtain i s where i: "a = (i,s)" by (metis surj_pair) note prems = Cons.prems note IH = Cons.IH show ?case proof (cases s) case (Receive ts) note si = Receive i then obtain A'' where A'': "A' = (i,receive\ts\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ X" using prems unlabel_Cons(1)[of i s A] by moura have *: "wf'\<^sub>s\<^sub>s\<^sub>t X (unlabel A)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X" "P D A" using prems si apply (force, force) using prems(4) si unfolding P_def by fastforce show ?thesis using IH[OF A''(2) *] A''(1,3) by simp next case (Send ts) note si = Send i then obtain A'' where A'': "A' = (i,send\ts\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" using prems by moura have *: "wf'\<^sub>s\<^sub>s\<^sub>t (X \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) (unlabel A)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "P D A" using prems si apply (force, force) using prems(4) si unfolding P_def by fastforce show ?thesis using IH[OF A''(2) *] A''(1) by simp next case (Equality ac t t') note si = Equality i then obtain A'' where A'': "A' = (i,\ac: t \ t'\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" "ac = Assign \ fv t' \ X" using prems unlabel_Cons(1)[of i s] by moura have *: "ac = Assign \ wf'\<^sub>s\<^sub>s\<^sub>t (X \ fv t) (unlabel A)" "ac = Check \ wf'\<^sub>s\<^sub>s\<^sub>t X (unlabel A)" "ac = Assign \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X \ fv t" "ac = Check \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X" "P D A" using prems si apply (force, force, force, force) using prems(4) si unfolding P_def by fastforce show ?thesis using IH[OF A''(2) *(1,3,5)] IH[OF A''(2) *(2,4,5)] A''(1,3) by (cases ac) simp_all next case (Insert t t') note si = Insert i hence A': "A' \ set (tr\<^sub>p\<^sub>c A (List.insert (i,t,t') D))" "fv t \ X" "fv t' \ X" using prems by auto have *: "wf'\<^sub>s\<^sub>s\<^sub>t X (unlabel A)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel (List.insert (i,t,t') D)) \ X" using prems si by (auto simp add: unlabel_def) have **: "P (List.insert (i,t,t') D) A" using prems(4) si unfolding P_def unlabel_def by fastforce show ?thesis using IH[OF A'(1) * **] A'(2,3) by simp next case (Delete t t') note si = Delete i define constr where "constr = (\Di. (map (\d. (i,\check: (pair (t,t')) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i,\[]\\\: [(pair (t,t'), pair (snd d))]\\<^sub>s\<^sub>t)) [d\dbproj i D. d \ set Di]))" from prems si obtain Di A'' where A'': "A' = constr Di@A''" "A'' \ set (tr\<^sub>p\<^sub>c A [d\D. d \ set Di])" "Di \ set (subseqs (dbproj i D))" unfolding constr_def by auto have *: "wf'\<^sub>s\<^sub>s\<^sub>t X (unlabel A)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel (filter (\d. d \ set Di) D)) \ X" using prems si apply simp using prems si by (fastforce simp add: unlabel_def) have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel (filter (\d. d \ set Di) D)) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" by (auto simp add: unlabel_def) hence **: "P [d\D. d \ set Di] A" using prems si unfolding P_def by fastforce have ***: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X" using prems si by auto show ?thesis using IH[OF A''(2) * **] A''(1) wf_pair_eqs_ineqs_map'[OF _ A''(3) ***] unfolding constr_def by simp next case (InSet ac t t') note si = InSet i then obtain d A'' where A'': "A' = (i,\ac: (pair (t,t')) \ (pair (snd d))\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" "d \ set D" using prems by (auto simp add: dbproj_def) have *: "ac = Assign \ wf'\<^sub>s\<^sub>s\<^sub>t (X \ fv t \ fv t') (unlabel A)" "ac = Check \ wf'\<^sub>s\<^sub>s\<^sub>t X (unlabel A)" "ac = Assign \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X \ fv t \ fv t'" "ac = Check \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X" "P D A" using prems si apply (force, force, force, force) using prems(4) si unfolding P_def by fastforce have **: "fv (pair (snd d)) \ X" using A''(3) prems(3) fv_pair_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subset by fast have ***: "fv (pair (t,t')) = fv t \ fv t'" unfolding pair_def by auto show ?thesis using IH[OF A''(2) *(1,3,5)] IH[OF A''(2) *(2,4,5)] A''(1) ** *** by (cases ac) (simp_all add: Un_assoc) next case (NegChecks Y F F') note si = NegChecks i then obtain A'' where A'': "A' = (map (\G. (i,\Y\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D))))@A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" using prems by moura have *: "wf'\<^sub>s\<^sub>s\<^sub>t X (unlabel A)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X" using prems si by auto have "bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel ((i,\Y\\\: F \\: F'\)#A))" "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel ((i,\Y\\\: F \\: F'\)#A))" by auto hence **: "P D A" using prems si unfolding P_def by blast show ?thesis using IH[OF A''(2) * **] A''(1) wf_pair_negchecks_map' by simp qed qed qed lemma tr_par_wf: assumes "A' \ set (tr\<^sub>p\<^sub>c A [])" and "wf\<^sub>s\<^sub>s\<^sub>t (unlabel A)" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" shows "wf\<^sub>l\<^sub>s\<^sub>t {} A'" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>t A')" and "fv\<^sub>l\<^sub>s\<^sub>t A' \ bvars\<^sub>l\<^sub>s\<^sub>t A' = {}" using tr_par_wf'[OF _ _ _ _ assms(1)] tr_par_wf_trms[OF assms(1,3)] tr_par_vars_disj[OF assms(1)] assms(2) by fastforce+ lemma tr_par_proj: assumes "B \ set (tr\<^sub>p\<^sub>c A D)" shows "proj n B \ set (tr\<^sub>p\<^sub>c (proj n A) (proj n D))" using assms proof (induction A D arbitrary: B rule: tr\<^sub>p\<^sub>c.induct) case (5 i t s S D) note prems = "5.prems" note IH = "5.IH" have IH': "proj n B \ set (tr\<^sub>p\<^sub>c (proj n S) (proj n (List.insert (i,t,s) D)))" using prems IH by auto show ?case proof (cases "(i = ln n) \ (i = \)") case True thus ?thesis using IH' proj_list_insert(1,2)[of n "(t,s)" D] proj_list_Cons(1,2)[of n _ S] by auto next case False then obtain m where "i = ln m" "n \ m" by (cases i) simp_all thus ?thesis using IH' proj_list_insert(3)[of n _ "(t,s)" D] proj_list_Cons(3)[of n _ "insert\t,s\" S] by auto qed next case (6 i t s S D) note prems = "6.prems" note IH = "6.IH" define constr where "constr = (\Di D. (map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\dbproj i D. d \ set Di]))" obtain Di B' where B': "B = constr Di D@B'" "Di \ set (subseqs (dbproj i D))" "B' \ set (tr\<^sub>p\<^sub>c S [d\D. d \ set Di])" using prems constr_def by fastforce hence "proj n B' \ set (tr\<^sub>p\<^sub>c (proj n S) (proj n [d\D. d \ set Di]))" using IH by simp hence IH': "proj n B' \ set (tr\<^sub>p\<^sub>c (proj n S) [d\proj n D. d \ set Di])" by (metis proj_filter) show ?case proof (cases "(i = ln n) \ (i = \)") case True hence "proj n B = constr Di D@proj n B'" "Di \ set (subseqs (dbproj i (proj n D)))" using B'(1,2) proj_dbproj(1,2)[of n D] unfolding proj_def constr_def by auto moreover have "constr Di (proj n D) = constr Di D" using True proj_dbproj(1,2)[of n D] unfolding constr_def by presburger ultimately have "proj n B \ set (tr\<^sub>p\<^sub>c ((i, delete\t,s\)#proj n S) (proj n D))" using IH' unfolding constr_def by force thus ?thesis by (metis proj_list_Cons(1,2) True) next case False then obtain m where m: "i = ln m" "n \ m" by (cases i) simp_all hence *: "(ln n) \ i" by simp have "proj n B = proj n B'" using B'(1) False unfolding constr_def proj_def by auto moreover have "[d\proj n D. d \ set Di] = proj n D" using proj_subseq[OF _ m(2)[symmetric]] m(1) B'(2) by simp ultimately show ?thesis using m(1) IH' proj_list_Cons(3)[OF m(2), of _ S] by auto qed next case (7 i ac t s S D) note prems = "7.prems" note IH = "7.IH" define constr where "constr = ( \d::'lbl strand_label \ ('fun,'var) term \ ('fun,'var) term. (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t))" obtain d B' where B': "B = constr d#B'" "d \ set (dbproj i D)" "B' \ set (tr\<^sub>p\<^sub>c S D)" using prems constr_def by fastforce hence IH': "proj n B' \ set (tr\<^sub>p\<^sub>c (proj n S) (proj n D))" using IH by auto show ?case proof (cases "(i = ln n) \ (i = \)") case True hence "proj n B = constr d#proj n B'" "d \ set (dbproj i (proj n D))" using B' proj_list_Cons(1,2)[of n _ B'] unfolding constr_def by (force, metis proj_dbproj(1,2)) hence "proj n B \ set (tr\<^sub>p\<^sub>c ((i, InSet ac t s)#proj n S) (proj n D))" using IH' unfolding constr_def by auto thus ?thesis using proj_list_Cons(1,2)[of n _ S] True by metis next case False then obtain m where m: "i = ln m" "n \ m" by (cases i) simp_all hence "proj n B = proj n B'" using B'(1) proj_list_Cons(3) unfolding constr_def by auto thus ?thesis using IH' m proj_list_Cons(3)[OF m(2), of "InSet ac t s" S] unfolding constr_def by auto qed next case (8 i X F F' S D) note prems = "8.prems" note IH = "8.IH" define constr where "constr = (\D. map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D))))" obtain B' where B': "B = constr D@B'" "B' \ set (tr\<^sub>p\<^sub>c S D)" using prems constr_def by fastforce hence IH': "proj n B' \ set (tr\<^sub>p\<^sub>c (proj n S) (proj n D))" using IH by auto show ?case proof (cases "(i = ln n) \ (i = \)") case True hence "proj n B = constr (proj n D)@proj n B'" using B'(1,2) proj_dbproj(1,2)[of n D] unfolding proj_def constr_def by auto hence "proj n B \ set (tr\<^sub>p\<^sub>c ((i, NegChecks X F F')#proj n S) (proj n D))" using IH' unfolding constr_def by auto thus ?thesis using proj_list_Cons(1,2)[of n _ S] True by metis next case False then obtain m where m: "i = ln m" "n \ m" by (cases i) simp_all hence "proj n B = proj n B'" using B'(1) unfolding constr_def proj_def by auto thus ?thesis using IH' m proj_list_Cons(3)[OF m(2), of "NegChecks X F F'" S] unfolding constr_def by auto qed qed (force simp add: proj_def)+ lemma tr_par_preserves_par_comp: assumes "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A Sec" "A' \ set (tr\<^sub>p\<^sub>c A [])" shows "par_comp A' Sec" proof - let ?M = "\l. trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l A)" let ?N = "\l. trms_proj\<^sub>l\<^sub>s\<^sub>t l A'" have 0: "\l1 l2. l1 \ l2 \ GSMP_disjoint (?M l1) (?M l2) Sec" using assms(1) unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by simp_all { fix l1 l2::'lbl assume *: "l1 \ l2" hence "GSMP_disjoint (?M l1) (?M l2) Sec" using 0(1) by metis moreover have "pair ` snd ` set (proj n []) = {}" for n::'lbl unfolding proj_def by simp hence "?N l1 \ ?M l1" "?N l2 \ ?M l2" using tr_par_trms_subset[OF tr_par_proj[OF assms(2)]] by (metis Un_empty_right)+ ultimately have "GSMP_disjoint (?N l1) (?N l2) Sec" using GSMP_disjoint_subset by presburger } hence 1: "\l1 l2. l1 \ l2 \ GSMP_disjoint (trms_proj\<^sub>l\<^sub>s\<^sub>t l1 A') (trms_proj\<^sub>l\<^sub>s\<^sub>t l2 A') Sec" using 0(1) by metis have 2: "ground Sec" "\s \ Sec. \{} \\<^sub>c s" using assms(1) unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by metis+ show ?thesis using 1 2 unfolding par_comp_def by metis qed lemma tr_preserves_receives: assumes "E \ set (tr\<^sub>p\<^sub>c F D)" "(l, receive\t\) \ set F" shows "(l, receive\t\\<^sub>s\<^sub>t) \ set E" using assms by (induct F D arbitrary: E rule: tr\<^sub>p\<^sub>c.induct) auto lemma tr_preserves_last_receive: assumes "E \ set (tr\<^sub>p\<^sub>c F D)" "suffix [(l, receive\t\\<^sub>s\<^sub>t)] E" shows "\G. suffix ((l, receive\t\)#G) F \ list_all (Not \ is_Receive \ snd) G" (is "\G. ?P G F \ ?Q G") using assms proof (induction F D arbitrary: E rule: tr\<^sub>p\<^sub>c.induct) case (1 D) thus ?case by simp next case (2 i t' S D) note prems = "2.prems" note IH = "2.IH" obtain E' where E': "E = (i,send\t'\\<^sub>s\<^sub>t)#E'" "E' \ set (tr\<^sub>p\<^sub>c S D)" using prems(1) by auto obtain G where G: "?P G S" "?Q G" using suffix_Cons'[OF prems(2)[unfolded E'(1)]] IH[OF E'(2)] by blast show ?case by (metis suffix_Cons G) next case (3 i t' S D) note prems = "3.prems" note IH = "3.IH" obtain E' where E': "E = (i,receive\t'\\<^sub>s\<^sub>t)#E'" "E' \ set (tr\<^sub>p\<^sub>c S D)" using prems(1) by auto show ?case using suffix_Cons'[OF prems(2)[unfolded E'(1)]] proof assume "suffix [(l, receive\t\\<^sub>s\<^sub>t)] E'" then obtain G where G: "?P G S" "?Q G" using IH[OF E'(2)] by blast show ?thesis by (metis suffix_Cons G) next assume "(i, receive\t'\\<^sub>s\<^sub>t) = (l, receive\t\\<^sub>s\<^sub>t) \ E' = []" hence *: "i = l" "t' = t" "E' = []" by simp_all show ?thesis using tr_preserves_receives[OF E'(2)] unfolding * list_all_iff is_Receive_def by fastforce qed next case (4 i ac t' t'' S D) note prems = "4.prems" note IH = "4.IH" obtain E' where E': "E = (i,\ac: t' \ t''\\<^sub>s\<^sub>t)#E'" "E' \ set (tr\<^sub>p\<^sub>c S D)" using prems(1) by auto obtain G where G: "?P G S" "?Q G" using suffix_Cons'[OF prems(2)[unfolded E'(1)]] IH[OF E'(2)] by blast show ?case by (metis suffix_Cons G) next case (5 i t' s S D) note prems = "5.prems" note IH = "5.IH" have "E \ set (tr\<^sub>p\<^sub>c S (List.insert (i,t',s) D))" using prems(1) by auto thus ?case by (metis IH[OF _ prems(2)] suffix_Cons) next case (6 i t' s S D) note prems = "6.prems" note IH = "6.IH" define constr where "constr = (\Di. (map (\d. (i,\check: (pair (t',s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i,\[]\\\: [(pair (t',s), pair (snd d))]\\<^sub>s\<^sub>t)) (filter (\d. d \ set Di) (dbproj i D))))" obtain E' Di where E': "E = constr Di@E'" "E' \ set (tr\<^sub>p\<^sub>c S (filter (\d. d \ set Di) D))" "Di \ set (subseqs (dbproj i D))" using prems(1) unfolding constr_def by auto have "receive\t\\<^sub>s\<^sub>t \ snd ` set (constr Di)" unfolding constr_def by force hence "\suffix [(l, receive\t\\<^sub>s\<^sub>t)] (constr Di)" unfolding suffix_def by auto hence "1 \ length E'" using prems(2) E'(1) by (cases E') auto hence "suffix [(l, receive\t\\<^sub>s\<^sub>t)] E'" using suffix_length_suffix[OF prems(2) suffixI[OF E'(1)]] by simp thus ?case by (metis IH[OF E'(3,2)] suffix_Cons) next case (7 i ac t' s S D) note prems = "7.prems" note IH = "7.IH" define constr where "constr = ( \d::(('lbl strand_label \ ('fun,'var) term \ ('fun,'var) term)). (i,\ac: (pair (t',s)) \ (pair (snd d))\\<^sub>s\<^sub>t))" obtain E' d where E': "E = constr d#E'" "E' \ set (tr\<^sub>p\<^sub>c S D)" "d \ set (dbproj i D)" using prems(1) unfolding constr_def by auto have "receive\t\\<^sub>s\<^sub>t \ snd (constr d)" unfolding constr_def by force hence "suffix [(l, receive\t\\<^sub>s\<^sub>t)] E'" using prems(2) E'(1) suffix_Cons' by fastforce thus ?case by (metis IH[OF E'(2)] suffix_Cons) next case (8 i X G G' S D) note prems = "8.prems" note IH = "8.IH" define constr where "constr = map (\H. (i,\X\\\: (G@H)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G' (map snd (dbproj i D)))" obtain E' where E': "E = constr@E'" "E' \ set (tr\<^sub>p\<^sub>c S D)" using prems(1) constr_def by auto have "receive\t\\<^sub>s\<^sub>t \ snd ` set constr" unfolding constr_def by force hence "\suffix [(l, receive\t\\<^sub>s\<^sub>t)] constr" unfolding suffix_def by auto hence "1 \ length E'" using prems(2) E'(1) by (cases E') auto hence "suffix [(l, receive\t\\<^sub>s\<^sub>t)] E'" using suffix_length_suffix[OF prems(2) suffixI[OF E'(1)]] by simp thus ?case by (metis IH[OF E'(2)] suffix_Cons) qed lemma tr_leaking_prefix_exists: assumes "A' \ set (tr\<^sub>p\<^sub>c A [])" "prefix B A'" "ik\<^sub>s\<^sub>t (proj_unl n B) \\<^sub>s\<^sub>e\<^sub>t \ \ t" shows "\C D. prefix C B \ prefix D A \ C \ set (tr\<^sub>p\<^sub>c D []) \ (ik\<^sub>s\<^sub>t (proj_unl n C) \\<^sub>s\<^sub>e\<^sub>t \ \ t) \ (\{} \\<^sub>c t \ (\l s G. suffix ((l, receive\s\)#G) D \ list_all (Not \ is_Receive \ snd) G))" proof - let ?P = "\B C C'. B = C@C' \ (\n t. (n, receive\t\\<^sub>s\<^sub>t) \ set C') \ (C = [] \ (\n t. suffix [(n,receive\t\\<^sub>s\<^sub>t)] C))" have "\C C'. ?P B C C'" proof (induction B) case (Cons b B) then obtain C C' n s where *: "?P B C C'" "b = (n,s)" by moura show ?case proof (cases "C = []") case True note T = True show ?thesis proof (cases "\t. s = receive\t\\<^sub>s\<^sub>t") case True hence "?P (b#B) [b] C'" using * T by auto thus ?thesis by metis next case False hence "?P (b#B) [] (b#C')" using * T by auto thus ?thesis by metis qed next case False hence "?P (b#B) (b#C) C'" using * unfolding suffix_def by auto thus ?thesis by metis qed qed simp then obtain C C' where C: "B = C@C'" "\n t. (n, receive\t\\<^sub>s\<^sub>t) \ set C'" "C = [] \ (\n t. suffix [(n,receive\t\\<^sub>s\<^sub>t)] C)" by moura hence 1: "prefix C B" by simp hence 2: "prefix C A'" using assms(2) by simp have "\m t. (m,receive\t\\<^sub>s\<^sub>t) \ set B \ (m,receive\t\\<^sub>s\<^sub>t) \ set C" using C by auto hence "\t. receive\t\\<^sub>s\<^sub>t \ set (proj_unl n B) \ receive\t\\<^sub>s\<^sub>t \ set (proj_unl n C)" unfolding unlabel_def proj_def by force hence "ik\<^sub>s\<^sub>t (proj_unl n B) \ ik\<^sub>s\<^sub>t (proj_unl n C)" using ik\<^sub>s\<^sub>t_is_rcv_set by blast hence 3: "ik\<^sub>s\<^sub>t (proj_unl n C) \\<^sub>s\<^sub>e\<^sub>t \ \ t" by (metis ideduct_mono[OF assms(3)] subst_all_mono) have "\F. prefix F A \ E \ set (tr\<^sub>p\<^sub>c F D)" when "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E" "prefix E A'" "A' \ set (tr\<^sub>p\<^sub>c A D)" for D E m t using that proof (induction A D arbitrary: A' E rule: tr\<^sub>p\<^sub>c.induct) case (1 D) thus ?case by simp next case (2 i t' S D) note prems = "2.prems" note IH = "2.IH" obtain A'' where *: "A' = (i,send\t'\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c S D)" using prems(3) by auto have "E \ []" using prems(1) by auto then obtain E' where **: "E = (i,send\t'\\<^sub>s\<^sub>t)#E'" using *(1) prems(2) by (cases E) auto hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'" "prefix E' A''" using *(1) prems(1,2) suffix_Cons[of _ _ E'] by auto then obtain F where "prefix F S" "E' \ set (tr\<^sub>p\<^sub>c F D)" using *(2) ** IH by metis hence "prefix ((i,Send t')#F) ((i,Send t')#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,Send t')#F) D)" using ** by auto thus ?case by metis next case (3 i t' S D) note prems = "3.prems" note IH = "3.IH" obtain A'' where *: "A' = (i,receive\t'\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c S D)" using prems(3) by auto have "E \ []" using prems(1) by auto then obtain E' where **: "E = (i,receive\t'\\<^sub>s\<^sub>t)#E'" using *(1) prems(2) by (cases E) auto show ?case proof (cases "(m, receive\t\\<^sub>s\<^sub>t) = (i, receive\t'\\<^sub>s\<^sub>t)") case True note T = True show ?thesis proof (cases "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'") case True hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'" "prefix E' A''" using ** *(1) prems(1,2) by auto then obtain F where "prefix F S" "E' \ set (tr\<^sub>p\<^sub>c F D)" using *(2) ** IH by metis hence "prefix ((i,receive\t'\)#F) ((i,receive\t'\)#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,receive\t'\)#F) D)" using ** by auto thus ?thesis by metis next case False hence "E' = []" using **(1) T prems(1) suffix_Cons[of "[(m, receive\t\\<^sub>s\<^sub>t)]" "(m, receive\t\\<^sub>s\<^sub>t)" E'] by auto hence "prefix [(i,receive\t'\)] ((i,receive\t'\) # S) \ E \ set (tr\<^sub>p\<^sub>c [(i,receive\t'\)] D)" using * ** prems by auto thus ?thesis by metis qed next case False hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'" "prefix E' A''" using ** *(1) prems(1,2) suffix_Cons[of _ _ E'] by auto then obtain F where "prefix F S" "E' \ set (tr\<^sub>p\<^sub>c F D)" using *(2) ** IH by metis hence "prefix ((i,receive\t'\)#F) ((i,receive\t'\)#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,receive\t'\)#F) D)" using ** by auto thus ?thesis by metis qed next case (4 i ac t' t'' S D) note prems = "4.prems" note IH = "4.IH" obtain A'' where *: "A' = (i,\ac: t' \ t''\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c S D)" using prems(3) by auto have "E \ []" using prems(1) by auto then obtain E' where **: "E = (i,\ac: t' \ t''\\<^sub>s\<^sub>t)#E'" using *(1) prems(2) by (cases E) auto hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'" "prefix E' A''" using *(1) prems(1,2) suffix_Cons[of _ _ E'] by auto then obtain F where "prefix F S" "E' \ set (tr\<^sub>p\<^sub>c F D)" using *(2) ** IH by metis hence "prefix ((i,Equality ac t' t'')#F) ((i,Equality ac t' t'')#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,Equality ac t' t'')#F) D)" using ** by auto thus ?case by metis next case (5 i t' s S D) note prems = "5.prems" note IH = "5.IH" have *: "A' \ set (tr\<^sub>p\<^sub>c S (List.insert (i,t',s) D))" using prems(3) by auto have "E \ []" using prems(1) by auto hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E" "prefix E A'" using *(1) prems(1,2) suffix_Cons[of _ _ E] by auto then obtain F where "prefix F S" "E \ set (tr\<^sub>p\<^sub>c F (List.insert (i,t',s) D))" using * IH by metis hence "prefix ((i,insert\t',s\)#F) ((i,insert\t',s\)#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,insert\t',s\)#F) D)" by auto thus ?case by metis next case (6 i t' s S D) note prems = "6.prems" note IH = "6.IH" define constr where "constr = (\Di. (map (\d. (i,\check: (pair (t',s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i,\[]\\\: [(pair (t',s), pair (snd d))]\\<^sub>s\<^sub>t)) (filter (\d. d \ set Di) (dbproj i D))))" obtain A'' Di where *: "A' = constr Di@A''" "A'' \ set (tr\<^sub>p\<^sub>c S (filter (\d. d \ set Di) D))" "Di \ set (subseqs (dbproj i D))" using prems(3) constr_def by auto have ***: "(m, receive\t\\<^sub>s\<^sub>t) \ set (constr Di)" using constr_def by auto have "E \ []" using prems(1) by auto then obtain E' where **: "E = constr Di@E'" using *(1) prems(1,2) *** by (metis (mono_tags, lifting) Un_iff list.set_intros(1) prefixI prefix_def prefix_same_cases set_append suffix_def) hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'" "prefix E' A''" using *(1) prems(1,2) suffix_append[of "[(m,receive\t\\<^sub>s\<^sub>t)]" "constr Di" E'] *** by (metis (no_types, opaque_lifting) Nil_suffix append_Nil2 in_set_conv_decomp rev_exhaust snoc_suffix_snoc suffix_appendD, auto) then obtain F where "prefix F S" "E' \ set (tr\<^sub>p\<^sub>c F (filter (\d. d \ set Di) D))" using *(2,3) ** IH by metis hence "prefix ((i,delete\t',s\)#F) ((i,delete\t',s\)#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,delete\t',s\)#F) D)" using *(3) ** constr_def by auto thus ?case by metis next case (7 i ac t' s S D) note prems = "7.prems" note IH = "7.IH" define constr where "constr = ( \d::(('lbl strand_label \ ('fun,'var) term \ ('fun,'var) term)). (i,\ac: (pair (t',s)) \ (pair (snd d))\\<^sub>s\<^sub>t))" obtain A'' d where *: "A' = constr d#A''" "A'' \ set (tr\<^sub>p\<^sub>c S D)" "d \ set (dbproj i D)" using prems(3) constr_def by auto have "E \ []" using prems(1) by auto then obtain E' where **: "E = constr d#E'" using *(1) prems(2) by (cases E) auto hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'" "prefix E' A''" using *(1) prems(1,2) suffix_Cons[of _ _ E'] using constr_def by auto then obtain F where "prefix F S" "E' \ set (tr\<^sub>p\<^sub>c F D)" using *(2) ** IH by metis hence "prefix ((i,InSet ac t' s)#F) ((i,InSet ac t' s)#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,InSet ac t' s)#F) D)" using *(3) ** unfolding constr_def by auto thus ?case by metis next case (8 i X G G' S D) note prems = "8.prems" note IH = "8.IH" define constr where "constr = map (\H. (i,\X\\\: (G@H)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G' (map snd (dbproj i D)))" obtain A'' where *: "A' = constr@A''" "A'' \ set (tr\<^sub>p\<^sub>c S D)" using prems(3) constr_def by auto have ***: "(m, receive\t\\<^sub>s\<^sub>t) \ set constr" using constr_def by auto have "E \ []" using prems(1) by auto then obtain E' where **: "E = constr@E'" using *(1) prems(1,2) *** by (metis (mono_tags, lifting) Un_iff list.set_intros(1) prefixI prefix_def prefix_same_cases set_append suffix_def) hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'" "prefix E' A''" using *(1) prems(1,2) suffix_append[of "[(m,receive\t\\<^sub>s\<^sub>t)]" constr E'] *** by (metis (no_types, opaque_lifting) Nil_suffix append_Nil2 in_set_conv_decomp rev_exhaust snoc_suffix_snoc suffix_appendD, auto) then obtain F where "prefix F S" "E' \ set (tr\<^sub>p\<^sub>c F D)" using *(2) ** IH by metis hence "prefix ((i,NegChecks X G G')#F) ((i,NegChecks X G G')#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,NegChecks X G G')#F) D)" using ** constr_def by auto thus ?case by metis qed moreover have "prefix [] A" "[] \ set (tr\<^sub>p\<^sub>c [] [])" by auto moreover have "{} \\<^sub>c t" when "C = []" using 3 by (simp add: deducts_eq_if_empty_ik that) ultimately have 4: "\D. prefix D A \ C \ set (tr\<^sub>p\<^sub>c D []) \ (\{} \\<^sub>c t \ (\l s G. suffix ((l, receive\s\)#G) D \ list_all (Not \ is_Receive \ snd) G))" using C(3) assms(1) 2 by (meson tr_preserves_last_receive) show ?thesis by (metis 1 3 4) qed end context labeled_stateful_typing begin lemma tr_par_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p: assumes "A' \ set (tr\<^sub>p\<^sub>c A D)" "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel A)" and "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" (is "?P0 A D") and "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" (is "?P1 A D") and "\t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D. \t' \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D. (\\. Unifier \ t t') \ \ t = \ t'" (is "?P3 A D") shows "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A')" proof - have sublmm: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel A)" "?P0 A D" "?P1 A D" "?P3 A D" when p: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (a#A))" "?P0 (a#A) D" "?P1 (a#A) D" "?P3 (a#A) D" for a A D proof - show "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel A)" using p(1) by (simp add: unlabel_def tfr\<^sub>s\<^sub>s\<^sub>t_def) show "?P0 A D" using p(2) fv\<^sub>s\<^sub>s\<^sub>t_cons_subset unfolding unlabel_def by fastforce show "?P1 A D" using p(3) bvars\<^sub>s\<^sub>s\<^sub>t_cons_subset unfolding unlabel_def by fastforce have "setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ setops\<^sub>s\<^sub>s\<^sub>t (unlabel (a#A))" using setops\<^sub>s\<^sub>s\<^sub>t_cons_subset unfolding unlabel_def by auto thus "?P3 A D" using p(4) by blast qed show ?thesis using assms proof (induction A D arbitrary: A' rule: tr\<^sub>p\<^sub>c.induct) case 1 thus ?case by simp next case (2 i t A D) note prems = "2.prems" note IH = "2.IH" from prems(1) obtain A'' where A'': "A' = (i,send\t\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" by moura have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A'')" using IH[OF A''(2)] prems(5) sublmm[OF prems(2,3,4,5)] by meson thus ?case using A''(1) by simp next case (3 i t A D) note prems = "3.prems" note IH = "3.IH" from prems(1) obtain A'' where A'': "A' = (i,receive\t\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" by moura have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A'')" using IH[OF A''(2)] prems(5) sublmm[OF prems(2,3,4,5)] by meson thus ?case using A''(1) by simp next case (4 i ac t t' A D) note prems = "4.prems" note IH = "4.IH" from prems(1) obtain A'' where A'': "A' = (i,\ac: t \ t'\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" by moura have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A'')" using IH[OF A''(2)] prems(5) sublmm[OF prems(2,3,4,5)] by meson thus ?case using A''(1) prems(2) by simp next case (5 i t s A D) note prems = "5.prems" note IH = "5.IH" from prems(1) have A': "A' \ set (tr\<^sub>p\<^sub>c A (List.insert (i,t,s) D))" by simp have 1: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel A)" using sublmm[OF prems(2,3,4,5)] by simp have "pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel ((i,insert\t,s\)#A)) \ pair`snd`set D = pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair`snd`set (List.insert (i,t,s) D)" by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) hence 3: "?P3 A (List.insert (i,t,s) D)" using prems(5) by metis moreover have "?P1 A (List.insert (i,t,s) D)" using prems(3,4) bvars\<^sub>s\<^sub>s\<^sub>t_cons_subset[of "unlabel A" "insert\t,s\"] unfolding unlabel_def by fastforce ultimately have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A')" using IH[OF A' sublmm(1,2)[OF prems(2,3,4,5)] _ 3] by metis thus ?case using A'(1) by auto next case (6 i t s A D) note prems = "6.prems" note IH = "6.IH" define constr where constr: "constr \ (\Di. (map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) (filter (\d. d \ set Di) (dbproj i D))))" from prems(1) obtain Di A'' where A'': "A' = constr Di@A''" "A'' \ set (tr\<^sub>p\<^sub>c A (filter (\d. d \ set Di) D))" "Di \ set (subseqs (dbproj i D))" unfolding constr by fastforce define Q1 where "Q1 \ (\(F::(('fun,'var) term \ ('fun,'var) term) list) X. \x \ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) - set X. \a. \ (Var x) = TAtom a)" define Q2 where "Q2 \ (\(F::(('fun,'var) term \ ('fun,'var) term) list) X. \f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ T = [] \ (\s \ set T. s \ Var ` set X))" have "pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair`snd`set [d\D. d \ set Di] \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel ((i,delete\t,s\)#A)) \ pair`snd`set D" using subseqs_set_subset[OF A''(3)] by (force simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) moreover have "\a\M. \b\M. P a b" when "M \ N" "\a\N. \b\N. P a b" for M N::"('fun, 'var) terms" and P using that by blast ultimately have *: "?P3 A (filter (\d. d \ set Di) D)" using prems(5) by presburger have **: "?P1 A (filter (\d. d \ set Di) D)" using prems(4) bvars\<^sub>s\<^sub>s\<^sub>t_cons_subset[of "unlabel A" "delete\t,s\"] unfolding unlabel_def by fastforce have 1: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A'')" using IH[OF A''(3,2) sublmm(1,2)[OF prems(2,3,4,5)] ** *] by metis have 2: "\ac: u \ u'\\<^sub>s\<^sub>t \ set (unlabel A'') \ (\d \ set Di. u = pair (t,s) \ u' = pair (snd d))" when "\ac: u \ u'\\<^sub>s\<^sub>t \ set (unlabel A')" for ac u u' using that A''(1) unfolding constr unlabel_def by force have 3: "\X\\\: u\\<^sub>s\<^sub>t \ set (unlabel A'') \ (\d \ set (filter (\d. d \ set Di) D). u = [(pair (t,s), pair (snd d))] \ Q2 u X)" when "\X\\\: u\\<^sub>s\<^sub>t \ set (unlabel A')" for X u using that A''(1) unfolding Q2_def constr unlabel_def dbproj_def by force have 4: "\d\set D. (\\. Unifier \ (pair (t,s)) (pair (snd d))) \ \ (pair (t,s)) = \ (pair (snd d))" using prems(5) by (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) { fix ac u u' assume a: "\ac: u \ u'\\<^sub>s\<^sub>t \ set (unlabel A')" "\\. Unifier \ u u'" hence "\ac: u \ u'\\<^sub>s\<^sub>t \ set (unlabel A'') \ (\d \ set Di. u = pair (t,s) \ u' = pair (snd d))" using 2 by metis moreover { assume "\ac: u \ u'\\<^sub>s\<^sub>t \ set (unlabel A'')" hence "tfr\<^sub>s\<^sub>t\<^sub>p (\ac: u \ u'\\<^sub>s\<^sub>t)" using 1 Ball_set_list_all[of "unlabel A''" tfr\<^sub>s\<^sub>t\<^sub>p] by fast } moreover { fix d assume "d \ set Di" "u = pair (t,s)" "u' = pair (snd d)" hence "\\. Unifier \ u u' \ \ u = \ u'" using 4 dbproj_subseq_subset A''(3) by fast hence "tfr\<^sub>s\<^sub>t\<^sub>p (\ac: u \ u'\\<^sub>s\<^sub>t)" using Ball_set_list_all[of "unlabel A''" tfr\<^sub>s\<^sub>t\<^sub>p] by simp hence "\ u = \ u'" using tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def[of "unlabel A''"] using a(2) unfolding unlabel_def by auto } ultimately have "\ u = \ u'" using tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def[of "unlabel A''"] a(2) unfolding unlabel_def by auto } moreover { fix u U assume "\U\\\: u\\<^sub>s\<^sub>t \ set (unlabel A')" hence "\U\\\: u\\<^sub>s\<^sub>t \ set (unlabel A'') \ (\d \ set (filter (\d. d \ set Di) D). u = [(pair (t,s), pair (snd d))] \ Q2 u U)" using 3 by metis hence "Q1 u U \ Q2 u U" using 1 4 subseqs_set_subset[OF A''(3)] tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def[of "unlabel A''"] unfolding Q1_def Q2_def by blast } ultimately show ?case using tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def[of "unlabel A'"] unfolding Q1_def Q2_def unlabel_def by blast next case (7 i ac t s A D) note prems = "7.prems" note IH = "7.IH" from prems(1) obtain d A'' where A'': "A' = (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" "d \ set (dbproj i D)" by moura have 1: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A'')" using IH[OF A''(2) sublmm(1,2,3)[OF prems(2,3,4,5)] sublmm(4)[OF prems(2,3,4,5)]] by metis have 2: "\ (pair (t,s)) = \ (pair (snd d))" when "\\. Unifier \ (pair (t,s)) (pair (snd d))" using that prems(2,5) A''(3) unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def dbproj_def) show ?case using A''(1) 1 2 by fastforce next case (8 i X F F' A D) note prems = "8.prems" note IH = "8.IH" define constr where "constr = map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" define Q1 where "Q1 \ (\(F::(('fun,'var) term \ ('fun,'var) term) list) X. \x \ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) - set X. \a. \ (Var x) = TAtom a)" define Q2 where "Q2 \ (\(M::('fun,'var) terms) X. \f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t M \ T = [] \ (\s \ set T. s \ Var ` set X))" have Q2_subset: "Q2 M' X" when "M' \ M" "Q2 M X" for X M M' using that unfolding Q2_def by auto have Q2_supset: "Q2 (M \ M') X" when "Q2 M X" "Q2 M' X" for X M M' using that unfolding Q2_def by auto from prems obtain A'' where A'': "A' = constr@A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" using constr_def by moura have 0: "constr = [(i,\X\\\: F\\<^sub>s\<^sub>t)]" when "F' = []" using that unfolding constr_def by simp have 1: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A'')" using IH[OF A''(2) sublmm(1,2,3)[OF prems(2,3,4,5)] sublmm(4)[OF prems(2,3,4,5)]] by metis have 2: "(F' = [] \ Q1 F X) \ Q2 (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair ` set F') X" using prems(2) unfolding Q1_def Q2_def by simp have 3: "F' = [] \ Q1 F X \ list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel constr)" using 0 2 tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def[of "unlabel constr"] unfolding Q1_def by auto { fix c assume "c \ set (unlabel constr)" hence "\G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D))). c = \X\\\: (F@G)\\<^sub>s\<^sub>t" unfolding constr_def unlabel_def by force } moreover { fix G assume G: "G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" and c: "\X\\\: (F@G)\\<^sub>s\<^sub>t \ set (unlabel constr)" and e: "Q2 (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair ` set F') X" have d_Q2: "Q2 (pair ` set (map snd D)) X" unfolding Q2_def proof (intro allI impI) fix f T assume "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (pair ` set (map snd D))" then obtain d where d: "d \ set (map snd D)" "Fun f T \ subterms (pair d)" by force hence "fv (pair d) \ set X = {}" using prems(4) unfolding pair_def by (force simp add: unlabel_def) thus "T = [] \ (\s \ set T. s \ Var ` set X)" by (metis fv_disj_Fun_subterm_param_cases d(2)) qed have "trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F@G) \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair ` set F' \ pair ` set (map snd D)" using tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_trms_subset[OF G] unfolding dbproj_def by force hence "Q2 (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F@G)) X" using Q2_subset[OF _ Q2_supset[OF e d_Q2]] by metis hence "tfr\<^sub>s\<^sub>t\<^sub>p (\X\\\: (F@G)\\<^sub>s\<^sub>t)" by (metis Q2_def tfr\<^sub>s\<^sub>t\<^sub>p.simps(2)) } ultimately have 4: "Q2 (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair ` set F') X \ list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel constr)" using Ball_set by blast have 5: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel constr)" using 2 3 4 by metis show ?case using 1 5 A''(1) by (simp add: unlabel_def) qed qed lemma tr_par_tfr: assumes "A' \ set (tr\<^sub>p\<^sub>c A [])" and "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel A)" and "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" shows "tfr\<^sub>s\<^sub>t (unlabel A')" proof - have *: "trms\<^sub>l\<^sub>s\<^sub>t A' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" using tr_par_trms_subset[OF assms(1)] by simp hence "SMP (trms\<^sub>l\<^sub>s\<^sub>t A') \ SMP (trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A))" using SMP_mono by simp moreover have "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A))" using assms(2) unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by fast ultimately have 1: "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>t A')" by (metis tfr_subset(2)[OF _ *]) have **: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel A)" using assms(2) unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by fast have "pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ SMP (trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)) - Var`\" using setops\<^sub>s\<^sub>s\<^sub>t_are_pairs unfolding pair_def by auto hence "\ t = \ t'" when "\\. Unifier \ t t'" "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "t' \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" for t t' using that assms(2) unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def tfr\<^sub>s\<^sub>e\<^sub>t_def by blast moreover have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel []) = {}" "pair ` snd ` set [] = {}" by auto ultimately have 2: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A')" using tr_par_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p[OF assms(1) ** assms(3)] by simp show ?thesis by (metis 1 2 tfr\<^sub>s\<^sub>t_def) qed lemma tr_par_preserves_typing_cond: assumes "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A Sec" "typing_cond\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "A' \ set (tr\<^sub>p\<^sub>c A [])" shows "typing_cond (unlabel A')" proof - have "wf'\<^sub>s\<^sub>s\<^sub>t {} (unlabel A)" "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t (unlabel A))" using assms(2) unfolding typing_cond\<^sub>s\<^sub>s\<^sub>t_def by simp_all hence 1: "wf\<^sub>s\<^sub>t {} (unlabel A')" "fv\<^sub>s\<^sub>t (unlabel A') \ bvars\<^sub>s\<^sub>t (unlabel A') = {}" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t (unlabel A'))" "Ana_invar_subst (ik\<^sub>s\<^sub>t (unlabel A') \ assignment_rhs\<^sub>s\<^sub>t (unlabel A'))" using tr_par_wf[OF assms(3)] Ana_invar_subst' by metis+ have 2: "tfr\<^sub>s\<^sub>t (unlabel A')" by (metis tr_par_tfr assms(2,3) typing_cond\<^sub>s\<^sub>s\<^sub>t_def) show ?thesis by (metis 1 2 typing_cond_def) qed end subsection \Theorem: Semantic Equivalence of Translation\ context labeled_stateful_typed_model begin context begin text \ An alternative version of the translation that does not perform database-state projections. It is used as an intermediate step in the proof of semantic equivalence/correctness. \ private fun tr'\<^sub>p\<^sub>c:: "('fun,'var,'lbl) labeled_stateful_strand \ ('fun,'var,'lbl) labeleddbstatelist \ ('fun,'var,'lbl) labeled_strand list" where "tr'\<^sub>p\<^sub>c [] D = [[]]" | "tr'\<^sub>p\<^sub>c ((i,send\ts\)#A) D = map ((#) (i,send\ts\\<^sub>s\<^sub>t)) (tr'\<^sub>p\<^sub>c A D)" | "tr'\<^sub>p\<^sub>c ((i,receive\ts\)#A) D = map ((#) (i,receive\ts\\<^sub>s\<^sub>t)) (tr'\<^sub>p\<^sub>c A D)" | "tr'\<^sub>p\<^sub>c ((i,\ac: t \ t'\)#A) D = map ((#) (i,\ac: t \ t'\\<^sub>s\<^sub>t)) (tr'\<^sub>p\<^sub>c A D)" | "tr'\<^sub>p\<^sub>c ((i,insert\t,s\)#A) D = tr'\<^sub>p\<^sub>c A (List.insert (i,(t,s)) D)" | "tr'\<^sub>p\<^sub>c ((i,delete\t,s\)#A) D = ( concat (map (\Di. map (\B. (map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\D. d \ set Di])@B) (tr'\<^sub>p\<^sub>c A [d\D. d \ set Di])) (subseqs D)))" | "tr'\<^sub>p\<^sub>c ((i,\ac: t \ s\)#A) D = concat (map (\B. map (\d. (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#B) D) (tr'\<^sub>p\<^sub>c A D))" | "tr'\<^sub>p\<^sub>c ((i,\X\\\: F \\: F'\)#A) D = map ((@) (map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd D)))) (tr'\<^sub>p\<^sub>c A D)" subsubsection \Part 1\ private lemma tr'_par_iff_unlabel_tr: assumes "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. p = q \ i = j" shows "(\C \ set (tr'\<^sub>p\<^sub>c A D). B = unlabel C) \ B \ set (tr (unlabel A) (unlabel D))" (is "?A \ ?B") proof { fix C have "C \ set (tr'\<^sub>p\<^sub>c A D) \ unlabel C \ set (tr (unlabel A) (unlabel D))" using assms proof (induction A D arbitrary: C rule: tr'\<^sub>p\<^sub>c.induct) case (5 i t s S D) hence "unlabel C \ set (tr (unlabel S) (unlabel (List.insert (i, t, s) D)))" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) moreover have "insert (i,t,s) (set D) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i,insert\t,s\)#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "\(j,p) \ insert (i,t,s) (set D). \(k,q) \ insert (i,t,s) (set D). p = q \ j = k" using "5.prems"(2) by blast hence "unlabel (List.insert (i, t, s) D) = (List.insert (t, s) (unlabel D))" using map_snd_list_insert_distrib[of "(i,t,s)" D] unfolding unlabel_def by simp ultimately show ?case by auto next case (6 i t s S D) let ?f1 = "\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t" let ?g1 = "\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t" let ?f2 = "\d. (i, ?f1 (snd d))" let ?g2 = "\d. (i, ?g1 (snd d))" define constr1 where "constr1 = (\Di. (map ?f1 Di)@(map ?g1 [d\unlabel D. d \ set Di]))" define constr2 where "constr2 = (\Di. (map ?f2 Di)@(map ?g2 [d\D. d \ set Di]))" obtain C' Di where C': "Di \ set (subseqs D)" "C = constr2 Di@C'" "C' \ set (tr'\<^sub>p\<^sub>c S [d\D. d \ set Di])" using "6.prems"(1) unfolding constr2_def by moura have 0: "set [d\D. d \ set Di] \ set D" "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S)" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence 1: "\(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set [d\D. d \ set Di]. \(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set [d\D. d \ set Di]. p = q \ j = k" using "6.prems"(2) by blast have "\(i,p) \ set D \ set Di. \(j,q) \ set D \ set Di. p = q \ i = j" using "6.prems"(2) subseqs_set_subset(1)[OF C'(1)] by blast hence 2: "unlabel [d\D. d \ set Di] = [d\unlabel D. d \ set (unlabel Di)]" using unlabel_filter_eq[of D "set Di"] unfolding unlabel_def by simp have 3: "\f g::('a \ 'a \ 'c). \A B::(('b \ 'a \ 'a) list). map snd ((map (\d. (i, f (snd d))) A)@(map (\d. (i, g (snd d))) B)) = map f (map snd A)@map g (map snd B)" by simp have "unlabel (constr2 Di) = constr1 (unlabel Di)" using 2 3[of ?f1 Di ?g1 "[d\D. d \ set Di]"] by (simp add: constr1_def constr2_def unlabel_def) hence 4: "unlabel C = constr1 (unlabel Di)@unlabel C'" using C'(2) unlabel_append by metis have "unlabel Di \ set (map unlabel (subseqs D))" using C'(1) unfolding unlabel_def by simp hence 5: "unlabel Di \ set (subseqs (unlabel D))" using map_subseqs[of snd D] unfolding unlabel_def by simp show ?case using "6.IH"[OF C'(1,3) 1] 2 4 5 unfolding constr1_def by auto next case (7 i ac t s S D) obtain C' d where C': "C = (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#C'" "C' \ set (tr'\<^sub>p\<^sub>c S D)" "d \ set D" using "7.prems"(1) by moura have "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i,InSet ac t s)#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "\(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D. \(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D. p = q \ j = k" using "7.prems"(2) by blast hence "unlabel C' \ set (tr (unlabel S) (unlabel D))" using "7.IH"[OF C'(2)] by auto thus ?case using C' unfolding unlabel_def by force next case (8 i X F F' S D) obtain C' where C': "C = map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd D))@C'" "C' \ set (tr'\<^sub>p\<^sub>c S D)" using "8.prems"(1) by moura have "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i,NegChecks X F F')#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "\(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D. \(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D. p = q \ j = k" using "8.prems"(2) by blast hence "unlabel C' \ set (tr (unlabel S) (unlabel D))" using "8.IH"[OF C'(2)] by auto thus ?case using C' unfolding unlabel_def by auto qed (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) } thus "?A \ ?B" by blast show "?B \ ?A" using assms proof (induction A arbitrary: B D) case (Cons a A) obtain ia sa where a: "a = (ia,sa)" by moura have "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A)" using a by (cases sa) (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence 1: "\(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. \(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. p = q \ j = k" using Cons.prems(2) by blast show ?case proof (cases sa) case (Send t) then obtain B' where B': "B = send\t\\<^sub>s\<^sub>t#B'" "B' \ set (tr (unlabel A) (unlabel D))" using Cons.prems(1) a by auto thus ?thesis using Cons.IH[OF B'(2) 1] a B'(1) Send by auto next case (Receive t) then obtain B' where B': "B = receive\t\\<^sub>s\<^sub>t#B'" "B' \ set (tr (unlabel A) (unlabel D))" using Cons.prems(1) a by auto thus ?thesis using Cons.IH[OF B'(2) 1] a B'(1) Receive by auto next case (Equality ac t t') then obtain B' where B': "B = \ac: t \ t'\\<^sub>s\<^sub>t#B'" "B' \ set (tr (unlabel A) (unlabel D))" using Cons.prems(1) a by auto thus ?thesis using Cons.IH[OF B'(2) 1] a B'(1) Equality by auto next case (Insert t s) hence B: "B \ set (tr (unlabel A) (List.insert (t,s) (unlabel D)))" using Cons.prems(1) a by auto let ?P = "\i. List.insert (t,s) (unlabel D) = unlabel (List.insert (i,t,s) D)" { obtain j where j: "?P j" "j = ia \ (j,t,s) \ set D" using labeled_list_insert_eq_ex_cases[of "(t,s)" D ia] by moura hence "j = ia" using Cons.prems(2) a Insert by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "?P ia" using j(1) by metis } hence j: "?P ia" by metis have 2: "\(k1, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set (List.insert (ia,t,s) D). \(k2, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set (List.insert (ia,t,s) D). p = q \ k1 = k2" using Cons.prems(2) a Insert by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) show ?thesis using Cons.IH[OF _ 2] j(1) B Insert a by auto next case (Delete t s) define c where "c \ (\(i::'lbl strand_label) Di. map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di@ map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\D. d \ set Di])" define d where "d \ (\Di. map (\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t) Di@ map (\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) [d\unlabel D. d \ set Di])" obtain B' Di where B': "B = d Di@B'" "Di \ set (subseqs (unlabel D))" "B' \ set (tr (unlabel A) [d\unlabel D. d \ set Di])" using Cons.prems(1) a Delete unfolding d_def by auto obtain Di' where Di': "Di' \ set (subseqs D)" "unlabel Di' = Di" using unlabel_subseqsD[OF B'(2)] by moura have 2: "\(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set [d\D. d \ set Di']. \(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set [d\D. d \ set Di']. p = q \ j = k" using 1 subseqs_subset[OF Di'(1)] filter_is_subset[of "\d. d \ set Di'"] by blast have "set Di' \ set D" by (rule subseqs_subset[OF Di'(1)]) hence "\(j, p)\set D \ set Di'. \(k, q)\set D \ set Di'. p = q \ j = k" using Cons.prems(2) by blast hence 3: "[d\unlabel D. d \ set Di] = unlabel [d\D. d \ set Di']" using Di'(2) unlabel_filter_eq[of D "set Di'"] unfolding unlabel_def by auto obtain C where C: "C \ set (tr'\<^sub>p\<^sub>c A [d\D. d \ set Di'])" "B' = unlabel C" using 3 Cons.IH[OF _ 2] B'(3) by auto hence 4: "c ia Di'@C \ set (tr'\<^sub>p\<^sub>c (a#A) D)" using Di'(1) a Delete unfolding c_def by auto have "unlabel (c ia Di') = d Di" using Di' 3 unfolding c_def d_def unlabel_def by auto hence 5: "B = unlabel (c ia Di'@C)" using B'(1) C(2) unlabel_append[of "c ia Di'" C] by simp show ?thesis using 4 5 by blast next case (InSet ac t s) then obtain B' d where B': "B = \ac: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t#B'" "B' \ set (tr (unlabel A) (unlabel D))" "d \ set (unlabel D)" using Cons.prems(1) a by auto thus ?thesis using Cons.IH[OF _ 1] a InSet unfolding unlabel_def by auto next case (NegChecks X F F') then obtain B' where B': "B = map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (unlabel D))@B'" "B' \ set (tr (unlabel A) (unlabel D))" using Cons.prems(1) a by auto thus ?thesis using Cons.IH[OF _ 1] a NegChecks unfolding unlabel_def by auto qed qed simp qed subsubsection \Part 2\ private lemma tr_par_iff_tr'_par: assumes "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. (\\. Unifier \ (pair p) (pair q)) \ i = j" (is "?R3 A D") and "\(l,t,s) \ set D. (fv t \ fv s) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" (is "?R4 A D") and "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" (is "?R5 A D") shows "(\B \ set (tr\<^sub>p\<^sub>c A D). \M; unlabel B\\<^sub>d \) \ (\C \ set (tr'\<^sub>p\<^sub>c A D). \M; unlabel C\\<^sub>d \)" (is "?P \ ?Q") proof { fix B assume "B \ set (tr\<^sub>p\<^sub>c A D)" "\M; unlabel B\\<^sub>d \" hence ?Q using assms proof (induction A D arbitrary: B M rule: tr\<^sub>p\<^sub>c.induct) case (1 D) thus ?case by simp next case (2 i ts S D) note prems = "2.prems" note IH = "2.IH" obtain B' where B': "B = (i,send\ts\\<^sub>s\<^sub>t)#B'" "B' \ set (tr\<^sub>p\<^sub>c S D)" using prems(1) by moura have 1: "\M; unlabel B'\\<^sub>d \" using prems(2) B'(1) by simp have 4: "?R3 S D" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force have 7: "\t \ set ts. M \ t \ \" using prems(2) B'(1) by simp obtain C where C: "C \ set (tr'\<^sub>p\<^sub>c S D)" "\M; unlabel C\\<^sub>d \" using IH[OF B'(2) 1 4 5 6] by moura hence "((i,send\ts\\<^sub>s\<^sub>t)#C) \ set (tr'\<^sub>p\<^sub>c ((i,Send ts)#S) D)" "\M; unlabel ((i,send\ts\\<^sub>s\<^sub>t)#C)\\<^sub>d \" using 7 by auto thus ?case by metis next case (3 i ts S D) note prems = "3.prems" note IH = "3.IH" obtain B' where B': "B = (i,receive\ts\\<^sub>s\<^sub>t)#B'" "B' \ set (tr\<^sub>p\<^sub>c S D)" using prems(1) by moura have 1: "\(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; unlabel B'\\<^sub>d \ " using prems(2) B'(1) by simp have 4: "?R3 S D" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force obtain C where C: "C \ set (tr'\<^sub>p\<^sub>c S D)" "\(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; unlabel C\\<^sub>d \" using IH[OF B'(2) 1 4 5 6] by moura hence "((i,receive\ts\\<^sub>s\<^sub>t)#C) \ set (tr'\<^sub>p\<^sub>c ((i,receive\ts\)#S) D)" "\(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; unlabel ((i,receive\ts\\<^sub>s\<^sub>t)#C)\\<^sub>d \" by auto thus ?case by auto next case (4 i ac t t' S D) note prems = "4.prems" note IH = "4.IH" obtain B' where B': "B = (i,\ac: t \ t'\\<^sub>s\<^sub>t)#B'" "B' \ set (tr\<^sub>p\<^sub>c S D)" using prems(1) by moura have 1: "\M; unlabel B'\\<^sub>d \ " using prems(2) B'(1) by simp have 4: "?R3 S D" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force have 7: "t \ \ = t' \ \" using prems(2) B'(1) by simp obtain C where C: "C \ set (tr'\<^sub>p\<^sub>c S D)" "\M; unlabel C\\<^sub>d \" using IH[OF B'(2) 1 4 5 6] by moura hence "((i,\ac: t \ t'\\<^sub>s\<^sub>t)#C) \ set (tr'\<^sub>p\<^sub>c ((i,Equality ac t t')#S) D)" "\M; unlabel ((i,\ac: t \ t'\\<^sub>s\<^sub>t)#C)\\<^sub>d \" using 7 by auto thus ?case by metis next case (5 i t s S D) note prems = "5.prems" note IH = "5.IH" have B: "B \ set (tr\<^sub>p\<^sub>c S (List.insert (i,t,s) D))" using prems(1) by simp have 1: "\M; unlabel B\\<^sub>d \ " using prems(2) B(1) by simp have 4: "?R3 S (List.insert (i,t,s) D)" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S (List.insert (i,t,s) D)" using prems(4,5) by force have 6: "?R5 S D" using prems(5) by force show ?case using IH[OF B(1) 1 4 5 6] by simp next case (6 i t s S D) note prems = "6.prems" note IH = "6.IH" let ?cl1 = "\Di. map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di" let ?cu1 = "\Di. map (\d. \check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t) Di" let ?cl2 = "\Di. map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\dbproj i D. d\set Di]" let ?cu2 = "\Di. map (\d. \[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\dbproj i D. d\set Di]" let ?dl1 = "\Di. map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di" let ?du1 = "\Di. map (\d. \check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t) Di" let ?dl2 = "\Di. map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\D. d\set Di]" let ?du2 = "\Di. map (\d. \[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\D. d\set Di]" define c where c: "c = (\Di. ?cl1 Di@?cl2 Di)" define d where d: "d = (\Di. ?dl1 Di@?dl2 Di)" obtain B' Di where B': "Di \ set (subseqs (dbproj i D))" "B = c Di@B'" "B' \ set (tr\<^sub>p\<^sub>c S [d\D. d \ set Di])" using prems(1) c by moura have 0: "ik\<^sub>s\<^sub>t (unlabel (c Di)) = {}" "ik\<^sub>s\<^sub>t (unlabel (d Di)) = {}" "unlabel (?cl1 Di) = ?cu1 Di" "unlabel (?cl2 Di) = ?cu2 Di" "unlabel (?dl1 Di) = ?du1 Di" "unlabel (?dl2 Di) = ?du2 Di" unfolding c d unlabel_def by force+ have 1: "\M; unlabel B'\\<^sub>d \ " using prems(2) B'(2) 0(1) unfolding unlabel_def by auto { fix j p k q assume "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set [d\D. d \ set Di]" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set [d\D. d \ set Di]" hence "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" using dbproj_subseq_subset[OF B'(1)] by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence 4: "?R3 S [d\D. d \ set Di]" by blast have 5: "?R4 S (filter (\d. d \ set Di) D)" using prems(4) by force have 6: "?R5 S D" using prems(5) by force obtain C where C: "C \ set (tr'\<^sub>p\<^sub>c S [d\D . d \ set Di])" "\M; unlabel C\\<^sub>d \" using IH[OF B'(1,3) 1 4 5 6] by moura have 7: "\M; unlabel (c Di)\\<^sub>d \" "\M; unlabel B'\\<^sub>d \" using prems(2) B'(2) 0(1) strand_sem_split(3,4)[of M "unlabel (c Di)" "unlabel B'"] unfolding c unlabel_def by auto have "\M; unlabel (?cl2 Di)\\<^sub>d \" using 7(1) 0(1) unfolding c unlabel_def by auto hence "\M; ?cu2 Di\\<^sub>d \" by (metis 0(4)) moreover { fix j p k q assume "(j, p) \ {(i, t, s)} \ set D \ set Di" "(k, q) \ {(i, t, s)} \ set D \ set Di" hence "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" using dbproj_subseq_subset[OF B'(1)] by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence "\(j, p) \ {(i, t, s)} \ set D \ set Di. \(k, q) \ {(i, t, s)} \ set D \ set Di. (\\. Unifier \ (pair p) (pair q)) \ j = k" by blast ultimately have "\M; ?du2 Di\\<^sub>d \" using labeled_sat_ineq_lift by simp hence "\M; unlabel (?dl2 Di)\\<^sub>d \" by (metis 0(6)) moreover have "\M; unlabel (?cl1 Di)\\<^sub>d \" using 7(1) unfolding c unlabel_def by auto hence "\M; unlabel (?dl1 Di)\\<^sub>d \" by (metis 0(3,5)) ultimately have "\M; unlabel (d Di)\\<^sub>d \" using 0(2) unfolding c d unlabel_def by force hence 8: "\M; unlabel (d Di@C)\\<^sub>d \" using 0(2) C(2) unfolding unlabel_def by auto have 9: "d Di@C \ set (tr'\<^sub>p\<^sub>c ((i,delete\t,s\)#S) D)" using C(1) dbproj_subseq_in_subseqs[OF B'(1)] unfolding d unlabel_def by auto show ?case by (metis 8 9) next case (7 i ac t s S D) note prems = "7.prems" note IH = "7.IH" obtain B' d where B': "B = (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#B'" "B' \ set (tr\<^sub>p\<^sub>c S D)" "d \ set (dbproj i D)" using prems(1) by moura have 1: "\M; unlabel B'\\<^sub>d \ " using prems(2) B'(1) by simp { fix j p k q assume "(j,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D" "(k,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D" hence "(j,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, InSet ac t s)#S) \ set D" "(k,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, InSet ac t s)#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence 4: "?R3 S D" by blast have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force have 7: "pair (t,s) \ \ = pair (snd d) \ \" using prems(2) B'(1) by simp obtain C where C: "C \ set (tr'\<^sub>p\<^sub>c S D)" "\M; unlabel C\\<^sub>d \" using IH[OF B'(2) 1 4 5 6] by moura hence "((i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#C) \ set (tr'\<^sub>p\<^sub>c ((i,InSet ac t s)#S) D)" "\M; unlabel ((i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#C)\\<^sub>d \" using 7 B'(3) unfolding dbproj_def by auto thus ?case by metis next case (8 i X F F' S D) note prems = "8.prems" note IH = "8.IH" let ?cl = "map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" let ?cu = "map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" let ?dl = "map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd D))" let ?du = "map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd D))" define c where c: "c = ?cl" define d where d: "d = ?dl" obtain B' where B': "B = c@B'" "B' \ set (tr\<^sub>p\<^sub>c S D)" using prems(1) c by moura have 0: "ik\<^sub>s\<^sub>t (unlabel c) = {}" "ik\<^sub>s\<^sub>t (unlabel d) = {}" "unlabel ?cl = ?cu" "unlabel ?dl = ?du" unfolding c d unlabel_def by force+ have "ik\<^sub>s\<^sub>t (unlabel c) = {}" unfolding c unlabel_def by force hence 1: "\M; unlabel B'\\<^sub>d \ " using prems(2) B'(1) unfolding unlabel_def by auto have "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, NegChecks X F F')#S)" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence 4: "?R3 S D" using prems(3) by blast have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force obtain C where C: "C \ set (tr'\<^sub>p\<^sub>c S D)" "\M; unlabel C\\<^sub>d \" using IH[OF B'(2) 1 4 5 6] by moura have 7: "\M; unlabel c\\<^sub>d \" "\M; unlabel B'\\<^sub>d \" using prems(2) B'(1) 0(1) strand_sem_split(3,4)[of M "unlabel c" "unlabel B'"] unfolding c unlabel_def by auto have 8: "d@C \ set (tr'\<^sub>p\<^sub>c ((i,NegChecks X F F')#S) D)" using C(1) unfolding d unlabel_def by auto have "\M; unlabel ?cl\\<^sub>d \" using 7(1) unfolding c unlabel_def by auto hence "\M; ?cu\\<^sub>d \" by (metis 0(3)) moreover { fix j p k q assume "(j, p) \ ((\(t,s). (i,t,s)) ` set F') \ set D" "(k, q) \ ((\(t,s). (i,t,s)) ` set F') \ set D" hence "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, NegChecks X F F')#S) \ set D" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, NegChecks X F F')#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence "\(j, p) \ ((\(t,s). (i,t,s)) ` set F') \ set D. \(k, q) \ ((\(t,s). (i,t,s)) ` set F') \ set D. (\\. Unifier \ (pair p) (pair q)) \ j = k" by blast moreover have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (map snd D) \ set X = {}" using prems(4) by fastforce ultimately have "\M; ?du\\<^sub>d \" using labeled_sat_ineq_dbproj_sem_equiv[of i] by simp hence "\M; unlabel ?dl\\<^sub>d \" by (metis 0(4)) hence "\M; unlabel d\\<^sub>d \" using 0(2) unfolding c d unlabel_def by force hence 9: "\M; unlabel (d@C)\\<^sub>d \" using 0(2) C(2) unfolding unlabel_def by auto show ?case by (metis 8 9) qed } thus "?P \ ?Q" by metis { fix C assume "C \ set (tr'\<^sub>p\<^sub>c A D)" "\M; unlabel C\\<^sub>d \" hence ?P using assms proof (induction A D arbitrary: C M rule: tr'\<^sub>p\<^sub>c.induct) case (1 D) thus ?case by simp next case (2 i ts S D) note prems = "2.prems" note IH = "2.IH" obtain C' where C': "C = (i,send\ts\\<^sub>s\<^sub>t)#C'" "C' \ set (tr'\<^sub>p\<^sub>c S D)" using prems(1) by moura have 1: "\M; unlabel C'\\<^sub>d \ " using prems(2) C'(1) by simp have 4: "?R3 S D" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force have 7: "\t \ set ts. M \ t \ \" using prems(2) C'(1) by simp obtain B where B: "B \ set (tr\<^sub>p\<^sub>c S D)" "\M; unlabel B\\<^sub>d \" using IH[OF C'(2) 1 4 5 6] by moura hence "((i,send\ts\\<^sub>s\<^sub>t)#B) \ set (tr\<^sub>p\<^sub>c ((i,Send ts)#S) D)" "\M; unlabel ((i,send\ts\\<^sub>s\<^sub>t)#B)\\<^sub>d \" using 7 by auto thus ?case by metis next case (3 i ts S D) note prems = "3.prems" note IH = "3.IH" obtain C' where C': "C = (i,receive\ts\\<^sub>s\<^sub>t)#C'" "C' \ set (tr'\<^sub>p\<^sub>c S D)" using prems(1) by moura have 1: "\(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; unlabel C'\\<^sub>d \ " using prems(2) C'(1) by simp have 4: "?R3 S D" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force obtain B where B: "B \ set (tr\<^sub>p\<^sub>c S D)" "\(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; unlabel B\\<^sub>d \" using IH[OF C'(2) 1 4 5 6] by moura hence "((i,receive\ts\\<^sub>s\<^sub>t)#B) \ set (tr\<^sub>p\<^sub>c ((i,receive\ts\)#S) D)" "\(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; unlabel ((i,receive\ts\\<^sub>s\<^sub>t)#B)\\<^sub>d \" by auto thus ?case by auto next case (4 i ac t t' S D) note prems = "4.prems" note IH = "4.IH" obtain C' where C': "C = (i,\ac: t \ t'\\<^sub>s\<^sub>t)#C'" "C' \ set (tr'\<^sub>p\<^sub>c S D)" using prems(1) by moura have 1: "\M; unlabel C'\\<^sub>d \ " using prems(2) C'(1) by simp have 4: "?R3 S D" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force have 7: "t \ \ = t' \ \" using prems(2) C'(1) by simp obtain B where B: "B \ set (tr\<^sub>p\<^sub>c S D)" "\M; unlabel B\\<^sub>d \" using IH[OF C'(2) 1 4 5 6] by moura hence "((i,\ac: t \ t'\\<^sub>s\<^sub>t)#B) \ set (tr\<^sub>p\<^sub>c ((i,Equality ac t t')#S) D)" "\M; unlabel ((i,\ac: t \ t'\\<^sub>s\<^sub>t)#B)\\<^sub>d \" using 7 by auto thus ?case by metis next case (5 i t s S D) note prems = "5.prems" note IH = "5.IH" have C: "C \ set (tr'\<^sub>p\<^sub>c S (List.insert (i,t,s) D))" using prems(1) by simp have 1: "\M; unlabel C\\<^sub>d \ " using prems(2) C(1) by simp have 4: "?R3 S (List.insert (i,t,s) D)" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S (List.insert (i,t,s) D)" using prems(4,5) by force have 6: "?R5 S (List.insert (i,t,s) D)" using prems(5) by force show ?case using IH[OF C(1) 1 4 5 6] by simp next case (6 i t s S D) note prems = "6.prems" note IH = "6.IH" let ?dl1 = "\Di. map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di" let ?du1 = "\Di. map (\d. \check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t) Di" let ?dl2 = "\Di. map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\dbproj i D. d\set Di]" let ?du2 = "\Di. map (\d. \[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\dbproj i D. d\set Di]" let ?cl1 = "\Di. map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di" let ?cu1 = "\Di. map (\d. \check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t) Di" let ?cl2 = "\Di. map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\D. d\set Di]" let ?cu2 = "\Di. map (\d. \[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\D. d\set Di]" define c where c: "c = (\Di. ?cl1 Di@?cl2 Di)" define d where d: "d = (\Di. ?dl1 Di@?dl2 Di)" obtain C' Di where C': "Di \ set (subseqs D)" "C = c Di@C'" "C' \ set (tr'\<^sub>p\<^sub>c S [d\D. d \ set Di])" using prems(1) c by moura have 0: "ik\<^sub>s\<^sub>t (unlabel (c Di)) = {}" "ik\<^sub>s\<^sub>t (unlabel (d Di)) = {}" "unlabel (?cl1 Di) = ?cu1 Di" "unlabel (?cl2 Di) = ?cu2 Di" "unlabel (?dl1 Di) = ?du1 Di" "unlabel (?dl2 Di) = ?du2 Di" unfolding c d unlabel_def by force+ have 1: "\M; unlabel C'\\<^sub>d \ " using prems(2) C'(2) 0(1) unfolding unlabel_def by auto { fix j p k q assume "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set [d\D. d \ set Di]" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set [d\D. d \ set Di]" hence "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence 4: "?R3 S [d\D. d \ set Di]" by blast have 5: "?R4 S (filter (\d. d \ set Di) D)" using prems(4) by force have 6: "?R5 S D" using prems(5) by force obtain B where B: "B \ set (tr\<^sub>p\<^sub>c S [d\D. d \ set Di])" "\M; unlabel B\\<^sub>d \" using IH[OF C'(1,3) 1 4 5 6] by moura have 7: "\M; unlabel (c Di)\\<^sub>d \" "\M; unlabel C'\\<^sub>d \" using prems(2) C'(2) 0(1) strand_sem_split(3,4)[of M "unlabel (c Di)" "unlabel C'"] unfolding c unlabel_def by auto { fix j p k q assume "(j, p) \ {(i, t, s)} \ set D" "(k, q) \ {(i, t, s)} \ set D" hence "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence "\(j, p) \ {(i, t, s)} \ set D. \(k, q) \ {(i, t, s)} \ set D. (\\. Unifier \ (pair p) (pair q)) \ j = k" by blast moreover have "\M; unlabel (?cl1 Di)\\<^sub>d \" using 7(1) unfolding c unlabel_append by auto hence "\M; ?cu1 Di\\<^sub>d \" by (metis 0(3)) ultimately have *: "Di \ set (subseqs (dbproj i D))" using labeled_sat_eqs_subseqs[OF C'(1)] by simp hence 8: "d Di@B \ set (tr\<^sub>p\<^sub>c ((i,delete\t,s\)#S) D)" using B(1) unfolding d unlabel_def by auto have "\M; unlabel (?cl2 Di)\\<^sub>d \" using 7(1) 0(1) unfolding c unlabel_def by auto hence "\M; ?cu2 Di\\<^sub>d \" by (metis 0(4)) hence "\M; ?du2 Di\\<^sub>d \" by (metis labeled_sat_ineq_dbproj) hence "\M; unlabel (?dl2 Di)\\<^sub>d \" by (metis 0(6)) moreover have "\M; unlabel (?cl1 Di)\\<^sub>d \" using 7(1) unfolding c unlabel_def by auto hence "\M; unlabel (?dl1 Di)\\<^sub>d \" by (metis 0(3,5)) ultimately have "\M; unlabel (d Di)\\<^sub>d \" using 0(2) unfolding c d unlabel_def by force hence 9: "\M; unlabel (d Di@B)\\<^sub>d \" using 0(2) B(2) unfolding unlabel_def by auto show ?case by (metis 8 9) next case (7 i ac t s S D) note prems = "7.prems" note IH = "7.IH" obtain C' d where C': "C = (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#C'" "C' \ set (tr'\<^sub>p\<^sub>c S D)" "d \ set D" using prems(1) by moura have 1: "\M; unlabel C'\\<^sub>d \ " using prems(2) C'(1) by simp { fix j p k q assume "(j,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D" "(k,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D" hence "(j,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, InSet ac t s)#S) \ set D" "(k,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, InSet ac t s)#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence 4: "?R3 S D" by blast have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force obtain B where B: "B \ set (tr\<^sub>p\<^sub>c S D)" "\M; unlabel B\\<^sub>d \" using IH[OF C'(2) 1 4 5 6] by moura have 7: "pair (t,s) \ \ = pair (snd d) \ \" using prems(2) C'(1) by simp have "(i,t,s) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, InSet ac t s)#S) \ set D" "(fst d, snd d) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, InSet ac t s)#S) \ set D" using C'(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "\\. Unifier \ (pair (t,s)) (pair (snd d)) \ i = fst d" using prems(3) by blast hence "fst d = i" using 7 by auto hence 8: "d \ set (dbproj i D)" using C'(3) unfolding dbproj_def by auto have 9: "((i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#B) \ set (tr\<^sub>p\<^sub>c ((i,InSet ac t s)#S) D)" using B 8 by auto have 10: "\M; unlabel ((i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#B)\\<^sub>d \" using B 7 8 by auto show ?case by (metis 9 10) next case (8 i X F F' S D) note prems = "8.prems" note IH = "8.IH" let ?dl = "map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" let ?du = "map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" let ?cl = "map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd D))" let ?cu = "map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd D))" define c where c: "c = ?cl" define d where d: "d = ?dl" obtain C' where C': "C = c@C'" "C' \ set (tr'\<^sub>p\<^sub>c S D)" using prems(1) c by moura have 0: "ik\<^sub>s\<^sub>t (unlabel c) = {}" "ik\<^sub>s\<^sub>t (unlabel d) = {}" "unlabel ?cl = ?cu" "unlabel ?dl = ?du" unfolding c d unlabel_def by force+ have "ik\<^sub>s\<^sub>t (unlabel c) = {}" unfolding c unlabel_def by force hence 1: "\M; unlabel C'\\<^sub>d \ " using prems(2) C'(1) unfolding unlabel_def by auto have "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, NegChecks X F F')#S)" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence 4: "?R3 S D" using prems(3) by blast have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force obtain B where B: "B \ set (tr\<^sub>p\<^sub>c S D)" "\M; unlabel B\\<^sub>d \" using IH[OF C'(2) 1 4 5 6] by moura have 7: "\M; unlabel c\\<^sub>d \" "\M; unlabel C'\\<^sub>d \" using prems(2) C'(1) 0(1) strand_sem_split(3,4)[of M "unlabel c" "unlabel C'"] unfolding c unlabel_def by auto have 8: "d@B \ set (tr\<^sub>p\<^sub>c ((i,NegChecks X F F')#S) D)" using B(1) unfolding d unlabel_def by auto have "\M; unlabel ?cl\\<^sub>d \" using 7(1) unfolding c unlabel_def by auto hence "\M; ?cu\\<^sub>d \" by (metis 0(3)) moreover { fix j p k q assume "(j, p) \ ((\(t,s). (i,t,s)) ` set F') \ set D" "(k, q) \ ((\(t,s). (i,t,s)) ` set F') \ set D" hence "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, NegChecks X F F')#S) \ set D" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, NegChecks X F F')#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence "\(j, p) \ ((\(t,s). (i,t,s)) ` set F') \ set D. \(k, q) \ ((\(t,s). (i,t,s)) ` set F') \ set D. (\\. Unifier \ (pair p) (pair q)) \ j = k" by blast moreover have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (map snd D) \ set X = {}" using prems(4) by fastforce ultimately have "\M; ?du\\<^sub>d \" using labeled_sat_ineq_dbproj_sem_equiv[of i] by simp hence "\M; unlabel ?dl\\<^sub>d \" by (metis 0(4)) hence "\M; unlabel d\\<^sub>d \" using 0(2) unfolding c d unlabel_def by force hence 9: "\M; unlabel (d@B)\\<^sub>d \" using 0(2) B(2) unfolding unlabel_def by auto show ?case by (metis 8 9) qed } thus "?Q \ ?P" by metis qed subsubsection \Part 3\ private lemma tr'_par_sem_equiv: assumes "\(l,t,s) \ set D. (fv t \ fv s) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" and "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" "ground M" and "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. (\\. Unifier \ (pair p) (pair q)) \ i = j" (is "?R A D") and \: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\M; set (unlabel D) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \; unlabel A\\<^sub>s \ \ (\B \ set (tr'\<^sub>p\<^sub>c A D). \M; unlabel B\\<^sub>d \)" (is "?P \ ?Q") proof - have 1: "\(t,s) \ set (unlabel D). (fv t \ fv s) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" using assms(1) unfolding unlabel_def by force have 2: "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. p = q \ i = j" using assms(4) subst_apply_term_empty by blast show ?thesis by (metis tr_sem_equiv'[OF 1 assms(2,3) \] tr'_par_iff_unlabel_tr[OF 2]) qed subsubsection \Part 4\ lemma tr_par_sem_equiv: assumes "\(l,t,s) \ set D. (fv t \ fv s) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" and "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" "ground M" and "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. (\\. Unifier \ (pair p) (pair q)) \ i = j" and \: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\M; set (unlabel D) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \; unlabel A\\<^sub>s \ \ (\B \ set (tr\<^sub>p\<^sub>c A D). \M; unlabel B\\<^sub>d \)" (is "?P \ ?Q") using tr_par_iff_tr'_par[OF assms(4,1,2), of M \] tr'_par_sem_equiv[OF assms] by metis end end subsection \Theorem: The Stateful Compositionality Result, on the Constraint Level\ theorem (in labeled_stateful_typed_model) par_comp_constr_stateful_typed: assumes \: "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ Sec" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" and \: "\ \\<^sub>s unlabel \" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "(\n. \ \\<^sub>s proj_unl n \) \ (\\' l' ts. prefix \' \ \ suffix [(l', receive\ts\)] \' \ (\' leaks Sec under \))" proof - let ?P = "\n A D. \(i, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A) \ set D. \(j, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A) \ set D. (\\. Unifier \ (pair p) (pair q)) \ i = j" have 1: "\(l, t, t')\set []. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel \) = {}" "fv\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel \) = {}" "ground {}" using \(2) by simp_all have 2: "\n. \(l, t, t')\set []. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n \) = {}" "\n. fv\<^sub>s\<^sub>s\<^sub>t (proj_unl n \) \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n \) = {}" using 1 sst_vars_proj_subset[of _ \] by fast+ note 3 = par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj[OF \(1)] have 4: "\{}; set (unlabel []) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \'; unlabel \\\<^sub>s \' \ (\B\set (tr\<^sub>p\<^sub>c \ []). \{}; unlabel B\\<^sub>d \')" when \': "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \'" for \' using tr_par_sem_equiv[OF 1 _ \'] \(1) unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def constr_sem_d_def by auto obtain \' where \': "\' \ set (tr\<^sub>p\<^sub>c \ [])" "\ \ \unlabel \'\" using 4[OF \(2)] \(1) unfolding constr_sem_d_def by moura have \': "(\n. (\ \ \proj_unl n \'\)) \ (\\'' l' ts. prefix \'' \' \ suffix [(l', receive\ts\\<^sub>s\<^sub>t)] \'' \ (strand_leaks\<^sub>l\<^sub>s\<^sub>t \'' Sec \))" using par_comp_constr_typed[OF tr_par_preserves_par_comp[OF \(1) \'(1)] \'(2) \(2-)] by blast show ?thesis proof (cases "\n. (\ \ \proj_unl n \'\)") case True { fix n assume "\ \ \proj_unl n \'\" hence "\{}; {}; unlabel (proj n \)\\<^sub>s \" using tr_par_proj[OF \'(1), of n] tr_par_sem_equiv[OF 2(1,2) 1(3) _ \(2), of n] 3(1) unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def proj_def constr_sem_d_def by force } thus ?thesis using True \(1,2,3) \(1) by metis next case False then obtain \''::"('fun,'var,'lbl) labeled_strand" where \'': "prefix \'' \'" "strand_leaks\<^sub>l\<^sub>s\<^sub>t \'' Sec \" using \' by blast have "\t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t \'' \. \l. (\ \ \unlabel (proj l \'')\) \ ik\<^sub>s\<^sub>t (unlabel (proj l \'')) \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" proof - obtain s m where sm: "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t \'' \" "\{}; proj_unl m \''@[send\[s]\\<^sub>s\<^sub>t]\\<^sub>d \" using \'' unfolding strand_leaks\<^sub>l\<^sub>s\<^sub>t_def constr_sem_d_def by blast show ?thesis using strand_sem_split(3,4)[OF sm(2)] sm(1) unfolding constr_sem_d_def by auto qed then obtain s m where sm: "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t \'' \" "\ \ \unlabel (proj m \'')\" "ik\<^sub>s\<^sub>t (unlabel (proj m \'')) \\<^sub>s\<^sub>e\<^sub>t \ \ s \ \" by moura hence s': "\{} \\<^sub>c s \ \" "s \ \ = s" using \(1) subst_ground_ident[of s \] unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto \ \ We now need to show that there is some prefix \B\ of \\''\ that also leaks and where \B \ set (tr C D)\ for some prefix \C\ of \\\ \ obtain B::"('fun,'var,'lbl) labeled_strand" and C G::"('fun,'var,'lbl) labeled_stateful_strand" where BC: "prefix B \'" "prefix C \" "B \ set (tr\<^sub>p\<^sub>c C [])" "ik\<^sub>s\<^sub>t (unlabel (proj m B)) \\<^sub>s\<^sub>e\<^sub>t \ \ s \ \" "prefix B \''" "\l t. suffix ((l, receive\t\)#G) C" and G: "list_all (Not \ is_Receive \ snd) G" using tr_leaking_prefix_exists[OF \'(1) \''(1) sm(3)] prefix_order.order_trans[OF _ \''(1)] s' by blast obtain C' where C': "C = C'@G" "\l t. suffix [(l, receive\t\)] C'" using BC(6) unfolding suffix_def by (metis append_Cons append_assoc append_self_conv2) have "\{}; unlabel (proj m B)\\<^sub>d \" using sm(2) BC(5) unfolding prefix_def unlabel_def proj_def constr_sem_d_def by auto hence BC': "\ \ \proj_unl m B@[send\[s]\\<^sub>s\<^sub>t]\" using BC(4) unfolding constr_sem_d_def by auto have BC'': "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t B \" using BC(5) sm(1) declassified_prefix_subset by auto have "\n. \ \\<^sub>s (proj_unl n C'@[Send [s]])" proof - have 5: "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n C) Sec" for n using \(1) BC(2) par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_split(1)[THEN par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj] unfolding prefix_def by auto have "fv\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel \) = {}" "fv\<^sub>s\<^sub>s\<^sub>t (unlabel C) \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel \)" "bvars\<^sub>s\<^sub>s\<^sub>t (unlabel C) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel \)" using \(2) BC(2) sst_vars_append_subset(1,2)[of "unlabel C"] unfolding prefix_def unlabel_def by auto hence "fv\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) = {}" for n using sst_vars_proj_subset[of _ C] sst_vars_proj_subset[of _ \] by blast hence 6: "\(l, t, t')\set []. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) = {}" "fv\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) = {}" "ground {}" for n using 2 by auto have 7: "?P n C []" for n using 5 unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by simp obtain n where n: "\ \\<^sub>s proj_unl n C" "ik\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) \\<^sub>s\<^sub>e\<^sub>t \ \ s \ \" using s'(2) tr_par_proj[OF BC(3), of m] BC'(1) tr_par_sem_equiv[OF 6 7 \(2), of m] tr_par_deduct_iff[OF tr_par_proj(1)[OF BC(3)], of \ m s] unfolding proj_def constr_sem_d_def by auto have "ik\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) = ik\<^sub>s\<^sub>s\<^sub>t (proj_unl n C')" using C'(1) G unfolding ik\<^sub>s\<^sub>s\<^sub>t_def unlabel_def proj_def list_all_iff by fastforce hence 8: "ik\<^sub>s\<^sub>s\<^sub>t (proj_unl n C') \\<^sub>s\<^sub>e\<^sub>t \ \ s \ \" using n(2) by simp have 9: "\ \\<^sub>s proj_unl n C'" using n(1) C'(1) strand_sem_append_stateful by simp show ?thesis using 8 9 strand_sem_append_stateful by auto qed moreover have "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t C \" by (metis tr_par_declassified_eq BC(3) BC'') hence "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t C' \" using ideduct_mono[of "\{set ts |ts. \\, receive\ts\\ \ set C'} \\<^sub>s\<^sub>e\<^sub>t \" _ "\{set ts |ts. \\, receive\ts\\ \ set (C'@G)} \\<^sub>s\<^sub>e\<^sub>t \"] unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_alt_def C'(1) by auto moreover have "prefix C' \" using BC(2) C' unfolding prefix_def by auto ultimately show ?thesis using C'(2) unfolding strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by meson qed qed theorem (in labeled_stateful_typing) par_comp_constr_stateful: assumes \: "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ Sec" "typing_cond\<^sub>s\<^sub>s\<^sub>t (unlabel \)" and \: "\ \\<^sub>s unlabel \" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\\\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\) \ (\\<^sub>\ \\<^sub>s unlabel \) \ ((\n. \\<^sub>\ \\<^sub>s proj_unl n \) \ (\\' l' ts. prefix \' \ \ suffix [(l', receive\ts\)] \' \ (\' leaks Sec under \\<^sub>\)))" proof - let ?P = "\n A D. \(i, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A) \ set D. \(j, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A) \ set D. (\\. Unifier \ (pair p) (pair q)) \ i = j" have 1: "\(l, t, t')\set []. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel \) = {}" "fv\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel \) = {}" "ground {}" using \(2) unfolding typing_cond\<^sub>s\<^sub>s\<^sub>t_def by simp_all have 2: "\n. \(l, t, t')\set []. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n \) = {}" "\n. fv\<^sub>s\<^sub>s\<^sub>t (proj_unl n \) \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n \) = {}" using 1(1,2) sst_vars_proj_subset[of _ \] by fast+ have 3: "\n. par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n \) Sec" using par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj[OF \(1)] by metis have 4: "\{}; set (unlabel []) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \'; unlabel \\\<^sub>s \' \ (\B\set (tr\<^sub>p\<^sub>c \ []). \{}; unlabel B\\<^sub>d \')" when \': "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \'" for \' using tr_par_sem_equiv[OF 1 _ \'] \(1) unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def constr_sem_d_def by auto obtain \' where \': "\' \ set (tr\<^sub>p\<^sub>c \ [])" "\ \ \unlabel \'\" using 4[OF \(2)] \(1) unfolding constr_sem_d_def by moura obtain \\<^sub>\ where \\<^sub>\: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" "\\<^sub>\ \ \unlabel \'\" "(\n. (\\<^sub>\ \ \proj_unl n \'\)) \ (\\'' l' ts. prefix \'' \' \ suffix [(l', receive\ts\\<^sub>s\<^sub>t)] \'' \ (strand_leaks\<^sub>l\<^sub>s\<^sub>t \'' Sec \\<^sub>\))" using par_comp_constr[OF tr_par_preserves_par_comp[OF \(1) \'(1)] tr_par_preserves_typing_cond[OF \ \'(1)] \'(2) \(2)] by moura have \\<^sub>\': "\\<^sub>\ \\<^sub>s unlabel \" using 4[OF \\<^sub>\(1)] \'(1) \\<^sub>\(4) unfolding constr_sem_d_def by auto show ?thesis proof (cases "\n. (\\<^sub>\ \ \proj_unl n \'\)") case True { fix n assume "\\<^sub>\ \ \proj_unl n \'\" hence "\{}; {}; unlabel (proj n \)\\<^sub>s \\<^sub>\" using tr_par_proj[OF \'(1), of n] tr_par_sem_equiv[OF 2(1,2) 1(3) _ \\<^sub>\(1), of n] 3(1) unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def proj_def constr_sem_d_def by force } thus ?thesis using True \\<^sub>\(1,2,3) \\<^sub>\' by metis next case False then obtain \''::"('fun,'var,'lbl) labeled_strand" where \'': "prefix \'' \'" "strand_leaks\<^sub>l\<^sub>s\<^sub>t \'' Sec \\<^sub>\" using \\<^sub>\ by blast moreover { fix ts l assume *: "\{}; unlabel (proj l \'')@[send\ts\\<^sub>s\<^sub>t]\\<^sub>d \\<^sub>\" have "\\<^sub>\ \ \unlabel (proj l \'')\" "\t \ set ts. ik\<^sub>s\<^sub>t (unlabel (proj l \'')) \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\ \ t \ \\<^sub>\" using strand_sem_split(3,4)[OF *] unfolding constr_sem_d_def by auto } ultimately have "\t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t \'' \\<^sub>\. \l. (\\<^sub>\ \ \unlabel (proj l \'')\) \ ik\<^sub>s\<^sub>t (unlabel (proj l \'')) \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\ \ t \ \\<^sub>\" unfolding strand_leaks\<^sub>l\<^sub>s\<^sub>t_def constr_sem_d_def by force then obtain s m where sm: "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t \'' \\<^sub>\" "\\<^sub>\ \ \unlabel (proj m \'')\" "ik\<^sub>s\<^sub>t (unlabel (proj m \'')) \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\ \ s \ \\<^sub>\" by moura hence s': "\{} \\<^sub>c s \ \\<^sub>\" "s \ \\<^sub>\ = s" using \(1) subst_ground_ident[of s \\<^sub>\] unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto \ \ We now need to show that there is some prefix \B\ of \\''\ that also leaks and where \B \ set (tr C D)\ for some prefix \C\ of \\\ \ obtain B::"('fun,'var,'lbl) labeled_strand" and C G::"('fun,'var,'lbl) labeled_stateful_strand" where BC: "prefix B \'" "prefix C \" "B \ set (tr\<^sub>p\<^sub>c C [])" "ik\<^sub>s\<^sub>t (unlabel (proj m B)) \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\ \ s \ \\<^sub>\" "prefix B \''" "\l t. suffix ((l, receive\t\)#G) C" and G: "list_all (Not \ is_Receive \ snd) G" using tr_leaking_prefix_exists[OF \'(1) \''(1) sm(3)] prefix_order.order_trans[OF _ \''(1)] s' by blast obtain C' where C': "C = C'@G" "\l t. suffix [(l, receive\t\)] C'" using BC(6) unfolding suffix_def by (metis append_Cons append_assoc append_self_conv2) have "\{}; unlabel (proj m B)\\<^sub>d \\<^sub>\" using sm(2) BC(5) unfolding prefix_def unlabel_def proj_def constr_sem_d_def by auto hence BC': "\\<^sub>\ \ \proj_unl m B@[send\[s]\\<^sub>s\<^sub>t]\" using BC(4) unfolding constr_sem_d_def by auto have BC'': "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t B \\<^sub>\" using BC(5) sm(1) declassified_prefix_subset by auto have "\n. \\<^sub>\ \\<^sub>s (proj_unl n C'@[Send [s]])" proof - have 5: "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n C) Sec" for n using \(1) BC(2) par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_split(1)[THEN par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj] unfolding prefix_def by auto have "fv\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel \) = {}" "fv\<^sub>s\<^sub>s\<^sub>t (unlabel C) \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel \)" "bvars\<^sub>s\<^sub>s\<^sub>t (unlabel C) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel \)" using \(2) BC(2) sst_vars_append_subset(1,2)[of "unlabel C"] unfolding typing_cond\<^sub>s\<^sub>s\<^sub>t_def prefix_def unlabel_def by auto hence "fv\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) = {}" for n using sst_vars_proj_subset[of _ C] sst_vars_proj_subset[of _ \] by blast hence 6: "\(l, t, t')\set []. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) = {}" "fv\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) = {}" "ground {}" for n using 2 by auto have 7: "?P n C []" for n using 5 unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by simp obtain n where n: "\\<^sub>\ \\<^sub>s proj_unl n C" "ik\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\ \ s \ \\<^sub>\" using s'(2) tr_par_proj[OF BC(3), of m] BC'(1) tr_par_sem_equiv[OF 6 7 \\<^sub>\(1), of m] tr_par_deduct_iff[OF tr_par_proj(1)[OF BC(3)], of \\<^sub>\ m s] unfolding proj_def constr_sem_d_def by auto have "ik\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) = ik\<^sub>s\<^sub>s\<^sub>t (proj_unl n C')" using C'(1) G unfolding ik\<^sub>s\<^sub>s\<^sub>t_def unlabel_def proj_def list_all_iff by fastforce hence 8: "ik\<^sub>s\<^sub>s\<^sub>t (proj_unl n C') \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\ \ s \ \\<^sub>\" using n(2) by simp have 9: "\\<^sub>\ \\<^sub>s proj_unl n C'" using n(1) C'(1) strand_sem_append_stateful by simp show ?thesis using 8 9 strand_sem_append_stateful by auto qed moreover have "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t C \\<^sub>\" by (metis tr_par_declassified_eq BC(3) BC'') hence "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t C' \\<^sub>\" using ideduct_mono[of "\{set ts |ts. \\, receive\ts\\ \ set C'} \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\" _ "\{set ts |ts. \\, receive\ts\\ \ set (C'@G)} \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\"] unfolding declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_alt_def C'(1) by auto moreover have "prefix C' \" using BC(2) C' unfolding prefix_def by auto ultimately show ?thesis using \\<^sub>\(1,2,3) \\<^sub>\' C'(2) unfolding strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by meson qed qed subsection \Theorem: The Stateful Compositionality Result, on the Protocol Level\ context labeled_stateful_typing begin context begin subsubsection \Definitions: Labeled Protocols\ text \ We state our result on the level of protocol traces (i.e., the constraints reachable in a symbolic execution of the actual protocol). Hence, we do not need to convert protocol strands to intruder constraints in the following well-formedness definitions. \ private definition wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s::"('fun,'var,'lbl) labeled_strand set \ bool" where "wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s \ \ (\\ \ \. wf\<^sub>l\<^sub>s\<^sub>t {} \) \ (\\ \ \. \\' \ \. fv\<^sub>l\<^sub>s\<^sub>t \ \ bvars\<^sub>l\<^sub>s\<^sub>t \' = {})" private definition wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s':: "('fun,'var,'lbl) labeled_strand set \ ('fun,'var,'lbl) labeled_strand \ bool" where "wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s' \ \ \ (\\' \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>l\<^sub>s\<^sub>t \) (unlabel \')) \ (\\' \ \. \\'' \ \. fv\<^sub>l\<^sub>s\<^sub>t \' \ bvars\<^sub>l\<^sub>s\<^sub>t \'' = {}) \ (\\' \ \. fv\<^sub>l\<^sub>s\<^sub>t \' \ bvars\<^sub>l\<^sub>s\<^sub>t \ = {}) \ (\\' \ \. fv\<^sub>l\<^sub>s\<^sub>t \ \ bvars\<^sub>l\<^sub>s\<^sub>t \' = {})" private definition typing_cond_prot where "typing_cond_prot \

\ wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s \

\ tfr\<^sub>s\<^sub>e\<^sub>t (\(trms\<^sub>l\<^sub>s\<^sub>t ` \

)) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\(trms\<^sub>l\<^sub>s\<^sub>t ` \

)) \ (\\ \ \

. list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel \)) \ Ana_invar_subst (\(ik\<^sub>s\<^sub>t ` unlabel ` \

) \ \(assignment_rhs\<^sub>s\<^sub>t ` unlabel ` \

))" private definition par_comp_prot where "par_comp_prot \

Sec \ (\l1 l2. l1 \ l2 \ GSMP_disjoint (\\ \ \

. trms_proj\<^sub>l\<^sub>s\<^sub>t l1 \) (\\ \ \

. trms_proj\<^sub>l\<^sub>s\<^sub>t l2 \) Sec) \ ground Sec \ (\s \ Sec. \{} \\<^sub>c s) \ typing_cond_prot \

" subsubsection \Lemmata: Labeled Protocols\ private lemma wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s_eqs_wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s': "wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s S = wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s' S []" unfolding wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s_def wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s'_def unlabel_def by auto private lemma par_comp_prot_impl_par_comp: assumes "par_comp_prot \

Sec" "\ \ \

" shows "par_comp \ Sec" proof - have *: "\l1 l2. l1 \ l2 \ GSMP_disjoint (\\ \ \

. trms_proj\<^sub>l\<^sub>s\<^sub>t l1 \) (\\ \ \

. trms_proj\<^sub>l\<^sub>s\<^sub>t l2 \) Sec" using assms(1) unfolding par_comp_prot_def by metis { fix l1 l2::'lbl assume **: "l1 \ l2" hence ***: "GSMP_disjoint (\\ \ \

. trms_proj\<^sub>l\<^sub>s\<^sub>t l1 \) (\\ \ \

. trms_proj\<^sub>l\<^sub>s\<^sub>t l2 \) Sec" using * by auto have "GSMP_disjoint (trms_proj\<^sub>l\<^sub>s\<^sub>t l1 \) (trms_proj\<^sub>l\<^sub>s\<^sub>t l2 \) Sec" using GSMP_disjoint_subset[OF ***] assms(2) by auto } hence "\l1 l2. l1 \ l2 \ GSMP_disjoint (trms_proj\<^sub>l\<^sub>s\<^sub>t l1 \) (trms_proj\<^sub>l\<^sub>s\<^sub>t l2 \) Sec" by metis thus ?thesis using assms unfolding par_comp_prot_def par_comp_def by metis qed private lemma typing_cond_prot_impl_typing_cond: assumes "typing_cond_prot \

" "\ \ \

" shows "typing_cond (unlabel \)" proof - have 1: "wf\<^sub>s\<^sub>t {} (unlabel \)" "fv\<^sub>l\<^sub>s\<^sub>t \ \ bvars\<^sub>l\<^sub>s\<^sub>t \ = {}" using assms unfolding typing_cond_prot_def wf\<^sub>l\<^sub>s\<^sub>t\<^sub>s_def by auto have "tfr\<^sub>s\<^sub>e\<^sub>t (\(trms\<^sub>l\<^sub>s\<^sub>t ` \

))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\(trms\<^sub>l\<^sub>s\<^sub>t ` \

))" "trms\<^sub>l\<^sub>s\<^sub>t \ \ \(trms\<^sub>l\<^sub>s\<^sub>t ` \

)" "SMP (trms\<^sub>l\<^sub>s\<^sub>t \) - Var`\ \ SMP (\(trms\<^sub>l\<^sub>s\<^sub>t ` \

)) - Var`\" using assms SMP_mono[of "trms\<^sub>l\<^sub>s\<^sub>t \" "\(trms\<^sub>l\<^sub>s\<^sub>t ` \

)"] unfolding typing_cond_prot_def by (metis, metis, auto) hence 2: "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>t \)" and 3: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>t \)" unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by (meson subsetD)+ have 4: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel \)" using assms unfolding typing_cond_prot_def by auto have "subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t (unlabel \) \ assignment_rhs\<^sub>s\<^sub>t (unlabel \)) \ subterms\<^sub>s\<^sub>e\<^sub>t (\(ik\<^sub>s\<^sub>t ` unlabel ` \

) \ \(assignment_rhs\<^sub>s\<^sub>t ` unlabel ` \

))" using assms(2) by auto hence 5: "Ana_invar_subst (ik\<^sub>s\<^sub>t (unlabel \) \ assignment_rhs\<^sub>s\<^sub>t (unlabel \))" using assms SMP_mono unfolding typing_cond_prot_def Ana_invar_subst_def by (meson subsetD) show ?thesis using 1 2 3 4 5 unfolding typing_cond_def tfr\<^sub>s\<^sub>t_def by blast qed subsubsection \Theorem: Parallel Compositionality for Labeled Protocols\ private definition component_prot where "component_prot n P \ (\l \ P. \s \ set l. has_LabelN n s \ has_LabelS s)" private definition composed_prot where "composed_prot \

\<^sub>i \ {\. \n. proj n \ \ \

\<^sub>i n}" private definition component_secure_prot where "component_secure_prot n P Sec attack \ (\\ \ P. suffix [(ln n, Send1 (Fun attack []))] \ \ (\\\<^sub>\. (interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)) \ \(\\<^sub>\ \ \proj_unl n \\) \ (\\'. prefix \' \ \ (\t \ Sec-declassified\<^sub>l\<^sub>s\<^sub>t \' \\<^sub>\. \(\\<^sub>\ \ \proj_unl n \'@[Send1 t]\)))))" private definition component_leaks where "component_leaks n \ Sec \ (\\' \\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\) \ prefix \' \ \ (\t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t \' \\<^sub>\. (\\<^sub>\ \ \proj_unl n \'@[Send1 t]\)))" private definition unsat where "unsat \ \ (\\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ \(\ \ \unlabel \\))" private theorem par_comp_constr_prot: assumes P: "P = composed_prot Pi" "par_comp_prot P Sec" "\n. component_prot n (Pi n)" and left_secure: "component_secure_prot n (Pi n) Sec attack" shows "\\ \ P. suffix [(ln n, Send1 (Fun attack []))] \ \ unsat \ \ (\m. n \ m \ component_leaks m \ Sec)" proof - { fix \ \' assume \: "\ = \'@[(ln n, Send1 (Fun attack []))]" "\ \ P" let ?P = "\\' \\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\) \ prefix \' \ \ (\t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t \' \\<^sub>\. \m. n \ m \ (\\<^sub>\ \ \proj_unl m \'@[Send1 t]\))" have tcp: "typing_cond_prot P" using P(2) unfolding par_comp_prot_def by simp have par_comp: "par_comp \ Sec" "typing_cond (unlabel \)" using par_comp_prot_impl_par_comp[OF P(2) \(2)] typing_cond_prot_impl_typing_cond[OF tcp \(2)] by metis+ have "unlabel (proj n \) = proj_unl n \" "proj_unl n \ = proj_unl n (proj n \)" "\A. A \ Pi n \ proj n A = A" "proj n \ = (proj n \')@[(ln n, Send1 (Fun attack []))]" using P(1,3) \ by (auto simp add: proj_def unlabel_def component_prot_def composed_prot_def) moreover have "proj n \ \ Pi n" using P(1) \ unfolding composed_prot_def by blast moreover { fix A assume "prefix A \" hence *: "prefix (proj n A) (proj n \)" unfolding proj_def prefix_def by force hence "proj_unl n A = proj_unl n (proj n A)" "\I. declassified\<^sub>l\<^sub>s\<^sub>t A I = declassified\<^sub>l\<^sub>s\<^sub>t (proj n A) I" unfolding proj_def declassified\<^sub>l\<^sub>s\<^sub>t_alt_def by auto hence "\B. prefix B (proj n \) \ proj_unl n A = proj_unl n B \ (\I. declassified\<^sub>l\<^sub>s\<^sub>t A I = declassified\<^sub>l\<^sub>s\<^sub>t B I)" using * by metis } ultimately have *: "\\\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\) \ \(\\<^sub>\ \ \proj_unl n \\) \ (\\'. prefix \' \ \ (\t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t \' \\<^sub>\. \(\\<^sub>\ \ \proj_unl n \'@[Send1 t]\)))" using left_secure unfolding component_secure_prot_def composed_prot_def suffix_def by metis { fix \ assume \: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "\ \ \unlabel \\" obtain \\<^sub>\ where \\<^sub>\: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" "\\'. prefix \' \ \ (strand_leaks\<^sub>l\<^sub>s\<^sub>t \' Sec \\<^sub>\)" using par_comp_constr[OF par_comp \(2,1)] * by moura hence "\\'. prefix \' \ \ (\t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t \' \\<^sub>\. \m. n \ m \ (\\<^sub>\ \ \proj_unl m \'@[Send1 t]\))" using \\<^sub>\(4) * unfolding strand_leaks\<^sub>l\<^sub>s\<^sub>t_def by metis hence ?P using \\<^sub>\(1,2,3) by auto } hence "unsat \ \ (\m. n \ m \ component_leaks m \ Sec)" by (metis unsat_def component_leaks_def) } thus ?thesis unfolding suffix_def by metis qed subsubsection \Theorem: Parallel Compositionality for Stateful Protocols\ private abbreviation wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t V \ \ wf'\<^sub>s\<^sub>s\<^sub>t V (unlabel \)" text \ We state our result on the level of protocol traces (i.e., the constraints reachable in a symbolic execution of the actual protocol). Hence, we do not need to convert protocol strands to intruder constraints in the following well-formedness definitions. \ private definition wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s::"('fun,'var,'lbl) labeled_stateful_strand set \ bool" where "wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s \ \ (\\ \ \. wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t {} \) \ (\\ \ \. \\' \ \. fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' = {})" private definition wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s':: "('fun,'var,'lbl) labeled_stateful_strand set \ ('fun,'var,'lbl) labeled_stateful_strand \ bool" where "wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s' \ \ \ (\\' \ \. wf'\<^sub>s\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) (unlabel \')) \ (\\' \ \. \\'' \ \. fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \'' = {}) \ (\\' \ \. fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}) \ (\\' \ \. fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' = {})" private definition typing_cond_prot_stateful where "typing_cond_prot_stateful \

\ wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s \

\ tfr\<^sub>s\<^sub>e\<^sub>t (\(trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ` \

) \ pair ` \(setops\<^sub>s\<^sub>s\<^sub>t ` unlabel ` \

)) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\(trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ` \

)) \ (\S \ \

. list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel S))" private definition par_comp_prot_stateful where "par_comp_prot_stateful \

Sec \ (\l1 l2. l1 \ l2 \ GSMP_disjoint (\\ \ \

. trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \)) (\\ \ \

. trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \)) Sec) \ ground Sec \ (\s \ Sec. \{} \\<^sub>c s) \ (\(i,p) \ \\ \ \

. setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \(j,q) \ \\ \ \

. setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. (\\. Unifier \ (pair p) (pair q)) \ i = j) \ typing_cond_prot_stateful \

" private definition component_secure_prot_stateful where "component_secure_prot_stateful n P Sec attack \ (\\ \ P. suffix [(ln n, Send [Fun attack []])] \ \ (\\\<^sub>\. (interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)) \ \(\\<^sub>\ \\<^sub>s (proj_unl n \)) \ (\\'. prefix \' \ \ (\t \ Sec-declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\. \(\\<^sub>\ \\<^sub>s (proj_unl n \'@[Send [t]]))))))" private definition component_leaks_stateful where "component_leaks_stateful n \ Sec \ (\\' \\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\) \ prefix \' \ \ (\t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\. (\\<^sub>\ \\<^sub>s (proj_unl n \'@[Send [t]]))))" private definition unsat_stateful where "unsat_stateful \ \ (\\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ \(\ \\<^sub>s unlabel \))" private lemma wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s_eqs_wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s': "wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s S = wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s' S []" unfolding wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s_def wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s'_def unlabel_def wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by simp private lemma par_comp_prot_impl_par_comp_stateful: assumes "par_comp_prot_stateful \

Sec" "\ \ \

" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ Sec" proof - have *: "\l1 l2. l1 \ l2 \ GSMP_disjoint (\\ \ \

. trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \)) (\\ \ \

. trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \)) Sec" using assms(1) unfolding par_comp_prot_stateful_def by argo { fix l1 l2::'lbl assume **: "l1 \ l2" hence ***: "GSMP_disjoint (\\ \ \

. trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \)) (\\ \ \

. trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \)) Sec" using * by auto have "GSMP_disjoint (trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \)) (trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \)) Sec" using GSMP_disjoint_subset[OF ***] assms(2) by auto } hence "\l1 l2. l1 \ l2 \ GSMP_disjoint (trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \)) (trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \)) Sec" by metis moreover have "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. (\\. Unifier \ (pair p) (pair q)) \ i = j" using assms(1,2) unfolding par_comp_prot_stateful_def by blast ultimately show ?thesis using assms unfolding par_comp_prot_stateful_def par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by fast qed private lemma typing_cond_prot_impl_typing_cond_stateful: assumes "typing_cond_prot_stateful \

" "\ \ \

" shows "typing_cond\<^sub>s\<^sub>s\<^sub>t (unlabel \)" proof - have 1: "wf'\<^sub>s\<^sub>s\<^sub>t {} (unlabel \)" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" using assms unfolding typing_cond_prot_stateful_def wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s_def by auto have "tfr\<^sub>s\<^sub>e\<^sub>t (\(trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ` \

) \ pair ` \(setops\<^sub>s\<^sub>s\<^sub>t ` unlabel ` \

))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\(trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ` \

))" "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ \(trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ` \

)" "SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel \)) - Var`\ \ SMP (\(trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ` \

) \ pair ` \(setops\<^sub>s\<^sub>s\<^sub>t ` unlabel ` \

)) - Var`\" using assms SMP_mono[of "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel \)" "\(trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ` \

) \ pair ` \(setops\<^sub>s\<^sub>s\<^sub>t ` unlabel ` \

)"] unfolding typing_cond_prot_stateful_def by (metis, metis, auto) hence 2: "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel \))" and 3: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by (meson subsetD)+ have 4: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel \)" using assms unfolding typing_cond_prot_stateful_def by auto show ?thesis using 1 2 3 4 unfolding typing_cond\<^sub>s\<^sub>s\<^sub>t_def tfr\<^sub>s\<^sub>s\<^sub>t_def by blast qed private theorem par_comp_constr_prot_stateful: assumes P: "P = composed_prot Pi" "par_comp_prot_stateful P Sec" "\n. component_prot n (Pi n)" and left_secure: "component_secure_prot_stateful n (Pi n) Sec attack" shows "\\ \ P. suffix [(ln n, Send [Fun attack []])] \ \ unsat_stateful \ \ (\m. n \ m \ component_leaks_stateful m \ Sec)" proof - { fix \ \' assume \: "\ = \'@[(ln n, Send [Fun attack []])]" "\ \ P" let ?P = "\\' \\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\) \ prefix \' \ \ (\t \ Sec-declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\. \m. n \ m \ (\\<^sub>\ \\<^sub>s (proj_unl m \'@[Send [t]])))" have tcp: "typing_cond_prot_stateful P" using P(2) unfolding par_comp_prot_stateful_def by simp have par_comp: "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ Sec" "typing_cond\<^sub>s\<^sub>s\<^sub>t (unlabel \)" using par_comp_prot_impl_par_comp_stateful[OF P(2) \(2)] typing_cond_prot_impl_typing_cond_stateful[OF tcp \(2)] by metis+ have "unlabel (proj n \) = proj_unl n \" "proj_unl n \ = proj_unl n (proj n \)" "\A. A \ Pi n \ proj n A = A" "proj n \ = (proj n \')@[(ln n, Send [Fun attack []])]" using P(1,3) \ by (auto simp add: proj_def unlabel_def component_prot_def composed_prot_def) moreover have "proj n \ \ Pi n" using P(1) \ unfolding composed_prot_def by blast moreover { fix A assume "prefix A \" hence *: "prefix (proj n A) (proj n \)" unfolding proj_def prefix_def by force hence "proj_unl n A = proj_unl n (proj n A)" "\I. declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A) I" by (simp, metis declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj_eq) hence "\B. prefix B (proj n \) \ proj_unl n A = proj_unl n B \ (\I. declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t B I)" using * by metis } ultimately have *: "\\\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\) \ \(\\<^sub>\ \\<^sub>s (proj_unl n \)) \ (\\'. prefix \' \ \ (\t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\. \(\\<^sub>\ \\<^sub>s (proj_unl n \'@[Send [t]]))))" using left_secure unfolding component_secure_prot_stateful_def composed_prot_def suffix_def by metis { fix \ assume \: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "\ \\<^sub>s unlabel \" obtain \\<^sub>\ where \\<^sub>\: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" "\\'. prefix \' \ \ (\' leaks Sec under \\<^sub>\)" using par_comp_constr_stateful[OF par_comp \(2,1)] * by moura hence "\\'. prefix \' \ \ (\t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\. \m. n \ m \ (\\<^sub>\ \\<^sub>s (proj_unl m \'@[Send [t]])))" using \\<^sub>\(4) * unfolding strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by metis hence ?P using \\<^sub>\(1,2,3) by auto } hence "unsat_stateful \ \ (\m. n \ m \ component_leaks_stateful m \ Sec)" by (metis unsat_stateful_def component_leaks_stateful_def) } thus ?thesis unfolding suffix_def by metis qed end end subsection \Automated Compositionality Conditions\ definition comp_GSMP_disjoint where "comp_GSMP_disjoint public arity Ana \ A' B' A B C \ let B\ = B \\<^sub>s\<^sub>e\<^sub>t var_rename (max_var_set (fv\<^sub>s\<^sub>e\<^sub>t A)) in has_all_wt_instances_of \ A' A \ has_all_wt_instances_of \ B' B\ \ finite_SMP_representation arity Ana \ A \ finite_SMP_representation arity Ana \ B\ \ (\t \ A. \s \ B\. \ t = \ s \ mgu t s \ None \ (intruder_synth' public arity {} t) \ (\u \ C. is_wt_instance_of_cond \ t u))" definition comp_par_comp'\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "comp_par_comp'\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ pair_fun A M C \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity C \ (\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. if i = j then True else (let s = pair' pair_fun p; t = pair' pair_fun q in mgu s (t \ var_rename (max_var s)) = None))" definition comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ pair_fun A M C \ let L = remdups (map (the_LabelN \ fst) (filter (Not \ has_LabelS) A)); MP0 = \B. trms\<^sub>s\<^sub>s\<^sub>t B \ (pair' pair_fun) ` setops\<^sub>s\<^sub>s\<^sub>t B; pr = \l. MP0 (proj_unl l A) in length L > 1 \ comp_par_comp'\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ pair_fun A M C \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (MP0 (unlabel A)) \ (\i \ set L. \j \ set L. if i = j then True else comp_GSMP_disjoint public arity Ana \ (pr i) (pr j) (M i) (M j) C)" lemma comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>tI: fixes pair_fun A MP0 pr defines "MP0 \ \B. trms\<^sub>s\<^sub>s\<^sub>t B \ (pair' pair_fun) ` setops\<^sub>s\<^sub>s\<^sub>t B" and "pr \ \l. MP0 (proj_unl l A)" assumes L_def: "L = remdups (map (the_LabelN \ fst) (filter (Not \ has_LabelS) A))" and L_gt: "length L > 1" and cpc': "comp_par_comp'\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ pair_fun A M C" and MP0_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (MP0 (unlabel A))" and GSMP_disj: "\i \ set L. \j \ set L. if i = j then True else comp_GSMP_disjoint public arity Ana \ (pr i) (pr j) (M i) (M j) C" shows "comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ pair_fun A M C" using assms unfolding comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by presburger lemma comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>tI': fixes pair_fun A MP0 pr Ms defines "MP0 \ \B. trms\<^sub>s\<^sub>s\<^sub>t B \ (pair' pair_fun) ` setops\<^sub>s\<^sub>s\<^sub>t B" and "pr \ \l. MP0 (proj_unl l A)" and "M \ \l. case find ((=) l \ fst) Ms of Some M \ set (snd M) | None \ {}" assumes L_def: "map fst Ms = remdups (map (the_LabelN \ fst) (filter (Not \ has_LabelS) A))" and L_gt: "length (map fst Ms) > 1" and cpc': "comp_par_comp'\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ pair_fun A M C" and MP0_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (MP0 (unlabel A))" and GSMP_disj: "\i \ set (map fst Ms). \j \ set (map fst Ms). if i = j then True else comp_GSMP_disjoint public arity Ana \ (pr i) (pr j) (M i) (M j) C" shows "comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ pair_fun A M C" by (rule comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>tI[OF L_def L_gt cpc' MP0_wf[unfolded MP0_def] GSMP_disj[unfolded pr_def MP0_def]]) locale labeled_stateful_typed_model' = labeled_typed_model' arity public Ana \ label_witness1 label_witness2 + stateful_typed_model' arity public Ana \ Pair for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,(('fun,'atom::finite) term_type \ nat)) term \ (('fun,(('fun,'atom) term_type \ nat)) term list \ ('fun,(('fun,'atom) term_type \ nat)) term list)" and \::"('fun,(('fun,'atom) term_type \ nat)) term \ ('fun,'atom) term_type" and Pair::"'fun" and label_witness1::"'lbl" and label_witness2::"'lbl" begin sublocale labeled_stateful_typed_model by unfold_locales lemma GSMP_disjoint_if_comp_GSMP_disjoint: defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" assumes AB'_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity A'" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity B'" and C_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity C" and AB'_disj: "comp_GSMP_disjoint public arity Ana \ A' B' A B C" shows "GSMP_disjoint A' B' (f C - {m. {} \\<^sub>c m})" using GSMP_disjointI[of A' B' A B] AB'_wf AB'_disj C_wf unfolding wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s'_def comp_GSMP_disjoint_def f_def wf\<^sub>t\<^sub>r\<^sub>m_code list_all_iff Let_def by blast lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_if_comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t: defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" assumes A: "comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ Pair A M C" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A (f C - {m. {} \\<^sub>c m})" proof (unfold par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def; intro conjI) let ?Sec = "f C - {m. {} \\<^sub>c m}" let ?L = "remdups (map (the_LabelN \ fst) (filter (Not \ has_LabelS) A))" let ?N1 = "\B. trms\<^sub>s\<^sub>s\<^sub>t B \ (pair' Pair) ` setops\<^sub>s\<^sub>s\<^sub>t B" let ?N2 = "\B. trms\<^sub>s\<^sub>s\<^sub>t B \ pair ` setops\<^sub>s\<^sub>s\<^sub>t B" let ?pr = "\l. ?N1 (proj_unl l A)" let ?\ = "\p. var_rename (max_var (pair p))" note defs = pair_code wf\<^sub>t\<^sub>r\<^sub>m_code wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s'_def list_all_iff trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t setops_list\<^sub>s\<^sub>s\<^sub>t_is_setops\<^sub>s\<^sub>s\<^sub>t have 0: "length ?L > 1" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (?N1 (unlabel A))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity C" (* "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t C) C" *) (* "is_TComp_var_instance_closed \ C" *) "\i \ set ?L. \j \ set ?L. i \ j \ comp_GSMP_disjoint public arity Ana \ (?pr i) (?pr j) (M i) (M j) C" "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. i \ j \ mgu (pair p) (pair q \ ?\ p) = None" using A unfolding comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def comp_par_comp'\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def pair_code by meson+ have L_in_iff: "l \ set ?L \ (\a \ set A. has_LabelN l a)" for l by force have A_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A))" using 0(2) unfolding defs by auto hence A_proj_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l A))" for l using trms\<^sub>s\<^sub>s\<^sub>t_proj_subset(1)[of l A] setops\<^sub>s\<^sub>s\<^sub>t_proj_subset(1)[of l A] by blast hence A_proj_wf_trms': "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (?N1 (proj_unl l A))" for l unfolding defs by auto note C_wf_trms = 0(3)[unfolded list_all_iff wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s'_def wf\<^sub>t\<^sub>r\<^sub>m_code[symmetric]] (* note 1 = has_all_wt_instances_ofD'[OF wf_trms_subterms[OF C_wf_trms] C_wf_trms 0(4)] *) have 2: "GSMP (?N2 (proj_unl l A)) \ GSMP (?N2 (proj_unl l' A))" when "l \ set ?L" for l l' using that L_in_iff GSMP_mono[of "?N2 (proj_unl l A)" "?N2 (proj_unl l' A)"] trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subset_if_no_label[of l A] setops\<^sub>s\<^sub>s\<^sub>t_unlabel_subset_if_no_label[of l A] unfolding list_ex_iff by fast have 3: "GSMP_disjoint (?N2 (proj_unl l1 A)) (?N2 (proj_unl l2 A)) ?Sec" when "l1 \ set ?L" "l2 \ set ?L" "l1 \ l2" for l1 l2 proof - have "GSMP_disjoint (?N1 (proj_unl l1 A)) (?N1 (proj_unl l2 A)) ?Sec" using 0(4) that GSMP_disjoint_if_comp_GSMP_disjoint[ OF A_proj_wf_trms'[of l1] A_proj_wf_trms'[of l2] 0(3), of "M l1" "M l2"] unfolding f_def by blast thus ?thesis unfolding pair_code trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t setops_list\<^sub>s\<^sub>s\<^sub>t_is_setops\<^sub>s\<^sub>s\<^sub>t by simp qed obtain a1 a2 where a: "a1 \ set ?L" "a2 \ set ?L" "a1 \ a2" using remdups_ex2[OF 0(1)] by moura show "ground ?Sec" unfolding f_def by fastforce { fix i p j q assume p: "(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" and q: "(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" and pq: "\\. Unifier \ (pair p) (pair q)" have "\\. Unifier \ (pair p) (pair q \ ?\ p)" using pq vars_term_disjoint_imp_unifier[OF var_rename_fv_disjoint[of "pair p"], of _ "pair q"] by (metis (no_types, lifting) subst_subst_compose var_rename_inv_comp) hence "i = j" using 0(5) mgu_None_is_subst_neq[of "pair p" "pair q \ ?\ p"] p q by fast } thus "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. (\\. Unifier \ (pair p) (pair q)) \ i = j" by blast show "\l1 l2. l1 \ l2 \ GSMP_disjoint (?N2 (proj_unl l1 A)) (?N2 (proj_unl l2 A)) ?Sec" using 2 3 3[OF a] unfolding GSMP_disjoint_def by blast (* show "\s \ ?Sec. \s' \ subterms s. {} \\<^sub>c s' \ s' \ ?Sec" proof (intro ballI) fix s s' assume s: "s \ ?Sec" and s': "s' \ s" then obtain t \ where t: "t \ C" "s = t \ \" "fv s = {}" "\{} \\<^sub>c s" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" unfolding f_def by blast obtain m \ where m: "m \ C" "s' = m \ \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using TComp_var_and_subterm_instance_closed_has_subterms_instances[ OF 0(5,4) C_wf_trms in_subterms_Union[OF t(1)] s'[unfolded t(2)] \] by blast thus "{} \\<^sub>c s' \ s' \ ?Sec" using ground_subterm[OF t(3) s'] unfolding f_def by blast qed *) show "\s \ ?Sec. \{} \\<^sub>c s" by simp qed end locale labeled_stateful_typing' = labeled_stateful_typed_model' arity public Ana \ Pair label_witness1 label_witness2 + stateful_typing_result' arity public Ana \ Pair for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,(('fun,'atom::finite) term_type \ nat)) term \ (('fun,(('fun,'atom) term_type \ nat)) term list \ ('fun,(('fun,'atom) term_type \ nat)) term list)" and \::"('fun,(('fun,'atom) term_type \ nat)) term \ ('fun,'atom) term_type" and Pair::"'fun" and label_witness1::"'lbl" and label_witness2::"'lbl" begin sublocale labeled_stateful_typing by unfold_locales lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_if_comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t': defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" assumes a: "comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ Pair A M C" and B: "\b \ set B. \a \ set A. \\. b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" (is "\b \ set B. \a \ set A. \\. b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \ ?D \") shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t B (f C - {m. {} \\<^sub>c m})" proof (unfold par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def; intro conjI) define N1 where "N1 \ \B::('fun, ('fun,'atom) term_type \ nat) stateful_strand. trms\<^sub>s\<^sub>s\<^sub>t B \ (pair' Pair) ` setops\<^sub>s\<^sub>s\<^sub>t B" define N2 where "N2 \ \B::('fun, ('fun,'atom) term_type \ nat) stateful_strand. trms\<^sub>s\<^sub>s\<^sub>t B \ pair ` setops\<^sub>s\<^sub>s\<^sub>t B" define L where "L \ \A::('fun, ('fun,'atom) term_type \ nat, 'lbl) labeled_stateful_strand. remdups (map (the_LabelN \ fst) (filter (Not \ has_LabelS) A))" define \ where "\ \ \p. var_rename (max_var (pair p::('fun, ('fun,'atom) term_type \ nat) term)) ::('fun, ('fun,'atom) term_type \ nat) subst" let ?Sec = "f C - {m. {} \\<^sub>c m}" have 0: "length (L A) > 1" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (N1 (unlabel A))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity C" (* "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t C) C" *) (* "is_TComp_var_instance_closed \ C" *) "\i \ set (L A). \j \ set (L A). i \ j \ comp_GSMP_disjoint public arity Ana \ (N1 (proj_unl i A)) (N1 (proj_unl j A)) (M i) (M j) C" "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. i \ j \ mgu (pair p) (pair q \ \ p) = None" using a unfolding comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def comp_par_comp'\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def pair_code L_def N1_def \_def by meson+ note 1 = trms\<^sub>s\<^sub>s\<^sub>t_proj_subset(1) setops\<^sub>s\<^sub>s\<^sub>t_proj_subset(1) have N1_iff_N2: "N1 A = N2 A" for A unfolding pair_code trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t setops_list\<^sub>s\<^sub>s\<^sub>t_is_setops\<^sub>s\<^sub>s\<^sub>t N1_def N2_def by simp have N2_proj_subset: "N2 (proj_unl l A) \ N2 (unlabel A)" for l::'lbl and A::"('fun, ('fun,'atom) term_type \ nat, 'lbl) labeled_stateful_strand" using 1(1)[of l A] image_mono[OF 1(2)[of l A], of pair] unfolding N2_def by blast have L_in_iff: "l \ set (L A) \ (\a \ set A. has_LabelN l a)" for l A unfolding L_def by force have L_B_subset_A: "l \ set (L A)" when l: "l \ set (L B)" for l using L_in_iff[of l B] L_in_iff[of l A] B l by fastforce note B_setops = setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_wt_instance_ex[OF B] have B_proj: "\b \ set (proj l B). \a \ set (proj l A). \\. b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \ ?D \" for l using proj_instance_ex[OF B] by fast have B': "\t \ N2 (unlabel B). \s \ N2 (unlabel A). \\. t = s \ \ \ ?D \" using trms\<^sub>s\<^sub>s\<^sub>t_setops\<^sub>s\<^sub>s\<^sub>t_wt_instance_ex[OF B] unfolding N2_def by blast have B'_proj: "\t \ N2 (proj_unl l B). \s \ N2 (proj_unl l A). \\. t = s \ \ \ ?D \" for l using trms\<^sub>s\<^sub>s\<^sub>t_setops\<^sub>s\<^sub>s\<^sub>t_wt_instance_ex[OF B_proj] unfolding N2_def by presburger have A_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (N2 (unlabel A))" using N1_iff_N2[of "unlabel A"] 0(2) unfolding wf\<^sub>t\<^sub>r\<^sub>m_code wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s'_def list_all_iff by auto hence A_proj_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (N2 (proj_unl l A))" for l using 1[of l] unfolding N2_def by blast hence A_proj_wf_trms': "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (N1 (proj_unl l A))" for l using N1_iff_N2[of "proj_unl l A"] unfolding wf\<^sub>t\<^sub>r\<^sub>m_code wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s'_def list_all_iff by presburger note C_wf_trms = 0(3)[unfolded list_all_iff wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s'_def wf\<^sub>t\<^sub>r\<^sub>m_code[symmetric]] have 2: "GSMP (N2 (proj_unl l A)) \ GSMP (N2 (proj_unl l' A))" when "l \ set (L A)" for l l' and A::"('fun, ('fun,'atom) term_type \ nat, 'lbl) labeled_stateful_strand" using that L_in_iff[of _ A] GSMP_mono[of "N2 (proj_unl l A)" "N2 (proj_unl l' A)"] trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subset_if_no_label[of l A] setops\<^sub>s\<^sub>s\<^sub>t_unlabel_subset_if_no_label[of l A] unfolding list_ex_iff N2_def by fast have 3: "GSMP (N2 (proj_unl l B)) \ GSMP (N2 (proj_unl l A))" (is "?X \ ?Y") for l proof fix t assume "t \ ?X" hence t: "t \ SMP (N2 (proj_unl l B))" "fv t = {}" unfolding GSMP_def by simp_all have "t \ SMP (N2 (proj_unl l A))" using t(1) B'_proj[of l] SMP_wt_instances_subset[of "N2 (proj_unl l B)" "N2 (proj_unl l A)"] by metis thus "t \ ?Y" using t(2) unfolding GSMP_def by fast qed have "GSMP_disjoint (N2 (proj_unl l1 A)) (N2 (proj_unl l2 A)) ?Sec" when "l1 \ set (L A)" "l2 \ set (L A)" "l1 \ l2" for l1 l2 proof - have "GSMP_disjoint (N1 (proj_unl l1 A)) (N1 (proj_unl l2 A)) ?Sec" using 0(4) that GSMP_disjoint_if_comp_GSMP_disjoint[ OF A_proj_wf_trms'[of l1] A_proj_wf_trms'[of l2] 0(3), of "M l1" "M l2"] unfolding f_def by blast thus ?thesis using N1_iff_N2 by simp qed hence 4: "GSMP_disjoint (N2 (proj_unl l1 B)) (N2 (proj_unl l2 B)) ?Sec" when "l1 \ set (L A)" "l2 \ set (L A)" "l1 \ l2" for l1 l2 using that 3 unfolding GSMP_disjoint_def by blast { fix i p j q assume p: "(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B" and q: "(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B" and pq: "\\. Unifier \ (pair p) (pair q)" obtain p' \p where p': "(i,p') \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "p = p' \\<^sub>p \p" "pair p = pair p' \ \p" using p B_setops unfolding pair_def by auto obtain q' \q where q': "(j,q') \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "q = q' \\<^sub>p \q" "pair q = pair q' \ \q" using q B_setops unfolding pair_def by auto obtain \ where "Unifier \ (pair p) (pair q)" using pq by blast hence "\\. Unifier \ (pair p') (pair q' \ \ p')" using p'(3) q'(3) var_rename_inv_comp[of "pair q'"] subst_subst_compose vars_term_disjoint_imp_unifier[ OF var_rename_fv_disjoint[of "pair p'"], of "\p \\<^sub>s \" "pair q'" "var_rename_inv (max_var_set (fv (pair p'))) \\<^sub>s \q \\<^sub>s \"] unfolding \_def by fastforce hence "i = j" using mgu_None_is_subst_neq[of "pair p'" "pair q' \ \ p'"] p'(1) q'(1) 0(5) unfolding \_def by fast } thus "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B. (\\. Unifier \ (pair p) (pair q)) \ i = j" by blast obtain a1 a2 where a: "a1 \ set (L A)" "a2 \ set (L A)" "a1 \ a2" using remdups_ex2[OF 0(1)[unfolded L_def]] unfolding L_def by moura show "\l1 l2. l1 \ l2 \ GSMP_disjoint (N2 (proj_unl l1 B)) (N2 (proj_unl l2 B)) ?Sec" using 2[of _ B] 4 4[OF a] L_B_subset_A unfolding GSMP_disjoint_def by blast show "ground ?Sec" unfolding f_def by fastforce (* show "\s \ ?Sec. \s' \ subterms s. {} \\<^sub>c s' \ s' \ ?Sec" proof (intro ballI) fix s s' assume s: "s \ ?Sec" and s': "s' \ s" then obtain t \ where t: "t \ C" "s = t \ \" "fv s = {}" "\{} \\<^sub>c s" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" unfolding f_def by blast obtain m \ where m: "m \ C" "s' = m \ \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using TComp_var_and_subterm_instance_closed_has_subterms_instances[ OF 0(5,4) C_wf_trms in_subterms_Union[OF t(1)] s'[unfolded t(2)] \] by blast thus "{} \\<^sub>c s' \ s' \ ?Sec" using ground_subterm[OF t(3) s'] unfolding f_def by blast qed *) show "\s \ ?Sec. \{} \\<^sub>c s" by simp qed end end diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Strands.thy b/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Strands.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Strands.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Strands.thy @@ -1,2814 +1,2814 @@ (* Title: Stateful_Strands.thy Author: Andreas Viktor Hess, DTU SPDX-License-Identifier: BSD-3-Clause *) section \Stateful Strands\ theory Stateful_Strands imports Strands_and_Constraints begin subsection \Stateful Constraints\ datatype (funs\<^sub>s\<^sub>s\<^sub>t\<^sub>p: 'a, vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p: 'b) stateful_strand_step = Send (the_msgs: "('a,'b) term list") ("send\_\" 80) | Receive (the_msgs: "('a,'b) term list") ("receive\_\" 80) | Equality (the_check: poscheckvariant) (the_lhs: "('a,'b) term") (the_rhs: "('a,'b) term") ("\_: _ \ _\" [80,80]) | Insert (the_elem_term: "('a,'b) term") (the_set_term: "('a,'b) term") ("insert\_,_\" 80) | Delete (the_elem_term: "('a,'b) term") (the_set_term: "('a,'b) term") ("delete\_,_\" 80) | InSet (the_check: poscheckvariant) (the_elem_term: "('a,'b) term") (the_set_term: "('a,'b) term") ("\_: _ \ _\" [80,80]) | NegChecks (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p: "'b list") (the_eqs: "(('a,'b) term \ ('a,'b) term) list") (the_ins: "(('a,'b) term \ ('a,'b) term) list") ("\_\\\: _ \\: _\" [80,80]) where "bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Send _) = []" | "bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Receive _) = []" | "bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Equality _ _ _) = []" | "bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Insert _ _) = []" | "bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Delete _ _) = []" | "bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (InSet _ _ _) = []" type_synonym ('a,'b) stateful_strand = "('a,'b) stateful_strand_step list" type_synonym ('a,'b) dbstatelist = "(('a,'b) term \ ('a,'b) term) list" type_synonym ('a,'b) dbstate = "(('a,'b) term \ ('a,'b) term) set" abbreviation "is_Assignment x \ (is_Equality x \ is_InSet x) \ the_check x = Assign" abbreviation "is_Check x \ ((is_Equality x \ is_InSet x) \ the_check x = Check) \ is_NegChecks x" abbreviation "is_Check_or_Assignment x \ is_Equality x \ is_InSet x \ is_NegChecks x" abbreviation "is_Update x \ is_Insert x \ is_Delete x" abbreviation InSet_select ("select\_,_\") where "select\t,s\ \ InSet Assign t s" abbreviation InSet_check ("\_ in _\") where "\t in s\ \ InSet Check t s" abbreviation Equality_assign ("\_ := _\") where "\t := s\ \ Equality Assign t s" abbreviation Equality_check ("\_ == _\") where "\t == s\ \ Equality Check t s" abbreviation NegChecks_Inequality1 ("\_ != _\") where "\t != s\ \ NegChecks [] [(t,s)] []" abbreviation NegChecks_Inequality2 ("\_\_ != _\") where "\x\t != s\ \ NegChecks [x] [(t,s)] []" abbreviation NegChecks_Inequality3 ("\_,_\_ != _\") where "\x,y\t != s\ \ NegChecks [x,y] [(t,s)] []" abbreviation NegChecks_Inequality4 ("\_,_,_\_ != _\") where "\x,y,z\t != s\ \ NegChecks [x,y,z] [(t,s)] []" abbreviation NegChecks_NotInSet1 ("\_ not in _\") where "\t not in s\ \ NegChecks [] [] [(t,s)]" abbreviation NegChecks_NotInSet2 ("\_\_ not in _\") where "\x\t not in s\ \ NegChecks [x] [] [(t,s)]" abbreviation NegChecks_NotInSet3 ("\_,_\_ not in _\") where "\x,y\t not in s\ \ NegChecks [x,y] [] [(t,s)]" abbreviation NegChecks_NotInSet4 ("\_,_,_\_ not in _\") where "\x,y,z\t not in s\ \ NegChecks [x,y,z] [] [(t,s)]" fun trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p where "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Send ts) = set ts" | "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Receive ts) = set ts" | "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Equality _ t t') = {t,t'}" | "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Insert t t') = {t,t'}" | "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Delete t t') = {t,t'}" | "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (InSet _ t t') = {t,t'}" | "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (NegChecks _ F F') = trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F'" definition trms\<^sub>s\<^sub>s\<^sub>t where "trms\<^sub>s\<^sub>s\<^sub>t S \ \(trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p ` set S)" declare trms\<^sub>s\<^sub>s\<^sub>t_def[simp] fun trms_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p where "trms_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Send ts) = ts" | "trms_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Receive ts) = ts" | "trms_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Equality _ t t') = [t,t']" | "trms_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Insert t t') = [t,t']" | "trms_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Delete t t') = [t,t']" | "trms_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (InSet _ t t') = [t,t']" | "trms_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (NegChecks _ F F') = concat (map (\(t,t'). [t,t']) (F@F'))" definition trms_list\<^sub>s\<^sub>s\<^sub>t where "trms_list\<^sub>s\<^sub>s\<^sub>t S \ remdups (concat (map trms_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p S))" definition ik\<^sub>s\<^sub>s\<^sub>t where "ik\<^sub>s\<^sub>s\<^sub>t A \ {t | t ts. Receive ts \ set A \ t \ set ts}" definition bvars\<^sub>s\<^sub>s\<^sub>t::"('a,'b) stateful_strand \ 'b set" where "bvars\<^sub>s\<^sub>s\<^sub>t S \ \(set (map (set \ bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p) S))" fun fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p::"('a,'b) stateful_strand_step \ 'b set" where "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Send ts) = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" | "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Receive ts) = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" | "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Equality _ t t') = fv t \ fv t'" | "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Insert t t') = fv t \ fv t'" | "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Delete t t') = fv t \ fv t'" | "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (InSet _ t t') = fv t \ fv t'" | "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (NegChecks X F F') = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' - set X" definition fv\<^sub>s\<^sub>s\<^sub>t::"('a,'b) stateful_strand \ 'b set" where "fv\<^sub>s\<^sub>s\<^sub>t S \ \(set (map fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p S))" fun fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p where "fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (send\ts\) = concat (map fv_list ts)" | "fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (receive\ts\) = concat (map fv_list ts)" | "fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\_: t \ s\) = fv_list t@fv_list s" | "fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (insert\t,s\) = fv_list t@fv_list s" | "fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (delete\t,s\) = fv_list t@fv_list s" | "fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\_: t \ s\) = fv_list t@fv_list s" | "fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: F \\: F'\) = filter (\x. x \ set X) (fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F@F'))" definition fv_list\<^sub>s\<^sub>s\<^sub>t where "fv_list\<^sub>s\<^sub>s\<^sub>t S \ remdups (concat (map fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p S))" declare bvars\<^sub>s\<^sub>s\<^sub>t_def[simp] declare fv\<^sub>s\<^sub>s\<^sub>t_def[simp] definition vars\<^sub>s\<^sub>s\<^sub>t::"('a,'b) stateful_strand \ 'b set" where "vars\<^sub>s\<^sub>s\<^sub>t S \ \(set (map vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p S))" abbreviation wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p::"('a,'b) stateful_strand_step \ 'b set" where "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p x \ case x of NegChecks _ _ _ \ {} | Equality Check _ _ \ {} | InSet Check _ _ \ {} | Delete _ _ \ {} | _ \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p x" definition wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t::"('a,'b) stateful_strand \ 'b set" where "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S \ \(set (map wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p S))" abbreviation wfvarsoccs\<^sub>s\<^sub>s\<^sub>t\<^sub>p where "wfvarsoccs\<^sub>s\<^sub>s\<^sub>t\<^sub>p x \ case x of Send ts \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) | Equality Assign s t \ fv s | InSet Assign s t \ fv s \ fv t | _ \ {}" definition wfvarsoccs\<^sub>s\<^sub>s\<^sub>t where "wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S \ \(set (map wfvarsoccs\<^sub>s\<^sub>s\<^sub>t\<^sub>p S))" fun wf'\<^sub>s\<^sub>s\<^sub>t::"'b set \ ('a,'b) stateful_strand \ bool" where "wf'\<^sub>s\<^sub>s\<^sub>t V [] = True" | "wf'\<^sub>s\<^sub>s\<^sub>t V (Receive ts#S) = (fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ V \ wf'\<^sub>s\<^sub>s\<^sub>t V S)" | "wf'\<^sub>s\<^sub>s\<^sub>t V (Send ts#S) = wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) S" | "wf'\<^sub>s\<^sub>s\<^sub>t V (Equality Assign t t'#S) = (fv t' \ V \ wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t) S)" | "wf'\<^sub>s\<^sub>s\<^sub>t V (Equality Check _ _#S) = wf'\<^sub>s\<^sub>s\<^sub>t V S" | "wf'\<^sub>s\<^sub>s\<^sub>t V (Insert t s#S) = (fv t \ V \ fv s \ V \ wf'\<^sub>s\<^sub>s\<^sub>t V S)" | "wf'\<^sub>s\<^sub>s\<^sub>t V (Delete _ _#S) = wf'\<^sub>s\<^sub>s\<^sub>t V S" | "wf'\<^sub>s\<^sub>s\<^sub>t V (InSet Assign t s#S) = wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t \ fv s) S" | "wf'\<^sub>s\<^sub>s\<^sub>t V (InSet Check _ _#S) = wf'\<^sub>s\<^sub>s\<^sub>t V S" | "wf'\<^sub>s\<^sub>s\<^sub>t V (NegChecks _ _ _#S) = wf'\<^sub>s\<^sub>s\<^sub>t V S" abbreviation "wf\<^sub>s\<^sub>s\<^sub>t S \ wf'\<^sub>s\<^sub>s\<^sub>t {} S \ fv\<^sub>s\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>s\<^sub>t S = {}" fun subst_apply_stateful_strand_step:: "('a,'b) stateful_strand_step \ ('a,'b) subst \ ('a,'b) stateful_strand_step" (infix "\\<^sub>s\<^sub>s\<^sub>t\<^sub>p" 51) where "send\ts\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = send\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\" | "receive\ts\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\" | "\a: t \ s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \a: (t \ \) \ (s \ \)\" | "\a: t \ s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \a: (t \ \) \ (s \ \)\" | "insert\t,s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = insert\t \ \, s \ \\" | "delete\t,s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = delete\t \ \, s \ \\" | "\X\\\: F \\: G\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \X\\\: (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \\: (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)\" definition subst_apply_stateful_strand:: "('a,'b) stateful_strand \ ('a,'b) subst \ ('a,'b) stateful_strand" (infix "\\<^sub>s\<^sub>s\<^sub>t" 51) where "S \\<^sub>s\<^sub>s\<^sub>t \ \ map (\x. x \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) S" fun dbupd\<^sub>s\<^sub>s\<^sub>t::"('f,'v) stateful_strand \ ('f,'v) subst \ ('f,'v) dbstate \ ('f,'v) dbstate" where "dbupd\<^sub>s\<^sub>s\<^sub>t [] I D = D" | "dbupd\<^sub>s\<^sub>s\<^sub>t (Insert t s#A) I D = dbupd\<^sub>s\<^sub>s\<^sub>t A I (insert ((t,s) \\<^sub>p I) D)" | "dbupd\<^sub>s\<^sub>s\<^sub>t (Delete t s#A) I D = dbupd\<^sub>s\<^sub>s\<^sub>t A I (D - {((t,s) \\<^sub>p I)})" | "dbupd\<^sub>s\<^sub>s\<^sub>t (_#A) I D = dbupd\<^sub>s\<^sub>s\<^sub>t A I D" fun db'\<^sub>s\<^sub>s\<^sub>t::"('f,'v) stateful_strand \ ('f,'v) subst \ ('f,'v) dbstatelist \ ('f,'v) dbstatelist" where "db'\<^sub>s\<^sub>s\<^sub>t [] I D = D" | "db'\<^sub>s\<^sub>s\<^sub>t (Insert t s#A) I D = db'\<^sub>s\<^sub>s\<^sub>t A I (List.insert ((t,s) \\<^sub>p I) D)" | "db'\<^sub>s\<^sub>s\<^sub>t (Delete t s#A) I D = db'\<^sub>s\<^sub>s\<^sub>t A I (List.removeAll ((t,s) \\<^sub>p I) D)" | "db'\<^sub>s\<^sub>s\<^sub>t (_#A) I D = db'\<^sub>s\<^sub>s\<^sub>t A I D" definition db\<^sub>s\<^sub>s\<^sub>t where "db\<^sub>s\<^sub>s\<^sub>t S I \ db'\<^sub>s\<^sub>s\<^sub>t S I []" fun setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p where "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Insert t s) = {(t,s)}" | "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Delete t s) = {(t,s)}" | "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (InSet _ t s) = {(t,s)}" | "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (NegChecks _ _ F') = set F'" | "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p _ = {}" text \The set-operations of a stateful strand\ definition setops\<^sub>s\<^sub>s\<^sub>t where "setops\<^sub>s\<^sub>s\<^sub>t S \ \(setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p ` set S)" fun setops_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p where "setops_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Insert t s) = [(t,s)]" | "setops_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (Delete t s) = [(t,s)]" | "setops_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (InSet _ t s) = [(t,s)]" | "setops_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p (NegChecks _ _ F') = F'" | "setops_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p _ = []" text \The set-operations of a stateful strand (list variant)\ definition setops_list\<^sub>s\<^sub>s\<^sub>t where "setops_list\<^sub>s\<^sub>s\<^sub>t S \ remdups (concat (map setops_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p S))" subsection \Small Lemmata\ lemma is_Check_or_Assignment_iff[simp]: "is_Check x \ is_Assignment x \ is_Check_or_Assignment x" by (cases x) (blast intro: poscheckvariant.exhaust)+ lemma subst_apply_stateful_strand_step_Inequality[simp]: "\t != s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \t \ \ != s \ \\" "\x\t != s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \x\t \ rm_vars {x} \ != s \ rm_vars {x} \\" "\x,y\t != s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \x,y\t \ rm_vars {x,y} \ != s \ rm_vars {x,y} \\" "\x,y,z\t != s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \x,y,z\t \ rm_vars {x,y,z} \ != s \ rm_vars {x,y,z} \\" by simp_all lemma subst_apply_stateful_strand_step_NotInSet[simp]: "\t not in s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \t \ \ not in s \ \\" "\x\t not in s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \x\t \ rm_vars {x} \ not in s \ rm_vars {x} \\" "\x,y\t not in s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \x,y\t \ rm_vars {x,y} \ not in s \ rm_vars {x,y} \\" "\x,y,z\t not in s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \x,y,z\t \ rm_vars {x,y,z} \ not in s \ rm_vars {x,y,z} \\" by simp_all lemma trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t: "trms\<^sub>s\<^sub>s\<^sub>t S = set (trms_list\<^sub>s\<^sub>s\<^sub>t S)" unfolding trms\<^sub>s\<^sub>t_def trms_list\<^sub>s\<^sub>s\<^sub>t_def proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp lemma setops_list\<^sub>s\<^sub>s\<^sub>t_is_setops\<^sub>s\<^sub>s\<^sub>t: "setops\<^sub>s\<^sub>s\<^sub>t S = set (setops_list\<^sub>s\<^sub>s\<^sub>t S)" unfolding setops\<^sub>s\<^sub>s\<^sub>t_def setops_list\<^sub>s\<^sub>s\<^sub>t_def proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp lemma fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p: "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a = set (fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p a)" proof (cases a) case (NegChecks X F G) thus ?thesis using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_append[of F G] fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_append[of F G] fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s[of "F@G"] by auto qed (simp_all add: fv_list\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s fv_list_is_fv) lemma fv_list\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t: "fv\<^sub>s\<^sub>s\<^sub>t S = set (fv_list\<^sub>s\<^sub>s\<^sub>t S)" unfolding fv\<^sub>s\<^sub>s\<^sub>t_def fv_list\<^sub>s\<^sub>s\<^sub>t_def by (induct S) (simp_all add: fv_list\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p) lemma trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_finite[simp]: "finite (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) auto lemma trms\<^sub>s\<^sub>s\<^sub>t_finite[simp]: "finite (trms\<^sub>s\<^sub>s\<^sub>t S)" using trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_finite unfolding trms\<^sub>s\<^sub>s\<^sub>t_def by (induct S) auto lemma vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_finite[simp]: "finite (vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) auto lemma vars\<^sub>s\<^sub>s\<^sub>t_finite[simp]: "finite (vars\<^sub>s\<^sub>s\<^sub>t S)" using vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_finite unfolding vars\<^sub>s\<^sub>s\<^sub>t_def by (induct S) auto lemma fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_finite[simp]: "finite (fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) auto lemma fv\<^sub>s\<^sub>s\<^sub>t_finite[simp]: "finite (fv\<^sub>s\<^sub>s\<^sub>t S)" using fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_finite unfolding fv\<^sub>s\<^sub>s\<^sub>t_def by (induct S) auto lemma bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_finite[simp]: "finite (set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p x))" by (rule finite_set) lemma bvars\<^sub>s\<^sub>s\<^sub>t_finite[simp]: "finite (bvars\<^sub>s\<^sub>s\<^sub>t S)" using bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_finite unfolding bvars\<^sub>s\<^sub>s\<^sub>t_def by (induct S) auto lemma subst_sst_nil[simp]: "[] \\<^sub>s\<^sub>s\<^sub>t \ = []" by (simp add: subst_apply_stateful_strand_def) lemma db\<^sub>s\<^sub>s\<^sub>t_nil[simp]: "db\<^sub>s\<^sub>s\<^sub>t [] \ = []" by (simp add: db\<^sub>s\<^sub>s\<^sub>t_def) lemma ik\<^sub>s\<^sub>s\<^sub>t_nil[simp]: "ik\<^sub>s\<^sub>s\<^sub>t [] = {}" by (simp add: ik\<^sub>s\<^sub>s\<^sub>t_def) lemma in_ik\<^sub>s\<^sub>s\<^sub>t_iff: "t \ ik\<^sub>s\<^sub>s\<^sub>t A \ (\ts. receive\ts\ \ set A \ t \ set ts)" unfolding ik\<^sub>s\<^sub>s\<^sub>t_def by blast lemma ik\<^sub>s\<^sub>s\<^sub>t_append[simp]: "ik\<^sub>s\<^sub>s\<^sub>t (A@B) = ik\<^sub>s\<^sub>s\<^sub>t A \ ik\<^sub>s\<^sub>s\<^sub>t B" by (auto simp add: ik\<^sub>s\<^sub>s\<^sub>t_def) lemma ik\<^sub>s\<^sub>s\<^sub>t_concat: "ik\<^sub>s\<^sub>s\<^sub>t (concat xs) = \(ik\<^sub>s\<^sub>s\<^sub>t ` set xs)" by (induct xs) auto lemma ik\<^sub>s\<^sub>s\<^sub>t_subst: "ik\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>s\<^sub>t \) = ik\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" proof (induction A) case (Cons a A) have "ik\<^sub>s\<^sub>s\<^sub>t ([a] \\<^sub>s\<^sub>s\<^sub>t \) = ik\<^sub>s\<^sub>s\<^sub>t [a] \\<^sub>s\<^sub>e\<^sub>t \" proof (cases a) case (Receive ts) thus ?thesis using in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of _ "[a]"] in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of _ "[a] \\<^sub>s\<^sub>s\<^sub>t \"] unfolding subst_apply_stateful_strand_def by auto qed (simp_all add: ik\<^sub>s\<^sub>s\<^sub>t_def subst_apply_stateful_strand_def) thus ?case using Cons.IH ik\<^sub>s\<^sub>s\<^sub>t_append by (metis append_Cons append_Nil image_Un map_append subst_apply_stateful_strand_def) qed simp lemma ik\<^sub>s\<^sub>s\<^sub>t_set_subset: "set A \ set B \ ik\<^sub>s\<^sub>s\<^sub>t A \ ik\<^sub>s\<^sub>s\<^sub>t B" unfolding ik\<^sub>s\<^sub>s\<^sub>t_def by blast lemma ik\<^sub>s\<^sub>s\<^sub>t_prefix_subset: "prefix A B \ ik\<^sub>s\<^sub>s\<^sub>t A \ ik\<^sub>s\<^sub>s\<^sub>t B" (is "?P A B \ ?P' A B") "prefix A (C@D) \ \prefix A C \ ik\<^sub>s\<^sub>s\<^sub>t C \ ik\<^sub>s\<^sub>s\<^sub>t A" (is "?Q \ ?Q' \ ?Q''") proof - show "?P A B \ ?P' A B" for A B by (metis set_mono_prefix ik\<^sub>s\<^sub>s\<^sub>t_set_subset) thus "?Q \ ?Q' \ ?Q''" by (metis prefixI prefix_same_cases) qed lemma ik\<^sub>s\<^sub>s\<^sub>t_snoc_no_receive_empty: assumes "\a \ set A. \is_Receive a" shows "ik\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I = {}" using assms in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of _ A] by fastforce lemma ik\<^sub>s\<^sub>s\<^sub>t_snoc_no_receive_eq: assumes "\s. a = receive\s\" shows "ik\<^sub>s\<^sub>s\<^sub>t (A@[a]) \\<^sub>s\<^sub>e\<^sub>t \ = ik\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" using assms ik\<^sub>s\<^sub>s\<^sub>t_snoc_no_receive_empty[of "[a]" \] ik\<^sub>s\<^sub>s\<^sub>t_append[of A "[a]"] unfolding is_Receive_def by auto lemma db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t: "set (db'\<^sub>s\<^sub>s\<^sub>t A I D) = dbupd\<^sub>s\<^sub>s\<^sub>t A I (set D)" (is "?A = ?B") proof show "?A \ ?B" proof fix t s show "(t,s) \ ?A \ (t,s) \ ?B" by (induct rule: db'\<^sub>s\<^sub>s\<^sub>t.induct) auto qed show "?B \ ?A" proof fix t s show "(t,s) \ ?B \ (t,s) \ ?A" by (induct arbitrary: D rule: dbupd\<^sub>s\<^sub>s\<^sub>t.induct) auto qed qed lemma db\<^sub>s\<^sub>s\<^sub>t_no_upd: assumes "\a \ set A. \is_Insert a \ \is_Delete a" shows "db'\<^sub>s\<^sub>s\<^sub>t A I D = D" using assms proof (induction A) case (Cons a A) thus ?case by (cases a) auto qed simp lemma db\<^sub>s\<^sub>s\<^sub>t_no_upd_append: assumes "\b \ set B. \is_Insert b \ \is_Delete b" shows "db'\<^sub>s\<^sub>s\<^sub>t A = db'\<^sub>s\<^sub>s\<^sub>t (A@B)" using assms proof (induction A) case Nil thus ?case by (simp add: db\<^sub>s\<^sub>s\<^sub>t_no_upd) next case (Cons a A) thus ?case by (cases a) simp_all qed lemma db\<^sub>s\<^sub>s\<^sub>t_append: "db'\<^sub>s\<^sub>s\<^sub>t (A@B) I D = db'\<^sub>s\<^sub>s\<^sub>t B I (db'\<^sub>s\<^sub>s\<^sub>t A I D)" proof (induction A arbitrary: D) case (Cons a A) thus ?case by (cases a) auto qed simp lemma db\<^sub>s\<^sub>s\<^sub>t_in_cases: assumes "(t,s) \ set (db'\<^sub>s\<^sub>s\<^sub>t A I D)" shows "(t,s) \ set D \ (\t' s'. insert\t',s'\ \ set A \ t = t' \ I \ s = s' \ I)" using assms proof (induction A arbitrary: D) case (Cons a A) thus ?case by (cases a) fastforce+ qed simp lemma db\<^sub>s\<^sub>s\<^sub>t_in_cases': assumes "(t,s) \ set (db'\<^sub>s\<^sub>s\<^sub>t A I D)" and "(t,s) \ set D" shows "\B C t' s'. A = B@insert\t',s'\#C \ t = t' \ I \ s = s' \ I \ (\t'' s''. delete\t'',s''\ \ set C \ t \ t'' \ I \ s \ s'' \ I)" using assms(1) proof (induction A rule: List.rev_induct) case (snoc a A) note * = snoc db\<^sub>s\<^sub>s\<^sub>t_append[of A "[a]" I D] thus ?case proof (cases a) case (Insert t' s') thus ?thesis using * by (cases "(t,s) \ set (db'\<^sub>s\<^sub>s\<^sub>t A I D)") force+ next case (Delete t' s') hence **: "t \ t' \ I \ s \ s' \ I" using * by simp have "(t,s) \ set (db'\<^sub>s\<^sub>s\<^sub>t A I D)" using * Delete by force then obtain B C u v where B: "A = B@insert\u,v\#C" "t = u \ I" "s = v \ I" "\t' s'. delete\t',s'\ \ set C \ t \ t' \ I \ s \ s' \ I" using snoc.IH by moura have "A@[a] = B@insert\u,v\#(C@[a])" "\t' s'. delete\t',s'\ \ set (C@[a]) \ t \ t' \ I \ s \ s' \ I" using B(1,4) Delete ** by auto thus ?thesis using B(2,3) by blast qed force+ qed (simp add: assms(2)) lemma db\<^sub>s\<^sub>s\<^sub>t_filter: "db'\<^sub>s\<^sub>s\<^sub>t A I D = db'\<^sub>s\<^sub>s\<^sub>t (filter is_Update A) I D" by (induct A I D rule: db'\<^sub>s\<^sub>s\<^sub>t.induct) simp_all lemma db\<^sub>s\<^sub>s\<^sub>t_subst_swap: assumes "\x \ fv\<^sub>s\<^sub>s\<^sub>t A. I x = J x" shows "db'\<^sub>s\<^sub>s\<^sub>t A I D = db'\<^sub>s\<^sub>s\<^sub>t A J D" using assms proof (induction A arbitrary: D) case (Cons a A) hence "\x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a. I x = J x" "\D. db'\<^sub>s\<^sub>s\<^sub>t A I D = db'\<^sub>s\<^sub>s\<^sub>t A J D" by auto thus ?case by (cases a) (simp_all add: term_subst_eq[of _ I J]) qed simp lemma dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd: assumes "\a \ set A. \is_Insert a \ \is_Delete a" shows "dbupd\<^sub>s\<^sub>s\<^sub>t A I D = D" using assms proof (induction A) case (Cons a A) thus ?case by (cases a) auto qed simp lemma dbupd\<^sub>s\<^sub>s\<^sub>t_no_deletes: assumes "list_all (\a. \is_Delete a) A" shows "dbupd\<^sub>s\<^sub>s\<^sub>t A I D = D \ {(t \ I, s \ I) | t s. insert\t,s\ \ set A}" (is "?Q A D") using assms proof (induction A arbitrary: D) case (Cons a A) hence IH: "?Q A D" for D by auto have "\is_Delete a" using Cons.prems by simp thus ?case using IH by (cases a) auto qed simp lemma dbupd\<^sub>s\<^sub>s\<^sub>t_append: "dbupd\<^sub>s\<^sub>s\<^sub>t (A@B) I D = dbupd\<^sub>s\<^sub>s\<^sub>t B I (dbupd\<^sub>s\<^sub>s\<^sub>t A I D)" proof (induction A arbitrary: D) case (Cons a A) thus ?case by (cases a) auto qed simp lemma dbupd\<^sub>s\<^sub>s\<^sub>t_filter: "dbupd\<^sub>s\<^sub>s\<^sub>t A I D = dbupd\<^sub>s\<^sub>s\<^sub>t (filter is_Update A) I D" by (induct A I D rule: dbupd\<^sub>s\<^sub>s\<^sub>t.induct) simp_all lemma dbupd\<^sub>s\<^sub>s\<^sub>t_in_cases: assumes "(t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D" shows "(t,s) \ D \ (\t' s'. insert\t',s'\ \ set A \ t = t' \ I \ s = s' \ I)" (is ?P) and "\u v B. suffix (delete\u,v\#B) A \ (t,s) = (u,v) \\<^sub>p I \ (\u' v'. (t,s) = (u',v') \\<^sub>p I \ insert\u',v'\ \ set B)" (is ?Q) proof - show ?P using assms proof (induction A arbitrary: D) case (Cons a A) thus ?case by (cases a) fastforce+ qed simp show ?Q using assms proof (induction A arbitrary: D rule: List.rev_induct) case (snoc a A) note 0 = snoc.IH snoc.prems note 1 = suffix_snoc[of _ A a] have 2: "dbupd\<^sub>s\<^sub>s\<^sub>t (A@[a]) I D = dbupd\<^sub>s\<^sub>s\<^sub>t A I D" when "\is_Update a" using that dbupd\<^sub>s\<^sub>s\<^sub>t_append[of A "[a]" I D] by (cases a) auto have 3: "suffix (delete\u,v\#B) A \ suffix (delete\u,v\#B@[a]) (A@[a])" when "\is_Update a" for u v B using that by simp have 4: "\C. B = C@[a] \ suffix (delete\u,v\#C) A" when a: "\is_Delete a" "suffix (delete\u,v\#B) (A@[a])" for u v B proof - have a': "a \ delete\u,v\" using a(1) by force obtain C where C: "delete\u,v\#B = C@[a]" "suffix C A" using 1 a(2) by blast show ?thesis using a' C by (cases C) auto qed note 5 = dbupd\<^sub>s\<^sub>s\<^sub>t_append[of A "[a]" I] show ?case proof (cases "is_Update a") case True then obtain u v where "a = insert\u,v\ \ a = delete\u,v\" by (cases a) auto thus ?thesis proof assume a: "a = insert\u,v\" hence a': "\is_Delete a" by simp have 6: "insert\u,v\ \ set B" when B: "suffix (delete\u',v'\#B) (A@[a])" for u' v' B using 4[OF a' B] unfolding a by fastforce have 7: "(t,s) = (u,v) \\<^sub>p I \ (t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D" using snoc.prems 5 a by auto show ?thesis proof (cases "(t,s) = (u,v) \\<^sub>p I") case True have "insert\u,v\ \ set B" when B: "suffix (delete\u',v'\#B) (A@[a])" for u' v' B using 4[OF a' B] unfolding a by fastforce thus ?thesis using True by blast next case False hence 8: "(t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D" using 7 by blast have "\u'' v''. (t,s) = (u'',v'') \\<^sub>p I \ insert\u'',v''\ \ set B" when B: "suffix (delete\u',v'\#B) (A @ [a])" "(t,s) = (u',v') \\<^sub>p I" for u' v' B proof - obtain C where C: "B = C@[a]" "suffix (delete\u',v'\#C) A" using 4[OF a' B(1)] by blast thus ?thesis using snoc.IH[OF 8] B(2) unfolding a by fastforce qed thus ?thesis by blast qed next assume a: "a = delete\u,v\" hence "(t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D - {((u,v) \\<^sub>p I)}" using snoc.prems 5 by auto hence 6: "(t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D" "(t,s) \ (u,v) \\<^sub>p I" by (blast,blast) have "(\C. B = C@[a] \ suffix (delete\u',v'\#C) A) \ (B = [] \ u' = u \ v' = v)" when B: "suffix (delete\u',v'\#B) (A@[a])" for B u' v' proof - obtain C where C: "delete\u',v'\#B = C@[a]" "suffix C A" using B 1 by blast show ?thesis proof (cases "B = []") case True thus ?thesis using C unfolding a by simp next case False then obtain b B' where B': "B = B'@[b]" by (meson rev_exhaust) show ?thesis using C unfolding a B' by auto qed qed hence "\C. B = C@[a] \ suffix (delete\u',v'\#C) A" when "suffix (delete\u',v'\#B) (A@[a])" "(t,s) = (u',v') \\<^sub>p I" for B u' v' using that 6 by blast thus ?thesis using snoc.IH[OF 6(1)] unfolding a by fastforce qed next case False have "\u' v'. (t,s) = (u',v') \\<^sub>p I \ insert\u',v'\ \ set B" when B: "suffix (delete\u,v\#B) (A@[a])" "(t,s) = (u,v) \\<^sub>p I" for u v B proof - obtain C where C: "B = C@[a]" "suffix (delete\u,v\#C) A" using 4[OF _ B(1)] False by blast show ?thesis using B(2) snoc.IH[OF snoc.prems[unfolded 2[OF False]]] C by fastforce qed thus ?thesis by blast qed qed simp qed lemma dbupd\<^sub>s\<^sub>s\<^sub>t_in_iff: "(t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D \ ((\u v B. suffix (delete\u,v\#B) A \ (t,s) = (u,v) \\<^sub>p I \ (\u' v'. (t,s) = (u',v') \\<^sub>p I \ insert\u',v'\ \ set B)) \ ((t,s) \ D \ (\u v. (t,s) = (u,v) \\<^sub>p I \ insert\u,v\ \ set A)))" (is "?P A D \ ?Q1 A \ ?Q2 A D") proof show "?P A D \ ?Q1 A \ ?Q2 A D" using dbupd\<^sub>s\<^sub>s\<^sub>t_in_cases by fast show "?Q1 A \ ?Q2 A D \ ?P A D" proof (induction A arbitrary: D) case (Cons a A) have Q1: "?Q1 A" using Cons.prems suffix_Cons[of _ a A] by blast show ?case proof (cases "is_Update a") case False thus ?thesis using Q1 Cons.IH Cons.prems by (cases a) auto next case True then obtain t' s' where "a = insert\t',s'\ \ a = delete\t',s'\" by (cases a) auto thus ?thesis proof assume a: "a = insert\t',s'\" hence "?Q2 A (insert ((t',s') \\<^sub>p I) D)" using Cons.prems by auto thus ?thesis using Q1 Cons.IH unfolding a by auto next assume a: "a = delete\t',s'\" hence "?Q2 A (D - {(t',s') \\<^sub>p I})" using Cons.prems by auto thus ?thesis using Q1 Cons.IH unfolding a by auto qed qed qed simp qed lemma dbupd\<^sub>s\<^sub>s\<^sub>t_in_cases': fixes A::"('a,'b) stateful_strand" assumes "(t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D" and "(t,s) \ D" shows "\B C t' s'. A = B@insert\t',s'\#C \ t = t' \ I \ s = s' \ I \ (\t'' s''. delete\t'',s''\ \ set C \ t \ t'' \ I \ s \ s'' \ I)" using assms(1) proof (induction A rule: List.rev_induct) case (snoc a A) note 0 = dbupd\<^sub>s\<^sub>s\<^sub>t_append[of A "[a]" I D] have 1: "(t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D" when "\is_Update a" using that snoc.prems 0 by (cases a) auto show ?case proof (cases "is_Update a") case False obtain B C t' s' where B: "A = B@insert\t',s'\#C" "t = t' \ I" "s = s' \ I" "\t'' s''. delete\t'',s''\ \ set C \ t \ t'' \ I \ s \ s'' \ I" using snoc.IH[OF 1[OF False]] by blast have "A@[a] = B@insert\t',s'\#(C@[a])" "\t'' s''. delete\t'',s''\ \ set (C@[a]) \ t \ t'' \ I \ s \ s'' \ I" using False B(1,4) by auto thus ?thesis using B(2,3) by blast next case True then obtain t' s' where "a = insert\t',s'\ \ a = delete\t',s'\" by (cases a) auto thus ?thesis proof assume a: "a = insert\t',s'\" hence "dbupd\<^sub>s\<^sub>s\<^sub>t (A@[a]) I D = insert ((t',s') \\<^sub>p I) (dbupd\<^sub>s\<^sub>s\<^sub>t A I D)" using 0 by simp hence "(t,s) = (t',s') \\<^sub>p I \ (t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D" using snoc.prems by blast thus ?thesis proof assume 2: "(t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D" show ?thesis using snoc.IH[OF 2] unfolding a by force qed (force simp add: a) next assume a: "a = delete\t',s'\" hence 2: "t \ t' \ I \ s \ s' \ I" using 0 snoc.prems by simp have "(t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D" using 0 snoc.prems a by force then obtain B C u v where B: "A = B@insert\u,v\#C" "t = u \ I" "s = v \ I" "\t' s'. delete\t',s'\ \ set C \ t \ t' \ I \ s \ s' \ I" using snoc.IH by moura have "A@[a] = B@insert\u,v\#(C@[a])" "\t' s'. delete\t',s'\ \ set (C@[a]) \ t \ t' \ I \ s \ s' \ I" using B(1,4) a 2 by auto thus ?thesis using B(2,3) by blast qed qed qed (simp add: assms(2)) lemma dbupd\<^sub>s\<^sub>s\<^sub>t_mono: assumes "D \ E" shows "dbupd\<^sub>s\<^sub>s\<^sub>t A I D \ dbupd\<^sub>s\<^sub>s\<^sub>t A I E" using assms proof (induction A arbitrary: D E) case (Cons a A) thus ?case proof (cases a) case (Insert t s) have "insert ((t,s) \\<^sub>p I) D \ insert ((t,s) \\<^sub>p I) E" using Cons.prems by fast thus ?thesis using Cons.IH unfolding Insert by simp next case (Delete t s) have "D - {(t,s) \\<^sub>p I} \ E - {(t,s) \\<^sub>p I}" using Cons.prems by fast thus ?thesis using Cons.IH unfolding Delete by simp qed auto qed simp lemma dbupd\<^sub>s\<^sub>s\<^sub>t_db_narrow: assumes "(t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I (D \ E)" and "(t,s) \ D" shows "(t,s) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I E" using assms proof (induction A arbitrary: D E) case (Cons a A) thus ?case proof (cases a) case (Delete t' s') thus ?thesis using Cons.prems Cons.IH[of "D - {(t',s') \\<^sub>p I}" "E - {(t',s') \\<^sub>p I}"] by (simp add: Un_Diff) qed auto qed simp lemma dbupd\<^sub>s\<^sub>s\<^sub>t_set_term_neq_in_iff: assumes f: "f \ k" and A: "\t s. insert\t,s\ \ set A \ (\g ts. s = Fun g ts)" shows "(t,Fun f ts) \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D \ (t,Fun f ts) \ dbupd\<^sub>s\<^sub>s\<^sub>t (filter (\a. \s ss. a = insert\s,Fun k ss\) A) I D" (is "?P A D \ ?P (?f A) D") proof show "?P A D \ ?P (?f A) D" using A proof (induction A arbitrary: D) case (Cons a A) have IH: "?P A D \ ?P (?f A) D" for D using Cons.prems(2) Cons.IH by simp show ?case proof (cases "is_Update a") case True then obtain u s where "a = insert\u,s\ \ a = delete\u,s\" by (cases a) auto thus ?thesis proof assume a: "a = insert\u,s\" obtain g ss where s: "s = Fun g ss" using a Cons.prems(2) by fastforce have 0: "?P A (insert ((u, s) \\<^sub>p I) D)" using a Cons.prems(1) by fastforce show ?thesis proof (cases "g = k") case g: True have "?f (a#A) = ?f A" unfolding a s g by force moreover have "(t,Fun f ts) \ (u, Fun g ss) \\<^sub>p I" using f unfolding g by auto ultimately show ?thesis using IH[OF 0] dbupd\<^sub>s\<^sub>s\<^sub>t_db_narrow[of t "Fun f ts" "?f A" I "{(u, s) \\<^sub>p I}" D] unfolding a s g by force next case g: False have "?f (a#A) = a#?f A" using g unfolding a s by force thus ?thesis using Cons.prems Cons.IH g unfolding a s by force qed next assume a: "a = delete\u,s\" hence "?f (a#A) = a#?f A" by auto thus ?thesis using Cons.prems Cons.IH unfolding a by fastforce qed next case a: False hence "?P A D" using Cons.prems(1) by (cases a) auto hence "?P (?f A) D" using Cons.IH Cons.prems(2) a by fastforce thus ?thesis using a by (cases a) auto qed qed simp have "dbupd\<^sub>s\<^sub>s\<^sub>t (?f A) I D \ dbupd\<^sub>s\<^sub>s\<^sub>t A I D" proof (induction A arbitrary: D) case (Cons a A) show ?case proof (cases a) case (Insert t s) have "?f (a#A) = a#?f A \ ?f (a#A) = ?f A" unfolding Insert by force hence "dbupd\<^sub>s\<^sub>s\<^sub>t (?f (a#A)) I D \ dbupd\<^sub>s\<^sub>s\<^sub>t (?f A) I (insert ((t,s) \\<^sub>p I) D)" using dbupd\<^sub>s\<^sub>s\<^sub>t_mono[of D "insert ((t, s) \\<^sub>p I) D"] unfolding Insert by auto thus ?thesis using Cons.IH unfolding Insert by fastforce qed (use Cons.prems Cons.IH in auto) qed simp thus "?P (?f A) D \ ?P A D" by blast qed lemma dbupd\<^sub>s\<^sub>s\<^sub>t_subst_const_swap: fixes t s defines "fvs \ \A D. fv\<^sub>s\<^sub>s\<^sub>t A \ fv t \ fv s \ \(fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r ` D)" assumes "(t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t A \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" (is "?in \ A D") and "\x \ fvs A D. \ x = \ x \ (\(\ x \ t) \ \(\ x \ s) \ \(\ x \ t) \ \(\ x \ s) \ (\(u,v) \ D. \(\ x \ u) \ \(\ x \ v) \ \(\ x \ u) \ \(\ x \ v)) \ (\u v. insert\u,v\ \ set A \ delete\u,v\ \ set A \ \(\ x \ u) \ \(\ x \ v) \ \(\ x \ u) \ \(\ x \ v)))" (is "?A \ \ D") and "\x \ fvs A D. \c. \ x = Fun c []" (is "?B \") and "\x \ fvs A D. \c. \ x = Fun c []" (is "?B \") and "\x \ fvs A D. \y \ fvs A D. \ x = \ y \ \ x = \ y" (is "?C \ \ A D") shows "(t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t A \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" (is "?in \ A D") using assms(2-) proof (induction A arbitrary: D rule: List.rev_induct) case Nil then obtain u v where u: "(u,v) \ D" "t \ \ = u \ \" "s \ \ = v \ \" by auto let ?X = "fv t \ fv u" let ?Y = "fv s \ fv v" have 0: "fv u \ fvs [] D" "fv v \ fvs [] D" "fv t \ fvs [] D" "fv s \ fvs [] D" using u(1) unfolding fvs_def by (blast, blast, blast, blast) have 1: "\x \ ?X. \ x = \ x \ (\(\ x \ t) \ \(\ x \ u))" "\x \ ?Y. \ x = \ x \ (\(\ x \ s) \ \(\ x \ v))" using Nil.prems(2) u(1) unfolding fvs_def by (blast,blast) have 2: "\x \ ?X. \c. \ x = Fun c []" "\x \ ?X. \c. \ x = Fun c []" "\x \ ?Y. \c. \ x = Fun c []" "\x \ ?Y. \c. \ x = Fun c []" using Nil.prems(3,4) 0 by (blast,blast,blast,blast) have 3: "\x \ ?X. \y \ ?X. \ x = \ y \ \ x = \ y" "\x \ ?Y. \y \ ?Y. \ x = \ y \ \ x = \ y" using Nil.prems(5) 0 by (blast,blast) have "t \ \ = u \ \" "s \ \ = v \ \" using subst_const_swap_eq'[OF u(2) 1(1) 2(1,2) 3(1)] subst_const_swap_eq'[OF u(3) 1(2) 2(3,4) 3(2)] by argo+ thus ?case using u(1) by force next case (snoc a A) have 0: "fvs A D \ fvs (A@[a]) D" "set A \ set (A@[a])" unfolding fvs_def by auto note 1 = dbupd\<^sub>s\<^sub>s\<^sub>t_append[of A "[a]"] have IH: "(t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t A \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \) \ (t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t A \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" using snoc.IH[of D] snoc.prems(2-) 0 by blast let ?q0 = "\t s \ \. \x \ fv t \ fv s. \ x = \ x \ (\(\ x \ t) \ \(\ x \ s))" let ?q1 = "\t s \. \x \ fv t \ fv s. \c. \ x = Fun c []" let ?q2 = "\t s \ \. \x \ fv t \ fv s. \y \ fv t \ fv s. \ x = \ y \ \ x = \ y" show ?case proof (cases "is_Update a") case False hence "dbupd\<^sub>s\<^sub>s\<^sub>t (A@[a]) \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \) = dbupd\<^sub>s\<^sub>s\<^sub>t A \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" "dbupd\<^sub>s\<^sub>s\<^sub>t (A@[a]) \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \) = dbupd\<^sub>s\<^sub>s\<^sub>t A \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" using 1 by (cases a; auto)+ thus ?thesis using IH snoc.prems(1) by blast next case True then obtain u v where u: "a = insert\u,v\ \ a = delete\u,v\" by (cases a) auto have uv_in: "insert\u,v\ \ set (A@[a]) \ delete\u,v\ \ set (A@[a])" using u by force hence fv_uv: "fv u \ fvs (A@[a]) D" "fv v \ fvs (A@[a]) D" unfolding fvs_def by (force,force) have fv_ts: "fv t \ fvs (A@[a]) D" "fv s \ fvs (A@[a]) D" unfolding fvs_def by (blast,blast) have q0: "?q0 t u \ \" "?q0 s v \ \" "?q0 t u \ \" "?q0 s v \ \" proof - show "?q0 t u \ \" "?q0 s v \ \" using snoc.prems(2) 0 fv_ts fv_uv uv_in by (blast,blast) show "?q0 t u \ \" proof fix x assume "x \ fv t \ fv u" hence "x \ fvs (A@[a]) D" using fv_ts(1) fv_uv(1) by blast thus "\ x = \ x \ (\(\ x \ t) \ \(\ x \ u))" using snoc.prems(2) uv_in by auto qed show "?q0 s v \ \" proof fix x assume "x \ fv s \ fv v" hence "x \ fvs (A@[a]) D" using fv_ts(2) fv_uv(2) by blast thus "\ x = \ x \ (\(\ x \ s) \ \(\ x \ v))" using snoc.prems(2) uv_in by auto qed qed have q1: "?q1 t u \" "?q1 t u \" "?q1 s v \" "?q1 s v \" using snoc.prems(3,4) 0 fv_ts fv_uv by (blast,blast,blast,blast) have q2: "?q2 t u \ \" "?q2 s v \ \" "?q2 t u \ \" "?q2 s v \ \" using snoc.prems(5) 0 fv_ts fv_uv by (blast,blast,blast,blast) from u show ?thesis proof assume a: "a = insert\u,v\" show ?thesis proof (cases "(t \ \, s \ \) = (u,v) \\<^sub>p \") case True hence "(t \ \, s \ \) = (u,v) \\<^sub>p \" using subst_const_swap_eq'[OF _ q0(1) q1(1,2) q2(1)] subst_const_swap_eq'[OF _ q0(2) q1(3,4) q2(2)] by fast thus ?thesis using 1 unfolding a by simp next case False hence "(t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t A \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" using snoc.prems(1) 1 unfolding a by force hence "(t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t A \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" using IH by blast thus ?thesis using 1 unfolding a by simp qed next assume a: "a = delete\u,v\" have "(t \ \, s \ \) \ (u,v) \\<^sub>p \" using snoc.prems(1) dbupd\<^sub>s\<^sub>s\<^sub>t_append[of A "[a]"] unfolding a by fastforce hence 2: "(t \ \, s \ \) \ (u,v) \\<^sub>p \" using subst_const_swap_eq'[OF _ q0(3) q1(2,1) q2(3)] subst_const_swap_eq'[OF _ q0(4) q1(4,3) q2(4)] by fast have "(t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t A \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" using snoc.prems(1) 1 unfolding a by fastforce hence 3: "(t \ \, s \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t A \ (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \)" using IH by blast show ?thesis using 2 3 dbupd\<^sub>s\<^sub>s\<^sub>t_append[of A "[a]"] unfolding a by auto qed qed qed lemma subst_sst_cons: "a#A \\<^sub>s\<^sub>s\<^sub>t \ = (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)#(A \\<^sub>s\<^sub>s\<^sub>t \)" by (simp add: subst_apply_stateful_strand_def) lemma subst_sst_snoc: "A@[a] \\<^sub>s\<^sub>s\<^sub>t \ = (A \\<^sub>s\<^sub>s\<^sub>t \)@[a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]" by (simp add: subst_apply_stateful_strand_def) lemma subst_sst_append[simp]: "A@B \\<^sub>s\<^sub>s\<^sub>t \ = (A \\<^sub>s\<^sub>s\<^sub>t \)@(B \\<^sub>s\<^sub>s\<^sub>t \)" by (simp add: subst_apply_stateful_strand_def) lemma subst_sst_list_all: "list_all is_Send S \ list_all is_Send (S \\<^sub>s\<^sub>s\<^sub>t \)" "list_all is_Receive S \ list_all is_Receive (S \\<^sub>s\<^sub>s\<^sub>t \)" "list_all is_Equality S \ list_all is_Equality (S \\<^sub>s\<^sub>s\<^sub>t \)" "list_all is_Insert S \ list_all is_Insert (S \\<^sub>s\<^sub>s\<^sub>t \)" "list_all is_Delete S \ list_all is_Delete (S \\<^sub>s\<^sub>s\<^sub>t \)" "list_all is_InSet S \ list_all is_InSet (S \\<^sub>s\<^sub>s\<^sub>t \)" "list_all is_NegChecks S \ list_all is_NegChecks (S \\<^sub>s\<^sub>s\<^sub>t \)" "list_all is_Assignment S \ list_all is_Assignment (S \\<^sub>s\<^sub>s\<^sub>t \)" "list_all is_Check S \ list_all is_Check (S \\<^sub>s\<^sub>s\<^sub>t \)" "list_all is_Update S \ list_all is_Update (S \\<^sub>s\<^sub>s\<^sub>t \)" "list_all is_Check_or_Assignment S \ list_all is_Check_or_Assignment (S \\<^sub>s\<^sub>s\<^sub>t \)" proof (induction S) case (Cons x S) note * = list_all_def subst_apply_stateful_strand_def { case 1 thus ?case using Cons.IH(1) unfolding * by (cases x) auto } { case 2 thus ?case using Cons.IH(2) unfolding * by (cases x) auto } { case 3 thus ?case using Cons.IH(3) unfolding * by (cases x) auto } { case 4 thus ?case using Cons.IH(4) unfolding * by (cases x) auto } { case 5 thus ?case using Cons.IH(5) unfolding * by (cases x) auto } { case 6 thus ?case using Cons.IH(6) unfolding * by (cases x) auto } { case 7 thus ?case using Cons.IH(7) unfolding * by (cases x) auto } { case 8 thus ?case using Cons.IH(8) unfolding * by (cases x) fastforce+ } { case 9 thus ?case using Cons.IH(9) unfolding * by (cases x) auto } { case 10 thus ?case using Cons.IH(10) unfolding * by (cases x) auto } { case 11 thus ?case using Cons.IH(11) unfolding * by (cases x) auto } qed simp_all lemma subst_sstp_id_subst: "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p Var = a" by (cases a) auto lemma subst_sst_id_subst: "A \\<^sub>s\<^sub>s\<^sub>t Var = A" by (induct A) (simp, metis subst_sstp_id_subst subst_sst_cons) lemma sst_vars_append_subset: "fv\<^sub>s\<^sub>s\<^sub>t A \ fv\<^sub>s\<^sub>s\<^sub>t (A@B)" "bvars\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t (A@B)" "fv\<^sub>s\<^sub>s\<^sub>t B \ fv\<^sub>s\<^sub>s\<^sub>t (A@B)" "bvars\<^sub>s\<^sub>s\<^sub>t B \ bvars\<^sub>s\<^sub>s\<^sub>t (A@B)" by auto lemma sst_vars_disj_cons[simp]: "fv\<^sub>s\<^sub>s\<^sub>t (a#A) \ bvars\<^sub>s\<^sub>s\<^sub>t (a#A) = {} \ fv\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t A = {}" unfolding fv\<^sub>s\<^sub>s\<^sub>t_def bvars\<^sub>s\<^sub>s\<^sub>t_def by auto lemma fv\<^sub>s\<^sub>s\<^sub>t_cons_subset[simp]: "fv\<^sub>s\<^sub>s\<^sub>t A \ fv\<^sub>s\<^sub>s\<^sub>t (a#A)" by auto lemma fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases[simp]: "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (send\ts\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (receive\ts\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\c: t \ s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (s \ \)" "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (insert\t,s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (s \ \)" "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (delete\t,s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (s \ \)" "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\c: t \ s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (s \ \)" "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: F \\: G\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) - set X" by simp_all lemma vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_cases[simp]: "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (send\ts\) = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (receive\ts\) = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\c: t \ s\) = fv t \ fv s" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (insert\t,s\) = fv t \ fv s" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (delete\t,s\) = fv t \ fv s" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\c: t \ s\) = fv t \ fv s" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: F \\: G\) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G \ set X" (is ?A) "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: [(t,s)] \\: []\) = fv t \ fv s \ set X" (is ?B) "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: [] \\: [(t,s)]\) = fv t \ fv s \ set X" (is ?C) proof show ?A ?B ?C by auto qed simp_all lemma vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases[simp]: "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (send\ts\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (receive\ts\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\c: t \ s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (s \ \)" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (insert\t,s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (s \ \)" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (delete\t,s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (s \ \)" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\c: t \ s\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (s \ \)" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: F \\: G\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ set X" (is ?A) "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: [(t,s)] \\: []\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ rm_vars (set X) \) \ fv (s \ rm_vars (set X) \) \ set X" (is ?B) "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: [] \\: [(t,s)]\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ rm_vars (set X) \) \ fv (s \ rm_vars (set X) \) \ set X" (is ?C) proof show ?A ?B ?C by auto qed simp_all lemma bvars\<^sub>s\<^sub>s\<^sub>t_cons_subset: "bvars\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t (a#A)" by auto lemma bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst: "bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" by (cases a) auto lemma bvars\<^sub>s\<^sub>s\<^sub>t_subst: "bvars\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>s\<^sub>t \) = bvars\<^sub>s\<^sub>s\<^sub>t A" using bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst[of _ \] by (induct A) (simp_all add: subst_apply_stateful_strand_def) lemma bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_set_cases[simp]: "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (send\ts\)) = {}" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (receive\ts\)) = {}" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\c: t \ s\)) = {}" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (insert\t,s\)) = {}" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (delete\t,s\)) = {}" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\c: t \ s\)) = {}" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: F \\: G\)) = set X" by simp_all lemma bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegChecks: "\is_NegChecks a \ bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a = []" by (cases a) simp_all lemma bvars\<^sub>s\<^sub>s\<^sub>t_NegChecks: "bvars\<^sub>s\<^sub>s\<^sub>t A = bvars\<^sub>s\<^sub>s\<^sub>t (filter is_NegChecks A)" proof (induction A) case (Cons a A) thus ?case by (cases a) fastforce+ qed simp lemma vars\<^sub>s\<^sub>s\<^sub>t_append[simp]: "vars\<^sub>s\<^sub>s\<^sub>t (A@B) = vars\<^sub>s\<^sub>s\<^sub>t A \ vars\<^sub>s\<^sub>s\<^sub>t B" by (simp add: vars\<^sub>s\<^sub>s\<^sub>t_def) lemma vars\<^sub>s\<^sub>s\<^sub>t_Nil[simp]: "vars\<^sub>s\<^sub>s\<^sub>t [] = {}" by (simp add: vars\<^sub>s\<^sub>s\<^sub>t_def) lemma vars\<^sub>s\<^sub>s\<^sub>t_Cons: "vars\<^sub>s\<^sub>s\<^sub>t (a#A) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \ vars\<^sub>s\<^sub>s\<^sub>t A" by (simp add: vars\<^sub>s\<^sub>s\<^sub>t_def) lemma fv\<^sub>s\<^sub>s\<^sub>t_Cons: "fv\<^sub>s\<^sub>s\<^sub>t (a#A) = fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \ fv\<^sub>s\<^sub>s\<^sub>t A" unfolding fv\<^sub>s\<^sub>s\<^sub>t_def by simp lemma bvars\<^sub>s\<^sub>s\<^sub>t_Cons: "bvars\<^sub>s\<^sub>s\<^sub>t (a#A) = set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) \ bvars\<^sub>s\<^sub>s\<^sub>t A" unfolding bvars\<^sub>s\<^sub>s\<^sub>t_def by auto lemma vars\<^sub>s\<^sub>s\<^sub>t_Cons'[simp]: "vars\<^sub>s\<^sub>s\<^sub>t (send\ts\#A) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (send\ts\) \ vars\<^sub>s\<^sub>s\<^sub>t A" "vars\<^sub>s\<^sub>s\<^sub>t (receive\ts\#A) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (receive\ts\) \ vars\<^sub>s\<^sub>s\<^sub>t A" "vars\<^sub>s\<^sub>s\<^sub>t (\a: t \ s\#A) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\a: t \ s\) \ vars\<^sub>s\<^sub>s\<^sub>t A" "vars\<^sub>s\<^sub>s\<^sub>t (insert\t,s\#A) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (insert\t,s\) \ vars\<^sub>s\<^sub>s\<^sub>t A" "vars\<^sub>s\<^sub>s\<^sub>t (delete\t,s\#A) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (delete\t,s\) \ vars\<^sub>s\<^sub>s\<^sub>t A" "vars\<^sub>s\<^sub>s\<^sub>t (\a: t \ s\#A) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\a: t \ s\) \ vars\<^sub>s\<^sub>s\<^sub>t A" "vars\<^sub>s\<^sub>s\<^sub>t (\X\\\: F \\: G\#A) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: F \\: G\) \ vars\<^sub>s\<^sub>s\<^sub>t A" by (simp_all add: vars\<^sub>s\<^sub>s\<^sub>t_def) lemma fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_if_no_bvars: assumes a: "bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a = []" shows "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a)" proof (cases a) case (NegChecks X F G) hence "set X = {}" using a by fastforce thus ?thesis using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst[of _ \] unfolding NegChecks by simp qed (auto simp add: subst_list_set_fv subst_apply_fv_unfold) lemma fv\<^sub>s\<^sub>s\<^sub>t_subst_if_no_bvars: assumes A: "bvars\<^sub>s\<^sub>s\<^sub>t A = {}" shows "fv\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>s\<^sub>s\<^sub>t A)" using assms proof (induction A) case (Cons a A) thus ?case using fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_if_no_bvars[of a \] fv\<^sub>s\<^sub>s\<^sub>t_Cons[of a A] bvars\<^sub>s\<^sub>s\<^sub>t_Cons[of a A] subst_sst_cons[of a A \] by simp qed simp lemma vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p: fixes x::"('a,'b) stateful_strand_step" shows "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p x \ set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p x)" proof (cases x) case (NegChecks X F G) thus ?thesis by (induct F) force+ qed simp_all lemma vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t: fixes S::"('a,'b) stateful_strand" shows "vars\<^sub>s\<^sub>s\<^sub>t S = fv\<^sub>s\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>s\<^sub>t S" proof (induction S) case (Cons x S) thus ?case using vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p[of x] by (auto simp add: vars\<^sub>s\<^sub>s\<^sub>t_def) qed simp lemma vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck[simp]: "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: F \\: G\) = set X \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" by (simp_all add: sup_commute sup_left_commute vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p) lemma bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck[simp]: "bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: F \\: G\) = X" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\[]\\\: F \\: G\)) = {}" by simp_all lemma fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck[simp]: "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\X\\\: F \\: G\) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G - set X" "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\[]\\\: F \\: G\) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\t != s\) = fv t \ fv s" "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\t not in s\) = fv t \ fv s" by simp_all lemma fv\<^sub>s\<^sub>s\<^sub>t_append[simp]: "fv\<^sub>s\<^sub>s\<^sub>t (A@B) = fv\<^sub>s\<^sub>s\<^sub>t A \ fv\<^sub>s\<^sub>s\<^sub>t B" by simp lemma bvars\<^sub>s\<^sub>s\<^sub>t_append[simp]: "bvars\<^sub>s\<^sub>s\<^sub>t (A@B) = bvars\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>s\<^sub>t B" by auto lemma fv\<^sub>s\<^sub>s\<^sub>t_mono: "set A \ set B \ fv\<^sub>s\<^sub>s\<^sub>t A \ fv\<^sub>s\<^sub>s\<^sub>t B" by auto lemma fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_subterm_trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p: assumes "x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" shows "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a)" using assms var_is_subterm proof (cases a) case (NegChecks X F F') hence "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' - set X" using assms by simp thus ?thesis using NegChecks var_is_subterm by fastforce qed force+ lemma fv\<^sub>s\<^sub>s\<^sub>t_is_subterm_trms\<^sub>s\<^sub>s\<^sub>t: "x \ fv\<^sub>s\<^sub>s\<^sub>t A \ Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t A)" proof (induction A) case (Cons a A) thus ?case using fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_subterm_trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p by (cases "x \ fv\<^sub>s\<^sub>s\<^sub>t A") auto qed simp lemma var_subterm_trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p: assumes "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a)" shows "x \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" using assms vars_iff_subtermeq proof (cases a) case (NegChecks X F F') hence "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F')" using assms by simp thus ?thesis using NegChecks vars_iff_subtermeq by force qed force+ lemma var_subterm_trms\<^sub>s\<^sub>s\<^sub>t_is_vars\<^sub>s\<^sub>s\<^sub>t: "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t A) \ x \ vars\<^sub>s\<^sub>s\<^sub>t A" proof (induction A) case (Cons a A) show ?case proof (cases "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t A)") case True thus ?thesis using Cons.IH by (simp add: vars\<^sub>s\<^sub>s\<^sub>t_def) next case False thus ?thesis using Cons.prems var_subterm_trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p by (fastforce simp add: vars\<^sub>s\<^sub>s\<^sub>t_def) qed qed simp lemma var_trms\<^sub>s\<^sub>s\<^sub>t_is_vars\<^sub>s\<^sub>s\<^sub>t: "Var x \ trms\<^sub>s\<^sub>s\<^sub>t A \ x \ vars\<^sub>s\<^sub>s\<^sub>t A" by (meson var_subterm_trms\<^sub>s\<^sub>s\<^sub>t_is_vars\<^sub>s\<^sub>s\<^sub>t UN_I term.order_refl) lemma ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset: "ik\<^sub>s\<^sub>s\<^sub>t A \ trms\<^sub>s\<^sub>s\<^sub>t A" by (force simp add: ik\<^sub>s\<^sub>s\<^sub>t_def) lemma var_subterm_ik\<^sub>s\<^sub>s\<^sub>t_is_vars\<^sub>s\<^sub>s\<^sub>t: "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>s\<^sub>t A) \ x \ vars\<^sub>s\<^sub>s\<^sub>t A" using var_subterm_trms\<^sub>s\<^sub>s\<^sub>t_is_vars\<^sub>s\<^sub>s\<^sub>t ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset by fast lemma var_subterm_ik\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t: assumes "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>s\<^sub>t A)" shows "x \ fv\<^sub>s\<^sub>s\<^sub>t A" proof - obtain ts where ts: "Receive ts \ set A" "Var x \\<^sub>s\<^sub>e\<^sub>t set ts" using assms unfolding ik\<^sub>s\<^sub>s\<^sub>t_def by moura hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ fv\<^sub>s\<^sub>s\<^sub>t A" unfolding fv\<^sub>s\<^sub>s\<^sub>t_def by force thus ?thesis using ts(2) subterm_is_var by fastforce qed lemma fv_ik\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t: assumes "x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>s\<^sub>t A)" shows "x \ fv\<^sub>s\<^sub>s\<^sub>t A" using var_subterm_ik\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t assms var_is_subterm by fastforce lemma fv_trms\<^sub>s\<^sub>s\<^sub>t_subset: "fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t S) \ vars\<^sub>s\<^sub>s\<^sub>t S" "fv\<^sub>s\<^sub>s\<^sub>t S \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t S)" proof (induction S) case (Cons x S) have *: "fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t (x#S)) = fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p x) \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t S)" "fv\<^sub>s\<^sub>s\<^sub>t (x#S) = fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p x \ fv\<^sub>s\<^sub>s\<^sub>t S" "vars\<^sub>s\<^sub>s\<^sub>t (x#S) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p x \ vars\<^sub>s\<^sub>s\<^sub>t S" unfolding trms\<^sub>s\<^sub>s\<^sub>t_def fv\<^sub>s\<^sub>s\<^sub>t_def vars\<^sub>s\<^sub>s\<^sub>t_def by auto { case 1 show ?case using Cons.IH(1) proof (cases x) case (NegChecks X F G) hence "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p x = trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G \ set X" by (simp, meson vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_cases(7)) hence "fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p x) \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p x" using fv_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s[of F] fv_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s[of G] by auto thus ?thesis using Cons.IH(1) *(1,3) by blast qed auto } { case 2 show ?case using Cons.IH(2) proof (cases x) case (NegChecks X F G) hence "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p x = trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p x = (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G) - set X" by auto hence "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p x)" using fv_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s[of F] fv_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s[of G] by auto thus ?thesis using Cons.IH(2) *(1,2) by blast qed auto } qed simp_all lemma fv_ik_subset_fv_sst'[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>s\<^sub>t S) \ fv\<^sub>s\<^sub>s\<^sub>t S" unfolding ik\<^sub>s\<^sub>s\<^sub>t_def by (induct S) auto lemma fv_ik_subset_vars_sst'[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>s\<^sub>t S) \ vars\<^sub>s\<^sub>s\<^sub>t S" using fv_ik_subset_fv_sst' fv_trms\<^sub>s\<^sub>s\<^sub>t_subset by fast lemma ik\<^sub>s\<^sub>s\<^sub>t_var_is_fv: "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>s\<^sub>t A) \ x \ fv\<^sub>s\<^sub>s\<^sub>t A" by (meson fv_ik_subset_fv_sst'[of A] fv_subset_subterms subsetCE term.set_intros(3)) lemma vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases': assumes x: "x \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" shows "x \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p s \ x \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p s)" using x vars_term_subst[of _ \] vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_cases(1,2,3,4,5,6) vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases(1,2)[of _ \] vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases(3,6)[of _ _ _ \] vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases(4,5)[of _ _ \] proof (cases s) case (NegChecks X F G) let ?\' = "rm_vars (set X) \" have "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\') \ x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\') \ x \ set X" using vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases(7)[of X F G \] x NegChecks by simp hence "x \ fv\<^sub>s\<^sub>e\<^sub>t (?\' ` fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ x \ fv\<^sub>s\<^sub>e\<^sub>t (?\' ` fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G) \ x \ set X" using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst[of _ ?\'] by blast hence "x \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ x \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G) \ x \ set X" using rm_vars_fv\<^sub>s\<^sub>e\<^sub>t_subst by fast thus ?thesis using NegChecks vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_cases(7)[of X F G] by auto qed simp_all lemma vars\<^sub>s\<^sub>s\<^sub>t_subst_cases: assumes "x \ vars\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" shows "x \ vars\<^sub>s\<^sub>s\<^sub>t S \ x \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` vars\<^sub>s\<^sub>s\<^sub>t S)" using assms proof (induction S) case (Cons s S) thus ?case proof (cases "x \ vars\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)") case False note * = subst_sst_cons[of s S \] vars\<^sub>s\<^sub>s\<^sub>t_Cons[of "s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" "S \\<^sub>s\<^sub>s\<^sub>t \"] vars\<^sub>s\<^sub>s\<^sub>t_Cons[of s S] have **: "x \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" using Cons.prems False * by simp show ?thesis using vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases'[OF **] * by auto qed (auto simp add: vars\<^sub>s\<^sub>s\<^sub>t_def) qed simp lemma subset_subst_pairs_diff_exists: fixes \::"('a,'b) subst" and D D'::"('a,'b) dbstate" shows "\Di. Di \ D \ Di \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ = (D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \) - D'" by (metis (no_types, lifting) Diff_subset subset_image_iff) lemma subset_subst_pairs_diff_exists': fixes \::"('a,'b) subst" and D::"('a,'b) dbstate" assumes "finite D" shows "\Di. Di \ D \ Di \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \ {d \\<^sub>p \} \ d \\<^sub>p \ \ (D - Di) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" using assms proof (induction D rule: finite_induct) case (insert d' D) then obtain Di where IH: "Di \ D" "Di \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \ {d \\<^sub>p \}" "d \\<^sub>p \ \ (D - Di) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" by moura show ?case proof (cases "d' \\<^sub>p \ = d \\<^sub>p \") case True hence "insert d' Di \ insert d' D" "insert d' Di \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \ {d \\<^sub>p \}" "d \\<^sub>p \ \ (insert d' D - insert d' Di) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" using IH by auto thus ?thesis by metis next case False hence "Di \ insert d' D" "Di \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \ {d \\<^sub>p \}" "d \\<^sub>p \ \ (insert d' D - Di) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" using IH by auto thus ?thesis by metis qed qed simp lemma stateful_strand_step_subst_inI[intro]: "send\ts\ \ set A \ send\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\ \ set (A \\<^sub>s\<^sub>s\<^sub>t \)" "receive\ts\ \ set A \ receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\ \ set (A \\<^sub>s\<^sub>s\<^sub>t \)" "\c: t \ s\ \ set A \ \c: (t \ \) \ (s \ \)\ \ set (A \\<^sub>s\<^sub>s\<^sub>t \)" "insert\t, s\ \ set A \ insert\t \ \, s \ \\ \ set (A \\<^sub>s\<^sub>s\<^sub>t \)" "delete\t, s\ \ set A \ delete\t \ \, s \ \\ \ set (A \\<^sub>s\<^sub>s\<^sub>t \)" "\c: t \ s\ \ set A \ \c: (t \ \) \ (s \ \)\ \ set (A \\<^sub>s\<^sub>s\<^sub>t \)" "\X\\\: F \\: G\ \ set A \ \X\\\: (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \\: (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)\ \ set (A \\<^sub>s\<^sub>s\<^sub>t \)" "\t != s\ \ set A \ \t \ \ != s \ \\ \ set (A \\<^sub>s\<^sub>s\<^sub>t \)" "\t not in s\ \ set A \ \t \ \ not in s \ \\ \ set (A \\<^sub>s\<^sub>s\<^sub>t \)" proof (induction A) case (Cons a A) note * = subst_sst_cons[of a A \] { case 1 thus ?case using Cons.IH(1) * by (cases a) auto } { case 2 thus ?case using Cons.IH(2) * by (cases a) auto } { case 3 thus ?case using Cons.IH(3) * by (cases a) auto } { case 4 thus ?case using Cons.IH(4) * by (cases a) auto } { case 5 thus ?case using Cons.IH(5) * by (cases a) auto } { case 6 thus ?case using Cons.IH(6) * by (cases a) auto } { case 7 thus ?case using Cons.IH(7) * by (cases a) auto } { case 8 thus ?case using Cons.IH(8) * by (cases a) auto } { case 9 thus ?case using Cons.IH(9) * by (cases a) auto } qed simp_all lemma stateful_strand_step_cases_subst: "is_Send a = is_Send (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "is_Receive a = is_Receive (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "is_Equality a = is_Equality (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "is_Insert a = is_Insert (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "is_Delete a = is_Delete (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "is_InSet a = is_InSet (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "is_NegChecks a = is_NegChecks (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "is_Assignment a = is_Assignment (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "is_Check a = is_Check (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "is_Update a = is_Update (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "is_Check_or_Assignment a = is_Check_or_Assignment (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" by (cases a; simp_all)+ lemma stateful_strand_step_substD: "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = send\ts\ \ \ts'. ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \ a = send\ts'\" "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = receive\ts\ \ \ts'. ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \ a = receive\ts'\" "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \c: t \ s\ \ \t' s'. t = t' \ \ \ s = s' \ \ \ a = \c: t' \ s'\" "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = insert\t,s\ \ \t' s'. t = t' \ \ \ s = s' \ \ \ a = insert\t',s'\" "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = delete\t,s\ \ \t' s'. t = t' \ \ \ s = s' \ \ \ a = delete\t',s'\" "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \c: t \ s\ \ \t' s'. t = t' \ \ \ s = s' \ \ \ a = \c: t' \ s'\" "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \X\\\: F \\: G\ \ \F' G'. F = F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \ \ G = G' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \ \ a = \X\\\: F' \\: G'\" "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \t != s\ \ \t' s'. t = t' \ \ \ s = s' \ \ \ a = \t' != s'\" "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \t not in s\ \ \t' s'. t = t' \ \ \ s = s' \ \ \ a = \t' not in s'\" by (cases a; auto simp add: subst_apply_pairs_def; fail)+ lemma stateful_strand_step_mem_substD: "send\ts\ \ set (S \\<^sub>s\<^sub>s\<^sub>t \) \ \ts'. ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \ send\ts'\ \ set S" "receive\ts\ \ set (S \\<^sub>s\<^sub>s\<^sub>t \) \ \ts'. ts = ts' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \ receive\ts'\ \ set S" "\c: t \ s\ \ set (S \\<^sub>s\<^sub>s\<^sub>t \) \ \t' s'. t = t' \ \ \ s = s' \ \ \ \c: t' \ s'\ \ set S" "insert\t,s\ \ set (S \\<^sub>s\<^sub>s\<^sub>t \) \ \t' s'. t = t' \ \ \ s = s' \ \ \ insert\t',s'\ \ set S" "delete\t,s\ \ set (S \\<^sub>s\<^sub>s\<^sub>t \) \ \t' s'. t = t' \ \ \ s = s' \ \ \ delete\t',s'\ \ set S" "\c: t \ s\ \ set (S \\<^sub>s\<^sub>s\<^sub>t \) \ \t' s'. t = t' \ \ \ s = s' \ \ \ \c: t' \ s'\ \ set S" "\X\\\: F \\: G\ \ set (S \\<^sub>s\<^sub>s\<^sub>t \) \ \F' G'. F = F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \ \ G = G' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \ \ \X\\\: F' \\: G'\ \ set S" "\t != s\ \ set (S \\<^sub>s\<^sub>s\<^sub>t \) \ \t' s'. t = t' \ \ \ s = s' \ \ \ \t' != s'\ \ set S" "\t not in s\ \ set (S \\<^sub>s\<^sub>s\<^sub>t \) \ \t' s'. t = t' \ \ \ s = s' \ \ \ \t' not in s'\ \ set S" proof (induction S) case (Cons a S) have *: "x \ set (S \\<^sub>s\<^sub>s\<^sub>t \)" when "x \ set (a#S \\<^sub>s\<^sub>s\<^sub>t \)" "x \ a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" for x using that by (simp add: subst_apply_stateful_strand_def) { case 1 thus ?case using Cons.IH(1)[OF *] by (cases a) auto } { case 2 thus ?case using Cons.IH(2)[OF *] by (cases a) auto } { case 3 thus ?case using Cons.IH(3)[OF *] by (cases a) auto } { case 4 thus ?case using Cons.IH(4)[OF *] by (cases a) auto } { case 5 thus ?case using Cons.IH(5)[OF *] by (cases a) auto } { case 6 thus ?case using Cons.IH(6)[OF *] by (cases a) auto } { case 7 thus ?case using Cons.IH(7)[OF *] by (cases a) auto } { case 8 show ?case proof (cases a) case (NegChecks Y F' G') thus ?thesis proof (cases "\t != s\ = a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \") case True thus ?thesis using NegChecks stateful_strand_step_substD(8)[of a \ t s] by force qed (use 8 Cons.IH(8)[OF *] in auto) qed (use 8 Cons.IH(8)[OF *] in simp_all) } { case 9 show ?case proof (cases a) case (NegChecks Y F' G') thus ?thesis proof (cases "\t not in s\ = a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \") case True thus ?thesis using NegChecks stateful_strand_step_substD(9)[of a \ t s] by force qed (use 9 Cons.IH(9)[OF *] in auto) qed (use 9 Cons.IH(9)[OF *] in simp_all) } qed simp_all lemma stateful_strand_step_fv_subset_cases: "send\ts\ \ set S \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ fv\<^sub>s\<^sub>s\<^sub>t S" "receive\ts\ \ set S \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ fv\<^sub>s\<^sub>s\<^sub>t S" "\c: t \ s\ \ set S \ fv t \ fv s \ fv\<^sub>s\<^sub>s\<^sub>t S" "insert\t,s\ \ set S \ fv t \ fv s \ fv\<^sub>s\<^sub>s\<^sub>t S" "delete\t,s\ \ set S \ fv t \ fv s \ fv\<^sub>s\<^sub>s\<^sub>t S" "\c: t \ s\ \ set S \ fv t \ fv s \ fv\<^sub>s\<^sub>s\<^sub>t S" "\X\\\: F \\: G\ \ set S \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G - set X \ fv\<^sub>s\<^sub>s\<^sub>t S" "\t != s\ \ set S \ fv t \ fv s \ fv\<^sub>s\<^sub>s\<^sub>t S" "\t not in s\ \ set S \ fv t \ fv s \ fv\<^sub>s\<^sub>s\<^sub>t S" proof (induction S) case (Cons a S) { case 1 thus ?case using Cons.IH(1) by auto } { case 2 thus ?case using Cons.IH(2) by auto } { case 3 thus ?case using Cons.IH(3) by auto } { case 4 thus ?case using Cons.IH(4) by auto } { case 5 thus ?case using Cons.IH(5) by auto } { case 6 thus ?case using Cons.IH(6) by auto } { case 7 thus ?case using Cons.IH(7) by fastforce } { case 8 thus ?case using Cons.IH(8) by fastforce } { case 9 thus ?case using Cons.IH(9) by fastforce } qed simp_all lemma trms\<^sub>s\<^sub>s\<^sub>t_nil[simp]: "trms\<^sub>s\<^sub>s\<^sub>t [] = {}" unfolding trms\<^sub>s\<^sub>s\<^sub>t_def by simp lemma trms\<^sub>s\<^sub>s\<^sub>t_mono: "set M \ set N \ trms\<^sub>s\<^sub>s\<^sub>t M \ trms\<^sub>s\<^sub>s\<^sub>t N" by auto lemma trms\<^sub>s\<^sub>s\<^sub>t_memI[intro?]: "send\ts\ \ set S \ t \ set ts \ t \ trms\<^sub>s\<^sub>s\<^sub>t S" "receive\ts\ \ set S \ t \ set ts \ t \ trms\<^sub>s\<^sub>s\<^sub>t S" "\ac: t \ s\ \ set S \ t \ trms\<^sub>s\<^sub>s\<^sub>t S" "\ac: t \ s\ \ set S \ s \ trms\<^sub>s\<^sub>s\<^sub>t S" "insert\t,s\ \ set S \ t \ trms\<^sub>s\<^sub>s\<^sub>t S" "insert\t,s\ \ set S \ s \ trms\<^sub>s\<^sub>s\<^sub>t S" "delete\t,s\ \ set S \ t \ trms\<^sub>s\<^sub>s\<^sub>t S" "delete\t,s\ \ set S \ s \ trms\<^sub>s\<^sub>s\<^sub>t S" "\X\\\: F \\: G\ \ set S \ t \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ t \ trms\<^sub>s\<^sub>s\<^sub>t S" "\X\\\: F \\: G\ \ set S \ t \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G \ t \ trms\<^sub>s\<^sub>s\<^sub>t S" unfolding trms\<^sub>s\<^sub>s\<^sub>t_def by fastforce+ lemma trms\<^sub>s\<^sub>s\<^sub>t_in: assumes "t \ trms\<^sub>s\<^sub>s\<^sub>t S" shows "\a \ set S. t \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" using assms unfolding trms\<^sub>s\<^sub>s\<^sub>t_def by simp lemma trms\<^sub>s\<^sub>s\<^sub>t_cons: "trms\<^sub>s\<^sub>s\<^sub>t (a#A) = trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \ trms\<^sub>s\<^sub>s\<^sub>t A" unfolding trms\<^sub>s\<^sub>s\<^sub>t_def by force lemma trms\<^sub>s\<^sub>s\<^sub>t_append[simp]: "trms\<^sub>s\<^sub>s\<^sub>t (A@B) = trms\<^sub>s\<^sub>s\<^sub>t A \ trms\<^sub>s\<^sub>s\<^sub>t B" unfolding trms\<^sub>s\<^sub>s\<^sub>t_def by force lemma trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst: assumes "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) \ subst_domain \ = {}" shows "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>s\<^sub>e\<^sub>t \" proof (cases a) case (NegChecks X F G) hence "rm_vars (set X) \ = \" using assms rm_vars_apply'[of \ "set X"] by auto hence "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>s\<^sub>e\<^sub>t \ = (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \\<^sub>s\<^sub>e\<^sub>t \) \ (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G \\<^sub>s\<^sub>e\<^sub>t \)" using NegChecks image_Un by simp_all thus ?thesis by (metis trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst) qed simp_all lemma trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst': assumes "\is_NegChecks a" shows "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>s\<^sub>e\<^sub>t \" using assms by (cases a) simp_all lemma trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst'': fixes t::"('a,'b) term" and \::"('a,'b) subst" assumes "t \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" shows "\s \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p b. t = s \ rm_vars (set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b)) \" proof (cases "is_NegChecks b") case True then obtain X F G where *: "b = NegChecks X F G" by (cases b) moura+ thus ?thesis using assms trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst[of _ "rm_vars (set X) \"] by auto next case False hence "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \\<^sub>s\<^sub>e\<^sub>t rm_vars (set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b)) \" using trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst' bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegChecks by fastforce thus ?thesis using assms by fast qed lemma trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst''': fixes t::"('a,'b) term" and \ \::"('a,'b) subst" assumes "t \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \\<^sub>s\<^sub>e\<^sub>t \" shows "\s \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p b. t = s \ rm_vars (set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b)) \ \\<^sub>s \" proof - obtain s where s: "s \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "t = s \ \" using assms by moura show ?thesis using trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst''[OF s(1)] s(2) by auto qed lemma trms\<^sub>s\<^sub>s\<^sub>t_subst: assumes "bvars\<^sub>s\<^sub>s\<^sub>t S \ subst_domain \ = {}" shows "trms\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \) = trms\<^sub>s\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \" using assms proof (induction S) case (Cons a S) hence IH: "trms\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \) = trms\<^sub>s\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \" and *: "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) \ subst_domain \ = {}" by auto show ?case using trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst[OF *] IH by (auto simp add: subst_apply_stateful_strand_def) qed simp lemma trms\<^sub>s\<^sub>s\<^sub>t_subst_cons: "trms\<^sub>s\<^sub>s\<^sub>t (a#A \\<^sub>s\<^sub>s\<^sub>t \) = trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ trms\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>s\<^sub>t \)" using subst_sst_cons[of a A \] trms\<^sub>s\<^sub>s\<^sub>t_cons[of a A] trms\<^sub>s\<^sub>s\<^sub>t_append by simp lemma (in intruder_model) wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst: assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>s\<^sub>e\<^sub>t \)" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \))" using assms proof (cases a) case (NegChecks X F G) hence *: "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)) \ (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \))" by simp have "trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>s\<^sub>e\<^sub>t \ = (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \\<^sub>s\<^sub>e\<^sub>t \) \ (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G \\<^sub>s\<^sub>e\<^sub>t \)" using NegChecks image_Un by simp hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \\<^sub>s\<^sub>e\<^sub>t \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G \\<^sub>s\<^sub>e\<^sub>t \)" using * assms by auto hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \\<^sub>s\<^sub>e\<^sub>t rm_vars (set X) \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G \\<^sub>s\<^sub>e\<^sub>t rm_vars (set X) \)" using wf_trms_subst_rm_vars[of \ "trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" "set X"] wf_trms_subst_rm_vars[of \ "trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" "set X"] by fast+ thus ?thesis using * trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst[of _ "rm_vars (set X) \"] by auto qed auto lemma trms\<^sub>s\<^sub>s\<^sub>t_fv_vars\<^sub>s\<^sub>s\<^sub>t_subset: "t \ trms\<^sub>s\<^sub>s\<^sub>t A \ fv t \ vars\<^sub>s\<^sub>s\<^sub>t A" proof (induction A) case (Cons a A) thus ?case by (cases a) auto qed simp lemma trms\<^sub>s\<^sub>s\<^sub>t_fv_subst_subset: assumes "t \ trms\<^sub>s\<^sub>s\<^sub>t S" "subst_domain \ \ bvars\<^sub>s\<^sub>s\<^sub>t S = {}" shows "fv (t \ \) \ vars\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" using assms proof (induction S) case (Cons s S) show ?case proof (cases "t \ trms\<^sub>s\<^sub>s\<^sub>t S") case True hence "fv (t \ \) \ vars\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" using Cons.IH Cons.prems by auto thus ?thesis using subst_sst_cons[of s S \] unfolding vars\<^sub>s\<^sub>s\<^sub>t_def by auto next case False hence *: "t \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p s" "subst_domain \ \ set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p s) = {}" using Cons.prems by auto hence "fv (t \ \) \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" proof (cases s) case (NegChecks X F G) hence **: "t \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ t \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" using *(1) by auto have ***: "rm_vars (set X) \ = \" using *(2) NegChecks rm_vars_apply' by auto have "fv (t \ \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" using ** *** trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_fv_subst_subset[of t _ \] by auto thus ?thesis using *(2) using NegChecks vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases(7)[of X F G \] by blast qed auto thus ?thesis using subst_sst_cons[of s S \] unfolding vars\<^sub>s\<^sub>s\<^sub>t_def by auto qed qed simp lemma trms\<^sub>s\<^sub>s\<^sub>t_fv_subst_subset': assumes "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t S)" "fv t \ bvars\<^sub>s\<^sub>s\<^sub>t S = {}" "fv (t \ \) \ bvars\<^sub>s\<^sub>s\<^sub>t S = {}" shows "fv (t \ \) \ fv\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" using assms proof (induction S) case (Cons s S) show ?case proof (cases "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t S)") case True hence "fv (t \ \) \ fv\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" using Cons.IH Cons.prems by auto thus ?thesis using subst_sst_cons[of s S \] unfolding vars\<^sub>s\<^sub>s\<^sub>t_def by auto next case False hence 0: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p s)" "fv t \ set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p s) = {}" "fv (t \ \) \ set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p s) = {}" using Cons.prems by auto note 1 = UN_Un UN_insert fv\<^sub>s\<^sub>e\<^sub>t.simps subst_apply_fv_subset subst_apply_fv_unfold subst_apply_term_empty sup_bot.comm_neutral fv_subterms_set fv_subset[OF 0(1)] note 2 = subst_apply_fv_union have "fv (t \ \) \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" proof (cases s) case (Send ts) have "fv t \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" using fv_subset[OF 0(1)] unfolding Send fv_subterms_set by simp hence "fv (t \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" by (metis subst_apply_fv_subset subst_apply_fv_unfold_set) thus ?thesis using Send by simp next case (Receive ts) have "fv t \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" using fv_subset[OF 0(1)] unfolding Receive fv_subterms_set by simp hence "fv (t \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" by (metis subst_apply_fv_subset subst_apply_fv_unfold_set) thus ?thesis using Receive by simp next case (NegChecks X F G) hence 3: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G)" using 0(1) by auto have "t \ rm_vars (set X) \ = t \ \" using 0(2) NegChecks rm_vars_ident[of t] by auto hence "fv (t \ \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" using 3 trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_fv_subst_subset'[of t _ "rm_vars (set X) \"] by fastforce thus ?thesis using 0(2,3) NegChecks fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases(7)[of X F G \] by auto qed (metis (no_types, lifting) 1 2 trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p.simps(3) fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases(3), metis (no_types, lifting) 1 2 trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p.simps(4) fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases(4), metis (no_types, lifting) 1 2 trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p.simps(5) fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases(5), metis (no_types, lifting) 1 2 trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p.simps(6) fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases(6)) thus ?thesis using subst_sst_cons[of s S \] unfolding fv\<^sub>s\<^sub>s\<^sub>t_def by auto qed qed simp lemma trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_funs_term_cases: assumes "t \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "f \ funs_term t" shows "(\u \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p s. f \ funs_term u) \ (\x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p s. f \ funs_term (\ x))" using assms proof (cases s) case (NegChecks X F G) hence "t \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ t \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" using assms(1) by auto hence "(\u\trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F. f \ funs_term u) \ (\x\fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F. f \ funs_term (rm_vars (set X) \ x)) \ (\u\trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G. f \ funs_term u) \ (\x\fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G. f \ funs_term (rm_vars (set X) \ x))" using trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_funs_term_cases[OF _ assms(2), of _ "rm_vars (set X) \"] by meson hence "(\u \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G. f \ funs_term u) \ (\x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G. f \ funs_term (rm_vars (set X) \ x))" by blast thus ?thesis proof assume "\x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G. f \ funs_term (rm_vars (set X) \ x)" then obtain x where x: "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" "f \ funs_term (rm_vars (set X) \ x)" by auto hence "x \ set X" "rm_vars (set X) \ x = \ x" by auto thus ?thesis using x by (auto simp add: assms NegChecks) qed (auto simp add: assms NegChecks) qed (use assms funs_term_subst[of _ \] in auto) lemma trms\<^sub>s\<^sub>s\<^sub>t_funs_term_cases: assumes "t \ trms\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" "f \ funs_term t" shows "(\u \ trms\<^sub>s\<^sub>s\<^sub>t S. f \ funs_term u) \ (\x \ fv\<^sub>s\<^sub>s\<^sub>t S. f \ funs_term (\ x))" using assms(1) proof (induction S) case (Cons s S) thus ?case proof (cases "t \ trms\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)") case False hence "t \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" using Cons.prems(1) subst_sst_cons[of s S \] trms\<^sub>s\<^sub>s\<^sub>t_cons by auto thus ?thesis using trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_funs_term_cases[OF _ assms(2)] by fastforce qed auto qed simp lemma fv\<^sub>s\<^sub>s\<^sub>t_is_subterm_trms\<^sub>s\<^sub>s\<^sub>t_subst: assumes "x \ fv\<^sub>s\<^sub>s\<^sub>t T" and "bvars\<^sub>s\<^sub>s\<^sub>t T \ subst_domain \ = {}" shows "\ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t (T \\<^sub>s\<^sub>s\<^sub>t \))" using trms\<^sub>s\<^sub>s\<^sub>t_subst[OF assms(2)] subterms_subst_subset'[of \ "trms\<^sub>s\<^sub>s\<^sub>t T"] fv\<^sub>s\<^sub>s\<^sub>t_is_subterm_trms\<^sub>s\<^sub>s\<^sub>t[OF assms(1)] -by (metis (no_types, lifting) image_iff subset_iff subst_apply_term.simps(1)) +by (metis (no_types, lifting) image_iff subset_iff eval_term.simps(1)) lemma fv\<^sub>s\<^sub>s\<^sub>t_subst_fv_subset: assumes "x \ fv\<^sub>s\<^sub>s\<^sub>t S" "x \ bvars\<^sub>s\<^sub>s\<^sub>t S" "fv (\ x) \ bvars\<^sub>s\<^sub>s\<^sub>t S = {}" shows "fv (\ x) \ fv\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" using assms proof (induction S) case (Cons a S) note 1 = fv_subst_subset[of _ _ \] note 2 = subst_apply_fv_union subst_apply_fv_unfold[of _ \] fv_subset image_eqI note 3 = fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst_cases note 4 = fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p.simps from Cons show ?case proof (cases "x \ fv\<^sub>s\<^sub>s\<^sub>t S") case False hence 5: "x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" " fv (\ x) \ set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) = {}" "x \ set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a)" using Cons.prems by auto hence "fv (\ x) \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" proof (cases a) case (Send ts) thus ?thesis using 1 5(1) 3(1) 4(1) by auto next case (Receive ts) thus ?thesis using 1 5(1) 3(2) 4(2) by auto next case (NegChecks X F G) let ?\ = "rm_vars (set X) \" have *: "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" using NegChecks 5(1) by auto have **: "fv (\ x) \ set X = {}" using NegChecks 5(2) by simp have ***: "\ x = ?\ x" using NegChecks 5(3) by auto have "fv (\ x) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\)" using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_fv_subset[of x _ ?\] * *** by auto thus ?thesis using NegChecks ** by auto qed (metis (full_types) 2 5(1) 3(3) 4(3), metis (full_types) 2 5(1) 3(4) 4(4), metis (full_types) 2 5(1) 3(5) 4(5), metis (full_types) 2 5(1) 3(6) 4(6)) thus ?thesis by (auto simp add: subst_sst_cons[of a S \]) qed (auto simp add: subst_sst_cons[of a S \]) qed simp lemma (in intruder_model) wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_trms\<^sub>s\<^sub>s\<^sub>t_subst: assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \)" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>s\<^sub>t \))" using assms proof (induction A) case (Cons a A) hence IH: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>s\<^sub>t \))" and *: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>s\<^sub>e\<^sub>t \)" by auto have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \))" by (rule wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst[OF *]) thus ?case using IH trms\<^sub>s\<^sub>s\<^sub>t_subst_cons[of a A \] by blast qed simp lemma fv\<^sub>s\<^sub>s\<^sub>t_subst_obtain_var: assumes "x \ fv\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" shows "\y \ fv\<^sub>s\<^sub>s\<^sub>t S. x \ fv (\ y)" using assms proof (induction S) case (Cons s S) hence "x \ fv\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \) \ \y \ fv\<^sub>s\<^sub>s\<^sub>t S. x \ fv (\ y)" using bvars\<^sub>s\<^sub>s\<^sub>t_cons_subset[of S s] by blast thus ?case proof (cases "x \ fv\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)") case False hence *: "x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" using Cons.prems(1) subst_sst_cons[of s S \] by fastforce have "\y \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p s. x \ fv (\ y)" proof (cases s) case (NegChecks X F G) hence "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" and **: "x \ set X" using * by simp_all then obtain y where y: "y \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ y \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" "x \ fv ((rm_vars (set X) \) y)" using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_obtain_var[of _ _ "rm_vars (set X) \"] by blast have "y \ set X" proof assume y_in: "y \ set X" hence "(rm_vars (set X) \) y = Var y" by auto hence "x = y" using y(2) by simp thus False using ** y_in by metis qed thus ?thesis using NegChecks y by auto qed (use * fv_subst_obtain_var in force)+ thus ?thesis by auto qed auto qed simp lemma fv\<^sub>s\<^sub>s\<^sub>t_subst_subset_range_vars_if_subset_domain: assumes "fv\<^sub>s\<^sub>s\<^sub>t S \ subst_domain \" shows "fv\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \) \ range_vars \" using assms fv\<^sub>s\<^sub>s\<^sub>t_subst_obtain_var[of _ S \] subst_dom_vars_in_subst[of _ \] subst_fv_imgI[of \] by (metis (no_types) in_mono subsetI) lemma fv\<^sub>s\<^sub>s\<^sub>t_in_fv_trms\<^sub>s\<^sub>s\<^sub>t: "x \ fv\<^sub>s\<^sub>s\<^sub>t S \ x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t S)" proof (induction S) case (Cons s S) thus ?case proof (cases "x \ fv\<^sub>s\<^sub>s\<^sub>t S") case False hence *: "x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p s" using Cons.prems by simp hence "x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p s)" proof (cases s) case (NegChecks X F G) hence "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" using * by simp_all thus ?thesis using * fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_in_fv_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s[of x] NegChecks by auto qed auto thus ?thesis by simp qed simp qed simp lemma fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_ground_subst_compose: assumes "subst_domain \ = subst_domain \" and "range_vars \ = {}" "range_vars \ = {}" shows "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \) = fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \)" proof - note 0 = fv_ground_subst_compose have 1: "range_vars \ \ set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) = {}" "range_vars \ \ set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) = {}" using assms(2,3) by (blast,blast) note 2 = 0[OF assms, of _ \] show ?thesis proof (cases a) case (NegChecks X F G) have 3: "range_vars \ \ set X = {}" "range_vars (rm_vars (set X) \) = {}" "range_vars \ \ set X = {}" "range_vars (rm_vars (set X) \) = {}" using assms(2,3) rm_vars_img_fv_subset[of "set X"] by auto have 4: "subst_domain (rm_vars (set X) \) = subst_domain (rm_vars (set X) \)" using assms(1) rm_vars_dom[of "set X"] by blast have 5: "fv (t \ rm_vars (set X) (\ \\<^sub>s \)) = fv (t \ rm_vars (set X) (\ \\<^sub>s \))" for t using 2[of t] rm_vars_comp[OF 3(1), of t \] rm_vars_comp[OF 3(3), of t \] 0[OF 4 3(2,4), of t "rm_vars (set X) \"] by argo have 6: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (H \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) (\ \\<^sub>s \)) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (H \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) (\ \\<^sub>s \))" for H proof - have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r (h \\<^sub>p rm_vars (set X) (\ \\<^sub>s \)) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r (h \\<^sub>p rm_vars (set X) (\ \\<^sub>s \))" for h proof - obtain s t where h: "h = (s,t)" by (metis surj_pair) show ?thesis using 5[of s] 5[of t] unfolding h by fast qed thus ?thesis unfolding subst_apply_pairs_def by auto qed show ?thesis using 5 6 unfolding NegChecks by simp qed (use 2 in auto) qed lemma fv\<^sub>s\<^sub>s\<^sub>t_ground_subst_compose: assumes "subst_domain \ = subst_domain \" and "range_vars \ = {}" "range_vars \ = {}" shows "fv\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \) = fv\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" by (induct S) (auto simp add: fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_ground_subst_compose[OF assms] fv\<^sub>s\<^sub>s\<^sub>t_Cons subst_sst_cons) lemma stateful_strand_step_subst_comp: assumes "range_vars \ \ set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p x) = {}" shows "x \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \ = (x \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" proof (cases x) case (NegChecks X F G) hence *: "range_vars \ \ set X = {}" using assms by simp have "H \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) (\ \\<^sub>s \) = (H \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \" for H using pairs_subst_comp rm_vars_comp[OF *] by (induct H) (auto simp add: subst_apply_pairs_def) thus ?thesis using NegChecks by simp qed simp_all lemma stateful_strand_subst_comp: assumes "range_vars \ \ bvars\<^sub>s\<^sub>s\<^sub>t S = {}" shows "S \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ = (S \\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>s\<^sub>t \" using assms proof (induction S) case (Cons s S) hence IH: "S \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \ = (S \\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>s\<^sub>t \" using Cons by auto have "s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \ = (s \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" using Cons.prems stateful_strand_step_subst_comp[of \ s \] unfolding range_vars_alt_def by auto thus ?case using IH by (simp add: subst_apply_stateful_strand_def) qed simp lemma subst_apply_bvars_disj_NegChecks: assumes "set X \ subst_domain \ = {}" shows "NegChecks X F G \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = NegChecks X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" proof - have "rm_vars (set X) \ = \" using assms rm_vars_apply'[of \ "set X"] by auto thus ?thesis by simp qed lemma subst_apply_NegChecks_no_bvars[simp]: "\[]\\\: F \\: F'\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \[]\\\: (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \\: (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)\" "\[]\\\: [] \\: F'\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \[]\\\: [] \\: (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)\" "\[]\\\: F \\: []\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \[]\\\: (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \\: []\" "\[]\\\: [] \\: [(t,s)]\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \[]\\\: [] \\: ([(t \ \,s \ \)])\" "\[]\\\: [(t,s)] \\: []\ \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ = \[]\\\: ([(t \ \,s \ \)]) \\: []\" by simp_all lemma setops\<^sub>s\<^sub>s\<^sub>t_mono: "set M \ set N \ setops\<^sub>s\<^sub>s\<^sub>t M \ setops\<^sub>s\<^sub>s\<^sub>t N" by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) lemma setops\<^sub>s\<^sub>s\<^sub>t_nil[simp]: "setops\<^sub>s\<^sub>s\<^sub>t [] = {}" by (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) lemma setops\<^sub>s\<^sub>s\<^sub>t_cons[simp]: "setops\<^sub>s\<^sub>s\<^sub>t (a#A) = setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \ setops\<^sub>s\<^sub>s\<^sub>t A" by (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) lemma setops\<^sub>s\<^sub>s\<^sub>t_cons_subset[simp]: "setops\<^sub>s\<^sub>s\<^sub>t A \ setops\<^sub>s\<^sub>s\<^sub>t (a#A)" using setops\<^sub>s\<^sub>s\<^sub>t_cons[of a A] by blast lemma setops\<^sub>s\<^sub>s\<^sub>t_append: "setops\<^sub>s\<^sub>s\<^sub>t (A@B) = setops\<^sub>s\<^sub>s\<^sub>t A \ setops\<^sub>s\<^sub>s\<^sub>t B" proof (induction A) case (Cons a A) thus ?case by (cases a) (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) qed (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) lemma setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p_member_iff: "(t,s) \ setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p x \ (x = Insert t s \ x = Delete t s \ (\ac. x = InSet ac t s) \ (\X F F'. x = NegChecks X F F' \ (t,s) \ set F'))" by (cases x) auto lemma setops\<^sub>s\<^sub>s\<^sub>t_member_iff: "(t,s) \ setops\<^sub>s\<^sub>s\<^sub>t A \ (Insert t s \ set A \ Delete t s \ set A \ (\ac. InSet ac t s \ set A) \ (\X F F'. NegChecks X F F' \ set A \ (t,s) \ set F'))" (is "?P \ ?Q") proof (induction A) case (Cons a A) thus ?case proof (cases "(t, s) \ setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p a") case True thus ?thesis using setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p_member_iff[of t s a] by auto qed auto qed simp lemma setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst: assumes "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) \ subst_domain \ = {}" shows "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" proof (cases a) case (NegChecks X F G) hence "rm_vars (set X) \ = \" using assms rm_vars_apply'[of \ "set X"] by auto hence "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = set (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ = set G \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" using NegChecks image_Un by simp_all thus ?thesis by (simp add: subst_apply_pairs_def) qed simp_all lemma setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst': assumes "\is_NegChecks a" shows "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p a \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" using assms by (cases a) auto lemma setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst'': fixes t::"('a,'b) term \ ('a,'b) term" and \::"('a,'b) subst" assumes t: "t \ setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" shows "\s \ setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p b. t = s \\<^sub>p rm_vars (set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b)) \" proof (cases "is_NegChecks b") case True then obtain X F G where b: "b = NegChecks X F G" by (cases b) moura+ hence "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p b = set G" "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = set (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b)) \)" by simp_all thus ?thesis using t subst_apply_pairs_pset_subst[of G] by blast next case False hence "setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) = setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \\<^sub>p\<^sub>s\<^sub>e\<^sub>t rm_vars (set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b)) \" using setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst' bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegChecks by fastforce thus ?thesis using t by blast qed lemma setops\<^sub>s\<^sub>s\<^sub>t_subst: assumes "bvars\<^sub>s\<^sub>s\<^sub>t S \ subst_domain \ = {}" shows "setops\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \) = setops\<^sub>s\<^sub>s\<^sub>t S \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" using assms proof (induction S) case (Cons a S) have "bvars\<^sub>s\<^sub>s\<^sub>t S \ subst_domain \ = {}" and *: "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a) \ subst_domain \ = {}" using Cons.prems by auto hence IH: "setops\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \) = setops\<^sub>s\<^sub>s\<^sub>t S \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" using Cons.IH by auto show ?case using setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst[OF *] IH unfolding setops\<^sub>s\<^sub>s\<^sub>t_def by (auto simp add: subst_apply_stateful_strand_def) qed (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) lemma setops\<^sub>s\<^sub>s\<^sub>t_subst': fixes p::"('a,'b) term \ ('a,'b) term" and \::"('a,'b) subst" assumes "p \ setops\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" shows "\s \ setops\<^sub>s\<^sub>s\<^sub>t S. \X. set X \ bvars\<^sub>s\<^sub>s\<^sub>t S \ p = s \\<^sub>p rm_vars (set X) \" using assms proof (induction S) case (Cons a S) note 0 = setops\<^sub>s\<^sub>s\<^sub>t_cons[of a S] bvars\<^sub>s\<^sub>s\<^sub>t_Cons[of a S] note 1 = setops\<^sub>s\<^sub>s\<^sub>t_cons[of "a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" "S \\<^sub>s\<^sub>s\<^sub>t \"] subst_sst_cons[of a S \] have "p \ setops\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \) \ p \ setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" using Cons.prems 1 by auto thus ?case proof assume *: "p \ setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" show ?thesis using setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst''[OF *] 0 by blast next assume *: "p \ setops\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" show ?thesis using Cons.IH[OF *] 0 by blast qed qed simp subsection \Stateful Constraint Semantics\ context intruder_model begin definition negchecks_model where "negchecks_model (\::('a,'b) subst) (D::('a,'b) dbstate) X F G \ (\\. subst_domain \ = set X \ ground (subst_range \) \ (\(t,s) \ set F. t \ \ \\<^sub>s \ \ s \ \ \\<^sub>s \) \ (\(t,s) \ set G. (t,s) \\<^sub>p \ \\<^sub>s \ \ D))" fun strand_sem_stateful:: "('fun,'var) terms \ ('fun,'var) dbstate \ ('fun,'var) stateful_strand \ ('fun,'var) subst \ bool" ("\_; _; _\\<^sub>s") where "\M; D; []\\<^sub>s = (\\. True)" | "\M; D; Send ts#S\\<^sub>s = (\\. (\t \ set ts. M \ t \ \) \ \M; D; S\\<^sub>s \)" | "\M; D; Receive ts#S\\<^sub>s = (\\. \(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; D; S\\<^sub>s \)" | "\M; D; Equality _ t t'#S\\<^sub>s = (\\. t \ \ = t' \ \ \ \M; D; S\\<^sub>s \)" | "\M; D; Insert t s#S\\<^sub>s = (\\. \M; insert ((t,s) \\<^sub>p \) D; S\\<^sub>s \)" | "\M; D; Delete t s#S\\<^sub>s = (\\. \M; D - {(t,s) \\<^sub>p \}; S\\<^sub>s \)" | "\M; D; InSet _ t s#S\\<^sub>s = (\\. (t,s) \\<^sub>p \ \ D \ \M; D; S\\<^sub>s \)" | "\M; D; NegChecks X F F'#S\\<^sub>s = (\\. negchecks_model \ D X F F' \ \M; D; S\\<^sub>s \)" lemmas strand_sem_stateful_induct = strand_sem_stateful.induct[case_names Nil ConsSnd ConsRcv ConsEq ConsIns ConsDel ConsIn ConsNegChecks] abbreviation constr_sem_stateful (infix "\\<^sub>s" 91) where "\ \\<^sub>s A \ \{}; {}; A\\<^sub>s \" lemma stateful_strand_sem_NegChecks_no_bvars: "\M; D; [\t not in s\]\\<^sub>s \ \ (t \ \, s \ \) \ D" "\M; D; [\t != s\]\\<^sub>s \ \ t \ \ \ s \ \" by (simp_all add: negchecks_model_def empty_dom_iff_empty_subst) lemma strand_sem_ik_mono_stateful: "\M; D; A\\<^sub>s \ \ \M \ M'; D; A\\<^sub>s \" proof (induction A arbitrary: M M' D rule: strand_sem_stateful.induct) case (2 M D ts S) hence "\t \ set ts. M \ t \ \" "\M \ M'; D; S\\<^sub>s \" by auto thus ?case using ideduct_mono[of M _ "M \ M'"] strand_sem_stateful.simps(2)[of "M \ M'" D ts S] by fastforce next case (3 M D ts S) hence "\((set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M) \ M'; D; S\\<^sub>s \" by simp hence "\(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ (M \ M'); D; S\\<^sub>s \" by (metis Un_assoc) thus ?case by simp qed simp_all lemma strand_sem_append_stateful: "\M; D; A@B\\<^sub>s \ \ \M; D; A\\<^sub>s \ \ \M \ (ik\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \); dbupd\<^sub>s\<^sub>s\<^sub>t A \ D; B\\<^sub>s \" (is "?P \ ?Q \ ?R") proof - have 1: "?P \ ?Q" by (induct A rule: strand_sem_stateful.induct) auto have 2: "?P \ ?R" proof (induction A arbitrary: M D B) case (Cons a A) thus ?case proof (cases a) case (Receive ts) have "(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ (M \ (ik\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \)) = M \ (ik\<^sub>s\<^sub>s\<^sub>t (a#A) \\<^sub>s\<^sub>e\<^sub>t \)" "dbupd\<^sub>s\<^sub>s\<^sub>t A \ D = dbupd\<^sub>s\<^sub>s\<^sub>t (a#A) \ D" using Receive by (auto simp add: ik\<^sub>s\<^sub>s\<^sub>t_def) thus ?thesis using Cons Receive by (metis (no_types, lifting) Un_assoc append_Cons strand_sem_stateful.simps(3)) qed (auto simp add: ik\<^sub>s\<^sub>s\<^sub>t_def) qed (simp add: ik\<^sub>s\<^sub>s\<^sub>t_def) have 3: "?Q \ ?R \ ?P" proof (induction A arbitrary: M D) case (Cons a A) thus ?case proof (cases a) case (Receive ts) have "(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ (M \ (ik\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \)) = M \ (ik\<^sub>s\<^sub>s\<^sub>t (a#A) \\<^sub>s\<^sub>e\<^sub>t \)" "dbupd\<^sub>s\<^sub>s\<^sub>t A \ D = dbupd\<^sub>s\<^sub>s\<^sub>t (a#A) \ D" using Receive by (auto simp add: ik\<^sub>s\<^sub>s\<^sub>t_def) thus ?thesis using Cons Receive by (metis (no_types, lifting) Un_assoc append_Cons strand_sem_stateful.simps(3)) qed (auto simp add: ik\<^sub>s\<^sub>s\<^sub>t_def) qed (simp add: ik\<^sub>s\<^sub>s\<^sub>t_def) show ?thesis by (metis 1 2 3) qed lemma negchecks_model_db_subset: fixes F F'::"(('a,'b) term \ ('a,'b) term) list" assumes "D' \ D" and "negchecks_model \ D X F F'" shows "negchecks_model \ D' X F F'" proof - have "\(s,t) \ set F'. (s,t) \\<^sub>p \ \\<^sub>s \ \ D'" when "\(s,t) \ set F'. (s,t) \\<^sub>p \ \\<^sub>s \ \ D" for \::"('a,'b) subst" using that assms(1) by blast thus ?thesis using assms(2) unfolding negchecks_model_def by meson qed lemma negchecks_model_db_supset: fixes F F'::"(('a,'b) term \ ('a,'b) term) list" assumes "D' \ D" and "\f \ set F'. \\. subst_domain \ = set X \ ground (subst_range \) \ f \\<^sub>p (\ \\<^sub>s \) \ D - D'" and "negchecks_model \ D' X F F'" shows "negchecks_model \ D X F F'" proof - have "\(s,t) \ set F'. (s,t) \\<^sub>p \ \\<^sub>s \ \ D" when "\(s,t) \ set F'. (s,t) \\<^sub>p \ \\<^sub>s \ \ D'" "subst_domain \ = set X \ ground (subst_range \)" for \::"('a,'b) subst" using that assms(1,2) by blast thus ?thesis using assms(3) unfolding negchecks_model_def by meson qed lemma negchecks_model_subst: fixes F F'::"(('a,'b) term \ ('a,'b) term) list" assumes "(subst_domain \ \ range_vars \) \ set X = {}" shows "negchecks_model (\ \\<^sub>s \) D X F F' \ negchecks_model \ D X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" proof - have 0: "\ \\<^sub>s (\ \\<^sub>s \) = \ \\<^sub>s (\ \\<^sub>s \)" when \: "subst_domain \ = set X" "ground (subst_range \)" for \ by (metis (no_types, lifting) \ subst_compose_assoc assms(1) inf_sup_aci(1) subst_comp_eq_if_disjoint_vars sup_inf_absorb range_vars_alt_def) { fix \::"('a,'b) subst" and t t' assume \: "subst_domain \ = set X" "ground (subst_range \)" and *: "\(s,t) \ set F. s \ (\ \\<^sub>s (\ \\<^sub>s \)) \ t \ (\ \\<^sub>s (\ \\<^sub>s \))" obtain f where f: "f \ set F" "fst f \ \ \\<^sub>s (\ \\<^sub>s \) \ snd f \ \ \\<^sub>s (\ \\<^sub>s \)" using * unfolding case_prod_unfold by (induct F) auto hence "(fst f \ \) \ \ \\<^sub>s \ \ (snd f \ \) \ \ \\<^sub>s \" using 0[OF \] by simp moreover have "(fst f \ \, snd f \ \) \ set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using f(1) by (auto simp add: subst_apply_pairs_def) ultimately have "\(s,t) \ set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \). s \ (\ \\<^sub>s \) \ t \ (\ \\<^sub>s \)" using f(1) by fastforce } moreover { fix \::"('a,'b) subst" and t t' assume \: "subst_domain \ = set X" "ground (subst_range \)" and *: "\(s,t) \ set F'. (s,t) \\<^sub>p \ \\<^sub>s (\ \\<^sub>s \) \ D" obtain f where f: "f \ set F'" "f \\<^sub>p \ \\<^sub>s (\ \\<^sub>s \) \ D" using * by (induct F') auto hence "f \\<^sub>p \ \\<^sub>p \ \\<^sub>s \ \ D" using 0[OF \] by (metis subst_pair_compose) moreover have "f \\<^sub>p \ \ set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using f(1) by (auto simp add: subst_apply_pairs_def) ultimately have "\(s,t) \ set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \). (s,t) \\<^sub>p \ \\<^sub>s \ \ D" using f(1) by (metis (no_types, lifting) case_prodI2) } moreover { fix \::"('a,'b) subst" and t t' assume \: "subst_domain \ = set X" "ground (subst_range \)" and *: "\(s,t) \ set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \). s \ (\ \\<^sub>s \) \ t \ (\ \\<^sub>s \)" obtain f where f: "f \ set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" "fst f \ \ \\<^sub>s \ \ snd f \ \ \\<^sub>s \" using * by (induct F) auto then obtain g where g: "g \ set F" "f = g \\<^sub>p \" by (auto simp add: subst_apply_pairs_def) have "fst g \ \ \\<^sub>s (\ \\<^sub>s \) \ snd g \ \ \\<^sub>s (\ \\<^sub>s \)" using f(2) g 0[OF \] by (simp add: prod.case_eq_if) hence "\(s,t) \ set F. s \ (\ \\<^sub>s (\ \\<^sub>s \)) \ t \ (\ \\<^sub>s (\ \\<^sub>s \))" using g by fastforce } moreover { fix \::"('a,'b) subst" and t t' assume \: "subst_domain \ = set X" "ground (subst_range \)" and *: "\(s,t) \ set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \). (s,t) \\<^sub>p (\ \\<^sub>s \) \ D" obtain f where f: "f \ set (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" "f \\<^sub>p \ \\<^sub>s \ \ D" using * by (induct F') auto then obtain g where g: "g \ set F'" "f = g \\<^sub>p \" by (auto simp add: subst_apply_pairs_def) have "g \\<^sub>p \ \\<^sub>s (\ \\<^sub>s \) \ D" using f(2) g 0[OF \] by (simp add: prod.case_eq_if) hence "\(s,t) \ set F'. (s,t) \\<^sub>p (\ \\<^sub>s (\ \\<^sub>s \)) \ D" using g by (metis (no_types, lifting) case_prodI2) } ultimately show ?thesis using assms unfolding negchecks_model_def by meson qed lemma strand_sem_subst_stateful: fixes \::"('fun,'var) subst" assumes "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>s\<^sub>t S = {}" shows "\M; D; S\\<^sub>s (\ \\<^sub>s \) \ \M; D; S \\<^sub>s\<^sub>s\<^sub>t \\\<^sub>s \" proof note [simp] = subst_sst_cons[of _ _ \] subst_subst_compose[of _ \ \] have "(subst_domain \ \ range_vars \) \ (subst_domain \ \ range_vars \) = {}" when \: "(subst_domain \ \ range_vars \) \ set X = {}" and \: "subst_domain \ = set X" "ground (subst_range \)" for X and \::"('fun,'var) subst" using \ \ unfolding range_vars_alt_def by auto hence 0: "\ \\<^sub>s \ = \ \\<^sub>s \" when \: "(subst_domain \ \ range_vars \) \ set X = {}" and \: "subst_domain \ = set X" "ground (subst_range \)" for \ X by (metis \ \ subst_comp_eq_if_disjoint_vars) show "\M; D; S\\<^sub>s (\ \\<^sub>s \) \ \M; D; S \\<^sub>s\<^sub>s\<^sub>t \\\<^sub>s \" using assms proof (induction S arbitrary: M D rule: strand_sem_stateful_induct) case (ConsSnd M D ts S) hence "\t \ set ts. M \ t \ \ \ \" and IH: "\M; D; S \\<^sub>s\<^sub>s\<^sub>t \\\<^sub>s \" by auto hence "\t \ set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \). M \ t \ \" by simp thus ?case using IH by simp next case (ConsRcv M D ts S) hence "\(set ts \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \) \ M; D; S \\<^sub>s\<^sub>s\<^sub>t \\\<^sub>s \" by simp hence "\(set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \) \ M; D; S \\<^sub>s\<^sub>s\<^sub>t \\\<^sub>s \" by (metis list.set_map subst_comp_all) thus ?case by simp next case (ConsNegChecks M D X F F' S) hence *: "\M; D; S \\<^sub>s\<^sub>s\<^sub>t \\\<^sub>s \" and **: "(subst_domain \ \ range_vars \) \ set X = {}" unfolding bvars\<^sub>s\<^sub>s\<^sub>t_def negchecks_model_def by (force, auto) have "negchecks_model (\ \\<^sub>s \) D X F F'" using ConsNegChecks by auto hence "negchecks_model \ D X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using 0[OF **] negchecks_model_subst[OF **] by blast moreover have "rm_vars (set X) \ = \" using ConsNegChecks.prems(2) by force ultimately show ?case using * by auto qed simp_all show "\M; D; S \\<^sub>s\<^sub>s\<^sub>t \\\<^sub>s \ \ \M; D; S\\<^sub>s (\ \\<^sub>s \)" using assms proof (induction S arbitrary: M D rule: strand_sem_stateful_induct) case (ConsSnd M D ts S) hence "\t \ set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \). M \ t \ \" and IH: "\M; D; S\\<^sub>s (\ \\<^sub>s \)" by auto hence "\t \ set ts. M \ t \ \ \ \" by simp thus ?case using IH by simp next case (ConsRcv M D ts S) hence "\(set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \) \ M; D; S\\<^sub>s (\ \\<^sub>s \)" by simp hence "\(set ts \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \) \ M; D; S\\<^sub>s (\ \\<^sub>s \)" by (metis list.set_map subst_comp_all) thus ?case by simp next case (ConsNegChecks M D X F F' S) have \: "rm_vars (set X) \ = \" using ConsNegChecks.prems(2) by force hence *: "\M; D; S\\<^sub>s (\ \\<^sub>s \)" and **: "(subst_domain \ \ range_vars \) \ set X = {}" using ConsNegChecks unfolding bvars\<^sub>s\<^sub>s\<^sub>t_def negchecks_model_def by auto have "negchecks_model \ D X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) (F' \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using ConsNegChecks.prems(1) \ by (auto simp add: subst_compose_assoc negchecks_model_def) hence "negchecks_model (\ \\<^sub>s \) D X F F'" using 0[OF **] negchecks_model_subst[OF **] by blast thus ?case using * by auto qed simp_all qed lemma strand_sem_receive_prepend_stateful: assumes "\M; D; S\\<^sub>s \" and "list_all is_Receive S'" shows "\M; D; S@S'\\<^sub>s \" using assms(2) proof (induction S' rule: List.rev_induct) case (snoc x S') hence "\t. x = receive\t\" "\M; D; S@S'\\<^sub>s \" unfolding list_all_iff is_Receive_def by auto thus ?case using strand_sem_append_stateful[of M D "S@S'" "[x]" \] by auto qed (simp add: assms(1)) lemma negchecks_model_model_swap: fixes I J::"('a,'b) subst" assumes "\x \ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G) - set X. I x = J x" and "negchecks_model I D X F G" shows "negchecks_model J D X F G" proof - have 0: "\x \ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G). (\ \\<^sub>s I) x = (\ \\<^sub>s J) x" when "subst_domain \ = set X" "ground (subst_range \)" for \::"('a,'b) subst" using that assms(1) by (metis DiffI ground_subst_range_empty_fv subst_comp_notin_dom_eq subst_ground_ident_compose(1)) have 1: "s \ (\ \\<^sub>s J) \ t \ (\ \\<^sub>s J)" when "s \ (\ \\<^sub>s I) \ t \ (\ \\<^sub>s I)" "(s,t) \ set F" "subst_domain \ = set X" "ground (subst_range \)" for \::"('a,'b) subst" and s t using that(1,2) 0[OF that(3,4)] term_subst_eq_conv[of _ "\ \\<^sub>s I" "\ \\<^sub>s J"] UnCI fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_inI(4,5)[OF that(2)] by metis have 2: "(s,t) \\<^sub>p (\ \\<^sub>s J) \ D" when "(s,t) \\<^sub>p (\ \\<^sub>s I) \ D" "(s,t) \ set G" "subst_domain \ = set X" "ground (subst_range \)" for \::"('a,'b) subst" and s t using that(1,2) 0[OF that(3,4)] fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_inI(4,5)[of s t G] term_subst_eq_conv[of s "\ \\<^sub>s I" "\ \\<^sub>s J"] term_subst_eq_conv[of t "\ \\<^sub>s I" "\ \\<^sub>s J"] by simp have 3: "(\(s,t) \ set F. s \ \ \\<^sub>s J \ t \ \ \\<^sub>s J) \ (\(s,t) \ set G. (s, t) \\<^sub>p \ \\<^sub>s J \ D)" when "subst_domain \ = set X" "ground (subst_range \)" "(\(s,t) \ set F. s \ \ \\<^sub>s I \ t \ \ \\<^sub>s I) \ (\(s,t) \ set G. (s, t) \\<^sub>p \ \\<^sub>s I \ D)" for \::"('a,'b) subst" using 1[OF _ _ that(1,2)] 2[OF _ _ that(1,2)] that(3) by blast thus ?thesis using assms(2) unfolding negchecks_model_def by simp qed lemma strand_sem_model_swap: assumes "\x \ fv\<^sub>s\<^sub>s\<^sub>t S. I x = J x" and "\M; D; S\\<^sub>s I" shows "\M; D; S\\<^sub>s J" using assms(2,1) proof (induction S arbitrary: M D rule: strand_sem_stateful_induct) case (ConsSnd M D ts S) hence *: "\M; D; S\\<^sub>s J" "\t \ set ts. M \ t \ I" "\x \ fv\<^sub>s\<^sub>e\<^sub>t (set ts). I x = J x" by (fastforce, fastforce, fastforce) have "\t \ set ts. M \ t \ J" using *(2,3) term_subst_eq_conv[of _ I J] by (metis fv_subset subsetD) thus ?case using *(1) by auto next case (ConsRcv M D ts S) hence *: "\(set ts \\<^sub>s\<^sub>e\<^sub>t I) \ M; D; S\\<^sub>s J" "\x \ fv\<^sub>s\<^sub>e\<^sub>t (set ts). I x = J x" by (fastforce, fastforce) have "set ts \\<^sub>s\<^sub>e\<^sub>t I = set ts \\<^sub>s\<^sub>e\<^sub>t J" using *(2) term_subst_eq_conv[of _ I J] by (meson fv_subset image_cong subsetD) thus ?case using *(1) by simp next case (ConsEq M D ac t t' S) thus ?case using term_subst_eq_conv[of t I J] term_subst_eq_conv[of t' I J] by force next case (ConsIns M D t s S) thus ?case using term_subst_eq_conv[of t I J] term_subst_eq_conv[of s I J] by force next case (ConsDel M D t s S) thus ?case using term_subst_eq_conv[of t I J] term_subst_eq_conv[of s I J] by force next case (ConsIn M D uv t s S) thus ?case using term_subst_eq_conv[of t I J] term_subst_eq_conv[of s I J] by force next case (ConsNegChecks M D X F F' S) hence "\M; D; S\\<^sub>s J" "negchecks_model I D X F F'" "\x\fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' - set X. I x = J x" by (fastforce, fastforce, fastforce) thus ?case using negchecks_model_model_swap[of F F' X I J D] by fastforce qed simp lemma strand_sem_receive_send_append: assumes A: "\M; D; A\\<^sub>s I" shows "\M; D; A@[receive\[t]\,send\[t]\]\\<^sub>s I" proof - have "M \ (ik\<^sub>s\<^sub>s\<^sub>t (A@[receive\[t]\]) \\<^sub>s\<^sub>e\<^sub>t I) \ t \ I" using in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of t "A@[receive\[t]\]"] by auto thus ?thesis using A strand_sem_append_stateful[of M D A "[receive\[t]\]" I] strand_sem_append_stateful[of M D "A@[receive\[t]\]" "[send\[t]\]" I] by force qed lemma strand_sem_stateful_if_no_send_or_check: assumes A: "list_all (\a. \is_Send a \ \is_Check_or_Assignment a) A" shows "\M; D; A\\<^sub>s I" using A proof (induction A rule: List.rev_induct) case (snoc a A) hence IH: "\M; D; A\\<^sub>s I" and a: "\is_Send a" "\is_Check_or_Assignment a" by auto from a have "\M; D; [a]\\<^sub>s I" for M D by (cases a) auto thus ?case using IH strand_sem_append_stateful[of M D A "[a]" I] by blast qed simp lemma strand_sem_stateful_if_sends_deduct: assumes "list_all is_Send A" and "\ts. send\ts\ \ set A \ (\t \ set ts. M \ t \ I)" shows "\M; D; A\\<^sub>s I" using assms proof (induction A rule: List.rev_induct) case (snoc a A) hence IH: "\M; D; A\\<^sub>s I" by auto obtain ts where a: "a = send\ts\" "\t \ set ts. M \ t \ I" using snoc.prems by (cases a) auto have "\M \ (ik\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I); D; [a]\\<^sub>s I" for D using ideduct_mono[of M _ "M \ (ik\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I)"] a by auto thus ?case using IH strand_sem_append_stateful[of M D A "[a]" I] by blast qed simp lemma strand_sem_stateful_if_checks: assumes "list_all is_Check_or_Assignment A" and "\ac t s. \ac: t \ s\ \ set A \ t \ I = s \ I" and "\ac t s. \ac: t \ s\ \ set A \ (t \ I, s \ I) \ D" and "\X F G. \X\\\: F \\: G\ \ set A \ negchecks_model I D X F G" shows "\M; D; A\\<^sub>s I" using assms proof (induction A rule: List.rev_induct) case (snoc a A) hence IH: "\M; D; A\\<^sub>s I" and a: "is_Check_or_Assignment a" by auto have 0: "dbupd\<^sub>s\<^sub>s\<^sub>t A I D = D" using snoc.prems(1) dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd[of A I D] unfolding list_all_iff by auto have "\M; D; [a]\\<^sub>s I" for M using a snoc.prems(2,3,4) by (cases a) auto thus ?case using IH strand_sem_append_stateful[of M D A "[a]" I] unfolding 0 by blast qed simp lemma strand_sem_stateful_sends_deduct: assumes A: "\M; D; A\\<^sub>s I" and ts: "send\ts\ \ set A" and t: "t \ set ts" shows "M \ (ik\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I) \ t \ I" using A ts proof (induction A arbitrary: M D rule: List.rev_induct) case (snoc a A) have 0: "\M; D; A\\<^sub>s I" using strand_sem_append_stateful snoc.prems(1) by fast have 1: "M \ (ik\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I) \ M \ (ik\<^sub>s\<^sub>s\<^sub>t (A@[a]) \\<^sub>s\<^sub>e\<^sub>t I)" by auto have 2: "M \ (ik\<^sub>s\<^sub>s\<^sub>t (A@[send\ts\]) \\<^sub>s\<^sub>e\<^sub>t I) = M \ (ik\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I)" using in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of _ A] in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of _ "A@[send\ts\]"] by fastforce show ?case proof (cases "send\ts\ \ set A") case True show ?thesis by (rule ideduct_mono[OF snoc.IH[OF 0 True] 1]) next case False hence a: "a = send\ts\" using snoc.prems(2) by force show ?thesis using strand_sem_append_stateful[of M D A "[a]" I] snoc.prems(1) t unfolding a 2 by auto qed qed simp end subsection \Well-Formedness Lemmata\ lemma wfvarsocc\<^sub>s\<^sub>s\<^sub>t_subset_wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t[simp]: "wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S" by (induction S) (auto simp add: wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def split: stateful_strand_step.split poscheckvariant.split) lemma wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_append: "wfvarsoccs\<^sub>s\<^sub>s\<^sub>t (S@S') = wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S'" by (simp add: wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def) lemma wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_union[simp]: "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t (S@T) = wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t T" by (simp add: wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def) lemma wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_singleton: "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t [s] = wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p s" by (simp add: wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def) lemma ik\<^sub>s\<^sub>s\<^sub>t_fv_subset_wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>s\<^sub>t S) \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S" using in_ik\<^sub>s\<^sub>s\<^sub>t_iff[of _ S] unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by force lemma wf\<^sub>s\<^sub>s\<^sub>t_prefix[dest]: "wf'\<^sub>s\<^sub>s\<^sub>t V (S@S') \ wf'\<^sub>s\<^sub>s\<^sub>t V S" by (induct S rule: wf'\<^sub>s\<^sub>s\<^sub>t.induct) auto lemma wf\<^sub>s\<^sub>s\<^sub>t_vars_mono: "wf'\<^sub>s\<^sub>s\<^sub>t V S \ wf'\<^sub>s\<^sub>s\<^sub>t (V \ W) S" proof (induction S arbitrary: V) case (Cons x S) thus ?case proof (cases x) case (Send ts) have "wf'\<^sub>s\<^sub>s\<^sub>t (V \ W \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) S" using Cons.prems(1) Cons.IH Send by (metis sup_assoc sup_commute wf'\<^sub>s\<^sub>s\<^sub>t.simps(3)) thus ?thesis by (simp add: Send) next case (Equality a t t') show ?thesis proof (cases a) case Assign hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t \ W) S" "fv t' \ V \ W" using Equality Cons.prems(1) Cons.IH by auto thus ?thesis using Equality Assign by (simp add: sup_commute sup_left_commute) next case Check thus ?thesis using Equality Cons by auto qed next case (InSet a t t') show ?thesis proof (cases a) case Assign hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t \ fv t' \ W) S" using InSet Cons.prems(1) Cons.IH by auto thus ?thesis using InSet Assign by (simp add: sup_commute sup_left_commute) next case Check thus ?thesis using InSet Cons by auto qed qed auto qed simp lemma wf\<^sub>s\<^sub>s\<^sub>tI[intro]: "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S \ V \ wf'\<^sub>s\<^sub>s\<^sub>t V S" proof (induction S) case (Cons x S) thus ?case proof (cases x) case (Send ts) hence "wf'\<^sub>s\<^sub>s\<^sub>t V S" "V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) = V" using Cons unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto thus ?thesis using Send by simp next case (Equality a t t') show ?thesis proof (cases a) case Assign hence "wf'\<^sub>s\<^sub>s\<^sub>t V S" "fv t' \ V" using Equality Cons unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto thus ?thesis using wf\<^sub>s\<^sub>s\<^sub>t_vars_mono Equality Assign by simp next case Check thus ?thesis using Equality Cons unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto qed next case (InSet a t t') show ?thesis proof (cases a) case Assign hence "wf'\<^sub>s\<^sub>s\<^sub>t V S" "fv t \ fv t' \ V" using InSet Cons unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto thus ?thesis using wf\<^sub>s\<^sub>s\<^sub>t_vars_mono InSet Assign by (simp add: Un_assoc) next case Check thus ?thesis using InSet Cons unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto qed qed (simp_all add: wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def) qed (simp add: wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def) lemma wf\<^sub>s\<^sub>s\<^sub>tI'[intro]: assumes "\((\x. case x of Receive ts \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) | Equality Assign _ t' \ fv t' | Insert t t' \ fv t \ fv t' | _ \ {}) ` set S) \ V" shows "wf'\<^sub>s\<^sub>s\<^sub>t V S" using assms proof (induction S) case (Cons x S) thus ?case proof (cases x) case (Equality a t t') thus ?thesis using Cons by (cases a) (auto simp add: wf\<^sub>s\<^sub>s\<^sub>t_vars_mono) next case (InSet a t t') thus ?thesis using Cons by (cases a) (auto simp add: wf\<^sub>s\<^sub>s\<^sub>t_vars_mono Un_assoc) qed (simp_all add: wf\<^sub>s\<^sub>s\<^sub>t_vars_mono) qed simp lemma wf\<^sub>s\<^sub>s\<^sub>t_append_exec: "wf'\<^sub>s\<^sub>s\<^sub>t V (S@S') \ wf'\<^sub>s\<^sub>s\<^sub>t (V \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S) S'" proof (induction S arbitrary: V) case (Cons x S V) thus ?case proof (cases x) case (Send ts) hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S) S'" using Cons.prems Cons.IH by simp thus ?thesis using Send unfolding wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def by (auto simp add: sup_assoc) next case (Equality a t t') show ?thesis proof (cases a) case Assign hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S) S'" using Equality Cons.prems Cons.IH by auto thus ?thesis using Equality Assign unfolding wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def by (auto simp add: sup_assoc) next case Check hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S) S'" using Equality Cons.prems Cons.IH by auto thus ?thesis using Equality Check unfolding wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def by (auto simp add: sup_assoc) qed next case (InSet a t t') show ?thesis proof (cases a) case Assign hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t \ fv t' \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S) S'" using InSet Cons.prems Cons.IH by auto thus ?thesis using InSet Assign unfolding wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def by (auto simp add: sup_assoc) next case Check hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S) S'" using InSet Cons.prems Cons.IH by auto thus ?thesis using InSet Check unfolding wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def by (auto simp add: sup_assoc) qed qed (auto simp add: wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def) qed (simp add: wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def) lemma wf\<^sub>s\<^sub>s\<^sub>t_append: "wf'\<^sub>s\<^sub>s\<^sub>t X S \ wf'\<^sub>s\<^sub>s\<^sub>t Y T \ wf'\<^sub>s\<^sub>s\<^sub>t (X \ Y) (S@T)" proof (induction X S rule: wf'\<^sub>s\<^sub>s\<^sub>t.induct) case 1 thus ?case by (metis wf\<^sub>s\<^sub>s\<^sub>t_vars_mono Un_commute append_Nil) next case 3 thus ?case by (metis append_Cons Un_commute Un_assoc wf'\<^sub>s\<^sub>s\<^sub>t.simps(3)) next case (4 V t t' S) hence *: "fv t' \ V" and "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t \ Y) (S @ T)" by simp_all hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ Y \ fv t) (S @ T)" by (metis Un_commute Un_assoc) thus ?case using * by auto next case (8 V t t' S) hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t \ fv t' \ Y) (S @ T)" by simp_all hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ Y \ fv t \ fv t') (S @ T)" by (metis Un_commute Un_assoc) thus ?case by auto qed auto lemma wf\<^sub>s\<^sub>s\<^sub>t_append_suffix: "wf'\<^sub>s\<^sub>s\<^sub>t V S \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S \ V \ wf'\<^sub>s\<^sub>s\<^sub>t V (S@S')" proof (induction V S rule: wf'\<^sub>s\<^sub>s\<^sub>t.induct) case (2 V ts S) hence *: "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ V" "wf'\<^sub>s\<^sub>s\<^sub>t V S" by simp_all hence "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S \ V" using "2.prems"(2) unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by fastforce thus ?case using "2.IH" * by simp next case (3 V ts S) hence *: "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) S" by simp_all hence "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S \ (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts))" using "3.prems"(2) unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto thus ?case using "3.IH" * by simp next case (4 V t t' S) hence *: "fv t' \ V" "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t) S" by simp_all moreover have "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (\t := t'\) = fv t \ fv t'" by simp moreover have "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t (\t := t'\#S) = fv t \ fv t' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S" unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto ultimately have "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S \ (V \ fv t)" using "4.prems"(2) by blast thus ?case using "4.IH" * by simp next case (6 V t t' S) hence *: "fv t \ fv t' \ V" "wf'\<^sub>s\<^sub>s\<^sub>t V S" by simp_all moreover have "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (insert\t,t'\) = fv t \ fv t'" by simp moreover have "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t (insert\t,t'\#S) = fv t \ fv t' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S" unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto ultimately have "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S \ V" using "6.prems"(2) by blast thus ?case using "6.IH" * by simp next case (8 V t t' S) hence *: "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t \ fv t') S" by simp_all moreover have "vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (select\t,t'\) = fv t \ fv t'" by simp moreover have "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t (select\t,t'\#S) = fv t \ fv t' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S" unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto ultimately have "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S \ (V \ fv t \ fv t')" using "8.prems"(2) by blast thus ?case using "8.IH" * by simp qed (simp_all add: wf\<^sub>s\<^sub>s\<^sub>tI wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def) lemma wf\<^sub>s\<^sub>s\<^sub>t_append_suffix': assumes "wf'\<^sub>s\<^sub>s\<^sub>t V S" and "\((\x. case x of Receive ts \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) | Equality Assign _ t' \ fv t' | Insert t t' \ fv t \ fv t' | _ \ {}) ` set S') \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S \ V" shows "wf'\<^sub>s\<^sub>s\<^sub>t V (S@S')" using assms by (induction V S rule: wf'\<^sub>s\<^sub>s\<^sub>t.induct) (auto simp add: wf\<^sub>s\<^sub>s\<^sub>tI' wf\<^sub>s\<^sub>s\<^sub>t_vars_mono wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def) lemma wf\<^sub>s\<^sub>s\<^sub>t_subst_apply: "wf'\<^sub>s\<^sub>s\<^sub>t V S \ wf'\<^sub>s\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) (S \\<^sub>s\<^sub>s\<^sub>t \)" proof (induction S arbitrary: V rule: wf'\<^sub>s\<^sub>s\<^sub>t.induct) case (2 V ts S) hence "wf'\<^sub>s\<^sub>s\<^sub>t V S" "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ V" by simp_all hence "wf'\<^sub>s\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) (S \\<^sub>s\<^sub>s\<^sub>t \)" "fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using "2.IH" subst_apply_fv_subset by (simp, force) thus ?case by (simp add: subst_apply_stateful_strand_def) next case (3 V ts S) hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) S" by simp hence "wf'\<^sub>s\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)))) (S \\<^sub>s\<^sub>s\<^sub>t \)" using "3.IH" by metis hence "wf'\<^sub>s\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)) (S \\<^sub>s\<^sub>s\<^sub>t \)" using subst_apply_fv_unfold_set[of \ ts] by fastforce thus ?case by (simp add: subst_apply_stateful_strand_def) next case (4 V t t' S) hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t) S" "fv t' \ V" by auto hence "wf'\<^sub>s\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ fv t))) (S \\<^sub>s\<^sub>s\<^sub>t \)" and *: "fv (t' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using "4.IH" subst_apply_fv_subset by force+ hence "wf'\<^sub>s\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv (t \ \)) (S \\<^sub>s\<^sub>s\<^sub>t \)" by (metis subst_apply_fv_union) thus ?case using * by (simp add: subst_apply_stateful_strand_def) next case (6 V t t' S) hence "wf'\<^sub>s\<^sub>s\<^sub>t V S" "fv t \ fv t' \ V" by auto hence "wf'\<^sub>s\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) (S \\<^sub>s\<^sub>s\<^sub>t \)" "fv (t \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" "fv (t' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using "6.IH" subst_apply_fv_subset by force+ thus ?case by (simp add: sup_assoc subst_apply_stateful_strand_def) next case (8 V t t' S) hence "wf'\<^sub>s\<^sub>s\<^sub>t (V \ fv t \ fv t') S" by auto hence "wf'\<^sub>s\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ fv t \ fv t'))) (S \\<^sub>s\<^sub>s\<^sub>t \)" using "8.IH" subst_apply_fv_subset by force hence "wf'\<^sub>s\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv (t \ \) \ fv (t' \ \)) (S \\<^sub>s\<^sub>s\<^sub>t \)" by (metis subst_apply_fv_union) thus ?case by (simp add: subst_apply_stateful_strand_def) qed (auto simp add: subst_apply_stateful_strand_def) lemma wf\<^sub>s\<^sub>s\<^sub>t_induct[consumes 1, case_names Nil ConsSnd ConsRcv ConsEq ConsEq2 ConsIn ConsIns ConsDel ConsNegChecks]: fixes S::"('a,'b) stateful_strand" assumes "wf'\<^sub>s\<^sub>s\<^sub>t V S" "P []" "\ts S. \wf'\<^sub>s\<^sub>s\<^sub>t V S; P S\ \ P (S@[Send ts])" "\ts S. \wf'\<^sub>s\<^sub>s\<^sub>t V S; P S; fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ V \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S\ \ P (S@[Receive ts])" "\t t' S. \wf'\<^sub>s\<^sub>s\<^sub>t V S; P S; fv t' \ V \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S\ \ P (S@[Equality Assign t t'])" "\t t' S. \wf'\<^sub>s\<^sub>s\<^sub>t V S; P S\ \ P (S@[Equality Check t t'])" "\ac t t' S. \wf'\<^sub>s\<^sub>s\<^sub>t V S; P S\ \ P (S@[InSet ac t t'])" "\t t' S. \wf'\<^sub>s\<^sub>s\<^sub>t V S; P S; fv t \ fv t' \ V \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S\ \ P (S@[Insert t t'])" "\t t' S. \wf'\<^sub>s\<^sub>s\<^sub>t V S; P S\ \ P (S@[Delete t t'])" "\X F G S. \wf'\<^sub>s\<^sub>s\<^sub>t V S; P S\ \ P (S@[NegChecks X F G])" shows "P S" using assms proof (induction S rule: List.rev_induct) case (snoc x S) hence *: "wf'\<^sub>s\<^sub>s\<^sub>t V S" "wf'\<^sub>s\<^sub>s\<^sub>t (V \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S) [x]" by (metis wf\<^sub>s\<^sub>s\<^sub>t_prefix, metis wf\<^sub>s\<^sub>s\<^sub>t_append_exec) have IH: "P S" using snoc.IH[OF *(1)] snoc.prems by auto note ** = snoc.prems(3-)[OF *(1) IH] *(2) show ?case proof (cases x) case (Send ts) thus ?thesis using **(1) by blast next case (Receive ts) thus ?thesis using **(2,9) by simp next case (Equality ac t t') thus ?thesis using **(3,4,9) by (cases ac) auto next case (Insert t t') thus ?thesis using **(6,9) by force next case (Delete t t') thus ?thesis using **(7) by presburger next case (NegChecks X F G) thus ?thesis using **(8) by presburger next case (InSet ac t t') thus ?thesis using **(5) by (cases ac) auto qed qed simp lemma wf\<^sub>s\<^sub>s\<^sub>t_strand_first_Send_var_split: assumes "wf'\<^sub>s\<^sub>s\<^sub>t {} S" "\v \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S. t \ \ \ \ v" shows "\S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f. \(\w \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e. t \ \ \ \ w) \ ( (\ts. S = S\<^sub>p\<^sub>r\<^sub>e@send\ts\#S\<^sub>s\<^sub>u\<^sub>f \ t \ \ \\<^sub>s\<^sub>e\<^sub>t set ts \\<^sub>s\<^sub>e\<^sub>t \) \ (\s u. S = S\<^sub>p\<^sub>r\<^sub>e@\assign: s \ u\#S\<^sub>s\<^sub>u\<^sub>f \ t \ \ \ s \ \ \ \(t \ \ \\<^sub>s\<^sub>e\<^sub>t \ ` fv u)) \ (\s u. S = S\<^sub>p\<^sub>r\<^sub>e@\assign: s \ u\#S\<^sub>s\<^sub>u\<^sub>f \ (t \ \ \ s \ \ \ t \ \ \ u \ \)))" (is "\S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f. ?P S\<^sub>p\<^sub>r\<^sub>e \ ?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f") using assms proof (induction S rule: wf\<^sub>s\<^sub>s\<^sub>t_induct) case (ConsSnd ts' S) show ?case proof (cases "\w \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S. t \ \ \ \ w") case True then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsSnd.IH by moura thus ?thesis by fastforce next case False then obtain v where v: "v \ fv\<^sub>s\<^sub>e\<^sub>t (set ts')" "t \ \ \ \ v" using ConsSnd.prems unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto then obtain t' where t': "t' \ set ts'" "v \ fv t'" by auto have "t \ \ \ t' \ \" using v(2) t'(2) subst_mono[of "Var v" t' \] vars_iff_subtermeq[of v] term.order_trans by auto hence "t \ \ \\<^sub>s\<^sub>e\<^sub>t set ts' \\<^sub>s\<^sub>e\<^sub>t \" using v(1) t'(1) by blast thus ?thesis using False v by blast qed next case (ConsRcv t' S) have "fv\<^sub>s\<^sub>e\<^sub>t (set t') \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S" using ConsRcv.hyps wfvarsocc\<^sub>s\<^sub>s\<^sub>t_subset_wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t[of S] by blast hence "\v \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S. t \ \ \ \ v" using ConsRcv.prems unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by fastforce then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsRcv.IH by moura thus ?case by fastforce next case (ConsEq s s' S) have *: "fv s' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S" using ConsEq.hyps wfvarsocc\<^sub>s\<^sub>s\<^sub>t_subset_wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t[of S] by blast show ?case proof (cases "\v \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S. t \ \ \ \ v") case True then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsEq.IH by moura thus ?thesis by fastforce next case False then obtain v where "v \ fv s" "t \ \ \ \ v" and **: "fv s' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S" using ConsEq.prems * unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto hence "t \ \ \ s \ \" using vars_iff_subtermeq[of v s] subst_mono[of "Var v" s \] term.order_trans by auto thus ?thesis using False ** by fastforce qed next case (ConsEq2 s s' S) have "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t (S@[Equality Check s s']) = wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S" unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by auto hence "\v \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S. t \ \ \ \ v" using ConsEq2.prems by metis then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsEq2.IH by moura thus ?case by fastforce next case (ConsNegChecks X F G S) hence "\v \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S. t \ \ \ \ v" unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by simp then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsNegChecks.IH by moura thus ?case by fastforce next case (ConsIn ac s s' S) show ?case proof (cases "\v \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S. t \ \ \ \ v") case True then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsIn.IH by moura thus ?thesis by fastforce next case False hence ac: "ac = assign" using ConsIn.prems unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by (cases ac) auto obtain v where "v \ fv s \ fv s'" "t \ \ \ \ v" using ConsIn.prems False unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def ac by auto hence *: "t \ \ \ s \ \ \ t \ \ \ s' \ \" using vars_iff_subtermeq[of v s] vars_iff_subtermeq[of v s'] subst_mono[of "Var v" s \] subst_mono[of "Var v" s' \] term.order_trans by auto show ?thesis using * False unfolding ac by fast qed next case (ConsIns s s' S) have "fv s \ fv s' \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S" using ConsIns.hyps wfvarsocc\<^sub>s\<^sub>s\<^sub>t_subset_wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t[of S] by blast hence "\v \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S. t \ \ \ \ v" using ConsIns.prems unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by fastforce then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsIns.IH by moura thus ?case by fastforce next case (ConsDel s s' S) hence "\v \ wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S. t \ \ \ \ v" unfolding wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by simp then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsDel.IH by moura thus ?case by fastforce qed (simp add: wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def) lemma wf\<^sub>s\<^sub>s\<^sub>t_vars_mono': "wf'\<^sub>s\<^sub>s\<^sub>t V S \ V \ W \ wf'\<^sub>s\<^sub>s\<^sub>t W S" by (metis Diff_partition wf\<^sub>s\<^sub>s\<^sub>t_vars_mono) lemma wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_receives_only_eq: assumes "list_all is_Receive S" shows "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S = fv\<^sub>s\<^sub>s\<^sub>t S" using assms proof (induction S) case (Cons a A) obtain ts where a: "a = receive\ts\" using Cons.prems by (cases a) auto have IH: "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t A = fv\<^sub>s\<^sub>s\<^sub>t A" using Cons.prems Cons.IH by simp show ?case using IH unfolding a wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def by simp qed (simp add: wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t_def) lemma wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_receives_only_empty: assumes "list_all is_Receive S" shows "wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S = {}" using assms proof (induction S) case (Cons a A) obtain ts where a: "a = receive\ts\" using Cons.prems by (cases a) auto have IH: "wfvarsoccs\<^sub>s\<^sub>s\<^sub>t A = {}" using Cons.prems Cons.IH by simp show ?case using IH unfolding a wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def by simp qed (simp add: wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def) lemma wf\<^sub>s\<^sub>s\<^sub>t_sends_only: assumes "list_all is_Send S" shows "wf'\<^sub>s\<^sub>s\<^sub>t V S" using assms proof (induction S arbitrary: V) case (Cons s S) thus ?case by (cases s) auto qed simp lemma wf\<^sub>s\<^sub>s\<^sub>t_sends_only_prepend: assumes "wf'\<^sub>s\<^sub>s\<^sub>t V S" and "list_all is_Send S'" shows "wf'\<^sub>s\<^sub>s\<^sub>t V (S'@S)" using wf\<^sub>s\<^sub>s\<^sub>t_append[OF wf\<^sub>s\<^sub>s\<^sub>t_sends_only[OF assms(2), of "{}"] assms(1)] by simp lemma wf\<^sub>s\<^sub>s\<^sub>t_receives_only_fv_subset: assumes "wf'\<^sub>s\<^sub>s\<^sub>t V S" and "list_all is_Receive S" shows "fv\<^sub>s\<^sub>s\<^sub>t S \ V" using assms proof (induction rule: wf\<^sub>s\<^sub>s\<^sub>t_induct) case (ConsRcv ts S) thus ?case using wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_receives_only_empty[of S] by auto qed auto lemma wf\<^sub>s\<^sub>s\<^sub>t_append_suffix'': assumes "wf'\<^sub>s\<^sub>s\<^sub>t V S" and "wfrestrictedvars\<^sub>s\<^sub>s\<^sub>t S' \ wfvarsoccs\<^sub>s\<^sub>s\<^sub>t S \ V" shows "wf'\<^sub>s\<^sub>s\<^sub>t V (S@S')" using assms by (induction V S rule: wf'\<^sub>s\<^sub>s\<^sub>t.induct) (auto simp add: wf\<^sub>s\<^sub>s\<^sub>tI' wf\<^sub>s\<^sub>s\<^sub>t_vars_mono wfvarsoccs\<^sub>s\<^sub>s\<^sub>t_def) end diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Strands_and_Constraints.thy b/thys/Stateful_Protocol_Composition_and_Typing/Strands_and_Constraints.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/Strands_and_Constraints.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/Strands_and_Constraints.thy @@ -1,2896 +1,2896 @@ (* Title: Strands_and_Constraints.thy Author: Andreas Viktor Hess, DTU SPDX-License-Identifier: BSD-3-Clause *) section \Strands and Symbolic Intruder Constraints\ theory Strands_and_Constraints imports Messages More_Unification Intruder_Deduction begin subsection \Constraints, Strands and Related Definitions\ datatype poscheckvariant = Assign ("assign") | Check ("check") text \ A strand (or constraint) step is either a message transmission (either a message being sent \Send\ or being received \Receive\) or a check on messages (a positive check \Equality\---which can be either an "assignment" or just a check---or a negative check \Inequality\) \ datatype (funs\<^sub>s\<^sub>t\<^sub>p: 'a, vars\<^sub>s\<^sub>t\<^sub>p: 'b) strand_step = Send "('a,'b) term list" ("send\_\\<^sub>s\<^sub>t" 80) | Receive "('a,'b) term list" ("receive\_\\<^sub>s\<^sub>t" 80) | Equality poscheckvariant "('a,'b) term" "('a,'b) term" ("\_: _ \ _\\<^sub>s\<^sub>t" [80,80]) | Inequality (bvars\<^sub>s\<^sub>t\<^sub>p: "'b list") "(('a,'b) term \ ('a,'b) term) list" ("\_\\\: _\\<^sub>s\<^sub>t" [80,80]) where "bvars\<^sub>s\<^sub>t\<^sub>p (Send _) = []" | "bvars\<^sub>s\<^sub>t\<^sub>p (Receive _) = []" | "bvars\<^sub>s\<^sub>t\<^sub>p (Equality _ _ _) = []" abbreviation "Send1 t \ Send [t]" abbreviation "Receive1 t \ Receive [t]" text \ A strand is a finite sequence of strand steps (constraints and strands share the same datatype) \ type_synonym ('a,'b) strand = "('a,'b) strand_step list" type_synonym ('a,'b) strands = "('a,'b) strand set" abbreviation "trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ \(t,t') \ set F. {t,t'}" fun trms\<^sub>s\<^sub>t\<^sub>p::"('a,'b) strand_step \ ('a,'b) terms" where "trms\<^sub>s\<^sub>t\<^sub>p (Send ts) = set ts" | "trms\<^sub>s\<^sub>t\<^sub>p (Receive ts) = set ts" | "trms\<^sub>s\<^sub>t\<^sub>p (Equality _ t t') = {t,t'}" | "trms\<^sub>s\<^sub>t\<^sub>p (Inequality _ F) = trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" lemma vars\<^sub>s\<^sub>t\<^sub>p_unfold[simp]: "vars\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t\<^sub>p x) \ set (bvars\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) auto text \The set of terms occurring in a strand\ definition trms\<^sub>s\<^sub>t where "trms\<^sub>s\<^sub>t S \ \(trms\<^sub>s\<^sub>t\<^sub>p ` set S)" fun trms_list\<^sub>s\<^sub>t\<^sub>p::"('a,'b) strand_step \ ('a,'b) term list" where "trms_list\<^sub>s\<^sub>t\<^sub>p (Send ts) = ts" | "trms_list\<^sub>s\<^sub>t\<^sub>p (Receive ts) = ts" | "trms_list\<^sub>s\<^sub>t\<^sub>p (Equality _ t t') = [t,t']" | "trms_list\<^sub>s\<^sub>t\<^sub>p (Inequality _ F) = concat (map (\(t,t'). [t,t']) F)" text \The set of terms occurring in a strand (list variant)\ definition trms_list\<^sub>s\<^sub>t where "trms_list\<^sub>s\<^sub>t S \ remdups (concat (map trms_list\<^sub>s\<^sub>t\<^sub>p S))" text \The set of variables occurring in a sent message\ definition fv\<^sub>s\<^sub>n\<^sub>d::"('a,'b) strand_step \ 'b set" where "fv\<^sub>s\<^sub>n\<^sub>d x \ case x of Send t \ fv\<^sub>s\<^sub>e\<^sub>t (set t) | _ \ {}" text \The set of variables occurring in a received message\ definition fv\<^sub>r\<^sub>c\<^sub>v::"('a,'b) strand_step \ 'b set" where "fv\<^sub>r\<^sub>c\<^sub>v x \ case x of Receive t \ fv\<^sub>s\<^sub>e\<^sub>t (set t) | _ \ {}" text \The set of variables occurring in an equality constraint\ definition fv\<^sub>e\<^sub>q::"poscheckvariant \ ('a,'b) strand_step \ 'b set" where "fv\<^sub>e\<^sub>q ac x \ case x of Equality ac' s t \ if ac = ac' then fv s \ fv t else {} | _ \ {}" text \The set of variables occurring at the left-hand side of an equality constraint\ definition fv_l\<^sub>e\<^sub>q::"poscheckvariant \ ('a,'b) strand_step \ 'b set" where "fv_l\<^sub>e\<^sub>q ac x \ case x of Equality ac' s t \ if ac = ac' then fv s else {} | _ \ {}" text \The set of variables occurring at the right-hand side of an equality constraint\ definition fv_r\<^sub>e\<^sub>q::"poscheckvariant \ ('a,'b) strand_step \ 'b set" where "fv_r\<^sub>e\<^sub>q ac x \ case x of Equality ac' s t \ if ac = ac' then fv t else {} | _ \ {}" text \The free variables of inequality constraints\ definition fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q::"('a,'b) strand_step \ 'b set" where "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x \ case x of Inequality X F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X | _ \ {}" fun fv\<^sub>s\<^sub>t\<^sub>p::"('a,'b) strand_step \ 'b set" where "fv\<^sub>s\<^sub>t\<^sub>p (Send t) = fv\<^sub>s\<^sub>e\<^sub>t (set t)" | "fv\<^sub>s\<^sub>t\<^sub>p (Receive t) = fv\<^sub>s\<^sub>e\<^sub>t (set t)" | "fv\<^sub>s\<^sub>t\<^sub>p (Equality _ t t') = fv t \ fv t'" | "fv\<^sub>s\<^sub>t\<^sub>p (Inequality X F) = (\(t,t') \ set F. fv t \ fv t') - set X" text \The set of free variables of a strand\ definition fv\<^sub>s\<^sub>t::"('a,'b) strand \ 'b set" where "fv\<^sub>s\<^sub>t S \ \(set (map fv\<^sub>s\<^sub>t\<^sub>p S))" text \The set of bound variables of a strand\ definition bvars\<^sub>s\<^sub>t::"('a,'b) strand \ 'b set" where "bvars\<^sub>s\<^sub>t S \ \(set (map (set \ bvars\<^sub>s\<^sub>t\<^sub>p) S))" text \The set of all variables occurring in a strand\ definition vars\<^sub>s\<^sub>t::"('a,'b) strand \ 'b set" where "vars\<^sub>s\<^sub>t S \ \(set (map vars\<^sub>s\<^sub>t\<^sub>p S))" abbreviation wfrestrictedvars\<^sub>s\<^sub>t\<^sub>p::"('a,'b) strand_step \ 'b set" where "wfrestrictedvars\<^sub>s\<^sub>t\<^sub>p x \ case x of Inequality _ _ \ {} | Equality Check _ _ \ {} | _ \ vars\<^sub>s\<^sub>t\<^sub>p x" text \The variables of a strand whose occurrences might be restricted by well-formedness constraints\ definition wfrestrictedvars\<^sub>s\<^sub>t::"('a,'b) strand \ 'b set" where "wfrestrictedvars\<^sub>s\<^sub>t S \ \(set (map wfrestrictedvars\<^sub>s\<^sub>t\<^sub>p S))" abbreviation wfvarsoccs\<^sub>s\<^sub>t\<^sub>p where "wfvarsoccs\<^sub>s\<^sub>t\<^sub>p x \ case x of Send t \ fv\<^sub>s\<^sub>e\<^sub>t (set t) | Equality Assign s t \ fv s | _ \ {}" text \The variables of a strand that occur in sent messages or in assignments\ definition wfvarsoccs\<^sub>s\<^sub>t where "wfvarsoccs\<^sub>s\<^sub>t S \ \(set (map wfvarsoccs\<^sub>s\<^sub>t\<^sub>p S))" text \The variables occurring at the right-hand side of assignment steps\ fun assignment_rhs\<^sub>s\<^sub>t where "assignment_rhs\<^sub>s\<^sub>t [] = {}" | "assignment_rhs\<^sub>s\<^sub>t (Equality Assign t t'#S) = insert t' (assignment_rhs\<^sub>s\<^sub>t S)" | "assignment_rhs\<^sub>s\<^sub>t (x#S) = assignment_rhs\<^sub>s\<^sub>t S" text \The set of function symbols occurring in a strand\ definition funs\<^sub>s\<^sub>t::"('a,'b) strand \ 'a set" where "funs\<^sub>s\<^sub>t S \ \(set (map funs\<^sub>s\<^sub>t\<^sub>p S))" fun subst_apply_strand_step::"('a,'b) strand_step \ ('a,'b) subst \ ('a,'b) strand_step" (infix "\\<^sub>s\<^sub>t\<^sub>p" 51) where "Send t \\<^sub>s\<^sub>t\<^sub>p \ = Send (t \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" | "Receive t \\<^sub>s\<^sub>t\<^sub>p \ = Receive (t \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" | "Equality a t t' \\<^sub>s\<^sub>t\<^sub>p \ = Equality a (t \ \) (t' \ \)" | "Inequality X F \\<^sub>s\<^sub>t\<^sub>p \ = Inequality X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" text \Substitution application for strands\ definition subst_apply_strand::"('a,'b) strand \ ('a,'b) subst \ ('a,'b) strand" (infix "\\<^sub>s\<^sub>t" 51) where "S \\<^sub>s\<^sub>t \ \ map (\x. x \\<^sub>s\<^sub>t\<^sub>p \) S" text \The semantics of inequality constraints\ definition "ineq_model (\::('a,'b) subst) X F \ (\\. subst_domain \ = set X \ ground (subst_range \) \ (\(t,t') \ set F. t \ \ \\<^sub>s \ \ t' \ \ \\<^sub>s \))" fun simple\<^sub>s\<^sub>t\<^sub>p where "simple\<^sub>s\<^sub>t\<^sub>p (Receive t) = True" | "simple\<^sub>s\<^sub>t\<^sub>p (Send [Var v]) = True" | "simple\<^sub>s\<^sub>t\<^sub>p (Inequality X F) = (\\. ineq_model \ X F)" | "simple\<^sub>s\<^sub>t\<^sub>p _ = False" text \Simple constraints\ definition simple where "simple S \ list_all simple\<^sub>s\<^sub>t\<^sub>p S" text \The intruder knowledge of a constraint\ fun ik\<^sub>s\<^sub>t::"('a,'b) strand \ ('a,'b) terms" where "ik\<^sub>s\<^sub>t [] = {}" | "ik\<^sub>s\<^sub>t (Receive t#S) = set t \ (ik\<^sub>s\<^sub>t S)" | "ik\<^sub>s\<^sub>t (_#S) = ik\<^sub>s\<^sub>t S" text \Strand well-formedness\ fun wf\<^sub>s\<^sub>t::"'b set \ ('a,'b) strand \ bool" where "wf\<^sub>s\<^sub>t V [] = True" | "wf\<^sub>s\<^sub>t V (Receive ts#S) = (fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ V \ wf\<^sub>s\<^sub>t V S)" | "wf\<^sub>s\<^sub>t V (Send ts#S) = wf\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) S" | "wf\<^sub>s\<^sub>t V (Equality Assign s t#S) = (fv t \ V \ wf\<^sub>s\<^sub>t (V \ fv s) S)" | "wf\<^sub>s\<^sub>t V (Equality Check s t#S) = wf\<^sub>s\<^sub>t V S" | "wf\<^sub>s\<^sub>t V (Inequality _ _#S) = wf\<^sub>s\<^sub>t V S" text \Well-formedness of constraint states\ definition wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r::"('a,'b) strand \ ('a,'b) subst \ bool" where "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \ \ (wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>s\<^sub>t {} S \ subst_domain \ \ vars\<^sub>s\<^sub>t S = {} \ range_vars \ \ bvars\<^sub>s\<^sub>t S = {} \ fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {})" declare trms\<^sub>s\<^sub>t_def[simp] declare fv\<^sub>s\<^sub>n\<^sub>d_def[simp] declare fv\<^sub>r\<^sub>c\<^sub>v_def[simp] declare fv\<^sub>e\<^sub>q_def[simp] declare fv_l\<^sub>e\<^sub>q_def[simp] declare fv_r\<^sub>e\<^sub>q_def[simp] declare fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q_def[simp] declare fv\<^sub>s\<^sub>t_def[simp] declare vars\<^sub>s\<^sub>t_def[simp] declare bvars\<^sub>s\<^sub>t_def[simp] declare wfrestrictedvars\<^sub>s\<^sub>t_def[simp] declare wfvarsoccs\<^sub>s\<^sub>t_def[simp] lemmas wf\<^sub>s\<^sub>t_induct = wf\<^sub>s\<^sub>t.induct[case_names Nil ConsRcv ConsSnd ConsEq ConsEq2 ConsIneq] lemmas ik\<^sub>s\<^sub>t_induct = ik\<^sub>s\<^sub>t.induct[case_names Nil ConsRcv ConsSnd ConsEq ConsIneq] lemmas assignment_rhs\<^sub>s\<^sub>t_induct = assignment_rhs\<^sub>s\<^sub>t.induct[case_names Nil ConsEq2 ConsSnd ConsRcv ConsEq ConsIneq] subsubsection \Lexicographical measure on strands\ definition size\<^sub>s\<^sub>t::"('a,'b) strand \ nat" where "size\<^sub>s\<^sub>t S \ size_list (\x. Max (insert 0 (size ` trms\<^sub>s\<^sub>t\<^sub>p x))) S" definition measure\<^sub>s\<^sub>t::"((('a, 'b) strand \ ('a,'b) subst) \ ('a, 'b) strand \ ('a,'b) subst) set" where "measure\<^sub>s\<^sub>t \ measures [\(S,\). card (fv\<^sub>s\<^sub>t S), \(S,\). size\<^sub>s\<^sub>t S]" lemma measure\<^sub>s\<^sub>t_alt_def: "((s,x),(t,y)) \ measure\<^sub>s\<^sub>t = (card (fv\<^sub>s\<^sub>t s) < card (fv\<^sub>s\<^sub>t t) \ (card (fv\<^sub>s\<^sub>t s) = card (fv\<^sub>s\<^sub>t t) \ size\<^sub>s\<^sub>t s < size\<^sub>s\<^sub>t t))" by (simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) lemma measure\<^sub>s\<^sub>t_trans: "trans measure\<^sub>s\<^sub>t" by (simp add: trans_def measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) subsubsection \Some lemmata\ lemma trms_list\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>t: "trms\<^sub>s\<^sub>t S = set (trms_list\<^sub>s\<^sub>t S)" unfolding trms\<^sub>s\<^sub>t_def trms_list\<^sub>s\<^sub>t_def proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp lemma subst_apply_strand_step_def: "s \\<^sub>s\<^sub>t\<^sub>p \ = (case s of Send t \ Send (t \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) | Receive t \ Receive (t \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) | Equality a t t' \ Equality a (t \ \) (t' \ \) | Inequality X F \ Inequality X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \))" by (cases s) simp_all lemma subst_apply_strand_nil[simp]: "[] \\<^sub>s\<^sub>t \ = []" unfolding subst_apply_strand_def by simp lemma finite_funs\<^sub>s\<^sub>t\<^sub>p[simp]: "finite (funs\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) auto lemma finite_funs\<^sub>s\<^sub>t[simp]: "finite (funs\<^sub>s\<^sub>t S)" unfolding funs\<^sub>s\<^sub>t_def by simp lemma finite_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s[simp]: "finite (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s x)" by (induct x) auto lemma finite_trms\<^sub>s\<^sub>t\<^sub>p[simp]: "finite (trms\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) auto lemma finite_vars\<^sub>s\<^sub>t\<^sub>p[simp]: "finite (vars\<^sub>s\<^sub>t\<^sub>p x)" by auto lemma finite_bvars\<^sub>s\<^sub>t\<^sub>p[simp]: "finite (set (bvars\<^sub>s\<^sub>t\<^sub>p x))" by rule lemma finite_fv\<^sub>s\<^sub>n\<^sub>d[simp]: "finite (fv\<^sub>s\<^sub>n\<^sub>d x)" by (cases x) auto lemma finite_fv\<^sub>r\<^sub>c\<^sub>v[simp]: "finite (fv\<^sub>r\<^sub>c\<^sub>v x)" by (cases x) auto lemma finite_fv\<^sub>s\<^sub>t\<^sub>p[simp]: "finite (fv\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) auto lemma finite_vars\<^sub>s\<^sub>t[simp]: "finite (vars\<^sub>s\<^sub>t S)" by simp lemma finite_bvars\<^sub>s\<^sub>t[simp]: "finite (bvars\<^sub>s\<^sub>t S)" by simp lemma finite_fv\<^sub>s\<^sub>t[simp]: "finite (fv\<^sub>s\<^sub>t S)" by simp lemma finite_wfrestrictedvars\<^sub>s\<^sub>t\<^sub>p[simp]: "finite (wfrestrictedvars\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) (auto split: poscheckvariant.splits) lemma finite_wfrestrictedvars\<^sub>s\<^sub>t[simp]: "finite (wfrestrictedvars\<^sub>s\<^sub>t S)" using finite_wfrestrictedvars\<^sub>s\<^sub>t\<^sub>p by auto lemma finite_wfvarsoccs\<^sub>s\<^sub>t\<^sub>p[simp]: "finite (wfvarsoccs\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) (auto split: poscheckvariant.splits) lemma finite_wfvarsoccs\<^sub>s\<^sub>t[simp]: "finite (wfvarsoccs\<^sub>s\<^sub>t S)" using finite_wfvarsoccs\<^sub>s\<^sub>t\<^sub>p by auto lemma finite_ik\<^sub>s\<^sub>t[simp]: "finite (ik\<^sub>s\<^sub>t S)" by (induct S rule: ik\<^sub>s\<^sub>t.induct) simp_all lemma finite_assignment_rhs\<^sub>s\<^sub>t[simp]: "finite (assignment_rhs\<^sub>s\<^sub>t S)" by (induct S rule: assignment_rhs\<^sub>s\<^sub>t.induct) simp_all lemma ik\<^sub>s\<^sub>t_is_rcv_set: "ik\<^sub>s\<^sub>t A = {t | ts t. Receive ts \ set A \ t \ set ts}" by (induct A rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik\<^sub>s\<^sub>t_snoc_no_receive_eq: assumes "\ts. a = receive\ts\\<^sub>s\<^sub>t" shows "ik\<^sub>s\<^sub>t (A@[a]) \\<^sub>s\<^sub>e\<^sub>t \ = ik\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" using assms unfolding ik\<^sub>s\<^sub>t_is_rcv_set by (metis (no_types, lifting) Un_iff append_Nil2 set_ConsD set_append) lemma ik\<^sub>s\<^sub>tD[dest]: "t \ ik\<^sub>s\<^sub>t S \ \ts. t \ set ts \ Receive ts \ set S" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik\<^sub>s\<^sub>tD'[dest]: "t \ ik\<^sub>s\<^sub>t S \ t \ trms\<^sub>s\<^sub>t S" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik\<^sub>s\<^sub>tD''[dest]: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t S) \ t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik\<^sub>s\<^sub>t_subterm_exD: assumes "t \ ik\<^sub>s\<^sub>t S" shows "\x \ set S. t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t\<^sub>p x)" using assms ik\<^sub>s\<^sub>tD by force lemma assignment_rhs\<^sub>s\<^sub>tD[dest]: "t \ assignment_rhs\<^sub>s\<^sub>t S \ \t'. Equality Assign t' t \ set S" by (induct S rule: assignment_rhs\<^sub>s\<^sub>t.induct) auto lemma assignment_rhs\<^sub>s\<^sub>tD'[dest]: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (assignment_rhs\<^sub>s\<^sub>t S) \ t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" by (induct S rule: assignment_rhs\<^sub>s\<^sub>t.induct) auto lemma bvars\<^sub>s\<^sub>t_split: "bvars\<^sub>s\<^sub>t (S@S') = bvars\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S'" unfolding bvars\<^sub>s\<^sub>t_def by auto lemma bvars\<^sub>s\<^sub>t_singleton: "bvars\<^sub>s\<^sub>t [x] = set (bvars\<^sub>s\<^sub>t\<^sub>p x)" unfolding bvars\<^sub>s\<^sub>t_def by auto lemma strand_fv_bvars_disjointD: assumes "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" "Inequality X F \ set S" shows "set X \ bvars\<^sub>s\<^sub>t S" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ fv\<^sub>s\<^sub>t S" using assms by (induct S) fastforce+ lemma strand_fv_bvars_disjoint_unfold: assumes "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" "Inequality X F \ set S" "Inequality Y G \ set S" shows "set Y \ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X) = {}" proof - have "set X \ bvars\<^sub>s\<^sub>t S" "set Y \ bvars\<^sub>s\<^sub>t S" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ fv\<^sub>s\<^sub>t S" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G - set Y \ fv\<^sub>s\<^sub>t S" using strand_fv_bvars_disjointD[OF assms(1)] assms(2,3) by auto thus ?thesis using assms(1) by fastforce qed lemma strand_subst_hom[iff]: "(S@S') \\<^sub>s\<^sub>t \ = (S \\<^sub>s\<^sub>t \)@(S' \\<^sub>s\<^sub>t \)" "(x#S) \\<^sub>s\<^sub>t \ = (x \\<^sub>s\<^sub>t\<^sub>p \)#(S \\<^sub>s\<^sub>t \)" unfolding subst_apply_strand_def by auto lemma strand_subst_comp: "range_vars \ \ bvars\<^sub>s\<^sub>t S = {} \ S \\<^sub>s\<^sub>t \ \\<^sub>s \ = ((S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \)" proof (induction S) case (Cons x S) have *: "range_vars \ \ bvars\<^sub>s\<^sub>t S = {}" "range_vars \ \ (set (bvars\<^sub>s\<^sub>t\<^sub>p x)) = {}" using Cons bvars\<^sub>s\<^sub>t_split[of "[x]" S] append_Cons inf_sup_absorb by (metis (no_types, lifting) Int_iff Un_commute disjoint_iff_not_equal self_append_conv2, metis append_self_conv2 bvars\<^sub>s\<^sub>t_singleton inf_bot_right inf_left_commute) hence IH: "S \\<^sub>s\<^sub>t \ \\<^sub>s \ = (S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \" using Cons.IH by auto have "(x#S \\<^sub>s\<^sub>t \ \\<^sub>s \) = (x \\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \)#(S \\<^sub>s\<^sub>t \ \\<^sub>s \)" by (metis strand_subst_hom(2)) hence "... = (x \\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \)#((S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \)" by (metis IH) hence "... = ((x \\<^sub>s\<^sub>t\<^sub>p \) \\<^sub>s\<^sub>t\<^sub>p \)#((S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \)" using rm_vars_comp[OF *(2)] proof (induction x) case (Inequality X F) thus ?case by (induct F) (auto simp add: subst_apply_pairs_def subst_apply_strand_step_def) qed (simp_all add: subst_apply_strand_step_def) thus ?case using IH by auto qed (simp add: subst_apply_strand_def) lemma strand_substI[intro]: "subst_domain \ \ fv\<^sub>s\<^sub>t S = {} \ S \\<^sub>s\<^sub>t \ = S" "subst_domain \ \ vars\<^sub>s\<^sub>t S = {} \ S \\<^sub>s\<^sub>t \ = S" proof - show "subst_domain \ \ vars\<^sub>s\<^sub>t S = {} \ S \\<^sub>s\<^sub>t \ = S" proof (induction S) case (Cons x S) hence "S \\<^sub>s\<^sub>t \ = S" by auto moreover have "vars\<^sub>s\<^sub>t\<^sub>p x \ subst_domain \ = {}" using Cons.prems by auto hence "x \\<^sub>s\<^sub>t\<^sub>p \ = x" proof (induction x) case (Send ts) thus ?case by (induct ts) auto next case (Receive ts) thus ?case by (induct ts) auto next case (Inequality X F) thus ?case by (induct F) (force simp add: subst_apply_pairs_def)+ qed auto ultimately show ?case by simp qed (simp add: subst_apply_strand_def) show "subst_domain \ \ fv\<^sub>s\<^sub>t S = {} \ S \\<^sub>s\<^sub>t \ = S" proof (induction S) case (Cons x S) hence "S \\<^sub>s\<^sub>t \ = S" by auto moreover have "fv\<^sub>s\<^sub>t\<^sub>p x \ subst_domain \ = {}" using Cons.prems by auto hence "x \\<^sub>s\<^sub>t\<^sub>p \ = x" proof (induction x) case (Send ts) thus ?case by (induct ts) auto next case (Receive ts) thus ?case by (induct ts) auto next case (Inequality X F) thus ?case by (induct F) (force simp add: subst_apply_pairs_def)+ qed auto ultimately show ?case by simp qed (simp add: subst_apply_strand_def) qed lemma strand_substI': "fv\<^sub>s\<^sub>t S = {} \ S \\<^sub>s\<^sub>t \ = S" "vars\<^sub>s\<^sub>t S = {} \ S \\<^sub>s\<^sub>t \ = S" by (metis inf_bot_right strand_substI(1), metis inf_bot_right strand_substI(2)) lemma strand_subst_set: "(set (S \\<^sub>s\<^sub>t \)) = ((\x. x \\<^sub>s\<^sub>t\<^sub>p \) ` (set S))" by (auto simp add: subst_apply_strand_def) lemma strand_map_inv_set_snd_rcv_subst: assumes "finite (M::('a,'b) terms)" shows "set ((map Send1 (inv set M)) \\<^sub>s\<^sub>t \) = Send1 ` (M \\<^sub>s\<^sub>e\<^sub>t \)" (is ?A) "set ((map Receive1 (inv set M)) \\<^sub>s\<^sub>t \) = Receive1 ` (M \\<^sub>s\<^sub>e\<^sub>t \)" (is ?B) proof - { fix f::"('a,'b) term \ ('a,'b) strand_step" assume f: "f = Send1 \ f = Receive1" from assms have "set ((map f (inv set M)) \\<^sub>s\<^sub>t \) = f ` (M \\<^sub>s\<^sub>e\<^sub>t \)" proof (induction rule: finite_induct) case empty thus ?case unfolding inv_def by auto next case (insert m M) have "set (map f (inv set (insert m M)) \\<^sub>s\<^sub>t \) = insert (f m \\<^sub>s\<^sub>t\<^sub>p \) (set (map f (inv set M) \\<^sub>s\<^sub>t \))" by (simp add: insert.hyps(1) inv_set_fset subst_apply_strand_def) thus ?case using f insert.IH by auto qed } thus "?A" "?B" by auto qed lemma strand_ground_subst_vars_subset: assumes "ground (subst_range \)" shows "vars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ vars\<^sub>s\<^sub>t S" proof (induction S) case (Cons x S) have "vars\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ vars\<^sub>s\<^sub>t\<^sub>p x" using ground_subst_fv_subset[OF assms] proof (cases x) case (Inequality X F) let ?\ = "rm_vars (set X) \" have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" proof (induction F) case (Cons f F) obtain t t' where f: "f = (t,t')" by (metis surj_pair) hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (f#F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\) = fv (t \ ?\) \ fv (t' \ ?\) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (f#F) = fv t \ fv t' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" by (auto simp add: subst_apply_pairs_def) thus ?case using ground_subst_fv_subset[OF ground_subset[OF rm_vars_img_subset assms, of "set X"]] Cons.IH by (metis (no_types, lifting) Un_mono) qed (simp add: subst_apply_pairs_def) moreover have "vars\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ set X" "vars\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ set X" using Inequality by (auto simp add: subst_apply_pairs_def) ultimately show ?thesis by auto qed auto thus ?case using Cons.IH by auto qed (simp add: subst_apply_strand_def) lemma ik_union_subset: "\(P ` ik\<^sub>s\<^sub>t S) \ (\x \ (set S). \(P ` trms\<^sub>s\<^sub>t\<^sub>p x))" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik_snd_empty[simp]: "ik\<^sub>s\<^sub>t (map Send X) = {}" by (induct "map Send X" arbitrary: X rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik_snd_empty'[simp]: "ik\<^sub>s\<^sub>t [Send t] = {}" by simp lemma ik_append[iff]: "ik\<^sub>s\<^sub>t (S@S') = ik\<^sub>s\<^sub>t S \ ik\<^sub>s\<^sub>t S'" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik_cons: "ik\<^sub>s\<^sub>t (x#S) = ik\<^sub>s\<^sub>t [x] \ ik\<^sub>s\<^sub>t S" using ik_append[of "[x]" S] by simp lemma assignment_rhs_append[iff]: "assignment_rhs\<^sub>s\<^sub>t (S@S') = assignment_rhs\<^sub>s\<^sub>t S \ assignment_rhs\<^sub>s\<^sub>t S'" by (induct S rule: assignment_rhs\<^sub>s\<^sub>t.induct) auto lemma eqs_rcv_map_empty: "assignment_rhs\<^sub>s\<^sub>t (map Receive M) = {}" by auto lemma ik_rcv_map: assumes "ts \ set L" shows "set ts \ ik\<^sub>s\<^sub>t (map Receive L)" proof - { fix L L' have "set ts \ ik\<^sub>s\<^sub>t [Receive ts]" by auto hence "set ts \ ik\<^sub>s\<^sub>t (map Receive L@Receive ts#map Receive L')" using ik_append by auto hence "set ts \ ik\<^sub>s\<^sub>t (map Receive (L@ts#L'))" by auto } thus ?thesis using assms split_list_last by force qed lemma ik_subst: "ik\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) = ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \" by (induct rule: ik\<^sub>s\<^sub>t_induct) auto lemma ik_rcv_map': assumes "t \ ik\<^sub>s\<^sub>t (map Receive L)" shows "\ts \ set L. t \ set ts" using assms by force lemma ik_append_subset[simp]: "ik\<^sub>s\<^sub>t S \ ik\<^sub>s\<^sub>t (S@S')" "ik\<^sub>s\<^sub>t S' \ ik\<^sub>s\<^sub>t (S@S')" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma assignment_rhs_append_subset[simp]: "assignment_rhs\<^sub>s\<^sub>t S \ assignment_rhs\<^sub>s\<^sub>t (S@S')" "assignment_rhs\<^sub>s\<^sub>t S' \ assignment_rhs\<^sub>s\<^sub>t (S@S')" by (induct S rule: assignment_rhs\<^sub>s\<^sub>t.induct) auto lemma trms\<^sub>s\<^sub>t_cons: "trms\<^sub>s\<^sub>t (x#S) = trms\<^sub>s\<^sub>t\<^sub>p x \ trms\<^sub>s\<^sub>t S" by simp lemma trm_strand_subst_cong: "t \ trms\<^sub>s\<^sub>t S \ t \ \ \ trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ (\X F. Inequality X F \ set S \ t \ rm_vars (set X) \ \ trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \))" (is "t \ trms\<^sub>s\<^sub>t S \ ?P t \ S") "t \ trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ (\t'. t = t' \ \ \ t' \ trms\<^sub>s\<^sub>t S) \ (\X F. Inequality X F \ set S \ (\t' \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F. t = t' \ rm_vars (set X) \))" (is "t \ trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ ?Q t \ S") proof - show "t \ trms\<^sub>s\<^sub>t S \ ?P t \ S" proof (induction S) case (Cons x S) show ?case proof (cases "t \ trms\<^sub>s\<^sub>t S") case True hence "?P t \ S" using Cons by simp thus ?thesis by (cases x) (metis (no_types, lifting) Un_iff list.set_intros(2) strand_subst_hom(2) trms\<^sub>s\<^sub>t_cons)+ next case False hence "t \ trms\<^sub>s\<^sub>t\<^sub>p x" using Cons.prems by auto thus ?thesis proof (induction x) case (Inequality X F) hence "t \ rm_vars (set X) \ \ trms\<^sub>s\<^sub>t\<^sub>p (Inequality X F \\<^sub>s\<^sub>t\<^sub>p \)" by (induct F) (auto simp add: subst_apply_pairs_def subst_apply_strand_step_def) thus ?case by fastforce qed (auto simp add: subst_apply_strand_step_def) qed qed simp show "t \ trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ ?Q t \ S" proof (induction S) case (Cons x S) show ?case proof (cases "t \ trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)") case True hence "?Q t \ S" using Cons by simp thus ?thesis by (cases x) force+ next case False hence "t \ trms\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \)" using Cons.prems by auto thus ?thesis proof (induction x) case (Inequality X F) hence "t \ trms\<^sub>s\<^sub>t\<^sub>p (Inequality X F) \\<^sub>s\<^sub>e\<^sub>t rm_vars (set X) \" by (induct F) (force simp add: subst_apply_pairs_def)+ thus ?case by fastforce qed (auto simp add: subst_apply_strand_step_def) qed qed simp qed subsection \Lemmata: Free Variables of Strands\ lemma fv_trm_snd_rcv[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t\<^sub>p (Send ts)) = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t\<^sub>p (Receive ts)) = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" by simp_all lemma in_strand_fv_subset: "x \ set S \ vars\<^sub>s\<^sub>t\<^sub>p x \ vars\<^sub>s\<^sub>t S" by fastforce lemma in_strand_fv_subset_snd: "Send ts \ set S \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ \(set (map fv\<^sub>s\<^sub>n\<^sub>d S))" by fastforce lemma in_strand_fv_subset_rcv: "Receive ts \ set S \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ \(set (map fv\<^sub>r\<^sub>c\<^sub>v S))" by fastforce lemma fv\<^sub>s\<^sub>n\<^sub>dE: assumes "v \ \(set (map fv\<^sub>s\<^sub>n\<^sub>d S))" obtains ts where "send\ts\\<^sub>s\<^sub>t \ set S" "v \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" proof - have "\ts. send\ts\\<^sub>s\<^sub>t \ set S \ v \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" by (metis (no_types, lifting) assms UN_E empty_iff set_map strand_step.case_eq_if fv\<^sub>s\<^sub>n\<^sub>d_def strand_step.collapse(1)) thus ?thesis by (metis that) qed lemma fv\<^sub>r\<^sub>c\<^sub>vE: assumes "v \ \(set (map fv\<^sub>r\<^sub>c\<^sub>v S))" obtains ts where "receive\ts\\<^sub>s\<^sub>t \ set S" "v \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" proof - have "\ts. receive\ts\\<^sub>s\<^sub>t \ set S \ v \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" by (metis (no_types, lifting) assms UN_E empty_iff set_map strand_step.case_eq_if fv\<^sub>r\<^sub>c\<^sub>v_def strand_step.collapse(2)) thus ?thesis by (metis that) qed lemma vars\<^sub>s\<^sub>t\<^sub>pI[intro]: "x \ fv\<^sub>s\<^sub>t\<^sub>p s \ x \ vars\<^sub>s\<^sub>t\<^sub>p s" by (induct s rule: fv\<^sub>s\<^sub>t\<^sub>p.induct) auto lemma vars\<^sub>s\<^sub>tI[intro]: "x \ fv\<^sub>s\<^sub>t S \ x \ vars\<^sub>s\<^sub>t S" using vars\<^sub>s\<^sub>t\<^sub>pI by fastforce lemma fv\<^sub>s\<^sub>t_subset_vars\<^sub>s\<^sub>t[simp]: "fv\<^sub>s\<^sub>t S \ vars\<^sub>s\<^sub>t S" using vars\<^sub>s\<^sub>tI by force lemma vars\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>t: "vars\<^sub>s\<^sub>t S = fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S" proof (induction S) case (Cons x S) thus ?case proof (induction x) case (Inequality X F) thus ?case by (induct F) auto qed auto qed simp lemma fv\<^sub>s\<^sub>t\<^sub>p_is_subterm_trms\<^sub>s\<^sub>t\<^sub>p: "x \ fv\<^sub>s\<^sub>t\<^sub>p a \ Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t\<^sub>p a)" using var_is_subterm by (cases a) force+ lemma fv\<^sub>s\<^sub>t_is_subterm_trms\<^sub>s\<^sub>t: "x \ fv\<^sub>s\<^sub>t A \ Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t A)" proof (induction A) case (Cons a A) thus ?case using fv\<^sub>s\<^sub>t\<^sub>p_is_subterm_trms\<^sub>s\<^sub>t\<^sub>p by (cases "x \ fv\<^sub>s\<^sub>t A") auto qed simp lemma vars_st_snd_map: "vars\<^sub>s\<^sub>t (map Send tss) = fv\<^sub>s\<^sub>e\<^sub>t (Fun f ` set tss)" by auto lemma vars_st_rcv_map: "vars\<^sub>s\<^sub>t (map Receive tss) = fv\<^sub>s\<^sub>e\<^sub>t (Fun f ` set tss)" by auto lemma vars_snd_rcv_union: "vars\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>s\<^sub>n\<^sub>d x \ fv\<^sub>r\<^sub>c\<^sub>v x \ fv\<^sub>e\<^sub>q assign x \ fv\<^sub>e\<^sub>q check x \ fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x \ set (bvars\<^sub>s\<^sub>t\<^sub>p x)" proof (cases x) case (Equality ac t t') thus ?thesis by (cases ac) auto qed auto lemma fv_snd_rcv_union: "fv\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>s\<^sub>n\<^sub>d x \ fv\<^sub>r\<^sub>c\<^sub>v x \ fv\<^sub>e\<^sub>q assign x \ fv\<^sub>e\<^sub>q check x \ fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x" proof (cases x) case (Equality ac t t') thus ?thesis by (cases ac) auto qed auto lemma fv_snd_rcv_empty[simp]: "fv\<^sub>s\<^sub>n\<^sub>d x = {} \ fv\<^sub>r\<^sub>c\<^sub>v x = {}" by (cases x) simp_all lemma vars_snd_rcv_strand[iff]: "vars\<^sub>s\<^sub>t (S::('a,'b) strand) = (\(set (map fv\<^sub>s\<^sub>n\<^sub>d S))) \ (\(set (map fv\<^sub>r\<^sub>c\<^sub>v S))) \ (\(set (map (fv\<^sub>e\<^sub>q assign) S))) \ (\(set (map (fv\<^sub>e\<^sub>q check) S))) \ (\(set (map fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q S))) \ bvars\<^sub>s\<^sub>t S" unfolding bvars\<^sub>s\<^sub>t_def proof (induction S) case (Cons x S) have "\s V. vars\<^sub>s\<^sub>t\<^sub>p (s::('a,'b) strand_step) \ V = fv\<^sub>s\<^sub>n\<^sub>d s \ fv\<^sub>r\<^sub>c\<^sub>v s \ fv\<^sub>e\<^sub>q assign s \ fv\<^sub>e\<^sub>q check s \ fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q s \ set (bvars\<^sub>s\<^sub>t\<^sub>p s) \ V" by (metis vars_snd_rcv_union) thus ?case using Cons.IH by (auto simp add: sup_assoc sup_left_commute) qed simp lemma fv_snd_rcv_strand[iff]: "fv\<^sub>s\<^sub>t (S::('a,'b) strand) = (\(set (map fv\<^sub>s\<^sub>n\<^sub>d S))) \ (\(set (map fv\<^sub>r\<^sub>c\<^sub>v S))) \ (\(set (map (fv\<^sub>e\<^sub>q assign) S))) \ (\(set (map (fv\<^sub>e\<^sub>q check) S))) \ (\(set (map fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q S)))" unfolding bvars\<^sub>s\<^sub>t_def proof (induction S) case (Cons x S) have "\s V. fv\<^sub>s\<^sub>t\<^sub>p (s::('a,'b) strand_step) \ V = fv\<^sub>s\<^sub>n\<^sub>d s \ fv\<^sub>r\<^sub>c\<^sub>v s \ fv\<^sub>e\<^sub>q assign s \ fv\<^sub>e\<^sub>q check s \ fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q s \ V" by (metis fv_snd_rcv_union) thus ?case using Cons.IH by (auto simp add: sup_assoc sup_left_commute) qed simp lemma vars_snd_rcv_strand2[iff]: "wfrestrictedvars\<^sub>s\<^sub>t (S::('a,'b) strand) = (\(set (map fv\<^sub>s\<^sub>n\<^sub>d S))) \ (\(set (map fv\<^sub>r\<^sub>c\<^sub>v S))) \ (\(set (map (fv\<^sub>e\<^sub>q assign) S)))" by (induct S) (auto simp add: split: strand_step.split poscheckvariant.split) lemma fv_snd_rcv_strand_subset[simp]: "\(set (map fv\<^sub>s\<^sub>n\<^sub>d S)) \ fv\<^sub>s\<^sub>t S" "\(set (map fv\<^sub>r\<^sub>c\<^sub>v S)) \ fv\<^sub>s\<^sub>t S" "\(set (map (fv\<^sub>e\<^sub>q ac) S)) \ fv\<^sub>s\<^sub>t S" "\(set (map fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q S)) \ fv\<^sub>s\<^sub>t S" "wfvarsoccs\<^sub>s\<^sub>t S \ fv\<^sub>s\<^sub>t S" proof - show "\(set (map fv\<^sub>s\<^sub>n\<^sub>d S)) \ fv\<^sub>s\<^sub>t S" "\(set (map fv\<^sub>r\<^sub>c\<^sub>v S)) \ fv\<^sub>s\<^sub>t S" "\(set (map fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q S)) \ fv\<^sub>s\<^sub>t S" using fv_snd_rcv_strand[of S] by auto show "\(set (map (fv\<^sub>e\<^sub>q ac) S)) \ fv\<^sub>s\<^sub>t S" by (induct S) (auto split: strand_step.split poscheckvariant.split) show "wfvarsoccs\<^sub>s\<^sub>t S \ fv\<^sub>s\<^sub>t S" by (induct S) (auto split: strand_step.split poscheckvariant.split) qed lemma vars_snd_rcv_strand_subset2[simp]: "\(set (map fv\<^sub>s\<^sub>n\<^sub>d S)) \ wfrestrictedvars\<^sub>s\<^sub>t S" "\(set (map fv\<^sub>r\<^sub>c\<^sub>v S)) \ wfrestrictedvars\<^sub>s\<^sub>t S" "\(set (map (fv\<^sub>e\<^sub>q assign) S)) \ wfrestrictedvars\<^sub>s\<^sub>t S" "wfvarsoccs\<^sub>s\<^sub>t S \ wfrestrictedvars\<^sub>s\<^sub>t S" by (induction S) (auto split: strand_step.split poscheckvariant.split) lemma wfrestrictedvars\<^sub>s\<^sub>t_subset_vars\<^sub>s\<^sub>t: "wfrestrictedvars\<^sub>s\<^sub>t S \ vars\<^sub>s\<^sub>t S" by (induction S) (auto split: strand_step.split poscheckvariant.split) lemma subst_sends_strand_step_fv_to_img: "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>s\<^sub>t\<^sub>p x \ range_vars \" using subst_sends_fv_to_img[of _ \] proof (cases x) case (Inequality X F) let ?\ = "rm_vars (set X) \" have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ range_vars ?\" proof (induction F) case (Cons f F) thus ?case using subst_sends_fv_to_img[of _ ?\] by (auto simp add: subst_apply_pairs_def) qed (auto simp add: subst_apply_pairs_def) hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ range_vars \" using rm_vars_img_subset[of "set X" \] fv_set_mono unfolding range_vars_alt_def by blast+ thus ?thesis using Inequality by (auto simp add: subst_apply_strand_step_def) qed (auto simp add: subst_apply_strand_step_def) lemma subst_sends_strand_fv_to_img: "fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t S \ range_vars \" proof (induction S) case (Cons x S) have *: "fv\<^sub>s\<^sub>t (x#S \\<^sub>s\<^sub>t \) = fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" "fv\<^sub>s\<^sub>t (x#S) \ range_vars \ = fv\<^sub>s\<^sub>t\<^sub>p x \ fv\<^sub>s\<^sub>t S \ range_vars \" by auto thus ?case using Cons.IH subst_sends_strand_step_fv_to_img[of x \] by auto qed simp lemma ineq_apply_subst: assumes "subst_domain \ \ set X = {}" shows "(Inequality X F) \\<^sub>s\<^sub>t\<^sub>p \ = Inequality X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using rm_vars_apply'[OF assms] by (simp add: subst_apply_strand_step_def) lemma fv_strand_step_subst: assumes "P = fv\<^sub>s\<^sub>t\<^sub>p \ P = fv\<^sub>r\<^sub>c\<^sub>v \ P = fv\<^sub>s\<^sub>n\<^sub>d \ P = fv\<^sub>e\<^sub>q ac \ P = fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q" and "set (bvars\<^sub>s\<^sub>t\<^sub>p x) \ (subst_domain \ \ range_vars \) = {}" shows "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (P x)) = P (x \\<^sub>s\<^sub>t\<^sub>p \)" proof (cases x) case (Send ts) hence "vars\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "fv\<^sub>s\<^sub>n\<^sub>d x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" by auto thus ?thesis using assms Send subst_apply_fv_unfold[of _ \] by fastforce next case (Receive ts) hence "vars\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "fv\<^sub>r\<^sub>c\<^sub>v x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" by auto thus ?thesis using assms Receive subst_apply_fv_unfold[of _ \] by fastforce next case (Equality ac' t t') show ?thesis proof (cases "ac = ac'") case True hence "vars\<^sub>s\<^sub>t\<^sub>p x = fv t \ fv t'" "fv\<^sub>e\<^sub>q ac x = fv t \ fv t'" using Equality by auto thus ?thesis using assms Equality subst_apply_fv_unfold[of _ \] True by auto next case False hence "vars\<^sub>s\<^sub>t\<^sub>p x = fv t \ fv t'" "fv\<^sub>e\<^sub>q ac x = {}" using Equality by auto thus ?thesis using assms Equality subst_apply_fv_unfold[of _ \] False by auto qed next case (Inequality X F) hence 1: "set X \ (subst_domain \ \ range_vars \) = {}" "x \\<^sub>s\<^sub>t\<^sub>p \ = Inequality X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" "rm_vars (set X) \ = \" using assms ineq_apply_subst[of \ X F] rm_vars_apply'[of \ "set X"] unfolding range_vars_alt_def by force+ have 2: "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" using Inequality by auto hence "fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) - set X" using fv\<^sub>s\<^sub>e\<^sub>t_subst_img_eq[OF 1(1), of "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F"] by simp hence 3: "fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" by (metis fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_step_subst) have 4: "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" using 1(2) by auto show ?thesis using assms(1) Inequality subst_apply_fv_unfold[of _ \] 1(2) 2 3 4 unfolding fv\<^sub>e\<^sub>q_def fv\<^sub>r\<^sub>c\<^sub>v_def fv\<^sub>s\<^sub>n\<^sub>d_def by (metis (no_types) Sup_empty image_empty fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.simps fv\<^sub>s\<^sub>e\<^sub>t.simps fv\<^sub>s\<^sub>t\<^sub>p.simps(4) strand_step.simps(20)) qed lemma fv_strand_subst: assumes "P = fv\<^sub>s\<^sub>t\<^sub>p \ P = fv\<^sub>r\<^sub>c\<^sub>v \ P = fv\<^sub>s\<^sub>n\<^sub>d \ P = fv\<^sub>e\<^sub>q ac \ P = fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q" and "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (\(set (map P S)))) = \(set (map P (S \\<^sub>s\<^sub>t \)))" using assms(2) proof (induction S) case (Cons x S) hence *: "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" "set (bvars\<^sub>s\<^sub>t\<^sub>p x) \ (subst_domain \ \ range_vars \) = {}" unfolding bvars\<^sub>s\<^sub>t_def by force+ hence **: "fv\<^sub>s\<^sub>e\<^sub>t (\ ` P x) = P (x \\<^sub>s\<^sub>t\<^sub>p \)" using fv_strand_step_subst[OF assms(1), of x \] by auto have "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (\(set (map P (x#S))))) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` P x) \ (\(set (map P ((S \\<^sub>s\<^sub>t \)))))" using Cons unfolding range_vars_alt_def bvars\<^sub>s\<^sub>t_def by force hence "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (\(set (map P (x#S))))) = P (x \\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (\(set (map P S))))" using ** by simp thus ?case using Cons.IH[OF *(1)] unfolding bvars\<^sub>s\<^sub>t_def by simp qed simp lemma fv_strand_subst2: assumes "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (wfrestrictedvars\<^sub>s\<^sub>t S)) = wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" by (metis (no_types, lifting) assms fv\<^sub>s\<^sub>e\<^sub>t.simps vars_snd_rcv_strand2 fv_strand_subst UN_Un image_Un) lemma fv_strand_subst': assumes "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (fv\<^sub>s\<^sub>t S)) = fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" by (metis assms fv_strand_subst fv\<^sub>s\<^sub>t_def) lemma fv_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s: "fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" by auto lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_in_fv_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s: "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F)" using fv_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s[of F] by blast lemma trms\<^sub>s\<^sub>t_append: "trms\<^sub>s\<^sub>t (A@B) = trms\<^sub>s\<^sub>t A \ trms\<^sub>s\<^sub>t B" by auto lemma trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst: "trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (a \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) = trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s a \\<^sub>s\<^sub>e\<^sub>t \" by (auto simp add: subst_apply_pairs_def) lemma trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_fv_subst_subset: "t \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv (t \ \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" by (force simp add: subst_apply_pairs_def) lemma trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_fv_subst_subset': fixes t::"('a,'b) term" and \::"('a,'b) subst" assumes "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F)" shows "fv (t \ \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" proof - { fix x assume "x \ fv t" hence "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" using fv_subset[OF assms] fv_subterms_set[of "trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F"] fv_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s[of F] by blast hence "fv (\ x) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_fv_subset by fast } thus ?thesis by (meson fv_subst_obtain_var subset_iff) qed lemma trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_funs_term_cases: assumes "t \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" "f \ funs_term t" shows "(\u \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F. f \ funs_term u) \ (\x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F. f \ funs_term (\ x))" using assms(1) proof (induction F) case (Cons g F) obtain s u where g: "g = (s,u)" by (metis surj_pair) show ?case proof (cases "t \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)") case False thus ?thesis using assms(2) Cons.prems g funs_term_subst[of _ \] by (auto simp add: subst_apply_pairs_def) qed (use Cons.IH in fastforce) qed simp lemma trm\<^sub>s\<^sub>t\<^sub>p_subst: assumes "subst_domain \ \ set (bvars\<^sub>s\<^sub>t\<^sub>p a) = {}" shows "trms\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>t\<^sub>p \) = trms\<^sub>s\<^sub>t\<^sub>p a \\<^sub>s\<^sub>e\<^sub>t \" proof - have "rm_vars (set (bvars\<^sub>s\<^sub>t\<^sub>p a)) \ = \" using assms by force thus ?thesis using assms by (auto simp add: subst_apply_pairs_def subst_apply_strand_step_def split: strand_step.splits) qed lemma trms\<^sub>s\<^sub>t_subst: assumes "subst_domain \ \ bvars\<^sub>s\<^sub>t A = {}" shows "trms\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>t \) = trms\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" using assms proof (induction A) case (Cons a A) have 1: "subst_domain \ \ bvars\<^sub>s\<^sub>t A = {}" "subst_domain \ \ set (bvars\<^sub>s\<^sub>t\<^sub>p a) = {}" using Cons.prems by auto hence IH: "trms\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ = trms\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>t \)" using Cons.IH by simp have "trms\<^sub>s\<^sub>t (a#A) = trms\<^sub>s\<^sub>t\<^sub>p a \ trms\<^sub>s\<^sub>t A" by auto hence 2: "trms\<^sub>s\<^sub>t (a#A) \\<^sub>s\<^sub>e\<^sub>t \ = (trms\<^sub>s\<^sub>t\<^sub>p a \\<^sub>s\<^sub>e\<^sub>t \) \ (trms\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \)" by (metis image_Un) have "trms\<^sub>s\<^sub>t (a#A \\<^sub>s\<^sub>t \) = (trms\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>t\<^sub>p \)) \ trms\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>t \)" by (auto simp add: subst_apply_strand_def) hence 3: "trms\<^sub>s\<^sub>t (a#A \\<^sub>s\<^sub>t \) = (trms\<^sub>s\<^sub>t\<^sub>p a \\<^sub>s\<^sub>e\<^sub>t \) \ trms\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>t \)" using trm\<^sub>s\<^sub>t\<^sub>p_subst[OF 1(2)] by auto show ?case using IH 2 3 by metis qed (simp add: subst_apply_strand_def) lemma strand_map_set_subst: assumes \: "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "\(set (map trms\<^sub>s\<^sub>t\<^sub>p (S \\<^sub>s\<^sub>t \))) = (\(set (map trms\<^sub>s\<^sub>t\<^sub>p S))) \\<^sub>s\<^sub>e\<^sub>t \" using assms proof (induction S) case (Cons x S) hence "bvars\<^sub>s\<^sub>t [x] \ subst_domain \ = {}" "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" unfolding bvars\<^sub>s\<^sub>t_def by force+ hence *: "subst_domain \ \ set (bvars\<^sub>s\<^sub>t\<^sub>p x) = {}" "\(set (map trms\<^sub>s\<^sub>t\<^sub>p (S \\<^sub>s\<^sub>t \))) = \(set (map trms\<^sub>s\<^sub>t\<^sub>p S)) \\<^sub>s\<^sub>e\<^sub>t \" using Cons.IH(1) bvars\<^sub>s\<^sub>t_singleton[of x] by auto hence "trms\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = (trms\<^sub>s\<^sub>t\<^sub>p x) \\<^sub>s\<^sub>e\<^sub>t \" proof (cases x) case (Inequality X F) thus ?thesis using rm_vars_apply'[of \ "set X"] * by (metis (no_types, lifting) image_cong trm\<^sub>s\<^sub>t\<^sub>p_subst) qed simp_all thus ?case using * subst_all_insert by auto qed simp lemma subst_apply_fv_subset_strand_trm: assumes P: "P = fv\<^sub>s\<^sub>t\<^sub>p \ P = fv\<^sub>r\<^sub>c\<^sub>v \ P = fv\<^sub>s\<^sub>n\<^sub>d \ P = fv\<^sub>e\<^sub>q ac \ P = fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q" and fv_sub: "fv t \ \(set (map P S)) \ V" and \: "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "fv (t \ \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using fv_strand_subst[OF P \] subst_apply_fv_subset[OF fv_sub, of \] by force lemma subst_apply_fv_subset_strand_trm2: assumes fv_sub: "fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" and \: "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "fv (t \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using fv_strand_subst2[OF \] subst_apply_fv_subset[OF fv_sub, of \] by force lemma subst_apply_fv_subset_strand: assumes P: "P = fv\<^sub>s\<^sub>t\<^sub>p \ P = fv\<^sub>r\<^sub>c\<^sub>v \ P = fv\<^sub>s\<^sub>n\<^sub>d \ P = fv\<^sub>e\<^sub>q ac \ P = fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q" and P_subset: "P x \ \(set (map P S)) \ V" and \: "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" "set (bvars\<^sub>s\<^sub>t\<^sub>p x) \ (subst_domain \ \ range_vars \) = {}" shows "P (x \\<^sub>s\<^sub>t\<^sub>p \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" proof (cases x) case (Send ts) hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" by auto hence **: "(P x = fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ \(set (map P S)) \ V" using P_subset by auto hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using subst_apply_fv_subset_strand_trm[OF P _ assms(3), of _ V] by fastforce hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)\ by force } ultimately show ?thesis by metis next case (Receive ts) hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "fv\<^sub>r\<^sub>c\<^sub>v x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" by auto hence **: "(P x = fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ \(set (map P S)) \ V" using P_subset by auto hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using subst_apply_fv_subset_strand_trm[OF P _ assms(3), of _ V] by fastforce hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)\ by blast } ultimately show ?thesis by metis next case (Equality ac' t t') show ?thesis proof (cases "ac' = ac") case True hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv t \ fv t'" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = fv t \ fv t'" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" using Equality by auto hence **: "(P x = fv t \ fv t' \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv t \ fv t'" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" hence "fv t \ \(set (map P S)) \ V" "fv t' \ \(set (map P S)) \ V" using P_subset by auto hence "fv (t \ \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" "fv (t' \ \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using P subst_apply_fv_subset_strand_trm assms by metis+ hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)\ by blast } ultimately show ?thesis by metis next case False hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv t \ fv t'" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" using Equality by auto hence **: "(P x = fv t \ fv t' \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv t \ fv t'" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" hence "fv t \ \(set (map P S)) \ V" "fv t' \ \(set (map P S)) \ V" using P_subset by auto hence "fv (t \ \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" "fv (t' \ \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using P subst_apply_fv_subset_strand_trm assms by metis+ hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)\ by blast } ultimately show ?thesis by metis qed next case (Inequality X F) hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" using \(2) ineq_apply_subst[of \ X F] by force+ hence **: "(P x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ \(set (map P S)) \ V" using P_subset by auto hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" proof (induction F) case (Cons f G) hence IH: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" by (metis (no_types, lifting) Diff_subset_conv UN_insert le_sup_iff list.simps(15) fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.simps) obtain t t' where f: "f = (t,t')" by (metis surj_pair) hence "fv t \ \(set (map P S)) \ (V \ set X)" "fv t' \ \(set (map P S)) \ (V \ set X)" using Cons.prems by auto hence "fv (t \ \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" "fv (t' \ \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" using subst_apply_fv_subset_strand_trm[OF P _ assms(3)] by blast+ thus ?case using f IH by (auto simp add: subst_apply_pairs_def) qed (simp add: subst_apply_pairs_def) moreover have "fv\<^sub>s\<^sub>e\<^sub>t (\ ` set X) = set X" using assms(4) Inequality by force ultimately have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" by auto hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X\ by blast } ultimately show ?thesis by metis qed lemma subst_apply_fv_subset_strand2: assumes P: "P = fv\<^sub>s\<^sub>t\<^sub>p \ P = fv\<^sub>r\<^sub>c\<^sub>v \ P = fv\<^sub>s\<^sub>n\<^sub>d \ P = fv\<^sub>e\<^sub>q ac \ P = fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q \ P = fv_r\<^sub>e\<^sub>q ac" and P_subset: "P x \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" and \: "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" "set (bvars\<^sub>s\<^sub>t\<^sub>p x) \ (subst_domain \ \ range_vars \) = {}" shows "P (x \\<^sub>s\<^sub>t\<^sub>p \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" proof (cases x) case (Send ts) hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv_r\<^sub>e\<^sub>q ac x = {}" "fv_r\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" by auto hence **: "(P x = fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using P_subset by auto hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using subst_apply_fv_subset_strand_trm2[OF _ assms(3), of _ V] by fastforce hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)\ by blast } ultimately show ?thesis by metis next case (Receive ts) hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "fv\<^sub>r\<^sub>c\<^sub>v x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv_r\<^sub>e\<^sub>q ac x = {}" "fv_r\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" by auto hence **: "(P x = fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using P_subset by auto hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using subst_apply_fv_subset_strand_trm2[OF _ assms(3), of _ V] by fastforce hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)\ by blast } ultimately show ?thesis by metis next case (Equality ac' t t') show ?thesis proof (cases "ac' = ac") case True hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv t \ fv t'" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = fv t \ fv t'" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv_r\<^sub>e\<^sub>q ac x = fv t'" "fv_r\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t' \ \)" using Equality by auto hence **: "(P x = fv t \ fv t' \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}) \ (P x = fv t' \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t' \ \))" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv t \ fv t'" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" hence "fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using P_subset by auto hence "fv (t \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" "fv (t' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using P subst_apply_fv_subset_strand_trm2 assms by blast+ hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)\ by blast } moreover { assume "P x = fv t'" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t' \ \)" hence "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using P_subset by auto hence "fv (t' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using P subst_apply_fv_subset_strand_trm2 assms by blast+ hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t' \ \)\ by blast } ultimately show ?thesis by metis next case False hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv t \ fv t'" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv_r\<^sub>e\<^sub>q ac x = {}" "fv_r\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" using Equality by auto hence **: "(P x = fv t \ fv t' \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}) \ (P x = fv t' \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t' \ \))" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv t \ fv t'" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" hence "fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using P_subset by auto hence "fv (t \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" "fv (t' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using P subst_apply_fv_subset_strand_trm2 assms by blast+ hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)\ by blast } moreover { assume "P x = fv t'" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t' \ \)" hence "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using P_subset by auto hence "fv (t' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using P subst_apply_fv_subset_strand_trm2 assms by blast+ hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t' \ \)\ by blast } ultimately show ?thesis by metis qed next case (Inequality X F) hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" "fv_r\<^sub>e\<^sub>q ac x = {}" "fv_r\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" using \(2) ineq_apply_subst[of \ X F] by force+ hence **: "(P x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using P_subset by auto hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" proof (induction F) case (Cons f G) hence IH: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" by (metis (no_types, lifting) Diff_subset_conv UN_insert le_sup_iff list.simps(15) fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.simps) obtain t t' where f: "f = (t,t')" by (metis surj_pair) hence "fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ (V \ set X)" "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ (V \ set X)" using Cons.prems by auto hence "fv (t \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" "fv (t' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" using subst_apply_fv_subset_strand_trm2[OF _ assms(3)] P by blast+ thus ?case using f IH by (auto simp add: subst_apply_pairs_def) qed (simp add: subst_apply_pairs_def) moreover have "fv\<^sub>s\<^sub>e\<^sub>t (\ ` set X) = set X" using assms(4) Inequality by force ultimately have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" by fastforce hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X\ by blast } ultimately show ?thesis by metis qed lemma strand_subst_fv_bounded_if_img_bounded: assumes "range_vars \ \ fv\<^sub>s\<^sub>t S" shows "fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t S" using subst_sends_strand_fv_to_img[of S \] assms by blast lemma strand_fv_subst_subset_if_subst_elim: assumes "subst_elim \ v" and "v \ fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "v \ fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" proof (cases "v \ fv\<^sub>s\<^sub>t S") case True thus ?thesis proof (induction S) case (Cons x S) have *: "v \ fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \)" using assms(1) proof (cases x) case (Inequality X F) hence "subst_elim (rm_vars (set X) \) v \ v \ set X" using assms(1) by blast moreover have "fv\<^sub>s\<^sub>t\<^sub>p (Inequality X F \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) - set X" using Inequality by auto ultimately have "v \ fv\<^sub>s\<^sub>t\<^sub>p (Inequality X F \\<^sub>s\<^sub>t\<^sub>p \)" by (induct F) (auto simp add: subst_elim_def subst_apply_pairs_def) thus ?thesis using Inequality by simp qed (simp_all add: subst_elim_def) moreover have "v \ fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" using Cons.IH proof (cases "v \ fv\<^sub>s\<^sub>t S") case False moreover have "v \ range_vars \" by (simp add: subst_elimD''[OF assms(1)] range_vars_alt_def) ultimately show ?thesis by (meson UnE subsetCE subst_sends_strand_fv_to_img) qed simp ultimately show ?case by auto qed simp next case False thus ?thesis using assms fv_strand_subst' unfolding subst_elim_def - by (metis (mono_tags, opaque_lifting) fv\<^sub>s\<^sub>e\<^sub>t.simps imageE mem_simps(8) subst_apply_term.simps(1)) + by (metis (mono_tags, opaque_lifting) fv\<^sub>s\<^sub>e\<^sub>t.simps imageE mem_simps(8) eval_term.simps(1)) qed lemma strand_fv_subst_subset_if_subst_elim': assumes "subst_elim \ v" "v \ fv\<^sub>s\<^sub>t S" "range_vars \ \ fv\<^sub>s\<^sub>t S" shows "fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t S" using strand_fv_subst_subset_if_subst_elim[OF assms(1)] assms(2) strand_subst_fv_bounded_if_img_bounded[OF assms(3)] by blast lemma fv_ik_is_fv_rcv: "fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t S) = \(set (map fv\<^sub>r\<^sub>c\<^sub>v S))" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma fv_ik_subset_fv_st[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t S) \ wfrestrictedvars\<^sub>s\<^sub>t S" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma fv_assignment_rhs_subset_fv_st[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (assignment_rhs\<^sub>s\<^sub>t S) \ wfrestrictedvars\<^sub>s\<^sub>t S" by (induct S rule: assignment_rhs\<^sub>s\<^sub>t.induct) force+ lemma fv_ik_subset_fv_st'[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t S) \ fv\<^sub>s\<^sub>t S" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik\<^sub>s\<^sub>t_var_is_fv: "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t A) \ x \ fv\<^sub>s\<^sub>t A" by (meson fv_ik_subset_fv_st'[of A] fv_subset_subterms subsetCE term.set_intros(3)) lemma fv_assignment_rhs_subset_fv_st'[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (assignment_rhs\<^sub>s\<^sub>t S) \ fv\<^sub>s\<^sub>t S" by (induct S rule: assignment_rhs\<^sub>s\<^sub>t.induct) auto lemma ik\<^sub>s\<^sub>t_assignment_rhs\<^sub>s\<^sub>t_wfrestrictedvars_subset: "fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>s\<^sub>t A) \ wfrestrictedvars\<^sub>s\<^sub>t A" using fv_ik_subset_fv_st[of A] fv_assignment_rhs_subset_fv_st[of A] by simp+ lemma strand_step_id_subst[iff]: "x \\<^sub>s\<^sub>t\<^sub>p Var = x" by (cases x) auto lemma strand_id_subst[iff]: "S \\<^sub>s\<^sub>t Var = S" using strand_step_id_subst by (induct S) auto lemma strand_subst_vars_union_bound[simp]: "vars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ vars\<^sub>s\<^sub>t S \ range_vars \" proof (induction S) case (Cons x S) moreover have "vars\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ vars\<^sub>s\<^sub>t\<^sub>p x \ range_vars \" using subst_sends_fv_to_img[of _ \] proof (cases x) case (Inequality X F) define \' where "\' \ rm_vars (set X) \" have 0: "range_vars \' \ range_vars \" using rm_vars_img[of "set X" \] by (auto simp add: \'_def subst_domain_def range_vars_alt_def) have "vars\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \') \ set X" "vars\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ set X" using Inequality by (auto simp add: \'_def) moreover have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \') \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ range_vars \" proof (induction F) case (Cons f G) obtain t t' where f: "f = (t,t')" by moura hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (f#G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \') = fv (t \ \') \ fv (t' \ \') \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \')" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (f#G) = fv t \ fv t' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" by (auto simp add: subst_apply_pairs_def) thus ?case using 0 Cons.IH subst_sends_fv_to_img[of t \'] subst_sends_fv_to_img[of t' \'] unfolding f by auto qed (simp add: subst_apply_pairs_def) ultimately show ?thesis by auto qed auto ultimately show ?case by auto qed simp lemma strand_vars_split: "vars\<^sub>s\<^sub>t (S@S') = vars\<^sub>s\<^sub>t S \ vars\<^sub>s\<^sub>t S'" "wfrestrictedvars\<^sub>s\<^sub>t (S@S') = wfrestrictedvars\<^sub>s\<^sub>t S \ wfrestrictedvars\<^sub>s\<^sub>t S'" "fv\<^sub>s\<^sub>t (S@S') = fv\<^sub>s\<^sub>t S \ fv\<^sub>s\<^sub>t S'" by auto lemma bvars_subst_ident: "bvars\<^sub>s\<^sub>t S = bvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" unfolding bvars\<^sub>s\<^sub>t_def by (induct S) (simp_all add: subst_apply_strand_step_def split: strand_step.splits) lemma strand_subst_subst_idem: assumes "subst_idem \" "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t S" "subst_domain \ \ fv\<^sub>s\<^sub>t S = {}" "range_vars \ \ bvars\<^sub>s\<^sub>t S = {}" "range_vars \ \ bvars\<^sub>s\<^sub>t S = {}" shows "(S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \ = (S \\<^sub>s\<^sub>t \)" and "(S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t (\ \\<^sub>s \) = (S \\<^sub>s\<^sub>t \)" proof - from assms(2,3) have "fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ subst_domain \ = {}" using subst_sends_strand_fv_to_img[of S \] by blast thus "(S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \ = (S \\<^sub>s\<^sub>t \)" by blast thus "(S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t (\ \\<^sub>s \) = (S \\<^sub>s\<^sub>t \)" by (metis assms(1,4,5) bvars_subst_ident strand_subst_comp subst_idem_def) qed lemma strand_subst_img_bound: assumes "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t S" and "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "range_vars \ \ fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" proof - have "subst_domain \ \ \(set (map fv\<^sub>s\<^sub>t\<^sub>p S))" by (metis (no_types) fv\<^sub>s\<^sub>t_def Un_subset_iff assms(1)) thus ?thesis unfolding range_vars_alt_def fv\<^sub>s\<^sub>t_def by (metis subst_range.simps fv_set_mono fv_strand_subst Int_commute assms(2) image_Un le_iff_sup) qed lemma strand_subst_img_bound': assumes "subst_domain \ \ range_vars \ \ vars\<^sub>s\<^sub>t S" and "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "range_vars \ \ vars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" proof - have "(subst_domain \ \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` subst_domain \)) \ vars\<^sub>s\<^sub>t S = subst_domain \ \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` subst_domain \)" using assms(1) by (metis inf.absorb_iff1 range_vars_alt_def subst_range.simps) hence "range_vars \ \ fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" using vars_snd_rcv_strand fv_snd_rcv_strand assms(2) strand_subst_img_bound unfolding range_vars_alt_def by (metis (no_types) inf_le2 inf_sup_distrib1 subst_range.simps sup_bot.right_neutral) thus "range_vars \ \ vars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" by (metis fv_snd_rcv_strand le_supI1 vars_snd_rcv_strand) qed lemma strand_subst_all_fv_subset: assumes "fv t \ fv\<^sub>s\<^sub>t S" "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "fv (t \ \) \ fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" using assms by (metis fv_strand_subst' Int_commute subst_apply_fv_subset) lemma strand_subst_not_dom_fixed: assumes "v \ fv\<^sub>s\<^sub>t S" and "v \ subst_domain \" shows "v \ fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" using assms proof (induction S) case (Cons x S') have 1: "\X. v \ subst_domain (rm_vars (set X) \)" using Cons.prems(2) rm_vars_dom_subset by force show ?case proof (cases "v \ fv\<^sub>s\<^sub>t S'") case True thus ?thesis using Cons.IH[OF _ Cons.prems(2)] by auto next case False hence 2: "v \ fv\<^sub>s\<^sub>t\<^sub>p x" using Cons.prems(1) by simp hence "v \ fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \)" using Cons.prems(2) subst_not_dom_fixed proof (cases x) case (Inequality X F) hence "v \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" using 2 by simp hence "v \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" using subst_not_dom_fixed[OF _ 1] by (induct F) (auto simp add: subst_apply_pairs_def) thus ?thesis using Inequality 2 by auto qed (force simp add: subst_domain_def)+ thus ?thesis by auto qed qed simp lemma strand_vars_unfold: "v \ vars\<^sub>s\<^sub>t S \ \S' x S''. S = S'@x#S'' \ v \ vars\<^sub>s\<^sub>t\<^sub>p x" proof (induction S) case (Cons x S) thus ?case proof (cases "v \ vars\<^sub>s\<^sub>t\<^sub>p x") case True thus ?thesis by blast next case False hence "v \ vars\<^sub>s\<^sub>t S" using Cons.prems by auto thus ?thesis using Cons.IH by (metis append_Cons) qed qed simp lemma strand_fv_unfold: "v \ fv\<^sub>s\<^sub>t S \ \S' x S''. S = S'@x#S'' \ v \ fv\<^sub>s\<^sub>t\<^sub>p x" proof (induction S) case (Cons x S) thus ?case proof (cases "v \ fv\<^sub>s\<^sub>t\<^sub>p x") case True thus ?thesis by blast next case False hence "v \ fv\<^sub>s\<^sub>t S" using Cons.prems by auto thus ?thesis using Cons.IH by (metis append_Cons) qed qed simp lemma subterm_if_in_strand_ik: "t \ ik\<^sub>s\<^sub>t S \ \ts. Receive ts \ set S \ t \\<^sub>s\<^sub>e\<^sub>t set ts" by (induct S rule: ik\<^sub>s\<^sub>t_induct) auto lemma fv_subset_if_in_strand_ik: "t \ ik\<^sub>s\<^sub>t S \ fv t \ \(set (map fv\<^sub>r\<^sub>c\<^sub>v S))" proof - assume "t \ ik\<^sub>s\<^sub>t S" then obtain ts where "Receive ts \ set S" "t \\<^sub>s\<^sub>e\<^sub>t set ts" by (metis subterm_if_in_strand_ik) hence "fv t \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" using subtermeq_vars_subset by auto thus ?thesis using in_strand_fv_subset_rcv[OF \Receive ts \ set S\] by auto qed lemma fv_subset_if_in_strand_ik': "t \ ik\<^sub>s\<^sub>t S \ fv t \ fv\<^sub>s\<^sub>t S" using fv_subset_if_in_strand_ik[of t S] fv_snd_rcv_strand_subset(2)[of S] by blast lemma vars_subset_if_in_strand_ik2: "t \ ik\<^sub>s\<^sub>t S \ fv t \ wfrestrictedvars\<^sub>s\<^sub>t S" using fv_subset_if_in_strand_ik[of t S] vars_snd_rcv_strand_subset2(2)[of S] by blast subsection \Lemmata: Simple Strands\ lemma simple_Cons[dest]: "simple (s#S) \ simple S" unfolding simple_def by auto lemma simple_split[dest]: assumes "simple (S@S')" shows "simple S" "simple S'" using assms unfolding simple_def by auto lemma simple_append[intro]: "\simple S; simple S'\ \ simple (S@S')" unfolding simple_def by auto lemma simple_append_sym[sym]: "simple (S@S') \ simple (S'@S)" by auto lemma not_simple_if_snd_fun: "Fun f T \ set ts \ S = S'@Send ts#S'' \ \simple S" by (metis simple_def list.set_cases list_all_append list_all_simps(1) simple\<^sub>s\<^sub>t\<^sub>p.simps(5,6)) lemma not_list_all_elim: "\list_all P A \ \B x C. A = B@x#C \ \P x \ list_all P B" proof (induction A rule: List.rev_induct) case (snoc a A) show ?case proof (cases "list_all P A") case True thus ?thesis using snoc.prems by auto next case False then obtain B x C where "A = B@x#C" "\P x" "list_all P B" using snoc.IH[OF False] by auto thus ?thesis by auto qed qed simp lemma not_simple\<^sub>s\<^sub>t\<^sub>p_elim: assumes "\simple\<^sub>s\<^sub>t\<^sub>p x" shows "(\ts. x = Send ts \ (\y. ts = [Var y])) \ (\a t t'. x = Equality a t t') \ (\X F. x = Inequality X F \ (\\. ineq_model \ X F))" using assms by (cases x) (fastforce elim: simple\<^sub>s\<^sub>t\<^sub>p.elims)+ lemma not_simple_elim: assumes "\simple S" shows "(\A B ts. S = A@Send ts#B \ (\x. ts = [Var x]) \ simple A) \ (\A B a t t'. S = A@Equality a t t'#B \ simple A) \ (\A B X F. S = A@Inequality X F#B \ (\\. ineq_model \ X F) \ simple A)" by (metis assms not_list_all_elim not_simple\<^sub>s\<^sub>t\<^sub>p_elim simple_def) lemma simple_snd_is_var: "\Send ts \ set S; simple S\ \ \v. ts = [Var v]" unfolding simple_def by (metis list_all_append list_all_simps(1) simple\<^sub>s\<^sub>t\<^sub>p.elims(2) split_list_first strand_step.distinct(1) strand_step.distinct(5) strand_step.inject(1)) subsection \Lemmata: Strand Measure\ lemma measure\<^sub>s\<^sub>t_wellfounded: "wf measure\<^sub>s\<^sub>t" unfolding measure\<^sub>s\<^sub>t_def by simp lemma strand_size_append[iff]: "size\<^sub>s\<^sub>t (S@S') = size\<^sub>s\<^sub>t S + size\<^sub>s\<^sub>t S'" by (induct S) (auto simp add: size\<^sub>s\<^sub>t_def) lemma strand_size_map_fun_lt[simp]: "size\<^sub>s\<^sub>t (map Send1 X) < size (Fun f X)" "size\<^sub>s\<^sub>t (map Send1 X) < size\<^sub>s\<^sub>t [Send [Fun f X]]" "size\<^sub>s\<^sub>t (map Receive1 X) < size\<^sub>s\<^sub>t [Receive [Fun f X]]" "size\<^sub>s\<^sub>t [Send X] < size\<^sub>s\<^sub>t [Send [Fun f X]]" "size\<^sub>s\<^sub>t [Receive X] < size\<^sub>s\<^sub>t [Receive [Fun f X]]" by (induct X) (auto simp add: size\<^sub>s\<^sub>t_def) lemma strand_size_rm_fun_lt[simp]: "size\<^sub>s\<^sub>t (S@S') < size\<^sub>s\<^sub>t (S@Send ts#S')" "size\<^sub>s\<^sub>t (S@S') < size\<^sub>s\<^sub>t (S@Receive ts#S')" by (induct S) (auto simp add: size\<^sub>s\<^sub>t_def) lemma strand_fv_card_map_fun_eq: "card (fv\<^sub>s\<^sub>t (S@Send [Fun f X]#S')) = card (fv\<^sub>s\<^sub>t (S@(map Send1 X)@S'))" proof - have "fv\<^sub>s\<^sub>t (S@Send [Fun f X]#S') = fv\<^sub>s\<^sub>t (S@(map Send1 X)@S')" by auto thus ?thesis by simp qed lemma strand_fv_card_rm_fun_le[simp]: "card (fv\<^sub>s\<^sub>t (S@S')) \ card (fv\<^sub>s\<^sub>t (S@Send [Fun f X]#S'))" by (force intro: card_mono) lemma strand_fv_card_rm_eq_le[simp]: "card (fv\<^sub>s\<^sub>t (S@S')) \ card (fv\<^sub>s\<^sub>t (S@Equality a t t'#S'))" by (force intro: card_mono) subsection \Lemmata: Well-formed Strands\ lemma wf_prefix[dest]: "wf\<^sub>s\<^sub>t V (S@S') \ wf\<^sub>s\<^sub>t V S" by (induct S rule: wf\<^sub>s\<^sub>t.induct) auto lemma wf_vars_mono[simp]: "wf\<^sub>s\<^sub>t V S \ wf\<^sub>s\<^sub>t (V \ W) S" proof (induction S arbitrary: V) case (Cons x S) thus ?case proof (cases x) case (Send ts) hence "wf\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ W) S" using Cons.prems(1) Cons.IH by simp thus ?thesis by (metis Send sup_assoc sup_commute wf\<^sub>s\<^sub>t.simps(3)) next case (Equality a t t') show ?thesis proof (cases a) case Assign hence "wf\<^sub>s\<^sub>t (V \ fv t \ W) S" "fv t' \ V \ W" using Equality Cons.prems(1) Cons.IH by auto thus ?thesis using Equality Assign by (simp add: sup_commute sup_left_commute) next case Check thus ?thesis using Equality Cons by auto qed qed auto qed simp lemma wf\<^sub>s\<^sub>tI[intro]: "wfrestrictedvars\<^sub>s\<^sub>t S \ V \ wf\<^sub>s\<^sub>t V S" proof (induction S) case (Cons x S) thus ?case proof (cases x) case (Send ts) hence "wf\<^sub>s\<^sub>t V S" "V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) = V" using Cons by auto thus ?thesis using Send by simp next case (Equality a t t') show ?thesis proof (cases a) case Assign hence "wf\<^sub>s\<^sub>t V S" "fv t' \ V" using Equality Cons by auto thus ?thesis using wf_vars_mono Equality Assign by simp next case Check thus ?thesis using Equality Cons by auto qed qed simp_all qed simp lemma wf\<^sub>s\<^sub>tI'[intro]: "\(fv\<^sub>r\<^sub>c\<^sub>v ` set S) \ \(fv_r\<^sub>e\<^sub>q assign ` set S) \ V \ wf\<^sub>s\<^sub>t V S" proof (induction S) case (Cons x S) thus ?case proof (cases x) case (Equality a t t') thus ?thesis using Cons by (cases a) auto qed simp_all qed simp lemma wf_append_exec: "wf\<^sub>s\<^sub>t V (S@S') \ wf\<^sub>s\<^sub>t (V \ wfvarsoccs\<^sub>s\<^sub>t S) S'" proof (induction S arbitrary: V) case (Cons x S V) thus ?case proof (cases x) case (Send ts) hence "wf\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfvarsoccs\<^sub>s\<^sub>t S) S'" using Cons.prems Cons.IH by simp thus ?thesis using Send by (auto simp add: sup_assoc) next case (Equality a t t') show ?thesis proof (cases a) case Assign hence "wf\<^sub>s\<^sub>t (V \ fv t \ wfvarsoccs\<^sub>s\<^sub>t S) S'" using Equality Cons.prems Cons.IH by auto thus ?thesis using Equality Assign by (auto simp add: sup_assoc) next case Check hence "wf\<^sub>s\<^sub>t (V \ wfvarsoccs\<^sub>s\<^sub>t S) S'" using Equality Cons.prems Cons.IH by auto thus ?thesis using Equality Check by (auto simp add: sup_assoc) qed qed auto qed simp lemma wf_append_suffix: "wf\<^sub>s\<^sub>t V S \ wfrestrictedvars\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V \ wf\<^sub>s\<^sub>t V (S@S')" proof (induction V S rule: wf\<^sub>s\<^sub>t_induct) case (ConsSnd V ts S) hence *: "wf\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) S" by simp_all hence "wfrestrictedvars\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>t S \ (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts))" using ConsSnd.prems(2) by fastforce thus ?case using ConsSnd.IH * by simp next case (ConsRcv V ts S) hence *: "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ V" "wf\<^sub>s\<^sub>t V S" by simp_all hence "wfrestrictedvars\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using ConsRcv.prems(2) by fastforce thus ?case using ConsRcv.IH * by simp next case (ConsEq V t t' S) hence *: "fv t' \ V" "wf\<^sub>s\<^sub>t (V \ fv t) S" by simp_all moreover have "vars\<^sub>s\<^sub>t\<^sub>p (Equality Assign t t') = fv t \ fv t'" by simp moreover have "wfrestrictedvars\<^sub>s\<^sub>t (Equality Assign t t'#S) = fv t \ fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S" by auto ultimately have "wfrestrictedvars\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>t S \ (V \ fv t)" using ConsEq.prems(2) by blast thus ?case using ConsEq.IH * by simp qed (simp_all add: wf\<^sub>s\<^sub>tI) lemma wf_append_suffix': assumes "wf\<^sub>s\<^sub>t V S" and "\(fv\<^sub>r\<^sub>c\<^sub>v ` set S') \ \(fv_r\<^sub>e\<^sub>q assign ` set S') \ wfvarsoccs\<^sub>s\<^sub>t S \ V" shows "wf\<^sub>s\<^sub>t V (S@S')" using assms proof (induction V S rule: wf\<^sub>s\<^sub>t_induct) case (ConsSnd V ts S) hence *: "wf\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) S" by simp_all have "wfvarsoccs\<^sub>s\<^sub>t (send\ts\\<^sub>s\<^sub>t#S) = fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfvarsoccs\<^sub>s\<^sub>t S" unfolding wfvarsoccs\<^sub>s\<^sub>t_def by simp hence "(\a\set S'. fv\<^sub>r\<^sub>c\<^sub>v a) \ (\a\set S'. fv_r\<^sub>e\<^sub>q assign a) \ wfvarsoccs\<^sub>s\<^sub>t S \ (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts))" using ConsSnd.prems(2) unfolding wfvarsoccs\<^sub>s\<^sub>t_def by auto thus ?case using ConsSnd.IH[OF *] by auto next case (ConsEq V t t' S) hence *: "fv t' \ V" "wf\<^sub>s\<^sub>t (V \ fv t) S" by simp_all have "wfvarsoccs\<^sub>s\<^sub>t (\assign: t \ t'\\<^sub>s\<^sub>t#S) = fv t \ wfvarsoccs\<^sub>s\<^sub>t S" unfolding wfvarsoccs\<^sub>s\<^sub>t_def by simp hence "(\a\set S'. fv\<^sub>r\<^sub>c\<^sub>v a) \ (\a\set S'. fv_r\<^sub>e\<^sub>q assign a) \ wfvarsoccs\<^sub>s\<^sub>t S \ (V \ fv t)" using ConsEq.prems(2) unfolding wfvarsoccs\<^sub>s\<^sub>t_def by auto thus ?case using ConsEq.IH[OF *(2)] *(1) by auto qed (auto simp add: wf\<^sub>s\<^sub>tI') lemma wf_send_compose: "wf\<^sub>s\<^sub>t V (S@(map Send1 X)@S') = wf\<^sub>s\<^sub>t V (S@Send1 (Fun f X)#S')" proof (induction S arbitrary: V) case Nil thus ?case proof (induction X arbitrary: V) case (Cons y Y) thus ?case by (simp add: sup_assoc) qed simp next case (Cons s S) thus ?case proof (cases s) case (Equality ac t t') thus ?thesis using Cons by (cases ac) auto qed auto qed lemma wf_snd_append[iff]: "wf\<^sub>s\<^sub>t V (S@[Send t]) = wf\<^sub>s\<^sub>t V S" by (induct S rule: wf\<^sub>s\<^sub>t.induct) simp_all lemma wf_snd_append': "wf\<^sub>s\<^sub>t V S \ wf\<^sub>s\<^sub>t V (Send t#S)" by simp lemma wf_rcv_append[dest]: "wf\<^sub>s\<^sub>t V (S@Receive t#S') \ wf\<^sub>s\<^sub>t V (S@S')" by (induct S rule: wf\<^sub>s\<^sub>t.induct) simp_all lemma wf_rcv_append'[intro]: "\wf\<^sub>s\<^sub>t V (S@S'); fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfrestrictedvars\<^sub>s\<^sub>t S \ V\ \ wf\<^sub>s\<^sub>t V (S@Receive ts#S')" proof (induction S rule: wf\<^sub>s\<^sub>t_induct) case (ConsRcv V t' S) hence "wf\<^sub>s\<^sub>t V (S@S')" "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" by (simp, fastforce) thus ?case using ConsRcv by auto next case (ConsEq V t' t'' S) hence "fv t'' \ V" by simp moreover have "wfrestrictedvars\<^sub>s\<^sub>t (Equality Assign t' t''#S) = fv t' \ fv t'' \ wfrestrictedvars\<^sub>s\<^sub>t S" by auto ultimately have "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfrestrictedvars\<^sub>s\<^sub>t S \ (V \ fv t')" using ConsEq.prems(2) by blast thus ?case using ConsEq by auto qed auto lemma wf_rcv_append''[intro]: "\wf\<^sub>s\<^sub>t V S; fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ \(set (map fv\<^sub>s\<^sub>n\<^sub>d S))\ \ wf\<^sub>s\<^sub>t V (S@[Receive ts])" by (induct S) (simp, metis vars_snd_rcv_strand_subset2(1) append_Nil2 le_supI1 order_trans wf_rcv_append') lemma wf_rcv_append'''[intro]: "\wf\<^sub>s\<^sub>t V S; fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfrestrictedvars\<^sub>s\<^sub>t S \ V\ \ wf\<^sub>s\<^sub>t V (S@[Receive ts])" by (simp add: wf_rcv_append'[of _ _ "[]"]) lemma wf_eq_append[dest]: "wf\<^sub>s\<^sub>t V (S@Equality a t t'#S') \ fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ V \ wf\<^sub>s\<^sub>t V (S@S')" proof (induction S rule: wf\<^sub>s\<^sub>t_induct) case (Nil V) hence "wf\<^sub>s\<^sub>t (V \ fv t) S'" by (cases a) auto moreover have "V \ fv t = V" using Nil by auto ultimately show ?case by simp next case (ConsRcv V us S) hence "wf\<^sub>s\<^sub>t V (S @ Equality a t t' # S')" "fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" "fv\<^sub>s\<^sub>e\<^sub>t (set us) \ V" by fastforce+ hence "wf\<^sub>s\<^sub>t V (S@S')" using ConsRcv.IH by auto thus ?case using \fv\<^sub>s\<^sub>e\<^sub>t (set us) \ V\ by simp next case (ConsEq V u u' S) hence "wf\<^sub>s\<^sub>t (V \ fv u) (S@Equality a t t'#S')" "fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ (V \ fv u)" "fv u' \ V" by auto hence "wf\<^sub>s\<^sub>t (V \ fv u) (S@S')" using ConsEq.IH by auto thus ?case using \fv u' \ V\ by simp qed auto lemma wf_eq_append'[intro]: "\wf\<^sub>s\<^sub>t V (S@S'); fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V\ \ wf\<^sub>s\<^sub>t V (S@Equality a t t'#S')" proof (induction S rule: wf\<^sub>s\<^sub>t_induct) case Nil thus ?case by (cases a) auto next case (ConsEq V u u' S) hence "wf\<^sub>s\<^sub>t (V \ fv u) (S@S')" "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V \ fv u" by fastforce+ thus ?case using ConsEq by auto next case (ConsEq2 V u u' S) hence "wf\<^sub>s\<^sub>t V (S@S')" by auto thus ?case using ConsEq2 by auto next case (ConsRcv V u S) hence "wf\<^sub>s\<^sub>t V (S@S')" "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" by fastforce+ thus ?case using ConsRcv by auto next case (ConsSnd V us S) hence "wf\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set us)) (S@S')" "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ (V \ fv\<^sub>s\<^sub>e\<^sub>t (set us))" by fastforce+ thus ?case using ConsSnd by auto qed auto lemma wf_eq_append''[intro]: "\wf\<^sub>s\<^sub>t V (S@S'); fv t' \ wfvarsoccs\<^sub>s\<^sub>t S \ V\ \ wf\<^sub>s\<^sub>t V (S@[Equality a t t']@S')" proof (induction S rule: wf\<^sub>s\<^sub>t_induct) case Nil thus ?case by (cases a) auto next case (ConsEq V u u' S) hence "wf\<^sub>s\<^sub>t (V \ fv u) (S@S')" "fv t' \ wfvarsoccs\<^sub>s\<^sub>t S \ V \ fv u" by fastforce+ thus ?case using ConsEq by auto next case (ConsEq2 V u u' S) hence "wf\<^sub>s\<^sub>t (V \ fv u) (S@S')" "fv t' \ wfvarsoccs\<^sub>s\<^sub>t S \ V \ fv u" by fastforce+ thus ?case using ConsEq2 by auto next case (ConsRcv V u S) hence "wf\<^sub>s\<^sub>t V (S@S')" "fv t' \ wfvarsoccs\<^sub>s\<^sub>t S \ V" by fastforce+ thus ?case using ConsRcv by auto next case (ConsSnd V us S) hence "wf\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set us)) (S@S')" "fv t' \ wfvarsoccs\<^sub>s\<^sub>t S \ (V \ fv\<^sub>s\<^sub>e\<^sub>t (set us))" by auto thus ?case using ConsSnd by auto qed auto lemma wf_eq_append'''[intro]: "\wf\<^sub>s\<^sub>t V S; fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V\ \ wf\<^sub>s\<^sub>t V (S@[Equality a t t'])" by (simp add: wf_eq_append'[of _ _ "[]"]) lemma wf_eq_check_append[dest]: "wf\<^sub>s\<^sub>t V (S@Equality Check t t'#S') \ wf\<^sub>s\<^sub>t V (S@S')" by (induct S rule: wf\<^sub>s\<^sub>t.induct) simp_all lemma wf_eq_check_append'[intro]: "wf\<^sub>s\<^sub>t V (S@S') \ wf\<^sub>s\<^sub>t V (S@Equality Check t t'#S')" by (induct S rule: wf\<^sub>s\<^sub>t.induct) auto lemma wf_eq_check_append''[intro]: "wf\<^sub>s\<^sub>t V S \ wf\<^sub>s\<^sub>t V (S@[Equality Check t t'])" by (induct S rule: wf\<^sub>s\<^sub>t.induct) auto lemma wf_ineq_append[dest]: "wf\<^sub>s\<^sub>t V (S@Inequality X F#S') \ wf\<^sub>s\<^sub>t V (S@S')" by (induct S rule: wf\<^sub>s\<^sub>t.induct) simp_all lemma wf_ineq_append'[intro]: "wf\<^sub>s\<^sub>t V (S@S') \ wf\<^sub>s\<^sub>t V (S@Inequality X F#S')" by (induct S rule: wf\<^sub>s\<^sub>t.induct) auto lemma wf_ineq_append''[intro]: "wf\<^sub>s\<^sub>t V S \ wf\<^sub>s\<^sub>t V (S@[Inequality X F])" by (induct S rule: wf\<^sub>s\<^sub>t.induct) auto lemma wf_Receive1_prefix: assumes "wf\<^sub>s\<^sub>t X S" and "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ X" shows "wf\<^sub>s\<^sub>t X (map Receive1 ts@S)" using assms by (induct ts) simp_all lemma wf_Send1_prefix: assumes "wf\<^sub>s\<^sub>t (X \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) S" shows "wf\<^sub>s\<^sub>t X (map Send1 ts@S)" using assms by (induct ts arbitrary: X) (simp, force simp add: Un_assoc) lemma wf_rcv_fv_single[elim]: "wf\<^sub>s\<^sub>t V (Receive ts#S') \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ V" by simp lemma wf_rcv_fv: "wf\<^sub>s\<^sub>t V (S@Receive ts#S') \ fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfvarsoccs\<^sub>s\<^sub>t S \ V" proof (induction S arbitrary: V) case (Cons a S) thus ?case by (cases a) (auto split!: poscheckvariant.split) qed simp lemma wf_eq_fv: "wf\<^sub>s\<^sub>t V (S@Equality Assign t t'#S') \ fv t' \ wfvarsoccs\<^sub>s\<^sub>t S \ V" proof (induction S arbitrary: V) case (Cons a S) thus ?case by (cases a) (auto split!: poscheckvariant.split) qed simp lemma wf_simple_fv_occurrence: assumes "wf\<^sub>s\<^sub>t {} S" "simple S" "v \ wfrestrictedvars\<^sub>s\<^sub>t S" shows "\S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f. S = S\<^sub>p\<^sub>r\<^sub>e@Send [Var v]#S\<^sub>s\<^sub>u\<^sub>f \ v \ wfrestrictedvars\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e" using assms proof (induction S rule: List.rev_induct) case (snoc x S) from \wf\<^sub>s\<^sub>t {} (S@[x])\ have "wf\<^sub>s\<^sub>t {} S" "wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>s\<^sub>t S) [x]" using wf_append_exec[THEN wf_vars_mono, of "{}" S "[x]" "wfrestrictedvars\<^sub>s\<^sub>t S - wfvarsoccs\<^sub>s\<^sub>t S"] vars_snd_rcv_strand_subset2(4)[of S] Diff_partition[of "wfvarsoccs\<^sub>s\<^sub>t S" "wfrestrictedvars\<^sub>s\<^sub>t S"] by auto from \simple (S@[x])\ have "simple S" "simple\<^sub>s\<^sub>t\<^sub>p x" unfolding simple_def by auto show ?case proof (cases "v \ wfrestrictedvars\<^sub>s\<^sub>t S") case False show ?thesis proof (cases x) case (Receive ts) hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfrestrictedvars\<^sub>s\<^sub>t S" using \wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>s\<^sub>t S) [x]\ by simp hence "v \ wfrestrictedvars\<^sub>s\<^sub>t S" using \v \ wfrestrictedvars\<^sub>s\<^sub>t (S@[x])\ \x = Receive ts\ by auto thus ?thesis using \x = Receive ts\ snoc.IH[OF \wf\<^sub>s\<^sub>t {} S\ \simple S\] by fastforce next case (Send ts) hence "v \ vars\<^sub>s\<^sub>t\<^sub>p x" using \v \ wfrestrictedvars\<^sub>s\<^sub>t (S@[x])\ False by auto from Send obtain w where "ts = [Var w]" using \simple\<^sub>s\<^sub>t\<^sub>p x\ by (cases ts) (simp, metis in_set_conv_decomp simple_snd_is_var snoc.prems(2)) hence "v = w" using \x = Send ts\ \v \ vars\<^sub>s\<^sub>t\<^sub>p x\ by simp thus ?thesis using \x = Send ts\ \v \ wfrestrictedvars\<^sub>s\<^sub>t S\ \ts = [Var w]\ by auto next case (Equality ac t t') thus ?thesis using snoc.prems(2) unfolding simple_def by auto next case (Inequality t t') thus ?thesis using False snoc.prems(3) by auto qed qed (use snoc.IH[OF \wf\<^sub>s\<^sub>t {} S\ \simple S\] in fastforce) qed simp lemma Unifier_strand_fv_subset: assumes g_in_ik: "t \ ik\<^sub>s\<^sub>t S" and \: "Unifier \ (Fun f X) t" and disj: "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "fv (Fun f X \ \) \ \(set (map fv\<^sub>r\<^sub>c\<^sub>v (S \\<^sub>s\<^sub>t \)))" by (metis (no_types) fv_subset_if_in_strand_ik[OF g_in_ik] disj \ fv_strand_subst subst_apply_fv_subset) lemma wf\<^sub>s\<^sub>t_induct'[consumes 1, case_names Nil ConsSnd ConsRcv ConsEq ConsEq2 ConsIneq]: fixes S::"('a,'b) strand" assumes "wf\<^sub>s\<^sub>t V S" "P []" "\ts S. \wf\<^sub>s\<^sub>t V S; P S\ \ P (S@[Send ts])" "\ts S. \wf\<^sub>s\<^sub>t V S; P S; fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ V \ wfvarsoccs\<^sub>s\<^sub>t S\ \ P (S@[Receive ts])" "\t t' S. \wf\<^sub>s\<^sub>t V S; P S; fv t' \ V \ wfvarsoccs\<^sub>s\<^sub>t S\ \ P (S@[Equality Assign t t'])" "\t t' S. \wf\<^sub>s\<^sub>t V S; P S\ \ P (S@[Equality Check t t'])" "\X F S. \wf\<^sub>s\<^sub>t V S; P S\ \ P (S@[Inequality X F])" shows "P S" using assms proof (induction S rule: List.rev_induct) case (snoc x S) hence *: "wf\<^sub>s\<^sub>t V S" "wf\<^sub>s\<^sub>t (V \ wfvarsoccs\<^sub>s\<^sub>t S) [x]" by (metis wf_prefix, metis wf_append_exec) have IH: "P S" using snoc.IH[OF *(1)] snoc.prems by auto note ** = snoc.prems(3,4,5,6,7)[OF *(1) IH] *(2) show ?case using **(1,2,4,5,6) proof (cases x) case (Equality ac t t') then show ?thesis using **(3,4,6) by (cases ac) auto qed auto qed simp lemma wf_subst_apply: "wf\<^sub>s\<^sub>t V S \ wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) (S \\<^sub>s\<^sub>t \)" proof (induction S arbitrary: V rule: wf\<^sub>s\<^sub>t_induct) case (ConsRcv V ts S) hence "wf\<^sub>s\<^sub>t V S" "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ V" by simp_all hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) (S \\<^sub>s\<^sub>t \)" "fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using ConsRcv.IH subst_apply_fv_subset by (simp, force) thus ?case by simp next case (ConsSnd V ts S) hence "wf\<^sub>s\<^sub>t (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)) S" by simp hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)))) (S \\<^sub>s\<^sub>t \)" using ConsSnd.IH by metis hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>s\<^sub>e\<^sub>t (set ts))) (S \\<^sub>s\<^sub>t \)" by simp hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)) (S \\<^sub>s\<^sub>t \)" by (metis subst_apply_fv_unfold_set) thus ?case by simp next case (ConsEq V t t' S) hence "wf\<^sub>s\<^sub>t (V \ fv t) S" "fv t' \ V" by auto hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ fv t))) (S \\<^sub>s\<^sub>t \)" and *: "fv (t' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using ConsEq.IH subst_apply_fv_subset by force+ hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv (t \ \)) (S \\<^sub>s\<^sub>t \)" using subst_apply_fv_union by metis thus ?case using * by simp qed simp_all lemma wf_unify: assumes wf: "wf\<^sub>s\<^sub>t V (S@Send [Fun f X]#S')" and g_in_ik: "t \ ik\<^sub>s\<^sub>t S" and \: "Unifier \ (Fun f X) t" and disj: "bvars\<^sub>s\<^sub>t (S@Send [Fun f X]#S') \ (subst_domain \ \ range_vars \) = {}" shows "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) ((S@S') \\<^sub>s\<^sub>t \)" using assms proof (induction S' arbitrary: V rule: List.rev_induct) case (snoc x S' V) have fun_fv_bound: "fv (Fun f X \ \) \ \(set (map fv\<^sub>r\<^sub>c\<^sub>v (S \\<^sub>s\<^sub>t \)))" using snoc.prems(4) bvars\<^sub>s\<^sub>t_split Unifier_strand_fv_subset[OF g_in_ik \] by auto hence "fv (Fun f X \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \))" using fv_ik_is_fv_rcv by metis hence "fv (Fun f X \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" using fv_ik_subset_fv_st[of "S \\<^sub>s\<^sub>t \"] by blast hence *: "fv ((Fun f X) \ \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" by fastforce from snoc.prems(1) have "wf\<^sub>s\<^sub>t V (S@Send [Fun f X]#S')" using wf_prefix[of V "S@Send [Fun f X]#S'" "[x]"] by simp hence **: "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) ((S@S') \\<^sub>s\<^sub>t \)" using snoc.IH[OF _ snoc.prems(2,3)] snoc.prems(4) by auto from snoc.prems(1) have ***: "wf\<^sub>s\<^sub>t (V \ wfvarsoccs\<^sub>s\<^sub>t (S@Send [Fun f X]#S')) [x]" using wf_append_exec[of V "(S@Send [Fun f X]#S')" "[x]"] by simp from snoc.prems(4) have disj': "bvars\<^sub>s\<^sub>t (S@S') \ (subst_domain \ \ range_vars \) = {}" "set (bvars\<^sub>s\<^sub>t\<^sub>p x) \ (subst_domain \ \ range_vars \) = {}" by auto show ?case proof (cases x) case (Send t) thus ?thesis using wf_snd_append[of "fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" "(S@S') \\<^sub>s\<^sub>t \"] ** by auto next case (Receive t) hence "fv\<^sub>s\<^sub>t\<^sub>p x \ V \ wfvarsoccs\<^sub>s\<^sub>t (S@Send [Fun f X]#S')" using *** by auto hence "fv\<^sub>s\<^sub>t\<^sub>p x \ V \ wfrestrictedvars\<^sub>s\<^sub>t (S@Send [Fun f X]#S')" using vars_snd_rcv_strand_subset2(4)[of "S@Send [Fun f X]#S'"] by blast hence "fv\<^sub>s\<^sub>t\<^sub>p x \ V \ fv (Fun f X) \ wfrestrictedvars\<^sub>s\<^sub>t (S@S')" by auto hence "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv ((Fun f X) \ \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" by (metis (no_types) inf_sup_aci(5) subst_apply_fv_subset_strand2 subst_apply_fv_union disj') hence "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" using * by blast hence "fv\<^sub>s\<^sub>e\<^sub>t (set t \\<^sub>s\<^sub>e\<^sub>t \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) " using \x = Receive t\ by auto hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) (((S@S') \\<^sub>s\<^sub>t \)@[Receive (t \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)])" using wf_rcv_append'''[OF **, of "t \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \"] by simp thus ?thesis using \x = Receive t\ by auto next case (Equality ac s s') show ?thesis proof (cases ac) case Assign hence "fv s' \ V \ wfvarsoccs\<^sub>s\<^sub>t (S@Send [Fun f X]#S')" using Equality *** by auto hence "fv s' \ V \ wfrestrictedvars\<^sub>s\<^sub>t (S@Send [Fun f X]#S')" using vars_snd_rcv_strand_subset2(4)[of "S@Send [Fun f X]#S'"] by blast hence "fv s' \ V \ fv (Fun f X) \ wfrestrictedvars\<^sub>s\<^sub>t (S@S')" by auto moreover have "fv s' = fv_r\<^sub>e\<^sub>q ac x" "fv (s' \ \) = fv_r\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \)" using Equality by simp_all ultimately have "fv (s' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv (Fun f X \ \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" using subst_apply_fv_subset_strand2[of "fv\<^sub>e\<^sub>q ac" ac x] by (metis disj'(1) subst_apply_fv_subset_strand_trm2 subst_apply_fv_union sup_commute) hence "fv (s' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" using * by blast hence "fv (s' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using \x = Equality ac s s'\ by auto hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) (((S@S') \\<^sub>s\<^sub>t \)@[Equality ac (s \ \) (s' \ \)])" using wf_eq_append'''[OF **] by metis thus ?thesis using \x = Equality ac s s'\ by auto next case Check thus ?thesis using wf_eq_check_append''[OF **] Equality by simp qed next case (Inequality t t') thus ?thesis using wf_ineq_append''[OF **] by simp qed qed (auto dest: wf_subst_apply) lemma wf_equality: assumes wf: "wf\<^sub>s\<^sub>t V (S@Equality ac t t'#S')" and \: "mgu t t' = Some \" and disj: "bvars\<^sub>s\<^sub>t (S@Equality ac t t'#S') \ (subst_domain \ \ range_vars \) = {}" shows "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) ((S@S') \\<^sub>s\<^sub>t \)" using assms proof (induction S' arbitrary: V rule: List.rev_induct) case Nil thus ?case using wf_prefix[of V S "[Equality ac t t']"] wf_subst_apply[of V S \] by auto next case (snoc x S' V) show ?case proof (cases ac) case Assign hence "fv t' \ V \ wfvarsoccs\<^sub>s\<^sub>t S" using wf_eq_fv[of V, of S t t' "S'@[x]"] snoc by auto hence "fv t' \ V \ wfrestrictedvars\<^sub>s\<^sub>t S" using vars_snd_rcv_strand_subset2(4)[of S] by blast hence "fv t' \ V \ wfrestrictedvars\<^sub>s\<^sub>t (S@S')" by force moreover have disj': "bvars\<^sub>s\<^sub>t (S@S') \ (subst_domain \ \ range_vars \) = {}" "set (bvars\<^sub>s\<^sub>t\<^sub>p x) \ (subst_domain \ \ range_vars \) = {}" "bvars\<^sub>s\<^sub>t (S@Equality ac t t'#S') \ (subst_domain \ \ range_vars \) = {}" using snoc.prems(3) by auto ultimately have "fv (t' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" by (metis inf_sup_aci(5) subst_apply_fv_subset_strand_trm2) moreover have "fv (t \ \) = fv (t' \ \)" by (metis MGU_is_Unifier[OF mgu_gives_MGU[OF \]]) ultimately have *: "fv (t \ \) \ fv (t' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" by simp from snoc.prems(1) have "wf\<^sub>s\<^sub>t V (S@Equality ac t t'#S')" using wf_prefix[of V "S@Equality ac t t'#S'"] by simp hence **: "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) ((S@S') \\<^sub>s\<^sub>t \)" by (metis snoc.IH \ disj'(3)) from snoc.prems(1) have ***: "wf\<^sub>s\<^sub>t (V \ wfvarsoccs\<^sub>s\<^sub>t (S@Equality ac t t'#S')) [x]" using wf_append_exec[of V "(S@Equality ac t t'#S')" "[x]"] by simp show ?thesis proof (cases x) case (Send t) thus ?thesis using wf_snd_append[of "fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" "(S@S') \\<^sub>s\<^sub>t \"] ** by auto next case (Receive s) hence "fv\<^sub>s\<^sub>t\<^sub>p x \ V \ wfvarsoccs\<^sub>s\<^sub>t (S@Equality ac t t'#S')" using *** by auto hence "fv\<^sub>s\<^sub>t\<^sub>p x \ V \ wfrestrictedvars\<^sub>s\<^sub>t (S@Equality ac t t'#S')" using vars_snd_rcv_strand_subset2(4)[of "S@Equality ac t t'#S'"] by blast hence "fv\<^sub>s\<^sub>t\<^sub>p x \ V \ fv t \ fv t' \ wfrestrictedvars\<^sub>s\<^sub>t (S@S')" by (cases ac) auto hence "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv (t \ \) \ fv (t' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" using subst_apply_fv_subset_strand2[of fv\<^sub>s\<^sub>t\<^sub>p] by (metis (no_types) inf_sup_aci(5) subst_apply_fv_union disj'(1,2)) hence "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" when "ac = Assign" using * that by blast hence "fv\<^sub>s\<^sub>e\<^sub>t (set s \\<^sub>s\<^sub>e\<^sub>t \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V))" when "ac = Assign" using \x = Receive s\ that by auto hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) (((S@S') \\<^sub>s\<^sub>t \)@[Receive (s \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)])" when "ac = Assign" using wf_rcv_append'''[OF **, of "s \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \"] that by simp thus ?thesis using \x = Receive s\ Assign by auto next case (Equality ac' s s') show ?thesis proof (cases ac') case Assign hence "fv s' \ V \ wfvarsoccs\<^sub>s\<^sub>t (S@Equality ac t t'#S')" using *** Equality by auto hence "fv s' \ V \ wfrestrictedvars\<^sub>s\<^sub>t (S@Equality ac t t'#S')" using vars_snd_rcv_strand_subset2(4)[of "S@Equality ac t t'#S'"] by blast hence "fv s' \ V \ fv t \ fv t' \ wfrestrictedvars\<^sub>s\<^sub>t (S@S')" by (cases ac) auto moreover have "fv s' = fv_r\<^sub>e\<^sub>q ac' x" "fv (s' \ \) = fv_r\<^sub>e\<^sub>q ac' (x \\<^sub>s\<^sub>t\<^sub>p \)" using Equality by simp_all ultimately have "fv (s' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv (t \ \) \ fv (t' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" using subst_apply_fv_subset_strand2[of "fv_r\<^sub>e\<^sub>q ac'" ac' x] by (metis disj'(1) subst_apply_fv_subset_strand_trm2 subst_apply_fv_union sup_commute) hence "fv (s' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" using * \ac = Assign\ by blast hence ****: "fv (s' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using \x = Equality ac' s s'\ \ac = Assign\ by auto thus ?thesis using \x = Equality ac' s s'\ ** **** wf_eq_append' \ac = Assign\ by (metis (no_types, lifting) append.assoc append_Nil2 strand_step.case(3) strand_subst_hom subst_apply_strand_step_def) next case Check thus ?thesis using wf_eq_check_append''[OF **] Equality by simp qed next case (Inequality s s') thus ?thesis using wf_ineq_append''[OF **] by simp qed qed (metis snoc.prems(1) wf_eq_check_append wf_subst_apply) qed lemma wf_rcv_prefix_ground: "wf\<^sub>s\<^sub>t {} ((map Receive M)@S) \ vars\<^sub>s\<^sub>t (map Receive M) = {}" by (induct M) auto lemma simple_wfvarsoccs\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>n\<^sub>d: assumes "simple S" shows "wfvarsoccs\<^sub>s\<^sub>t S = \(set (map fv\<^sub>s\<^sub>n\<^sub>d S))" using assms unfolding simple_def proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp lemma wf\<^sub>s\<^sub>t_simple_induct[consumes 2, case_names Nil ConsSnd ConsRcv ConsIneq]: fixes S::"('a,'b) strand" assumes "wf\<^sub>s\<^sub>t V S" "simple S" "P []" "\v S. \wf\<^sub>s\<^sub>t V S; simple S; P S\ \ P (S@[Send [Var v]])" "\ts S. \wf\<^sub>s\<^sub>t V S; simple S; P S; fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ V \ \(set (map fv\<^sub>s\<^sub>n\<^sub>d S))\ \ P (S@[Receive ts])" "\X F S. \wf\<^sub>s\<^sub>t V S; simple S; P S\ \ P (S@[Inequality X F])" shows "P S" using assms proof (induction S rule: wf\<^sub>s\<^sub>t_induct') case (ConsSnd t S) hence "P S" by auto obtain v where "t = [Var v]" using simple_snd_is_var[OF _ \simple (S@[Send t])\] by auto thus ?case using ConsSnd.prems(3)[OF \wf\<^sub>s\<^sub>t V S\ _ \P S\] \simple (S@[Send t])\ by auto next case (ConsRcv t S) thus ?case using simple_wfvarsoccs\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>n\<^sub>d[of "S@[Receive t]"] by auto qed (auto simp add: simple_def) lemma wf_trm_stp_dom_fv_disjoint: "\wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \; t \ trms\<^sub>s\<^sub>t S\ \ subst_domain \ \ fv t = {}" unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by force lemma wf_constr_bvars_disj: "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \ \ (subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" unfolding range_vars_alt_def wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by fastforce lemma wf_constr_bvars_disj': assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t S" shows "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" (is ?A) and "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) = {}" (is ?B) proof - have "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" using assms(1) unfolding range_vars_alt_def wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by fastforce+ thus ?A and ?B using assms(2) bvars_subst_ident[of S \] by blast+ qed lemma (in intruder_model) wf_simple_strand_first_Send_var_split: assumes "wf\<^sub>s\<^sub>t {} S" "simple S" "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ = \ v" shows "\v S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f. S = S\<^sub>p\<^sub>r\<^sub>e@Send [Var v]#S\<^sub>s\<^sub>u\<^sub>f \ t \ \ = \ v \ \(\w \ wfrestrictedvars\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e. t \ \ = \ w)" (is "?P S") using assms proof (induction S rule: wf\<^sub>s\<^sub>t_simple_induct) case (ConsSnd v S) show ?case proof (cases "\w \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ = \ w") case True thus ?thesis using ConsSnd.IH by fastforce next case False thus ?thesis using ConsSnd.prems by auto qed next case (ConsRcv t' S) have "fv\<^sub>s\<^sub>e\<^sub>t (set t') \ wfrestrictedvars\<^sub>s\<^sub>t S" using ConsRcv.hyps(3) vars_snd_rcv_strand_subset2(1) by force hence "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ = \ v" using ConsRcv.prems(1) by fastforce hence "?P S" by (metis ConsRcv.IH) thus ?case by fastforce next case (ConsIneq X F S) moreover have "wfrestrictedvars\<^sub>s\<^sub>t (S @ [Inequality X F]) = wfrestrictedvars\<^sub>s\<^sub>t S" by auto ultimately have "?P S" by blast thus ?case by fastforce qed simp lemma (in intruder_model) wf_strand_first_Send_var_split: assumes "wf\<^sub>s\<^sub>t {} S" "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ \ \ v" shows "\S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f. \(\w \ wfrestrictedvars\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e. t \ \ \ \ w) \ ((\t'. S = S\<^sub>p\<^sub>r\<^sub>e@Send t'#S\<^sub>s\<^sub>u\<^sub>f \ t \ \ \\<^sub>s\<^sub>e\<^sub>t set t' \\<^sub>s\<^sub>e\<^sub>t \) \ (\t' t''. S = S\<^sub>p\<^sub>r\<^sub>e@Equality Assign t' t''#S\<^sub>s\<^sub>u\<^sub>f \ t \ \ \ t' \ \))" (is "\S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f. ?P S\<^sub>p\<^sub>r\<^sub>e \ ?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f") using assms proof (induction S rule: wf\<^sub>s\<^sub>t_induct') case (ConsSnd ts' S) show ?case proof (cases "\w \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ \ \ w") case True then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsSnd.IH by moura thus ?thesis by fastforce next case False then obtain v where v: "v \ fv\<^sub>s\<^sub>e\<^sub>t (set ts')" "t \ \ \ \ v" using ConsSnd.prems by auto then obtain t' where t': "t' \ set ts'" "v \ fv t'" by auto have "t \ \ \ t' \ \" using v(2) t'(2) subst_mono[of "Var v" t' \] vars_iff_subtermeq[of v] term.order_trans by auto hence "t \ \ \\<^sub>s\<^sub>e\<^sub>t set ts' \\<^sub>s\<^sub>e\<^sub>t \" using v(1) t'(1) by blast thus ?thesis using False v by auto qed next case (ConsRcv t' S) have "fv\<^sub>s\<^sub>e\<^sub>t (set t') \ wfrestrictedvars\<^sub>s\<^sub>t S" using ConsRcv.hyps vars_snd_rcv_strand_subset2(4)[of S] by blast hence "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ \ \ v" using ConsRcv.prems by fastforce then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsRcv.IH by moura thus ?case by fastforce next case (ConsEq s s' S) have *: "fv s' \ wfrestrictedvars\<^sub>s\<^sub>t S" using ConsEq.hyps vars_snd_rcv_strand_subset2(4)[of S] by blast show ?case proof (cases "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ \ \ v") case True then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsEq.IH by moura thus ?thesis by fastforce next case False then obtain v where "v \ fv s" "t \ \ \ \ v" using ConsEq.prems * by auto hence "t \ \ \ s \ \" using vars_iff_subtermeq[of v s] subst_mono[of "Var v" s \] term.order_trans by auto thus ?thesis using False by fastforce qed next case (ConsEq2 s s' S) have "wfrestrictedvars\<^sub>s\<^sub>t (S@[Equality Check s s']) = wfrestrictedvars\<^sub>s\<^sub>t S" by auto hence "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ \ \ v" using ConsEq2.prems by metis then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsEq2.IH by moura thus ?case by fastforce next case (ConsIneq X F S) hence "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ \ \ v" by fastforce then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsIneq.IH by moura thus ?case by fastforce qed simp subsection \Constraint Semantics\ context intruder_model begin subsubsection \Definitions\ text \The constraint semantics in which the intruder is limited to composition only\ fun strand_sem_c::"('fun,'var) terms \ ('fun,'var) strand \ ('fun,'var) subst \ bool" ("\_; _\\<^sub>c") where "\M; []\\<^sub>c = (\\. True)" | "\M; Send ts#S\\<^sub>c = (\\. (\t \ set ts. M \\<^sub>c t \ \) \ \M; S\\<^sub>c \)" | "\M; Receive ts#S\\<^sub>c = (\\. \(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; S\\<^sub>c \)" | "\M; Equality _ t t'#S\\<^sub>c = (\\. t \ \ = t' \ \ \ \M; S\\<^sub>c \)" | "\M; Inequality X F#S\\<^sub>c = (\\. ineq_model \ X F \ \M; S\\<^sub>c \)" definition constr_sem_c ("_ \\<^sub>c \_,_\") where "\ \\<^sub>c \S,\\ \ (\ supports \ \ \{}; S\\<^sub>c \)" abbreviation constr_sem_c' ("_ \\<^sub>c \_\" 90) where "\ \\<^sub>c \S\ \ \ \\<^sub>c \S,Var\" text \The full constraint semantics\ fun strand_sem_d::"('fun,'var) terms \ ('fun,'var) strand \ ('fun,'var) subst \ bool" ("\_; _\\<^sub>d") where "\M; []\\<^sub>d = (\\. True)" | "\M; Send ts#S\\<^sub>d = (\\. (\t \ set ts. M \ t \ \) \ \M; S\\<^sub>d \)" | "\M; Receive ts#S\\<^sub>d = (\\. \(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M; S\\<^sub>d \)" | "\M; Equality _ t t'#S\\<^sub>d = (\\. t \ \ = t' \ \ \ \M; S\\<^sub>d \)" | "\M; Inequality X F#S\\<^sub>d = (\\. ineq_model \ X F \ \M; S\\<^sub>d \)" definition constr_sem_d ("_ \ \_,_\") where "\ \ \S,\\ \ (\ supports \ \ \{}; S\\<^sub>d \)" abbreviation constr_sem_d' ("_ \ \_\" 90) where "\ \ \S\ \ \ \ \S,Var\" lemmas strand_sem_induct = strand_sem_c.induct[case_names Nil ConsSnd ConsRcv ConsEq ConsIneq] subsubsection \Lemmata\ lemma strand_sem_d_if_c: "\ \\<^sub>c \S,\\ \ \ \ \S,\\" proof - assume *: "\ \\<^sub>c \S,\\" { fix M have "\M; S\\<^sub>c \ \ \M; S\\<^sub>d \" proof (induction S rule: strand_sem_induct) case (ConsSnd M ts S) hence "\t \ set ts. M \\<^sub>c t \ \" "\M; S\\<^sub>d \" by auto thus ?case using strand_sem_d.simps(2)[of M _ S] by auto qed (auto simp add: ineq_model_def) } thus ?thesis using * by (simp add: constr_sem_c_def constr_sem_d_def) qed lemma strand_sem_mono_ik: "\M \ M'; \M; S\\<^sub>c \\ \ \M'; S\\<^sub>c \" (is "\?A'; ?A''\ \ ?A") "\M \ M'; \M; S\\<^sub>d \\ \ \M'; S\\<^sub>d \" (is "\?B'; ?B''\ \ ?B") proof - show "\?A'; ?A''\ \ ?A" proof (induction M S arbitrary: M M' rule: strand_sem_induct) case (ConsRcv M ts S) thus ?case using ConsRcv.IH[of "(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M" "(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M'"] by auto next case (ConsSnd M ts S) hence "\t \ set ts. M \\<^sub>c t \ \" "\M'; S\\<^sub>c \" by auto hence "\t \ set ts. M' \\<^sub>c t \ \" using ideduct_synth_mono \M \ M'\ by metis thus ?case using \\M'; S\\<^sub>c \\ by simp qed auto show "\?B'; ?B''\ \ ?B" proof (induction M S arbitrary: M M' rule: strand_sem_induct) case (ConsRcv M ts S) thus ?case using ConsRcv.IH[of "(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M" "(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M'"] by auto next case (ConsSnd M ts S) hence "\t \ set ts. M \ t \ \" "\M'; S\\<^sub>d \" by auto hence "\t \ set ts. M' \ t \ \" using ideduct_mono \M \ M'\ by metis thus ?case using \\M'; S\\<^sub>d \\ by simp qed auto qed context begin private lemma strand_sem_split_left: "\M; S@S'\\<^sub>c \ \ \M; S\\<^sub>c \" "\M; S@S'\\<^sub>d \ \ \M; S\\<^sub>d \" proof (induct S arbitrary: M) case (Cons x S) { case 1 thus ?case using Cons by (cases x) simp_all } { case 2 thus ?case using Cons by (cases x) simp_all } qed simp_all private lemma strand_sem_split_right: "\M; S@S'\\<^sub>c \ \ \M \ (ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \); S'\\<^sub>c \" "\M; S@S'\\<^sub>d \ \ \M \ (ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \); S'\\<^sub>d \" proof (induction S arbitrary: M rule: ik\<^sub>s\<^sub>t_induct) case (ConsRcv ts S) { case 1 thus ?case using ConsRcv.IH(1)[of "(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M"] by (simp add: Un_commute Un_left_commute image_Un) } { case 2 thus ?case using ConsRcv.IH(2)[of "(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M"] by (simp add: Un_commute Un_left_commute image_Un) } qed simp_all lemmas strand_sem_split[dest] = strand_sem_split_left(1) strand_sem_split_right(1) strand_sem_split_left(2) strand_sem_split_right(2) end lemma strand_sem_Send_split[dest]: "\\M; map Send T\\<^sub>c \; ts \ set T\ \ \M; [Send ts]\\<^sub>c \" (is "\?A'; ?A''\ \ ?A") "\\M; map Send T\\<^sub>d \; ts \ set T\ \ \M; [Send ts]\\<^sub>d \" (is "\?B'; ?B''\ \ ?B") "\\M; map Send T@S\\<^sub>c \; ts \ set T\ \ \M; Send ts#S\\<^sub>c \" (is "\?C'; ?C''\ \ ?C") "\\M; map Send T@S\\<^sub>d \; ts \ set T\ \ \M; Send ts#S\\<^sub>d \" (is "\?D'; ?D''\ \ ?D") "\\M; map Send1 T'\\<^sub>c \; t \ set T'\ \ \M; [Send1 t]\\<^sub>c \" (is "\?E'; ?E''\ \ ?E") "\\M; map Send1 T'\\<^sub>d \; t \ set T'\ \ \M; [Send1 t]\\<^sub>d \" (is "\?F'; ?F''\ \ ?F") "\\M; map Send1 T'@S\\<^sub>c \; t \ set T'\ \ \M; Send1 t#S\\<^sub>c \" (is "\?G'; ?G''\ \ ?G") "\\M; map Send1 T'@S\\<^sub>d \; t \ set T'\ \ \M; Send1 t#S\\<^sub>d \" (is "\?H'; ?H''\ \ ?H") proof - show A: "\?A'; ?A''\ \ ?A" by (induct "map Send T" arbitrary: T rule: strand_sem_c.induct) auto show B: "\?B'; ?B''\ \ ?B" by (induct "map Send T" arbitrary: T rule: strand_sem_d.induct) auto show "\?C'; ?C''\ \ ?C" "\?D'; ?D''\ \ ?D" using list.set_map list.simps(8) set_empty ik_snd_empty sup_bot.right_neutral by (metis (no_types, lifting) A strand_sem_split(1,2) strand_sem_c.simps(2), metis (no_types, lifting) B strand_sem_split(3,4) strand_sem_d.simps(2)) show "\?E'; ?E''\ \ ?E" by (induct "map Send1 T'" arbitrary: T' rule: strand_sem_c.induct) auto show "\?F'; ?F''\ \ ?F" by (induct "map Send1 T'" arbitrary: T' rule: strand_sem_c.induct) auto show "\?G'; ?G''\ \ ?G" proof (induction "map Send1 T'" arbitrary: T' rule: strand_sem_c.induct) case (2 M ts S') obtain t' T'' where ts: "ts = [t']" "T' = t'#T''" "S' = map Send1 T''" using "2.hyps"(2) by blast thus ?case using "2.prems" "2.hyps"(1) proof (cases "t = t'") case True have "ik\<^sub>s\<^sub>t (map Send1 T') \\<^sub>s\<^sub>e\<^sub>t \ = {}" by force hence "\M; [Send1 t]\\<^sub>c \" "\M; S\\<^sub>c \" using "2.prems"(1) unfolding ts(2) True by auto thus ?thesis by simp qed auto qed auto show "\?H'; ?H''\ \ ?H" proof (induction "map Send1 T'" arbitrary: T' rule: strand_sem_c.induct) case (2 M ts S') obtain t' T'' where ts: "ts = [t']" "T' = t'#T''" "S' = map Send1 T''" using "2.hyps"(2) by blast thus ?case using "2.prems" "2.hyps"(1) proof (cases "t = t'") case True have "ik\<^sub>s\<^sub>t (map Send1 T') \\<^sub>s\<^sub>e\<^sub>t \ = {}" by force hence "\M; [Send1 t]\\<^sub>d \" "\M; S\\<^sub>d \" using "2.prems"(1) unfolding ts(2) True by auto thus ?thesis by simp qed auto qed auto qed lemma strand_sem_Send_map: "(\ts. ts \ set T \ \M; [Send ts]\\<^sub>c \) \ \M; map Send T\\<^sub>c \" "(\ts. ts \ set T \ \M; [Send ts]\\<^sub>d \) \ \M; map Send T\\<^sub>d \" "(\t. t \ set T' \ \M; [Send1 t]\\<^sub>c \) \ \M; map Send1 T'\\<^sub>c \" "(\t. t \ set T' \ \M; [Send1 t]\\<^sub>d \) \ \M; map Send1 T'\\<^sub>d \" "\M; map Send1 T'\\<^sub>c \ \ \M; [Send T']\\<^sub>c \" "\M; map Send1 T'\\<^sub>d \ \ \M; [Send T']\\<^sub>d \" proof - show "(\ts. ts \ set T \ \M; [Send ts]\\<^sub>c \) \ \M; map Send T\\<^sub>c \" "(\ts. ts \ set T \ \M; [Send ts]\\<^sub>d \) \ \M; map Send T\\<^sub>d \" by (induct T) auto show "(\t. t \ set T' \ \M; [Send1 t]\\<^sub>c \) \ \M; map Send1 T'\\<^sub>c \" "(\t. t \ set T' \ \M; [Send1 t]\\<^sub>d \) \ \M; map Send1 T'\\<^sub>d \" by (induct T') auto show "\M; map Send1 T'\\<^sub>c \ \ \M; [Send T']\\<^sub>c \" "\M; map Send1 T'\\<^sub>d \ \ \M; [Send T']\\<^sub>d \" by (induct T') auto qed lemma strand_sem_Receive_map: "\M; map Receive T\\<^sub>c \" "\M; map Receive T\\<^sub>d \" "\M; map Receive1 T'\\<^sub>c \" "\M; map Receive1 T'\\<^sub>d \" "\M; [Receive T']\\<^sub>c \" "\M; [Receive T']\\<^sub>d \" proof - show "\M; map Receive T\\<^sub>c \" "\M; map Receive T\\<^sub>d \" by (induct T arbitrary: M) auto show "\M; map Receive1 T'\\<^sub>c \" "\M; map Receive1 T'\\<^sub>d \" by (induct T' arbitrary: M) auto show "\M; [Receive T']\\<^sub>c \" "\M; [Receive T']\\<^sub>d \" by (induct T' arbitrary: M) auto qed lemma strand_sem_append[intro]: "\\M; S\\<^sub>c \; \M \ (ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \); S'\\<^sub>c \\ \ \M; S@S'\\<^sub>c \" "\\M; S\\<^sub>d \; \M \ (ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \); S'\\<^sub>d \\ \ \M; S@S'\\<^sub>d \" proof (induction S arbitrary: M) case (Cons x S) { case 1 thus ?case using Cons proof (cases x) case (Receive ts) thus ?thesis using 1 Cons.IH(1)[of "(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M"] strand_sem_c.simps(3)[of M ts] image_Un[of "\t. t \ \" "set ts" "ik\<^sub>s\<^sub>t S"] by (metis (no_types, lifting) ik\<^sub>s\<^sub>t.simps(2) Un_assoc Un_commute append_Cons) qed auto } { case 2 thus ?case using Cons proof (cases x) case (Receive ts) thus ?thesis using 2 Cons.IH(2)[of "(set ts \\<^sub>s\<^sub>e\<^sub>t \) \ M"] strand_sem_d.simps(3)[of M ts] image_Un[of "\t. t \ \" "set ts" "ik\<^sub>s\<^sub>t S"] by (metis (no_types, lifting) ik\<^sub>s\<^sub>t.simps(2) Un_assoc Un_commute append_Cons) qed auto } qed simp_all lemma ineq_model_subst: fixes F::"(('a,'b) term \ ('a,'b) term) list" assumes "(subst_domain \ \ range_vars \) \ set X = {}" and "ineq_model (\ \\<^sub>s \) X F" shows "ineq_model \ X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" proof - { fix \::"('a,'b) subst" and t t' assume \: "subst_domain \ = set X" "ground (subst_range \)" and *: "\(s,t) \ set F. s \ (\ \\<^sub>s (\ \\<^sub>s \)) \ t \ (\ \\<^sub>s (\ \\<^sub>s \))" obtain f where f: "f \ set F" "fst f \ \ \\<^sub>s (\ \\<^sub>s \) \ snd f \ \ \\<^sub>s (\ \\<^sub>s \)" using * by (induct F) force+ have "\ \\<^sub>s (\ \\<^sub>s \) = \ \\<^sub>s (\ \\<^sub>s \)" by (metis (no_types, lifting) \ subst_compose_assoc assms(1) inf_sup_aci(1) subst_comp_eq_if_disjoint_vars sup_inf_absorb range_vars_alt_def) hence "(fst f \ \) \ \ \\<^sub>s \ \ (snd f \ \) \ \ \\<^sub>s \" using f by auto moreover have "(fst f \ \, snd f \ \) \ set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using f(1) by (auto simp add: subst_apply_pairs_def) ultimately have "\(s,t) \ set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \). s \ (\ \\<^sub>s \) \ t \ (\ \\<^sub>s \)" using f(1) Bex_set by fastforce } thus ?thesis using assms unfolding ineq_model_def by simp qed lemma ineq_model_subst': fixes F::"(('a,'b) term \ ('a,'b) term) list" assumes "(subst_domain \ \ range_vars \) \ set X = {}" and "ineq_model \ X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" shows "ineq_model (\ \\<^sub>s \) X F" proof - { fix \::"('a,'b) subst" and t t' assume \: "subst_domain \ = set X" "ground (subst_range \)" and *: "\(s,t) \ set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \). s \ (\ \\<^sub>s \) \ t \ (\ \\<^sub>s \)" obtain f where f: "f \ set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" "fst f \ \ \\<^sub>s \ \ snd f \ \ \\<^sub>s \" using * by (induct F) auto then obtain g where g: "g \ set F" "f = g \\<^sub>p \" by (auto simp add: subst_apply_pairs_def) have "\ \\<^sub>s (\ \\<^sub>s \) = \ \\<^sub>s (\ \\<^sub>s \)" by (metis (no_types, lifting) \ subst_compose_assoc assms(1) inf_sup_aci(1) subst_comp_eq_if_disjoint_vars sup_inf_absorb range_vars_alt_def) hence "fst g \ \ \\<^sub>s (\ \\<^sub>s \) \ snd g \ \ \\<^sub>s (\ \\<^sub>s \)" using f(2) g by (simp add: prod.case_eq_if) hence "\(s,t) \ set F. s \ (\ \\<^sub>s (\ \\<^sub>s \)) \ t \ (\ \\<^sub>s (\ \\<^sub>s \))" using g Bex_set by fastforce } thus ?thesis using assms unfolding ineq_model_def by simp qed lemma ineq_model_ground_subst: fixes F::"(('a,'b) term \ ('a,'b) term) list" assumes "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ subst_domain \" and "ground (subst_range \)" and "ineq_model \ X F" shows "ineq_model (\ \\<^sub>s \) X F" proof - { fix \::"('a,'b) subst" and t t' assume \: "subst_domain \ = set X" "ground (subst_range \)" and *: "\(s,t) \ set F. s \ (\ \\<^sub>s \) \ t \ (\ \\<^sub>s \ )" obtain f where f: "f \ set F" "fst f \ \ \\<^sub>s \ \ snd f \ \ \\<^sub>s \" using * by (induct F) force+ hence "fv (fst f) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" "fv (snd f) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" by auto hence "fv (fst f) - set X \ subst_domain \" "fv (snd f) - set X \ subst_domain \" using assms(1) by auto hence "fv (fst f \ \) \ subst_domain \" "fv (snd f \ \) \ subst_domain \" using \ by (simp_all add: range_vars_alt_def subst_fv_unfold_ground_img) hence "fv (fst f \ \ \\<^sub>s \) = {}" "fv (snd f \ \ \\<^sub>s \) = {}" using assms(2) by (simp_all add: subst_fv_dom_ground_if_ground_img) hence "fst f \ \ \\<^sub>s (\ \\<^sub>s \) \ snd f \ \ \\<^sub>s (\ \\<^sub>s \)" using f(2) subst_ground_ident by fastforce hence "\(s,t) \ set F. s \ (\ \\<^sub>s (\ \\<^sub>s \)) \ t \ (\ \\<^sub>s (\ \\<^sub>s \))" using f(1) Bex_set by fastforce } thus ?thesis using assms unfolding ineq_model_def by simp qed context begin private lemma strand_sem_subst_c: assumes "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "\M; S\\<^sub>c (\ \\<^sub>s \) \ \M; S \\<^sub>s\<^sub>t \\\<^sub>c \" using assms proof (induction S arbitrary: \ M rule: strand_sem_induct) case (ConsSnd M ts S) hence "\M; S \\<^sub>s\<^sub>t \\\<^sub>c \" "\t \ set ts. M \\<^sub>c t \ (\ \\<^sub>s \)" by auto hence "\t \ set ts. M \\<^sub>c (t \ \) \ \" using subst_comp_all[of \ \ M] subst_subst_compose[of _ \ \] by simp hence "\t \ set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \). M \\<^sub>c t \ \" by fastforce thus ?case using \\M; S \\<^sub>s\<^sub>t \\\<^sub>c \\ unfolding subst_apply_strand_def by simp next case (ConsRcv M ts S) have *: "\(set ts \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \) \ M; S\\<^sub>c (\ \\<^sub>s \)" using ConsRcv.prems(1) by simp have "bvars\<^sub>s\<^sub>t (Receive ts#S) = bvars\<^sub>s\<^sub>t S" by auto hence **: "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" using ConsRcv.prems(2) by blast have "\M; Receive (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)#(S \\<^sub>s\<^sub>t \)\\<^sub>c \" using ConsRcv.IH[OF * **] by (metis (no_types) image_set strand_sem_c.simps(3) subst_comp_all) thus ?case by simp next case (ConsIneq M X F S) hence *: "\M; S \\<^sub>s\<^sub>t \\\<^sub>c \" and ***: "(subst_domain \ \ range_vars \) \ set X = {}" unfolding bvars\<^sub>s\<^sub>t_def ineq_model_def by auto have **: "ineq_model (\ \\<^sub>s \) X F" using ConsIneq by (auto simp add: subst_compose_assoc ineq_model_def) have "\\. subst_domain \ = set X \ ground (subst_range \) \ (subst_domain \ \ range_vars \) \ (subst_domain \ \ range_vars \) = {}" using * ** *** unfolding range_vars_alt_def by auto hence "\\. subst_domain \ = set X \ ground (subst_range \) \ \ \\<^sub>s \ = \ \\<^sub>s \" by (metis subst_comp_eq_if_disjoint_vars) hence "ineq_model \ X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using ineq_model_subst[OF *** **] by blast moreover have "rm_vars (set X) \ = \" using ConsIneq.prems(2) by force ultimately show ?case using * by auto qed simp_all private lemma strand_sem_subst_c': assumes "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "\M; S \\<^sub>s\<^sub>t \\\<^sub>c \ \ \M; S\\<^sub>c (\ \\<^sub>s \)" using assms proof (induction S arbitrary: \ M rule: strand_sem_induct) case (ConsSnd M ts S) hence "\M; [Send ts] \\<^sub>s\<^sub>t \\\<^sub>c \" "\M; S \\<^sub>s\<^sub>t \\\<^sub>c \" by auto hence "\M; S\\<^sub>c (\ \\<^sub>s \)" using ConsSnd.IH[OF _] ConsSnd.prems(2) by auto moreover have "\M; [Send ts]\\<^sub>c (\ \\<^sub>s \)" proof - have "\M; [send\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\<^sub>s\<^sub>t]\\<^sub>c \" using \\M; [Send ts] \\<^sub>s\<^sub>t \\\<^sub>c \\ by simp hence "\t \ set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \). M \\<^sub>c t \ \" by simp hence "\t \ set ts. M \\<^sub>c t \ \ \ \" by auto hence "\t \ set ts. M \\<^sub>c t \ (\ \\<^sub>s \)" using subst_subst_compose by metis thus "\M; [Send ts]\\<^sub>c (\ \\<^sub>s \)" by auto qed ultimately show ?case by auto next case (ConsRcv M ts S) hence "\((set ts \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) \ M); S \\<^sub>s\<^sub>t \\\<^sub>c \" by (simp add: subst_all_insert) hence "\((set ts \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \) \ M); S \\<^sub>s\<^sub>t \\\<^sub>c \" by (metis subst_comp_all) thus ?case using ConsRcv.IH ConsRcv.prems(2) by auto next case (ConsIneq M X F S) have \: "rm_vars (set X) \ = \" using ConsIneq.prems(2) by force hence *: "\M; S\\<^sub>c (\ \\<^sub>s \)" and ***: "(subst_domain \ \ range_vars \) \ set X = {}" using ConsIneq unfolding bvars\<^sub>s\<^sub>t_def ineq_model_def by auto have **: "ineq_model \ X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using ConsIneq.prems(1) \ by (auto simp add: subst_compose_assoc ineq_model_def) have "\\. subst_domain \ = set X \ ground (subst_range \) \ (subst_domain \ \ range_vars \) \ (subst_domain \ \ range_vars \) = {}" using * ** *** unfolding range_vars_alt_def by auto hence "\\. subst_domain \ = set X \ ground (subst_range \) \ \ \\<^sub>s \ = \ \\<^sub>s \" by (metis subst_comp_eq_if_disjoint_vars) hence "ineq_model (\ \\<^sub>s \) X F" using ineq_model_subst'[OF *** **] by blast thus ?case using * by auto next case ConsEq thus ?case unfolding bvars\<^sub>s\<^sub>t_def by auto qed simp_all private lemma strand_sem_subst_d: assumes "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "\M; S\\<^sub>d (\ \\<^sub>s \) \ \M; S \\<^sub>s\<^sub>t \\\<^sub>d \" using assms proof (induction S arbitrary: \ M rule: strand_sem_induct) case (ConsSnd M ts S) hence "\M; S \\<^sub>s\<^sub>t \\\<^sub>d \" "\t \ set ts. M \ t \ (\ \\<^sub>s \)" by auto hence "\t \ set ts. M \ (t \ \) \ \" using subst_comp_all[of \ \ M] subst_subst_compose[of _ \ \] by simp hence "\t \ set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \). M \ t \ \" by simp thus ?case using \\M; S \\<^sub>s\<^sub>t \\\<^sub>d \\ by simp next case (ConsRcv M ts S) have "\(set ts \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \) \ M; S\\<^sub>d (\ \\<^sub>s \)" using ConsRcv.prems(1) by simp hence *: "\(set ts \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) \ M; S\\<^sub>d (\ \\<^sub>s \)" by (metis subst_comp_all) have "bvars\<^sub>s\<^sub>t (Receive ts#S) = bvars\<^sub>s\<^sub>t S" by auto hence **: "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" using ConsRcv.prems(2) by blast have "\M; Receive (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)#(S \\<^sub>s\<^sub>t \)\\<^sub>d \" using ConsRcv.IH[OF * **] by simp thus ?case by simp next case (ConsIneq M X F S) hence *: "\M; S \\<^sub>s\<^sub>t \\\<^sub>d \" and ***: "(subst_domain \ \ range_vars \) \ set X = {}" unfolding bvars\<^sub>s\<^sub>t_def ineq_model_def by auto have **: "ineq_model (\ \\<^sub>s \) X F" using ConsIneq by (auto simp add: subst_compose_assoc ineq_model_def) have "\\. subst_domain \ = set X \ ground (subst_range \) \ (subst_domain \ \ range_vars \) \ (subst_domain \ \ range_vars \) = {}" using * ** *** unfolding range_vars_alt_def by auto hence "\\. subst_domain \ = set X \ ground (subst_range \) \ \ \\<^sub>s \ = \ \\<^sub>s \" by (metis subst_comp_eq_if_disjoint_vars) hence "ineq_model \ X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using ineq_model_subst[OF *** **] by blast moreover have "rm_vars (set X) \ = \" using ConsIneq.prems(2) by force ultimately show ?case using * by auto next case ConsEq thus ?case unfolding bvars\<^sub>s\<^sub>t_def by auto qed simp_all private lemma strand_sem_subst_d': assumes "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "\M; S \\<^sub>s\<^sub>t \\\<^sub>d \ \ \M; S\\<^sub>d (\ \\<^sub>s \)" using assms proof (induction S arbitrary: \ M rule: strand_sem_induct) case (ConsSnd M ts S) hence "\M; [Send ts] \\<^sub>s\<^sub>t \\\<^sub>d \" "\M; S \\<^sub>s\<^sub>t \\\<^sub>d \" by auto hence "\M; S\\<^sub>d (\ \\<^sub>s \)" using ConsSnd.IH[OF _] ConsSnd.prems(2) by auto moreover have "\M; [Send ts]\\<^sub>d (\ \\<^sub>s \)" proof - have "\M; [send\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\<^sub>s\<^sub>t]\\<^sub>d \" using \\M; [Send ts] \\<^sub>s\<^sub>t \\\<^sub>d \\ by simp hence "\t \ set (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \). M \ t \ \" by simp hence "\t \ set ts. M \ t \ \ \ \" by auto hence "\t \ set ts. M \ t \ (\ \\<^sub>s \)" using subst_subst_compose by metis thus "\M; [Send ts]\\<^sub>d (\ \\<^sub>s \)" by auto qed ultimately show ?case by auto next case (ConsRcv M ts S) hence "\((set ts \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) \ M); S \\<^sub>s\<^sub>t \\\<^sub>d \" by (simp add: subst_all_insert) hence "\((set ts \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \) \ M); S \\<^sub>s\<^sub>t \\\<^sub>d \" by (metis subst_comp_all) thus ?case using ConsRcv.IH ConsRcv.prems(2) by auto next case (ConsIneq M X F S) have \: "rm_vars (set X) \ = \" using ConsIneq.prems(2) by force hence *: "\M; S\\<^sub>d (\ \\<^sub>s \)" and ***: "(subst_domain \ \ range_vars \) \ set X = {}" using ConsIneq unfolding bvars\<^sub>s\<^sub>t_def ineq_model_def by auto have **: "ineq_model \ X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using ConsIneq.prems(1) \ by (auto simp add: subst_compose_assoc ineq_model_def) have "\\. subst_domain \ = set X \ ground (subst_range \) \ (subst_domain \ \ range_vars \) \ (subst_domain \ \ range_vars \) = {}" using * ** *** unfolding range_vars_alt_def by auto hence "\\. subst_domain \ = set X \ ground (subst_range \) \ \ \\<^sub>s \ = \ \\<^sub>s \" by (metis subst_comp_eq_if_disjoint_vars) hence "ineq_model (\ \\<^sub>s \) X F" using ineq_model_subst'[OF *** **] by blast thus ?case using * by auto next case ConsEq thus ?case unfolding bvars\<^sub>s\<^sub>t_def by auto qed simp_all lemmas strand_sem_subst = strand_sem_subst_c strand_sem_subst_c' strand_sem_subst_d strand_sem_subst_d' end lemma strand_sem_subst_subst_idem: assumes \: "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "\\M; S \\<^sub>s\<^sub>t \\\<^sub>c (\ \\<^sub>s \); subst_idem \\ \ \M; S\\<^sub>c (\ \\<^sub>s \)" using strand_sem_subst(2)[OF assms, of M "\ \\<^sub>s \"] subst_compose_assoc[of \ \ \] unfolding subst_idem_def by argo lemma strand_sem_subst_comp: assumes "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" and "\M; S\\<^sub>c \" "subst_domain \ \ (vars\<^sub>s\<^sub>t S \ fv\<^sub>s\<^sub>e\<^sub>t M) = {}" shows "\M; S\\<^sub>c (\ \\<^sub>s \)" proof - from assms(3) have "subst_domain \ \ vars\<^sub>s\<^sub>t S = {}" "subst_domain \ \ fv\<^sub>s\<^sub>e\<^sub>t M = {}" by auto hence "S \\<^sub>s\<^sub>t \ = S" "M \\<^sub>s\<^sub>e\<^sub>t \ = M" using strand_substI set_subst_ident[of M \] by (blast, blast) thus ?thesis using assms(2) by (auto simp add: strand_sem_subst(2)[OF assms(1)]) qed lemma strand_sem_c_imp_ineqs_neq: assumes "\M; S\\<^sub>c \" "Inequality X [(t,t')] \ set S" shows "t \ t' \ (\\. subst_domain \ = set X \ ground (subst_range \) \ t \ \ \ t' \ \ \ t \ \ \ \ \ t' \ \ \ \)" using assms proof (induction rule: strand_sem_induct) case (ConsIneq M Y F S) thus ?case proof (cases "Inequality X [(t,t')] \ set S") case False hence "X = Y" "F = [(t,t')]" using ConsIneq by auto hence *: "\\. subst_domain \ = set X \ ground (subst_range \) \ t \ \ \ \ \ t' \ \ \ \" using ConsIneq by (auto simp add: ineq_model_def) then obtain \ where \: "subst_domain \ = set X" "ground (subst_range \)" "t \ \ \ \ \ t' \ \ \ \" using interpretation_subst_exists'[of "set X"] by moura hence "t \ t'" by auto moreover have "\\ \. t \ \ \ \ \ t' \ \ \ \ \ t \ \ \ t' \ \" by auto ultimately show ?thesis using * by auto qed simp qed simp_all lemma strand_sem_c_imp_ineq_model: assumes "\M; S\\<^sub>c \" "Inequality X F \ set S" shows "ineq_model \ X F" using assms by (induct S rule: strand_sem_induct) force+ lemma strand_sem_wf_simple_fv_sat: assumes "wf\<^sub>s\<^sub>t {} S" "simple S" "\{}; S\\<^sub>c \" shows "\v. v \ wfrestrictedvars\<^sub>s\<^sub>t S \ ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c \ v" using assms proof (induction S rule: wf\<^sub>s\<^sub>t_simple_induct) case (ConsRcv t S) have "v \ wfrestrictedvars\<^sub>s\<^sub>t S" using ConsRcv.hyps(3) ConsRcv.prems(1) vars_snd_rcv_strand2 by fastforce moreover have "\{}; S\\<^sub>c \" using \\{}; S@[Receive t]\\<^sub>c \\ by blast moreover have "ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \ ik\<^sub>s\<^sub>t (S@[Receive t]) \\<^sub>s\<^sub>e\<^sub>t \" by auto ultimately show ?case using ConsRcv.IH ideduct_synth_mono by meson next case (ConsIneq X F S) hence "v \ wfrestrictedvars\<^sub>s\<^sub>t S" by fastforce moreover have "\{}; S\\<^sub>c \" using \\{}; S@[Inequality X F]\\<^sub>c \\ by blast moreover have "ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \ ik\<^sub>s\<^sub>t (S@[Inequality X F]) \\<^sub>s\<^sub>e\<^sub>t \" by auto ultimately show ?case using ConsIneq.IH ideduct_synth_mono by meson next case (ConsSnd w S) hence *: "\{}; S\\<^sub>c \" "ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c \ w" by auto have **: "ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \ ik\<^sub>s\<^sub>t (S@[Send [Var w]]) \\<^sub>s\<^sub>e\<^sub>t \" by simp show ?case proof (cases "v = w") case True thus ?thesis using *(2) ideduct_synth_mono[OF _ **] by meson next case False hence "v \ wfrestrictedvars\<^sub>s\<^sub>t S" using ConsSnd.prems(1) by auto thus ?thesis using ConsSnd.IH[OF _ *(1)] ideduct_synth_mono[OF _ **] by metis qed qed simp lemma strand_sem_wf_ik_or_assignment_rhs_fun_subterm: assumes "wf\<^sub>s\<^sub>t {} A" "\{}; A\\<^sub>c \" "Var x \ ik\<^sub>s\<^sub>t A" "\ x = Fun f T" "t\<^sub>i \ set T" "\ik\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t\<^sub>i" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" obtains S where "Fun f S \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t A) \ Fun f S \ subterms\<^sub>s\<^sub>e\<^sub>t (assignment_rhs\<^sub>s\<^sub>t A)" "Fun f T = Fun f S \ \" proof - have "x \ wfrestrictedvars\<^sub>s\<^sub>t A" by (metis (no_types) assms(3) set_rev_mp term.set_intros(3) vars_subset_if_in_strand_ik2) moreover have "Fun f T \ \ = Fun f T" by (metis subst_ground_ident interpretation_grounds_all assms(4,7)) ultimately obtain A\<^sub>p\<^sub>r\<^sub>e A\<^sub>s\<^sub>u\<^sub>f where *: "\(\w \ wfrestrictedvars\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e. Fun f T \ \ w)" "(\t. A = A\<^sub>p\<^sub>r\<^sub>e@Send t#A\<^sub>s\<^sub>u\<^sub>f \ Fun f T \\<^sub>s\<^sub>e\<^sub>t set t \\<^sub>s\<^sub>e\<^sub>t \) \ (\t t'. A = A\<^sub>p\<^sub>r\<^sub>e@Equality Assign t t'#A\<^sub>s\<^sub>u\<^sub>f \ Fun f T \ t \ \)" using wf_strand_first_Send_var_split[OF assms(1)] assms(4) subtermeqI' by metis moreover { fix ts assume **: "A = A\<^sub>p\<^sub>r\<^sub>e@Send ts#A\<^sub>s\<^sub>u\<^sub>f" "Fun f T \\<^sub>s\<^sub>e\<^sub>t set ts \\<^sub>s\<^sub>e\<^sub>t \" hence ***: "\t \ set ts. ik\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" "\ik\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t\<^sub>i" using assms(2,6) by (auto intro: ideduct_synth_mono) then obtain t where t: "t \ set ts" "Fun f T \ t \ \" "ik\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" using **(2) by blast obtain s where s: "s \ ik\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e" "Fun f T \ s \ \" using t(3,2) ***(2) assms(5) by (induct rule: intruder_synth_induct) auto then obtain g S where gS: "Fun g S \ s" "Fun f T = Fun g S \ \" using subterm_subst_not_img_subterm[OF s(2)] *(1) by force hence ?thesis using that **(1) s(1) by force } moreover { fix t t' assume **: "A = A\<^sub>p\<^sub>r\<^sub>e@Equality Assign t t'#A\<^sub>s\<^sub>u\<^sub>f" "Fun f T \ t \ \" with assms(2) have "t \ \ = t' \ \" by auto hence "Fun f T \ t' \ \" using **(2) by auto from assms(1) **(1) have "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e" using wf_eq_fv[of "{}" A\<^sub>p\<^sub>r\<^sub>e t t' A\<^sub>s\<^sub>u\<^sub>f] vars_snd_rcv_strand_subset2(4)[of A\<^sub>p\<^sub>r\<^sub>e] by blast then obtain g S where gS: "Fun g S \ t'" "Fun f T = Fun g S \ \" using subterm_subst_not_img_subterm[OF \Fun f T \ t' \ \\] *(1) by fastforce hence ?thesis using that **(1) by auto } ultimately show ?thesis by auto qed lemma ineq_model_not_unif_is_sat_ineq: assumes "\\. Unifier \ t t'" shows "ineq_model \ X [(t, t')]" using assms list.set_intros(1)[of "(t,t')" "[]"] unfolding ineq_model_def by blast lemma strand_sem_not_unif_is_sat_ineq: assumes "\\. Unifier \ t t'" shows "\M; [Inequality X [(t,t')]]\\<^sub>c \" "\M; [Inequality X [(t,t')]]\\<^sub>d \" using ineq_model_not_unif_is_sat_ineq[OF assms] strand_sem_c.simps(1,5)[of M] strand_sem_d.simps(1,5)[of M] by presburger+ lemma ineq_model_singleI[intro]: assumes "\\. subst_domain \ = set X \ ground (subst_range \) \ t \ \ \ \ \ t' \ \ \ \" shows "ineq_model \ X [(t,t')]" using assms unfolding ineq_model_def by auto lemma ineq_model_singleE: assumes "ineq_model \ X [(t,t')]" shows "\\. subst_domain \ = set X \ ground (subst_range \) \ t \ \ \ \ \ t' \ \ \ \" using assms unfolding ineq_model_def by auto lemma ineq_model_single_iff: fixes F::"(('a,'b) term \ ('a,'b) term) list" shows "ineq_model \ X F \ ineq_model \ X [(Fun f (Fun c []#map fst F),Fun f (Fun c []#map snd F))]" (is "?A \ ?B") proof - let ?P = "\\ f. fst f \ (\ \\<^sub>s \) \ snd f \ (\ \\<^sub>s \)" let ?Q = "\\ t t'. t \ (\ \\<^sub>s \) \ t' \ (\ \\<^sub>s \)" let ?T = "\g. Fun c []#map g F" let ?S = "\\ g. map (\x. x \ (\ \\<^sub>s \)) (Fun c []#map g F)" let ?t = "Fun f (?T fst)" let ?t' = "Fun f (?T snd)" have len: "\g h. length (?T g) = length (?T h)" "\g h \. length (?S \ g) = length (?T h)" "\g h \. length (?S \ g) = length (?T h)" "\g h \ \. length (?S \ g) = length (?S \ h)" by simp_all { fix \::"('a,'b) subst" assume \: "subst_domain \ = set X" "ground (subst_range \)" have "list_ex (?P \) F \ ?Q \ ?t ?t'" proof assume "list_ex (?P \) F" then obtain a where a: "a \ set F" "?P \ a" by (metis (mono_tags, lifting) Bex_set) thus "?Q \ ?t ?t'" by auto qed (fastforce simp add: Bex_set) } thus ?thesis unfolding ineq_model_def case_prod_unfold by auto qed subsection \Constraint Semantics (Alternative, Equivalent Version)\ text \These are the constraint semantics used in the CSF 2017 paper\ fun strand_sem_c'::"('fun,'var) terms \ ('fun,'var) strand \ ('fun,'var) subst \ bool" ("\_; _\\<^sub>c''") where "\M; []\\<^sub>c' = (\\. True)" | "\M; Send ts#S\\<^sub>c' = (\\. (\t \ set ts. M \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \) \ \M; S\\<^sub>c' \)" | "\M; Receive ts#S\\<^sub>c' = \set ts \ M; S\\<^sub>c'" | "\M; Equality _ t t'#S\\<^sub>c' = (\\. t \ \ = t' \ \ \ \M; S\\<^sub>c' \)" | "\M; Inequality X F#S\\<^sub>c' = (\\. ineq_model \ X F \ \M; S\\<^sub>c' \)" fun strand_sem_d'::"('fun,'var) terms \ ('fun,'var) strand \ ('fun,'var) subst \ bool" ("\_; _\\<^sub>d''") where "\M; []\\<^sub>d' = (\\. True)" | "\M; Send ts#S\\<^sub>d' = (\\. (\t \ set ts. M \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \) \ \M; S\\<^sub>d' \)" | "\M; Receive ts#S\\<^sub>d' = \set ts \ M; S\\<^sub>d'" | "\M; Equality _ t t'#S\\<^sub>d' = (\\. t \ \ = t' \ \ \ \M; S\\<^sub>d' \)" | "\M; Inequality X F#S\\<^sub>d' = (\\. ineq_model \ X F \ \M; S\\<^sub>d' \)" lemma strand_sem_eq_defs: "\M; \\\<^sub>c' \ = \M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>c \" "\M; \\\<^sub>d' \ = \M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>d \" proof - have 1: "\M; \\\<^sub>c' \ \ \M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>c \" proof (induction \ arbitrary: M rule: strand_sem_induct) case (ConsRcv M ts S) thus ?case by (fastforce simp add: image_Un[of "\t. t \ \"]) qed simp_all have 2: "\M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>c \ \ \M; \\\<^sub>c' \" proof (induction \ arbitrary: M rule: strand_sem_c'.induct) case (3 M ts S) thus ?case by (fastforce simp add: image_Un[of "\t. t \ \"]) qed simp_all have 3: "\M; \\\<^sub>d' \ \ \M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>d \" proof (induction \ arbitrary: M rule: strand_sem_induct) case (ConsRcv M ts S) thus ?case by (fastforce simp add: image_Un[of "\t. t \ \"]) qed simp_all have 4: "\M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>d \ \ \M; \\\<^sub>d' \" proof (induction \ arbitrary: M rule: strand_sem_d'.induct) case (3 M ts S) thus ?case by (fastforce simp add: image_Un[of "\t. t \ \"]) qed simp_all show "\M; \\\<^sub>c' \ = \M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>c \" "\M; \\\<^sub>d' \ = \M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>d \" by (metis 1 2, metis 3 4) qed lemma strand_sem_split'[dest]: "\M; S@S'\\<^sub>c' \ \ \M; S\\<^sub>c' \" "\M; S@S'\\<^sub>c' \ \ \M \ ik\<^sub>s\<^sub>t S; S'\\<^sub>c' \" "\M; S@S'\\<^sub>d' \ \ \M; S\\<^sub>d' \" "\M; S@S'\\<^sub>d' \ \ \M \ ik\<^sub>s\<^sub>t S; S'\\<^sub>d' \" using strand_sem_eq_defs[of M "S@S'" \] strand_sem_eq_defs[of M S \] strand_sem_eq_defs[of "M \ ik\<^sub>s\<^sub>t S" S' \] strand_sem_split(2,4) by (auto simp add: image_Un) lemma strand_sem_append'[intro]: "\M; S\\<^sub>c' \ \ \M \ ik\<^sub>s\<^sub>t S; S'\\<^sub>c' \ \ \M; S@S'\\<^sub>c' \" "\M; S\\<^sub>d' \ \ \M \ ik\<^sub>s\<^sub>t S; S'\\<^sub>d' \ \ \M; S@S'\\<^sub>d' \" using strand_sem_eq_defs[of M "S@S'" \] strand_sem_eq_defs[of M S \] strand_sem_eq_defs[of "M \ ik\<^sub>s\<^sub>t S" S' \] by (auto simp add: image_Un) end subsection \Dual Strands\ fun dual\<^sub>s\<^sub>t::"('a,'b) strand \ ('a,'b) strand" where "dual\<^sub>s\<^sub>t [] = []" | "dual\<^sub>s\<^sub>t (Receive t#S) = Send t#(dual\<^sub>s\<^sub>t S)" | "dual\<^sub>s\<^sub>t (Send t#S) = Receive t#(dual\<^sub>s\<^sub>t S)" | "dual\<^sub>s\<^sub>t (x#S) = x#(dual\<^sub>s\<^sub>t S)" lemma dual\<^sub>s\<^sub>t_append: "dual\<^sub>s\<^sub>t (A@B) = (dual\<^sub>s\<^sub>t A)@(dual\<^sub>s\<^sub>t B)" by (induct A rule: dual\<^sub>s\<^sub>t.induct) auto lemma dual\<^sub>s\<^sub>t_self_inverse: "dual\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t S) = S" proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp lemma dual\<^sub>s\<^sub>t_trms_eq: "trms\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t S) = trms\<^sub>s\<^sub>t S" proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp lemma dual\<^sub>s\<^sub>t_fv: "fv\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t A) = fv\<^sub>s\<^sub>t A" by (induct A rule: dual\<^sub>s\<^sub>t.induct) auto lemma dual\<^sub>s\<^sub>t_bvars: "bvars\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t A) = bvars\<^sub>s\<^sub>t A" by (induct A rule: dual\<^sub>s\<^sub>t.induct) fastforce+ end diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Typed_Model.thy b/thys/Stateful_Protocol_Composition_and_Typing/Typed_Model.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/Typed_Model.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/Typed_Model.thy @@ -1,2208 +1,2208 @@ (* Title: Typed_Model.thy Author: Andreas Viktor Hess, DTU SPDX-License-Identifier: BSD-3-Clause *) section \The Typed Model\ theory Typed_Model imports Lazy_Intruder begin text \Term types\ type_synonym ('f,'v) term_type = "('f,'v) term" text \Constructors for term types\ abbreviation (input) TAtom::"'v \ ('f,'v) term_type" where "TAtom a \ Var a" abbreviation (input) TComp::"['f, ('f,'v) term_type list] \ ('f,'v) term_type" where "TComp f ts \ Fun f ts" text \ The typed model extends the intruder model with a typing function \\\ that assigns types to terms. \ locale typed_model = intruder_model arity public Ana for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,'var) term \ (('fun,'var) term list \ ('fun,'var) term list)" + fixes \::"('fun,'var) term \ ('fun,'atom::finite) term_type" assumes const_type: "\c. arity c = 0 \ \a. \ts. \ (Fun c ts) = TAtom a" and fun_type: "\f ts. arity f > 0 \ \ (Fun f ts) = TComp f (map \ ts)" and \_wf: "\x f ts. TComp f ts \ \ (Var x) \ arity f > 0" "\x. wf\<^sub>t\<^sub>r\<^sub>m (\ (Var x))" begin subsection \Definitions\ text \The set of atomic types\ abbreviation "\\<^sub>a \ UNIV::('atom set)" text \Well-typed substitutions\ definition wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t where "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ (\v. \ (Var v) = \ (\ v))" text \The set of sub-message patterns (SMP)\ inductive_set SMP::"('fun,'var) terms \ ('fun,'var) terms" for M where MP[intro]: "t \ M \ t \ SMP M" | Subterm[intro]: "\t \ SMP M; t' \ t\ \ t' \ SMP M" | Substitution[intro]: "\t \ SMP M; wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \; wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)\ \ (t \ \) \ SMP M" | Ana[intro]: "\t \ SMP M; Ana t = (K,T); k \ set K\ \ k \ SMP M" text \ Type-flaw resistance for sets: Unifiable sub-message patterns must have the same type (unless they are variables) \ definition tfr\<^sub>s\<^sub>e\<^sub>t where "tfr\<^sub>s\<^sub>e\<^sub>t M \ (\s \ SMP M - (Var`\). \t \ SMP M - (Var`\). (\\. Unifier \ s t) \ \ s = \ t)" text \ Type-flaw resistance for strand steps: - The terms in a satisfiable equality step must have the same types - Inequality steps must satisfy the conditions of the "inequality lemma"\ fun tfr\<^sub>s\<^sub>t\<^sub>p where "tfr\<^sub>s\<^sub>t\<^sub>p (Equality a t t') = ((\\. Unifier \ t t') \ \ t = \ t')" | "tfr\<^sub>s\<^sub>t\<^sub>p (Inequality X F) = ( (\x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X. \a. \ (Var x) = TAtom a) \ (\f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ T = [] \ (\s \ set T. s \ Var ` set X)))" | "tfr\<^sub>s\<^sub>t\<^sub>p _ = True" text \ Type-flaw resistance for strands: - The set of terms in strands must be type-flaw resistant - The steps of strands must be type-flaw resistant \ definition tfr\<^sub>s\<^sub>t where "tfr\<^sub>s\<^sub>t S \ tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S) \ list_all tfr\<^sub>s\<^sub>t\<^sub>p S" subsection \Small Lemmata\ lemma tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def: "list_all tfr\<^sub>s\<^sub>t\<^sub>p S \ ((\a t t'. Equality a t t' \ set S \ (\\. Unifier \ t t') \ \ t = \ t') \ (\X F. Inequality X F \ set S \ (\x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X. \a. \ (Var x) = TAtom a) \ (\f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ T = [] \ (\s \ set T. s \ Var ` set X))))" (is "?P S \ ?Q S") proof show "?P S \ ?Q S" proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp show "?Q S \ ?P S" proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp qed lemma \_wf'': "TComp f T \ \ t \ arity f > 0" proof (induction t) case (Var x) thus ?case using \_wf(1)[of f T x] by blast next case (Fun g S) thus ?case using fun_type[of g S] const_type[of g] by (cases "arity g") auto qed lemma \_wf': "wf\<^sub>t\<^sub>r\<^sub>m t \ wf\<^sub>t\<^sub>r\<^sub>m (\ t)" proof (induction t) case (Fun f T) hence *: "arity f = length T" "\t. t \ set T \ wf\<^sub>t\<^sub>r\<^sub>m (\ t)" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto { assume "arity f = 0" hence ?case using const_type[of f] by auto } moreover { assume "arity f > 0" hence ?case using fun_type[of f] * by force } ultimately show ?case by auto qed (metis \_wf(2)) lemma fun_type_inv: assumes "\ t = TComp f T" shows "arity f > 0" using \_wf''(1)[of f T t] assms by simp_all lemma fun_type_inv_wf: assumes "\ t = TComp f T" "wf\<^sub>t\<^sub>r\<^sub>m t" shows "arity f = length T" using \_wf'[OF assms(2)] assms(1) unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto lemma const_type_inv: "\ (Fun c X) = TAtom a \ arity c = 0" by (rule ccontr, simp add: fun_type) lemma const_type_inv_wf: assumes "\ (Fun c X) = TAtom a" and "wf\<^sub>t\<^sub>r\<^sub>m (Fun c X)" shows "X = []" by (metis assms const_type_inv length_0_conv subtermeqI' wf\<^sub>t\<^sub>r\<^sub>m_def) lemma const_type': "\c \ \. \a \ \\<^sub>a. \X. \ (Fun c X) = TAtom a" using const_type by simp lemma fun_type': "\f \ \\<^sub>f. \X. \ (Fun f X) = TComp f (map \ X)" using fun_type by simp lemma fun_type_id_eq: "\ (Fun f X) = TComp g Y \ f = g" by (metis const_type fun_type neq0_conv "term.inject"(2) "term.simps"(4)) lemma fun_type_length_eq: "\ (Fun f X) = TComp g Y \ length X = length Y" by (metis fun_type fun_type_id_eq fun_type_inv(1) length_map term.inject(2)) lemma pgwt_type_map: assumes "public_ground_wf_term t" shows "\ t = TAtom a \ \f. t = Fun f []" "\ t = TComp g Y \ \X. t = Fun g X \ map \ X = Y" proof - let ?A = "\ t = TAtom a \ (\f. t = Fun f [])" let ?B = "\ t = TComp g Y \ (\X. t = Fun g X \ map \ X = Y)" have "?A \ ?B" proof (cases "\ t") case (Var a) obtain f X where "t = Fun f X" "arity f = length X" using pgwt_fun[OF assms(1)] pgwt_arity[OF assms(1)] by fastforce+ thus ?thesis using const_type_inv \\ t = TAtom a\ by auto next case (Fun g Y) obtain f X where *: "t = Fun f X" using pgwt_fun[OF assms(1)] by force hence "f = g" "map \ X = Y" using fun_type_id_eq \\ t = TComp g Y\ fun_type[OF fun_type_inv(1)[OF \\ t = TComp g Y\]] by fastforce+ thus ?thesis using *(1) \\ t = TComp g Y\ by auto qed thus "\ t = TAtom a \ \f. t = Fun f []" "\ t = TComp g Y \ \X. t = Fun g X \ map \ X = Y" by auto qed lemma wt_subst_Var[simp]: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" by (metis wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) lemma wt_subst_trm: "(\v. v \ fv t \ \ (Var v) = \ (\ v)) \ \ t = \ (t \ \)" proof (induction t) case (Fun f X) hence *: "\x. x \ set X \ \ x = \ (x \ \)" by auto show ?case proof (cases "f \ \\<^sub>f") case True hence "\X. \ (Fun f X) = TComp f (map \ X)" using fun_type' by auto thus ?thesis using * by auto next case False hence "\a \ \\<^sub>a. \X. \ (Fun f X) = TAtom a" using const_type' by auto thus ?thesis by auto qed qed auto lemma wt_subst_trm': "\wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \; \ s = \ t\ \ \ (s \ \) = \ (t \ \)" by (metis wt_subst_trm wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) lemma wt_subst_trm'': "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ \ t = \ (t \ \)" by (metis wt_subst_trm wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) lemma wt_subst_compose: assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" proof - have "\v. \ (\ v) = \ (\ v \ \)" using wt_subst_trm \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by metis moreover have "\v. \ (Var v) = \ (\ v)" using \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by metis ultimately have "\v. \ (Var v) = \ (\ v \ \)" by metis thus ?thesis unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def subst_compose_def by metis qed lemma wt_subst_TAtom_Var_cases: assumes \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and x: "\ (Var x) = TAtom a" shows "(\y. \ x = Var y) \ (\c. \ x = Fun c [])" proof (cases "(\y. \ x = Var y)") case False then obtain c T where c: "\ x = Fun c T" by (cases "\ x") simp_all hence "wf\<^sub>t\<^sub>r\<^sub>m (Fun c T)" using \(2) by fastforce hence "T = []" using const_type_inv_wf[of c T a] x c wt_subst_trm''[OF \(1), of "Var x"] by fastforce thus ?thesis using c by blast qed simp lemma wt_subst_TAtom_fv: assumes \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "\x. wf\<^sub>t\<^sub>r\<^sub>m (\ x)" and "\x \ fv t - X. \a. \ (Var x) = TAtom a" shows "\x \ fv (t \ \) - fv\<^sub>s\<^sub>e\<^sub>t (\ ` X). \a. \ (Var x) = TAtom a" using assms(3) proof (induction t) case (Var x) thus ?case proof (cases "x \ X") case False with Var obtain a where "\ (Var x) = TAtom a" by moura hence *: "\ (\ x) = TAtom a" "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" using \ unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by auto show ?thesis proof (cases "\ x") case (Var y) thus ?thesis using * by auto next case (Fun f T) hence "T = []" using * const_type_inv[of f T a] unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto thus ?thesis using Fun by auto qed qed auto qed fastforce lemma wt_subst_TAtom_subterms_subst: assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "\x \ fv t. \a. \ (Var x) = TAtom a" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\ ` fv t)" shows "subterms (t \ \) = subterms t \\<^sub>s\<^sub>e\<^sub>t \" using assms(2,3) proof (induction t) case (Var x) obtain a where a: "\ (Var x) = TAtom a" using Var.prems(1) by moura hence "\ (\ x) = TAtom a" using wt_subst_trm''[OF assms(1), of "Var x"] by simp hence "(\y. \ x = Var y) \ (\c. \ x = Fun c [])" using const_type_inv_wf Var.prems(2) by (cases "\ x") auto thus ?case by auto next case (Fun f T) have "subterms (t \ \) = subterms t \\<^sub>s\<^sub>e\<^sub>t \" when "t \ set T" for t using that Fun.prems(1,2) Fun.IH[OF that] by auto thus ?case by auto qed lemma wt_subst_TAtom_subterms_set_subst: assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M. \a. \ (Var x) = TAtom a" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\ ` fv\<^sub>s\<^sub>e\<^sub>t M)" shows "subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) = subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" proof show "subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) \ subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" proof fix t assume "t \ subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" then obtain s where s: "s \ M" "t \ subterms (s \ \)" by auto thus "t \ subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" using assms(2,3) wt_subst_TAtom_subterms_subst[OF assms(1), of s] by auto qed show "subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \ \ subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" proof fix t assume "t \ subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" then obtain s where s: "s \ M" "t \ subterms s \\<^sub>s\<^sub>e\<^sub>t \" by auto thus "t \ subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" using assms(2,3) wt_subst_TAtom_subterms_subst[OF assms(1), of s] by auto qed qed lemma wt_subst_subst_upd: assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "\ (Var x) = \ t" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\(x := t))" using assms unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by (metis fun_upd_other fun_upd_same) lemma wt_subst_const_fv_type_eq: assumes "\x \ fv t. \a. \ (Var x) = TAtom a" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "\x \ fv (t \ \). \y \ fv t. \ (Var x) = \ (Var y)" using assms(1) proof (induction t) case (Var x) then obtain a where a: "\ (Var x) = TAtom a" by moura show ?case proof (cases "\ x") case (Fun f T) hence "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" "\ (Fun f T) = TAtom a" using a wt_subst_trm''[OF \(1), of "Var x"] \(2) by fastforce+ thus ?thesis using const_type_inv_wf Fun by fastforce qed (use a wt_subst_trm''[OF \(1), of "Var x"] in simp) qed fastforce lemma TComp_term_cases: assumes "wf\<^sub>t\<^sub>r\<^sub>m t" "\ t = TComp f T" shows "(\v. t = Var v) \ (\T'. t = Fun f T' \ T = map \ T' \ T' \ [])" proof (cases "\v. t = Var v") case False then obtain T' where T': "t = Fun f T'" "T = map \ T'" using assms fun_type[OF fun_type_inv(1)[OF assms(2)]] fun_type_id_eq by (cases t) force+ thus ?thesis using assms fun_type_inv(1) fun_type_inv_wf by fastforce qed metis lemma TAtom_term_cases: assumes "wf\<^sub>t\<^sub>r\<^sub>m t" "\ t = TAtom \" shows "(\v. t = Var v) \ (\f. t = Fun f [])" using assms const_type_inv unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by (cases t) auto lemma subtermeq_imp_subtermtypeeq: assumes "wf\<^sub>t\<^sub>r\<^sub>m t" "s \ t" shows "\ s \ \ t" using assms(2,1) proof (induction t) case (Fun f T) thus ?case proof (cases "s = Fun f T") case False then obtain x where x: "x \ set T" "s \ x" using Fun.prems(1) by auto hence "wf\<^sub>t\<^sub>r\<^sub>m x" using wf_trm_subtermeq[OF Fun.prems(2)] Fun_param_is_subterm[of _ T f] by auto hence "\ s \ \ x" using Fun.IH[OF x] by simp moreover have "arity f > 0" using x fun_type_inv_wf Fun.prems by (metis length_pos_if_in_set term.order_refl wf\<^sub>t\<^sub>r\<^sub>m_def) ultimately show ?thesis using x Fun.prems fun_type[of f T] by auto qed simp qed simp lemma subterm_funs_term_in_type: assumes "wf\<^sub>t\<^sub>r\<^sub>m t" "Fun f T \ t" "\ (Fun f T) = TComp f (map \ T)" shows "f \ funs_term (\ t)" using assms(2,1,3) proof (induction t) case (Fun f' T') hence [simp]: "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" by (metis wf_trm_subtermeq) { fix a assume \: "\ (Fun f' T') = TAtom a" hence "Fun f T = Fun f' T'" using Fun TAtom_term_cases subtermeq_Var_const by metis hence False using Fun.prems(3) \ by simp } moreover { fix g S assume \: "\ (Fun f' T') = TComp g S" hence "g = f'" "S = map \ T'" using Fun.prems(2) fun_type_id_eq[OF \] fun_type[OF fun_type_inv(1)[OF \]] by auto hence \': "\ (Fun f' T') = TComp f' (map \ T')" using \ by auto hence "g \ funs_term (\ (Fun f' T'))" using \ by auto moreover { assume "Fun f T \ Fun f' T'" then obtain x where "x \ set T'" "Fun f T \ x" using Fun.prems(1) by auto hence "f \ funs_term (\ x)" using Fun.IH[OF _ _ _ Fun.prems(3), of x] wf_trm_subtermeq[OF \wf\<^sub>t\<^sub>r\<^sub>m (Fun f' T')\, of x] by force moreover have "\ x \ set (map \ T')" using \' \x \ set T'\ by auto ultimately have "f \ funs_term (\ (Fun f' T'))" using \' by auto } ultimately have ?case by (cases "Fun f T = Fun f' T'") (auto simp add: \g = f'\) } ultimately show ?case by (cases "\ (Fun f' T')") auto qed simp lemma wt_subst_fv_termtype_subterm: assumes "x \ fv (\ y)" and "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "wf\<^sub>t\<^sub>r\<^sub>m (\ y)" shows "\ (Var x) \ \ (Var y)" using subtermeq_imp_subtermtypeeq[OF assms(3) var_is_subterm[OF assms(1)]] wt_subst_trm''[OF assms(2), of "Var y"] by auto lemma wt_subst_fv\<^sub>s\<^sub>e\<^sub>t_termtype_subterm: assumes "x \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` Y)" and "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "\y \ Y. \ (Var x) \ \ (Var y)" using wt_subst_fv_termtype_subterm[OF _ assms(2), of x] assms(1,3) by fastforce lemma funs_term_type_iff: assumes t: "wf\<^sub>t\<^sub>r\<^sub>m t" and f: "arity f > 0" shows "f \ funs_term (\ t) \ (f \ funs_term t \ (\x \ fv t. f \ funs_term (\ (Var x))))" (is "?P t \ ?Q t") using t proof (induction t) case (Fun g T) hence IH: "?P s \ ?Q s" when "s \ set T" for s using that wf_trm_subterm[OF _ Fun_param_is_subterm] by blast have 0: "arity g = length T" using Fun.prems unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto show ?case proof (cases "f = g") case True thus ?thesis using fun_type[OF f] by simp next case False have "?P (Fun g T) \ (\s \ set T. ?P s)" proof assume *: "?P (Fun g T)" hence "\ (Fun g T) = TComp g (map \ T)" using const_type[of g] fun_type[of g] by force thus "\s \ set T. ?P s" using False * by force next assume *: "\s \ set T. ?P s" hence "\ (Fun g T) = TComp g (map \ T)" using 0 const_type[of g] fun_type[of g] by force thus "?P (Fun g T)" using False * by force qed thus ?thesis using False f IH by auto qed qed simp lemma funs_term_type_iff': assumes M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and f: "arity f > 0" shows "f \ \(funs_term ` \ ` M) \ (f \ \(funs_term ` M) \ (\x \ fv\<^sub>s\<^sub>e\<^sub>t M. f \ funs_term (\ (Var x))))" (is "?A \ ?B") proof assume ?A then obtain t where "t \ M" "wf\<^sub>t\<^sub>r\<^sub>m t" "f \ funs_term (\ t)" using M by moura thus ?B using funs_term_type_iff[OF _ f, of t] by auto next assume ?B then obtain t where "t \ M" "wf\<^sub>t\<^sub>r\<^sub>m t" "f \ funs_term t \ (\x \ fv t. f \ funs_term (\ (Var x)))" using M by auto thus ?A using funs_term_type_iff[OF _ f, of t] by blast qed lemma Ana_subterm_type: assumes "Ana t = (K,M)" and "wf\<^sub>t\<^sub>r\<^sub>m t" and "m \ set M" shows "\ m \ \ t" proof - have "m \ t" using Ana_subterm[OF assms(1)] assms(3) by auto thus ?thesis using subtermeq_imp_subtermtypeeq[OF assms(2)] by simp qed lemma wf_trm_TAtom_subterms: assumes "wf\<^sub>t\<^sub>r\<^sub>m t" "\ t = TAtom \" shows "subterms t = {t}" using assms const_type_inv unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by (cases t) force+ lemma wf_trm_TComp_subterm: assumes "wf\<^sub>t\<^sub>r\<^sub>m s" "t \ s" obtains f T where "\ s = TComp f T" proof (cases s) case (Var x) thus ?thesis using \t \ s\ by simp next case (Fun g S) hence "length S > 0" using assms Fun_subterm_inside_params[of t g S] by auto hence "arity g > 0" by (metis \wf\<^sub>t\<^sub>r\<^sub>m s\ \s = Fun g S\ term.order_refl wf\<^sub>t\<^sub>r\<^sub>m_def) thus ?thesis using fun_type \s = Fun g S\ that by auto qed lemma SMP_empty[simp]: "SMP {} = {}" proof (rule ccontr) assume "SMP {} \ {}" then obtain t where "t \ SMP {}" by auto thus False by (induct t rule: SMP.induct) auto qed lemma SMP_I: assumes "s \ M" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "t \ s \ \" "\v. wf\<^sub>t\<^sub>r\<^sub>m (\ v)" shows "t \ SMP M" using SMP.Substitution[OF SMP.MP[OF assms(1)] assms(2)] SMP.Subterm[of "s \ \" M t] assms(3,4) by (cases "t = s \ \") simp_all lemma SMP_wf_trm: assumes "t \ SMP M" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" shows "wf\<^sub>t\<^sub>r\<^sub>m t" using assms(1) by (induct t rule: SMP.induct) (use assms(2) in blast, use wf_trm_subtermeq in blast, use wf_trm_subst in blast, use Ana_keys_wf' in blast) lemma SMP_ikI[intro]: "t \ ik\<^sub>s\<^sub>t S \ t \ SMP (trms\<^sub>s\<^sub>t S)" by force lemma MP_setI[intro]: "x \ set S \ trms\<^sub>s\<^sub>t\<^sub>p x \ trms\<^sub>s\<^sub>t S" by force lemma SMP_setI[intro]: "x \ set S \ trms\<^sub>s\<^sub>t\<^sub>p x \ SMP (trms\<^sub>s\<^sub>t S)" by force lemma SMP_subset_I: assumes M: "\t \ M. \s \. s \ N \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = s \ \" shows "SMP M \ SMP N" proof fix t show "t \ SMP M \ t \ SMP N" proof (induction t rule: SMP.induct) case (MP t) then obtain s \ where s: "s \ N" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t = s \ \" using M by moura show ?case using SMP_I[OF s(1,2), of "s \ \"] s(3,4) wf_trm_subst_range_iff by fast qed (auto intro!: SMP.Substitution[of _ N]) qed lemma SMP_union: "SMP (A \ B) = SMP A \ SMP B" proof show "SMP (A \ B) \ SMP A \ SMP B" proof fix t assume "t \ SMP (A \ B)" thus "t \ SMP A \ SMP B" by (induct rule: SMP.induct) blast+ qed { fix t assume "t \ SMP A" hence "t \ SMP (A \ B)" by (induct rule: SMP.induct) blast+ } moreover { fix t assume "t \ SMP B" hence "t \ SMP (A \ B)" by (induct rule: SMP.induct) blast+ } ultimately show "SMP A \ SMP B \ SMP (A \ B)" by blast qed lemma SMP_append[simp]: "SMP (trms\<^sub>s\<^sub>t (S@S')) = SMP (trms\<^sub>s\<^sub>t S) \ SMP (trms\<^sub>s\<^sub>t S')" (is "?A = ?B") using SMP_union by simp lemma SMP_mono: "A \ B \ SMP A \ SMP B" proof - assume "A \ B" then obtain C where "B = A \ C" by moura thus "SMP A \ SMP B" by (simp add: SMP_union) qed lemma SMP_Union: "SMP (\m \ M. f m) = (\m \ M. SMP (f m))" proof show "SMP (\m\M. f m) \ (\m\M. SMP (f m))" proof fix t assume "t \ SMP (\m\M. f m)" thus "t \ (\m\M. SMP (f m))" by (induct t rule: SMP.induct) force+ qed show "(\m\M. SMP (f m)) \ SMP (\m\M. f m)" proof fix t assume "t \ (\m\M. SMP (f m))" then obtain m where "m \ M" "t \ SMP (f m)" by moura thus "t \ SMP (\m\M. f m)" using SMP_mono[of "f m" "\m\M. f m"] by auto qed qed lemma SMP_singleton_ex: "t \ SMP M \ (\m \ M. t \ SMP {m})" "m \ M \ t \ SMP {m} \ t \ SMP M" using SMP_Union[of "\t. {t}" M] by auto lemma SMP_Cons: "SMP (trms\<^sub>s\<^sub>t (x#S)) = SMP (trms\<^sub>s\<^sub>t [x]) \ SMP (trms\<^sub>s\<^sub>t S)" using SMP_append[of "[x]" S] by auto lemma SMP_Nil[simp]: "SMP (trms\<^sub>s\<^sub>t []) = {}" proof - { fix t assume "t \ SMP (trms\<^sub>s\<^sub>t [])" hence False by induct auto } thus ?thesis by blast qed lemma SMP_subset_union_eq: assumes "M \ SMP N" shows "SMP N = SMP (M \ N)" proof - { fix t assume "t \ SMP (M \ N)" hence "t \ SMP N" using assms by (induction rule: SMP.induct) blast+ } thus ?thesis using SMP_union by auto qed lemma SMP_subterms_subset: "subterms\<^sub>s\<^sub>e\<^sub>t M \ SMP M" proof fix t assume "t \ subterms\<^sub>s\<^sub>e\<^sub>t M" then obtain m where "m \ M" "t \ m" by auto thus "t \ SMP M" using SMP_I[of _ _ Var] by auto qed lemma SMP_SMP_subset: "N \ SMP M \ SMP N \ SMP M" by (metis SMP_mono SMP_subset_union_eq Un_commute Un_upper2) lemma wt_subst_rm_vars: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (rm_vars X \)" using rm_vars_dom unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by auto lemma wt_subst_SMP_subset: assumes "trms\<^sub>s\<^sub>t S \ SMP S'" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ SMP S'" proof fix t assume *: "t \ trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" show "t \ SMP S'" using trm_strand_subst_cong(2)[OF *] proof assume "\t'. t = t' \ \ \ t' \ trms\<^sub>s\<^sub>t S" thus "t \ SMP S'" using assms SMP.Substitution by auto next assume "\X F. Inequality X F \ set S \ (\t'\trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F. t = t' \ rm_vars (set X) \)" then obtain X F t' where **: "Inequality X F \ set S" "t'\trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" "t = t' \ rm_vars (set X) \" by force then obtain s where s: "s \ trms\<^sub>s\<^sub>t\<^sub>p (Inequality X F)" "t = s \ rm_vars (set X) \" by moura hence "s \ SMP (trms\<^sub>s\<^sub>t S)" using **(1) by force hence "t \ SMP (trms\<^sub>s\<^sub>t S)" using SMP.Substitution[OF _ wt_subst_rm_vars[OF assms(2)] wf_trms_subst_rm_vars'[OF assms(3)]] unfolding s(2) by blast thus "t \ SMP S'" by (metis SMP_union SMP_subset_union_eq UnCI assms(1)) qed qed lemma MP_subset_SMP: "\(trms\<^sub>s\<^sub>t\<^sub>p ` set S) \ SMP (trms\<^sub>s\<^sub>t S)" "trms\<^sub>s\<^sub>t S \ SMP (trms\<^sub>s\<^sub>t S)" "M \ SMP M" by auto lemma SMP_fun_map_snd_subset: "SMP (trms\<^sub>s\<^sub>t (map Send1 X)) \ SMP (trms\<^sub>s\<^sub>t [Send1 (Fun f X)])" proof fix t assume "t \ SMP (trms\<^sub>s\<^sub>t (map Send1 X))" thus "t \ SMP (trms\<^sub>s\<^sub>t [Send1 (Fun f X)])" proof (induction t rule: SMP.induct) case (MP t) hence "t \ set X" by auto hence "t \ Fun f X" by (metis subtermI') thus ?case using SMP.Subterm[of "Fun f X" "trms\<^sub>s\<^sub>t [Send1 (Fun f X)]" t] using SMP.MP by auto qed blast+ qed lemma SMP_wt_subst_subset: assumes "t \ SMP (M \\<^sub>s\<^sub>e\<^sub>t \)" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "t \ SMP M" using assms wf_trm_subst_range_iff[of \] by (induct t rule: SMP.induct) blast+ lemma SMP_wt_instances_subset: assumes "\t \ M. \s \ N. \\. t = s \ \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and "t \ SMP M" shows "t \ SMP N" proof - obtain m where m: "m \ M" "t \ SMP {m}" using SMP_singleton_ex(1)[OF assms(2)] by blast then obtain n \ where n: "n \ N" "m = n \ \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using assms(1) by fast have "t \ SMP (N \\<^sub>s\<^sub>e\<^sub>t \)" using n(1,2) SMP_singleton_ex(2)[of m "N \\<^sub>s\<^sub>e\<^sub>t \", OF _ m(2)] by fast thus ?thesis using SMP_wt_subst_subset[OF _ n(3,4)] by blast qed lemma SMP_consts: assumes "\t \ M. \c. t = Fun c []" and "\t \ M. Ana t = ([], [])" shows "SMP M = M" proof show "SMP M \ M" proof fix t show "t \ SMP M \ t \ M" apply (induction t rule: SMP.induct) by (use assms in auto) qed qed auto lemma SMP_subterms_eq: "SMP (subterms\<^sub>s\<^sub>e\<^sub>t M) = SMP M" proof show "SMP M \ SMP (subterms\<^sub>s\<^sub>e\<^sub>t M)" using SMP_mono[of M "subterms\<^sub>s\<^sub>e\<^sub>t M"] by blast show "SMP (subterms\<^sub>s\<^sub>e\<^sub>t M) \ SMP M" proof fix t show "t \ SMP (subterms\<^sub>s\<^sub>e\<^sub>t M) \ t \ SMP M" by (induction t rule: SMP.induct) blast+ qed qed lemma SMP_funs_term: assumes t: "t \ SMP M" "f \ funs_term t \ (\x \ fv t. f \ funs_term (\ (Var x)))" and f: "arity f > 0" and M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and Ana_f: "\s K T. Ana s = (K,T) \ f \ \(funs_term ` set K) \ f \ funs_term s" shows "f \ \(funs_term ` M) \ (\x \ fv\<^sub>s\<^sub>e\<^sub>t M. f \ funs_term (\ (Var x)))" using t proof (induction t rule: SMP.induct) case (Subterm t t') thus ?case by (metis UN_I vars_iff_subtermeq funs_term_subterms_eq(1) term.order_trans) next case (Substitution t \) show ?case using M SMP_wf_trm[OF Substitution.hyps(1)] wf_trm_subst[of \ t, OF Substitution.hyps(3)] funs_term_type_iff[OF _ f] wt_subst_trm''[OF Substitution.hyps(2), of t] Substitution.prems Substitution.IH by metis next case (Ana t K T t') thus ?case using Ana_f[OF Ana.hyps(2)] Ana_keys_fv[OF Ana.hyps(2)] by fastforce qed auto lemma id_type_eq: assumes "\ (Fun f X) = \ (Fun g Y)" shows "f \ \ \ g \ \" "f \ \\<^sub>f \ g \ \\<^sub>f" using assms const_type' fun_type' id_union_univ(1) by (metis UNIV_I UnE "term.distinct"(1))+ lemma fun_type_arg_cong: assumes "f \ \\<^sub>f" "g \ \\<^sub>f" "\ (Fun f (x#X)) = \ (Fun g (y#Y))" shows "\ x = \ y" "\ (Fun f X) = \ (Fun g Y)" using assms fun_type' by auto lemma fun_type_arg_cong': assumes "f \ \\<^sub>f" "g \ \\<^sub>f" "\ (Fun f (X@x#X')) = \ (Fun g (Y@y#Y'))" "length X = length Y" shows "\ x = \ y" using assms proof (induction X arbitrary: Y) case Nil thus ?case using fun_type_arg_cong(1)[of f g x X' y Y'] by auto next case (Cons x' X Y'') then obtain y' Y where "Y'' = y'#Y" by (metis length_Suc_conv) hence "\ (Fun f (X@x#X')) = \ (Fun g (Y@y#Y'))" "length X = length Y" using Cons.prems(3,4) fun_type_arg_cong(2)[OF Cons.prems(1,2), of x' "X@x#X'"] by auto thus ?thesis using Cons.IH[OF Cons.prems(1,2)] by auto qed lemma fun_type_param_idx: "\ (Fun f T) = Fun g S \ i < length T \ \ (T ! i) = S ! i" by (metis fun_type fun_type_id_eq fun_type_inv(1) nth_map term.inject(2)) lemma fun_type_param_ex: assumes "\ (Fun f T) = Fun g (map \ S)" "t \ set S" shows "\s \ set T. \ s = \ t" using fun_type_length_eq[OF assms(1)] length_map[of \ S] assms(2) fun_type_param_idx[OF assms(1)] nth_map in_set_conv_nth by metis lemma tfr_stp_all_split: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (x#S) \ list_all tfr\<^sub>s\<^sub>t\<^sub>p [x]" "list_all tfr\<^sub>s\<^sub>t\<^sub>p (x#S) \ list_all tfr\<^sub>s\<^sub>t\<^sub>p S" "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S') \ list_all tfr\<^sub>s\<^sub>t\<^sub>p S" "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S') \ list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@x#S') \ list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S')" by fastforce+ lemma tfr_stp_all_append: assumes "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" shows "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S')" using assms by fastforce lemma tfr_stp_all_wt_subst_apply: assumes "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "bvars\<^sub>s\<^sub>t S \ range_vars \ = {}" shows "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S \\<^sub>s\<^sub>t \)" using assms(1,4) proof (induction S) case (Cons x S) hence IH: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S \\<^sub>s\<^sub>t \)" using tfr_stp_all_split(2)[of x S] unfolding range_vars_alt_def by fastforce thus ?case proof (cases x) case (Equality a t t') hence "(\\. Unifier \ t t') \ \ t = \ t'" using Cons.prems by auto hence "(\\. Unifier \ (t \ \) (t' \ \)) \ \ (t \ \) = \ (t' \ \)" by (metis Unifier_comp' wt_subst_trm'[OF assms(2)]) moreover have "(x#S) \\<^sub>s\<^sub>t \ = Equality a (t \ \) (t' \ \)#(S \\<^sub>s\<^sub>t \)" using \x = Equality a t t'\ by auto ultimately show ?thesis using IH by auto next case (Inequality X F) let ?\ = "rm_vars (set X) \" let ?G = "F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\" let ?P = "\F X. \x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X. \a. \ (Var x) = TAtom a" let ?Q = "\F X. \f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ T = [] \ (\s \ set T. s \ Var ` set X)" have 0: "set X \ range_vars ?\ = {}" using Cons.prems(2) Inequality rm_vars_img_subset[of "set X"] by (auto simp add: subst_domain_def range_vars_alt_def) have 1: "?P F X \ ?Q F X" using Inequality Cons.prems by simp have 2: "fv\<^sub>s\<^sub>e\<^sub>t (?\ ` set X) = set X" by auto have "?P ?G X" when "?P F X" using that proof (induction F) case (Cons g G) obtain t t' where g: "g = (t,t')" by (metis surj_pair) have "\x \ (fv (t \ ?\) \ fv (t' \ ?\)) - set X. \a. \ (Var x) = Var a" proof - have *: "\x \ fv t - set X. \a. \ (Var x) = Var a" "\x \ fv t' - set X. \a. \ (Var x) = Var a" using g Cons.prems by simp_all have **: "\x. wf\<^sub>t\<^sub>r\<^sub>m (?\ x)" using \(2) wf_trm_subst_range_iff[of \] wf_trm_subst_rm_vars'[of \ _ "set X"] by simp show ?thesis using wt_subst_TAtom_fv[OF wt_subst_rm_vars[OF \(1)] ** *(1)] wt_subst_TAtom_fv[OF wt_subst_rm_vars[OF \(1)] ** *(2)] wt_subst_trm'[OF wt_subst_rm_vars[OF \(1), of "set X"]] 2 by blast qed moreover have "\x\fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\) - set X. \a. \ (Var x) = Var a" using Cons by auto ultimately show ?case using g by (auto simp add: subst_apply_pairs_def) qed (simp add: subst_apply_pairs_def) hence "?P ?G X \ ?Q ?G X" using 1 ineq_subterm_inj_cond_subst[OF 0, of "trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F"] trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst[of F ?\] by presburger moreover have "(x#S) \\<^sub>s\<^sub>t \ = Inequality X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\)#(S \\<^sub>s\<^sub>t \)" using \x = Inequality X F\ by auto ultimately show ?thesis using IH by simp qed auto qed simp lemma tfr_stp_all_same_type: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@Equality a t t'#S') \ Unifier \ t t' \ \ t = \ t'" by force+ lemma tfr_subset: "\A B. tfr\<^sub>s\<^sub>e\<^sub>t (A \ B) \ tfr\<^sub>s\<^sub>e\<^sub>t A" "\A B. tfr\<^sub>s\<^sub>e\<^sub>t B \ A \ B \ tfr\<^sub>s\<^sub>e\<^sub>t A" "\A B. tfr\<^sub>s\<^sub>e\<^sub>t B \ SMP A \ SMP B \ tfr\<^sub>s\<^sub>e\<^sub>t A" proof - show 1: "tfr\<^sub>s\<^sub>e\<^sub>t (A \ B) \ tfr\<^sub>s\<^sub>e\<^sub>t A" for A B using SMP_union[of A B] unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by simp fix A B assume B: "tfr\<^sub>s\<^sub>e\<^sub>t B" show "A \ B \ tfr\<^sub>s\<^sub>e\<^sub>t A" proof - assume "A \ B" then obtain C where "B = A \ C" by moura thus ?thesis using B 1 by blast qed show "SMP A \ SMP B \ tfr\<^sub>s\<^sub>e\<^sub>t A" proof - assume "SMP A \ SMP B" then obtain C where "SMP B = SMP A \ C" by moura thus ?thesis using B unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast qed qed lemma tfr_empty[simp]: "tfr\<^sub>s\<^sub>e\<^sub>t {}" unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by simp lemma tfr_consts_mono: assumes "\t \ M. \c. t = Fun c []" and "\t \ M. Ana t = ([], [])" and "tfr\<^sub>s\<^sub>e\<^sub>t N" shows "tfr\<^sub>s\<^sub>e\<^sub>t (N \ M)" proof - { fix s t assume *: "s \ SMP (N \ M) - range Var" "t \ SMP (N \ M) - range Var" "\\. Unifier \ s t" hence **: "is_Fun s" "is_Fun t" "s \ SMP N \ s \ M" "t \ SMP N \ t \ M" using assms(3) SMP_consts[OF assms(1,2)] SMP_union[of N M] by auto moreover have "\ s = \ t" when "s \ SMP N" "t \ SMP N" using that assms(3) *(3) **(1,2) unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast moreover have "\ s = \ t" when st: "s \ M" "t \ M" proof - obtain c d where "s = Fun c []" "t = Fun d []" using st assms(1) by moura hence "s = t" using *(3) by fast thus ?thesis by metis qed moreover have "\ s = \ t" when st: "s \ SMP N" "t \ M" proof - obtain c where "t = Fun c []" using st assms(1) by moura hence "s = t" using *(3) **(1,2) by auto thus ?thesis by metis qed moreover have "\ s = \ t" when st: "s \ M" "t \ SMP N" proof - obtain c where "s = Fun c []" using st assms(1) by moura hence "s = t" using *(3) **(1,2) by auto thus ?thesis by metis qed ultimately have "\ s = \ t" by metis } thus ?thesis by (metis tfr\<^sub>s\<^sub>e\<^sub>t_def) qed lemma dual\<^sub>s\<^sub>t_tfr\<^sub>s\<^sub>t\<^sub>p: "list_all tfr\<^sub>s\<^sub>t\<^sub>p S \ list_all tfr\<^sub>s\<^sub>t\<^sub>p (dual\<^sub>s\<^sub>t S)" proof (induction S) case (Cons x S) have "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Cons.prems by simp hence IH: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (dual\<^sub>s\<^sub>t S)" using Cons.IH by metis from Cons show ?case proof (cases x) case (Equality a t t') hence "(\\. Unifier \ t t') \ \ t = \ t'" using Cons by auto thus ?thesis using Equality IH by fastforce next case (Inequality X F) have "set (dual\<^sub>s\<^sub>t (x#S)) = insert x (set (dual\<^sub>s\<^sub>t S))" using Inequality by auto moreover have "(\x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X. \a. \ (Var x) = Var a) \ (\f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ T = [] \ (\s \ set T. s \ Var ` set X))" using Cons.prems Inequality by auto ultimately show ?thesis using Inequality IH by auto qed auto qed simp lemma subst_var_inv_wt: assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (subst_var_inv \ X)" using assms f_inv_into_f[of _ \ X] unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def subst_var_inv_def by presburger lemma subst_var_inv_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (subst_var_inv \ X))" using f_inv_into_f[of _ \ X] unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def subst_var_inv_def by auto lemma unify_list_wt_if_same_type: assumes "Unification.unify E B = Some U" "\(s,t) \ set E. wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t \ \ s = \ t" and "\(v,t) \ set B. \ (Var v) = \ t" shows "\(v,t) \ set U. \ (Var v) = \ t" using assms proof (induction E B arbitrary: U rule: Unification.unify.induct) case (2 f X g Y E B U) hence "wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)" "wf\<^sub>t\<^sub>r\<^sub>m (Fun g Y)" "\ (Fun f X) = \ (Fun g Y)" by auto from "2.prems"(1) obtain E' where *: "decompose (Fun f X) (Fun g Y) = Some E'" and [simp]: "f = g" "length X = length Y" "E' = zip X Y" and **: "Unification.unify (E'@E) B = Some U" by (auto split: option.splits) have "\(s,t) \ set E'. wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t \ \ s = \ t" proof - { fix s t assume "(s,t) \ set E'" then obtain X' X'' Y' Y'' where "X = X'@s#X''" "Y = Y'@t#Y''" "length X' = length Y'" using zip_arg_subterm_split[of s t X Y] \E' = zip X Y\ by metis hence "\ (Fun f (X'@s#X'')) = \ (Fun g (Y'@t#Y''))" by (metis \\ (Fun f X) = \ (Fun g Y)\) from \E' = zip X Y\ have "\(s,t) \ set E'. s \ Fun f X \ t \ Fun g Y" using zip_arg_subterm[of _ _ X Y] by blast with \(s,t) \ set E'\ have "wf\<^sub>t\<^sub>r\<^sub>m s" "wf\<^sub>t\<^sub>r\<^sub>m t" using wf_trm_subterm \wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)\ \wf\<^sub>t\<^sub>r\<^sub>m (Fun g Y)\ by (blast,blast) moreover have "f \ \\<^sub>f" proof (rule ccontr) assume "f \ \\<^sub>f" hence "f \ \" "arity f = 0" using const_arity_eq_zero[of f] by simp_all thus False using \wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)\ * \(s,t) \ set E'\ unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto qed hence "\ s = \ t" using fun_type_arg_cong' \f \ \\<^sub>f\ \\ (Fun f (X'@s#X'')) = \ (Fun g (Y'@t#Y''))\ \length X' = length Y'\ \f = g\ by metis ultimately have "wf\<^sub>t\<^sub>r\<^sub>m s" "wf\<^sub>t\<^sub>r\<^sub>m t" "\ s = \ t" by metis+ } thus ?thesis by blast qed moreover have "\(s,t) \ set E. wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t \ \ s = \ t" using "2.prems"(2) by auto ultimately show ?case using "2.IH"[OF * ** _ "2.prems"(3)] by fastforce next case (3 v t E B U) hence "\ (Var v) = \ t" "wf\<^sub>t\<^sub>r\<^sub>m t" by auto hence "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (subst v t)" and *: "\(v, t) \ set ((v,t)#B). \ (Var v) = \ t" "\t t'. (t,t') \ set E \ \ t = \ t'" using "3.prems"(2,3) unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def subst_def by auto show ?case proof (cases "t = Var v") assume "t = Var v" thus ?case using 3 by auto next assume "t \ Var v" hence "v \ fv t" using "3.prems"(1) by auto hence **: "Unification.unify (subst_list (subst v t) E) ((v, t)#B) = Some U" using Unification.unify.simps(3)[of v t E B] "3.prems"(1) \t \ Var v\ by auto have "\(s, t) \ set (subst_list (subst v t) E). wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t" using wf_trm_subst_singleton[OF _ \wf\<^sub>t\<^sub>r\<^sub>m t\] "3.prems"(2) unfolding subst_list_def subst_def by auto moreover have "\(s, t) \ set (subst_list (subst v t) E). \ s = \ t" using *(2)[THEN wt_subst_trm'[OF \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (subst v t)\]] by (simp add: subst_list_def) ultimately show ?thesis using "3.IH"(2)[OF \t \ Var v\ \v \ fv t\ ** _ *(1)] by auto qed next case (4 f X v E B U) hence "\ (Var v) = \ (Fun f X)" "wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)" by auto hence "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (subst v (Fun f X))" and *: "\(v, t) \ set ((v,(Fun f X))#B). \ (Var v) = \ t" "\t t'. (t,t') \ set E \ \ t = \ t'" using "4.prems"(2,3) unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def subst_def by auto have "v \ fv (Fun f X)" using "4.prems"(1) by force hence **: "Unification.unify (subst_list (subst v (Fun f X)) E) ((v, (Fun f X))#B) = Some U" using Unification.unify.simps(3)[of v "Fun f X" E B] "4.prems"(1) by auto have "\(s, t) \ set (subst_list (subst v (Fun f X)) E). wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t" using wf_trm_subst_singleton[OF _ \wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)\] "4.prems"(2) unfolding subst_list_def subst_def by auto moreover have "\(s, t) \ set (subst_list (subst v (Fun f X)) E). \ s = \ t" using *(2)[THEN wt_subst_trm'[OF \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (subst v (Fun f X))\]] by (simp add: subst_list_def) ultimately show ?case using "4.IH"[OF \v \ fv (Fun f X)\ ** _ *(1)] by auto qed auto lemma mgu_wt_if_same_type: assumes "mgu s t = Some \" "wf\<^sub>t\<^sub>r\<^sub>m s" "wf\<^sub>t\<^sub>r\<^sub>m t" "\ s = \ t" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" proof - let ?fv_disj = "\v t S. \(\(v',t') \ S - {(v,t)}. (insert v (fv t)) \ (insert v' (fv t')) \ {})" from assms(1) obtain \' where "Unification.unify [(s,t)] [] = Some \'" "subst_of \' = \" by (auto simp: mgu_def split: option.splits) hence "\(v,t) \ set \'. \ (Var v) = \ t" "distinct (map fst \')" using assms(2,3,4) unify_list_wt_if_same_type unify_list_distinct[of "[(s,t)]"] by auto thus "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using \subst_of \' = \\ unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def proof (induction \' arbitrary: \ rule: List.rev_induct) case (snoc tt \' \) then obtain v t where tt: "tt = (v,t)" by (metis surj_pair) hence \: "\ = subst v t \\<^sub>s subst_of \'" using snoc.prems(3) by simp have "\(v,t) \ set \'. \ (Var v) = \ t" "distinct (map fst \')" using snoc.prems(1,2) by auto then obtain \'' where \'': "subst_of \' = \''" "\v. \ (Var v) = \ (\'' v)" by (metis snoc.IH) hence "\ t = \ (t \ \'')" for t using wt_subst_trm by blast hence "\ (Var v) = \ (\'' v)" "\ t = \ (t \ \'')" using \''(2) by auto moreover have "\ (Var v) = \ t" using snoc.prems(1) tt by simp moreover have \2: "\ = Var(v := t) \\<^sub>s \'' " using \ \''(1) unfolding subst_def by simp ultimately have "\ (Var v) = \ (\ v)" unfolding subst_compose_def by simp have "subst_domain (subst v t) \ {v}" unfolding subst_def by (auto simp add: subst_domain_def) hence *: "subst_domain \ \ insert v (subst_domain \'')" using tt \ \''(1) snoc.prems(2) subst_domain_compose[of _ \''] by (auto simp add: subst_domain_def) have "v \ set (map fst \')" using tt snoc.prems(2) by auto hence "v \ subst_domain \''" using \''(1) subst_of_dom_subset[of \'] by auto { fix w assume "w \ subst_domain \''" hence "\ w = \'' w" using \2 \''(1) \v \ subst_domain \''\ unfolding subst_compose_def by auto hence "\ (Var w) = \ (\ w)" using \''(2) by simp } thus ?case using \\ (Var v) = \ (\ v)\ * by force qed simp qed lemma wt_Unifier_if_Unifier: assumes s_t: "wf\<^sub>t\<^sub>r\<^sub>m s" "wf\<^sub>t\<^sub>r\<^sub>m t" "\ s = \ t" and \: "Unifier \ s t" shows "\\. Unifier \ s t \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_always_unifies[OF \] mgu_gives_MGU[THEN MGU_is_Unifier[of s _ t]] mgu_wt_if_same_type[OF _ s_t] mgu_wf_trm[OF _ s_t(1,2)] wf_trm_subst_range_iff by fast end subsection \Automatically Proving Type-Flaw Resistance\ subsubsection \Definitions: Variable Renaming\ abbreviation "max_var t \ Max (insert 0 (snd ` fv t))" abbreviation "max_var_set X \ Max (insert 0 (snd ` X))" definition "var_rename n v \ Var (fst v, snd v + Suc n)" definition "var_rename_inv n v \ Var (fst v, snd v - Suc n)" subsubsection \Definitions: Computing a Finite Representation of the Sub-Message Patterns\ text \A sufficient requirement for a term to be a well-typed instance of another term\ definition is_wt_instance_of_cond where "is_wt_instance_of_cond \ t s \ ( \ t = \ s \ (case mgu t s of None \ False | Some \ \ inj_on \ (fv t) \ (\x \ fv t. is_Var (\ x))))" definition has_all_wt_instances_of where "has_all_wt_instances_of \ N M \ \t \ N. \s \ M. is_wt_instance_of_cond \ t s" text \This function computes a finite representation of the set of sub-message patterns\ definition SMP0 where "SMP0 Ana \ M \ let f = \t. Fun (the_Fun (\ t)) (map Var (zip (args (\ t)) [0.. t))])); g = \M'. map f (filter (\t. is_Var t \ is_Fun (\ t)) M')@ concat (map (fst \ Ana) M')@concat (map subterms_list M'); h = remdups \ g in while (\A. set (h A) \ set A) h M" text \These definitions are useful to refine an SMP representation set\ fun generalize_term where "generalize_term _ _ n (Var x) = (Var x, n)" | "generalize_term \ p n (Fun f T) = (let \ = \ (Fun f T) in if p \ then (Var (\, n), Suc n) else let (T',n') = foldr (\t (S,m). let (t',m') = generalize_term \ p m t in (t'#S,m')) T ([],n) in (Fun f T', n'))" definition generalize_terms where "generalize_terms \ p \ map (fst \ generalize_term \ p 0)" definition remove_superfluous_terms where "remove_superfluous_terms \ T \ let f = \S t R. \s \ set S - R. s \ t \ is_wt_instance_of_cond \ t s; g = \S t (U,R). if f S t R then (U, insert t R) else (t#U, R); h = \S. remdups (fst (foldr (g S) S ([],{}))) in while (\S. h S \ S) h T" subsubsection \Definitions: Checking Type-Flaw Resistance\ definition is_TComp_var_instance_closed where "is_TComp_var_instance_closed \ M \ \x \ fv\<^sub>s\<^sub>e\<^sub>t M. is_Fun (\ (Var x)) \ (\t \ M. is_Fun t \ \ t = \ (Var x) \ list_all is_Var (args t) \ distinct (args t))" definition finite_SMP_representation where "finite_SMP_representation arity Ana \ M \ (M = {} \ card M > 0) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity M \ has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M \ has_all_wt_instances_of \ (\((set \ fst \ Ana) ` M)) M \ is_TComp_var_instance_closed \ M" definition comp_tfr\<^sub>s\<^sub>e\<^sub>t where "comp_tfr\<^sub>s\<^sub>e\<^sub>t arity Ana \ M \ finite_SMP_representation arity Ana \ M \ (let \ = var_rename (max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M)) in \s \ M. \t \ M. is_Fun s \ is_Fun t \ \ s \ \ t \ mgu s (t \ \) = None)" fun comp_tfr\<^sub>s\<^sub>t\<^sub>p where "comp_tfr\<^sub>s\<^sub>t\<^sub>p \ (\_: t \ t'\\<^sub>s\<^sub>t) = (mgu t t' \ None \ \ t = \ t')" | "comp_tfr\<^sub>s\<^sub>t\<^sub>p \ (\X\\\: F\\<^sub>s\<^sub>t) = ( (\x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X. is_Var (\ (Var x))) \ (\u \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F). is_Fun u \ (args u = [] \ (\s \ set (args u). s \ Var ` set X))))" | "comp_tfr\<^sub>s\<^sub>t\<^sub>p _ _ = True" definition comp_tfr\<^sub>s\<^sub>t where "comp_tfr\<^sub>s\<^sub>t arity Ana \ M S \ list_all (comp_tfr\<^sub>s\<^sub>t\<^sub>p \) S \ list_all (wf\<^sub>t\<^sub>r\<^sub>m' arity) (trms_list\<^sub>s\<^sub>t S) \ has_all_wt_instances_of \ (trms\<^sub>s\<^sub>t S) M \ comp_tfr\<^sub>s\<^sub>e\<^sub>t arity Ana \ M" subsubsection \Small Lemmata\ lemma max_var_set_mono: assumes "finite N" and "M \ N" shows "max_var_set M \ max_var_set N" by (meson assms Max.subset_imp finite.insertI finite_imageI image_mono insert_mono insert_not_empty) lemma less_Suc_max_var_set: assumes z: "z \ X" and X: "finite X" shows "snd z < Suc (max_var_set X)" proof - have "snd z \ snd ` X" using z by simp hence "snd z \ Max (insert 0 (snd ` X))" using X by simp thus ?thesis using X by simp qed lemma (in typed_model) finite_SMP_representationD: assumes "finite_SMP_representation arity Ana \ M" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" and "has_all_wt_instances_of \ (\((set \ fst \ Ana) ` M)) M" and "is_TComp_var_instance_closed \ M" and "finite M" using assms wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[of M] unfolding finite_SMP_representation_def list_all_iff card_gt_0_iff by blast+ lemma (in typed_model) is_wt_instance_of_condD: assumes t_instance_s: "is_wt_instance_of_cond \ t s" obtains \ where "\ t = \ s" "mgu t s = Some \" "inj_on \ (fv t)" "\ ` (fv t) \ range Var" using t_instance_s unfolding is_wt_instance_of_cond_def Let_def by (cases "mgu t s") fastforce+ lemma (in typed_model) is_wt_instance_of_condD': assumes t_wf_trm: "wf\<^sub>t\<^sub>r\<^sub>m t" and s_wf_trm: "wf\<^sub>t\<^sub>r\<^sub>m s" and t_instance_s: "is_wt_instance_of_cond \ t s" shows "\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = s \ \" proof - obtain \ where s: "\ t = \ s" "mgu t s = Some \" "inj_on \ (fv t)" "\ ` (fv t) \ range Var" by (metis is_wt_instance_of_condD[OF t_instance_s]) have 0: "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m s" using s(1) t_wf_trm s_wf_trm by auto note 1 = mgu_wt_if_same_type[OF s(2) 0 s(1)] note 2 = conjunct1[OF mgu_gives_MGU[OF s(2)]] show ?thesis using s(1) inj_var_ran_unifiable_has_subst_match[OF 2 s(3,4)] wt_subst_compose[OF 1 subst_var_inv_wt[OF 1, of "fv t"]] wf_trms_subst_compose[OF mgu_wf_trms[OF s(2) 0] subst_var_inv_wf_trms[of \ "fv t"]] by auto qed lemma (in typed_model) is_wt_instance_of_condD'': assumes s_wf_trm: "wf\<^sub>t\<^sub>r\<^sub>m s" and t_instance_s: "is_wt_instance_of_cond \ t s" and t_var: "t = Var x" shows "\y. s = Var y \ \ (Var y) = \ (Var x)" proof - obtain \ where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and s: "Var x = s \ \" using is_wt_instance_of_condD'[OF _ s_wf_trm t_instance_s] t_var by auto obtain y where y: "s = Var y" using s by (cases s) auto show ?thesis using wt_subst_trm''[OF \] s y by metis qed lemma (in typed_model) has_all_wt_instances_ofD: assumes N_instance_M: "has_all_wt_instances_of \ N M" and t_in_N: "t \ N" obtains s \ where "s \ M" "\ t = \ s" "mgu t s = Some \" "inj_on \ (fv t)" "\ ` (fv t) \ range Var" by (metis t_in_N N_instance_M is_wt_instance_of_condD has_all_wt_instances_of_def) lemma (in typed_model) has_all_wt_instances_ofD': assumes N_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and M_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and N_instance_M: "has_all_wt_instances_of \ N M" and t_in_N: "t \ N" shows "\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t \ M \\<^sub>s\<^sub>e\<^sub>t \" using assms is_wt_instance_of_condD' unfolding has_all_wt_instances_of_def by fast lemma (in typed_model) has_all_wt_instances_ofD'': assumes N_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and M_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and N_instance_M: "has_all_wt_instances_of \ N M" and t_in_N: "Var x \ N" shows "\y. Var y \ M \ \ (Var y) = \ (Var x)" using assms is_wt_instance_of_condD'' unfolding has_all_wt_instances_of_def by fast lemma (in typed_model) has_all_instances_of_if_subset: assumes "N \ M" shows "has_all_wt_instances_of \ N M" unfolding has_all_wt_instances_of_def proof fix t assume t: "t \ N" hence "is_wt_instance_of_cond \ t t" using inj_onI[of "fv t"] mgu_same_empty[of t] unfolding is_wt_instance_of_cond_def by force thus "\s \ M. is_wt_instance_of_cond \ t s" using t assms by blast qed lemma (in typed_model) SMP_I': assumes N_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and M_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and N_instance_M: "has_all_wt_instances_of \ N M" and t_in_N: "t \ N" shows "t \ SMP M" using has_all_wt_instances_ofD'[OF N_wf_trms M_wf_trms N_instance_M t_in_N] SMP.Substitution[OF SMP.MP[of _ M]] by blast subsubsection \Lemma: Proving Type-Flaw Resistance\ locale typed_model' = typed_model arity public Ana \ for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,(('fun,'atom::finite) term_type \ nat)) term \ (('fun,(('fun,'atom) term_type \ nat)) term list \ ('fun,(('fun,'atom) term_type \ nat)) term list)" and \::"('fun,(('fun,'atom) term_type \ nat)) term \ ('fun,'atom) term_type" + assumes \_Var_fst: "\\ n m. \ (Var (\,n)) = \ (Var (\,m))" and Ana_const: "\c T. arity c = 0 \ Ana (Fun c T) = ([],[])" and Ana_subst'_or_Ana_keys_subterm: "(\f T \ K R. Ana (Fun f T) = (K,R) \ Ana (Fun f T \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \,R \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)) \ (\t K R k. Ana t = (K,R) \ k \ set K \ k \ t)" begin lemma var_rename_inv_comp: "t \ (var_rename n \\<^sub>s var_rename_inv n) = t" proof (induction t) case (Fun f T) hence "map (\t. t \ var_rename n \\<^sub>s var_rename_inv n) T = T" by (simp add: map_idI) - thus ?case by (metis subst_apply_term.simps(2)) -qed (simp add: var_rename_def var_rename_inv_def) + thus ?case by (metis eval_term.simps(2)) +qed (simp add: var_rename_def var_rename_inv_def subst_compose) lemma var_rename_fv_disjoint: "fv s \ fv (t \ var_rename (max_var s)) = {}" proof - have 1: "\v \ fv s. snd v \ max_var s" by simp have 2: "\v \ fv (t \ var_rename n). snd v > n" for n unfolding var_rename_def by (induct t) auto show ?thesis using 1 2 by force qed lemma var_rename_fv_set_disjoint: assumes "finite M" "s \ M" shows "fv s \ fv (t \ var_rename (max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M))) = {}" proof - have 1: "\v \ fv s. snd v \ max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M)" using assms proof (induction M rule: finite_induct) case (insert t M) thus ?case proof (cases "t = s") case False hence "\v \ fv s. snd v \ max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M)" using insert by simp moreover have "max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M) \ max_var_set (fv\<^sub>s\<^sub>e\<^sub>t (insert t M))" using insert.hyps(1) insert.prems by force ultimately show ?thesis by auto qed simp qed simp have 2: "\v \ fv (t \ var_rename n). snd v > n" for n unfolding var_rename_def by (induct t) auto show ?thesis using 1 2 by force qed lemma var_rename_fv_set_disjoint': assumes "finite M" shows "fv\<^sub>s\<^sub>e\<^sub>t M \ fv\<^sub>s\<^sub>e\<^sub>t (N \\<^sub>s\<^sub>e\<^sub>t var_rename (max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M))) = {}" using var_rename_fv_set_disjoint[OF assms] by auto lemma var_rename_is_renaming[simp]: "subst_range (var_rename n) \ range Var" "subst_range (var_rename_inv n) \ range Var" unfolding var_rename_def var_rename_inv_def by auto lemma var_rename_wt[simp]: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (var_rename n)" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (var_rename_inv n)" by (auto simp add: var_rename_def var_rename_inv_def wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def \_Var_fst) lemma var_rename_wt': assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "s = m \ \" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (var_rename_inv n \\<^sub>s \)" "s = m \ var_rename n \ var_rename_inv n \\<^sub>s \" using assms(2) wt_subst_compose[OF var_rename_wt(2)[of n] assms(1)] var_rename_inv_comp[of m n] by force+ lemma var_rename_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_range[simp]: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (var_rename n))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (var_rename_inv n))" using var_rename_is_renaming by fastforce+ lemma Fun_range_case: "(\f T. Fun f T \ M \ P f T) \ (\u \ M. case u of Fun f T \ P f T | _ \ True)" "(\f T. Fun f T \ M \ P f T) \ (\u \ M. is_Fun u \ P (the_Fun u) (args u))" by (auto split: "term.splits") lemma is_TComp_var_instance_closedD: assumes x: "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var x) = \ (Var y)" "\ (Var x) = TComp f T" and closed: "is_TComp_var_instance_closed \ M" shows "\g U. Fun g U \ M \ \ (Fun g U) = \ (Var x) \ (\u \ set U. is_Var u) \ distinct U" using assms unfolding is_TComp_var_instance_closed_def list_all_iff list_ex_iff by fastforce lemma is_TComp_var_instance_closedD': assumes "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var x) = \ (Var y)" "TComp f T \ \ (Var x)" and closed: "is_TComp_var_instance_closed \ M" and wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" shows "\g U. Fun g U \ M \ \ (Fun g U) = TComp f T \ (\u \ set U. is_Var u) \ distinct U" using assms(1,2) proof (induction "\ (Var x)" arbitrary: x) case (Fun g U) note IH = Fun.hyps(1) have g: "arity g > 0" using Fun.hyps(2) fun_type_inv[of "Var x"] \_Var_fst by simp_all then obtain V where V: "Fun g V \ M" "\ (Fun g V) = \ (Var x)" "\v \ set V. \x. v = Var x" "distinct V" "length U = length V" using is_TComp_var_instance_closedD[OF Fun.prems(1) Fun.hyps(2)[symmetric] closed(1)] by (metis Fun.hyps(2) fun_type_id_eq fun_type_length_eq is_VarE) hence U: "U = map \ V" using fun_type[OF g(1), of V] Fun.hyps(2) by simp hence 1: "\ v \ set U" when v: "v \ set V" for v using v by simp have 2: "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var z) = \ (Var y)" when z: "Var z \ set V" for z using V(1) fv_subset_subterms Fun_param_in_subterms[OF z] by fastforce show ?case proof (cases "TComp f T = \ (Var x)") case False then obtain u where u: "u \ set U" "TComp f T \ u" using Fun.prems(2) Fun.hyps(2) by moura then obtain y where y: "Var y \ set V" "\ (Var y) = u" using U V(3) \_Var_fst by auto show ?thesis using IH[OF _ 2[OF y(1)]] u y(2) by metis qed (use V in fastforce) qed simp lemma TComp_var_instance_wt_subst_exists: assumes gT: "\ (Fun g T) = TComp g (map \ U)" "wf\<^sub>t\<^sub>r\<^sub>m (Fun g T)" and U: "\u \ set U. \y. u = Var y" "distinct U" shows "\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ Fun g T = Fun g U \ \" proof - define the_i where "the_i \ \y. THE x. x < length U \ U ! x = Var y" define \ where \: "\ \ \y. if Var y \ set U then T ! the_i y else Var y" have g: "arity g > 0" using gT(1,2) fun_type_inv(1) by blast have UT: "length U = length T" using fun_type_length_eq gT(1) by fastforce have 1: "the_i y < length U \ U ! the_i y = Var y" when y: "Var y \ set U" for y using theI'[OF distinct_Ex1[OF U(2) y]] unfolding the_i_def by simp have 2: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using \ 1 gT(1) fun_type[OF g] UT unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by (metis (no_types, lifting) nth_map term.inject(2)) have "\i \ = T ! i" using \ 1 U(1) UT distinct_Ex1[OF U(2)] in_set_conv_nth - by (metis (no_types, lifting) subst_apply_term.simps(1)) + by (metis (no_types, lifting) eval_term.simps(1)) hence "T = map (\t. t \ \) U" by (simp add: UT nth_equalityI) hence 3: "Fun g T = Fun g U \ \" by simp have "subst_range \ \ set T" using \ 1 U(1) UT by (auto simp add: subst_domain_def) hence 4: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using gT(2) wf_trm_param by auto show ?thesis by (metis 2 3 4) qed lemma TComp_var_instance_closed_has_Var: assumes closed: "is_TComp_var_instance_closed \ M" and wf_M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and wf_\x: "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" and y_ex: "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var x) = \ (Var y)" and t: "t \ \ x" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var y) = \ t" proof (cases "\ (Var x)") case (Var a) hence "t = \ x" using t wf_\x \_wt by (metis (full_types) const_type_inv_wf fun_if_subterm subtermeq_Var_const(2) wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) thus ?thesis using y_ex wt_subst_trm''[OF \_wt, of "Var x"] by fastforce next case (Fun f T) hence \_\x: "\ (\ x) = TComp f T" using wt_subst_trm''[OF \_wt, of "Var x"] by auto show ?thesis proof (cases "t = \ x") case False hence t_subt_\x: "t \ \ x" using t(1) \_\x by fastforce obtain T' where T': "\ x = Fun f T'" using \_\x t_subt_\x fun_type_id_eq by (cases "\ x") auto obtain g S where gS: "Fun g S \ \ x" "t \ set S" using Fun_ex_if_subterm[OF t_subt_\x] by blast have gS_wf: "wf\<^sub>t\<^sub>r\<^sub>m (Fun g S)" by (rule wf_trm_subtermeq[OF wf_\x gS(1)]) hence "arity g > 0" using gS(2) by (metis length_pos_if_in_set wf_trm_arity) hence gS_\: "\ (Fun g S) = TComp g (map \ S)" using fun_type by blast obtain h U where hU: "Fun h U \ M" "\ (Fun h U) = Fun g (map \ S)" "\u \ set U. is_Var u" using is_TComp_var_instance_closedD'[OF y_ex _ closed wf_M] subtermeq_imp_subtermtypeeq[OF wf_\x] gS \_\x Fun gS_\ by metis obtain y where y: "Var y \ set U" "\ (Var y) = \ t" using hU(3) fun_type_param_ex[OF hU(2) gS(2)] by fast have "y \ fv\<^sub>s\<^sub>e\<^sub>t M" using hU(1) y(1) by force thus ?thesis using y(2) closed by metis qed (metis y_ex Fun \_\x) qed lemma TComp_var_instance_closed_has_Fun: assumes closed: "is_TComp_var_instance_closed \ M" and wf_M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and wf_\x: "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" and y_ex: "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var x) = \ (Var y)" and t: "t \ \ x" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and t_\: "\ t = TComp g T" and t_fun: "is_Fun t" shows "\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = m \ \ \ is_Fun m" proof - obtain T'' where T'': "t = Fun g T''" using t_\ t_fun fun_type_id_eq by blast have g: "arity g > 0" using t_\ fun_type_inv[of t] by simp_all have "TComp g T \ \ (Var x)" using \_wt t t_\ by (metis wf_\x subtermeq_imp_subtermtypeeq wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) then obtain U where U: "Fun g U \ M" "\ (Fun g U) = TComp g T" "\u \ set U. \y. u = Var y" "distinct U" "length T'' = length U" using is_TComp_var_instance_closedD'[OF y_ex _ closed wf_M] by (metis t_\ T'' fun_type_id_eq fun_type_length_eq is_VarE) hence UT': "T = map \ U" using fun_type[OF g, of U] by simp show ?thesis using TComp_var_instance_wt_subst_exists UT' T'' U(1,3,4) t t_\ wf_\x wf_trm_subtermeq by (metis term.disc(2)) qed lemma TComp_var_and_subterm_instance_closed_has_subterms_instances: assumes M_var_inst_cl: "is_TComp_var_instance_closed \ M" and M_subterms_cl: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" and M_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and t: "t \\<^sub>s\<^sub>e\<^sub>t M" and s: "s \ t \ \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ s = m \ \" using subterm_subst_unfold[OF s] proof assume "\s'. s' \ t \ s = s' \ \" then obtain s' where s': "s' \ t" "s = s' \ \" by blast then obtain \ where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "s' \ M \\<^sub>s\<^sub>e\<^sub>t \" using t has_all_wt_instances_ofD'[OF wf_trms_subterms[OF M_wf] M_wf M_subterms_cl] term.order_trans[of s' t] by blast then obtain m where m: "m \ M" "s' = m \ \" by blast have "s = m \ (\ \\<^sub>s \)" using s'(2) m(2) by simp thus ?thesis using m(1) wt_subst_compose[OF \(1) \(1)] wf_trms_subst_compose[OF \(2) \(2)] by blast next assume "\x \ fv t. s \ \ x" then obtain x where x: "x \ fv t" "s \ \ x" "s \ \ x" by blast note 0 = TComp_var_instance_closed_has_Var[OF M_var_inst_cl M_wf] note 1 = has_all_wt_instances_ofD''[OF wf_trms_subterms[OF M_wf] M_wf M_subterms_cl] have \x_wf: "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" and s_wf_trm: "wf\<^sub>t\<^sub>r\<^sub>m s" using \(2) wf_trm_subterm[OF _ x(2)] by fastforce+ have x_fv_ex: "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var x) = \ (Var y)" using x(1) s fv_subset_subterms[OF t] by auto obtain y where y: "y \ fv\<^sub>s\<^sub>e\<^sub>t M" "\ (Var y) = \ s" using 0[of \ x s, OF \x_wf x_fv_ex x(3) \(1)] by metis then obtain z where z: "Var z \ M" "\ (Var z) = \ s" using 1[of y] vars_iff_subtermeq_set[of y M] by metis define \ where "\ \ Var(z := s)::('fun, ('fun, 'atom) term \ nat) subst" have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "s = Var z \ \" using z(2) s_wf_trm unfolding \_def wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by force+ thus ?thesis using z(1) by blast qed context begin private lemma SMP_D_aux1: assumes "t \ SMP M" and closed: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" "is_TComp_var_instance_closed \ M" and wf_M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" shows "\x \ fv t. \y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var y) = \ (Var x)" using assms(1) proof (induction t rule: SMP.induct) case (MP t) show ?case proof fix x assume x: "x \ fv t" hence "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t M" using MP.hyps vars_iff_subtermeq by fastforce then obtain \ s where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and s: "s \ M" "Var x = s \ \" using has_all_wt_instances_ofD'[OF wf_trms_subterms[OF wf_M] wf_M closed(1)] by blast then obtain y where y: "s = Var y" by (cases s) auto thus "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var y) = \ (Var x)" using s wt_subst_trm''[OF \(1), of "Var y"] by force qed next case (Subterm t t') hence "fv t' \ fv t" using subtermeq_vars_subset by auto thus ?case using Subterm.IH by auto next case (Substitution t \) note IH = Substitution.IH show ?case proof fix x assume x: "x \ fv (t \ \)" then obtain y where y: "y \ fv t" "\ (Var x) \ \ (Var y)" using Substitution.hyps(2,3) by (metis subst_apply_img_var subtermeqI' subtermeq_imp_subtermtypeeq vars_iff_subtermeq wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def wf_trm_subst_rangeD) let ?P = "\x. \y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var y) = \ (Var x)" show "?P x" using y IH proof (induction "\ (Var y)" arbitrary: y t) case (Var a) hence "\ (Var x) = \ (Var y)" by auto thus ?case using Var(2,4) by auto next case (Fun f T) obtain z where z: "\w \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var z) = \ (Var w)" "\ (Var z) = \ (Var y)" using Fun.prems(1,3) by blast show ?case proof (cases "\ (Var x) = \ (Var y)") case True thus ?thesis using Fun.prems by auto next case False then obtain \ where \: "\ \ set T" "\ (Var x) \ \" using Fun.prems(2) Fun.hyps(2) by auto then obtain U where U: "Fun f U \ M" "\ (Fun f U) = \ (Var z)" "\u \ set U. \v. u = Var v" "distinct U" using is_TComp_var_instance_closedD'[OF z(1) _ closed(2) wf_M] Fun.hyps(2) z(2) by (metis fun_type_id_eq subtermeqI' is_VarE) hence 1: "\x \ fv (Fun f U). \y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var y) = \ (Var x)" by force have "arity f > 0" using U(2) z(2) Fun.hyps(2) fun_type_inv(1) by metis hence "\ (Fun f U) = TComp f (map \ U)" using fun_type by auto then obtain u where u: "Var u \ set U" "\ (Var u) = \" using \(1) U(2,3) z(2) Fun.hyps(2) by auto show ?thesis using Fun.hyps(1)[of u "Fun f U"] u \ 1 by force qed qed qed next case (Ana t K T k) have "fv k \ fv t" using Ana_keys_fv[OF Ana.hyps(2)] Ana.hyps(3) by auto thus ?case using Ana.IH by auto qed private lemma SMP_D_aux2: fixes t::"('fun, ('fun, 'atom) term \ nat) term" assumes t_SMP: "t \ SMP M" and t_Var: "\x. t = Var x" and M_SMP_repr: "finite_SMP_representation arity Ana \ M" shows "\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = m \ \" proof - have M_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and M_var_inst_cl: "is_TComp_var_instance_closed \ M" and M_subterms_cl: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" and M_Ana_cl: "has_all_wt_instances_of \ (\((set \ fst \ Ana) ` M)) M" using finite_SMP_representationD[OF M_SMP_repr] by blast+ have M_Ana_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\ ((set \ fst \ Ana) ` M))" proof fix k assume "k \ \((set \ fst \ Ana) ` M)" then obtain m where m: "m \ M" "k \ set (fst (Ana m))" by force thus "wf\<^sub>t\<^sub>r\<^sub>m k" using M_wf Ana_keys_wf'[of m "fst (Ana m)" _ k] surjective_pairing by blast qed note 0 = has_all_wt_instances_ofD'[OF wf_trms_subterms[OF M_wf] M_wf M_subterms_cl] note 1 = has_all_wt_instances_ofD'[OF M_Ana_wf M_wf M_Ana_cl] obtain x y where x: "t = Var x" and y: "y \ fv\<^sub>s\<^sub>e\<^sub>t M" "\ (Var y) = \ (Var x)" using t_Var SMP_D_aux1[OF t_SMP M_subterms_cl M_var_inst_cl M_wf] by fastforce then obtain m \ where m: "m \ M" "m \ \ = Var y" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using 0[of "Var y"] vars_iff_subtermeq_set[of y M] by fastforce obtain z where z: "m = Var z" using m(2) by (cases m) auto define \ where "\ \ Var(z := Var x)::('fun, ('fun, 'atom) term \ nat) subst" have "\ (Var z) = \ (Var x)" using y(2) m(2) z wt_subst_trm''[OF \, of m] by argo hence "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" unfolding \_def wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by force+ moreover have "t = m \ \" using x z unfolding \_def by simp ultimately show ?thesis using m(1) by blast qed private lemma SMP_D_aux3: assumes hyps: "t' \ t" and wf_t: "wf\<^sub>t\<^sub>r\<^sub>m t" and prems: "is_Fun t'" and IH: "((\f. t = Fun f []) \ (\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = m \ \)) \ (\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = m \ \ \ is_Fun m)" and M_SMP_repr: "finite_SMP_representation arity Ana \ M" shows "((\f. t' = Fun f []) \ (\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t' = m \ \)) \ (\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t' = m \ \ \ is_Fun m)" proof (cases "\f. t = Fun f [] \ t' = Fun f []") case True have M_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and M_var_inst_cl: "is_TComp_var_instance_closed \ M" and M_subterms_cl: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" and M_Ana_cl: "has_all_wt_instances_of \ (\((set \ fst \ Ana) ` M)) M" using finite_SMP_representationD[OF M_SMP_repr] by blast+ note 0 = has_all_wt_instances_ofD'[OF wf_trms_subterms[OF M_wf] M_wf M_subterms_cl] note 1 = TComp_var_instance_closed_has_Fun[OF M_var_inst_cl M_wf] note 2 = TComp_var_and_subterm_instance_closed_has_subterms_instances[ OF M_var_inst_cl M_subterms_cl M_wf] have wf_t': "wf\<^sub>t\<^sub>r\<^sub>m t'" using hyps wf_t wf_trm_subterm by blast obtain c where "t = Fun c [] \ t' = Fun c []" using True by moura thus ?thesis proof assume c: "t' = Fun c []" show ?thesis proof (cases "\f. t = Fun f []") case True hence "t = t'" using c hyps by force thus ?thesis using IH by fast next case False note F = this then obtain m \ where m: "m \ M" "t = m \ \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using IH by blast show ?thesis using subterm_subst_unfold[OF hyps[unfolded m(2)]] proof assume "\m'. m' \ m \ t' = m' \ \" then obtain m' where m': "m' \ m" "t' = m' \ \" by moura obtain n \ where n: "n \ M" "m' = n \ \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using 0[of m'] m(1) m'(1) by blast have "t' = n \ (\ \\<^sub>s \)" using m'(2) n(2) by auto thus ?thesis using c n(1) wt_subst_compose[OF \(1) \(1)] wf_trms_subst_compose[OF \(2) \(2)] by blast next assume "\x \ fv m. t' \ \ x" then obtain x where x: "x \ fv m" "t' \ \ x" "t' \ \ x" by moura have \x_wf: "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" using \(2) by fastforce have x_fv_ex: "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var x) = \ (Var y)" using x m by auto show ?thesis proof (cases "\ t'") case (Var a) show ?thesis using c m 2[OF _ hyps[unfolded m(2)] \] by fast next case (Fun g S) show ?thesis using c 1[of \ x t', OF \x_wf x_fv_ex x(3) \(1) Fun] by blast qed qed qed qed (use IH hyps in simp) next case False note F = False then obtain m \ where m: "m \ M" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "t = m \ \" "is_Fun m" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using IH by moura obtain f T where fT: "t' = Fun f T" "arity f > 0" "\ t' = TComp f (map \ T)" using F prems fun_type wf_trm_subtermeq[OF wf_t hyps] by (metis is_FunE length_greater_0_conv subtermeqI' wf\<^sub>t\<^sub>r\<^sub>m_def) have closed: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" "is_TComp_var_instance_closed \ M" using M_SMP_repr unfolding finite_SMP_representation_def by metis+ have M_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" using finite_SMP_representationD[OF M_SMP_repr] by blast show ?thesis proof (cases "\x \ fv m. t' \ \ x") case True then obtain x where x: "x \ fv m" "t' \ \ x" by moura have 1: "x \ fv\<^sub>s\<^sub>e\<^sub>t M" using m(1) x(1) by auto have 2: "is_Fun (\ x)" using prems x(2) by auto have 3: "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" using m(5) by (simp add: wf_trm_subst_rangeD) have "\(\f. \ x = Fun f [])" using F x(2) by auto hence "\f T. \ (Var x) = TComp f T" using 2 3 m(2) by (metis (no_types) fun_type is_FunE length_greater_0_conv subtermeqI' wf\<^sub>t\<^sub>r\<^sub>m_def wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) moreover have "\f T. \ t' = Fun f T" using False prems wf_trm_subtermeq[OF wf_t hyps] by (metis (no_types) fun_type is_FunE length_greater_0_conv subtermeqI' wf\<^sub>t\<^sub>r\<^sub>m_def) ultimately show ?thesis using TComp_var_instance_closed_has_Fun 1 x(2) m(2) prems closed 3 M_wf by metis next case False then obtain m' where m': "m' \ m" "t' = m' \ \" "is_Fun m'" using hyps m(3) subterm_subst_not_img_subterm by blast then obtain \ m'' where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "m'' \ M" "m' = m'' \ \" using m(1) has_all_wt_instances_ofD'[OF wf_trms_subterms[OF M_wf] M_wf closed(1)] by blast hence t'_m'': "t' = m'' \ \ \\<^sub>s \" using m'(2) by fastforce note \\ = wt_subst_compose[OF \(1) m(2)] wf_trms_subst_compose[OF \(2) m(5)] show ?thesis proof (cases "is_Fun m''") case True thus ?thesis using \(3,4) m'(2,3) m(4) fT t'_m'' \\ by blast next case False then obtain x where x: "m'' = Var x" by moura hence "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var x) = \ (Var y)" "t' \ (\ \\<^sub>s \) x" "\ (Var x) = Fun f (map \ T)" "wf\<^sub>t\<^sub>r\<^sub>m ((\ \\<^sub>s \) x)" - using \\ t'_m'' \(3) fv_subset[OF \(3)] fT(3) subst_apply_term.simps(1)[of x "\ \\<^sub>s \"] + using \\ t'_m'' \(3) fv_subset[OF \(3)] fT(3) eval_term.simps(1)[of _ "\ \\<^sub>s \"] wt_subst_trm''[OF \\(1), of "Var x"] - by (fastforce, blast, argo, fastforce) + by force+ thus ?thesis using x TComp_var_instance_closed_has_Fun[ of M "\ \\<^sub>s \" x t' f "map \ T", OF closed(2) M_wf _ _ _ \\(1) fT(3) prems] by blast qed qed qed lemma SMP_D: assumes "t \ SMP M" "is_Fun t" and M_SMP_repr: "finite_SMP_representation arity Ana \ M" shows "((\f. t = Fun f []) \ (\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = m \ \)) \ (\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = m \ \ \ is_Fun m)" proof - have wf_M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and closed: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" "has_all_wt_instances_of \ (\((set \ fst \ Ana) ` M)) M" "is_TComp_var_instance_closed \ M" using finite_SMP_representationD[OF M_SMP_repr] by blast+ show ?thesis using assms(1,2) proof (induction t rule: SMP.induct) case (MP t) moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range Var)" "t = t \ Var" by simp_all ultimately show ?case by blast next case (Subterm t t') hence t_fun: "is_Fun t" by auto note * = Subterm.hyps(2) SMP_wf_trm[OF Subterm.hyps(1) wf_M(1)] Subterm.prems Subterm.IH[OF t_fun] M_SMP_repr show ?case by (rule SMP_D_aux3[OF *]) next case (Substitution t \) have wf: "wf\<^sub>t\<^sub>r\<^sub>m t" by (metis Substitution.hyps(1) wf_M(1) SMP_wf_trm) hence wf': "wf\<^sub>t\<^sub>r\<^sub>m (t \ \)" using Substitution.hyps(3) wf_trm_subst by blast show ?case proof (cases "\ t") case (Var a) hence 1: "\ (t \ \) = TAtom a" using Substitution.hyps(2) by (metis wt_subst_trm'') then obtain c where c: "t \ \ = Fun c []" using TAtom_term_cases[OF wf' 1] Substitution.prems by fastforce hence "(\x. t = Var x) \ t = t \ \" by (cases t) auto thus ?thesis proof assume t_Var: "\x. t = Var x" then obtain x where x: "t = Var x" "\ x = Fun c []" "\ (Var x) = TAtom a" using c 1 wt_subst_trm''[OF Substitution.hyps(2), of t] by force obtain m \ where m: "m \ M" "t = m \ \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using SMP_D_aux2[OF Substitution.hyps(1) t_Var M_SMP_repr] by moura have "m \ (\ \\<^sub>s \) = Fun c []" using c m(2) by auto thus ?thesis using c m(1) wt_subst_compose[OF \(1) Substitution.hyps(2)] wf_trms_subst_compose[OF \(2) Substitution.hyps(3)] by metis qed (use c Substitution.IH in auto) next case (Fun f T) hence 1: "\ (t \ \) = TComp f T" using Substitution.hyps(2) by (metis wt_subst_trm'') have 2: "\(\f. t = Fun f [])" using Fun TComp_term_cases[OF wf] by auto obtain T'' where T'': "t \ \ = Fun f T''" using 1 2 fun_type_id_eq Fun Substitution.prems by fastforce have f: "arity f > 0" using fun_type_inv[OF 1] by metis show ?thesis proof (cases t) case (Fun g U) then obtain m \ where m: "m \ M" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "t = m \ \" "is_Fun m" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using Substitution.IH Fun 2 by moura have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" "t \ \ = m \ (\ \\<^sub>s \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using wt_subst_compose[OF m(2) Substitution.hyps(2)] m(3) wf_trms_subst_compose[OF m(5) Substitution.hyps(3)] by auto thus ?thesis using m(1,4) by metis next case (Var x) then obtain y where y: "y \ fv\<^sub>s\<^sub>e\<^sub>t M" "\ (Var y) = \ (Var x)" using SMP_D_aux1[OF Substitution.hyps(1) closed(1,3) wf_M] Fun by moura hence 3: "\ (Var y) = TComp f T" using Var Fun \_Var_fst by simp obtain h V where V: "Fun h V \ M" "\ (Fun h V) = \ (Var y)" "\u \ set V. \z. u = Var z" "distinct V" by (metis is_VarE is_TComp_var_instance_closedD[OF _ 3 closed(3)] y(1)) moreover have "length T'' = length V" using 3 V(2) fun_type_length_eq 1 T'' by metis ultimately have TV: "T = map \ V" by (metis fun_type[OF f(1)] 3 fun_type_id_eq term.inject(2)) obtain \ where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t \ \ = Fun h V \ \" using TComp_var_instance_wt_subst_exists 1 3 T'' TV V(2,3,4) wf' by (metis fun_type_id_eq) have 9: "\ (Fun h V) = \ (\ x)" using y(2) Substitution.hyps(2) V(2) 1 3 Var by auto show ?thesis using Var \ 9 V(1) by force qed qed next case (Ana t K T k) have 1: "is_Fun t" using Ana.hyps(2,3) by auto then obtain f U where U: "t = Fun f U" by moura have 2: "fv k \ fv t" using Ana_keys_fv[OF Ana.hyps(2)] Ana.hyps(3) by auto have wf_t: "wf\<^sub>t\<^sub>r\<^sub>m t" using SMP_wf_trm[OF Ana.hyps(1)] wf\<^sub>t\<^sub>r\<^sub>m_code wf_M by auto hence wf_k: "wf\<^sub>t\<^sub>r\<^sub>m k" using Ana_keys_wf'[OF Ana.hyps(2)] wf\<^sub>t\<^sub>r\<^sub>m_code Ana.hyps(3) by auto have wf_M_keys: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\((set \ fst \ Ana) ` M))" proof fix t assume "t \ (\((set \ fst \ Ana) ` M))" then obtain s where s: "s \ M" "t \ (set \ fst \ Ana) s" by blast obtain K R where KR: "Ana s = (K,R)" by (metis surj_pair) hence "t \ set K" using s(2) by simp thus "wf\<^sub>t\<^sub>r\<^sub>m t" using Ana_keys_wf'[OF KR] wf_M s(1) by blast qed show ?case using Ana_subst'_or_Ana_keys_subterm proof assume "\t K T k. Ana t = (K, T) \ k \ set K \ k \ t" hence *: "k \ t" using Ana.hyps(2,3) by auto show ?thesis by (rule SMP_D_aux3[OF * wf_t Ana.prems Ana.IH[OF 1] M_SMP_repr]) next assume Ana_subst': "\f T \ K M. Ana (Fun f T) = (K, M) \ Ana (Fun f T \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" have "arity f > 0" using Ana_const[of f U] U Ana.hyps(2,3) by fastforce hence "U \ []" using wf_t U unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by force then obtain m \ where m: "m \ M" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t = m \ \" "is_Fun m" using Ana.IH[OF 1] U by auto hence "Ana (t \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \,T \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Ana_subst' U Ana.hyps(2) by auto obtain Km Tm where Ana_m: "Ana m = (Km,Tm)" by moura hence "Ana (m \ \) = (Km \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \,Tm \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Ana_subst' U m(4) is_FunE[OF m(5)] Ana.hyps(2) by metis then obtain km where km: "km \ set Km" "k = km \ \" using Ana.hyps(2,3) m(4) by auto then obtain \ km' where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and km': "km' \ M" "km = km' \ \" using Ana_m m(1) has_all_wt_instances_ofD'[OF wf_M_keys wf_M closed(2), of km] by force have k\\: "k = km' \ \ \\<^sub>s \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using km(2) km' wt_subst_compose[OF \(1) m(2)] wf_trms_subst_compose[OF \(2) m(3)] by auto show ?case proof (cases "is_Fun km'") case True thus ?thesis using k\\ km'(1) by blast next case False note F = False then obtain x where x: "km' = Var x" by auto hence 3: "x \ fv\<^sub>s\<^sub>e\<^sub>t M" using fv_subset[OF km'(1)] by auto obtain kf kT where kf: "k = Fun kf kT" using Ana.prems by auto show ?thesis proof (cases "kT = []") case True thus ?thesis using k\\(1) k\\(2) k\\(3) kf km'(1) by blast next case False hence 4: "arity kf > 0" using wf_k kf TAtom_term_cases const_type by fastforce then obtain kT' where kT': "\ k = TComp kf kT'" by (simp add: fun_type kf) then obtain V where V: "Fun kf V \ M" "\ (Fun kf V) = \ (Var x)" "\u \ set V. \v. u = Var v" "distinct V" "is_Fun (Fun kf V)" using is_TComp_var_instance_closedD[OF _ _ closed(3), of x] x m(2) k\\(1) 3 wt_subst_trm''[OF k\\(2)] by (metis fun_type_id_eq term.disc(2) is_VarE) have 5: "kT' = map \ V" using fun_type[OF 4] x kT' k\\ m(2) V(2) by (metis term.inject(2) wt_subst_trm'') thus ?thesis using TComp_var_instance_wt_subst_exists wf_k kf 4 V(3,4) kT' V(1,5) by metis qed qed qed qed qed lemma SMP_D': fixes M defines "\ \ var_rename (max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M))" assumes M_SMP_repr: "finite_SMP_representation arity Ana \ M" and s: "s \ SMP M" "is_Fun s" "\f. s = Fun f []" and t: "t \ SMP M" "is_Fun t" "\f. t = Fun f []" obtains \ s0 \ t0 where "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "s0 \ M" "is_Fun s0" "s = s0 \ \" "\ s = \ s0" and "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t0 \ M" "is_Fun t0" "t = t0 \ \ \ \" "\ t = \ t0" proof - obtain \ s0 where s0: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "s0 \ M" "s = s0 \ \" "is_Fun s0" using s(3) SMP_D[OF s(1,2) M_SMP_repr] unfolding \_def by metis obtain \ t0 where t0: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t0 \ M" "t = t0 \ \ \ \" "is_Fun t0" using t(3) SMP_D[OF t(1,2) M_SMP_repr] var_rename_wt'[of _ t] wf_trms_subst_compose_Var_range(1)[OF _ var_rename_is_renaming(2)] unfolding \_def by metis have "\ s = \ s0" "\ t = \ (t0 \ \)" "\ (t0 \ \) = \ t0" using s0 t0 wt_subst_trm'' by (metis, metis, metis \_def var_rename_wt(1)) thus ?thesis using s0 t0 that by simp qed lemma SMP_D'': fixes t::"('fun, ('fun, 'atom) term \ nat) term" assumes t_SMP: "t \ SMP M" and M_SMP_repr: "finite_SMP_representation arity Ana \ M" shows "\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = m \ \" proof (cases "(\x. t = Var x) \ (\c. t = Fun c [])") case True have M_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and M_var_inst_cl: "is_TComp_var_instance_closed \ M" and M_subterms_cl: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" and M_Ana_cl: "has_all_wt_instances_of \ (\((set \ fst \ Ana) ` M)) M" using finite_SMP_representationD[OF M_SMP_repr] by blast+ have M_Ana_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\ ((set \ fst \ Ana) ` M))" proof fix k assume "k \ \((set \ fst \ Ana) ` M)" then obtain m where m: "m \ M" "k \ set (fst (Ana m))" by force thus "wf\<^sub>t\<^sub>r\<^sub>m k" using M_wf Ana_keys_wf'[of m "fst (Ana m)" _ k] surjective_pairing by blast qed show ?thesis using True proof assume "\x. t = Var x" then obtain x y where x: "t = Var x" and y: "y \ fv\<^sub>s\<^sub>e\<^sub>t M" "\ (Var y) = \ (Var x)" using SMP_D_aux1[OF t_SMP M_subterms_cl M_var_inst_cl M_wf] by fastforce then obtain m \ where m: "m \ M" "m \ \ = Var y" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using has_all_wt_instances_ofD'[OF wf_trms_subterms[OF M_wf] M_wf M_subterms_cl, of "Var y"] vars_iff_subtermeq_set[of y M] by fastforce obtain z where z: "m = Var z" using m(2) by (cases m) auto define \ where "\ \ Var(z := Var x)::('fun, ('fun, 'atom) term \ nat) subst" have "\ (Var z) = \ (Var x)" using y(2) m(2) z wt_subst_trm''[OF \, of m] by argo hence "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" unfolding \_def wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by force+ moreover have "t = m \ \" using x z unfolding \_def by simp ultimately show ?thesis using m(1) by blast qed (use SMP_D[OF t_SMP _ M_SMP_repr] in blast) qed (use SMP_D[OF t_SMP _ M_SMP_repr] in blast) end lemma tfr\<^sub>s\<^sub>e\<^sub>t_if_comp_tfr\<^sub>s\<^sub>e\<^sub>t: assumes "comp_tfr\<^sub>s\<^sub>e\<^sub>t arity Ana \ M" shows "tfr\<^sub>s\<^sub>e\<^sub>t M" proof - let ?\ = "var_rename (max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M))" have M_SMP_repr: "finite_SMP_representation arity Ana \ M" by (metis comp_tfr\<^sub>s\<^sub>e\<^sub>t_def assms) have M_finite: "finite M" using assms card_gt_0_iff unfolding comp_tfr\<^sub>s\<^sub>e\<^sub>t_def finite_SMP_representation_def by blast show ?thesis proof (unfold tfr\<^sub>s\<^sub>e\<^sub>t_def; intro ballI impI) fix s t assume "s \ SMP M - Var`\" "t \ SMP M - Var`\" hence st: "s \ SMP M" "is_Fun s" "t \ SMP M" "is_Fun t" by auto have "\(\\. Unifier \ s t)" when st_type_neq: "\ s \ \ t" proof (cases "\f. s = Fun f [] \ t = Fun f []") case False then obtain \ s0 \ t0 where s0: "s0 \ M" "is_Fun s0" "s = s0 \ \" "\ s = \ s0" and t0: "t0 \ M" "is_Fun t0" "t = t0 \ ?\ \ \" "\ t = \ t0" using SMP_D'[OF M_SMP_repr st(1,2) _ st(3,4)] by metis hence "\(\\. Unifier \ s0 (t0 \ ?\))" using assms mgu_None_is_subst_neq st_type_neq wt_subst_trm''[OF var_rename_wt(1)] unfolding comp_tfr\<^sub>s\<^sub>e\<^sub>t_def Let_def by metis thus ?thesis using vars_term_disjoint_imp_unifier[OF var_rename_fv_set_disjoint[OF M_finite]] s0(1) t0(1) unfolding s0(3) t0(3) by (metis (no_types, opaque_lifting) subst_subst_compose) qed (use st_type_neq st(2,4) in auto) thus "\ s = \ t" when "\\. Unifier \ s t" by (metis that) qed qed lemma tfr\<^sub>s\<^sub>e\<^sub>t_if_comp_tfr\<^sub>s\<^sub>e\<^sub>t': assumes "let N = SMP0 Ana \ M in set M \ set N \ comp_tfr\<^sub>s\<^sub>e\<^sub>t arity Ana \ (set N)" shows "tfr\<^sub>s\<^sub>e\<^sub>t (set M)" by (rule tfr_subset(2)[ OF tfr\<^sub>s\<^sub>e\<^sub>t_if_comp_tfr\<^sub>s\<^sub>e\<^sub>t[OF conjunct2[OF assms[unfolded Let_def]]] conjunct1[OF assms[unfolded Let_def]]]) lemma tfr\<^sub>s\<^sub>t\<^sub>p_is_comp_tfr\<^sub>s\<^sub>t\<^sub>p: "tfr\<^sub>s\<^sub>t\<^sub>p a = comp_tfr\<^sub>s\<^sub>t\<^sub>p \ a" proof (cases a) case (Equality ac t t') thus ?thesis using mgu_always_unifies[of t _ t'] mgu_gives_MGU[of t t'] by auto next case (Inequality X F) thus ?thesis using tfr\<^sub>s\<^sub>t\<^sub>p.simps(2)[of X F] comp_tfr\<^sub>s\<^sub>t\<^sub>p.simps(2)[of \ X F] Fun_range_case(2)[of "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F)"] unfolding is_Var_def by auto qed auto lemma tfr\<^sub>s\<^sub>t_if_comp_tfr\<^sub>s\<^sub>t: assumes "comp_tfr\<^sub>s\<^sub>t arity Ana \ M S" shows "tfr\<^sub>s\<^sub>t S" unfolding tfr\<^sub>s\<^sub>t_def proof have comp_tfr\<^sub>s\<^sub>e\<^sub>t_M: "comp_tfr\<^sub>s\<^sub>e\<^sub>t arity Ana \ M" using assms unfolding comp_tfr\<^sub>s\<^sub>t_def by blast have wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_S: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" and S_trms_instance_M: "has_all_wt_instances_of \ (trms\<^sub>s\<^sub>t S) M" using assms wf\<^sub>t\<^sub>r\<^sub>m_code wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code trms_list\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>t unfolding comp_tfr\<^sub>s\<^sub>t_def comp_tfr\<^sub>s\<^sub>e\<^sub>t_def finite_SMP_representation_def list_all_iff by blast+ show "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" using tfr_subset(3)[OF tfr\<^sub>s\<^sub>e\<^sub>t_if_comp_tfr\<^sub>s\<^sub>e\<^sub>t[OF comp_tfr\<^sub>s\<^sub>e\<^sub>t_M] SMP_SMP_subset] SMP_I'[OF wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_S wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_M S_trms_instance_M] by blast have "list_all (comp_tfr\<^sub>s\<^sub>t\<^sub>p \) S" by (metis assms comp_tfr\<^sub>s\<^sub>t_def) thus "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" by (induct S) (simp_all add: tfr\<^sub>s\<^sub>t\<^sub>p_is_comp_tfr\<^sub>s\<^sub>t\<^sub>p) qed lemma tfr\<^sub>s\<^sub>t_if_comp_tfr\<^sub>s\<^sub>t': assumes "comp_tfr\<^sub>s\<^sub>t arity Ana \ (set (SMP0 Ana \ (trms_list\<^sub>s\<^sub>t S))) S" shows "tfr\<^sub>s\<^sub>t S" by (rule tfr\<^sub>s\<^sub>t_if_comp_tfr\<^sub>s\<^sub>t[OF assms]) subsubsection \Lemmata for Checking Ground SMP (GSMP) Disjointness\ context begin private lemma ground_SMP_disjointI_aux1: fixes M::"('fun, ('fun, 'atom) term \ nat) term set" assumes f_def: "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and g_def: "g \ \M. {t \ M. fv t = {}}" shows "f (SMP M) = g (SMP M)" proof have "t \ f (SMP M)" when t: "t \ SMP M" "fv t = {}" for t proof - define \ where "\ \ Var::('fun, ('fun, 'atom) term \ nat) subst" have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t = t \ \" using subst_apply_term_empty[of t] that(2) wt_subst_Var wf_trm_subst_range_Var unfolding \_def by auto thus ?thesis using SMP.Substitution[OF t(1), of \] t(2) unfolding f_def by fastforce qed thus "g (SMP M) \ f (SMP M)" unfolding g_def by blast qed (use f_def g_def in blast) private lemma ground_SMP_disjointI_aux2: fixes M::"('fun, ('fun, 'atom) term \ nat) term set" assumes f_def: "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and M_SMP_repr: "finite_SMP_representation arity Ana \ M" shows "f M = f (SMP M)" proof have M_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and M_var_inst_cl: "is_TComp_var_instance_closed \ M" and M_subterms_cl: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" and M_Ana_cl: "has_all_wt_instances_of \ (\((set \ fst \ Ana) ` M)) M" using finite_SMP_representationD[OF M_SMP_repr] by blast+ show "f (SMP M) \ f M" proof fix t assume "t \ f (SMP M)" then obtain s \ where s: "t = s \ \" "s \ SMP M" "fv (s \ \) = {}" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" unfolding f_def by blast have t_wf: "wf\<^sub>t\<^sub>r\<^sub>m t" using SMP_wf_trm[OF s(2) M_wf] s(1) wf_trm_subst[OF \(2)] by blast obtain m \ where m: "m \ M" "s = m \ \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using SMP_D''[OF s(2) M_SMP_repr] by blast have "t = m \ (\ \\<^sub>s \)" "fv (m \ (\ \\<^sub>s \)) = {}" using s(1,3) m(2) by simp_all thus "t \ f M" using m(1) wt_subst_compose[OF \(1) \(1)] wf_trms_subst_compose[OF \(2) \(2)] unfolding f_def by blast qed qed (auto simp add: f_def) private lemma ground_SMP_disjointI_aux3: fixes A B C::"('fun, ('fun, 'atom) term \ nat) term set" defines "P \ \t s. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ Unifier \ t s" assumes f_def: "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and Q_def: "Q \ \t. intruder_synth' public arity {} t" and R_def: "R \ \t. \u \ C. is_wt_instance_of_cond \ t u" and AB: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s A" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s B" "fv\<^sub>s\<^sub>e\<^sub>t A \ fv\<^sub>s\<^sub>e\<^sub>t B = {}" and C: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s C" and ABC: "\t \ A. \s \ B. P t s \ Q t \ R t" shows "f A \ f B \ f C \ {m. {} \\<^sub>c m}" proof fix t assume "t \ f A \ f B" then obtain ta tb \a \b where ta: "t = ta \ \a" "ta \ A" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \a" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \a)" "fv (ta \ \a) = {}" and tb: "t = tb \ \b" "tb \ B" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \b" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \b)" "fv (tb \ \b) = {}" unfolding f_def by blast have ta_tb_wf: "wf\<^sub>t\<^sub>r\<^sub>m ta" "wf\<^sub>t\<^sub>r\<^sub>m tb" "fv ta \ fv tb = {}" "\ ta = \ tb" using ta(1,2) tb(1,2) AB fv_subset_subterms wt_subst_trm''[OF ta(3), of ta] wt_subst_trm''[OF tb(3), of tb] by (fast, fast, blast, simp) obtain \ where \: "Unifier \ ta tb" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using vars_term_disjoint_imp_unifier[OF ta_tb_wf(3), of \a \b] ta(1) tb(1) wt_Unifier_if_Unifier[OF ta_tb_wf(1,2,4)] by blast hence "Q ta \ R ta" using ABC ta(2) tb(2) unfolding P_def by blast+ thus "t \ f C \ {m. {} \\<^sub>c m}" proof show "Q ta \ ?thesis" using ta(1) pgwt_ground[of ta] pgwt_is_empty_synth[of ta] subst_ground_ident[of ta \a] unfolding Q_def f_def intruder_synth_code[symmetric] by simp next assume "R ta" then obtain ua \a where ua: "ta = ua \ \a" "ua \ C" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \a" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \a)" using \ ABC ta_tb_wf(1,2) ta(2) tb(2) C is_wt_instance_of_condD' unfolding P_def R_def by metis have "t = ua \ (\a \\<^sub>s \a)" "fv t = {}" using ua(1) ta(1,5) tb(1,5) by auto thus ?thesis using ua(2) wt_subst_compose[OF ua(3) ta(3)] wf_trms_subst_compose[OF ua(4) ta(4)] unfolding f_def by blast qed qed lemma ground_SMP_disjointI: fixes A B::"('fun, ('fun, 'atom) term \ nat) term set" and C defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and "g \ \M. {t \ M. fv t = {}}" and "Q \ \t. intruder_synth' public arity {} t" and "R \ \t. \u \ C. is_wt_instance_of_cond \ t u" assumes AB_fv_disj: "fv\<^sub>s\<^sub>e\<^sub>t A \ fv\<^sub>s\<^sub>e\<^sub>t B = {}" and A_SMP_repr: "finite_SMP_representation arity Ana \ A" and B_SMP_repr: "finite_SMP_representation arity Ana \ B" and C_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s C" and ABC: "\t \ A. \s \ B. \ t = \ s \ mgu t s \ None \ Q t \ R t" shows "g (SMP A) \ g (SMP B) \ f C \ {m. {} \\<^sub>c m}" proof - have AB_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s A" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s B" using A_SMP_repr B_SMP_repr finite_SMP_representationD(1) by (blast, blast) let ?P = "\t s. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ Unifier \ t s" have ABC': "\t \ A. \s \ B. ?P t s \ Q t \ R t" by (metis (no_types) ABC mgu_None_is_subst_neq wt_subst_trm'') show ?thesis using ground_SMP_disjointI_aux1[OF f_def g_def, of A] ground_SMP_disjointI_aux1[OF f_def g_def, of B] ground_SMP_disjointI_aux2[OF f_def A_SMP_repr] ground_SMP_disjointI_aux2[OF f_def B_SMP_repr] ground_SMP_disjointI_aux3[OF f_def Q_def R_def AB_wf AB_fv_disj C_wf ABC'] by argo qed end end end diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Typing_Result.thy b/thys/Stateful_Protocol_Composition_and_Typing/Typing_Result.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/Typing_Result.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/Typing_Result.thy @@ -1,3665 +1,3665 @@ (* Title: Typing_Result.thy Author: Andreas Viktor Hess, DTU SPDX-License-Identifier: BSD-3-Clause *) section \The Typing Result\ text\\label{sec:Typing-Result}\ theory Typing_Result imports Typed_Model begin subsection \Locale Setup\ locale typing_result = typed_model arity public Ana \ for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,'var) term \ (('fun,'var) term list \ ('fun,'var) term list)" and \::"('fun,'var) term \ ('fun,'atom::finite) term_type" + assumes infinite_typed_consts: "\a. infinite {c. \ (Fun c []) = TAtom a \ public c}" and no_private_funs[simp]: "\f. arity f > 0 \ public f" begin subsubsection \Minor Lemmata\ lemma fun_type_inv': assumes "\ t = TComp f T" shows "arity f > 0" "public f" using assms fun_type_inv by simp_all lemma infinite_public_consts[simp]: "infinite {c. public c \ arity c = 0}" proof - fix a::'atom define A where "A \ {c. \ (Fun c []) = TAtom a \ public c}" define B where "B \ {c. public c \ arity c = 0}" have "arity c = 0" when c: "c \ A" for c using c const_type_inv unfolding A_def by blast hence "A \ B" unfolding A_def B_def by blast hence "infinite B" using infinite_typed_consts[of a, unfolded A_def[symmetric]] by (metis infinite_super) thus ?thesis unfolding B_def by blast qed lemma infinite_fun_syms[simp]: "infinite {c. public c \ arity c > 0} \ infinite \\<^sub>f" "infinite \" "infinite \\<^sub>p\<^sub>u\<^sub>b" "infinite (UNIV::'fun set)" by (metis \\<^sub>f_unfold finite_Collect_conjI, metis infinite_public_consts finite_Collect_conjI, use infinite_public_consts \pub_unfold in \force simp add: Collect_conj_eq\, metis UNIV_I finite_subset subsetI infinite_public_consts(1)) lemma id_univ_proper_subset[simp]: "\\<^sub>f \ UNIV" "(\f. arity f > 0) \ \ \ UNIV" by (metis finite.emptyI inf_top.right_neutral top.not_eq_extremum disjoint_fun_syms infinite_fun_syms(2) inf_commute) (metis top.not_eq_extremum UNIV_I const_arity_eq_zero less_irrefl) lemma exists_fun_notin_funs_term: "\f::'fun. f \ funs_term t" by (metis UNIV_eq_I finite_fun_symbols infinite_fun_syms(4)) lemma exists_fun_notin_funs_terms: assumes "finite M" shows "\f::'fun. f \ \(funs_term ` M)" by (metis assms finite_fun_symbols infinite_fun_syms(4) ex_new_if_finite finite_UN) lemma exists_notin_funs\<^sub>s\<^sub>t: "\f. f \ funs\<^sub>s\<^sub>t (S::('fun,'var) strand)" by (metis UNIV_eq_I finite_funs\<^sub>s\<^sub>t infinite_fun_syms(4)) lemma infinite_typed_consts': "infinite {c. \ (Fun c []) = TAtom a \ public c \ arity c = 0}" proof - { fix c assume "\ (Fun c []) = TAtom a" "public c" hence "arity c = 0" using const_type[of c] fun_type[of c "[]"] by auto } hence "{c. \ (Fun c []) = TAtom a \ public c \ arity c = 0} = {c. \ (Fun c []) = TAtom a \ public c}" by auto thus "infinite {c. \ (Fun c []) = TAtom a \ public c \ arity c = 0}" using infinite_typed_consts[of a] by metis qed lemma atypes_inhabited: "\c. \ (Fun c []) = TAtom a \ wf\<^sub>t\<^sub>r\<^sub>m (Fun c []) \ public c \ arity c = 0" proof - obtain c where "\ (Fun c []) = TAtom a" "public c" "arity c = 0" using infinite_typed_consts'(1)[of a] not_finite_existsD by blast thus ?thesis using const_type_inv[OF \\ (Fun c []) = TAtom a\] unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto qed lemma atype_ground_term_ex: "\t. fv t = {} \ \ t = TAtom a \ wf\<^sub>t\<^sub>r\<^sub>m t" using atypes_inhabited[of a] by force lemma type_ground_inhabited: "\t'. fv t' = {} \ \ t = \ t'" proof - { fix \::"('fun, 'atom) term_type" assume "\f T. Fun f T \ \ \ 0 < arity f" hence "\t'. fv t' = {} \ \ = \ t'" proof (induction \) case (Fun f T) hence "arity f > 0" by auto from Fun.IH Fun.prems(1) have "\Y. map \ Y = T \ (\x \ set Y. fv x = {})" proof (induction T) case (Cons x X) hence "\g Y. Fun g Y \ Fun f X \ 0 < arity g" by auto hence "\Y. map \ Y = X \ (\x\set Y. fv x = {})" using Cons by auto moreover have "\t'. fv t' = {} \ x = \ t'" using Cons by auto ultimately obtain y Y where "fv y = {}" "\ y = x" "map \ Y = X" "\x\set Y. fv x = {}" using Cons by moura hence "map \ (y#Y) = x#X \ (\x\set (y#Y). fv x = {})" by auto thus ?case by meson qed simp then obtain Y where "map \ Y = T" "\x \ set Y. fv x = {}" by metis hence "fv (Fun f Y) = {}" "\ (Fun f Y) = TComp f T" using fun_type[OF \arity f > 0\] by auto thus ?case by (metis exI[of "\t. fv t = {} \ \ t = TComp f T" "Fun f Y"]) qed (metis atype_ground_term_ex) } thus ?thesis by (metis \_wf'') qed lemma type_wfttype_inhabited: assumes "\f T. Fun f T \ \ \ 0 < arity f" "wf\<^sub>t\<^sub>r\<^sub>m \" shows "\t. \ t = \ \ wf\<^sub>t\<^sub>r\<^sub>m t" using assms proof (induction \) case (Fun f Y) have IH: "\t. \ t = y \ wf\<^sub>t\<^sub>r\<^sub>m t" when y: "y \ set Y " for y proof - have "wf\<^sub>t\<^sub>r\<^sub>m y" using Fun y unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by (metis Fun_param_is_subterm term.le_less_trans) moreover have "Fun g Z \ y \ 0 < arity g" for g Z using Fun y by auto ultimately show ?thesis using Fun.IH[OF y] by auto qed from Fun have "arity f = length Y" "arity f > 0" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by force+ moreover from IH have "\X. map \ X = Y \ (\x \ set X. wf\<^sub>t\<^sub>r\<^sub>m x)" by (induct Y, simp_all, metis list.simps(9) set_ConsD) ultimately show ?case by (metis fun_type length_map wf_trmI) qed (use atypes_inhabited wf\<^sub>t\<^sub>r\<^sub>m_def in blast) lemma type_pgwt_inhabited: "wf\<^sub>t\<^sub>r\<^sub>m t \ \t'. \ t = \ t' \ public_ground_wf_term t'" proof - assume "wf\<^sub>t\<^sub>r\<^sub>m t" { fix \ assume "\ t = \" hence "\t'. \ t = \ t' \ public_ground_wf_term t'" using \wf\<^sub>t\<^sub>r\<^sub>m t\ proof (induction \ arbitrary: t) case (Var a t) then obtain c where "\ t = \ (Fun c [])" "arity c = 0" "public c" using const_type_inv[of _ "[]" a] infinite_typed_consts(1)[of a] not_finite_existsD by force thus ?case using PGWT[OF \public c\, of "[]"] by auto next case (Fun f Y t) have *: "arity f > 0" "public f" "arity f = length Y" using fun_type_inv[OF \\ t = TComp f Y\] fun_type_inv_wf[OF \\ t = TComp f Y\ \wf\<^sub>t\<^sub>r\<^sub>m t\] by auto have "\y. y \ set Y \ \t'. y = \ t' \ public_ground_wf_term t'" using Fun.prems(1) Fun.IH \_wf''[of _ _ t] \_wf'[OF \wf\<^sub>t\<^sub>r\<^sub>m t\] type_wfttype_inhabited by (metis Fun_param_is_subterm term.order_trans wf_trm_subtermeq) hence "\X. map \ X = Y \ (\x \ set X. public_ground_wf_term x)" by (induct Y, simp_all, metis list.simps(9) set_ConsD) then obtain X where X: "map \ X = Y" "\x. x \ set X \ public_ground_wf_term x" by moura hence "arity f = length X" using *(3) by auto have "\ t = \ (Fun f X)" "public_ground_wf_term (Fun f X)" using fun_type[OF *(1), of X] Fun.prems(1) X(1) apply simp using PGWT[OF *(2) \arity f = length X\ X(2)] by metis thus ?case by metis qed } thus ?thesis using \wf\<^sub>t\<^sub>r\<^sub>m t\ by auto qed end subsection \The Typing Result for the Composition-Only Intruder\ context typing_result begin subsubsection \Well-typedness and Type-Flaw Resistance Preservation\ context begin private lemma LI_preserves_tfr_stp_all_single: assumes "(S,\) \ (S',\')" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" shows "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" using assms proof (induction rule: LI_rel.induct) case (Compose S X f S' \) hence "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" by simp_all moreover have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (map Send1 X)" by (induct X) auto ultimately show ?case by simp next case (Unify S f Y \ X S' \) hence "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S')" by simp have "fv\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S') \ bvars\<^sub>s\<^sub>t (S@S') = {}" using Unify.prems(1) by (auto simp add: wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) moreover have "fv (Fun f X) \ fv\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S')" by auto moreover have "fv (Fun f Y) \ fv\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S')" using Unify.hyps(2) fv_subset_if_in_strand_ik'[of "Fun f Y" S] by force ultimately have bvars_disj: "bvars\<^sub>s\<^sub>t (S@S') \ fv (Fun f X) = {}" "bvars\<^sub>s\<^sub>t (S@S') \ fv (Fun f Y) = {}" by blast+ have "wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)" using Unify.prems(5) by simp moreover have "wf\<^sub>t\<^sub>r\<^sub>m (Fun f Y)" proof - obtain x where "x \ set S" "Fun f Y \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t\<^sub>p x)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p x)" using Unify.hyps(2) Unify.prems(5) by force+ thus ?thesis using wf_trm_subterm by auto qed moreover have "Fun f X \ SMP (trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S'))" "Fun f Y \ SMP (trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S'))" using SMP_append[of S "Send1 (Fun f X)#S'"] SMP_Cons[of "Send1 (Fun f X)" S'] SMP_ikI[OF Unify.hyps(2)] by auto hence "\ (Fun f X) = \ (Fun f Y)" using Unify.prems(4) mgu_gives_MGU[OF Unify.hyps(3)[symmetric]] unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast ultimately have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using mgu_wt_if_same_type[OF Unify.hyps(3)[symmetric]] by metis moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_wf_trm[OF Unify.hyps(3)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)\ \wf\<^sub>t\<^sub>r\<^sub>m (Fun f Y)\] by (metis wf_trm_subst_range_iff) moreover have "bvars\<^sub>s\<^sub>t (S@S') \ range_vars \ = {}" using mgu_vars_bounded[OF Unify.hyps(3)[symmetric]] bvars_disj by fast ultimately show ?case using tfr_stp_all_wt_subst_apply[OF \list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S')\] by metis next case (Equality S \ t t' a S' \) have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S')" "\ t = \ t'" using tfr_stp_all_same_type[of S a t t' S'] tfr_stp_all_split(5)[of S _ S'] MGU_is_Unifier[OF mgu_gives_MGU[OF Equality.hyps(2)[symmetric]]] Equality.prems(3) by blast+ moreover have "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m t'" using Equality.prems(5) by auto ultimately have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using mgu_wt_if_same_type[OF Equality.hyps(2)[symmetric]] by metis moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_wf_trm[OF Equality.hyps(2)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m t\ \wf\<^sub>t\<^sub>r\<^sub>m t'\] by (metis wf_trm_subst_range_iff) moreover have "fv\<^sub>s\<^sub>t (S@Equality a t t'#S') \ bvars\<^sub>s\<^sub>t (S@Equality a t t'#S') = {}" using Equality.prems(1) by (auto simp add: wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) hence "bvars\<^sub>s\<^sub>t (S@S') \ fv t = {}" "bvars\<^sub>s\<^sub>t (S@S') \ fv t' = {}" by auto hence "bvars\<^sub>s\<^sub>t (S@S') \ range_vars \ = {}" using mgu_vars_bounded[OF Equality.hyps(2)[symmetric]] by fast ultimately show ?case using tfr_stp_all_wt_subst_apply[OF \list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S')\] by metis qed private lemma LI_in_SMP_subset_single: assumes "(S,\) \ (S',\')" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" and "trms\<^sub>s\<^sub>t S \ SMP M" shows "trms\<^sub>s\<^sub>t S' \ SMP M" using assms proof (induction rule: LI_rel.induct) case (Compose S X f S' \) hence "SMP (trms\<^sub>s\<^sub>t [Send1 (Fun f X)]) \ SMP M" proof - have "SMP (trms\<^sub>s\<^sub>t [Send1 (Fun f X)]) \ SMP (trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S'))" using trms\<^sub>s\<^sub>t_append SMP_mono by auto thus ?thesis using SMP_union[of "trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S')" M] SMP_subset_union_eq[OF Compose.prems(6)] by auto qed thus ?case using Compose.prems(6) by auto next case (Unify S f Y \ X S' \) have "Fun f X \ SMP (trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S'))" by auto moreover have "MGU \ (Fun f X) (Fun f Y)" by (metis mgu_gives_MGU[OF Unify.hyps(3)[symmetric]]) moreover have "\x. x \ set S \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p x)" "wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)" using Unify.prems(4) by force+ moreover have "Fun f Y \ SMP (trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S'))" by (meson SMP_ikI Unify.hyps(2) contra_subsetD ik_append_subset(1)) ultimately have "wf\<^sub>t\<^sub>r\<^sub>m (Fun f Y)" "\ (Fun f X) = \ (Fun f Y)" using ik\<^sub>s\<^sub>t_subterm_exD[OF \Fun f Y \ ik\<^sub>s\<^sub>t S\] \tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S'))\ unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by (metis (full_types) SMP_wf_trm Unify.prems(4), blast) hence "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" by (metis mgu_wt_if_same_type[OF Unify.hyps(3)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)\]) moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_wf_trm[OF Unify.hyps(3)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)\ \wf\<^sub>t\<^sub>r\<^sub>m (Fun f Y)\] by simp ultimately have "trms\<^sub>s\<^sub>t ((S@Send1 (Fun f X)#S') \\<^sub>s\<^sub>t \) \ SMP M" using SMP.Substitution Unify.prems(6) wt_subst_SMP_subset by metis thus ?case by auto next case (Equality S \ t t' a S' \) hence "\ t = \ t'" using tfr_stp_all_same_type MGU_is_Unifier[OF mgu_gives_MGU[OF Equality.hyps(2)[symmetric]]] by metis moreover have "t \ SMP (trms\<^sub>s\<^sub>t (S@Equality a t t'#S'))" "t' \ SMP (trms\<^sub>s\<^sub>t (S@Equality a t t'#S'))" using Equality.prems(1) by auto moreover have "MGU \ t t'" using mgu_gives_MGU[OF Equality.hyps(2)[symmetric]] by metis moreover have "\x. x \ set S \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p x)" "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m t'" using Equality.prems(4) by force+ ultimately have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" by (metis mgu_wt_if_same_type[OF Equality.hyps(2)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m t\]) moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_wf_trm[OF Equality.hyps(2)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m t\ \wf\<^sub>t\<^sub>r\<^sub>m t'\] by simp ultimately have "trms\<^sub>s\<^sub>t ((S@Equality a t t'#S') \\<^sub>s\<^sub>t \) \ SMP M" using SMP.Substitution Equality.prems wt_subst_SMP_subset by metis thus ?case by auto qed private lemma LI_preserves_tfr_single: assumes "(S,\) \ (S',\')" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" shows "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S') \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S')" using assms proof (induction rule: LI_rel.induct) case (Compose S X f S' \) let ?SMPmap = "SMP (trms\<^sub>s\<^sub>t (S@map Send1 X@S')) - (Var`\)" have "?SMPmap \ SMP (trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S')) - (Var`\)" using SMP_fun_map_snd_subset[of X f] SMP_append[of "map Send1 X" S'] SMP_Cons[of "Send1 (Fun f X)" S'] SMP_append[of S "Send1 (Fun f X)#S'"] SMP_append[of S "map Send1 X@S'"] by auto hence "\s \ ?SMPmap. \t \ ?SMPmap. (\\. Unifier \ s t) \ \ s = \ t" using Compose unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by (meson subsetCE) thus ?case using LI_preserves_trm_wf[OF r_into_rtrancl[OF LI_rel.Compose[OF Compose.hyps]], of S'] Compose.prems(5) unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast next case (Unify S f Y \ X S' \) let ?SMP\ = "SMP (trms\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \)) - (Var`\)" have "SMP (trms\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \)) \ SMP (trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S'))" proof fix s assume "s \ SMP (trms\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \))" thus "s \ SMP (trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S'))" using LI_in_SMP_subset_single[ OF LI_rel.Unify[OF Unify.hyps] Unify.prems(1,2,4,5,6) MP_subset_SMP(2)[of "S@Send1 (Fun f X)#S'"]] by (metis SMP_union SMP_subset_union_eq Un_iff) qed hence "\s \ ?SMP\. \t \ ?SMP\. (\\. Unifier \ s t) \ \ s = \ t" using Unify.prems(4) unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by (meson Diff_iff subsetCE) thus ?case using LI_preserves_trm_wf[OF r_into_rtrancl[OF LI_rel.Unify[OF Unify.hyps]], of S'] Unify.prems(5) unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast next case (Equality S \ t t' a S' \) let ?SMP\ = "SMP (trms\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \)) - (Var`\)" have "SMP (trms\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \)) \ SMP (trms\<^sub>s\<^sub>t (S@Equality a t t'#S'))" proof fix s assume "s \ SMP (trms\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \))" thus "s \ SMP (trms\<^sub>s\<^sub>t (S@Equality a t t'#S'))" using LI_in_SMP_subset_single[ OF LI_rel.Equality[OF Equality.hyps] Equality.prems(1,2,4,5,6) MP_subset_SMP(2)[of "S@Equality a t t'#S'"]] by (metis SMP_union SMP_subset_union_eq Un_iff) qed hence "\s \ ?SMP\. \t \ ?SMP\. (\\. Unifier \ s t) \ \ s = \ t" using Equality.prems unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by (meson Diff_iff subsetCE) thus ?case using LI_preserves_trm_wf[OF r_into_rtrancl[OF LI_rel.Equality[OF Equality.hyps]], of _ S'] Equality.prems unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast qed private lemma LI_preserves_welltypedness_single: assumes "(S,\) \ (S',\')" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \' \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \')" using assms proof (induction rule: LI_rel.induct) case (Unify S f Y \ X S' \) have "wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)" using Unify.prems(5) unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by simp moreover have "wf\<^sub>t\<^sub>r\<^sub>m (Fun f Y)" proof - obtain x where "x \ set S" "Fun f Y \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t\<^sub>p x)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p x)" using Unify.hyps(2) Unify.prems(5) unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by force thus ?thesis using wf_trm_subterm by auto qed moreover have "Fun f X \ SMP (trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S'))" "Fun f Y \ SMP (trms\<^sub>s\<^sub>t (S@Send1 (Fun f X)#S'))" using SMP_append[of S "Send1 (Fun f X)#S'"] SMP_Cons[of "Send1 (Fun f X)" S'] SMP_ikI[OF Unify.hyps(2)] by auto hence "\ (Fun f X) = \ (Fun f Y)" using Unify.prems(4) mgu_gives_MGU[OF Unify.hyps(3)[symmetric]] unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast ultimately have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using mgu_wt_if_same_type[OF Unify.hyps(3)[symmetric]] by metis have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (meson mgu_wf_trm[OF Unify.hyps(3)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)\ \wf\<^sub>t\<^sub>r\<^sub>m (Fun f Y)\] wf_trm_subst_range_iff) hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using wf_trm_subst_range_iff wf_trm_subst \wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)\ unfolding subst_compose_def by (metis (no_types, lifting)) thus ?case by (metis wt_subst_compose[OF \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\]) next case (Equality S \ t t' a S' \) have "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m t'" using Equality.prems(5) by simp_all moreover have "\ t = \ t'" using \list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@Equality a t t'#S')\ MGU_is_Unifier[OF mgu_gives_MGU[OF Equality.hyps(2)[symmetric]]] by auto ultimately have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using mgu_wt_if_same_type[OF Equality.hyps(2)[symmetric]] by metis have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (meson mgu_wf_trm[OF Equality.hyps(2)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m t\ \wf\<^sub>t\<^sub>r\<^sub>m t'\] wf_trm_subst_range_iff) hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using wf_trm_subst_range_iff wf_trm_subst \wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)\ unfolding subst_compose_def by (metis (no_types, lifting)) thus ?case by (metis wt_subst_compose[OF \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\]) qed metis lemma LI_preserves_welltypedness: assumes "(S,\) \\<^sup>* (S',\')" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \'" (is "?A \'") and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \')" (is "?B \'") proof - have "?A \' \ ?B \'" using assms proof (induction S \ rule: converse_rtrancl_induct2) case (step S1 \1 S2 \2) hence "?A \2 \ ?B \2" using LI_preserves_welltypedness_single by presburger moreover have "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S2 \2" by (fact LI_preserves_wellformedness[OF r_into_rtrancl[OF step.hyps(1)] step.prems(1)]) moreover have "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S2)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S2)" using LI_preserves_tfr_single[OF step.hyps(1)] step.prems by presburger+ moreover have "list_all tfr\<^sub>s\<^sub>t\<^sub>p S2" using LI_preserves_tfr_stp_all_single[OF step.hyps(1)] step.prems by fastforce ultimately show ?case using step.IH by presburger qed simp thus "?A \'" "?B \'" by simp_all qed lemma LI_preserves_tfr: assumes "(S,\) \\<^sup>* (S',\')" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" shows "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S')" (is "?A S'") and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S')" (is "?B S'") and "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" (is "?C S'") proof - have "?A S' \ ?B S' \ ?C S'" using assms proof (induction S \ rule: converse_rtrancl_induct2) case (step S1 \1 S2 \2) have "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S2 \2" "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S2)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S2)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S2" using LI_preserves_wellformedness[OF r_into_rtrancl[OF step.hyps(1)] step.prems(1)] LI_preserves_tfr_single[OF step.hyps(1) step.prems(1,2)] LI_preserves_tfr_stp_all_single[OF step.hyps(1) step.prems(1,2)] step.prems(3,4,5,6) by metis+ moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \2" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \2)" using LI_preserves_welltypedness[OF r_into_rtrancl[OF step.hyps(1)] step.prems] by simp_all ultimately show ?case using step.IH by presburger qed blast thus "?A S'" "?B S'" "?C S'" by simp_all qed lemma LI_preproc_preserves_tfr: assumes "tfr\<^sub>s\<^sub>t S" shows "tfr\<^sub>s\<^sub>t (LI_preproc S)" unfolding tfr\<^sub>s\<^sub>t_def proof have S: "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using assms unfolding tfr\<^sub>s\<^sub>t_def by metis+ show "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t (LI_preproc S))" by (metis S(1) LI_preproc_trms_eq) show "list_all tfr\<^sub>s\<^sub>t\<^sub>p (LI_preproc S)" using S(2) proof (induction S) case (Cons x S) have IH: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (LI_preproc S)" using Cons by simp have x: "tfr\<^sub>s\<^sub>t\<^sub>p x" using Cons.prems by simp show ?case using x IH unfolding list_all_iff by (cases x) auto qed simp qed end subsubsection \Simple Constraints are Well-typed Satisfiable\ text \Proving the existence of a well-typed interpretation\ context begin lemma wt_interpretation_exists: obtains \::"('fun,'var) subst" where "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "subst_range \ \ public_ground_wf_terms" proof define \ where "\ = (\x. (SOME t. \ (Var x) = \ t \ public_ground_wf_term t))" { fix x t assume "\ x = t" hence "\ (Var x) = \ t \ public_ground_wf_term t" using someI_ex[of "\t. \ (Var x) = \ t \ public_ground_wf_term t", OF type_pgwt_inhabited[of "Var x"]] unfolding \_def wf\<^sub>t\<^sub>r\<^sub>m_def by simp } hence props: "\ v = t \ \ (Var v) = \ t \ public_ground_wf_term t" for v t by metis have "\ v \ Var v" for v using props pgwt_ground by fastforce hence "subst_domain \ = UNIV" by auto moreover have "ground (subst_range \)" by (simp add: props pgwt_ground) ultimately show "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" by metis show "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def using props by simp show "subst_range \ \ public_ground_wf_terms" by (auto simp add: props) qed lemma wt_grounding_subst_exists: "\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}" proof - obtain \ where \: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "subst_range \ \ public_ground_wf_terms" using wt_interpretation_exists by blast show ?thesis using pgwt_wellformed interpretation_grounds[OF \(1)] \(2,3) by blast qed private fun fresh_pgwt::"'fun set \ ('fun,'atom) term_type \ ('fun,'var) term" where "fresh_pgwt S (TAtom a) = Fun (SOME c. c \ S \ \ (Fun c []) = TAtom a \ public c) []" | "fresh_pgwt S (TComp f T) = Fun f (map (fresh_pgwt S) T)" private lemma fresh_pgwt_same_type: assumes "finite S" "wf\<^sub>t\<^sub>r\<^sub>m t" shows "\ (fresh_pgwt S (\ t)) = \ t" proof - let ?P = "\\::('fun,'atom) term_type. wf\<^sub>t\<^sub>r\<^sub>m \ \ (\f T. TComp f T \ \ \ 0 < arity f)" { fix \ assume "?P \" hence "\ (fresh_pgwt S \) = \" proof (induction \) case (Var a) let ?P = "\c. c \ S \ \ (Fun c []) = Var a \ public c" let ?Q = "\c. \ (Fun c []) = Var a \ public c" have " {c. ?Q c} - S = {c. ?P c}" by auto hence "infinite {c. ?P c}" using Diff_infinite_finite[OF assms(1) infinite_typed_consts[of a]] by metis hence "\c. ?P c" using not_finite_existsD by blast thus ?case using someI_ex[of ?P] by auto next case (Fun f T) have f: "0 < arity f" using Fun.prems fun_type_inv by auto have "\t. t \ set T \ ?P t" using Fun.prems wf_trm_subtermeq term.le_less_trans Fun_param_is_subterm by metis hence "\t. t \ set T \ \ (fresh_pgwt S t) = t" using Fun.prems Fun.IH by auto hence "map \ (map (fresh_pgwt S) T) = T" by (induct T) auto thus ?case using fun_type[OF f] by simp qed } thus ?thesis using assms(1) \_wf'[OF assms(2)] \_wf'' by auto qed private lemma fresh_pgwt_empty_synth: assumes "finite S" "wf\<^sub>t\<^sub>r\<^sub>m t" shows "{} \\<^sub>c fresh_pgwt S (\ t)" proof - let ?P = "\\::('fun,'atom) term_type. wf\<^sub>t\<^sub>r\<^sub>m \ \ (\f T. TComp f T \ \ \ 0 < arity f)" { fix \ assume "?P \" hence "{} \\<^sub>c fresh_pgwt S \" proof (induction \) case (Var a) let ?P = "\c. c \ S \ \ (Fun c []) = Var a \ public c" let ?Q = "\c. \ (Fun c []) = Var a \ public c" have " {c. ?Q c} - S = {c. ?P c}" by auto hence "infinite {c. ?P c}" using Diff_infinite_finite[OF assms(1) infinite_typed_consts[of a]] by metis hence "\c. ?P c" using not_finite_existsD by blast thus ?case using someI_ex[of ?P] intruder_synth.ComposeC[of "[]" _ "{}"] const_type_inv by auto next case (Fun f T) have f: "0 < arity f" "length T = arity f" "public f" using Fun.prems fun_type_inv unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto have "\t. t \ set T \ ?P t" using Fun.prems wf_trm_subtermeq term.le_less_trans Fun_param_is_subterm by metis hence "\t. t \ set T \ {} \\<^sub>c fresh_pgwt S t" using Fun.prems Fun.IH by auto moreover have "length (map (fresh_pgwt S) T) = arity f" using f(2) by auto ultimately show ?case using intruder_synth.ComposeC[of "map (fresh_pgwt S) T" f] f by auto qed } thus ?thesis using assms(1) \_wf'[OF assms(2)] \_wf'' by auto qed private lemma fresh_pgwt_has_fresh_const: assumes "finite S" "wf\<^sub>t\<^sub>r\<^sub>m t" obtains c where "Fun c [] \ fresh_pgwt S (\ t)" "c \ S" proof - let ?P = "\\::('fun,'atom) term_type. wf\<^sub>t\<^sub>r\<^sub>m \ \ (\f T. TComp f T \ \ \ 0 < arity f)" { fix \ assume "?P \" hence "\c. Fun c [] \ fresh_pgwt S \ \ c \ S" proof (induction \) case (Var a) let ?P = "\c. c \ S \ \ (Fun c []) = Var a \ public c" let ?Q = "\c. \ (Fun c []) = Var a \ public c" have " {c. ?Q c} - S = {c. ?P c}" by auto hence "infinite {c. ?P c}" using Diff_infinite_finite[OF assms(1) infinite_typed_consts[of a]] by metis hence "\c. ?P c" using not_finite_existsD by blast thus ?case using someI_ex[of ?P] by auto next case (Fun f T) have f: "0 < arity f" "length T = arity f" "public f" "T \ []" using Fun.prems fun_type_inv unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto obtain t' where t': "t' \ set T" by (meson all_not_in_conv f(4) set_empty) have "\t. t \ set T \ ?P t" using Fun.prems wf_trm_subtermeq term.le_less_trans Fun_param_is_subterm by metis hence "\t. t \ set T \ \c. Fun c [] \ fresh_pgwt S t \ c \ S" using Fun.prems Fun.IH by auto then obtain c where c: "Fun c [] \ fresh_pgwt S t'" "c \ S" using t' by metis thus ?case using t' by auto qed } thus ?thesis using that assms \_wf'[OF assms(2)] \_wf'' by blast qed private lemma fresh_pgwt_subterm_fresh: assumes "finite S" "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m s" "funs_term s \ S" shows "s \ subterms (fresh_pgwt S (\ t))" proof - let ?P = "\\::('fun,'atom) term_type. wf\<^sub>t\<^sub>r\<^sub>m \ \ (\f T. TComp f T \ \ \ 0 < arity f)" { fix \ assume "?P \" hence "s \ subterms (fresh_pgwt S \)" proof (induction \) case (Var a) let ?P = "\c. c \ S \ \ (Fun c []) = Var a \ public c" let ?Q = "\c. \ (Fun c []) = Var a \ public c" have " {c. ?Q c} - S = {c. ?P c}" by auto hence "infinite {c. ?P c}" using Diff_infinite_finite[OF assms(1) infinite_typed_consts[of a]] by metis hence "\c. ?P c" using not_finite_existsD by blast thus ?case using someI_ex[of ?P] assms(4) by auto next case (Fun f T) have f: "0 < arity f" "length T = arity f" "public f" using Fun.prems fun_type_inv unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto have "\t. t \ set T \ ?P t" using Fun.prems wf_trm_subtermeq term.le_less_trans Fun_param_is_subterm by metis hence "\t. t \ set T \ s \ subterms (fresh_pgwt S t)" using Fun.prems Fun.IH by auto moreover have "s \ fresh_pgwt S (Fun f T)" proof - obtain c where c: "Fun c [] \ fresh_pgwt S (Fun f T)" "c \ S" using fresh_pgwt_has_fresh_const[OF assms(1)] type_wfttype_inhabited Fun.prems by metis hence "\Fun c [] \ s" using assms(4) subtermeq_imp_funs_term_subset by force thus ?thesis using c(1) by auto qed ultimately show ?case by auto qed } thus ?thesis using assms(1) \_wf'[OF assms(2)] \_wf'' by auto qed private lemma wt_fresh_pgwt_term_exists: assumes "finite T" "wf\<^sub>t\<^sub>r\<^sub>m s" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s T" obtains t where "\ t = \ s" "{} \\<^sub>c t" "\s \ T. \u \ subterms s. u \ subterms t" proof - have finite_S: "finite (\(funs_term ` T))" using assms(1) by auto have 1: "\ (fresh_pgwt (\(funs_term ` T)) (\ s)) = \ s" using fresh_pgwt_same_type[OF finite_S assms(2)] by auto have 2: "{} \\<^sub>c fresh_pgwt (\(funs_term ` T)) (\ s)" using fresh_pgwt_empty_synth[OF finite_S assms(2)] by auto have 3: "\v \ T. \u \ subterms v. u \ subterms (fresh_pgwt (\(funs_term ` T)) (\ s))" using fresh_pgwt_subterm_fresh[OF finite_S assms(2)] assms(3) wf_trm_subtermeq subtermeq_imp_funs_term_subset by force show ?thesis by (rule that[OF 1 2 3]) qed lemma wt_bij_finite_subst_exists: assumes "finite (S::'var set)" "finite (T::('fun,'var) terms)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s T" shows "\\::('fun,'var) subst. subst_domain \ = S \ bij_betw \ (subst_domain \) (subst_range \) \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ {t. {} \\<^sub>c t} - T \ (\s \ subst_range \. \u \ subst_range \. (\v. v \ s \ v \ u) \ s = u) \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using assms proof (induction rule: finite_induct) case empty have "subst_domain Var = {}" "bij_betw Var (subst_domain Var) (subst_range Var)" "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range Var) \ {t. {} \\<^sub>c t} - T" "\s \ subst_range Var. \u \ subst_range Var. (\v. v \ s \ v \ u) \ s = u" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range Var)" unfolding bij_betw_def by auto thus ?case by (force simp add: subst_domain_def) next case (insert x S) then obtain \ where \: "subst_domain \ = S" "bij_betw \ (subst_domain \) (subst_range \)" "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ {t. {} \\<^sub>c t} - T" "\s \ subst_range \. \u \ subst_range \. (\v. v \ s \ v \ u) \ s = u" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (auto simp del: subst_range.simps) have *: "finite (T \ subst_range \)" using insert.prems(1) insert.hyps(1) \(1) by simp have **: "wf\<^sub>t\<^sub>r\<^sub>m (Var x)" by simp have ***: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (T \ subst_range \)" using assms(3) \(6) by blast obtain t where t: "\ t = \ (Var x)" "{} \\<^sub>c t" "\s \ T \ subst_range \. \u \ subterms s. u \ subterms t" using wt_fresh_pgwt_term_exists[OF * ** ***] by auto obtain \ where \: "\ \ \y. if x = y then t else \ y" by simp have t_ground: "fv t = {}" using t(2) pgwt_ground[of t] pgwt_is_empty_synth[of t] by auto hence x_dom: "x \ subst_domain \" "x \ subst_domain \" using insert.hyps(2) \(1) \ by auto moreover have "subst_range \ \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \)" by auto hence ground_imgs: "ground (subst_range \)" using \(3) pgwt_ground pgwt_is_empty_synth by force ultimately have x_img: "\ x \ subst_range \" using ground_subst_dom_iff_img by (auto simp add: subst_domain_def) have "ground (insert t (subst_range \))" using ground_imgs x_dom t_ground by auto have \_dom: "subst_domain \ = insert x (subst_domain \)" using \ t_ground by (auto simp add: subst_domain_def) have \_img: "subst_range \ = insert t (subst_range \)" proof show "subst_range \ \ insert t (subst_range \)" proof fix t' assume "t' \ subst_range \" then obtain y where "y \ subst_domain \" "t' = \ y" by auto thus "t' \ insert t (subst_range \)" using \ by (auto simp add: subst_domain_def) qed show "insert t (subst_range \) \ subst_range \" proof fix t' assume t': "t' \ insert t (subst_range \)" hence "fv t' = {}" using ground_imgs x_img t_ground by auto hence "t' \ Var x" by auto show "t' \ subst_range \" proof (cases "t' = t") case False hence "t' \ subst_range \" using t' by auto then obtain y where "\ y \ subst_range \" "t' = \ y" by auto hence "y \ subst_domain \" "t' \ Var y" using ground_subst_dom_iff_img[OF ground_imgs(1)] by (auto simp add: subst_domain_def simp del: subst_range.simps) hence "x \ y" using x_dom by auto hence "\ y = \ y" unfolding \ by auto thus ?thesis using \t' \ Var y\ \t' = \ y\ subst_imgI[of \ y] by auto qed (metis subst_imgI \ \t' \ Var x\) qed qed hence \_ground_img: "ground (subst_range \)" using ground_imgs t_ground by auto have "subst_domain \ = insert x S" using \_dom \(1) by auto moreover have "bij_betw \ (subst_domain \) (subst_range \)" proof (intro bij_betwI') fix y z assume *: "y \ subst_domain \" "z \ subst_domain \" hence "fv (\ y) = {}" "fv (\ z) = {}" using \_ground_img by auto { assume "\ y = \ z" hence "y = z" proof (cases "\ y \ subst_range \ \ \ z \ subst_range \") case True hence **: "y \ subst_domain \" "z \ subst_domain \" using \ \_dom True * t(3) by (metis Un_iff term.order_refl insertE)+ hence "y \ x" "z \ x" using x_dom by auto hence "\ y = \ y" "\ z = \ z" using \ by auto thus ?thesis using \\ y = \ z\ \(2) ** unfolding bij_betw_def inj_on_def by auto qed (metis \ * \\ y = \ z\ \_dom ground_imgs(1) ground_subst_dom_iff_img insertE) } thus "(\ y = \ z) = (y = z)" by auto next fix y assume "y \ subst_domain \" thus "\ y \ subst_range \" by auto next fix t assume "t \ subst_range \" thus "\z \ subst_domain \. t = \ z" by auto qed moreover have "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ {t. {} \\<^sub>c t} - T" proof - { fix s assume "s \ t" hence "s \ {t. {} \\<^sub>c t} - T" using t(2,3) by (metis Diff_eq_empty_iff Diff_iff Un_upper1 term.order_refl deduct_synth_subterm mem_Collect_eq) } thus ?thesis using \(3) \ \_img by auto qed moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using \ t(1) \(5) unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by auto moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using \ \(6) t(2) pgwt_is_empty_synth pgwt_wellformed wf_trm_subst_range_iff[of \] wf_trm_subst_range_iff[of \] by metis moreover have "\s\subst_range \. \u\subst_range \. (\v. v \ s \ v \ u) \ s = u" using \(4) \_img t(3) by (auto simp del: subst_range.simps) ultimately show ?case by blast qed private lemma wt_bij_finite_tatom_subst_exists_single: assumes "finite (S::'var set)" "finite (T::('fun,'var) terms)" and "\x. x \ S \ \ (Var x) = TAtom a" shows "\\::('fun,'var) subst. subst_domain \ = S \ bij_betw \ (subst_domain \) (subst_range \) \ subst_range \ \ ((\c. Fun c []) ` {c. \ (Fun c []) = TAtom a \ public c \ arity c = 0}) - T \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" proof - let ?U = "{c. \ (Fun c []) = TAtom a \ public c \ arity c = 0}" obtain \ where \: "subst_domain \ = S" "bij_betw \ (subst_domain \) (subst_range \)" "subst_range \ \ ((\c. Fun c []) ` ?U) - T" using bij_finite_const_subst_exists'[OF assms(1,2) infinite_typed_consts'[of a]] by auto { fix x assume "x \ subst_domain \" hence "\ (Var x) = \ (\ x)" by auto } moreover { fix x assume "x \ subst_domain \" hence "\c \ ?U. \ x = Fun c [] \ arity c = 0" using \ by auto hence "\ (\ x) = TAtom a" "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" using assms(3) const_type wf_trmI[of "[]"] by auto hence "\ (Var x) = \ (\ x)" "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" using assms(3) \(1) by force+ } ultimately have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using wf_trm_subst_range_iff[of \] unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by force+ thus ?thesis using \ by auto qed lemma wt_bij_finite_tatom_subst_exists: assumes "finite (S::'var set)" "finite (T::('fun,'var) terms)" and "\x. x \ S \ \a. \ (Var x) = TAtom a" shows "\\::('fun,'var) subst. subst_domain \ = S \ bij_betw \ (subst_domain \) (subst_range \) \ subst_range \ \ ((\c. Fun c []) ` \\<^sub>p\<^sub>u\<^sub>b) - T \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using assms proof (induction rule: finite_induct) case empty have "subst_domain Var = {}" "bij_betw Var (subst_domain Var) (subst_range Var)" "subst_range Var \ ((\c. Fun c []) ` \\<^sub>p\<^sub>u\<^sub>b) - T" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range Var)" unfolding bij_betw_def by auto thus ?case by (auto simp add: subst_domain_def) next case (insert x S) then obtain a where a: "\ (Var x) = TAtom a" by fastforce from insert obtain \ where \: "subst_domain \ = S" "bij_betw \ (subst_domain \) (subst_range \)" "subst_range \ \ ((\c. Fun c []) ` \\<^sub>p\<^sub>u\<^sub>b) - T" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by auto let ?S' = "{y \ S. \ (Var y) = TAtom a}" let ?T' = "T \ subst_range \" have *: "finite (insert x ?S')" using insert by simp have **: "finite ?T'" using insert.prems(1) insert.hyps(1) \(1) by simp have ***: "\y. y \ insert x ?S' \ \ (Var y) = TAtom a" using a by auto obtain \ where \: "subst_domain \ = insert x ?S'" "bij_betw \ (subst_domain \) (subst_range \)" "subst_range \ \ ((\c. Fun c []) ` \\<^sub>p\<^sub>u\<^sub>b) - ?T'" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using wt_bij_finite_tatom_subst_exists_single[OF * ** ***] const_type_inv[of _ "[]" a] by blast obtain \ where \: "\ \ \y. if x = y then \ y else \ y" by simp have x_dom: "x \ subst_domain \" "x \ subst_domain \" "x \ subst_domain \" using insert.hyps(2) \(1) \(1) \ by (auto simp add: subst_domain_def) moreover have ground_imgs: "ground (subst_range \)" "ground (subst_range \)" using pgwt_ground \(3) \(3) by auto ultimately have x_img: "\ x \ subst_range \" "\ x \ subst_range \" using ground_subst_dom_iff_img by (auto simp add: subst_domain_def) have "ground (insert (\ x) (subst_range \))" using ground_imgs x_dom by auto have \_dom: "subst_domain \ = insert x (subst_domain \)" using \(1) \ by (auto simp add: subst_domain_def) have \_img: "subst_range \ = insert (\ x) (subst_range \)" proof show "subst_range \ \ insert (\ x) (subst_range \)" proof fix t assume "t \ subst_range \" then obtain y where "y \ subst_domain \" "t = \ y" by auto thus "t \ insert (\ x) (subst_range \)" using \ by (auto simp add: subst_domain_def) qed show "insert (\ x) (subst_range \) \ subst_range \" proof fix t assume t: "t \ insert (\ x) (subst_range \)" hence "fv t = {}" using ground_imgs x_img(2) by auto hence "t \ Var x" by auto show "t \ subst_range \" proof (cases "t = \ x") case True thus ?thesis using subst_imgI \ \t \ Var x\ by metis next case False hence "t \ subst_range \" using t by auto then obtain y where "\ y \ subst_range \" "t = \ y" by auto hence "y \ subst_domain \" "t \ Var y" using ground_subst_dom_iff_img[OF ground_imgs(1)] by (auto simp add: subst_domain_def simp del: subst_range.simps) hence "x \ y" using x_dom by auto hence "\ y = \ y" unfolding \ by auto thus ?thesis using \t \ Var y\ \t = \ y\ subst_imgI[of \ y] by auto qed qed qed hence \_ground_img: "ground (subst_range \)" using ground_imgs x_img by auto have "subst_domain \ = insert x S" using \_dom \(1) by auto moreover have "bij_betw \ (subst_domain \) (subst_range \)" proof (intro bij_betwI') fix y z assume *: "y \ subst_domain \" "z \ subst_domain \" hence "fv (\ y) = {}" "fv (\ z) = {}" using \_ground_img by auto { assume "\ y = \ z" hence "y = z" proof (cases "\ y \ subst_range \ \ \ z \ subst_range \") case True hence **: "y \ subst_domain \" "z \ subst_domain \" using \ \_dom x_img(2) \(3) True by (metis (no_types) *(1) DiffE Un_upper2 insertE subsetCE, metis (no_types) *(2) DiffE Un_upper2 insertE subsetCE) hence "y \ x" "z \ x" using x_dom by auto hence "\ y = \ y" "\ z = \ z" using \ by auto thus ?thesis using \\ y = \ z\ \(2) ** unfolding bij_betw_def inj_on_def by auto qed (metis \ * \\ y = \ z\ \_dom ground_imgs(1) ground_subst_dom_iff_img insertE) } thus "(\ y = \ z) = (y = z)" by auto next fix y assume "y \ subst_domain \" thus "\ y \ subst_range \" by auto next fix t assume "t \ subst_range \" thus "\z \ subst_domain \. t = \ z" by auto qed moreover have "subst_range \ \ (\c. Fun c []) ` \\<^sub>p\<^sub>u\<^sub>b - T" using \(3) \(3) \ by (auto simp add: subst_domain_def) moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using \(4) \(4) \ unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by auto moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using \ \(5) \(5) wf_trm_subst_range_iff[of \] wf_trm_subst_range_iff[of \] wf_trm_subst_range_iff[of \] by presburger ultimately show ?case by blast qed theorem wt_sat_if_simple: assumes "simple S" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" and \': "\X F. Inequality X F \ set S \ ineq_model \' X F" "ground (subst_range \')" "subst_domain \' = {x \ vars\<^sub>s\<^sub>t S. \X F. Inequality X F \ set S \ x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X}" and tfr_stp_all: "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" shows "\\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ (\ \\<^sub>c \S, \\) \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" proof - from \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \\ have "wf\<^sub>s\<^sub>t {} S" "subst_idem \" and S_\_disj: "\v \ vars\<^sub>s\<^sub>t S. \ v = Var v" using subst_idemI[of \] unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by force+ obtain \::"('fun,'var) subst" where \: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "subst_range \ \ public_ground_wf_terms" using wt_interpretation_exists by blast hence \_deduct: "\x M. M \\<^sub>c \ x" and \_wf_trm: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using pgwt_deducible pgwt_wellformed by fastforce+ let ?P = "\\ X. subst_domain \ = set X \ ground (subst_range \)" let ?Sineqsvars = "{x \ vars\<^sub>s\<^sub>t S. \X F. Inequality X F \ set S \ x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ x \ set X}" let ?Strms = "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" have finite_vars: "finite ?Sineqsvars" "finite ?Strms" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s ?Strms" using wf_trm_subtermeq assms(5) by fastforce+ define Q1 where "Q1 = (\(F::(('fun,'var) term \ ('fun,'var) term) list) X. \x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X. \a. \ (Var x) = TAtom a)" define Q2 where "Q2 = (\(F::(('fun,'var) term \ ('fun,'var) term) list) X. \f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ T = [] \ (\s \ set T. s \ Var ` set X))" define Q1' where "Q1' = (\(t::('fun,'var) term) (t'::('fun,'var) term) X. \x \ (fv t \ fv t') - set X. \a. \ (Var x) = TAtom a)" define Q2' where "Q2' = (\(t::('fun,'var) term) (t'::('fun,'var) term) X. \f T. Fun f T \ subterms t \ subterms t' \ T = [] \ (\s \ set T. s \ Var ` set X))" have ex_P: "\X. \\. ?P \ X" using interpretation_subst_exists' by blast have tfr_ineq: "\X F. Inequality X F \ set S \ Q1 F X \ Q2 F X" using tfr_stp_all Q1_def Q2_def tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def[of S] by blast have S_fv_bvars_disj: "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" using \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \\ unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by metis hence ineqs_vars_not_bound: "\X F x. Inequality X F \ set S \ x \ ?Sineqsvars \ x \ set X" using strand_fv_bvars_disjoint_unfold by blast have \_vars_S_bvars_disj: "(subst_domain \ \ range_vars \) \ set X = {}" when "Inequality X F \ set S" for F X using wf_constr_bvars_disj[OF \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \\] strand_fv_bvars_disjointD(1)[OF S_fv_bvars_disj that] by blast obtain \::"('fun,'var) subst" where \_fv_dom: "subst_domain \ = ?Sineqsvars" and \_subterm_inj: "subterm_inj_on \ (subst_domain \)" and \_fresh_pub_img: "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ {t. {} \\<^sub>c t} - ?Strms" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wf_trm: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using wt_bij_finite_subst_exists[OF finite_vars] subst_inj_on_is_bij_betw subterm_inj_on_alt_def' by moura have \_bij_dom_img: "bij_betw \ (subst_domain \) (subst_range \)" by (metis \_subterm_inj subst_inj_on_is_bij_betw subterm_inj_on_alt_def) have "finite (subst_domain \)" by(metis \_fv_dom finite_vars(1)) hence \_finite_img: "finite (subst_range \)" using \_bij_dom_img bij_betw_finite by blast have \_img_subterms: "\s \ subst_range \. \u \ subst_range \. (\v. v \ s \ v \ u) \ s = u" by (metis \_subterm_inj subterm_inj_on_alt_def') have "subst_range \ \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \)" by auto hence "subst_range \ \ public_ground_wf_terms - ?Strms" and \_pgwt_img: "subst_range \ \ public_ground_wf_terms" "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ public_ground_wf_terms" using \_fresh_pub_img pgwt_is_empty_synth by blast+ have \_img_ground: "ground (subst_range \)" using \_pgwt_img pgwt_ground by auto hence \_inj: "inj \" using \_bij_dom_img subst_inj_is_bij_betw_dom_img_if_ground_img by auto have \_ineqs_fv_dom: "\X F. Inequality X F \ set S \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ subst_domain \" using \_fv_dom by fastforce have \_dom_bvars_disj: "\X F. Inequality X F \ set S \ subst_domain \ \ set X = {}" using ineqs_vars_not_bound \_fv_dom by fastforce have \'1: "\X F \. Inequality X F \ set S \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ subst_domain \'" using \'(3) ineqs_vars_not_bound by fastforce have \'2: "\X F. Inequality X F \ set S \ subst_domain \' \ set X = {}" using \'(3) ineqs_vars_not_bound by blast have doms_eq: "subst_domain \' = subst_domain \" using \'(3) \_fv_dom by simp have \_ineqs_neq: "ineq_model \ X F" when "Inequality X F \ set S" for X F proof - obtain a::"'fun" where a: "a \ \(funs_term ` subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \))" using exists_fun_notin_funs_terms[OF subterms_union_finite[OF \_finite_img]] by moura hence a': "\T. Fun a T \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \)" "\S. Fun a [] \ set (Fun a []#S)" "Fun a [] \ Var ` set X" by (meson a UN_I term.set_intros(1), auto) define t where "t \ Fun a (Fun a []#map fst F)" define t' where "t' \ Fun a (Fun a []#map snd F)" note F_in = that have t_fv: "fv t \ fv t' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" unfolding t_def t'_def by force have t_subterms: "subterms t \ subterms t' \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ {t, t', Fun a []}" unfolding t_def t'_def by force have "t \ \ \ \ \ t' \ \ \ \" when "?P \ X" for \ proof - have tfr_assms: "Q1 F X \ Q2 F X" using tfr_ineq F_in by metis have "Q1 F X \ \x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X. \c. \ x = Fun c []" proof fix x assume "Q1 F X" and x: "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" then obtain a where "\ (Var x) = TAtom a" unfolding Q1_def by moura hence a: "\ (\ x) = TAtom a" using \_wt unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by simp have "x \ subst_domain \" using \_ineqs_fv_dom x F_in by auto then obtain f T where fT: "\ x = Fun f T" by (meson \_img_ground ground_img_obtain_fun) hence "T = []" using \_wf_trm a TAtom_term_cases by fastforce thus "\c. \ x = Fun c []" using fT by metis qed hence 1: "Q1 F X \ \x \ (fv t \ fv t') - set X. \c. \ x = Fun c []" using t_fv by auto have 2: "\Q1 F X \ Q2 F X" by (metis tfr_assms) have 3: "subst_domain \ \ set X = {}" using \_dom_bvars_disj F_in by auto have 4: "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ (subterms t \ subterms t') = {}" proof - define M1 where "M1 \ {t, t', Fun a []}" define M2 where "M2 \ ?Strms" have "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ M2" using F_in unfolding M2_def by force moreover have "subterms t \ subterms t' \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ M1" using t_subterms unfolding M1_def by blast ultimately have *: "subterms t \ subterms t' \ M2 \ M1" by auto have "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ M1 = {}" "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ M2 = {}" using a' \_fresh_pub_img unfolding t_def t'_def M1_def M2_def by blast+ thus ?thesis using * by blast qed have 5: "(fv t \ fv t') - subst_domain \ \ set X" using \_ineqs_fv_dom[OF F_in] t_fv by auto have 6: "\\. ?P \ X \ t \ \ \ \' \ t' \ \ \ \'" by (metis t_def t'_def \'(1) F_in ineq_model_singleE ineq_model_single_iff) have 7: "fv t \ fv t' - set X \ subst_domain \'" using \'1 F_in t_fv by force have 8: "subst_domain \' \ set X = {}" using \'2 F_in by auto have 9: "Q1' t t' X" when "Q1 F X" using that t_fv unfolding Q1_def Q1'_def t_def t'_def by blast have 10: "Q2' t t' X" when "Q2 F X" unfolding Q2'_def proof (intro allI impI) fix f T assume "Fun f T \ subterms t \ subterms t'" moreover { assume "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F)" hence "T = [] \ (\s\set T. s \ Var ` set X)" by (metis Q2_def that) } moreover { assume "Fun f T = t" hence "T = [] \ (\s\set T. s \ Var ` set X)" unfolding t_def using a'(2,3) by simp } moreover { assume "Fun f T = t'" hence "T = [] \ (\s\set T. s \ Var ` set X)" unfolding t'_def using a'(2,3) by simp } moreover { assume "Fun f T = Fun a []" hence "T = [] \ (\s\set T. s \ Var ` set X)" by simp } ultimately show "T = [] \ (\s\set T. s \ Var ` set X)" using t_subterms by blast qed note 11 = \_subterm_inj \_img_ground 3 4 5 note 12 = 6 7 8 \'(2) doms_eq show "t \ \ \ \ \ t' \ \ \ \" using 1 2 9 10 that sat_ineq_subterm_inj_subst[OF 11 _ 12] unfolding Q1'_def Q2'_def by metis qed thus ?thesis by (metis t_def t'_def ineq_model_singleI ineq_model_single_iff) qed have \_ineqs_fv_dom': "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \ subst_domain \" when "Inequality X F \ set S" and "?P \ X" for F \ X using \_ineqs_fv_dom[OF that(1)] proof (induction F) case (Cons g G) obtain t t' where g: "g = (t,t')" by (metis surj_pair) hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (g#G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) = fv (t \ \) \ fv (t' \ \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (g#G) = fv t \ fv t' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" by (simp_all add: subst_apply_pairs_def) moreover have "fv (t \ \) = fv t - subst_domain \" "fv (t' \ \) = fv t' - subst_domain \" using g that(2) by (simp_all add: subst_fv_unfold_ground_img range_vars_alt_def) moreover have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \ subst_domain \" using Cons by auto ultimately show ?case using Cons.prems that(2) by auto qed (simp add: subst_apply_pairs_def) have \_ineqs_ground: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ((F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) = {}" when "Inequality X F \ set S" and "?P \ X" for F \ X using \_ineqs_fv_dom'[OF that] proof (induction F) case (Cons g G) obtain t t' where g: "g = (t,t')" by (metis surj_pair) hence "fv (t \ \) \ subst_domain \" "fv (t' \ \) \ subst_domain \" using Cons.prems by (auto simp add: subst_apply_pairs_def) hence "fv (t \ \ \ \) = {}" "fv (t' \ \ \ \) = {}" using subst_fv_dom_ground_if_ground_img[OF _ \_img_ground] by metis+ thus ?case using g Cons by (auto simp add: subst_apply_pairs_def) qed (simp add: subst_apply_pairs_def) from \_pgwt_img \_ineqs_neq have \_deduct: "M \\<^sub>c \ x" when "x \ subst_domain \" for x M using that pgwt_deducible by fastforce { fix M::"('fun,'var) terms" have "\M; S\\<^sub>c (\ \\<^sub>s \ \\<^sub>s \)" using \wf\<^sub>s\<^sub>t {} S\ \simple S\ S_\_disj \_ineqs_neq \_ineqs_fv_dom' \_vars_S_bvars_disj proof (induction S arbitrary: M rule: wf\<^sub>s\<^sub>t_simple_induct) case (ConsSnd v S) hence S_sat: "\M; S\\<^sub>c (\ \\<^sub>s \ \\<^sub>s \)" and "\ v = Var v" by auto hence *: "\M. M \\<^sub>c Var v \ (\ \\<^sub>s \ \\<^sub>s \)" using \_deduct \_deduct - by (metis ideduct_synth_subst_apply subst_apply_term.simps(1) + by (metis ideduct_synth_subst_apply eval_term.simps(1) subst_subst_compose trm_subst_ident') define M' where "M' \ M \ (ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s \)" have "\t \ set [Var v]. M' \\<^sub>c t \ (\ \\<^sub>s \ \\<^sub>s \)" using *[of M'] by simp thus ?case using strand_sem_append(1)[OF S_sat, of "[Send1 (Var v)]", unfolded M'_def[symmetric]] strand_sem_c.simps(1)[of M'] strand_sem_c.simps(2)[of M' "[Var v]" "[]"] by presburger next case (ConsIneq X F S) have dom_disj: "subst_domain \ \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F = {}" using ConsIneq.prems(1) subst_dom_vars_in_subst by force hence *: "F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \ = F" by blast have **: "ineq_model \ X F" by (meson ConsIneq.prems(2) in_set_conv_decomp) have "\x. x \ vars\<^sub>s\<^sub>t S \ x \ vars\<^sub>s\<^sub>t (S@[Inequality X F])" "\x. x \ set S \ x \ set (S@[Inequality X F])" by auto hence IH: "\M; S\\<^sub>c (\ \\<^sub>s \ \\<^sub>s \)" by (metis ConsIneq.IH ConsIneq.prems(1,2,3,4)) have "ineq_model (\ \\<^sub>s \) X F" proof - have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \ subst_domain \" when "?P \ X" for \ using ConsIneq.prems(3)[OF _ that] by simp hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ subst_domain \" using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_subset ex_P by (metis Diff_subset_conv Un_commute) thus ?thesis by (metis ineq_model_ground_subst[OF _ \_img_ground **]) qed hence "ineq_model (\ \\<^sub>s \ \\<^sub>s \) X F" using * ineq_model_subst' subst_compose_assoc ConsIneq.prems(4) by (metis UnCI list.set_intros(1) set_append) thus ?case using IH by (auto simp add: ineq_model_def) qed auto } moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \ \\<^sub>s \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \ \\<^sub>s \))" by (metis wt_subst_compose \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\, metis assms(4) \_wf_trm \_wf_trm wf_trm_subst subst_img_comp_subset') ultimately show ?thesis using interpretation_comp(1)[OF \interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\, of "\ \\<^sub>s \"] subst_idem_support[OF \subst_idem \\, of "\ \\<^sub>s \"] subst_compose_assoc unfolding constr_sem_c_def by metis qed end subsubsection \Theorem: Type-flaw resistant constraints are well-typed satisfiable (composition-only)\ text \ There exists well-typed models of satisfiable type-flaw resistant constraints in the semantics where the intruder is limited to composition only (i.e., he cannot perform decomposition/analysis of deducible messages). \ theorem wt_attack_if_tfr_attack: assumes "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "\ \\<^sub>c \S, \\" and "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" and "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "tfr\<^sub>s\<^sub>t S" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" obtains \\<^sub>\ where "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" and "\\<^sub>\ \\<^sub>c \S, \\" and "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" proof - have tfr: "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t (LI_preproc S))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t (LI_preproc S))" "list_all tfr\<^sub>s\<^sub>t\<^sub>p (LI_preproc S)" using assms(5,6) LI_preproc_preserves_tfr unfolding tfr\<^sub>s\<^sub>t_def by (metis, metis LI_preproc_trms_eq, metis) have wf_constr: "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r (LI_preproc S) \" by (metis LI_preproc_preserves_wellformedness assms(3)) obtain S' \' where *: "simple S'" "(LI_preproc S,\) \\<^sup>* (S',\')" "\{}; S'\\<^sub>c \" using LI_completeness[OF assms(3,2)] unfolding constr_sem_c_def by (meson term.order_refl) have **: "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S' \'" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \'" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S')" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \')" using LI_preserves_welltypedness[OF *(2) wf_constr assms(4,7) tfr] LI_preserves_wellformedness[OF *(2) wf_constr] LI_preserves_tfr[OF *(2) wf_constr assms(4,7) tfr] by metis+ define A where "A \ {x \ vars\<^sub>s\<^sub>t S'. \X F. Inequality X F \ set S' \ x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ x \ set X}" define B where "B \ UNIV - A" let ?\ = "rm_vars B \" have gr\: "ground (subst_range \)" "ground (subst_range ?\)" using assms(1) rm_vars_img_subset[of B \] by (auto simp add: subst_domain_def) { fix X F assume "Inequality X F \ set S'" hence *: "ineq_model \ X F" using strand_sem_c_imp_ineq_model[OF *(3)] by (auto simp del: subst_range.simps) hence "ineq_model ?\ X F" proof - { fix \ assume 1: "subst_domain \ = set X" "ground (subst_range \)" and 2: "list_ex (\f. fst f \ \ \\<^sub>s \ \ snd f \ \ \\<^sub>s \) F" have "list_ex (\f. fst f \ \ \\<^sub>s rm_vars B \ \ snd f \ \ \\<^sub>s rm_vars B \) F" using 2 proof (induction F) case (Cons g G) obtain t t' where g: "g = (t,t')" by (metis surj_pair) thus ?case using Cons Unifier_ground_rm_vars[OF gr\(1), of "t \ \" B "t' \ \"] by auto qed simp } thus ?thesis using * unfolding ineq_model_def list_ex_iff case_prod_unfold by simp qed } moreover have "subst_domain \ = UNIV" using assms(1) by metis hence "subst_domain ?\ = A" using rm_vars_dom[of B \] B_def by blast ultimately obtain \\<^sub>\ where "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "\\<^sub>\ \\<^sub>c \S', \'\" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" using wt_sat_if_simple[OF *(1) **(1,2,5,4) _ gr\(2) _ **(3)] A_def by (auto simp del: subst_range.simps) thus ?thesis using that LI_soundness[OF assms(3) *(2)] by metis qed text \ Contra-positive version: if a type-flaw resistant constraint does not have a well-typed model then it is unsatisfiable \ corollary secure_if_wt_secure: assumes "\(\\\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ (\\<^sub>\ \\<^sub>c \S, \\) \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\)" and "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "tfr\<^sub>s\<^sub>t S" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "\(\\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ (\ \\<^sub>c \S, \\))" using wt_attack_if_tfr_attack[OF _ _ assms(2,3,4,5,6)] assms(1) by metis end subsection \Lifting the Composition-Only Typing Result to the Full Intruder Model\ context typing_result begin subsubsection \Analysis Invariance\ definition (in typed_model) Ana_invar_subst where "Ana_invar_subst \ \ (\f T K M \. Fun f T \ (subterms\<^sub>s\<^sub>e\<^sub>t \) \ Ana (Fun f T) = (K, M) \ Ana (Fun f T \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \))" lemma (in typed_model) Ana_invar_subst_subset: assumes "Ana_invar_subst M" "N \ M" shows "Ana_invar_subst N" using assms unfolding Ana_invar_subst_def by blast lemma (in typed_model) Ana_invar_substD: assumes "Ana_invar_subst \" and "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t \" "Ana (Fun f T) = (K, M)" shows "Ana (Fun f T \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using assms Ana_invar_subst_def by blast end subsubsection \Preliminary Definitions\ text \Strands extended with "decomposition steps"\ datatype (funs\<^sub>e\<^sub>s\<^sub>t\<^sub>p: 'a, vars\<^sub>e\<^sub>s\<^sub>t\<^sub>p: 'b) extstrand_step = Step "('a,'b) strand_step" | Decomp "('a,'b) term" context typing_result begin context begin private fun trms\<^sub>e\<^sub>s\<^sub>t\<^sub>p where "trms\<^sub>e\<^sub>s\<^sub>t\<^sub>p (Step x) = trms\<^sub>s\<^sub>t\<^sub>p x" | "trms\<^sub>e\<^sub>s\<^sub>t\<^sub>p (Decomp t) = {t}" private abbreviation trms\<^sub>e\<^sub>s\<^sub>t where "trms\<^sub>e\<^sub>s\<^sub>t S \ \(trms\<^sub>e\<^sub>s\<^sub>t\<^sub>p ` set S)" private type_synonym ('a,'b) extstrand = "('a,'b) extstrand_step list" private type_synonym ('a,'b) extstrands = "('a,'b) extstrand set" private definition decomp::"('fun,'var) term \ ('fun,'var) strand" where "decomp t \ (case (Ana t) of (K,T) \ [send\[t]\\<^sub>s\<^sub>t,send\K\\<^sub>s\<^sub>t,receive\T\\<^sub>s\<^sub>t])" private fun to_st where "to_st [] = []" | "to_st (Step x#S) = x#(to_st S)" | "to_st (Decomp t#S) = (decomp t)@(to_st S)" private fun to_est where "to_est [] = []" | "to_est (x#S) = Step x#to_est S" private abbreviation "ik\<^sub>e\<^sub>s\<^sub>t A \ ik\<^sub>s\<^sub>t (to_st A)" private abbreviation "wf\<^sub>e\<^sub>s\<^sub>t V A \ wf\<^sub>s\<^sub>t V (to_st A)" private abbreviation "assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>s\<^sub>t (to_st A)" private abbreviation "vars\<^sub>e\<^sub>s\<^sub>t A \ vars\<^sub>s\<^sub>t (to_st A)" private abbreviation "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t A \ wfrestrictedvars\<^sub>s\<^sub>t (to_st A)" private abbreviation "bvars\<^sub>e\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>t (to_st A)" private abbreviation "fv\<^sub>e\<^sub>s\<^sub>t A \ fv\<^sub>s\<^sub>t (to_st A)" private abbreviation "funs\<^sub>e\<^sub>s\<^sub>t A \ funs\<^sub>s\<^sub>t (to_st A)" private definition wf\<^sub>s\<^sub>t\<^sub>s'::"('fun,'var) strands \ ('fun,'var) extstrand \ bool" where "wf\<^sub>s\<^sub>t\<^sub>s' \ \ \ (\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \) (dual\<^sub>s\<^sub>t S)) \ (\S \ \. \S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}) \ (\S \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t \ = {}) \ (\S \ \. fv\<^sub>s\<^sub>t (to_st \) \ bvars\<^sub>s\<^sub>t S = {})" private definition wf\<^sub>s\<^sub>t\<^sub>s::"('fun,'var) strands \ bool" where "wf\<^sub>s\<^sub>t\<^sub>s \ \ (\S \ \. wf\<^sub>s\<^sub>t {} (dual\<^sub>s\<^sub>t S)) \ (\S \ \. \S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {})" private inductive well_analyzed::"('fun,'var) extstrand \ bool" where Nil[simp]: "well_analyzed []" | Step: "well_analyzed A \ well_analyzed (A@[Step x])" | Decomp: "\well_analyzed A; t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) - (Var ` \)\ \ well_analyzed (A@[Decomp t])" private fun subst_apply_extstrandstep (infix "\\<^sub>e\<^sub>s\<^sub>t\<^sub>p" 51) where "subst_apply_extstrandstep (Step x) \ = Step (x \\<^sub>s\<^sub>t\<^sub>p \)" | "subst_apply_extstrandstep (Decomp t) \ = Decomp (t \ \)" private lemma subst_apply_extstrandstep'_simps[simp]: "(Step (send\ts\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \ = Step (send\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\<^sub>s\<^sub>t)" "(Step (receive\ts\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \ = Step (receive\ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\\<^sub>s\<^sub>t)" "(Step (\a: t \ t'\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \ = Step (\a: (t \ \) \ (t' \ \)\\<^sub>s\<^sub>t)" "(Step (\X\\\: F\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \ = Step (\X\\\: (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)\\<^sub>s\<^sub>t)" by simp_all private lemma vars\<^sub>e\<^sub>s\<^sub>t\<^sub>p_subst_apply_simps[simp]: "vars\<^sub>e\<^sub>s\<^sub>t\<^sub>p ((Step (send\ts\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "vars\<^sub>e\<^sub>s\<^sub>t\<^sub>p ((Step (receive\ts\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" "vars\<^sub>e\<^sub>s\<^sub>t\<^sub>p ((Step (\a: t \ t'\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" "vars\<^sub>e\<^sub>s\<^sub>t\<^sub>p ((Step (\X\\\: F\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \) = set X \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" by auto private definition subst_apply_extstrand (infix "\\<^sub>e\<^sub>s\<^sub>t" 51) where "S \\<^sub>e\<^sub>s\<^sub>t \ \ map (\x. x \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \) S" private abbreviation update\<^sub>s\<^sub>t::"('fun,'var) strands \ ('fun,'var) strand \ ('fun,'var) strands" where "update\<^sub>s\<^sub>t \ S \ (case S of Nil \ \ - {S} | Cons _ S' \ insert S' (\ - {S}))" private inductive_set decomps\<^sub>e\<^sub>s\<^sub>t:: "('fun,'var) terms \ ('fun,'var) terms \ ('fun,'var) subst \ ('fun,'var) extstrands" (* \: intruder knowledge \: additional messages *) for \ and \ and \ where Nil: "[] \ decomps\<^sub>e\<^sub>s\<^sub>t \ \ \" | Decomp: "\\ \ decomps\<^sub>e\<^sub>s\<^sub>t \ \ \; Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (\ \ \); Ana (Fun f T) = (K,M); M \ []; (\ \ ik\<^sub>e\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c Fun f T \ \; \k. k \ set K \ (\ \ ik\<^sub>e\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \\ \ \@[Decomp (Fun f T)] \ decomps\<^sub>e\<^sub>s\<^sub>t \ \ \" private fun decomp_rm\<^sub>e\<^sub>s\<^sub>t::"('fun,'var) extstrand \ ('fun,'var) extstrand" where "decomp_rm\<^sub>e\<^sub>s\<^sub>t [] = []" | "decomp_rm\<^sub>e\<^sub>s\<^sub>t (Decomp t#S) = decomp_rm\<^sub>e\<^sub>s\<^sub>t S" | "decomp_rm\<^sub>e\<^sub>s\<^sub>t (Step x#S) = Step x#(decomp_rm\<^sub>e\<^sub>s\<^sub>t S)" private inductive sem\<^sub>e\<^sub>s\<^sub>t_d::"('fun,'var) terms \ ('fun,'var) subst \ ('fun,'var) extstrand \ bool" where Nil[simp]: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ []" | Send: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ S \ \t \ set ts. (ik\<^sub>e\<^sub>s\<^sub>t S \ M\<^sub>0) \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \ \ sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ (S@[Step (send\ts\\<^sub>s\<^sub>t)])" | Receive: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ S \ sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ (S@[Step (receive\t\\<^sub>s\<^sub>t)])" | Equality: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ S \ t \ \ = t' \ \ \ sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ (S@[Step (\a: t \ t'\\<^sub>s\<^sub>t)])" | Inequality: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ S \ ineq_model \ X F \ sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ (S@[Step (\X\\\: F\\<^sub>s\<^sub>t)])" | Decompose: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ S \ (ik\<^sub>e\<^sub>s\<^sub>t S \ M\<^sub>0) \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \ \ Ana t = (K, M) \ (\k. k \ set K \ (ik\<^sub>e\<^sub>s\<^sub>t S \ M\<^sub>0) \\<^sub>s\<^sub>e\<^sub>t \ \ k \ \) \ sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ (S@[Decomp t])" private inductive sem\<^sub>e\<^sub>s\<^sub>t_c::"('fun,'var) terms \ ('fun,'var) subst \ ('fun,'var) extstrand \ bool" where Nil[simp]: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ []" | Send: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ S \ \t \ set ts. (ik\<^sub>e\<^sub>s\<^sub>t S \ M\<^sub>0) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \ \ sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ (S@[Step (send\ts\\<^sub>s\<^sub>t)])" | Receive: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ S \ sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ (S@[Step (receive\t\\<^sub>s\<^sub>t)])" | Equality: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ S \ t \ \ = t' \ \ \ sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ (S@[Step (\a: t \ t'\\<^sub>s\<^sub>t)])" | Inequality: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ S \ ineq_model \ X F \ sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ (S@[Step (\X\\\: F\\<^sub>s\<^sub>t)])" | Decompose: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ S \ (ik\<^sub>e\<^sub>s\<^sub>t S \ M\<^sub>0) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \ \ Ana t = (K, M) \ (\k. k \ set K \ (ik\<^sub>e\<^sub>s\<^sub>t S \ M\<^sub>0) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \) \ sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ (S@[Decomp t])" subsubsection \Preliminary Lemmata\ private lemma wf\<^sub>s\<^sub>t\<^sub>s_wf\<^sub>s\<^sub>t\<^sub>s': "wf\<^sub>s\<^sub>t\<^sub>s \ = wf\<^sub>s\<^sub>t\<^sub>s' \ []" by (simp add: wf\<^sub>s\<^sub>t\<^sub>s_def wf\<^sub>s\<^sub>t\<^sub>s'_def) private lemma decomp_ik: assumes "Ana t = (K,M)" shows "ik\<^sub>s\<^sub>t (decomp t) = set M" using ik_rcv_map ik_rcv_map' by (auto simp add: decomp_def inv_def assms) private lemma decomp_assignment_rhs_empty: assumes "Ana t = (K,M)" shows "assignment_rhs\<^sub>s\<^sub>t (decomp t) = {}" by (auto simp add: decomp_def inv_def assms) private lemma decomp_tfr\<^sub>s\<^sub>t\<^sub>p: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (decomp t)" by (auto simp add: decomp_def list_all_def) private lemma trms\<^sub>e\<^sub>s\<^sub>t_ikI: "t \ ik\<^sub>e\<^sub>s\<^sub>t A \ t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>e\<^sub>s\<^sub>t A)" proof (induction A rule: to_st.induct) case (2 x S) thus ?case by (cases x) auto next case (3 t' A) obtain K M where Ana: "Ana t' = (K,M)" by (metis surj_pair) show ?case using 3 decomp_ik[OF Ana] Ana_subterm[OF Ana] by auto qed simp private lemma trms\<^sub>e\<^sub>s\<^sub>t_ik_assignment_rhsI: "t \ ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \ t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>e\<^sub>s\<^sub>t A)" proof (induction A rule: to_st.induct) case (2 x S) thus ?case proof (cases x) case (Equality ac t t') thus ?thesis using 2 by (cases ac) auto qed auto next case (3 t' A) obtain K M where Ana: "Ana t' = (K,M)" by (metis surj_pair) show ?case using 3 decomp_ik[OF Ana] decomp_assignment_rhs_empty[OF Ana] Ana_subterm[OF Ana] by auto qed simp private lemma trms\<^sub>e\<^sub>s\<^sub>t_ik_subtermsI: assumes "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A)" shows "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>e\<^sub>s\<^sub>t A)" proof - obtain t' where "t' \ ik\<^sub>e\<^sub>s\<^sub>t A" "t \ t'" using trms\<^sub>e\<^sub>s\<^sub>t_ikI assms by auto thus ?thesis by (meson contra_subsetD in_subterms_subset_Union trms\<^sub>e\<^sub>s\<^sub>t_ikI) qed private lemma trms\<^sub>e\<^sub>s\<^sub>tD: assumes "t \ trms\<^sub>e\<^sub>s\<^sub>t A" shows "t \ trms\<^sub>s\<^sub>t (to_st A)" using assms proof (induction A) case (Cons a A) obtain K M where Ana: "Ana t = (K,M)" by (metis surj_pair) hence "t \ trms\<^sub>s\<^sub>t (decomp t)" unfolding decomp_def by force thus ?case using Cons.IH Cons.prems by (cases a) auto qed simp private lemma subst_apply_extstrand_nil[simp]: "[] \\<^sub>e\<^sub>s\<^sub>t \ = []" by (simp add: subst_apply_extstrand_def) private lemma subst_apply_extstrand_singleton[simp]: "[Step (receive\ts\\<^sub>s\<^sub>t)] \\<^sub>e\<^sub>s\<^sub>t \ = [Step (Receive (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \))]" "[Step (send\ts\\<^sub>s\<^sub>t)] \\<^sub>e\<^sub>s\<^sub>t \ = [Step (Send (ts \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \))]" "[Step (\a: t \ t'\\<^sub>s\<^sub>t)] \\<^sub>e\<^sub>s\<^sub>t \ = [Step (Equality a (t \ \) (t' \ \))]" "[Decomp t] \\<^sub>e\<^sub>s\<^sub>t \ = [Decomp (t \ \)]" unfolding subst_apply_extstrand_def by auto private lemma extstrand_subst_hom: "(S@S') \\<^sub>e\<^sub>s\<^sub>t \ = (S \\<^sub>e\<^sub>s\<^sub>t \)@(S' \\<^sub>e\<^sub>s\<^sub>t \)" "(x#S) \\<^sub>e\<^sub>s\<^sub>t \ = (x \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \)#(S \\<^sub>e\<^sub>s\<^sub>t \)" unfolding subst_apply_extstrand_def by auto private lemma decomp_vars: "wfrestrictedvars\<^sub>s\<^sub>t (decomp t) = fv t" "vars\<^sub>s\<^sub>t (decomp t) = fv t" "bvars\<^sub>s\<^sub>t (decomp t) = {}" "fv\<^sub>s\<^sub>t (decomp t) = fv t" proof - obtain K M where Ana: "Ana t = (K,M)" by (metis surj_pair) hence "decomp t = [send\[t]\\<^sub>s\<^sub>t,Send K,Receive M]" unfolding decomp_def by simp moreover have "\(set (map fv K)) = fv\<^sub>s\<^sub>e\<^sub>t (set K)" "\(set (map fv M)) = fv\<^sub>s\<^sub>e\<^sub>t (set M)" by auto moreover have "fv\<^sub>s\<^sub>e\<^sub>t (set K) \ fv t" "fv\<^sub>s\<^sub>e\<^sub>t (set M) \ fv t" using Ana_subterm[OF Ana(1)] Ana_keys_fv[OF Ana(1)] by (simp_all add: UN_least psubsetD subtermeq_vars_subset) ultimately show "wfrestrictedvars\<^sub>s\<^sub>t (decomp t) = fv t" "vars\<^sub>s\<^sub>t (decomp t) = fv t" "bvars\<^sub>s\<^sub>t (decomp t) = {}" "fv\<^sub>s\<^sub>t (decomp t) = fv t" by auto qed private lemma bvars\<^sub>e\<^sub>s\<^sub>t_cons: "bvars\<^sub>e\<^sub>s\<^sub>t (x#X) = bvars\<^sub>e\<^sub>s\<^sub>t [x] \ bvars\<^sub>e\<^sub>s\<^sub>t X" by (cases x) auto private lemma bvars\<^sub>e\<^sub>s\<^sub>t_append: "bvars\<^sub>e\<^sub>s\<^sub>t (A@B) = bvars\<^sub>e\<^sub>s\<^sub>t A \ bvars\<^sub>e\<^sub>s\<^sub>t B" proof (induction A) case (Cons x A) thus ?case using bvars\<^sub>e\<^sub>s\<^sub>t_cons[of x "A@B"] bvars\<^sub>e\<^sub>s\<^sub>t_cons[of x A] by force qed simp private lemma fv\<^sub>e\<^sub>s\<^sub>t_cons: "fv\<^sub>e\<^sub>s\<^sub>t (x#X) = fv\<^sub>e\<^sub>s\<^sub>t [x] \ fv\<^sub>e\<^sub>s\<^sub>t X" by (cases x) auto private lemma fv\<^sub>e\<^sub>s\<^sub>t_append: "fv\<^sub>e\<^sub>s\<^sub>t (A@B) = fv\<^sub>e\<^sub>s\<^sub>t A \ fv\<^sub>e\<^sub>s\<^sub>t B" proof (induction A) case (Cons x A) thus ?case using fv\<^sub>e\<^sub>s\<^sub>t_cons[of x "A@B"] fv\<^sub>e\<^sub>s\<^sub>t_cons[of x A] by auto qed simp private lemma bvars_decomp: "bvars\<^sub>e\<^sub>s\<^sub>t (A@[Decomp t]) = bvars\<^sub>e\<^sub>s\<^sub>t A" "bvars\<^sub>e\<^sub>s\<^sub>t (Decomp t#A) = bvars\<^sub>e\<^sub>s\<^sub>t A" using bvars\<^sub>e\<^sub>s\<^sub>t_append decomp_vars(3) by fastforce+ private lemma bvars_decomp_rm: "bvars\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) = bvars\<^sub>e\<^sub>s\<^sub>t A" using bvars_decomp by (induct A rule: decomp_rm\<^sub>e\<^sub>s\<^sub>t.induct) simp_all+ private lemma fv_decomp_rm: "fv\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ fv\<^sub>e\<^sub>s\<^sub>t A" by (induct A rule: decomp_rm\<^sub>e\<^sub>s\<^sub>t.induct) auto private lemma ik_assignment_rhs_decomp_fv: assumes "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" shows "fv\<^sub>e\<^sub>s\<^sub>t (A@[Decomp t]) = fv\<^sub>e\<^sub>s\<^sub>t A" proof - have "fv\<^sub>e\<^sub>s\<^sub>t (A@[Decomp t]) = fv\<^sub>e\<^sub>s\<^sub>t A \ fv t" using fv\<^sub>e\<^sub>s\<^sub>t_append decomp_vars by simp moreover have "fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \ fv\<^sub>e\<^sub>s\<^sub>t A" by force moreover have "fv t \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" using fv_subset_subterms[OF assms(1)] by simp ultimately show ?thesis by blast qed private lemma wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_decomp_rm\<^sub>e\<^sub>s\<^sub>t_subset: "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t A" by (induct A rule: decomp_rm\<^sub>e\<^sub>s\<^sub>t.induct) auto+ private lemma wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_eq_wfrestrictedvars\<^sub>s\<^sub>t: "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t A = wfrestrictedvars\<^sub>s\<^sub>t (to_st A)" by simp private lemma decomp_set_unfold: assumes "Ana t = (K, M)" shows "set (decomp t) = {send\[t]\\<^sub>s\<^sub>t,send\K\\<^sub>s\<^sub>t,receive\M\\<^sub>s\<^sub>t}" using assms unfolding decomp_def by auto private lemma ik\<^sub>e\<^sub>s\<^sub>t_finite: "finite (ik\<^sub>e\<^sub>s\<^sub>t A)" by (rule finite_ik\<^sub>s\<^sub>t) private lemma assignment_rhs\<^sub>e\<^sub>s\<^sub>t_finite: "finite (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" by (rule finite_assignment_rhs\<^sub>s\<^sub>t) private lemma to_est_append: "to_est (A@B) = to_est A@to_est B" by (induct A rule: to_est.induct) auto private lemma to_st_to_est_inv: "to_st (to_est A) = A" by (induct A rule: to_est.induct) auto private lemma to_st_append: "to_st (A@B) = (to_st A)@(to_st B)" by (induct A rule: to_st.induct) auto private lemma to_st_cons: "to_st (a#B) = (to_st [a])@(to_st B)" using to_st_append[of "[a]" B] by simp private lemma wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_split: "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t (x#S) = wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t [x] \ wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t S" "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t (S@S') = wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t S \ wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t S'" using to_st_cons[of x S] to_st_append[of S S'] by auto private lemma ik\<^sub>e\<^sub>s\<^sub>t_append: "ik\<^sub>e\<^sub>s\<^sub>t (A@B) = ik\<^sub>e\<^sub>s\<^sub>t A \ ik\<^sub>e\<^sub>s\<^sub>t B" by (metis ik_append to_st_append) private lemma assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append: "assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@B) = assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t B" by (metis assignment_rhs_append to_st_append) private lemma ik\<^sub>e\<^sub>s\<^sub>t_cons: "ik\<^sub>e\<^sub>s\<^sub>t (a#A) = ik\<^sub>e\<^sub>s\<^sub>t [a] \ ik\<^sub>e\<^sub>s\<^sub>t A" by (metis ik_append to_st_cons) private lemma ik\<^sub>e\<^sub>s\<^sub>t_append_subst: "ik\<^sub>e\<^sub>s\<^sub>t (A@B \\<^sub>e\<^sub>s\<^sub>t \) = ik\<^sub>e\<^sub>s\<^sub>t (A \\<^sub>e\<^sub>s\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t (B \\<^sub>e\<^sub>s\<^sub>t \)" "ik\<^sub>e\<^sub>s\<^sub>t (A@B) \\<^sub>s\<^sub>e\<^sub>t \ = (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t B \\<^sub>s\<^sub>e\<^sub>t \)" by (metis ik\<^sub>e\<^sub>s\<^sub>t_append extstrand_subst_hom(1), simp add: image_Un to_st_append) private lemma assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append_subst: "assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@B \\<^sub>e\<^sub>s\<^sub>t \) = assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A \\<^sub>e\<^sub>s\<^sub>t \) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t (B \\<^sub>e\<^sub>s\<^sub>t \)" "assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@B) \\<^sub>s\<^sub>e\<^sub>t \ = (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t B \\<^sub>s\<^sub>e\<^sub>t \)" by (metis assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append extstrand_subst_hom(1), use assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append in blast) private lemma ik\<^sub>e\<^sub>s\<^sub>t_cons_subst: "ik\<^sub>e\<^sub>s\<^sub>t (a#A \\<^sub>e\<^sub>s\<^sub>t \) = ik\<^sub>e\<^sub>s\<^sub>t ([a \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \]) \ ik\<^sub>e\<^sub>s\<^sub>t (A \\<^sub>e\<^sub>s\<^sub>t \)" "ik\<^sub>e\<^sub>s\<^sub>t (a#A) \\<^sub>s\<^sub>e\<^sub>t \ = (ik\<^sub>e\<^sub>s\<^sub>t [a] \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \)" by (metis ik\<^sub>e\<^sub>s\<^sub>t_cons extstrand_subst_hom(2), metis image_Un ik\<^sub>e\<^sub>s\<^sub>t_cons) private lemma decomp_rm\<^sub>e\<^sub>s\<^sub>t_append: "decomp_rm\<^sub>e\<^sub>s\<^sub>t (S@S') = (decomp_rm\<^sub>e\<^sub>s\<^sub>t S)@(decomp_rm\<^sub>e\<^sub>s\<^sub>t S')" by (induct S rule: decomp_rm\<^sub>e\<^sub>s\<^sub>t.induct) auto private lemma decomp_rm\<^sub>e\<^sub>s\<^sub>t_single[simp]: "decomp_rm\<^sub>e\<^sub>s\<^sub>t [Step (send\ts\\<^sub>s\<^sub>t)] = [Step (send\ts\\<^sub>s\<^sub>t)]" "decomp_rm\<^sub>e\<^sub>s\<^sub>t [Step (receive\ts\\<^sub>s\<^sub>t)] = [Step (receive\ts\\<^sub>s\<^sub>t)]" "decomp_rm\<^sub>e\<^sub>s\<^sub>t [Decomp t] = []" by auto private lemma decomp_rm\<^sub>e\<^sub>s\<^sub>t_ik_subset: "ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t S) \ ik\<^sub>e\<^sub>s\<^sub>t S" proof (induction S rule: decomp_rm\<^sub>e\<^sub>s\<^sub>t.induct) case (3 x S) thus ?case by (cases x) auto qed auto private lemma decomps\<^sub>e\<^sub>s\<^sub>t_ik_subset: "D \ decomps\<^sub>e\<^sub>s\<^sub>t M N \ \ ik\<^sub>e\<^sub>s\<^sub>t D \ subterms\<^sub>s\<^sub>e\<^sub>t (M \ N)" proof (induction D rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case (Decomp D f T K M') have "ik\<^sub>s\<^sub>t (decomp (Fun f T)) \ subterms (Fun f T)" "ik\<^sub>s\<^sub>t (decomp (Fun f T)) = ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun f T)]" using decomp_ik[OF Decomp.hyps(3)] Ana_subterm[OF Decomp.hyps(3)] by auto hence "ik\<^sub>s\<^sub>t (to_st [Decomp (Fun f T)]) \ subterms\<^sub>s\<^sub>e\<^sub>t (M \ N)" using in_subterms_subset_Union[OF Decomp.hyps(2)] by blast thus ?case using ik\<^sub>e\<^sub>s\<^sub>t_append[of D "[Decomp (Fun f T)]"] using Decomp.IH by auto qed simp private lemma decomps\<^sub>e\<^sub>s\<^sub>t_decomp_rm\<^sub>e\<^sub>s\<^sub>t_empty: "D \ decomps\<^sub>e\<^sub>s\<^sub>t M N \ \ decomp_rm\<^sub>e\<^sub>s\<^sub>t D = []" by (induct D rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) (auto simp add: decomp_rm\<^sub>e\<^sub>s\<^sub>t_append) private lemma decomps\<^sub>e\<^sub>s\<^sub>t_append: assumes "A \ decomps\<^sub>e\<^sub>s\<^sub>t S N \" "B \ decomps\<^sub>e\<^sub>s\<^sub>t S N \" shows "A@B \ decomps\<^sub>e\<^sub>s\<^sub>t S N \" using assms(2) proof (induction B rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case Nil show ?case using assms(1) by simp next case (Decomp B f X K T) hence "S \ ik\<^sub>e\<^sub>s\<^sub>t B \\<^sub>s\<^sub>e\<^sub>t \ \ S \ ik\<^sub>e\<^sub>s\<^sub>t (A@B) \\<^sub>s\<^sub>e\<^sub>t \" using ik\<^sub>e\<^sub>s\<^sub>t_append by auto thus ?case using decomps\<^sub>e\<^sub>s\<^sub>t.Decomp[OF Decomp.IH(1) Decomp.hyps(2,3,4)] ideduct_synth_mono[OF Decomp.hyps(5)] ideduct_synth_mono[OF Decomp.hyps(6)] by auto qed private lemma decomps\<^sub>e\<^sub>s\<^sub>t_subterms: assumes "A' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" shows "subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A') \ subterms\<^sub>s\<^sub>e\<^sub>t (M \ N)" using assms proof (induction A' rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case (Decomp D f X K T) hence "Fun f X \ subterms\<^sub>s\<^sub>e\<^sub>t (M \ N)" by auto hence "subterms\<^sub>s\<^sub>e\<^sub>t (set X) \ subterms\<^sub>s\<^sub>e\<^sub>t (M \ N)" using in_subterms_subset_Union[of "Fun f X" "M \ N"] params_subterms_Union[of X f] by blast moreover have "ik\<^sub>s\<^sub>t (to_st [Decomp (Fun f X)]) = set T" using Decomp.hyps(3) decomp_ik by simp hence "subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t (to_st [Decomp (Fun f X)])) \ subterms\<^sub>s\<^sub>e\<^sub>t (set X)" using Ana_fun_subterm[OF Decomp.hyps(3)] by auto ultimately show ?case using ik\<^sub>e\<^sub>s\<^sub>t_append[of D "[Decomp (Fun f X)]"] Decomp.IH by auto qed simp private lemma decomps\<^sub>e\<^sub>s\<^sub>t_assignment_rhs_empty: assumes "A' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" shows "assignment_rhs\<^sub>e\<^sub>s\<^sub>t A' = {}" using assms by (induction A' rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) (simp_all add: decomp_assignment_rhs_empty assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append) private lemma decomps\<^sub>e\<^sub>s\<^sub>t_finite_ik_append: assumes "finite M" "M \ decomps\<^sub>e\<^sub>s\<^sub>t A N \" shows "\D \ decomps\<^sub>e\<^sub>s\<^sub>t A N \. ik\<^sub>e\<^sub>s\<^sub>t D = (\m \ M. ik\<^sub>e\<^sub>s\<^sub>t m)" using assms proof (induction M rule: finite_induct) case empty moreover have "[] \ decomps\<^sub>e\<^sub>s\<^sub>t A N \" "ik\<^sub>s\<^sub>t (to_st []) = {}" using decomps\<^sub>e\<^sub>s\<^sub>t.Nil by auto ultimately show ?case by blast next case (insert m M) then obtain D where "D \ decomps\<^sub>e\<^sub>s\<^sub>t A N \" "ik\<^sub>e\<^sub>s\<^sub>t D = (\m\M. ik\<^sub>s\<^sub>t (to_st m))" by moura moreover have "m \ decomps\<^sub>e\<^sub>s\<^sub>t A N \" using insert.prems(1) by blast ultimately show ?case using decomps\<^sub>e\<^sub>s\<^sub>t_append[of D A N \ m] ik\<^sub>e\<^sub>s\<^sub>t_append[of D m] by blast qed private lemma decomp_snd_exists[simp]: "\D. decomp t = send\[t]\\<^sub>s\<^sub>t#D" by (metis (mono_tags, lifting) decomp_def prod.case surj_pair) private lemma decomp_nonnil[simp]: "decomp t \ []" using decomp_snd_exists[of t] by fastforce private lemma to_st_nil_inv[dest]: "to_st A = [] \ A = []" by (induct A rule: to_st.induct) auto private lemma well_analyzedD: assumes "well_analyzed A" "Decomp t \ set A" shows "\f T. t = Fun f T" using assms proof (induction A rule: well_analyzed.induct) case (Decomp A t') hence "\f T. t' = Fun f T" by (cases t') auto moreover have "Decomp t \ set A \ t = t'" using Decomp by auto ultimately show ?case using Decomp.IH by auto qed auto private lemma well_analyzed_inv: assumes "well_analyzed (A@[Decomp t])" shows "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) - (Var ` \)" using assms well_analyzed.cases[of "A@[Decomp t]"] by fastforce private lemma well_analyzed_split_left_single: "well_analyzed (A@[a]) \ well_analyzed A" by (induction "A@[a]" rule: well_analyzed.induct) auto private lemma well_analyzed_split_left: "well_analyzed (A@B) \ well_analyzed A" proof (induction B rule: List.rev_induct) case (snoc b B) thus ?case using well_analyzed_split_left_single[of "A@B" b] by simp qed simp private lemma well_analyzed_append: assumes "well_analyzed A" "well_analyzed B" shows "well_analyzed (A@B)" using assms(2,1) proof (induction B rule: well_analyzed.induct) case (Step B x) show ?case using well_analyzed.Step[OF Step.IH[OF Step.prems]] by simp next case (Decomp B t) thus ?case using well_analyzed.Decomp[OF Decomp.IH[OF Decomp.prems]] ik\<^sub>e\<^sub>s\<^sub>t_append assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append by auto qed simp_all private lemma well_analyzed_singleton: "well_analyzed [Step (send\ts\\<^sub>s\<^sub>t)]" "well_analyzed [Step (receive\ts\\<^sub>s\<^sub>t)]" "well_analyzed [Step (\a: t \ t'\\<^sub>s\<^sub>t)]" "well_analyzed [Step (\X\\\: F\\<^sub>s\<^sub>t)]" "\well_analyzed [Decomp t]" proof - show "well_analyzed [Step (send\ts\\<^sub>s\<^sub>t)]" "well_analyzed [Step (receive\ts\\<^sub>s\<^sub>t)]" "well_analyzed [Step (\a: t \ t'\\<^sub>s\<^sub>t)]" "well_analyzed [Step (\X\\\: F\\<^sub>s\<^sub>t)]" using well_analyzed.Step[OF well_analyzed.Nil] by simp_all show "\well_analyzed [Decomp t]" using well_analyzed.cases[of "[Decomp t]"] by auto qed private lemma well_analyzed_decomp_rm\<^sub>e\<^sub>s\<^sub>t_fv: "well_analyzed A \ fv\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) = fv\<^sub>e\<^sub>s\<^sub>t A" proof assume "well_analyzed A" thus "fv\<^sub>e\<^sub>s\<^sub>t A \ fv\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A)" proof (induction A rule: well_analyzed.induct) case Decomp thus ?case using ik_assignment_rhs_decomp_fv decomp_rm\<^sub>e\<^sub>s\<^sub>t_append by auto next case (Step A x) have "fv\<^sub>e\<^sub>s\<^sub>t (A@[Step x]) = fv\<^sub>e\<^sub>s\<^sub>t A \ fv\<^sub>s\<^sub>t\<^sub>p x" "fv\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t (A@[Step x])) = fv\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ fv\<^sub>s\<^sub>t\<^sub>p x" using fv\<^sub>e\<^sub>s\<^sub>t_append decomp_rm\<^sub>e\<^sub>s\<^sub>t_append by auto thus ?case using Step by auto qed simp qed (rule fv_decomp_rm) private lemma sem\<^sub>e\<^sub>s\<^sub>t_d_split_left: assumes "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ (\@\')" shows "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ \" using assms sem\<^sub>e\<^sub>s\<^sub>t_d.cases by (induction \' rule: List.rev_induct) fastforce+ private lemma sem\<^sub>e\<^sub>s\<^sub>t_d_eq_sem_st: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ \ = \M\<^sub>0; to_st \\\<^sub>d' \" proof show "\M\<^sub>0; to_st \\\<^sub>d' \ \ sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ \" proof (induction \ arbitrary: M\<^sub>0 rule: List.rev_induct) case Nil show ?case using to_st_nil_inv by simp next case (snoc a \) hence IH: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ \" and *: "\ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0; to_st [a]\\<^sub>d' \" using to_st_append by (auto simp add: sup.commute) thus ?case using snoc proof (cases a) case (Step b) thus ?thesis proof (cases b) case (Send t) thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_d.Send[OF IH] * Step by auto next case (Receive t) thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_d.Receive[OF IH] Step by auto next case (Equality a t t') thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_d.Equality[OF IH] * Step by auto next case (Inequality X F) thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_d.Inequality[OF IH] * Step by auto qed next case (Decomp t) obtain K M where Ana: "Ana t = (K,M)" by moura have "to_st [a] = decomp t" using Decomp by auto hence "to_st [a] = [send\[t]\\<^sub>s\<^sub>t,Send K,Receive M]" using Ana unfolding decomp_def by auto hence **: "ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" and "\ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0; [Send K]\\<^sub>d' \" using * by auto hence "\k. k \ set K \ ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ k \ \" using * strand_sem_Send_split(2) strand_sem_d.simps(2) unfolding strand_sem_eq_defs(2) list_all_iff by meson thus ?thesis using Decomp sem\<^sub>e\<^sub>s\<^sub>t_d.Decompose[OF IH ** Ana] by metis qed qed show "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ \ \ \M\<^sub>0; to_st \\\<^sub>d' \" proof (induction rule: sem\<^sub>e\<^sub>s\<^sub>t_d.induct) case Nil thus ?case by simp next case (Send M\<^sub>0 \ \ ts) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[send\ts\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (send\ts\\<^sub>s\<^sub>t)]"] by (simp add: sup.commute) next case (Receive M\<^sub>0 \ \ ts) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[receive\ts\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (receive\ts\\<^sub>s\<^sub>t)]"] by (simp add: sup.commute) next case (Equality M\<^sub>0 \ \ t t' a) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[\a: t \ t'\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (\a: t \ t'\\<^sub>s\<^sub>t)]"] by (simp add: sup.commute) next case (Inequality M\<^sub>0 \ \ X F) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[\X\\\: F\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (\X\\\: F\\<^sub>s\<^sub>t)]"] by (simp add: sup.commute) next case (Decompose M\<^sub>0 \ \ t K M) have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); decomp t\\<^sub>d' \" proof - have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [send\[t]\\<^sub>s\<^sub>t]\\<^sub>d' \" using Decompose.hyps(2) by (auto simp add: sup.commute) moreover have "\k. k \ set K \ M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \) \\<^sub>s\<^sub>e\<^sub>t \ \ k \ \" using Decompose by (metis sup.commute) hence "\k. k \ set K \ \M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [Send1 k]\\<^sub>d' \" by auto hence "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [Send K]\\<^sub>d' \" using strand_sem_Send_map(4)[of _ "M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \) \\<^sub>s\<^sub>e\<^sub>t \" \] strand_sem_Send_map(6) unfolding strand_sem_eq_defs(2) by auto moreover have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [Receive M]\\<^sub>d' \" by (metis strand_sem_Receive_map(6) strand_sem_eq_defs(2)) ultimately have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [send\[t]\\<^sub>s\<^sub>t,send\K\\<^sub>s\<^sub>t,receive\M\\<^sub>s\<^sub>t]\\<^sub>d' \" by auto thus ?thesis using Decompose.hyps(3) unfolding decomp_def by auto qed hence "\M\<^sub>0; to_st \@decomp t\\<^sub>d' \" using strand_sem_append'[of M\<^sub>0 "to_st \" \ "decomp t"] Decompose.IH by simp thus ?case using to_st_append[of \ "[Decomp t]"] by simp qed qed private lemma sem\<^sub>e\<^sub>s\<^sub>t_c_eq_sem_st: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ \ = \M\<^sub>0; to_st \\\<^sub>c' \" proof show "\M\<^sub>0; to_st \\\<^sub>c' \ \ sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ \" proof (induction \ arbitrary: M\<^sub>0 rule: List.rev_induct) case Nil show ?case using to_st_nil_inv by simp next case (snoc a \) hence IH: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ \" and *: "\ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0; to_st [a]\\<^sub>c' \" using to_st_append by (auto simp add: sup.commute) thus ?case using snoc proof (cases a) case (Step b) thus ?thesis proof (cases b) case (Send t) thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_c.Send[OF IH] * Step by auto next case (Receive t) thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_c.Receive[OF IH] Step by auto next case (Equality t) thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_c.Equality[OF IH] * Step by auto next case (Inequality t) thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_c.Inequality[OF IH] * Step by auto qed next case (Decomp t) obtain K M where Ana: "Ana t = (K,M)" by moura have "to_st [a] = decomp t" using Decomp by auto hence "to_st [a] = [send\[t]\\<^sub>s\<^sub>t,send\K\\<^sub>s\<^sub>t,receive\M\\<^sub>s\<^sub>t]" using Ana unfolding decomp_def by auto hence **: "ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" and "\ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0; [send\K\\<^sub>s\<^sub>t]\\<^sub>c' \" using * by auto hence "ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \" when k: "k \ set K" for k using * strand_sem_Send_split(5)[OF _ k] strand_sem_Send_map(5) unfolding strand_sem_eq_defs(1) by auto thus ?thesis using Decomp sem\<^sub>e\<^sub>s\<^sub>t_c.Decompose[OF IH ** Ana] by metis qed qed show "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ \ \ \M\<^sub>0; to_st \\\<^sub>c' \" proof (induction rule: sem\<^sub>e\<^sub>s\<^sub>t_c.induct) case Nil thus ?case by simp next case (Send M\<^sub>0 \ \ ts) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[send\ts\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (send\ts\\<^sub>s\<^sub>t)]"] by (simp add: sup.commute) next case (Receive M\<^sub>0 \ \ ts) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[receive\ts\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (receive\ts\\<^sub>s\<^sub>t)]"] by (simp add: sup.commute) next case (Equality M\<^sub>0 \ \ t t' a) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[\a: t \ t'\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (\a: t \ t'\\<^sub>s\<^sub>t)]"] by (simp add: sup.commute) next case (Inequality M\<^sub>0 \ \ X F) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[\X\\\: F\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (\X\\\: F\\<^sub>s\<^sub>t)]"] by (auto simp add: sup.commute) next case (Decompose M\<^sub>0 \ \ t K M) have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); decomp t\\<^sub>c' \" proof - have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [send\[t]\\<^sub>s\<^sub>t]\\<^sub>c' \" using Decompose.hyps(2) by (auto simp add: sup.commute) moreover have "\k. k \ set K \ M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \" using Decompose by (metis sup.commute) hence "\k. k \ set K \ \M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [Send1 k]\\<^sub>c' \" by auto hence "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [Send K]\\<^sub>c' \" using strand_sem_Send_map(3)[of K, of "M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \) \\<^sub>s\<^sub>e\<^sub>t \" \] strand_sem_Send_map(5) unfolding strand_sem_eq_defs(1) by auto moreover have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [Receive M]\\<^sub>c' \" by (metis strand_sem_Receive_map(5) strand_sem_eq_defs(1)) ultimately have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [send\[t]\\<^sub>s\<^sub>t,send\K\\<^sub>s\<^sub>t,receive\M\\<^sub>s\<^sub>t]\\<^sub>c' \" by auto thus ?thesis using Decompose.hyps(3) unfolding decomp_def by auto qed hence "\M\<^sub>0; to_st \@decomp t\\<^sub>c' \" using strand_sem_append'[of M\<^sub>0 "to_st \" \ "decomp t"] Decompose.IH by simp thus ?case using to_st_append[of \ "[Decomp t]"] by simp qed qed private lemma sem\<^sub>e\<^sub>s\<^sub>t_c_decomp_rm\<^sub>e\<^sub>s\<^sub>t_deduct_aux: assumes "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ A" "t \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" "t \ ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t \" shows "ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ t" using assms proof (induction M\<^sub>0 \ A arbitrary: t rule: sem\<^sub>e\<^sub>s\<^sub>t_c.induct) case (Send M\<^sub>0 \ A t') thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append ik\<^sub>e\<^sub>s\<^sub>t_append by auto next case (Receive M\<^sub>0 \ A t') hence "t \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" "t \ ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t \" using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append ik\<^sub>e\<^sub>s\<^sub>t_append by auto hence IH: "ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ t" using Receive.IH by auto show ?case using ideduct_mono[OF IH] decomp_rm\<^sub>e\<^sub>s\<^sub>t_append ik\<^sub>e\<^sub>s\<^sub>t_append by (metis Un_subset_iff Un_upper1 Un_upper2 image_mono) next case (Equality M\<^sub>0 \ A t') thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append ik\<^sub>e\<^sub>s\<^sub>t_append by auto next case (Inequality M\<^sub>0 \ A t') thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append ik\<^sub>e\<^sub>s\<^sub>t_append by auto next case (Decompose M\<^sub>0 \ A t' K M t) have *: "ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ t' \ \" using Decompose.hyps(2) proof (induction rule: intruder_synth_induct) case (AxiomC t'') moreover { assume "t'' \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" "t'' \ ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t \" hence ?case using Decompose.IH by auto } ultimately show ?case by force qed simp { fix k assume "k \ set K" hence "ik\<^sub>e\<^sub>s\<^sub>t A \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \" using Decompose.hyps by auto hence "ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ k \ \" proof (induction rule: intruder_synth_induct) case (AxiomC t'') moreover { assume "t'' \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" "t'' \ ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t \" hence ?case using Decompose.IH by auto } ultimately show ?case by force qed simp } hence **: "\k. k \ set (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ k" by auto show ?case proof (cases "t \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \") case True thus ?thesis using Decompose.IH Decompose.prems(2) decomp_rm\<^sub>e\<^sub>s\<^sub>t_append by auto next case False hence "t \ ik\<^sub>s\<^sub>t (decomp t') \\<^sub>s\<^sub>e\<^sub>t \" using Decompose.prems(1) ik\<^sub>e\<^sub>s\<^sub>t_append by auto hence ***: "t \ set (M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Decompose.hyps(3) decomp_ik by auto hence "M \ []" by auto hence ****: "Ana (t' \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Ana_subst[OF Decompose.hyps(3)] by auto have "ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ t" by (rule intruder_deduct.Decompose[OF * **** ** ***]) thus ?thesis using ideduct_mono decomp_rm\<^sub>e\<^sub>s\<^sub>t_append by auto qed qed simp private lemma sem\<^sub>e\<^sub>s\<^sub>t_c_decomp_rm\<^sub>e\<^sub>s\<^sub>t_deduct: assumes "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ A" "ik\<^sub>e\<^sub>s\<^sub>t A \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t" shows "ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ t" using assms(2) proof (induction t rule: intruder_synth_induct) case (AxiomC t) hence "t \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ \ t \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \" by auto moreover { assume "t \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" "t \ ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t \" hence ?case using ideduct_mono[OF intruder_deduct.Axiom] by auto } moreover { assume "t \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" "t \ ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t \" hence ?case using sem\<^sub>e\<^sub>s\<^sub>t_c_decomp_rm\<^sub>e\<^sub>s\<^sub>t_deduct_aux[OF assms(1)] by auto } ultimately show ?case by auto qed simp private lemma sem\<^sub>e\<^sub>s\<^sub>t_d_decomp_rm\<^sub>e\<^sub>s\<^sub>t_if_sem\<^sub>e\<^sub>s\<^sub>t_c: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ A \ sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ (decomp_rm\<^sub>e\<^sub>s\<^sub>t A)" proof (induction M\<^sub>0 \ A rule: sem\<^sub>e\<^sub>s\<^sub>t_c.induct) case (Send M\<^sub>0 \ A t) thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append sem\<^sub>e\<^sub>s\<^sub>t_d.Send[OF Send.IH] sem\<^sub>e\<^sub>s\<^sub>t_c_decomp_rm\<^sub>e\<^sub>s\<^sub>t_deduct unfolding list_all_iff by auto next case (Receive t) thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append sem\<^sub>e\<^sub>s\<^sub>t_d.Receive by auto next case (Equality M\<^sub>0 \ A t) thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append sem\<^sub>e\<^sub>s\<^sub>t_d.Equality[OF Equality.IH] sem\<^sub>e\<^sub>s\<^sub>t_c_decomp_rm\<^sub>e\<^sub>s\<^sub>t_deduct by auto next case (Inequality M\<^sub>0 \ A t) thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append sem\<^sub>e\<^sub>s\<^sub>t_d.Inequality[OF Inequality.IH] sem\<^sub>e\<^sub>s\<^sub>t_c_decomp_rm\<^sub>e\<^sub>s\<^sub>t_deduct by auto next case Decompose thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append by auto qed auto private lemma sem\<^sub>e\<^sub>s\<^sub>t_c_decomps\<^sub>e\<^sub>s\<^sub>t_append: assumes "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ A" "D \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \) \" shows "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ (A@D)" using assms(2,1) proof (induction D rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case (Decomp D f T K M) hence *: "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ (A @ D)" "ik\<^sub>e\<^sub>s\<^sub>t (A@D) \ {} \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c Fun f T \ \" "\k. k \ set K \ ik\<^sub>e\<^sub>s\<^sub>t (A @ D) \ {} \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \" using ik\<^sub>e\<^sub>s\<^sub>t_append by auto show ?case using sem\<^sub>e\<^sub>s\<^sub>t_c.Decompose[OF *(1,2) Decomp.hyps(3) *(3)] by simp qed auto private lemma decomps\<^sub>e\<^sub>s\<^sub>t_preserves_wf: assumes "D \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \" "wf\<^sub>e\<^sub>s\<^sub>t V A" shows "wf\<^sub>e\<^sub>s\<^sub>t V (A@D)" using assms proof (induction D rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case (Decomp D f T K M) have "wfrestrictedvars\<^sub>s\<^sub>t (decomp (Fun f T)) \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" using decomp_vars fv_subset_subterms[OF Decomp.hyps(2)] by fast hence "wfrestrictedvars\<^sub>s\<^sub>t (decomp (Fun f T)) \ wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t A" using ik\<^sub>s\<^sub>t_assignment_rhs\<^sub>s\<^sub>t_wfrestrictedvars_subset[of "to_st A"] by blast hence "wfrestrictedvars\<^sub>s\<^sub>t (decomp (Fun f T)) \ wfrestrictedvars\<^sub>s\<^sub>t (to_st (A@D)) \ V" using to_st_append[of A D] strand_vars_split(2)[of "to_st A" "to_st D"] by (metis le_supI1) thus ?case using wf_append_suffix[OF Decomp.IH[OF Decomp.prems], of "decomp (Fun f T)"] to_st_append[of "A@D" "[Decomp (Fun f T)]"] by auto qed auto private lemma decomps\<^sub>e\<^sub>s\<^sub>t_preserves_model_c: assumes "D \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \" "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ A" shows "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ (A@D)" using assms proof (induction D rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case (Decomp D f T K M) show ?case using sem\<^sub>e\<^sub>s\<^sub>t_c.Decompose[OF Decomp.IH[OF Decomp.prems] _ Decomp.hyps(3)] Decomp.hyps(5,6) ideduct_synth_mono ik\<^sub>e\<^sub>s\<^sub>t_append by (metis (mono_tags, lifting) List.append_assoc image_Un sup_ge1) qed auto private lemma decomps\<^sub>e\<^sub>s\<^sub>t_exist_aux: assumes "D \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "M \ ik\<^sub>e\<^sub>s\<^sub>t D \ t" "\(M \ (ik\<^sub>e\<^sub>s\<^sub>t D) \\<^sub>c t)" obtains D' where "D@D' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "M \ ik\<^sub>e\<^sub>s\<^sub>t (D@D') \\<^sub>c t" "M \ ik\<^sub>e\<^sub>s\<^sub>t D \ M \ ik\<^sub>e\<^sub>s\<^sub>t (D@D')" proof - have "\D' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \. M \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>c t" using assms(2) proof (induction t rule: intruder_deduct_induct) case (Compose X f) from Compose.IH have "\D \ decomps\<^sub>e\<^sub>s\<^sub>t M N \. \x \ set X. M \ ik\<^sub>e\<^sub>s\<^sub>t D \\<^sub>c x" proof (induction X) case (Cons t X) then obtain D' D'' where D': "D' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "M \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>c t" and D'': "D'' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "\x \ set X. M \ ik\<^sub>e\<^sub>s\<^sub>t D'' \\<^sub>c x" by moura hence "M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>c t" "\x \ set X. M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>c x" by (auto intro: ideduct_synth_mono simp add: ik\<^sub>e\<^sub>s\<^sub>t_append) thus ?case using decomps\<^sub>e\<^sub>s\<^sub>t_append[OF D'(1) D''(1)] by (metis set_ConsD) qed (auto intro: decomps\<^sub>e\<^sub>s\<^sub>t.Nil) thus ?case using intruder_synth.ComposeC[OF Compose.hyps(1,2)] by metis next case (Decompose t K T t\<^sub>i) have "\D \ decomps\<^sub>e\<^sub>s\<^sub>t M N \. \k \ set K. M \ ik\<^sub>e\<^sub>s\<^sub>t D \\<^sub>c k" using Decompose.IH proof (induction K) case (Cons t X) then obtain D' D'' where D': "D' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "M \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>c t" and D'': "D'' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "\x \ set X. M \ ik\<^sub>e\<^sub>s\<^sub>t D'' \\<^sub>c x" using assms(1) by moura hence "M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>c t" "\x \ set X. M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>c x" by (auto intro: ideduct_synth_mono simp add: ik\<^sub>e\<^sub>s\<^sub>t_append) thus ?case using decomps\<^sub>e\<^sub>s\<^sub>t_append[OF D'(1) D''(1)] by auto qed auto then obtain D' where D': "D' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "\k. k \ set K \ M \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>c k" by metis obtain D'' where D'': "D'' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "M \ ik\<^sub>e\<^sub>s\<^sub>t D'' \\<^sub>c t" by (metis Decompose.IH(1)) obtain f X where fX: "t = Fun f X" "t\<^sub>i \ set X" using Decompose.hyps(2,4) by (cases t) (auto dest: Ana_fun_subterm) from decomps\<^sub>e\<^sub>s\<^sub>t_append[OF D'(1) D''(1)] D'(2) D''(2) have *: "D'@D'' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "\k. k \ set K \ M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>c k" "M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>c t" by (auto intro: ideduct_synth_mono simp add: ik\<^sub>e\<^sub>s\<^sub>t_append) hence **: "\k. k \ set K \ M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \" using ideduct_synth_subst by auto have "t\<^sub>i \ ik\<^sub>s\<^sub>t (decomp t)" using Decompose.hyps(2,4) ik_rcv_map unfolding decomp_def by auto with *(3) fX(1) Decompose.hyps(2) show ?case proof (induction t rule: intruder_synth_induct) case (AxiomC t) hence t_in_subterms: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (M \ N)" using decomps\<^sub>e\<^sub>s\<^sub>t_ik_subset[OF *(1)] subset_subterms_Union by auto have "M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" using ideduct_synth_subst[OF intruder_synth.AxiomC[OF AxiomC.hyps(1)]] by metis moreover have "T \ []" using decomp_ik[OF \Ana t = (K,T)\] \t\<^sub>i \ ik\<^sub>s\<^sub>t (decomp t)\ by auto ultimately have "D'@D''@[Decomp (Fun f X)] \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" using AxiomC decomps\<^sub>e\<^sub>s\<^sub>t.Decomp[OF *(1) _ _ _ _ **] subset_subterms_Union t_in_subterms by (simp add: subset_eq) moreover have "decomp t = to_st [Decomp (Fun f X)]" using AxiomC.prems(1,2) by auto ultimately show ?case by (metis AxiomC.prems(3) UnCI intruder_synth.AxiomC ik\<^sub>e\<^sub>s\<^sub>t_append to_st_append) qed (auto intro!: fX(2) *(1)) qed (fastforce intro: intruder_synth.AxiomC assms(1)) hence "\D' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \. M \ ik\<^sub>e\<^sub>s\<^sub>t (D@D') \\<^sub>c t" by (auto intro: ideduct_synth_mono simp add: ik\<^sub>e\<^sub>s\<^sub>t_append) thus thesis using that[OF decomps\<^sub>e\<^sub>s\<^sub>t_append[OF assms(1)]] assms ik\<^sub>e\<^sub>s\<^sub>t_append by moura qed private lemma decomps\<^sub>e\<^sub>s\<^sub>t_ik_max_exist: assumes "finite A" "finite N" shows "\D \ decomps\<^sub>e\<^sub>s\<^sub>t A N \. \D' \ decomps\<^sub>e\<^sub>s\<^sub>t A N \. ik\<^sub>e\<^sub>s\<^sub>t D' \ ik\<^sub>e\<^sub>s\<^sub>t D" proof - let ?IK = "\M. \D \ M. ik\<^sub>e\<^sub>s\<^sub>t D" have "?IK (decomps\<^sub>e\<^sub>s\<^sub>t A N \) \ (\t \ A \ N. subterms t)" by (auto dest!: decomps\<^sub>e\<^sub>s\<^sub>t_ik_subset) hence "finite (?IK (decomps\<^sub>e\<^sub>s\<^sub>t A N \))" using subterms_union_finite[OF assms(1)] subterms_union_finite[OF assms(2)] infinite_super by auto then obtain M where M: "finite M" "M \ decomps\<^sub>e\<^sub>s\<^sub>t A N \" "?IK M = ?IK (decomps\<^sub>e\<^sub>s\<^sub>t A N \)" using finite_subset_Union by moura show ?thesis using decomps\<^sub>e\<^sub>s\<^sub>t_finite_ik_append[OF M(1,2)] M(3) by auto qed private lemma decomps\<^sub>e\<^sub>s\<^sub>t_exist: assumes "finite A" "finite N" shows "\D \ decomps\<^sub>e\<^sub>s\<^sub>t A N \. \t. A \ t \ A \ ik\<^sub>e\<^sub>s\<^sub>t D \\<^sub>c t" proof (rule ccontr) assume neg: "\(\D \ decomps\<^sub>e\<^sub>s\<^sub>t A N \. \t. A \ t \ A \ ik\<^sub>e\<^sub>s\<^sub>t D \\<^sub>c t)" obtain D where D: "D \ decomps\<^sub>e\<^sub>s\<^sub>t A N \" "\D' \ decomps\<^sub>e\<^sub>s\<^sub>t A N \. ik\<^sub>e\<^sub>s\<^sub>t D' \ ik\<^sub>e\<^sub>s\<^sub>t D" using decomps\<^sub>e\<^sub>s\<^sub>t_ik_max_exist[OF assms] by moura then obtain t where t: "A \ ik\<^sub>e\<^sub>s\<^sub>t D \ t" "\(A \ ik\<^sub>e\<^sub>s\<^sub>t D \\<^sub>c t)" using neg by (fastforce intro: ideduct_mono) obtain D' where D': "D@D' \ decomps\<^sub>e\<^sub>s\<^sub>t A N \" "A \ ik\<^sub>e\<^sub>s\<^sub>t (D@D') \\<^sub>c t" "A \ ik\<^sub>e\<^sub>s\<^sub>t D \ A \ ik\<^sub>e\<^sub>s\<^sub>t (D@D')" by (metis decomps\<^sub>e\<^sub>s\<^sub>t_exist_aux t D(1)) hence "ik\<^sub>e\<^sub>s\<^sub>t D \ ik\<^sub>e\<^sub>s\<^sub>t (D@D')" using ik\<^sub>e\<^sub>s\<^sub>t_append by auto moreover have "ik\<^sub>e\<^sub>s\<^sub>t (D@D') \ ik\<^sub>e\<^sub>s\<^sub>t D" using D(2) D'(1) by auto ultimately show False by simp qed private lemma decomps\<^sub>e\<^sub>s\<^sub>t_exist_subst: assumes "ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" and "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ A" "wf\<^sub>e\<^sub>s\<^sub>t {} A" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "Ana_invar_subst (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" and "well_analyzed A" shows "\D \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \. ik\<^sub>e\<^sub>s\<^sub>t (A@D) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" proof - have ik_eq: "ik\<^sub>e\<^sub>s\<^sub>t (A \\<^sub>e\<^sub>s\<^sub>t \) = ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" using assms(5,6) proof (induction A rule: List.rev_induct) case (snoc a A) hence "Ana_invar_subst (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" using Ana_invar_subst_subset[OF snoc.prems(1)] ik\<^sub>e\<^sub>s\<^sub>t_append assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append unfolding Ana_invar_subst_def by simp with snoc have IH: "ik\<^sub>e\<^sub>s\<^sub>t (A@[a] \\<^sub>e\<^sub>s\<^sub>t \) = (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t ([a] \\<^sub>e\<^sub>s\<^sub>t \)" "ik\<^sub>e\<^sub>s\<^sub>t (A@[a]) \\<^sub>s\<^sub>e\<^sub>t \ = (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t [a] \\<^sub>s\<^sub>e\<^sub>t \)" using well_analyzed_split_left[OF snoc.prems(2)] by (auto simp add: to_st_append ik\<^sub>e\<^sub>s\<^sub>t_append_subst) have "ik\<^sub>e\<^sub>s\<^sub>t [a \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \] = ik\<^sub>e\<^sub>s\<^sub>t [a] \\<^sub>s\<^sub>e\<^sub>t \" proof (cases a) case (Step b) thus ?thesis by (cases b) auto next case (Decomp t) then obtain f T where t: "t = Fun f T" using well_analyzedD[OF snoc.prems(2)] by force obtain K M where Ana_t: "Ana (Fun f T) = (K,M)" by (metis surj_pair) moreover have "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t ((ik\<^sub>e\<^sub>s\<^sub>t (A@[a]) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@[a])))" using t Decomp snoc.prems(2) by (auto dest: well_analyzed_inv simp add: ik\<^sub>e\<^sub>s\<^sub>t_append assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append) hence "Ana (Fun f T \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Ana_t snoc.prems(1) unfolding Ana_invar_subst_def by blast ultimately show ?thesis using Decomp t by (auto simp add: decomp_ik) qed thus ?case using IH unfolding subst_apply_extstrand_def by simp qed simp moreover have assignment_rhs_eq: "assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A \\<^sub>e\<^sub>s\<^sub>t \) = assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" using assms(5,6) proof (induction A rule: List.rev_induct) case (snoc a A) hence "Ana_invar_subst (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" using Ana_invar_subst_subset[OF snoc.prems(1)] ik\<^sub>e\<^sub>s\<^sub>t_append assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append unfolding Ana_invar_subst_def by simp hence "assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A \\<^sub>e\<^sub>s\<^sub>t \) = assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" using snoc.IH well_analyzed_split_left[OF snoc.prems(2)] by simp hence IH: "assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@[a] \\<^sub>e\<^sub>s\<^sub>t \) = (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t ([a] \\<^sub>e\<^sub>s\<^sub>t \)" "assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@[a]) \\<^sub>s\<^sub>e\<^sub>t \ = (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t [a] \\<^sub>s\<^sub>e\<^sub>t \)" by (metis assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append_subst(1), metis assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append_subst(2)) have "assignment_rhs\<^sub>e\<^sub>s\<^sub>t [a \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \] = assignment_rhs\<^sub>e\<^sub>s\<^sub>t [a] \\<^sub>s\<^sub>e\<^sub>t \" proof (cases a) case (Step b) thus ?thesis by (cases b) auto next case (Decomp t) then obtain f T where t: "t = Fun f T" using well_analyzedD[OF snoc.prems(2)] by force obtain K M where Ana_t: "Ana (Fun f T) = (K,M)" by (metis surj_pair) moreover have "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t ((ik\<^sub>e\<^sub>s\<^sub>t (A@[a]) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@[a])))" using t Decomp snoc.prems(2) by (auto dest: well_analyzed_inv simp add: ik\<^sub>e\<^sub>s\<^sub>t_append assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append) hence "Ana (Fun f T \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Ana_t snoc.prems(1) unfolding Ana_invar_subst_def by blast ultimately show ?thesis using Decomp t by (auto simp add: decomp_assignment_rhs_empty) qed thus ?case using IH unfolding subst_apply_extstrand_def by simp qed simp ultimately obtain D where D: "D \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) Var" "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D) \\<^sub>c t \ \" using decomps\<^sub>e\<^sub>s\<^sub>t_exist[OF ik\<^sub>e\<^sub>s\<^sub>t_finite assignment_rhs\<^sub>e\<^sub>s\<^sub>t_finite, of "A \\<^sub>e\<^sub>s\<^sub>t \" "A \\<^sub>e\<^sub>s\<^sub>t \"] ik\<^sub>e\<^sub>s\<^sub>t_append assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append assms(1) by force let ?P = "\D D'. \t. (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D) \\<^sub>c t \ (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t" have "\D' \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \. ?P D D'" using D(1) proof (induction D rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case Nil have "ik\<^sub>e\<^sub>s\<^sub>t [] = ik\<^sub>e\<^sub>s\<^sub>t [] \\<^sub>s\<^sub>e\<^sub>t \" by auto thus ?case by (metis decomps\<^sub>e\<^sub>s\<^sub>t.Nil) next case (Decomp D f T K M) obtain D' where D': "D' \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \" "?P D D'" using Decomp.IH by auto hence IH: "\k. k \ set K \ (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c k" "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c Fun f T" using Decomp.hyps(5,6) by auto have D'_ik: "ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \ \ subterms\<^sub>s\<^sub>e\<^sub>t ((ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)) \\<^sub>s\<^sub>e\<^sub>t \" "ik\<^sub>e\<^sub>s\<^sub>t D' \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" using decomps\<^sub>e\<^sub>s\<^sub>t_ik_subset[OF D'(1)] by (metis subst_all_mono, metis) show ?case using IH(2,1) Decomp.hyps(2,3,4) proof (induction "Fun f T" arbitrary: f T K M rule: intruder_synth_induct) case (AxiomC f T) then obtain s where s: "s \ ik\<^sub>e\<^sub>s\<^sub>t A \ ik\<^sub>e\<^sub>s\<^sub>t D'" "Fun f T = s \ \" using AxiomC.prems by blast hence fT_s_in: "Fun f T \ (subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)) \\<^sub>s\<^sub>e\<^sub>t \" "s \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" using AxiomC D'_ik subset_subterms_Union[of "ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A"] subst_all_mono[OF subset_subterms_Union, of \] by (metis (no_types) Un_iff image_eqI subset_Un_eq, metis (no_types) Un_iff subset_Un_eq) obtain Ks Ms where Ana_s: "Ana s = (Ks,Ms)" by moura have AD'_props: "wf\<^sub>e\<^sub>s\<^sub>t {} (A@D')" "\{}; to_st (A@D')\\<^sub>c \" using decomps\<^sub>e\<^sub>s\<^sub>t_preserves_model_c[OF D'(1) assms(2)] decomps\<^sub>e\<^sub>s\<^sub>t_preserves_wf[OF D'(1) assms(3)] sem\<^sub>e\<^sub>s\<^sub>t_c_eq_sem_st strand_sem_eq_defs(1) by auto show ?case proof (cases s) case (Var x) \ \In this case \\ x\ (is a subterm of something that) was derived from an "earlier intruder knowledge" because \A\ is well-formed and has \\\ as a model. So either the intruder composed \Fun f T\ himself (making \Decomp (Fun f T)\ unnecessary) or \Fun f T\ is an instance of something else in the intruder knowledge (in which case the "something" can be used in place of \Fun f T\)\ hence "Var x \ ik\<^sub>e\<^sub>s\<^sub>t (A@D')" "\ x = Fun f T" using s ik\<^sub>e\<^sub>s\<^sub>t_append by auto show ?thesis proof (cases "\m \ set M. ik\<^sub>e\<^sub>s\<^sub>t A \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c m") case True \ \All terms acquired by decomposing \Fun f T\ are already derivable. Hence there is no need to consider decomposition of \Fun f T\ at all.\ have *: "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t (D@[Decomp (Fun f T)]) = (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t D \ set M" using decomp_ik[OF \Ana (Fun f T) = (K,M)\] ik\<^sub>e\<^sub>s\<^sub>t_append[of D "[Decomp (Fun f T)]"] by auto { fix t' assume "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t D \ set M \\<^sub>c t'" hence "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t'" proof (induction t' rule: intruder_synth_induct) case (AxiomC t') thus ?case proof assume "t' \ set M" moreover have "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) = ik\<^sub>e\<^sub>s\<^sub>t A \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \" by auto ultimately show ?case using True by auto qed (metis D'(2) intruder_synth.AxiomC) qed auto } thus ?thesis using D'(1) * by metis next case False \ \Some term acquired by decomposition of \Fun f T\ cannot be derived in \\\<^sub>c\. \Fun f T\ must therefore be an instance of something else in the intruder knowledge, because of well-formedness.\ then obtain t\<^sub>i where t\<^sub>i: "t\<^sub>i \ set T" "\ik\<^sub>e\<^sub>s\<^sub>t (A@D') \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t\<^sub>i" using Ana_fun_subterm[OF \Ana (Fun f T) = (K,M)\] by (auto simp add: ik\<^sub>e\<^sub>s\<^sub>t_append) obtain S where fS: "Fun f S \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t (A@D')) \ Fun f S \ subterms\<^sub>s\<^sub>e\<^sub>t (assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@D'))" "\ x = Fun f S \ \" using strand_sem_wf_ik_or_assignment_rhs_fun_subterm[ OF AD'_props \Var x \ ik\<^sub>e\<^sub>s\<^sub>t (A@D')\ _ t\<^sub>i \interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\] \\ x = Fun f T\ by moura hence fS_in: "Fun f S \ \ \ ik\<^sub>e\<^sub>s\<^sub>t A \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \" "Fun f S \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" using imageI[OF s(1), of "\x. x \ \"] Var ik\<^sub>e\<^sub>s\<^sub>t_append[of A D'] assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append[of A D'] decomps\<^sub>e\<^sub>s\<^sub>t_subterms[OF D'(1)] decomps\<^sub>e\<^sub>s\<^sub>t_assignment_rhs_empty[OF D'(1)] by auto obtain KS MS where Ana_fS: "Ana (Fun f S) = (KS, MS)" by moura hence "K = KS \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" "M = MS \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using Ana_invar_substD[OF assms(5) fS_in(2)] s(2) fS(2) \s = Var x\ \Ana (Fun f T) = (K,M)\ by simp_all hence "MS \ []" using \M \ []\ by simp have "\k. k \ set KS \ ik\<^sub>e\<^sub>s\<^sub>t A \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \" using AxiomC.prems(1) \K = KS \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\ by (simp add: image_Un) hence D'': "D'@[Decomp (Fun f S)] \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \" using decomps\<^sub>e\<^sub>s\<^sub>t.Decomp[OF D'(1) fS_in(2) Ana_fS \MS \ []\] AxiomC.prems(1) intruder_synth.AxiomC[OF fS_in(1)] by simp moreover { fix t' assume "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t (D@[Decomp (Fun f T)]) \\<^sub>c t'" hence "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t (D'@[Decomp (Fun f S)]) \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t'" proof (induction t' rule: intruder_synth_induct) case (AxiomC t') hence "t' \ (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t D \ t' \ ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun f T)]" by (simp add: ik\<^sub>e\<^sub>s\<^sub>t_append) thus ?case proof assume "t' \ ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun f T)]" hence "t' \ ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun f S)] \\<^sub>s\<^sub>e\<^sub>t \" using decomp_ik \Ana (Fun f T) = (K,M)\ \Ana (Fun f S) = (KS,MS)\ \M = MS \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\ by simp thus ?case using ideduct_synth_mono[ OF intruder_synth.AxiomC[of t' "ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun f S)] \\<^sub>s\<^sub>e\<^sub>t \"], of "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t (D'@[Decomp (Fun f S)]) \\<^sub>s\<^sub>e\<^sub>t \)"] by (auto simp add: ik\<^sub>e\<^sub>s\<^sub>t_append) next assume "t' \ (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t D" hence "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t'" by (metis D'(2) intruder_synth.AxiomC) hence "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun f S)] \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t'" by (simp add: ideduct_synth_mono) thus ?case using ik\<^sub>e\<^sub>s\<^sub>t_append[of D' "[Decomp (Fun f S)]"] image_Un[of "\x. x \ \" "ik\<^sub>e\<^sub>s\<^sub>t D'" "ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun f S)]"] by (simp add: sup_aci(2)) qed qed auto } ultimately show ?thesis using D'' by auto qed next case (Fun g S) \ \Hence \Decomp (Fun f T)\ can be substituted for \Decomp (Fun g S)\\ hence KM: "K = Ks \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" "M = Ms \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" "set K = set Ks \\<^sub>s\<^sub>e\<^sub>t \" "set M = set Ms \\<^sub>s\<^sub>e\<^sub>t \" using fT_s_in(2) \Ana (Fun f T) = (K,M)\ Ana_s s(2) Ana_invar_substD[OF assms(5), of g S] by auto hence Ms_nonempty: "Ms \ []" using \M \ []\ by auto { fix t' assume "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t (D@[Decomp (Fun f T)]) \\<^sub>c t'" hence "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t (D'@[Decomp (Fun g S)]) \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t'" using AxiomC proof (induction t' rule: intruder_synth_induct) case (AxiomC t') hence "t' \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ \ t' \ ik\<^sub>e\<^sub>s\<^sub>t D \ t' \ set M" by (simp add: decomp_ik ik\<^sub>e\<^sub>s\<^sub>t_append) thus ?case proof (elim disjE) assume "t' \ ik\<^sub>e\<^sub>s\<^sub>t D" hence *: "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t'" using D'(2) by simp show ?case by (auto intro: ideduct_synth_mono[OF *] simp add: ik\<^sub>e\<^sub>s\<^sub>t_append_subst(2)) next assume "t' \ set M" hence "t' \ ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun g S)] \\<^sub>s\<^sub>e\<^sub>t \" using KM(2) Fun decomp_ik[OF Ana_s] by auto thus ?case by (simp add: image_Un ik\<^sub>e\<^sub>s\<^sub>t_append) qed (simp add: ideduct_synth_mono[OF intruder_synth.AxiomC]) qed auto } thus ?thesis using s Fun Ana_s AxiomC.prems(1) KM(3) fT_s_in decomps\<^sub>e\<^sub>s\<^sub>t.Decomp[OF D'(1) _ _ Ms_nonempty, of g S Ks] by (metis AxiomC.hyps image_Un image_eqI intruder_synth.AxiomC) qed next case (ComposeC T f) have *: "\m. m \ set M \ (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c m" using Ana_fun_subterm[OF \Ana (Fun f T) = (K, M)\] ComposeC.hyps(3) by auto have **: "ik\<^sub>e\<^sub>s\<^sub>t (D@[Decomp (Fun f T)]) = ik\<^sub>e\<^sub>s\<^sub>t D \ set M" using decomp_ik[OF \Ana (Fun f T) = (K, M)\] ik\<^sub>e\<^sub>s\<^sub>t_append by auto { fix t' assume "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t (D@[Decomp (Fun f T)]) \\<^sub>c t'" hence "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t'" by (induct rule: intruder_synth_induct) (auto simp add: D'(2) * **) } thus ?case using D'(1) by auto qed qed thus ?thesis using D(2) assms(1) by (auto simp add: ik\<^sub>e\<^sub>s\<^sub>t_append_subst(2)) qed private lemma decomps\<^sub>e\<^sub>s\<^sub>t_exist_subst_list: assumes "\t \ set ts. ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" and "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ A" "wf\<^sub>e\<^sub>s\<^sub>t {} A" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "Ana_invar_subst (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" and "well_analyzed A" shows "\D \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \. \t \ set ts. ik\<^sub>e\<^sub>s\<^sub>t (A@D) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" (is "\D \ ?A. ?B D ts") proof - note 0 = decomps\<^sub>e\<^sub>s\<^sub>t_exist_subst[OF _ assms(2-6)] show ?thesis using assms(1) proof (induction ts) case (Cons t ts) have 1: "ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" and 2: "\t \ set ts. ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" using Cons.prems by auto obtain D where D: "D \ ?A" "ik\<^sub>e\<^sub>s\<^sub>t (A@D) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" using 0[OF 1] by blast obtain D' where D': "D' \ ?A" "?B D' ts" using Cons.IH[OF 2] by auto have "ik\<^sub>e\<^sub>s\<^sub>t (A@D@D') \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" using ideduct_synth_mono[OF D(2)] ik\<^sub>e\<^sub>s\<^sub>t_append_subst(2)[of \ "A@D" D'] by fastforce hence "?B (D@D') (t#ts)" using D'(2) ideduct_synth_mono[of "ik\<^sub>e\<^sub>s\<^sub>t (A@D') \\<^sub>s\<^sub>e\<^sub>t \" _ "ik\<^sub>e\<^sub>s\<^sub>t (A@D@D') \\<^sub>s\<^sub>e\<^sub>t \"] ik\<^sub>e\<^sub>s\<^sub>t_append_subst(2)[of \] by auto thus ?case using decomps\<^sub>e\<^sub>s\<^sub>t_append[OF D(1) D'(1)] by blast qed (fastforce intro: decomps\<^sub>e\<^sub>s\<^sub>t.Nil) qed private lemma wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_nil: assumes "wf\<^sub>s\<^sub>t\<^sub>s' \ \" shows "wf\<^sub>s\<^sub>t\<^sub>s' (update\<^sub>s\<^sub>t \ []) \" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by auto private lemma wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_snd: assumes "wf\<^sub>s\<^sub>t\<^sub>s' \ \" "send\ts\\<^sub>s\<^sub>t#S \ \" shows "wf\<^sub>s\<^sub>t\<^sub>s' (update\<^sub>s\<^sub>t \ (send\ts\\<^sub>s\<^sub>t#S)) (\@[Step (receive\ts\\<^sub>s\<^sub>t)])" unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def proof (intro conjI) let ?S = "send\ts\\<^sub>s\<^sub>t#S" let ?A = "\@[Step (receive\ts\\<^sub>s\<^sub>t)]" have \: "\S'. S' \ update\<^sub>s\<^sub>t \ ?S \ S' = S \ S' \ \" by auto have 1: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \) (dual\<^sub>s\<^sub>t S)" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by auto moreover have 2: "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A = wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \ \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" using wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_split(2) by (auto simp add: Un_assoc) ultimately have 3: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by (metis wf_vars_mono) have 4: "\S \ \. \S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by simp have "wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" using 1 2 3 assms(2) by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by (metis 3 \) have "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" "\S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t S = {}" using 4 assms(2) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by force+ thus "\S \ update\<^sub>s\<^sub>t \ ?S. \S' \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" by (metis 4 \) have "\S' \ \. fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t ?S = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by metis+ hence 5: "fv\<^sub>e\<^sub>s\<^sub>t ?A = fv\<^sub>e\<^sub>s\<^sub>t \ \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "bvars\<^sub>e\<^sub>s\<^sub>t ?A = bvars\<^sub>e\<^sub>s\<^sub>t \" "\S' \ \. fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ bvars\<^sub>s\<^sub>t S' = {}" using to_st_append by fastforce+ have *: "\S \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using 5 assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fast hence "fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using assms(2) by metis hence "fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by (metis * \) have **: "\S \ \. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" using 5 assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fast hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t ?S = {}" using assms(2) by metis hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by fastforce thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by (metis ** \) qed private lemma wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_rcv: assumes "wf\<^sub>s\<^sub>t\<^sub>s' \ \" "receive\ts\\<^sub>s\<^sub>t#S \ \" shows "wf\<^sub>s\<^sub>t\<^sub>s' (update\<^sub>s\<^sub>t \ (receive\ts\\<^sub>s\<^sub>t#S)) (\@[Step (send\ts\\<^sub>s\<^sub>t)])" unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def proof (intro conjI) let ?S = "receive\ts\\<^sub>s\<^sub>t#S" let ?A = "\@[Step (send\ts\\<^sub>s\<^sub>t)]" have \: "\S'. S' \ update\<^sub>s\<^sub>t \ ?S \ S' = S \ S' \ \" by auto have 1: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \) (dual\<^sub>s\<^sub>t S)" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by auto moreover have 2: "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A = wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \ \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" using wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_split(2) by (auto simp add: Un_assoc) ultimately have 3: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by (metis wf_vars_mono) have 4: "\S \ \. \S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by simp have "wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" using 1 2 3 assms(2) by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by (metis 3 \) have "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" "\S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t S = {}" using 4 assms(2) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by force+ thus "\S \ update\<^sub>s\<^sub>t \ ?S. \S' \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" by (metis 4 \) have "\S' \ \. fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t ?S = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by metis+ hence 5: "fv\<^sub>e\<^sub>s\<^sub>t ?A = fv\<^sub>e\<^sub>s\<^sub>t \ \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "bvars\<^sub>e\<^sub>s\<^sub>t ?A = bvars\<^sub>e\<^sub>s\<^sub>t \" "\S' \ \. fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ bvars\<^sub>s\<^sub>t S' = {}" using to_st_append by fastforce+ have *: "\S \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using 5 assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fast hence "fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using assms(2) by metis hence "fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by (metis * \) have **: "\S \ \. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" using 5 assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fast hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t ?S = {}" using assms(2) by metis hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by fastforce thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by (metis ** \) qed private lemma wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_eq: assumes "wf\<^sub>s\<^sub>t\<^sub>s' \ \" "\a: t \ t'\\<^sub>s\<^sub>t#S \ \" shows "wf\<^sub>s\<^sub>t\<^sub>s' (update\<^sub>s\<^sub>t \ (\a: t \ t'\\<^sub>s\<^sub>t#S)) (\@[Step (\a: t \ t'\\<^sub>s\<^sub>t)])" unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def proof (intro conjI) let ?S = "\a: t \ t'\\<^sub>s\<^sub>t#S" let ?A = "\@[Step (\a: t \ t'\\<^sub>s\<^sub>t)]" have \: "\S'. S' \ update\<^sub>s\<^sub>t \ ?S \ S' = S \ S' \ \" by auto have 1: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \) (dual\<^sub>s\<^sub>t S)" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by auto moreover have 2: "a = Assign \ wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A = wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \ \ fv t \ fv t'" "a = Check \ wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A = wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \" using wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_split(2) by (auto simp add: Un_assoc) ultimately have 3: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by (cases a) (metis wf_vars_mono, metis) have 4: "\S \ \. \S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by simp have "wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" using 1 2 3 assms(2) by (cases a) auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by (metis 3 \) have "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" "\S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t S = {}" using 4 assms(2) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by force+ thus "\S \ update\<^sub>s\<^sub>t \ ?S. \S' \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" by (metis 4 \) have "\S' \ \. fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t ?S = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by metis+ hence 5: "fv\<^sub>e\<^sub>s\<^sub>t ?A = fv\<^sub>e\<^sub>s\<^sub>t \ \ fv t \ fv t'" "bvars\<^sub>e\<^sub>s\<^sub>t ?A = bvars\<^sub>e\<^sub>s\<^sub>t \" "\S' \ \. fv t \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv t' \ bvars\<^sub>s\<^sub>t S' = {}" using to_st_append by fastforce+ have *: "\S \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using 5 assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fast hence "fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using assms(2) by metis hence "fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by (metis * \) have **: "\S \ \. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" using 5 assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fast hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t ?S = {}" using assms(2) by metis hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by fastforce thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by (metis ** \) qed private lemma wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_ineq: assumes "wf\<^sub>s\<^sub>t\<^sub>s' \ \" "\X\\\: F\\<^sub>s\<^sub>t#S \ \" shows "wf\<^sub>s\<^sub>t\<^sub>s' (update\<^sub>s\<^sub>t \ (\X\\\: F\\<^sub>s\<^sub>t#S)) (\@[Step (\X\\\: F\\<^sub>s\<^sub>t)])" unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def proof (intro conjI) let ?S = "\X\\\: F\\<^sub>s\<^sub>t#S" let ?A = "\@[Step (\X\\\: F\\<^sub>s\<^sub>t)]" have \: "\S'. S' \ update\<^sub>s\<^sub>t \ ?S \ S' = S \ S' \ \" by auto have 1: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \) (dual\<^sub>s\<^sub>t S)" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by auto moreover have 2: "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A = wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \" using wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_split(2) by (auto simp add: Un_assoc) ultimately have 3: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by metis have 4: "\S \ \. \S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by simp have "wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" using 1 2 3 assms(2) by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by (metis 3 \) have "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" "\S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t S = {}" using 4 assms(2) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by force+ thus "\S \ update\<^sub>s\<^sub>t \ ?S. \S' \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" by (metis 4 \) have "\S' \ \. fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t ?S = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by metis+ moreover have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ fv\<^sub>s\<^sub>t (\X\\\: F\\<^sub>s\<^sub>t # S)" by auto ultimately have 5: "\S' \ \. (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X) \ bvars\<^sub>s\<^sub>t S' = {}" "fv\<^sub>e\<^sub>s\<^sub>t ?A = fv\<^sub>e\<^sub>s\<^sub>t \ \ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X)" "bvars\<^sub>e\<^sub>s\<^sub>t ?A = set X \ bvars\<^sub>e\<^sub>s\<^sub>t \" "\S \ \. fv\<^sub>s\<^sub>t S \ set X = {}" using to_st_append by (blast, force, force, force) have *: "\S \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using 5(3,4) assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by blast hence "fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using assms(2) by metis hence "fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by (metis * \) have **: "\S \ \. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" using 5(1,2) assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fast hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t ?S = {}" using assms(2) by metis hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by (metis ** \) qed private lemma trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq: assumes "x#S \ \" shows "\(trms\<^sub>s\<^sub>t ` update\<^sub>s\<^sub>t \ (x#S)) \ trms\<^sub>s\<^sub>t\<^sub>p x = \(trms\<^sub>s\<^sub>t ` \)" (is "?A = ?B") proof show "?B \ ?A" proof have "trms\<^sub>s\<^sub>t\<^sub>p x \ trms\<^sub>s\<^sub>t (x#S)" by auto hence "\t'. t' \ ?B \ t' \ trms\<^sub>s\<^sub>t\<^sub>p x \ t' \ ?A" by simp moreover { fix t' assume t': "t' \ ?B" "t' \ trms\<^sub>s\<^sub>t\<^sub>p x" then obtain S' where S': "t' \ trms\<^sub>s\<^sub>t S'" "S' \ \" by auto hence "S' = x#S \ S' \ update\<^sub>s\<^sub>t \ (x#S)" by auto moreover { assume "S' = x#S" hence "t' \ trms\<^sub>s\<^sub>t S" using S' t' by simp hence "t' \ ?A" by auto } ultimately have "t' \ ?A" using t' S' by auto } ultimately show "\t'. t' \ ?B \ t' \ ?A" by metis qed show "?A \ ?B" proof have "\t'. t' \ ?A \ t' \ trms\<^sub>s\<^sub>t\<^sub>p x \ trms\<^sub>s\<^sub>t\<^sub>p x \ ?B" using assms by force+ moreover { fix t' assume t': "t' \ ?A" "t' \ trms\<^sub>s\<^sub>t\<^sub>p x" then obtain S' where "t' \ trms\<^sub>s\<^sub>t S'" "S' \ update\<^sub>s\<^sub>t \ (x#S)" by auto hence "S' = S \ S' \ \" by auto moreover have "trms\<^sub>s\<^sub>t S \ ?B" using assms trms\<^sub>s\<^sub>t_cons[of x S] by blast ultimately have "t' \ ?B" using t' by fastforce } ultimately show "\t'. t' \ ?A \ t' \ ?B" by blast qed qed private lemma trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_snd: assumes "send\ts\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (send\ts\\<^sub>s\<^sub>t#S)" "\' = \@[Step (receive\ts\\<^sub>s\<^sub>t)]" shows "(\(trms\<^sub>s\<^sub>t ` \)) \ (trms\<^sub>e\<^sub>s\<^sub>t \) = (\(trms\<^sub>s\<^sub>t ` \')) \ (trms\<^sub>e\<^sub>s\<^sub>t \')" proof - have "(trms\<^sub>e\<^sub>s\<^sub>t \') = (trms\<^sub>e\<^sub>s\<^sub>t \) \ set ts" "\(trms\<^sub>s\<^sub>t ` \') \ set ts = \(trms\<^sub>s\<^sub>t ` \)" using to_st_append trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq[OF assms(1)] assms(2,3) by auto thus ?thesis by (metis (no_types, lifting) Un_commute Un_left_commute) qed private lemma trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_rcv: assumes "receive\ts\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (receive\ts\\<^sub>s\<^sub>t#S)" "\' = \@[Step (send\ts\\<^sub>s\<^sub>t)]" shows "(\(trms\<^sub>s\<^sub>t ` \)) \ (trms\<^sub>e\<^sub>s\<^sub>t \) = (\(trms\<^sub>s\<^sub>t ` \')) \ (trms\<^sub>e\<^sub>s\<^sub>t \')" proof - have "(trms\<^sub>e\<^sub>s\<^sub>t \') = (trms\<^sub>e\<^sub>s\<^sub>t \) \ set ts" "\(trms\<^sub>s\<^sub>t ` \') \ set ts = \(trms\<^sub>s\<^sub>t ` \)" using to_st_append trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq[OF assms(1)] assms(2,3) by auto thus ?thesis by (metis (no_types, lifting) Un_commute Un_left_commute) qed private lemma trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_eq: assumes "\a: t \ t'\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (\a: t \ t'\\<^sub>s\<^sub>t#S)" "\' = \@[Step (\a: t \ t'\\<^sub>s\<^sub>t)]" shows "(\(trms\<^sub>s\<^sub>t ` \)) \ (trms\<^sub>e\<^sub>s\<^sub>t \) = (\(trms\<^sub>s\<^sub>t ` \')) \ (trms\<^sub>e\<^sub>s\<^sub>t \')" proof - have "(trms\<^sub>e\<^sub>s\<^sub>t \') = (trms\<^sub>e\<^sub>s\<^sub>t \) \ {t,t'}" "\(trms\<^sub>s\<^sub>t ` \') \ {t,t'} = \(trms\<^sub>s\<^sub>t ` \)" using to_st_append trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq[OF assms(1)] assms(2,3) by auto thus ?thesis by (metis (no_types, lifting) Un_insert_left Un_insert_right sup_bot.right_neutral) qed private lemma trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_ineq: assumes "\X\\\: F\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (\X\\\: F\\<^sub>s\<^sub>t#S)" "\' = \@[Step (\X\\\: F\\<^sub>s\<^sub>t)]" shows "(\(trms\<^sub>s\<^sub>t ` \)) \ (trms\<^sub>e\<^sub>s\<^sub>t \) = (\(trms\<^sub>s\<^sub>t ` \')) \ (trms\<^sub>e\<^sub>s\<^sub>t \')" proof - have "(trms\<^sub>e\<^sub>s\<^sub>t \') = (trms\<^sub>e\<^sub>s\<^sub>t \) \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" "\(trms\<^sub>s\<^sub>t ` \') \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F = \(trms\<^sub>s\<^sub>t ` \)" using to_st_append trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq[OF assms(1)] assms(2,3) by auto thus ?thesis by (simp add: Un_commute sup_left_commute) qed private lemma ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset: assumes "x#S \ \" shows "\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` (update\<^sub>s\<^sub>t \ (x#S))) \ \(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)" (is ?A) "\(assignment_rhs\<^sub>s\<^sub>t ` (update\<^sub>s\<^sub>t \ (x#S))) \ \(assignment_rhs\<^sub>s\<^sub>t ` \)" (is ?B) proof - { fix t assume "t \ \(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` (update\<^sub>s\<^sub>t \ (x#S)))" then obtain S' where S': "S' \ update\<^sub>s\<^sub>t \ (x#S)" "t \ ik\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t S')" by auto have *: "ik\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t S) \ ik\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t (x#S))" using ik_append[of "dual\<^sub>s\<^sub>t [x]" "dual\<^sub>s\<^sub>t S"] dual\<^sub>s\<^sub>t_append[of "[x]" S] by auto hence "t \ \(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)" proof (cases "S' = S") case True thus ?thesis using * assms S' by auto next case False thus ?thesis using S' by auto qed } moreover { fix t assume "t \ \(assignment_rhs\<^sub>s\<^sub>t ` (update\<^sub>s\<^sub>t \ (x#S)))" then obtain S' where S': "S' \ update\<^sub>s\<^sub>t \ (x#S)" "t \ assignment_rhs\<^sub>s\<^sub>t S'" by auto have "assignment_rhs\<^sub>s\<^sub>t S \ assignment_rhs\<^sub>s\<^sub>t (x#S)" using assignment_rhs_append[of "[x]" S] by simp hence "t \ \(assignment_rhs\<^sub>s\<^sub>t ` \)" using assms S' by (cases "S' = S") auto } ultimately show ?A ?B by (metis subsetI)+ qed private lemma ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_snd: assumes "send\ts\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (send\ts\\<^sub>s\<^sub>t#S)" "\' = \@[Step (receive\ts\\<^sub>s\<^sub>t)]" shows "(\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \') \ (\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" (is ?A) "(\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \') \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" (is ?B) proof - { fix t' assume t'_in: "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \')" hence "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \) \ set ts" using assms ik\<^sub>e\<^sub>s\<^sub>t_append by auto moreover have "set ts \ \(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)" using assms(1) by force ultimately have "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) by auto } moreover { fix t' assume t'_in: "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \')" hence "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" using assms assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append by auto hence "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) by auto } ultimately show ?A ?B by (metis subsetI)+ qed private lemma ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_rcv: assumes "receive\t\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (receive\t\\<^sub>s\<^sub>t#S)" "\' = \@[Step (send\t\\<^sub>s\<^sub>t)]" shows "(\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \') \ (\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" (is ?A) "(\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \') \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" (is ?B) proof - { fix t' assume t'_in: "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \')" hence "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" using assms ik\<^sub>e\<^sub>s\<^sub>t_append by auto hence "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) by auto } moreover { fix t' assume t'_in: "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \')" hence "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" using assms assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append by auto hence "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) by auto } ultimately show ?A ?B by (metis subsetI)+ qed private lemma ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_eq: assumes "\a: t \ t'\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (\a: t \ t'\\<^sub>s\<^sub>t#S)" "\' = \@[Step (\a: t \ t'\\<^sub>s\<^sub>t)]" shows "(\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \') \ (\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" (is ?A) "(\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \') \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" (is ?B) proof - have 1: "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" when "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \')" for t' proof - have "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" using that assms ik\<^sub>e\<^sub>s\<^sub>t_append by auto thus ?thesis using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) by auto qed have 2: "t'' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" when "t'' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \')" "a = Assign" for t'' proof - have "t'' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \) \ {t'}" using that assms assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append by auto moreover have "t' \ \(assignment_rhs\<^sub>s\<^sub>t ` \)" using assms(1) that by force ultimately show ?thesis using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) that by auto qed have 3: "assignment_rhs\<^sub>e\<^sub>s\<^sub>t \' = assignment_rhs\<^sub>e\<^sub>s\<^sub>t \" (is ?C) "(\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (\(assignment_rhs\<^sub>s\<^sub>t ` \))" (is ?D) when "a = Check" proof - show ?C using that assms(2,3) by (simp add: assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append) show ?D using assms(1,2,3) ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset(2) by auto qed show ?A using 1 2 by (metis subsetI) show ?B using 1 2 3 by (cases a) blast+ qed private lemma ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_ineq: assumes "\X\\\: F\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (\X\\\: F\\<^sub>s\<^sub>t#S)" "\' = \@[Step (\X\\\: F\\<^sub>s\<^sub>t)]" shows "(\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \') \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" (is ?A) "(\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \') \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" (is ?B) proof - { fix t' assume t'_in: "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \')" hence "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" using assms ik\<^sub>e\<^sub>s\<^sub>t_append by auto hence "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) by auto } moreover { fix t' assume t'_in: "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \')" hence "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" using assms assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append by auto hence "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) by auto } ultimately show ?A ?B by (metis subsetI)+ qed subsubsection \Transition Systems Definitions\ inductive pts_symbolic:: "(('fun,'var) strands \ ('fun,'var) strand) \ (('fun,'var) strands \ ('fun,'var) strand) \ bool" (infix "\\<^sup>\" 50) where Nil[simp]: "[] \ \ \ (\,\) \\<^sup>\ (update\<^sub>s\<^sub>t \ [],\)" | Send[simp]: "send\t\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\ (update\<^sub>s\<^sub>t \ (send\t\\<^sub>s\<^sub>t#S),\@[receive\t\\<^sub>s\<^sub>t])" | Receive[simp]: "receive\t\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\ (update\<^sub>s\<^sub>t \ (receive\t\\<^sub>s\<^sub>t#S),\@[send\t\\<^sub>s\<^sub>t])" | Equality[simp]: "\a: t \ t'\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\ (update\<^sub>s\<^sub>t \ (\a: t \ t'\\<^sub>s\<^sub>t#S),\@[\a: t \ t'\\<^sub>s\<^sub>t])" | Inequality[simp]: "\X\\\: F\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\ (update\<^sub>s\<^sub>t \ (\X\\\: F\\<^sub>s\<^sub>t#S),\@[\X\\\: F\\<^sub>s\<^sub>t])" private inductive pts_symbolic_c:: "(('fun,'var) strands \ ('fun,'var) extstrand) \ (('fun,'var) strands \ ('fun,'var) extstrand) \ bool" (infix "\\<^sup>\\<^sub>c" 50) where Nil[simp]: "[] \ \ \ (\,\) \\<^sup>\\<^sub>c (update\<^sub>s\<^sub>t \ [],\)" | Send[simp]: "send\t\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\\<^sub>c (update\<^sub>s\<^sub>t \ (send\t\\<^sub>s\<^sub>t#S),\@[Step (receive\t\\<^sub>s\<^sub>t)])" | Receive[simp]: "receive\t\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\\<^sub>c (update\<^sub>s\<^sub>t \ (receive\t\\<^sub>s\<^sub>t#S),\@[Step (send\t\\<^sub>s\<^sub>t)])" | Equality[simp]: "\a: t \ t'\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\\<^sub>c (update\<^sub>s\<^sub>t \ (\a: t \ t'\\<^sub>s\<^sub>t#S),\@[Step (\a: t \ t'\\<^sub>s\<^sub>t)])" | Inequality[simp]: "\X\\\: F\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\\<^sub>c (update\<^sub>s\<^sub>t \ (\X\\\: F\\<^sub>s\<^sub>t#S),\@[Step (\X\\\: F\\<^sub>s\<^sub>t)])" | Decompose[simp]: "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \ \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \) \ (\,\) \\<^sup>\\<^sub>c (\,\@[Decomp (Fun f T)])" abbreviation pts_symbolic_rtrancl (infix "\\<^sup>\\<^sup>*" 50) where "a \\<^sup>\\<^sup>* b \ pts_symbolic\<^sup>*\<^sup>* a b" private abbreviation pts_symbolic_c_rtrancl (infix "\\<^sup>\\<^sub>c\<^sup>*" 50) where "a \\<^sup>\\<^sub>c\<^sup>* b \ pts_symbolic_c\<^sup>*\<^sup>* a b" lemma pts_symbolic_induct[consumes 1, case_names Nil Send Receive Equality Inequality]: assumes "(\,\) \\<^sup>\ (\',\')" and "\[] \ \; \' = update\<^sub>s\<^sub>t \ []; \' = \\ \ P" and "\t S. \send\t\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (send\t\\<^sub>s\<^sub>t#S); \' = \@[receive\t\\<^sub>s\<^sub>t]\ \ P" and "\t S. \receive\t\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (receive\t\\<^sub>s\<^sub>t#S); \' = \@[send\t\\<^sub>s\<^sub>t]\ \ P" and "\a t t' S. \\a: t \ t'\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (\a: t \ t'\\<^sub>s\<^sub>t#S); \' = \@[\a: t \ t'\\<^sub>s\<^sub>t]\ \ P" and "\X F S. \\X\\\: F\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (\X\\\: F\\<^sub>s\<^sub>t#S); \' = \@[\X\\\: F\\<^sub>s\<^sub>t]\ \ P" shows "P" apply (rule pts_symbolic.cases[OF assms(1)]) using assms(2,3,4,5,6) by simp_all private lemma pts_symbolic_c_induct[consumes 1, case_names Nil Send Receive Equality Inequality Decompose]: assumes "(\,\) \\<^sup>\\<^sub>c (\',\')" and "\[] \ \; \' = update\<^sub>s\<^sub>t \ []; \' = \\ \ P" and "\t S. \send\t\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (send\t\\<^sub>s\<^sub>t#S); \' = \@[Step (receive\t\\<^sub>s\<^sub>t)]\ \ P" and "\t S. \receive\t\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (receive\t\\<^sub>s\<^sub>t#S); \' = \@[Step (send\t\\<^sub>s\<^sub>t)]\ \ P" and "\a t t' S. \\a: t \ t'\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (\a: t \ t'\\<^sub>s\<^sub>t#S); \' = \@[Step (\a: t \ t'\\<^sub>s\<^sub>t)]\ \ P" and "\X F S. \\X\\\: F\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (\X\\\: F\\<^sub>s\<^sub>t#S); \' = \@[Step (\X\\\: F\\<^sub>s\<^sub>t)]\ \ P" and "\f T. \Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \ \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \); \' = \; \' = \@[Decomp (Fun f T)]\ \ P" shows "P" apply (rule pts_symbolic_c.cases[OF assms(1)]) using assms(2,3,4,5,6,7) by simp_all private lemma pts_symbolic_c_preserves_wf_prot: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" "wf\<^sub>s\<^sub>t\<^sub>s' \ \" shows "wf\<^sub>s\<^sub>t\<^sub>s' \' \'" using assms proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) from step.hyps(2) step.IH[OF step.prems] show ?case proof (induction rule: pts_symbolic_c_induct) case Decompose hence "fv\<^sub>e\<^sub>s\<^sub>t \2 = fv\<^sub>e\<^sub>s\<^sub>t \1" "bvars\<^sub>e\<^sub>s\<^sub>t \2 = bvars\<^sub>e\<^sub>s\<^sub>t \1" using bvars_decomp ik_assignment_rhs_decomp_fv by metis+ thus ?case using Decompose unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by (metis wf_vars_mono wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_split(2)) qed (metis wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_nil, metis wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_snd, metis wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_rcv, metis wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_eq, metis wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_ineq) qed metis private lemma pts_symbolic_c_preserves_wf_is: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" "wf\<^sub>s\<^sub>t\<^sub>s' \ \" "wf\<^sub>s\<^sub>t V (to_st \)" shows "wf\<^sub>s\<^sub>t V (to_st \')" using assms proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) hence "(\, \) \\<^sup>\\<^sub>c\<^sup>* (\2, \2)" by auto hence *: "wf\<^sub>s\<^sub>t\<^sub>s' \1 \1" "wf\<^sub>s\<^sub>t\<^sub>s' \2 \2" using pts_symbolic_c_preserves_wf_prot[OF _ step.prems(1)] step.hyps(1) by auto from step.hyps(2) step.IH[OF step.prems] show ?case proof (induction rule: pts_symbolic_c_induct) case Nil thus ?case by auto next case (Send ts S) hence "wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \1) (receive\ts\\<^sub>s\<^sub>t#(dual\<^sub>s\<^sub>t S))" using *(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fastforce hence "fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ wfrestrictedvars\<^sub>s\<^sub>t (to_st \1) \ V" using wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_eq_wfrestrictedvars\<^sub>s\<^sub>t by auto thus ?case using Send wf_rcv_append''' to_st_append by simp next case (Receive ts) thus ?case using wf_snd_append to_st_append by simp next case (Equality a t t' S) hence "wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \1) (\a: t \ t'\\<^sub>s\<^sub>t#(dual\<^sub>s\<^sub>t S))" using *(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fastforce hence "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t (to_st \1) \ V" when "a = Assign" using wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_eq_wfrestrictedvars\<^sub>s\<^sub>t that by auto thus ?case using Equality wf_eq_append''' to_st_append by (cases a) auto next case (Inequality t t' S) thus ?case using wf_ineq_append'' to_st_append by simp next case (Decompose f T) hence "fv (Fun f T) \ wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \1" by (metis fv_subterms_set fv_subset subset_trans ik\<^sub>s\<^sub>t_assignment_rhs\<^sub>s\<^sub>t_wfrestrictedvars_subset) hence "vars\<^sub>s\<^sub>t (decomp (Fun f T)) \ wfrestrictedvars\<^sub>s\<^sub>t (to_st \1) \ V" using decomp_vars[of "Fun f T"] wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_eq_wfrestrictedvars\<^sub>s\<^sub>t[of \1] by auto thus ?case using to_st_append[of \1 "[Decomp (Fun f T)]"] wf_append_suffix[OF Decompose.prems] Decompose.hyps(3) by (metis append_Nil2 decomp_vars(1,2) to_st.simps(1,3)) qed qed metis private lemma pts_symbolic_c_preserves_tfr\<^sub>s\<^sub>e\<^sub>t: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" and "tfr\<^sub>s\<^sub>e\<^sub>t ((\(trms\<^sub>s\<^sub>t ` \)) \ (trms\<^sub>e\<^sub>s\<^sub>t \))" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s ((\(trms\<^sub>s\<^sub>t ` \)) \ (trms\<^sub>e\<^sub>s\<^sub>t \))" shows "tfr\<^sub>s\<^sub>e\<^sub>t ((\(trms\<^sub>s\<^sub>t ` \')) \ (trms\<^sub>e\<^sub>s\<^sub>t \')) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s ((\(trms\<^sub>s\<^sub>t ` \')) \ (trms\<^sub>e\<^sub>s\<^sub>t \'))" using assms proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) from step.hyps(2) step.IH[OF step.prems] show ?case proof (induction rule: pts_symbolic_c_induct) case Nil hence "\(trms\<^sub>s\<^sub>t ` \1) = \(trms\<^sub>s\<^sub>t ` \2)" by force thus ?case using Nil by metis next case (Decompose f T) obtain t where t: "t \ ik\<^sub>e\<^sub>s\<^sub>t \1 \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1" "Fun f T \ t" using Decompose.hyps(1) by auto have t_wf: "wf\<^sub>t\<^sub>r\<^sub>m t" using Decompose.prems wf_trm_subterm[of _ t] trms\<^sub>e\<^sub>s\<^sub>t_ik_assignment_rhsI[OF t(1)] unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by (metis UN_E Un_iff) have "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>e\<^sub>s\<^sub>t \1)" using trms\<^sub>e\<^sub>s\<^sub>t_ik_assignment_rhsI t by auto hence "Fun f T \ SMP (trms\<^sub>e\<^sub>s\<^sub>t \1)" by (metis (no_types) SMP.MP SMP.Subterm UN_E t(2)) hence "{Fun f T} \ SMP (trms\<^sub>e\<^sub>s\<^sub>t \1)" using SMP.Subterm[of "Fun f T"] by auto moreover have "trms\<^sub>e\<^sub>s\<^sub>t \2 = insert (Fun f T) (trms\<^sub>e\<^sub>s\<^sub>t \1)" using Decompose.hyps(3) by auto ultimately have *: "SMP (trms\<^sub>e\<^sub>s\<^sub>t \1) = SMP (trms\<^sub>e\<^sub>s\<^sub>t \2)" using SMP_subset_union_eq[of "{Fun f T}"] by (simp add: Un_commute) hence "SMP ((\(trms\<^sub>s\<^sub>t ` \1)) \ (trms\<^sub>e\<^sub>s\<^sub>t \1)) = SMP ((\(trms\<^sub>s\<^sub>t ` \2)) \ (trms\<^sub>e\<^sub>s\<^sub>t \2))" using Decompose.hyps(2) SMP_union by auto moreover have "\t \ trms\<^sub>e\<^sub>s\<^sub>t \1. wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" using Decompose.prems wf_trm_subterm t(2) t_wf unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by auto hence "\t \ trms\<^sub>e\<^sub>s\<^sub>t \2. wf\<^sub>t\<^sub>r\<^sub>m t" by (metis * SMP.MP SMP_wf_trm) hence "\t \ (\(trms\<^sub>s\<^sub>t ` \2)) \ (trms\<^sub>e\<^sub>s\<^sub>t \2). wf\<^sub>t\<^sub>r\<^sub>m t" using Decompose.prems Decompose.hyps(2) unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by force ultimately show ?thesis using Decompose.prems unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by presburger qed (metis trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_snd, metis trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_rcv, metis trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_eq, metis trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_ineq) qed metis private lemma pts_symbolic_c_preserves_tfr\<^sub>s\<^sub>t\<^sub>p: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" "\S \ \ \ {to_st \}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" shows "\S \ \' \ {to_st \'}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using assms proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) from step.hyps(2) step.IH[OF step.prems] show ?case proof (induction rule: pts_symbolic_c_induct) case Nil have 1: "\S \ {to_st \2}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Nil by simp have 2: "\2 = \1 - {[]}" "\S \ \1. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Nil by simp_all have "\S \ \2. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" proof fix S assume "S \ \2" hence "S \ \1" using 2(1) by simp thus "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using 2(2) by simp qed thus ?case using 1 by auto next case (Send t S) have 1: "\S \ {to_st \2}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Send by (simp add: to_st_append) have 2: "\2 = insert S (\1 - {send\t\\<^sub>s\<^sub>t#S})" "\S \ \1. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Send by simp_all have 3: "\S \ \2. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" proof fix S' assume "S' \ \2" hence "S' \ \1 \ S' = S" using 2(1) by auto moreover have "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Send.hyps 2(2) by auto ultimately show "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" using 2(2) by blast qed thus ?case using 1 by auto next case (Receive t S) have 1: "\S \ {to_st \2}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Receive by (simp add: to_st_append) have 2: "\2 = insert S (\1 - {receive\t\\<^sub>s\<^sub>t#S})" "\S \ \1. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Receive by simp_all have 3: "\S \ \2. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" proof fix S' assume "S' \ \2" hence "S' \ \1 \ S' = S" using 2(1) by auto moreover have "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Receive.hyps 2(2) by auto ultimately show "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" using 2(2) by blast qed show ?case using 1 3 by auto next case (Equality a t t' S) have 1: "to_st \2 = to_st \1@[\a: t \ t'\\<^sub>s\<^sub>t]" "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \1)" using Equality by (simp_all add: to_st_append) have 2: "list_all tfr\<^sub>s\<^sub>t\<^sub>p [\a: t \ t'\\<^sub>s\<^sub>t]" using Equality by fastforce have 3: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \2)" using tfr_stp_all_append[of "to_st \1" "[\a: t \ t'\\<^sub>s\<^sub>t]"] 1 2 by metis hence 4: "\S \ {to_st \2}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Equality by simp have 5: "\2 = insert S (\1 - {\a: t \ t'\\<^sub>s\<^sub>t#S})" "\S \ \1. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Equality by simp_all have 6: "\S \ \2. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" proof fix S' assume "S' \ \2" hence "S' \ \1 \ S' = S" using 5(1) by auto moreover have "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Equality.hyps 5(2) by auto ultimately show "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" using 5(2) by blast qed thus ?case using 4 by auto next case (Inequality X F S) have 1: "to_st \2 = to_st \1@[\X\\\: F\\<^sub>s\<^sub>t]" "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \1)" using Inequality by (simp_all add: to_st_append) have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (\X\\\: F\\<^sub>s\<^sub>t#S)" using Inequality(1,4) by blast hence 2: "list_all tfr\<^sub>s\<^sub>t\<^sub>p [\X\\\: F\\<^sub>s\<^sub>t]" by simp have 3: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \2)" using tfr_stp_all_append[of "to_st \1" "[\X\\\: F\\<^sub>s\<^sub>t]"] 1 2 by metis hence 4: "\S \ {to_st \2}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Inequality by simp have 5: "\2 = insert S (\1 - {\X\\\: F\\<^sub>s\<^sub>t#S})" "\S \ \1. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Inequality by simp_all have 6: "\S \ \2. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" proof fix S' assume "S' \ \2" hence "S' \ \1 \ S' = S" using 5(1) by auto moreover have "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Inequality.hyps 5(2) by auto ultimately show "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" using 5(2) by blast qed thus ?case using 4 by auto next case (Decompose f T) hence 1: "\S \ \2. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" by blast have 2: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \1)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st [Decomp (Fun f T)])" using Decompose.prems decomp_tfr\<^sub>s\<^sub>t\<^sub>p by auto hence "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \1@to_st [Decomp (Fun f T)])" by auto hence "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \2)" using Decompose.hyps(3) to_st_append[of \1 "[Decomp (Fun f T)]"] by auto thus ?case using 1 by blast qed qed private lemma pts_symbolic_c_preserves_well_analyzed: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" "well_analyzed \" shows "well_analyzed \'" using assms proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) from step.hyps(2) step.IH[OF step.prems] show ?case proof (induction rule: pts_symbolic_c_induct) case Receive thus ?case by (metis well_analyzed_singleton(1) well_analyzed_append) next case Send thus ?case by (metis well_analyzed_singleton(2) well_analyzed_append) next case Equality thus ?case by (metis well_analyzed_singleton(3) well_analyzed_append) next case Inequality thus ?case by (metis well_analyzed_singleton(4) well_analyzed_append) next case (Decompose f T) hence "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \1 \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1) - (Var`\)" by auto thus ?case by (metis well_analyzed.Decomp Decompose.prems Decompose.hyps(3)) qed simp qed metis private lemma pts_symbolic_c_preserves_Ana_invar_subst: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" and "Ana_invar_subst ( (\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \) \ (ik\<^sub>e\<^sub>s\<^sub>t \)) \ (\(assignment_rhs\<^sub>s\<^sub>t ` \) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)))" shows "Ana_invar_subst ( (\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \') \ (ik\<^sub>e\<^sub>s\<^sub>t \')) \ (\(assignment_rhs\<^sub>s\<^sub>t ` \') \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \')))" using assms proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) from step.hyps(2) step.IH[OF step.prems] show ?case proof (induction rule: pts_symbolic_c_induct) case Nil hence "\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \1) = \(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \2)" "\(assignment_rhs\<^sub>s\<^sub>t ` \1) = \(assignment_rhs\<^sub>s\<^sub>t ` \2)" by force+ thus ?case using Nil by metis next case Send show ?case using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_snd[OF Send.hyps] Ana_invar_subst_subset[OF Send.prems] by (metis Un_mono) next case Receive show ?case using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_rcv[OF Receive.hyps] Ana_invar_subst_subset[OF Receive.prems] by (metis Un_mono) next case Equality show ?case using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_eq[OF Equality.hyps] Ana_invar_subst_subset[OF Equality.prems] by (metis Un_mono) next case Inequality show ?case using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_ineq[OF Inequality.hyps] Ana_invar_subst_subset[OF Inequality.prems] by (metis Un_mono) next case (Decompose f T) let ?X = "\(assignment_rhs\<^sub>s\<^sub>t`\2) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \2" let ?Y = "\(assignment_rhs\<^sub>s\<^sub>t`\1) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1" obtain K M where Ana: "Ana (Fun f T) = (K,M)" by moura hence *: "ik\<^sub>e\<^sub>s\<^sub>t \2 = ik\<^sub>e\<^sub>s\<^sub>t \1 \ set M" "assignment_rhs\<^sub>e\<^sub>s\<^sub>t \2 = assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1" using ik\<^sub>e\<^sub>s\<^sub>t_append assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append decomp_ik decomp_assignment_rhs_empty Decompose.hyps(3) by auto { fix g S assume "Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t`\2) \ ik\<^sub>e\<^sub>s\<^sub>t \2 \ ?X)" hence "Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \1) \ ik\<^sub>e\<^sub>s\<^sub>t \1 \ set M \ ?X)" using * Decompose.hyps(2) by auto hence "Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \1)) \ Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \1) \ Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (set M) \ Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (\(assignment_rhs\<^sub>s\<^sub>t`\1)) \ Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1)" using Decompose * Ana_fun_subterm[OF Ana] by auto moreover have "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \1 \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1)" using trms\<^sub>e\<^sub>s\<^sub>t_ik_subtermsI Decompose.hyps(1) by auto hence "subterms (Fun f T) \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \1 \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1)" by (metis in_subterms_subset_Union) hence "subterms\<^sub>s\<^sub>e\<^sub>t (set M) \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \1 \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1)" by (meson Un_upper2 Ana_subterm[OF Ana] subterms_subset_set psubsetE subset_trans) ultimately have "Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \1) \ ik\<^sub>e\<^sub>s\<^sub>t \1 \ ?Y)" by auto } thus ?case using Decompose unfolding Ana_invar_subst_def by metis qed qed private lemma pts_symbolic_c_preserves_constr_disj_vars: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" "wf\<^sub>s\<^sub>t\<^sub>s' \ \" "fv\<^sub>e\<^sub>s\<^sub>t \ \ bvars\<^sub>e\<^sub>s\<^sub>t \ = {}" shows "fv\<^sub>e\<^sub>s\<^sub>t \' \ bvars\<^sub>e\<^sub>s\<^sub>t \' = {}" using assms proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) have *: "\S. S \ \1 \ fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t \1 = {}" "\S. S \ \1 \ fv\<^sub>e\<^sub>s\<^sub>t \1 \ bvars\<^sub>s\<^sub>t S = {}" using pts_symbolic_c_preserves_wf_prot[OF step.hyps(1) step.prems(1)] unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by auto from step.hyps(2) step.IH[OF step.prems] show ?case proof (induction rule: pts_symbolic_c_induct) case Nil thus ?case by auto next case (Send ts S) hence "fv\<^sub>e\<^sub>s\<^sub>t \2 = fv\<^sub>e\<^sub>s\<^sub>t \1 \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "bvars\<^sub>e\<^sub>s\<^sub>t \2 = bvars\<^sub>e\<^sub>s\<^sub>t \1" "fv\<^sub>s\<^sub>t (send\ts\\<^sub>s\<^sub>t#S) = fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ fv\<^sub>s\<^sub>t S" using fv\<^sub>e\<^sub>s\<^sub>t_append bvars\<^sub>e\<^sub>s\<^sub>t_append by simp+ thus ?case using *(1)[OF Send(1)] Send(4) by auto next case (Receive ts S) hence "fv\<^sub>e\<^sub>s\<^sub>t \2 = fv\<^sub>e\<^sub>s\<^sub>t \1 \ fv\<^sub>s\<^sub>e\<^sub>t (set ts)" "bvars\<^sub>e\<^sub>s\<^sub>t \2 = bvars\<^sub>e\<^sub>s\<^sub>t \1" "fv\<^sub>s\<^sub>t (receive\ts\\<^sub>s\<^sub>t#S) = fv\<^sub>s\<^sub>e\<^sub>t (set ts) \ fv\<^sub>s\<^sub>t S" using fv\<^sub>e\<^sub>s\<^sub>t_append bvars\<^sub>e\<^sub>s\<^sub>t_append by simp+ thus ?case using *(1)[OF Receive(1)] Receive(4) by auto next case (Equality a t t' S) hence "fv\<^sub>e\<^sub>s\<^sub>t \2 = fv\<^sub>e\<^sub>s\<^sub>t \1 \ fv t \ fv t'" "bvars\<^sub>e\<^sub>s\<^sub>t \2 = bvars\<^sub>e\<^sub>s\<^sub>t \1" "fv\<^sub>s\<^sub>t (\a: t \ t'\\<^sub>s\<^sub>t#S) = fv t \ fv t' \ fv\<^sub>s\<^sub>t S" using fv\<^sub>e\<^sub>s\<^sub>t_append bvars\<^sub>e\<^sub>s\<^sub>t_append by fastforce+ thus ?case using *(1)[OF Equality(1)] Equality(4) by auto next case (Inequality X F S) hence "fv\<^sub>e\<^sub>s\<^sub>t \2 = fv\<^sub>e\<^sub>s\<^sub>t \1 \ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X)" "bvars\<^sub>e\<^sub>s\<^sub>t \2 = bvars\<^sub>e\<^sub>s\<^sub>t \1 \ set X" "fv\<^sub>s\<^sub>t (\X\\\: F\\<^sub>s\<^sub>t#S) = (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X) \ fv\<^sub>s\<^sub>t S" using fv\<^sub>e\<^sub>s\<^sub>t_append bvars\<^sub>e\<^sub>s\<^sub>t_append strand_vars_split(3)[of "[\X\\\: F\\<^sub>s\<^sub>t]" S] by auto+ moreover have "fv\<^sub>e\<^sub>s\<^sub>t \1 \ set X = {}" using *(2)[OF Inequality(1)] by auto ultimately show ?case using *(1)[OF Inequality(1)] Inequality(4) by auto next case (Decompose f T) thus ?case using Decompose(3,4) bvars_decomp ik_assignment_rhs_decomp_fv[OF Decompose(1)] by auto qed qed subsubsection \Theorem: The Typing Result Lifted to the Transition System Level\ private lemma wf\<^sub>s\<^sub>t\<^sub>s'_decomp_rm: assumes "well_analyzed A" "wf\<^sub>s\<^sub>t\<^sub>s' S (decomp_rm\<^sub>e\<^sub>s\<^sub>t A)" shows "wf\<^sub>s\<^sub>t\<^sub>s' S A" unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def proof (intro conjI) show "\S\S. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t A) (dual\<^sub>s\<^sub>t S)" by (metis (no_types) assms(2) wf\<^sub>s\<^sub>t\<^sub>s'_def wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_decomp_rm\<^sub>e\<^sub>s\<^sub>t_subset wf_vars_mono le_iff_sup) show "\Sa\S. \S'\S. fv\<^sub>s\<^sub>t Sa \ bvars\<^sub>s\<^sub>t S' = {}" by (metis assms(2) wf\<^sub>s\<^sub>t\<^sub>s'_def) show "\S\S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t A = {}" by (metis assms(2) wf\<^sub>s\<^sub>t\<^sub>s'_def bvars_decomp_rm) show "\S\S. fv\<^sub>e\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>t S = {}" by (metis assms wf\<^sub>s\<^sub>t\<^sub>s'_def well_analyzed_decomp_rm\<^sub>e\<^sub>s\<^sub>t_fv) qed private lemma decomps\<^sub>e\<^sub>s\<^sub>t_pts_symbolic_c: assumes "D \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \" shows "(S,A) \\<^sup>\\<^sub>c\<^sup>* (S,A@D)" using assms(1) proof (induction D rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case (Decomp B f X K T) have "subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t (A@B) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@B))" using ik\<^sub>e\<^sub>s\<^sub>t_append[of A B] assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append[of A B] by auto hence "Fun f X \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t (A@B) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@B))" using Decomp.hyps by auto hence "(S,A@B) \\<^sup>\\<^sub>c (S,A@B@[Decomp (Fun f X)])" using pts_symbolic_c.Decompose[of f X "A@B"] by simp thus ?case using Decomp.IH rtrancl_into_rtrancl rtranclp_rtrancl_eq[of pts_symbolic_c "(S,A)" "(S,A@B)"] by auto qed simp private lemma pts_symbolic_to_pts_symbolic_c: assumes "(\,to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \\<^sub>d)) \\<^sup>\\<^sup>* (\',\')" "sem\<^sub>e\<^sub>s\<^sub>t_d {} \ (to_est \')" "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ \\<^sub>d" and wf: "wf\<^sub>s\<^sub>t\<^sub>s' \ (decomp_rm\<^sub>e\<^sub>s\<^sub>t \\<^sub>d)" "wf\<^sub>e\<^sub>s\<^sub>t {} \\<^sub>d" and tar: "Ana_invar_subst ((\(ik\<^sub>s\<^sub>t` dual\<^sub>s\<^sub>t` \) \ (ik\<^sub>e\<^sub>s\<^sub>t \\<^sub>d)) \ (\(assignment_rhs\<^sub>s\<^sub>t` \) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \\<^sub>d)))" and wa: "well_analyzed \\<^sub>d" and \: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\\\<^sub>d'. \' = to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \\<^sub>d') \ (\,\\<^sub>d) \\<^sup>\\<^sub>c\<^sup>* (\',\\<^sub>d') \ sem\<^sub>e\<^sub>s\<^sub>t_c {} \ \\<^sub>d'" using assms(1,2) proof (induction rule: rtranclp_induct2) case refl thus ?case using assms by auto next case (step \1 \1 \2 \2) have "sem\<^sub>e\<^sub>s\<^sub>t_d {} \ (to_est \1)" using step.hyps(2) step.prems by (induct rule: pts_symbolic_induct, metis, (metis sem\<^sub>e\<^sub>s\<^sub>t_d_split_left to_est_append)+) then obtain \1d where \1d: "\1 = to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \1d)" "(\, \\<^sub>d) \\<^sup>\\<^sub>c\<^sup>* (\1, \1d)" "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ \1d" using step.IH by moura show ?case using step.hyps(2) proof (induction rule: pts_symbolic_induct) case Nil hence "(\, \\<^sub>d) \\<^sup>\\<^sub>c\<^sup>* (\2, \1d)" using \1d pts_symbolic_c.Nil[OF Nil.hyps(1), of \1d] by simp thus ?case using \1d Nil by auto next case (Send t S) hence "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ (\1d@[Step (receive\t\\<^sub>s\<^sub>t)])" using sem\<^sub>e\<^sub>s\<^sub>t_c.Receive[OF \1d(3)] by simp moreover have "(\1, \1d) \\<^sup>\\<^sub>c (\2, \1d@[Step (receive\t\\<^sub>s\<^sub>t)])" using Send.hyps(2) pts_symbolic_c.Send[OF Send.hyps(1), of \1d] by simp moreover have "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t (\1d@[Step (receive\t\\<^sub>s\<^sub>t)])) = \2" using Send.hyps(3) decomp_rm\<^sub>e\<^sub>s\<^sub>t_append \1d(1) by (simp add: to_st_append) ultimately show ?case using \1d(2) by auto next case (Equality a t t' S) hence "t \ \ = t' \ \" using step.prems sem\<^sub>e\<^sub>s\<^sub>t_d_eq_sem_st[of "{}" \ "to_est \2"] to_st_append to_est_append to_st_to_est_inv by auto hence "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ (\1d@[Step (\a: t \ t'\\<^sub>s\<^sub>t)])" using sem\<^sub>e\<^sub>s\<^sub>t_c.Equality[OF \1d(3)] by simp moreover have "(\1, \1d) \\<^sup>\\<^sub>c (\2, \1d@[Step (\a: t \ t'\\<^sub>s\<^sub>t)])" using Equality.hyps(2) pts_symbolic_c.Equality[OF Equality.hyps(1), of \1d] by simp moreover have "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t (\1d@[Step (\a: t \ t'\\<^sub>s\<^sub>t)])) = \2" using Equality.hyps(3) decomp_rm\<^sub>e\<^sub>s\<^sub>t_append \1d(1) by (simp add: to_st_append) ultimately show ?case using \1d(2) by auto next case (Inequality X F S) hence "ineq_model \ X F" using step.prems sem\<^sub>e\<^sub>s\<^sub>t_d_eq_sem_st[of "{}" \ "to_est \2"] to_st_append to_est_append to_st_to_est_inv by auto hence "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ (\1d@[Step (\X\\\: F\\<^sub>s\<^sub>t)])" using sem\<^sub>e\<^sub>s\<^sub>t_c.Inequality[OF \1d(3)] by simp moreover have "(\1, \1d) \\<^sup>\\<^sub>c (\2, \1d@[Step (\X\\\: F\\<^sub>s\<^sub>t)])" using Inequality.hyps(2) pts_symbolic_c.Inequality[OF Inequality.hyps(1), of \1d] by simp moreover have "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t (\1d@[Step (\X\\\: F\\<^sub>s\<^sub>t)])) = \2" using Inequality.hyps(3) decomp_rm\<^sub>e\<^sub>s\<^sub>t_append \1d(1) by (simp add: to_st_append) ultimately show ?case using \1d(2) by auto next case (Receive ts S) hence "\t \ set ts. ik\<^sub>s\<^sub>t \1 \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" using step.prems sem\<^sub>e\<^sub>s\<^sub>t_d_eq_sem_st[of "{}" \ "to_est \2"] strand_sem_split(4)[of "{}" \1 "[send\ts\\<^sub>s\<^sub>t]" \] to_st_append to_est_append to_st_to_est_inv by auto moreover have "ik\<^sub>s\<^sub>t \1 \\<^sub>s\<^sub>e\<^sub>t \ \ ik\<^sub>e\<^sub>s\<^sub>t \1d \\<^sub>s\<^sub>e\<^sub>t \" using \1d(1) decomp_rm\<^sub>e\<^sub>s\<^sub>t_ik_subset by auto ultimately have *: "\t \ set ts. ik\<^sub>e\<^sub>s\<^sub>t \1d \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" using ideduct_mono by auto have "wf\<^sub>s\<^sub>t\<^sub>s' \ \\<^sub>d" by (rule wf\<^sub>s\<^sub>t\<^sub>s'_decomp_rm[OF wa assms(4)]) hence **: "wf\<^sub>e\<^sub>s\<^sub>t {} \1d" by (rule pts_symbolic_c_preserves_wf_is[OF \1d(2) _ assms(5)]) have "Ana_invar_subst (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t`\1) \ (ik\<^sub>e\<^sub>s\<^sub>t \1d) \ (\(assignment_rhs\<^sub>s\<^sub>t`\1) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1d)))" using tar \1d(2) pts_symbolic_c_preserves_Ana_invar_subst by metis hence "Ana_invar_subst (ik\<^sub>e\<^sub>s\<^sub>t \1d)" "Ana_invar_subst (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1d)" using Ana_invar_subst_subset by blast+ moreover have "well_analyzed \1d" using pts_symbolic_c_preserves_well_analyzed[OF \1d(2) wa] by metis ultimately obtain D where D: "D \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \1d) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1d) \" "\t \ set ts. ik\<^sub>e\<^sub>s\<^sub>t (\1d@D) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" using decomps\<^sub>e\<^sub>s\<^sub>t_exist_subst_list[OF * \1d(3) ** assms(8)] unfolding Ana_invar_subst_def by auto have "(\, \\<^sub>d) \\<^sup>\\<^sub>c\<^sup>* (\1, \1d@D)" using \1d(2) decomps\<^sub>e\<^sub>s\<^sub>t_pts_symbolic_c[OF D(1), of \1] by auto hence "(\, \\<^sub>d) \\<^sup>\\<^sub>c\<^sup>* (\2, \1d@D@[Step (send\ts\\<^sub>s\<^sub>t)])" using Receive(2) pts_symbolic_c.Receive[OF Receive.hyps(1), of "\1d@D"] by auto moreover have "\2 = to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t (\1d@D@[Step (send\ts\\<^sub>s\<^sub>t)]))" using Receive.hyps(3) \1d(1) decomps\<^sub>e\<^sub>s\<^sub>t_decomp_rm\<^sub>e\<^sub>s\<^sub>t_empty[OF D(1)] decomp_rm\<^sub>e\<^sub>s\<^sub>t_append to_st_append by auto moreover have "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ (\1d@D@[Step (send\ts\\<^sub>s\<^sub>t)])" using D(2) sem\<^sub>e\<^sub>s\<^sub>t_c.Send[OF sem\<^sub>e\<^sub>s\<^sub>t_c_decomps\<^sub>e\<^sub>s\<^sub>t_append[OF \1d(3) D(1)]] by simp ultimately show ?case by auto qed qed private lemma pts_symbolic_c_to_pts_symbolic: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ \'" shows "(\,to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \)) \\<^sup>\\<^sup>* (\',to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \'))" "sem\<^sub>e\<^sub>s\<^sub>t_d {} \ (decomp_rm\<^sub>e\<^sub>s\<^sub>t \')" proof - show "(\,to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \)) \\<^sup>\\<^sup>* (\',to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \'))" using assms(1) proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) show ?case using step.hyps(2,1) step.IH proof (induction rule: pts_symbolic_c_induct) case Nil thus ?case using pts_symbolic.Nil[OF Nil.hyps(1), of "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \1)"] by simp next case (Send t S) thus ?case using pts_symbolic.Send[OF Send.hyps(1), of "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \1)"] by (simp add: decomp_rm\<^sub>e\<^sub>s\<^sub>t_append to_st_append) next case (Receive t S) thus ?case using pts_symbolic.Receive[OF Receive.hyps(1), of "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \1)"] by (simp add: decomp_rm\<^sub>e\<^sub>s\<^sub>t_append to_st_append) next case (Equality a t t' S) thus ?case using pts_symbolic.Equality[OF Equality.hyps(1), of "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \1)"] by (simp add: decomp_rm\<^sub>e\<^sub>s\<^sub>t_append to_st_append) next case (Inequality t t' S) thus ?case using pts_symbolic.Inequality[OF Inequality.hyps(1), of "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \1)"] by (simp add: decomp_rm\<^sub>e\<^sub>s\<^sub>t_append to_st_append) next case (Decompose t) thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append by simp qed qed simp qed (rule sem\<^sub>e\<^sub>s\<^sub>t_d_decomp_rm\<^sub>e\<^sub>s\<^sub>t_if_sem\<^sub>e\<^sub>s\<^sub>t_c[OF assms(2)]) private lemma pts_symbolic_to_pts_symbolic_c_from_initial: assumes "(\\<^sub>0,[]) \\<^sup>\\<^sup>* (\,\)" "\ \ \\\" "wf\<^sub>s\<^sub>t\<^sub>s' \\<^sub>0 []" and "Ana_invar_subst (\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \\<^sub>0) \ \(assignment_rhs\<^sub>s\<^sub>t ` \\<^sub>0))" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\\\<^sub>d. \ = to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \\<^sub>d) \ (\\<^sub>0,[]) \\<^sup>\\<^sub>c\<^sup>* (\,\\<^sub>d) \ (\ \\<^sub>c \to_st \\<^sub>d\)" using assms pts_symbolic_to_pts_symbolic_c[of \\<^sub>0 "[]" \ \ \] sem\<^sub>e\<^sub>s\<^sub>t_c_eq_sem_st[of "{}" \] sem\<^sub>e\<^sub>s\<^sub>t_d_eq_sem_st[of "{}" \] to_st_to_est_inv[of \] strand_sem_eq_defs by (auto simp add: constr_sem_c_def constr_sem_d_def simp del: subst_range.simps) private lemma pts_symbolic_c_to_pts_symbolic_from_initial: assumes "(\\<^sub>0,[]) \\<^sup>\\<^sub>c\<^sup>* (\,\)" "\ \\<^sub>c \to_st \\" shows "(\\<^sub>0,[]) \\<^sup>\\<^sup>* (\,to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \))" "\ \ \to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \)\" using assms pts_symbolic_c_to_pts_symbolic[of \\<^sub>0 "[]" \ \ \] sem\<^sub>e\<^sub>s\<^sub>t_c_eq_sem_st[of "{}" \] sem\<^sub>e\<^sub>s\<^sub>t_d_eq_sem_st[of "{}" \] strand_sem_eq_defs by (auto simp add: constr_sem_c_def constr_sem_d_def) private lemma to_st_trms_wf: assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>e\<^sub>s\<^sub>t A)" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t (to_st A))" using assms proof (induction A) case (Cons x A) hence IH: "\t \ trms\<^sub>s\<^sub>t (to_st A). wf\<^sub>t\<^sub>r\<^sub>m t" by auto with Cons show ?case proof (cases x) case (Decomp t) hence "wf\<^sub>t\<^sub>r\<^sub>m t" using Cons.prems by auto obtain K T where Ana_t: "Ana t = (K,T)" by moura hence "trms\<^sub>s\<^sub>t (decomp t) \ {t} \ set K \ set T" using decomp_set_unfold[OF Ana_t] by force moreover have "\t \ set T. wf\<^sub>t\<^sub>r\<^sub>m t" using Ana_subterm[OF Ana_t] \wf\<^sub>t\<^sub>r\<^sub>m t\ wf_trm_subterm by auto ultimately have "\t \ trms\<^sub>s\<^sub>t (decomp t). wf\<^sub>t\<^sub>r\<^sub>m t" using Ana_keys_wf'[OF Ana_t] \wf\<^sub>t\<^sub>r\<^sub>m t\ by auto thus ?thesis using IH Decomp by auto qed auto qed simp private lemma to_st_trms_SMP_subset: "trms\<^sub>s\<^sub>t (to_st A) \ SMP (trms\<^sub>e\<^sub>s\<^sub>t A)" proof fix t assume "t \ trms\<^sub>s\<^sub>t (to_st A)" thus "t \ SMP (trms\<^sub>e\<^sub>s\<^sub>t A)" proof (induction A) case (Cons x A) hence *: "t \ trms\<^sub>s\<^sub>t (to_st [x]) \ trms\<^sub>s\<^sub>t (to_st A)" using to_st_append[of "[x]" A] by auto have **: "trms\<^sub>s\<^sub>t (to_st A) \ trms\<^sub>s\<^sub>t (to_st (x#A))" "trms\<^sub>e\<^sub>s\<^sub>t A \ trms\<^sub>e\<^sub>s\<^sub>t (x#A)" using to_st_append[of "[x]" A] by auto show ?case proof (cases "t \ trms\<^sub>s\<^sub>t (to_st A)") case True thus ?thesis using Cons.IH SMP_mono[OF **(2)] by auto next case False hence ***: "t \ trms\<^sub>s\<^sub>t (to_st [x])" using * by auto thus ?thesis proof (cases x) case (Decomp t') hence ****: "t \ trms\<^sub>s\<^sub>t (decomp t')" "t' \ trms\<^sub>e\<^sub>s\<^sub>t (x#A)" using *** by auto obtain K T where Ana_t': "Ana t' = (K,T)" by moura hence "t \ {t'} \ set K \ set T" using decomp_set_unfold[OF Ana_t'] ****(1) by force moreover { assume "t = t'" hence ?thesis using SMP.MP[OF ****(2)] by simp } moreover { assume "t \ set K" hence ?thesis using SMP.Ana[OF SMP.MP[OF ****(2)] Ana_t'] by auto } moreover { assume "t \ set T" "t \ t'" hence "t \ t'" using Ana_subterm[OF Ana_t'] by blast hence ?thesis using SMP.Subterm[OF SMP.MP[OF ****(2)]] by auto } ultimately show ?thesis using Decomp by auto qed auto qed qed simp qed private lemma to_st_trms_tfr\<^sub>s\<^sub>e\<^sub>t: assumes "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>e\<^sub>s\<^sub>t A)" shows "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t (to_st A))" proof - have *: "trms\<^sub>s\<^sub>t (to_st A) \ SMP (trms\<^sub>e\<^sub>s\<^sub>t A)" using to_st_trms_wf to_st_trms_SMP_subset assms unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by auto have "trms\<^sub>s\<^sub>t (to_st A) = trms\<^sub>s\<^sub>t (to_st A) \ trms\<^sub>e\<^sub>s\<^sub>t A" by (blast dest!: trms\<^sub>e\<^sub>s\<^sub>tD) hence "SMP (trms\<^sub>e\<^sub>s\<^sub>t A) = SMP (trms\<^sub>s\<^sub>t (to_st A))" using SMP_subset_union_eq[OF *] by auto thus ?thesis using * assms unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by presburger qed theorem wt_attack_if_tfr_attack_pts: assumes "wf\<^sub>s\<^sub>t\<^sub>s \\<^sub>0" "tfr\<^sub>s\<^sub>e\<^sub>t (\(trms\<^sub>s\<^sub>t ` \\<^sub>0))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\(trms\<^sub>s\<^sub>t ` \\<^sub>0))" "\S \ \\<^sub>0. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" and "Ana_invar_subst (\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \\<^sub>0) \ \(assignment_rhs\<^sub>s\<^sub>t ` \\<^sub>0))" and "(\\<^sub>0,[]) \\<^sup>\\<^sup>* (\,\)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "\ \ \\, Var\" shows "\\\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ (\\<^sub>\ \ \\, Var\) \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" proof - have "(\(trms\<^sub>s\<^sub>t ` \\<^sub>0)) \ (trms\<^sub>e\<^sub>s\<^sub>t []) = \(trms\<^sub>s\<^sub>t ` \\<^sub>0)" "to_st [] = []" "list_all tfr\<^sub>s\<^sub>t\<^sub>p []" using assms by simp_all hence *: "tfr\<^sub>s\<^sub>e\<^sub>t ((\(trms\<^sub>s\<^sub>t ` \\<^sub>0)) \ (trms\<^sub>e\<^sub>s\<^sub>t []))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s ((\(trms\<^sub>s\<^sub>t ` \\<^sub>0)) \ (trms\<^sub>e\<^sub>s\<^sub>t []))" "wf\<^sub>s\<^sub>t\<^sub>s' \\<^sub>0 []" "\S \ \\<^sub>0 \ {to_st []}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using assms wf\<^sub>s\<^sub>t\<^sub>s_wf\<^sub>s\<^sub>t\<^sub>s' by (metis, metis, metis, simp) obtain \\<^sub>d where \\<^sub>d: "\ = to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \\<^sub>d)" "(\\<^sub>0,[]) \\<^sup>\\<^sub>c\<^sup>* (\,\\<^sub>d)" "\ \\<^sub>c \to_st \\<^sub>d\" using pts_symbolic_to_pts_symbolic_c_from_initial assms *(3) by metis hence "tfr\<^sub>s\<^sub>e\<^sub>t (\(trms\<^sub>s\<^sub>t ` \) \ (trms\<^sub>e\<^sub>s\<^sub>t \\<^sub>d))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\(trms\<^sub>s\<^sub>t ` \) \ (trms\<^sub>e\<^sub>s\<^sub>t \\<^sub>d))" using pts_symbolic_c_preserves_tfr\<^sub>s\<^sub>e\<^sub>t[OF _ *(1,2)] by blast+ hence "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>e\<^sub>s\<^sub>t \\<^sub>d)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>e\<^sub>s\<^sub>t \\<^sub>d)" unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by (metis DiffE DiffI SMP_union UnCI, metis UnCI) hence "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t (to_st \\<^sub>d))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t (to_st \\<^sub>d))" by (metis to_st_trms_tfr\<^sub>s\<^sub>e\<^sub>t, metis to_st_trms_wf) moreover have "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r (to_st \\<^sub>d) Var" proof - have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range Var)" "subst_domain Var \ vars\<^sub>e\<^sub>s\<^sub>t \\<^sub>d = {}" "range_vars Var \ bvars\<^sub>e\<^sub>s\<^sub>t \\<^sub>d = {}" by (simp_all add: range_vars_alt_def) moreover have "wf\<^sub>e\<^sub>s\<^sub>t {} \\<^sub>d" using pts_symbolic_c_preserves_wf_is[OF \\<^sub>d(2) *(3), of "{}"] by auto moreover have "fv\<^sub>s\<^sub>t (to_st \\<^sub>d) \ bvars\<^sub>e\<^sub>s\<^sub>t \\<^sub>d = {}" using pts_symbolic_c_preserves_constr_disj_vars[OF \\<^sub>d(2)] assms(1) wf\<^sub>s\<^sub>t\<^sub>s_wf\<^sub>s\<^sub>t\<^sub>s' by fastforce ultimately show ?thesis unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by simp qed moreover have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \\<^sub>d)" using pts_symbolic_c_preserves_tfr\<^sub>s\<^sub>t\<^sub>p[OF \\<^sub>d(2) *(4)] by blast moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range Var)" by simp_all ultimately obtain \\<^sub>\ where \\<^sub>\: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "\\<^sub>\ \\<^sub>c \to_st \\<^sub>d, Var\" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" using wt_attack_if_tfr_attack[OF assms(7) \\<^sub>d(3)] \tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t (to_st \\<^sub>d))\ \list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \\<^sub>d)\ unfolding tfr\<^sub>s\<^sub>t_def by metis hence "\\<^sub>\ \ \\, Var\" using pts_symbolic_c_to_pts_symbolic_from_initial \\<^sub>d by metis thus ?thesis using \\<^sub>\(1,3,4) by metis qed subsubsection \Corollary: The Typing Result on the Level of Constraints\ text \There exists well-typed models of satisfiable type-flaw resistant constraints\ corollary wt_attack_if_tfr_attack_d: assumes "wf\<^sub>s\<^sub>t {} \" "fv\<^sub>s\<^sub>t \ \ bvars\<^sub>s\<^sub>t \ = {}" "tfr\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t \)" and "Ana_invar_subst (ik\<^sub>s\<^sub>t \ \ assignment_rhs\<^sub>s\<^sub>t \)" and "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "\ \ \\\" shows "\\\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ (\\<^sub>\ \ \\\) \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" proof - { fix S A have "({S},A) \\<^sup>\\<^sup>* ({},A@dual\<^sub>s\<^sub>t S)" proof (induction S arbitrary: A) case Nil thus ?case using pts_symbolic.Nil[of "{[]}"] by auto next case (Cons x S) hence "({S}, A@dual\<^sub>s\<^sub>t [x]) \\<^sup>\\<^sup>* ({}, A@dual\<^sub>s\<^sub>t (x#S))" by (metis dual\<^sub>s\<^sub>t_append List.append_assoc List.append_Nil List.append_Cons) moreover have "({x#S}, A) \\<^sup>\ ({S}, A@dual\<^sub>s\<^sub>t [x])" using pts_symbolic.Send[of _ S "{x#S}"] pts_symbolic.Receive[of _ S "{x#S}"] pts_symbolic.Equality[of _ _ _ S "{x#S}"] pts_symbolic.Inequality[of _ _ S "{x#S}"] by (cases x) auto ultimately show ?case by simp qed } hence 0: "({dual\<^sub>s\<^sub>t \},[]) \\<^sup>\\<^sup>* ({},\)" using dual\<^sub>s\<^sub>t_self_inverse by (metis List.append_Nil) have "fv\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t \) \ bvars\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t \) = {}" using assms(2) dual\<^sub>s\<^sub>t_fv dual\<^sub>s\<^sub>t_bvars by metis+ hence 1: "wf\<^sub>s\<^sub>t\<^sub>s {dual\<^sub>s\<^sub>t \}" using assms(1,2) dual\<^sub>s\<^sub>t_self_inverse[of \] unfolding wf\<^sub>s\<^sub>t\<^sub>s_def by auto have "\(trms\<^sub>s\<^sub>t ` {\}) = trms\<^sub>s\<^sub>t \" "\(trms\<^sub>s\<^sub>t ` {dual\<^sub>s\<^sub>t \}) = trms\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t \)" by auto hence "tfr\<^sub>s\<^sub>e\<^sub>t (\(trms\<^sub>s\<^sub>t ` {\}))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\(trms\<^sub>s\<^sub>t ` {\}))" "(\(trms\<^sub>s\<^sub>t ` {\})) = \(trms\<^sub>s\<^sub>t ` {dual\<^sub>s\<^sub>t \})" using assms(3,4) unfolding tfr\<^sub>s\<^sub>t_def by (metis, metis, metis dual\<^sub>s\<^sub>t_trms_eq) hence 2: "tfr\<^sub>s\<^sub>e\<^sub>t (\(trms\<^sub>s\<^sub>t ` {dual\<^sub>s\<^sub>t \}))" and 3: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\(trms\<^sub>s\<^sub>t ` {dual\<^sub>s\<^sub>t \}))" by metis+ have 4: "\S \ {dual\<^sub>s\<^sub>t \}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using dual\<^sub>s\<^sub>t_tfr\<^sub>s\<^sub>t\<^sub>p assms(3) unfolding tfr\<^sub>s\<^sub>t_def by blast have "assignment_rhs\<^sub>s\<^sub>t \ = assignment_rhs\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t \)" by (induct \ rule: assignment_rhs\<^sub>s\<^sub>t.induct) auto hence 5: "Ana_invar_subst (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t`{dual\<^sub>s\<^sub>t \}) \ \(assignment_rhs\<^sub>s\<^sub>t`{dual\<^sub>s\<^sub>t \}))" using assms(5) dual\<^sub>s\<^sub>t_self_inverse[of \] by auto show ?thesis by (rule wt_attack_if_tfr_attack_pts[OF 1 2 3 4 5 0 assms(6,7)]) qed end end end diff --git a/thys/Weighted_Path_Order/WPO.thy b/thys/Weighted_Path_Order/WPO.thy --- a/thys/Weighted_Path_Order/WPO.thy +++ b/thys/Weighted_Path_Order/WPO.thy @@ -1,1715 +1,1715 @@ section \The Weighted Path Order\ text \This is a version of WPO that also permits multiset comparisons of lists of terms. It therefore generalizes RPO.\ theory WPO imports Knuth_Bendix_Order.Lexicographic_Extension First_Order_Terms.Subterm_and_Context Knuth_Bendix_Order.Order_Pair Polynomial_Factorization.Missing_List Status Precedence Multiset_Extension2 HOL.Zorn begin datatype order_tag = Lex | Mul locale wpo = fixes n :: nat and S NS :: "('f, 'v) term rel" and "prc" :: "('f \ nat \ 'f \ nat \ bool \ bool)" and prl :: "'f \ nat \ bool" and \\ :: "'f status" and c :: "'f \ nat \ order_tag" and ssimple :: bool and large :: "'f \ nat \ bool" begin fun wpo :: "('f, 'v) term \ ('f, 'v) term \ bool \ bool" where "wpo s t = (if (s,t) \ S then (True, True) else if (s,t) \ NS then (case s of Var x \ (False, (case t of Var y \ x = y | Fun g ts \ status \\ (g, length ts) = [] \ prl (g, length ts))) | Fun f ss \ if \ i \ set (status \\ (f, length ss)). snd (wpo (ss ! i) t) then (True, True) else (case t of Var _ \ (False, ssimple \ large (f, length ss)) | Fun g ts \ (case prc (f, length ss) (g, length ts) of (prs, prns) \ if prns \ (\ j \ set (status \\ (g, length ts)). fst (wpo s (ts ! j))) then if prs then (True, True) else let ss' = map (\ i. ss ! i) (status \\ (f, length ss)); ts' = map (\ i. ts ! i) (status \\ (g, length ts)); cf = c (f,length ss); cg = c (g,length ts) in if cf = Lex \ cg = Lex then lex_ext wpo n ss' ts' else if cf = Mul \ cg = Mul then mul_ext wpo ss' ts' else (length ss' \ 0 \ length ts' = 0, length ts' = 0) else (False, False)))) else (False, False))" declare wpo.simps [simp del] abbreviation wpo_s (infix "\" 50) where "s \ t \ fst (wpo s t)" abbreviation wpo_ns (infix "\" 50) where "s \ t \ snd (wpo s t)" abbreviation "WPO_S \ {(s,t). s \ t}" abbreviation "WPO_NS \ {(s,t). s \ t}" lemma wpo_s_imp_ns: "s \ t \ s \ t" using lex_ext_stri_imp_nstri unfolding wpo.simps[of s t] by (auto simp: Let_def mul_ext_stri_imp_nstri split: term.splits if_splits prod.splits) lemma S_imp_wpo_s: "(s,t) \ S \ s \ t" by (simp add: wpo.simps) end declare wpo.wpo.simps[code] definition strictly_simple_status :: "'f status \ ('f,'v)term rel \ bool" where "strictly_simple_status \ rel = (\ f ts i. i \ set (status \ (f,length ts)) \ (Fun f ts, ts ! i) \ rel)" definition trans_precedence where "trans_precedence prc = (\ f g h. (fst (prc f g) \ snd (prc g h) \ fst (prc f h)) \ (snd (prc f g) \ fst (prc g h) \ fst (prc f h)) \ (snd (prc f g) \ snd (prc g h) \ snd (prc f h)))" locale wpo_with_basic_assms = wpo + order_pair + irrefl_precedence + constrains S :: "('f, 'v) term rel" and NS :: _ and prc :: "'f \ nat \ 'f \ nat \ bool \ bool" and prl :: "'f \ nat \ bool" and ssimple :: bool and large :: "'f \ nat \ bool" and c :: "'f \ nat \ order_tag" and n :: nat and \\ :: "'f status" assumes subst_S: "(s,t) \ S \ (s \ \, t \ \) \ S" and subst_NS: "(s,t) \ NS \ (s \ \, t \ \) \ NS" and irrefl_S: "irrefl S" and S_imp_NS: "S \ NS" and ss_status: "ssimple \ i \ set (status \\ fn) \ simple_arg_pos S fn i" and large: "ssimple \ large fn \ fst (prc fn gm) \ snd (prc fn gm) \ status \\ gm = []" and large_trans: "ssimple \ large fn \ snd (prc gm fn) \ large gm" and ss_S_non_empty: "ssimple \ S \ {}" begin abbreviation "\ \ status \\" lemma ss_NS_not_UNIV: "ssimple \ NS \ UNIV" proof assume "ssimple" "NS = UNIV" with ss_S_non_empty obtain a b where "(a,b) \ S" "(b,a) \ NS" by auto from compat_S_NS_point[OF this] have "(a,a) \ S" . with irrefl_S show False unfolding irrefl_def by auto qed lemmas \ = status[of \\] lemma \E: "i \ set (\ (f, length ss)) \ ss ! i \ set ss" by (rule status_aux) lemma wpo_ns_imp_NS: "s \ t \ (s,t) \ NS" using S_imp_NS by (cases s, auto simp: wpo.simps[of _ t], cases t, auto simp: refl_NS_point split: if_splits) lemma wpo_s_imp_NS: "s \ t \ (s,t) \ NS" by (rule wpo_ns_imp_NS[OF wpo_s_imp_ns]) lemma wpo_least_1: assumes "prl (f,length ss)" and "(t, Fun f ss) \ NS" and "\ (f,length ss) = []" shows "t \ Fun f ss" proof (cases t) case (Var x) with assms show ?thesis by (simp add: wpo.simps) next case (Fun g ts) let ?f = "(f,length ss)" let ?g = "(g,length ts)" obtain s ns where "prc ?g ?f = (s,ns)" by force with prl[OF assms(1), of ?g] have prc: "prc ?g ?f = (s,True)" by auto show ?thesis using assms(2) unfolding Fun unfolding wpo.simps[of "Fun g ts" "Fun f ss"] term.simps assms(3) by (auto simp: prc lex_ext_least_1 mul_ext_def ns_mul_ext_bottom Let_def) qed lemma wpo_least_2: assumes "prl (f,length ss)" (is "prl ?f") and "(Fun f ss, t) \ S" and "\ (f,length ss) = []" shows "\ Fun f ss \ t" proof (cases t) case (Var x) with Var show ?thesis using assms(2-3) by (auto simp: wpo.simps split: if_splits) next case (Fun g ts) let ?g = "(g,length ts)" obtain s ns where "prc ?f ?g = (s,ns)" by force with prl2[OF assms(1), of ?g] have prc: "prc ?f ?g = (False,ns)" by auto show ?thesis using assms(2) assms(3) unfolding Fun by (simp add: wpo.simps[of _ "Fun g ts"] lex_ext_least_2 prc mul_ext_def s_mul_ext_bottom_strict Let_def) qed lemma wpo_least_3: assumes "prl (f,length ss)" (is "prl ?f") and ns: "Fun f ss \ t" and NS: "(u, Fun f ss) \ NS" and ss: "\ (f,length ss) = []" and S: "\ x. (Fun f ss, x) \ S" and u: "u = Var x" shows "u \ t" proof (cases "(Fun f ss, t) \ S \ (u, Fun f ss) \ S \ (u, t) \ S") case True with wpo_ns_imp_NS[OF ns] NS compat_NS_S_point compat_S_NS_point have "(u, t) \ S" by blast from wpo_s_imp_ns[OF S_imp_wpo_s[OF this]] show ?thesis . next case False from trans_NS_point[OF NS wpo_ns_imp_NS[OF ns]] have utA: "(u, t) \ NS" . show ?thesis proof (cases t) case t: (Var y) with ns False ss have *: "ssimple" "large (f,length ss)" by (auto simp: wpo.simps split: if_splits) show ?thesis proof (cases "x = y") case True thus ?thesis using ns * False utA ss unfolding wpo.simps[of u t] wpo.simps[of "Fun f ss" t] unfolding t u term.simps by (auto split: if_splits) next case False from utA[unfolded t u] have "(Var x, Var y) \ NS" . from False subst_NS[OF this, of "\ z. if z = x then v else w" for v w] have "(v,w) \ NS" for v w by auto hence "NS = UNIV" by auto with ss_NS_not_UNIV[OF \ssimple\] have False by auto thus ?thesis .. qed next case (Fun g ts) let ?g = "(g,length ts)" obtain s ns where "prc ?f ?g = (s,ns)" by force with prl2[OF \prl ?f\, of ?g] have prc: "prc ?f ?g = (False,ns)" by auto show ?thesis proof (cases "\ ?g") case Nil with False Fun assms prc have "prc ?f ?g = (False,True)" by (auto simp: wpo.simps split: if_splits) with prl3[OF \prl ?f\, of ?g] have "prl ?g" by auto show ?thesis using utA unfolding Fun by (rule wpo_least_1[OF \prl ?g\], simp add: Nil) next case (Cons t1 tts) have "\ wpo_s (Fun f ss) (ts ! t1)" by (rule wpo_least_2[OF \prl ?f\ S ss]) with \wpo_ns (Fun f ss) t\ False Fun Cons have False by (simp add: ss wpo.simps split: if_splits) then show ?thesis .. qed qed qed (* Transitivity / compatibility of the orders *) lemma wpo_compat': assumes SN: "\ f. prl f \ SN S" (* it is not clear whether this is really required *) shows "(s \ t \ t \ u \ s \ u) \ (s \ t \ t \ u \ s \ u) \ (s \ t \ t \ u \ s \ u)" (is "?tran s t u") proof (induct "(s,t,u)" arbitrary: s t u rule: wf_induct[OF wf_measures[of "[\ (s,t,u). size s, \ (s,t,u). size t, \ (s,t,u). size u]"]]) case 1 note ind = 1[simplified] show "?tran s t u" proof (cases "(s,t) \ S \ (t,u) \ S \ (s,u) \ S") case True { assume st: "wpo_ns s t" and tu: "wpo_ns t u" from wpo_ns_imp_NS[OF st] wpo_ns_imp_NS[OF tu] True compat_NS_S_point compat_S_NS_point have "(s,u) \ S" by blast from S_imp_wpo_s[OF this] have "wpo_s s u" . } with wpo_s_imp_ns show ?thesis by blast next case False then have stS: "(s,t) \ S" and tuS: "(t,u) \ S" and suS: "(s,u) \ S" by auto show ?thesis proof (cases t) note [simp] = wpo.simps[of s u] wpo.simps[of s t] wpo.simps[of t u] case (Var x) note wpo.simps[simp] show ?thesis proof safe assume "wpo_s t u" with Var tuS show "wpo_s s u" by (auto split: if_splits) next assume gr: "wpo_s s t" and ge: "wpo_ns t u" from wpo_s_imp_NS[OF gr] have stA: "(s,t) \ NS" . from wpo_ns_imp_NS[OF ge] have tuA: "(t,u) \ NS" . from trans_NS_point[OF stA tuA] have suA: "(s,u) \ NS" . show "wpo_s s u" proof (cases u) case (Var y) with ge \t = Var x\ tuS have "t = u" by (auto split: if_splits) with gr show ?thesis by auto next case (Fun h us) let ?h = "(h,length us)" from Fun ge Var tuS have us: "\ ?h = []" and pri: "prl ?h" by (auto split: if_splits) from gr Var tuS ge stS obtain f ss where s: "s = Fun f ss" by (cases s, auto split: if_splits) let ?f = "(f,length ss)" from s gr Var False obtain i where i: "i \ set (\ ?f)" and sit: "ss ! i \ t" by (auto split: if_splits) from trans_NS_point[OF wpo_ns_imp_NS[OF sit] tuA] have siu: "(ss ! i,u) \ NS" . from wpo_least_1[OF pri siu[unfolded Fun us] us] have "ss ! i \ u" unfolding Fun us . with i have "\ i \ set (\ ?f). ss ! i \ u" by blast with s suA show ?thesis by simp qed next assume ge1: "wpo_ns s t" and ge2: "wpo_ns t u" show "wpo_ns s u" proof (cases u) case (Var y) with ge2 \t = Var x\ tuS have "t = u" by (auto split: if_splits) with ge1 show ?thesis by auto next case (Fun h us) let ?h = "(h,length us)" from Fun ge2 Var tuS have us: "\ ?h = []" and pri: "prl ?h" by (auto split: if_splits) show ?thesis unfolding Fun us by (rule wpo_least_1[OF pri trans_NS_point[OF wpo_ns_imp_NS[OF ge1] wpo_ns_imp_NS[OF ge2[unfolded Fun us]]] us]) qed qed next case (Fun g ts) let ?g = "(g,length ts)" let ?ts = "set (\ ?g)" let ?t = "Fun g ts" from Fun have t: "t = ?t" . show ?thesis proof (cases s) case (Var x) show ?thesis proof safe assume gr: "wpo_s s t" with Var Fun stS show "wpo_s s u" by (auto simp: wpo.simps split: if_splits) next assume ge: "wpo_ns s t" and gr: "wpo_s t u" with Var Fun stS have pri: "prl ?g" and "\ ?g = []" by (auto simp: wpo.simps split: if_splits) with gr Fun show "wpo_s s u" using wpo_least_2[OF pri, of u] False by auto next assume ge1: "wpo_ns s t" and ge2: "wpo_ns t u" with Var Fun stS have pri: "prl ?g" and empty: "\ ?g = []" by (auto simp: wpo.simps split: if_splits) from wpo_ns_imp_NS[OF ge1] Var Fun empty have ns: "(Var x, Fun g ts) \ NS" by simp show "wpo_ns s u" proof (rule wpo_least_3[OF pri ge2[unfolded Fun empty] wpo_ns_imp_NS[OF ge1[unfolded Fun empty]] empty _ Var], rule) fix v assume S: "(Fun g ts, v) \ S" from compat_NS_S_point[OF ns S] have xv: "(Var x, v) \ S" . from assms[OF pri] have SN: "SN S" . from SN_imp_minimal[OF SN, rule_format, of undefined UNIV] obtain s where "\ u. (s,u) \ S" by blast with subst_S[OF xv, of "\ _. s"] show False by auto qed qed next case (Fun f ss) let ?s = "Fun f ss" let ?f = "(f,length ss)" let ?ss = "set (\ ?f)" from Fun have s: "s = ?s" . let ?s1 = "\ i \ ?ss. ss ! i \ t" let ?t1 = "\ j \ ?ts. ts ! j \ u" let ?ls = "length ss" let ?lt = "length ts" obtain ps pns where prc: "prc ?f ?g = (ps,pns)" by force let ?tran2 = "\ a b c. ((wpo_ns a b) \ (wpo_s b c) \ (wpo_s a c)) \ ((wpo_s a b) \ (wpo_ns b c) \ (wpo_s a c)) \ ((wpo_ns a b) \ (wpo_ns b c) \ (wpo_ns a c)) \ ((wpo_s a b) \ (wpo_s b c) \ (wpo_s a c))" from s have "\ s' \ set ss. size s' < size s" by (auto simp: size_simps) with ind have ind2: "\ s' t' u'. \s' \ set ss\ \ ?tran s' t' u'" by blast with wpo_s_imp_ns have ind3: "\ us s' t' u'. \s' \ set ss; t' \ set ts\ \ ?tran2 s' t' u'" by blast let ?mss = "map (\ i. ss ! i) (\ ?f)" let ?mts = "map (\ j. ts ! j) (\ ?g)" have ind3': "\ us s' t' u'. \s' \ set ?mss; t' \ set ?mts\ \ ?tran2 s' t' u'" by (rule ind3, auto simp: status_aux) { assume ge1: "s \ t" and ge2: "t \ u" from wpo_ns_imp_NS[OF ge1] have stA: "(s,t) \ NS" . from wpo_s_imp_NS[OF ge2] have tuA: "(t,u) \ NS" . from trans_NS_point[OF stA tuA] have suA: "(s,u) \ NS" . have "s \ u" proof (cases ?s1) case True from this obtain i where i: "i \ ?ss" and ges: "ss ! i \ t" by auto from \E[OF i] have s': "ss ! i \ set ss" . with i s s' ind2[of "ss ! i" t u, simplified] ges ge2 have "ss ! i \ u" by auto then have "ss ! i \ u" by (rule wpo_s_imp_ns) with i s suA show ?thesis by (cases u, auto simp: wpo.simps split: if_splits) next case False show ?thesis proof (cases ?t1) case True from this obtain j where j: "j \ ?ts" and ges: "ts ! j \ u" by auto from \E[OF j] have t': "ts ! j \ set ts" by auto from j t' t stS False ge1 s have ge1': "s \ ts ! j" unfolding wpo.simps[of s t] by (auto split: if_splits prod.splits) from t' s t ge1' ges ind[rule_format, of s "ts ! j" u, simplified] show "s \ u" using suA size_simps supt.intros unfolding wpo.simps[of s u] by (auto split: if_splits) next case False from t this ge2 tuS obtain h us where u: "u = Fun h us" by (cases u, auto simp: wpo.simps split: if_splits) let ?u = "Fun h us" let ?h = "(h,length us)" let ?us = "set (\ ?h)" let ?mus = "map (\ k. us ! k) (\ ?h)" from s t u ge1 ge2 have ge1: "?s \ ?t" and ge2: "?t \ ?u" by auto from stA stS s t have stAS: "((?s,?t) \ S) = False" "((?s,?t) \ NS) = True" by auto from tuA tuS t u have tuAS: "((?t,?u) \ S) = False" "((?t,?u) \ NS) = True" by auto note ge1 = ge1[unfolded wpo.simps[of ?s ?t] stAS, simplified] note ge2 = ge2[unfolded wpo.simps[of ?t ?u] tuAS, simplified] obtain ps2 pns2 where prc2: "prc ?g ?h = (ps2,pns2)" by force obtain ps3 pns3 where prc3: "prc ?f ?h = (ps3,pns3)" by force from \\ ?s1\ t ge1 have st': "\ j \ ?ts. ?s \ ts ! j" by (auto split: if_splits prod.splits) from \\ ?t1\ t u ge2 tuS have tu': "\ k \ ?us. ?t \ us ! k" by (auto split: if_splits prod.splits) from \\ ?s1\ s t ge1 stS st' have fg: "pns" by (cases ?thesis, auto simp: prc) from \\ ?t1\ u ge2 tu' have gh: "pns2" by (cases ?thesis, auto simp: prc2) from \\ ?s1\ have "?s1 = False" by simp note ge1 = ge1[unfolded this[unfolded t] if_False term.simps prc split] from \\ ?t1\ have "?t1 = False" by simp note ge2 = ge2[unfolded this[unfolded u] if_False term.simps prc2 split] note compat = prc_compat[OF prc prc2 prc3] from fg gh compat have fh: "pns3" by simp { fix k assume k: "k \ ?us" from \E[OF this] have "size (us ! k) < size u" unfolding u using size_simps by auto with tu'[folded t] \s \ t\ ind[rule_format, of s t "us ! k"] k have "s \ us ! k" by blast } note su' = this show ?thesis proof (cases ps3) case True with su' s u fh prc3 suA show ?thesis by (auto simp: wpo.simps) next case False from False fg gh compat have nfg: "\ ps" and ngh: "\ ps2" and *: "ps = False" "ps2 = False" by blast+ note ge1 = ge1[unfolded * if_False] note ge2 = ge2[unfolded * if_False] show ?thesis proof (cases "c ?f") case Mul note cf = this show ?thesis proof (cases "c ?g") case Mul note cg = this show ?thesis proof (cases "c ?h") case Mul note ch = this from ge1[unfolded cf cg] have mul1: "snd (mul_ext wpo ?mss ?mts)" by (auto split: if_splits) from ge2[unfolded cg ch] have mul2: "fst (mul_ext wpo ?mts ?mus)" by (auto split: if_splits) from mul1 mul2 mul_ext_compat[OF ind3', of ?mss ?mts ?mus] have "fst (mul_ext wpo ?mss ?mus)" by auto with s u fh su' prc3 cf ch suA show ?thesis unfolding wpo.simps[of s u] by simp next case Lex note ch = this from gh u ge2 tu' prc2 ngh cg ch have us_e: "?mus = []" by simp from gh u ge2 tu' prc2 ngh cg ch have ts_ne: "?mts \ []" by (auto split: if_splits) from ns_mul_ext_bottom_uniqueness[of "mset ?mts"] have "\f. snd (mul_ext f [] ?mts) \ ?mts = []" unfolding mul_ext_def by (simp add: Let_def) with ts_ne fg \\ ?s1\ t ge1 st' prc nfg cf cg have ss_ne: "?mss \ []" by (cases ss) auto from us_e ss_ne s u fh su' prc3 cf cg ch suA show ?thesis unfolding wpo.simps[of s u] by simp qed next case Lex note cg = this from fg \\ ?s1\ t ge1 st' prc nfg cf cg have ts_e: "?mts = []" by simp with gh \\ ?t1\ u ge2 tu' prc2 ngh cg show ?thesis by (cases "c ?h") (simp_all add: lex_ext_least_2) qed next case Lex note cf = this show ?thesis proof (cases "c ?g") case Mul note cg = this from fg \\ ?s1\ t ge1 st' prc nfg cf cg have ts_e: "?mts = []" by simp with gh \\ ?t1\ u ge2 tu' prc2 ngh cg show ?thesis by (cases "c ?h") (auto simp: Let_def s_mul_ext_def s_mul_ext_bottom mul_ext_def elim: mult2_alt_sE) next case Lex note cg = this show ?thesis proof (cases "c ?h") case Mul note ch = this from gh u ge2 tu' ngh cg ch have us_e: "?mus = []" by simp from gh u ge2 tu' ngh cg ch have ts_ne: "?mts \ []" by simp from lex_ext_iff[of _ _ "[]" ?mts] have "\f. snd (lex_ext f n [] ?mts) \ ?mts = []" by simp with ts_ne fg t ge1 st' nfg cf cg have ss_ne: "?mss \ []" by auto from us_e ss_ne s u fh su' prc3 cf cg ch suA show ?thesis unfolding wpo.simps[of s u] by simp next case Lex note ch = this from fg t ge1 st' nfg cf cg have lex1: "snd (lex_ext wpo n ?mss ?mts)" by auto from gh u ge2 tu' ngh cg ch have lex2: "fst (lex_ext wpo n ?mts ?mus)" by auto from lex1 lex2 lex_ext_compat[OF ind3', of ?mss ?mts ?mus] have "fst (lex_ext wpo n ?mss ?mus)" by auto with s u fh su' prc3 cf cg ch suA show ?thesis unfolding wpo.simps[of s u] by simp qed qed qed qed qed qed } moreover { assume ge1: "s \ t" and ge2: "t \ u" from wpo_s_imp_NS[OF ge1] have stA: "(s,t) \ NS" . from wpo_ns_imp_NS[OF ge2] have tuA: "(t,u) \ NS" . from trans_NS_point[OF stA tuA] have suA: "(s,u) \ NS" . have "s \ u" proof (cases ?s1) case True from True obtain i where i: "i \ ?ss" and ges: "ss ! i \ t" by auto from \E[OF i] have s': "ss ! i \ set ss" by auto with s s' ind2[of "ss ! i" t u, simplified] ges ge2 have "ss ! i \ u" by auto with i s' s suA show ?thesis by (cases u, auto simp: wpo.simps split: if_splits) next case False show ?thesis proof (cases ?t1) case True from this obtain j where j: "j \ ?ts" and ges: "ts ! j \ u" by auto from \E[OF j] have t': "ts ! j \ set ts" . from j t' t stS False ge1 s have ge1': "s \ ts ! j" unfolding wpo.simps[of s t] by (auto split: if_splits prod.splits) from t' s t ge1' ges ind[rule_format, of s "ts ! j" u, simplified] show "s \ u" using suA size_simps supt.intros unfolding wpo.simps[of s u] by (auto split: if_splits) next case False show ?thesis proof (cases u) case u: (Fun h us) let ?u = "Fun h us" let ?h = "(h,length us)" let ?us = "set (\ ?h)" let ?mss = "map (\ i. ss ! i) (\ ?f)" let ?mts = "map (\ j. ts ! j) (\ ?g)" let ?mus = "map (\ k. us ! k) (\ ?h)" note \E = \E[of _ f ss] \E[of _ g ts] \E[of _ h us] from s t u ge1 ge2 have ge1: "?s \ ?t" and ge2: "?t \ ?u" by auto from stA stS s t have stAS: "((?s,?t) \ S) = False" "((?s,?t) \ NS) = True" by auto from tuA tuS t u have tuAS: "((?t,?u) \ S) = False" "((?t,?u) \ NS) = True" by auto note ge1 = ge1[unfolded wpo.simps[of ?s ?t] stAS, simplified] note ge2 = ge2[unfolded wpo.simps[of ?t ?u] tuAS, simplified] let ?lu = "length us" obtain ps2 pns2 where prc2: "prc ?g ?h = (ps2,pns2)" by force obtain ps3 pns3 where prc3: "prc ?f ?h = (ps3,pns3)" by force from \\ ?s1\ t ge1 have st': "\ j \ ?ts. ?s \ ts ! j" by (auto split: if_splits prod.splits) from \\ ?t1\ t u ge2 tuS have tu': "\ k \ ?us. ?t \ us ! k" by (auto split: if_splits prod.splits) from \\ ?s1\ s t ge1 stS st' have fg: "pns" by (cases ?thesis, auto simp: prc) from \\ ?t1\ u ge2 tu' have gh: "pns2" by (cases ?thesis, auto simp: prc2) from \\ ?s1\ have "?s1 = False" by simp note ge1 = ge1[unfolded this[unfolded t] if_False term.simps prc split] from \\ ?t1\ have "?t1 = False" by simp note ge2 = ge2[unfolded this[unfolded u] if_False term.simps prc2 split] note compat = prc_compat[OF prc prc2 prc3] from fg gh compat have fh: "pns3" by simp { fix k assume k: "k \ ?us" from \E(3)[OF this] have "size (us ! k) < size u" unfolding u using size_simps by auto with tu'[folded t] wpo_s_imp_ns[OF \s \ t\] ind[rule_format, of s t "us ! k"] k have "s \ us ! k" by blast } note su' = this show ?thesis proof (cases ps3) case True with su' s u fh prc3 suA show ?thesis by (auto simp: wpo.simps) next case False from False fg gh compat have nfg: "\ ps" and ngh: "\ ps2" and *: "ps = False" "ps2 = False" by blast+ note ge1 = ge1[unfolded * if_False] note ge2 = ge2[unfolded * if_False] show ?thesis proof (cases "c ?f") case Mul note cf = this show ?thesis proof (cases "c ?g") case Mul note cg = this show ?thesis proof (cases "c ?h") case Mul note ch = this from fg t ge1 st' nfg cf cg have mul1: "fst (mul_ext wpo ?mss ?mts)" by auto from gh u ge2 tu' ngh cg ch have mul2: "snd (mul_ext wpo ?mts ?mus)" by auto from mul1 mul2 mul_ext_compat[OF ind3', of ?mss ?mts ?mus] have "fst (mul_ext wpo ?mss ?mus)" by auto with s u fh su' prc3 cf ch suA show ?thesis unfolding wpo.simps[of s u] by simp next case Lex note ch = this from gh u ge2 tu' ngh cg ch have us_e: "?mus = []" by simp from fg t ge1 st' nfg cf cg s_mul_ext_bottom_strict have ss_ne: "?mss \ []" by (cases ?mss) (auto simp: Let_def mul_ext_def) from us_e ss_ne s u fh su' prc3 cf cg ch suA show ?thesis unfolding wpo.simps[of s u] by simp qed next case Lex note cg = this from fg t ge1 st' prc nfg cf cg s_mul_ext_bottom_strict have ss_ne: "?mss \ []" by (auto simp: mul_ext_def) from fg t ge1 st' nfg cf cg have ts_e: "?mts = []" by simp show ?thesis proof (cases "c ?h") case Mul note ch = this with gh u ge2 tu' ngh cg ch ns_mul_ext_bottom_uniqueness have "?mus = []" by simp with ss_ne s u fh su' prc3 cf cg ch s_mul_ext_bottom suA show ?thesis unfolding wpo.simps[of s u] by (simp add: Let_def mul_ext_def s_mul_ext_def mult2_alt_s_def) next case Lex note ch = this from lex_ext_iff[of _ _ "[]" ?mus] have "\f. snd (lex_ext f n [] ?mus) \ ?mus = []" by simp with ts_e gh u ge2 tu' ngh cg ch have "?mus = []" by simp with ss_ne s u fh su' prc3 cf cg ch s_mul_ext_bottom suA show ?thesis unfolding wpo.simps[of s u] by (simp add: mul_ext_def) qed qed next case Lex note cf = this show ?thesis proof (cases "c ?g") case Mul note cg = this from fg t ge1 st' nfg cf cg have ss_ne: "?mss \ []" by simp from fg t ge1 st' nfg cf cg have ts_e: "?mts = []" by simp show ?thesis proof (cases "c ?h") case Mul note ch = this from ts_e gh u ge2 tu' ngh cg ch ns_mul_ext_bottom_uniqueness[of "mset ?mus"] have "?mus = []" by (simp add: mul_ext_def Let_def) with ss_ne s u fh su' prc3 cf cg ch s_mul_ext_bottom suA show ?thesis unfolding wpo.simps[of s u] by (simp add: mul_ext_def) next case Lex note ch = this from gh u ge2 tu' prc2 ngh cg ch have "?mus = []" by simp with ss_ne s u fh su' prc3 cf cg ch suA show ?thesis unfolding wpo.simps[of s u] by (simp add: lex_ext_iff) qed next case Lex note cg = this show ?thesis proof (cases "c ?h") case Mul note ch = this from gh u ge2 tu' ngh cg ch have us_e: "?mus = []" by simp have "\f. fst (lex_ext f n ?mss ?mts) \ ?mss \ []" by (cases ?mss) (simp_all add: lex_ext_iff) with fg t ge1 st' prc nfg cf cg have ss_ne: "?mss \ []" by simp with us_e s u fh su' prc3 cf cg ch suA show ?thesis unfolding wpo.simps[of s u] by simp next case Lex note ch = this from fg t ge1 st' nfg cf cg have lex1: "fst (lex_ext wpo n ?mss ?mts)" by auto from gh u ge2 tu' ngh cg ch have lex2: "snd (lex_ext wpo n ?mts ?mus)" by auto from lex1 lex2 lex_ext_compat[OF ind3', of ?mss ?mts ?mus] have "fst (lex_ext wpo n ?mss ?mus)" by auto with s u fh su' prc3 cf cg ch suA show ?thesis unfolding wpo.simps[of s u] by simp qed qed qed qed next case (Var z) from ge2 \\ ?t1\ tuS have "ssimple" "large ?g" unfolding Var t by (auto simp: wpo.simps split: if_splits) from large[OF this, of ?f] have large: "fst (prc ?g ?f) \ snd (prc ?g ?f) \ \ ?f = []" by auto obtain fgs fgns where prc_fg: "prc ?f ?g = (fgs,fgns)" by (cases "prc ?f ?g", auto) from ge1 \\ ?s1\ stS have weak_fg: "snd (prc ?f ?g)" unfolding s t using prc_fg by (auto simp: wpo.simps split: if_splits) have prc_irrefl: "\ fst (prc ?f ?f)" using prc_refl by simp from large have False proof assume "fst (prc ?g ?f)" with weak_fg have "fst (prc ?f ?f)" by (metis prc_compat prod.collapse) with prc_irrefl show False by auto next assume weak: "snd (prc ?g ?f) \ \ ?f = []" let ?mss = "map (\ i. ss ! i) (\ ?f)" let ?mts = "map (\ j. ts ! j) (\ ?g)" { assume "fst (prc ?f ?g)" with weak have "fst (prc ?f ?f)" by (metis prc_compat prod.collapse) with prc_irrefl have False by auto } hence "\ fst (prc ?f ?g)" by auto with ge1 \\ ?s1\ stS prc_fg have "fst (lex_ext wpo n ?mss ?mts) \ fst (mul_ext wpo ?mss ?mts) \ ?mss \ []" unfolding wpo.simps[of s t] unfolding s t by (auto simp: Let_def split: if_splits) with weak have "fst (lex_ext wpo n [] ?mts) \ fst (mul_ext wpo [] ?mts)" by auto thus False using lex_ext_least_2 by (auto simp: mul_ext_def Let_def s_mul_ext_bottom_strict) qed thus ?thesis .. qed qed qed } moreover { assume ge1: "s \ t" and ge2: "t \ u" and ngt1: "\ s \ t" and ngt2: "\ t \ u" from wpo_ns_imp_NS[OF ge1] have stA: "(s,t) \ NS" . from wpo_ns_imp_NS[OF ge2] have tuA: "(t,u) \ NS" . from trans_NS_point[OF stA tuA] have suA: "(s,u) \ NS" . from ngt1 stA have "\ ?s1" unfolding s t by (auto simp: wpo.simps split: if_splits) from ngt2 tuA have "\ ?t1" unfolding t by (cases u, auto simp: wpo.simps split: if_splits) have "s \ u" proof (cases u) case u: (Var x) from t \\ ?t1\ ge2 tuA ngt2 have large: "ssimple" "large ?g" unfolding u by (auto simp: wpo.simps split: if_splits) from s t ngt1 ge1 have "snd (prc ?f ?g)" by (auto simp: wpo.simps split: if_splits prod.splits) from large_trans[OF large this] suA large show ?thesis unfolding wpo.simps[of s u] using s u by auto next case u: (Fun h us) let ?u = "Fun h us" let ?h = "(h,length us)" let ?us = "set (\ ?h)" let ?mss = "map (\ i. ss ! i) (\ ?f)" let ?mts = "map (\ j. ts ! j) (\ ?g)" let ?mus = "map (\ k. us ! k) (\ ?h)" from s t u ge1 ge2 have ge1: "?s \ ?t" and ge2: "?t \ ?u" by auto from stA stS s t have stAS: "((?s,?t) \ S) = False" "((?s,?t) \ NS) = True" by auto from tuA tuS t u have tuAS: "((?t,?u) \ S) = False" "((?t,?u) \ NS) = True" by auto note ge1 = ge1[unfolded wpo.simps[of ?s ?t] stAS, simplified] note ge2 = ge2[unfolded wpo.simps[of ?t ?u] tuAS, simplified] from s t u ngt1 ngt2 have ngt1: "\ ?s \ ?t" and ngt2: "\ ?t \ ?u" by auto note ngt1 = ngt1[unfolded wpo.simps[of ?s ?t] stAS, simplified] note ngt2 = ngt2[unfolded wpo.simps[of ?t ?u] tuAS, simplified] from \\ ?s1\ t ge1 have st': "\ j \ ?ts. ?s \ ts ! j" by (cases ?thesis, auto) from \\ ?t1\ u ge2 have tu': "\ k \ ?us. ?t \ us ! k" by (cases ?thesis, auto) let ?lu = "length us" obtain ps2 pns2 where prc2: "prc ?g ?h = (ps2,pns2)" by force obtain ps3 pns3 where prc3: "prc ?f ?h = (ps3,pns3)" by force from \\ ?s1\ t ge1 st' have fg: "pns" by (cases ?thesis, auto simp: prc) from \\ ?t1\ u ge2 tu' have gh: "pns2" by (cases ?thesis, auto simp: prc2) note compat = prc_compat[OF prc prc2 prc3] from \\ ?s1\ have "?s1 = False" by simp note ge1 = ge1[unfolded this[unfolded t] if_False term.simps prc split] from \\ ?t1\ have "?t1 = False" by simp note ge2 = ge2[unfolded this[unfolded u] if_False term.simps prc2 split] from compat fg gh have fh: pns3 by blast { fix k assume k: "k \ ?us" from \E[OF this] have "size (us ! k) < size u" unfolding u using size_simps by auto with tu'[folded t] \s \ t\ ind[rule_format, of s t "us ! k"] k have "s \ us ! k" by blast } note su' = this from \\ ?s1\ st' ge1 ngt1 s t have nfg: "\ ps" by (simp, cases ?thesis, simp, cases ps, auto simp: prc fg) from \\ ?t1\ tu' ge2 ngt2 t u have ngh: "\ ps2" by (simp, cases ?thesis, simp, cases ps2, auto simp: prc2 gh) show "s \ u" proof (cases "c ?f") case Mul note cf = this show ?thesis proof (cases "c ?g") case Mul note cg = this show ?thesis proof (cases "c ?h") case Mul note ch = this from fg t ge1 st' nfg cf cg have mul1: "snd (mul_ext wpo ?mss ?mts)" by auto from gh u ge2 tu' ngh cg ch have mul2: "snd (mul_ext wpo ?mts ?mus)" by auto from mul1 mul2 mul_ext_compat[OF ind3', of ?mss ?mts ?mus] have "snd (mul_ext wpo ?mss ?mus)" by auto with s u fh su' prc3 cf ch suA show ?thesis unfolding wpo.simps[of s u] by simp next case Lex note ch = this from gh u ge2 tu' ngh cg ch have us_e: "?mus = []" by simp with s u fh su' prc3 cf cg ch suA show ?thesis unfolding wpo.simps[of s u] by simp qed next case Lex note cg = this from fg t ge1 st' nfg cf cg have ts_e: "?mts = []" by simp show ?thesis proof (cases "c ?h") case Mul note ch = this with gh u ge2 tu' ngh cg ch have "?mus = []" by simp with s u fh su' prc3 cf cg ch ns_mul_ext_bottom suA show ?thesis unfolding wpo.simps[of s u] by (simp add: ns_mul_ext_def mul_ext_def Let_def mult2_alt_ns_def) next case Lex note ch = this have "\f. snd (lex_ext f n [] ?mus) \ ?mus = []" by (simp_all add: lex_ext_iff) with ts_e gh u ge2 tu' ngh cg ch have "?mus = []" by simp with s u fh su' prc3 cf cg ch suA show ?thesis unfolding wpo.simps[of s u] by simp qed qed next case Lex note cf = this show ?thesis proof (cases "c ?g") case Mul note cg = this from fg t ge1 st' prc nfg cf cg have ts_e: "?mts = []" by simp show ?thesis proof (cases "c ?h") case Mul note ch = this with ts_e gh u ge2 tu' ngh cg ch ns_mul_ext_bottom_uniqueness[of "mset ?mus"] have "?mus = []" by (simp add: Let_def mul_ext_def) with s u fh su' prc3 cf cg ch suA show ?thesis unfolding wpo.simps[of s u] by simp next case Lex note ch = this with gh u ge2 tu' prc2 ngh cg ch have "?mus = []" by simp with s u fh su' prc3 cf cg ch suA show ?thesis unfolding wpo.simps[of s u] by (simp add: lex_ext_least_1) qed next case Lex note cg = this show ?thesis proof (cases "c ?h") case Mul note ch = this with gh u ge2 tu' ngh cg ch have "?mus = []" by simp with s u fh su' prc3 cf cg ch suA show ?thesis unfolding wpo.simps[of s u] by (simp add: lex_ext_least_1) next case Lex note ch = this from st' ge1 s t fg nfg cf cg have lex1: "snd (lex_ext wpo n ?mss ?mts)" by (auto simp: prc) from tu' ge2 t u gh ngh cg ch have lex2: "snd (lex_ext wpo n ?mts ?mus)" by (auto simp: prc2) from lex1 lex2 lex_ext_compat[OF ind3', of ?mss ?mts ?mus] have "snd (lex_ext wpo n ?mss ?mus)" by auto with fg gh su' s u fh cf cg ch suA show ?thesis unfolding wpo.simps[of s u] by (auto simp: prc3) qed qed qed qed } ultimately show ?thesis using wpo_s_imp_ns by auto qed qed qed qed context assumes ssimple: "strictly_simple_status \\ NS" begin lemma NS_arg': assumes i: "i \ set (\ (f,length ts))" shows "(Fun f ts, ts ! i) \ NS" using assms ssimple unfolding simple_arg_pos_def strictly_simple_status_def by simp lemma wpo_ns_refl': shows "s \ s" proof (induct s) case (Fun f ss) { fix i assume si: "i \ set (\ (f,length ss))" from NS_arg'[OF this] have "(Fun f ss, ss ! i) \ NS" . with si Fun[OF status_aux[OF si]] have "wpo_s (Fun f ss) (ss ! i)" unfolding wpo.simps[of "Fun f ss" "ss ! i"] by auto } note wpo_s = this let ?ss = "map (\ i. ss ! i) (\ (f,length ss))" have rec11: "snd (lex_ext wpo n ?ss ?ss)" by (rule all_nstri_imp_lex_nstri, insert \E[of _ f ss], auto simp: Fun) have rec12: "snd (mul_ext wpo ?ss ?ss)" unfolding mul_ext_def Let_def snd_conv by (intro ns_mul_ext_refl_local, unfold locally_refl_def, auto simp: in_multiset_in_set[of ?ss] intro!: Fun status_aux) from rec11 rec12 show ?case using refl_NS_point wpo_s by (cases "c (f,length ss)", auto simp: wpo.simps[of "Fun f ss" "Fun f ss"] prc_refl) qed (simp add: wpo.simps refl_NS_point) lemma wpo_stable': fixes \ :: "('f,'v)subst" shows "(s \ t \ s \ \ \ t \ \) \ (s \ t \ s \ \ \ t \ \)" (is "?p s t") proof (induct "(s,t)" arbitrary:s t rule: wf_induct[OF wf_measure[of "\ (s,t). size s + size t"]]) case (1 s t) from 1 have "\ s' t'. size s' + size t' < size s + size t \ ?p s' t'" by auto note IH = this[rule_format] let ?s = "s \ \" let ?t = "t \ \" note simps = wpo.simps[of s t] wpo.simps[of ?s ?t] show "?case" proof (cases "((s,t) \ S \ (?s,?t) \ S) \ ((s,t) \ NS \ \ wpo_ns s t)") case True then show ?thesis proof assume "(s,t) \ S \ (?s,?t) \ S" with subst_S[of s t \] have "(?s,?t) \ S" by blast from S_imp_wpo_s[OF this] have "wpo_s ?s ?t" . with wpo_s_imp_ns[OF this] show ?thesis by blast next assume "(s,t) \ NS \ \ wpo_ns s t" with wpo_ns_imp_NS have st: "\ wpo_ns s t" by auto with wpo_s_imp_ns have "\ wpo_s s t" by auto with st show ?thesis by blast qed next case False then have not: "((s,t) \ S) = False" "((?s,?t) \ S) = False" and stA: "(s,t) \ NS" and ns: "wpo_ns s t" by auto from subst_NS[OF stA] have sstsA: "(?s,?t) \ NS" by auto from stA sstsA have id: "((s,t) \ NS) = True" "((?s,?t) \ NS) = True" by auto note simps = simps[unfolded id not if_False if_True] show ?thesis proof (cases s) case (Var x) note s = this show ?thesis proof (cases t) case (Var y) note t = this show ?thesis unfolding simps(1) unfolding s t using wpo_ns_refl'[of "\ y"] by auto next case (Fun g ts) note t = this let ?g = "(g,length ts)" show ?thesis proof (cases "\ x") case (Var y) then show ?thesis unfolding simps unfolding s t by simp next case (Fun f ss) let ?f = "(f, length ss)" show ?thesis proof (cases "prl ?g") case False then show ?thesis unfolding simps unfolding s t Fun by auto next case True obtain s ns where "prc ?f ?g = (s,ns)" by force with prl[OF True, of ?f] have prc: "prc ?f ?g = (s, True)" by auto show ?thesis unfolding simps unfolding s t Fun by (auto simp: Fun prc mul_ext_def ns_mul_ext_bottom Let_def intro!: all_nstri_imp_lex_nstri[of "[]", simplified]) qed qed qed next case (Fun f ss) note s = this let ?f = "(f,length ss)" let ?ss = "set (\ ?f)" { fix i assume i: "i \ ?ss" and ns: "wpo_ns (ss ! i) t" from IH[of "ss ! i" t] \E[OF i] ns have "wpo_ns (ss ! i \ \) ?t" using s by (auto simp: size_simps) then have "wpo_s ?s ?t" using i sstsA \[of f "length ss"] unfolding simps unfolding s by force with wpo_s_imp_ns[OF this] have ?thesis by blast } note si_arg = this show ?thesis proof (cases t) case t: (Var y) show ?thesis proof (cases "\i\?ss. wpo_ns (ss ! i) t") case True then obtain i where si: "i \ ?ss" and ns: "wpo_ns (ss ! i) t" unfolding s t by auto from si_arg[OF this] show ?thesis . next case False with ns[unfolded simps] s t have ssimple and largef: "large ?f" by (auto split: if_splits) from False s t not have "\ wpo_s s t" unfolding wpo.simps[of s t] by auto moreover have "wpo_ns ?s ?t" proof (cases "\ y") case (Var z) show ?thesis unfolding wpo.simps[of ?s ?t] not id unfolding s t using Var \ssimple\ largef by auto next case (Fun g ts) let ?g = "(g,length ts)" obtain ps pns where prc: "prc ?f ?g = (ps,pns)" by (cases "prc ?f ?g", auto) from prc_stri_imp_nstri[of ?f ?g] prc have ps: "ps \ pns" by auto { fix j assume "j \ set (\ ?g)" with set_status_nth[OF refl this] ss_status[OF \ssimple\ this] t Fun have "(t \ \, ts ! j) \ S" by (auto simp: simple_arg_pos_def) with sstsA have S: "(s \ \, ts ! j) \ S" by (metis compat_NS_S_point) hence "wpo_s (s \ \) (ts ! j)" by (rule S_imp_wpo_s) } note ssimple = this from large[OF \ssimple\ largef, of ?g, unfolded prc] have "ps \ pns \ \ ?g = []" by auto thus ?thesis using ssimple unfolding wpo.simps[of ?s ?t] not id unfolding s t using Fun prc ps by (auto simp: lex_ext_least_1 mul_ext_def Let_def ns_mul_ext_bottom) qed ultimately show ?thesis by blast qed next case (Fun g ts) note t = this let ?g = "(g,length ts)" let ?ts = "set (\ ?g)" obtain prs prns where p: "prc ?f ?g = (prs, prns)" by force note ns = ns[unfolded simps, unfolded s t p term.simps split] show ?thesis proof (cases "\ i \ ?ss. wpo_ns (ss ! i) t") case True with si_arg show ?thesis by blast next case False then have id: "(\ i \ ?ss. wpo_ns (ss ! i) (Fun g ts)) = False" unfolding t by auto note ns = ns[unfolded this if_False] let ?mss = "map (\ s . s \ \) ss" let ?mts = "map (\ t . t \ \) ts" from ns have prns and s_tj: "\ j. j \ ?ts \ wpo_s (Fun f ss) (ts ! j)" by (auto split: if_splits) { fix j assume j: "j \ ?ts" from \E[OF this] have "size s + size (ts ! j) < size s + size t" unfolding t by (auto simp: size_simps) from IH[OF this] s_tj[OF j, folded s] have wpo: "wpo_s ?s (ts ! j \ \)" by auto from j \[of g "length ts"] have "j < length ts" by auto with wpo have "wpo_s ?s (?mts ! j)" by auto } note ss_ts = this note \E = \E[of _ f ss] \E[of _ g ts] show ?thesis proof (cases prs) case True with ss_ts sstsA p \prns\ have "wpo_s ?s ?t" unfolding simps unfolding s t by (auto split: if_splits) with wpo_s_imp_ns[OF this] show ?thesis by blast next case False let ?mmss = "map ((!) ss) (\ ?f)" let ?mmts = "map ((!) ts) (\ ?g)" let ?Mmss = "map ((!) ?mss) (\ ?f)" let ?Mmts = "map ((!) ?mts) (\ ?g)" have id_map: "?Mmss = map (\ t. t \ \) ?mmss" "?Mmts = map (\ t. t \ \) ?mmts" unfolding map_map o_def by (auto simp: set_status_nth) let ?ls = "length (\ ?f)" let ?lt = "length (\ ?g)" { fix si tj assume *: "si \ set ?mmss" "tj \ set ?mmts" have "(wpo_s si tj \ wpo_s (si \ \) (tj \ \)) \ (wpo_ns si tj \ wpo_ns (si \ \) (tj \ \))" proof (intro IH add_strict_mono) from *(1) have "si \ set ss" using set_status_nth[of _ _ _ \\] by auto then show "size si < size s" unfolding s by (auto simp: termination_simp) from *(2) have "tj \ set ts" using set_status_nth[of _ _ _ \\] by auto then show "size tj < size t" unfolding t by (auto simp: termination_simp) qed hence "wpo_s si tj \ wpo_s (si \ \) (tj \ \)" "wpo_ns si tj \ wpo_ns (si \ \) (tj \ \)" by blast+ } note IH' = this { fix i assume "i < ?ls" "i < ?lt" then have i_f: "i < length (\ ?f)" and i_g: "i < length (\ ?g)" by auto with \[of f "length ss"] \[of g "length ts"] have i: "\ ?f ! i < length ss" "\ ?g ! i < length ts" unfolding set_conv_nth by auto then have "size (ss ! (\ ?f ! i)) < size s" "size (ts ! (\ ?g ! i)) < size t" unfolding s t by (auto simp: size_simps) then have "size (ss ! (\ ?f ! i)) + size (ts ! (\ ?g ! i)) < size s + size t" by simp from IH[OF this] i i_f i_g have "(wpo_s (?mmss ! i) (?mmts ! i) \ wpo_s (?mss ! (\ ?f ! i)) (?mts ! (\ ?g ! i)))" "(wpo_ns (?mmss ! i) (?mmts ! i) \ wpo_ns (?mss ! (\ ?f ! i)) (?mts ! (\ ?g ! i)))" by auto } note IH = this consider (Lex) "c ?f = Lex" "c ?g = Lex" | (Mul) "c ?f = Mul" "c ?g = Mul" | (Diff) "c ?f \ c ?g" by (cases "c ?f"; cases "c ?g", auto) thus ?thesis proof cases case Lex from Lex False ns have "snd (lex_ext wpo n ?mmss ?mmts)" by (auto split: if_splits) from this[unfolded lex_ext_iff snd_conv] have len: "(?ls = ?lt \ ?lt \ n)" and choice: "(\i< ?ls. i < ?lt \ (\j wpo_s (?mmss ! i) (?mmts ! i)) \ (\i< ?lt. wpo_ns (?mmss ! i) (?mmts ! i)) \ ?lt \ ?ls" (is "?stri \ ?nstri") by auto from choice have "?stri \ (\ ?stri \ ?nstri)" by blast then show ?thesis proof assume ?stri then obtain i where i: "i < ?ls" "i < ?lt" and NS: "(\jjprns\ have "wpo_s ?s ?t" unfolding simps unfolding s t by (auto split: if_splits) with wpo_s_imp_ns[OF this] show ?thesis by blast next assume "\ ?stri \ ?nstri" then have ?nstri and nstri: "\ ?stri" by blast+ with IH have "(\i< ?lt. wpo_ns (?Mmss ! i) (?Mmts ! i)) \ ?lt \ ?ls" by auto with len have "snd (lex_ext wpo n ?Mmss ?Mmts)" unfolding lex_ext_iff by auto with Lex ss_ts sstsA p \prns\ have ns: "wpo_ns ?s ?t" unfolding simps unfolding s t by (auto split: if_splits) { assume "wpo_s s t" from Lex this[unfolded simps, unfolded s t term.simps p split id] False have "fst (lex_ext wpo n ?mmss ?mmts)" by (auto split: if_splits) from this[unfolded lex_ext_iff fst_conv] nstri have "(\i< ?lt. wpo_ns (?mmss ! i) (?mmts ! i)) \ ?lt < ?ls" by auto with IH have "(\i< ?lt. wpo_ns (?Mmss ! i) (?Mmts ! i)) \ ?lt < ?ls" by auto then have "fst (lex_ext wpo n ?Mmss ?Mmts)" using len unfolding lex_ext_iff by auto with Lex ss_ts sstsA p \prns\ have ns: "wpo_s ?s ?t" unfolding simps unfolding s t by (auto split: if_splits) } with ns show ?thesis by blast qed next case Diff thus ?thesis using ns ss_ts sstsA p \prns\ unfolding simps unfolding s t by (auto simp: Let_def split: if_splits) next case Mul from Mul False ns have ge: "snd (mul_ext wpo ?mmss ?mmts)" by (auto split: if_splits) have ge: "snd (mul_ext wpo ?Mmss ?Mmts)" unfolding id_map by (rule nstri_mul_ext_map[OF _ _ ge], (intro IH', auto)+) { assume gr: "fst (mul_ext wpo ?mmss ?mmts)" have gr\: "fst (mul_ext wpo ?Mmss ?Mmts)" unfolding id_map by (rule stri_mul_ext_map[OF _ _ gr], (intro IH', auto)+) } note gr = this from ge gr show ?thesis using ss_ts \prns\ unfolding simps - unfolding s t term.simps p split subst_apply_term.simps length_map Mul + unfolding s t term.simps p split eval_term.simps length_map Mul by (simp add: id_map id) qed qed qed qed qed qed qed lemma subterm_wpo_s_arg': assumes i: "i \ set (\ (f,length ss))" shows "Fun f ss \ ss ! i" proof - have refl: "ss ! i \ ss ! i" by (rule wpo_ns_refl') with i have "\ t \ set (\ (f,length ss)). ss ! i \ ss ! i" by auto with NS_arg'[OF i] i show ?thesis by (auto simp: wpo.simps split: if_splits) qed context fixes f s t bef aft assumes ctxt_NS: "(s,t) \ NS \ (Fun f (bef @ s # aft), Fun f (bef @ t # aft)) \ NS" and SN: "\f. prl f \ SN S" begin lemma wpo_ns_pre_mono': defines "\f \ \ (f, Suc (length bef + length aft))" assumes rel: "(wpo_ns s t)" shows "(\j\set \f. Fun f (bef @ s # aft) \ (bef @ t # aft) ! j) \ (Fun f (bef @ s # aft), (Fun f (bef @ t # aft))) \ NS \ (\ i < length \f. ((map ((!) (bef @ s # aft)) \f) ! i) \ ((map ((!) (bef @ t # aft)) \f) ! i))" (is "_ \ _ \ ?three") proof - let ?ss = "bef @ s # aft" let ?ts = "bef @ t # aft" let ?s = "Fun f ?ss" let ?t = "Fun f ?ts" let ?len = "Suc (length bef + length aft)" let ?f = "(f, ?len)" let ?\ = "\ ?f" from wpo_ns_imp_NS[OF rel] have stA: "(s,t) \ NS" . have ?three unfolding \f_def proof (intro allI impI) fix i assume "i < length ?\" then have id: "\ ss. (map ((!) ss) ?\) ! i = ss ! (?\ ! i)" by auto show "wpo_ns ((map ((!) ?ss) ?\) ! i) ((map ((!) ?ts) ?\) ! i)" proof (cases "?\ ! i = length bef") case True then show ?thesis unfolding id using rel by auto next case False from append_Cons_nth_not_middle[OF this, of s aft t] wpo_ns_refl' show ?thesis unfolding id by auto qed qed have "\j\set ?\. wpo_s ?s ((bef @ t # aft) ! j)" (is ?one) proof fix j assume j: "j \ set ?\" then have "j \ set (\ (f,length ?ss))" by simp from subterm_wpo_s_arg'[OF this] have s: "wpo_s ?s (?ss ! j)" . show "wpo_s ?s (?ts ! j)" proof (cases "j = length bef") case False then have "?ss ! j = ?ts ! j" by (rule append_Cons_nth_not_middle) with s show ?thesis by simp next case True with s have "wpo_s ?s s" by simp with rel wpo_compat'[OF SN] have "wpo_s ?s t" by fast with True show ?thesis by simp qed qed with \?three\ ctxt_NS[OF stA] show ?thesis unfolding \f_def by auto qed lemma wpo_ns_mono': assumes rel: "s \ t" shows "Fun f (bef @ s # aft) \ Fun f (bef @ t # aft)" proof - let ?ss = "bef @ s # aft" let ?ts = "bef @ t # aft" let ?s = "Fun f ?ss" let ?t = "Fun f ?ts" let ?len = "Suc (length bef + length aft)" let ?f = "(f, ?len)" let ?\ = "\ ?f" from wpo_ns_pre_mono'[OF rel] have id: "(\j\set ?\. wpo_s ?s ((bef @ t # aft) ! j)) = True" "((?s,?t) \ NS) = True" "length ?ss = ?len" "length ?ts = ?len" by auto have "snd (lex_ext wpo n (map ((!) ?ss) ?\) (map ((!) ?ts) ?\))" by (rule all_nstri_imp_lex_nstri, intro allI impI, insert wpo_ns_pre_mono'[OF rel], auto) moreover have "snd (mul_ext wpo (map ((!) ?ss) ?\) (map ((!) ?ts) ?\))" by (rule all_nstri_imp_mul_nstri, intro allI impI, insert wpo_ns_pre_mono'[OF rel], auto) ultimately show ?thesis unfolding wpo.simps[of ?s ?t] term.simps id prc_refl using order_tag.exhaust by (auto simp: Let_def) qed end end end locale wpo_with_assms = wpo_with_basic_assms + SN_order_pair + precedence + constrains S :: "('f, 'v) term rel" and NS :: _ and prc :: "'f \ nat \ 'f \ nat \ bool \ bool" and prl :: "'f \ nat \ bool" and ssimple :: bool and large :: "'f \ nat \ bool" and c :: "'f \ nat \ order_tag" and n :: nat and \\ :: "'f status" assumes ctxt_NS: "(s,t) \ NS \ (Fun f (bef @ s # aft), Fun f (bef @ t # aft)) \ NS" and ws_status: "i \ set (status \\ fn) \ simple_arg_pos NS fn i" begin lemma ssimple: "strictly_simple_status \\ NS" using ws_status set_status_nth unfolding strictly_simple_status_def simple_arg_pos_def by fastforce lemma trans_prc: "trans_precedence prc" unfolding trans_precedence_def proof (intro allI, goal_cases) case (1 f g h) show ?case using prc_compat[of f g _ _ h] by (cases "prc f g"; cases "prc g h"; cases "prc f h", auto) qed lemma NS_arg: assumes i: "i \ set (\ (f,length ts))" shows "(Fun f ts, ts ! i) \ NS" using NS_arg'[OF ssimple i] . lemma NS_subterm: assumes all: "\ f k. set (\ (f,k)) = {0 ..< k}" shows "s \ t \ (s,t) \ NS" proof (induct s t rule: supteq.induct) case (refl) from refl_NS show ?case unfolding refl_on_def by blast next case (subt s ss t f) from subt(1) obtain i where i: "i < length ss" and s: "s = ss ! i" unfolding set_conv_nth by auto from NS_arg[of i f ss, unfolded all] s i have "(Fun f ss, s) \ NS" by auto from trans_NS_point[OF this subt(3)] show ?case . qed lemma wpo_ns_refl: "s \ s" using wpo_ns_refl'[OF ssimple] . lemma Var_not_S[simp]: "(Var x, t) \ S" proof assume st: "(Var x, t) \ S" from SN_imp_minimal[OF SN, rule_format, of undefined UNIV] obtain s where "\ u. (s,u) \ S" by blast with subst_S[OF st, of "\ _. s"] show False by auto qed (* Transitivity / compatibility of the orders *) lemma wpo_compat: " (s \ t \ t \ u \ s \ u) \ (s \ t \ t \ u \ s \ u) \ (s \ t \ t \ u \ s \ u)" by (rule wpo_compat'[OF SN]) lemma subterm_wpo_s_arg: assumes i: "i \ set (\ (f,length ss))" shows "Fun f ss \ ss ! i" by (rule subterm_wpo_s_arg'[OF ssimple i]) lemma subterm_wpo_ns_arg: assumes i: "i \ set (\ (f,length ss))" shows "Fun f ss \ ss ! i" by (rule wpo_s_imp_ns[OF subterm_wpo_s_arg[OF i]]) lemma wpo_ns_mono: assumes rel: "s \ t" shows "Fun f (bef @ s # aft) \ Fun f (bef @ t # aft)" by (rule wpo_ns_mono'[OF ssimple ctxt_NS SN rel]) lemma wpo_ns_pre_mono: fixes f and bef aft :: "('f,'v)term list" defines "\f \ \ (f, Suc (length bef + length aft))" assumes rel: "(wpo_ns s t)" shows "(\j\set \f. Fun f (bef @ s # aft) \ (bef @ t # aft) ! j) \ (Fun f (bef @ s # aft), (Fun f (bef @ t # aft))) \ NS \ (\ i < length \f. ((map ((!) (bef @ s # aft)) \f) ! i) \ ((map ((!) (bef @ t # aft)) \f) ! i))" unfolding \f_def by (rule wpo_ns_pre_mono'[OF ssimple ctxt_NS SN rel]) lemma wpo_stable: fixes \ :: "('f,'v)subst" shows "(s \ t \ s \ \ \ t \ \) \ (s \ t \ s \ \ \ t \ \)" by (rule wpo_stable'[OF ssimple]) lemma WPO_S_SN: "SN WPO_S" proof - { fix t :: "('f,'v)term" let ?S = "\x. SN_on WPO_S {x}" note iff = SN_on_all_reducts_SN_on_conv[of WPO_S] { fix x have "?S (Var x)" unfolding iff[of "Var x"] proof (intro allI impI) fix s assume "(Var x, s) \ WPO_S" then have False by (cases s, auto simp: wpo.simps split: if_splits) then show "?S s" .. qed } note var_SN = this have "?S t" proof (induct t) case (Var x) show ?case by (rule var_SN) next case (Fun f ts) let ?Slist = "\ f ys. \ i \ set (\ f). ?S (ys ! i)" let ?r3 = "{((f,ab), (g,ab')). ((c f = c g) \ (?Slist f ab \ (c f = Mul \ fst (mul_ext wpo (map ((!) ab) (\ f)) (map ((!) ab') (\ g)))) \ (c f = Lex \ fst (lex_ext wpo n (map ((!) ab) (\ f)) (map ((!) ab') (\ g)))))) \ ((c f \ c g) \ (map ((!) ab) (\ f) \ [] \ (map ((!) ab') (\ g)) = []))}" let ?r0 = "lex_two {(f,g). fst (prc f g)} {(f,g). snd (prc f g)} ?r3" { fix ab { assume "\S. S 0 = ab \ (\i. (S i, S (Suc i)) \ ?r3)" then obtain S where S0: "S 0 = ab" and SS: "\i. (S i, S (Suc i)) \ ?r3" by auto let ?Sf = "\i. fst (fst (S i))" let ?Sn = "\i. snd (fst (S i))" let ?Sfn = "\ i. fst (S i)" let ?Sts = "\i. snd (S i)" let ?Sts\ = "\i. map ((!) (?Sts i)) (\ (?Sfn i))" have False proof (cases "\i. c (?Sfn i) = Mul") case True let ?r' = "{((f,ys), (g,xs)). (\ yi \set ((map ((!) ys) (\ f))). SN_on WPO_S {yi}) \ fst (mul_ext wpo (map ((!) ys) (\ f)) (map ((!) xs) (\ g)))}" { fix i from True[rule_format, of i] and True[rule_format, of "Suc i"] and SS[rule_format, of i] have "(S i, S (Suc i)) \ ?r'" by auto } then have Hf: "\ SN_on ?r' {S 0}" unfolding SN_on_def by auto from mul_ext_SN[of wpo, rule_format, OF wpo_ns_refl] and wpo_compat wpo_s_imp_ns have tmp: "SN {(ys, xs). (\y\set ys. SN_on {(s, t). wpo_s s t} {y}) \ fst (mul_ext wpo ys xs)}" (is "SN ?R") by blast have id: "?r' = inv_image ?R (\ (f,ys). map ((!) ys) (\ f))" by auto from SN_inv_image[OF tmp] have "SN ?r'" unfolding id . from SN_on_subset2[OF subset_UNIV[of "{S 0}"], OF this] have "SN_on ?r' {(S 0)}" . with Hf show ?thesis .. next case False note HMul = this show ?thesis proof (cases "\i. c (?Sfn i) = Lex") case True let ?r' = "{((f,ys), (g,xs)). (\ yi \set ((map ((!) ys) (\ f))). SN_on WPO_S {yi}) \ fst (lex_ext wpo n (map ((!) ys) (\ f)) (map ((!) xs) (\ g)))}" { fix i from SS[rule_format, of i] True[rule_format, of i] True[rule_format, of "Suc i"] have "(S i, S (Suc i)) \ ?r'" by auto } then have Hf: "\ SN_on ?r' {S 0}" unfolding SN_on_def by auto from wpo_compat have "\ x y z. wpo_ns x y \ wpo_s y z \ wpo_s x z" by blast from lex_ext_SN[of "wpo" n, OF this] have tmp: "SN {(ys, xs). (\y\set ys. SN_on WPO_S {y}) \ fst (lex_ext wpo n ys xs)}" (is "SN ?R") . have id: "?r' = inv_image ?R (\ (f,ys). map ((!) ys) (\ f))" by auto from SN_inv_image[OF tmp] have "SN ?r'" unfolding id . then have "SN_on ?r' {(S 0)}" unfolding SN_defs by blast with Hf show False .. next case False note HLex = this from HMul and HLex have "\i. c (?Sfn i) \ c (?Sfn (Suc i))" proof (cases ?thesis, simp) case False then have T: "\i. c (?Sfn i) = c (?Sfn (Suc i))" by simp { fix i have "c (?Sfn i) = c (?Sfn 0)" proof (induct i) case (Suc j) then show ?case by (simp add: T[rule_format, of j]) qed simp } then show ?thesis using HMul HLex by (cases "c (?Sfn 0)") auto qed then obtain i where Hdiff: "c (?Sfn i) \ c (?Sfn (Suc i))" by auto from Hdiff have Hf: "?Sts\ (Suc i) = []" using SS[rule_format, of i] by auto show ?thesis proof (cases "c (?Sfn (Suc i)) = c (?Sfn (Suc (Suc i)))") case False then show ?thesis using Hf and SS[rule_format, of "Suc i"] by auto next case True show ?thesis proof (cases "c (?Sfn (Suc i))") case Mul with True and SS[rule_format, of "Suc i"] have "fst (mul_ext wpo (?Sts\ (Suc i)) (?Sts\ (Suc (Suc i))))" by auto with Hf and s_mul_ext_bottom_strict show ?thesis by (simp add: Let_def mul_ext_def s_mul_ext_bottom_strict) next case Lex with True and SS[rule_format, of "Suc i"] have "fst (lex_ext wpo n (?Sts\ (Suc i)) (?Sts\ (Suc (Suc i))))" by auto with Hf show ?thesis by (simp add: lex_ext_iff) qed qed qed qed } } then have "SN ?r3" unfolding SN_on_def by blast have SN1:"SN ?r0" proof (rule lex_two[OF _ prc_SN \SN ?r3\]) let ?r' = "{(f,g). fst (prc f g)}" let ?r = "{(f,g). snd (prc f g)}" { fix a1 a2 a3 assume "(a1,a2) \ ?r" "(a2,a3) \ ?r'" then have "(a1,a3) \ ?r'" by (cases "prc a1 a2", cases "prc a2 a3", cases "prc a1 a3", insert prc_compat[of a1 a2 _ _ a3], force) } then show "?r O ?r' \ ?r'" by auto qed let ?m = "\ (f,ts). ((f,length ts), ((f, length ts), ts))" let ?r = "{(a,b). (?m a, ?m b) \ ?r0}" have SN_r: "SN ?r" using SN_inv_image[OF SN1, of ?m] unfolding inv_image_def by fast let ?SA = "(\ x y. (x,y) \ S)" let ?NSA = "(\ x y. (x,y) \ NS)" let ?rr = "lex_two S NS ?r" define rr where "rr = ?rr" from lex_two[OF compat_NS_S SN SN_r] have SN_rr: "SN rr" unfolding rr_def by auto let ?rrr = "inv_image rr (\ (f,ts). (Fun f ts, (f,ts)))" have SN_rrr: "SN ?rrr" by (rule SN_inv_image[OF SN_rr]) let ?ind = "\ (f,ts). ?Slist (f,length ts) ts \ ?S (Fun f ts)" have "?ind (f,ts)" proof (rule SN_induct[OF SN_rrr, of ?ind]) fix fts assume ind: "\ gss. (fts,gss) \ ?rrr \ ?ind gss" obtain f ts where Pair: "fts = (f,ts)" by force let ?f = "(f,length ts)" note ind = ind[unfolded Pair] show "?ind fts" unfolding Pair split proof (intro impI) assume ts: "?Slist ?f ts" let ?t = "Fun f ts" show "?S ?t" proof (simp only: iff[of ?t], intro allI impI) fix s assume "(?t,s) \ WPO_S" then have "?t \ s" by simp then show "?S s" proof (induct s, simp add: var_SN) case (Fun g ss) note IH = this let ?s = "Fun g ss" let ?g = "(g,length ss)" from Fun have t_gr_s: "?t \ ?s" by auto show "?S ?s" proof (cases "\ i \ set (\ ?f). ts ! i \ ?s") case True then obtain i where "i \ set (\ ?f)" and ge: "ts ! i \ ?s" by auto with ts have "?S (ts ! i)" by auto show "?S ?s" proof (unfold iff[of ?s], intro allI impI) fix u assume "(?s,u) \ WPO_S" with wpo_compat ge have u: "ts ! i \ u" by blast with \?S (ts ! i)\[unfolded iff[of "ts ! i"]] show "?S u" by simp qed next case False note oFalse = this from wpo_s_imp_NS[OF t_gr_s] have t_NS_s: "(?t,?s) \ NS" . show ?thesis proof (cases "(?t,?s) \ S") case True then have "((f,ts),(g,ss)) \ ?rrr" unfolding rr_def by auto with ind have ind: "?ind (g,ss)" by auto { fix i assume i: "i \ set (\ ?g)" have "?s \ ss ! i" by (rule subterm_wpo_ns_arg[OF i]) with t_gr_s have ts: "?t \ ss ! i" using wpo_compat by blast have "?S (ss ! i)" using IH(1)[OF \E[OF i] ts] by auto } note SN_ss = this from ind SN_ss show ?thesis by auto next case False with t_NS_s oFalse have id: "(?t,?s) \ S = False" "(?t,?s) \ NS = True" by simp_all let ?ls = "length ss" let ?lt = "length ts" let ?f = "(f,?lt)" let ?g = "(g,?ls)" obtain s1 ns1 where prc1: "prc ?f ?g = (s1,ns1)" by force note t_gr_s = t_gr_s[unfolded wpo.simps[of ?t ?s], unfolded term.simps id if_True if_False prc1 split] from oFalse t_gr_s have f_ge_g: "ns1" by (cases ?thesis, auto) from oFalse t_gr_s f_ge_g have small_ss: "\ i \ set (\ ?g). ?t \ ss ! i" by (cases ?thesis, auto) with Fun \E[of _ g ss] have ss_S: "?Slist ?g ss" by auto show ?thesis proof (cases s1) case True then have "((f,ts),(g,ss)) \ ?r" by (simp add: prc1) with t_NS_s have "((f,ts),(g,ss)) \ ?rrr" unfolding rr_def by auto with ind have "?ind (g,ss)" by auto with ss_S show ?thesis by auto next case False consider (Diff) "c ?f \ c ?g" | (Lex) "c ?f = Lex" "c ?g = Lex" | (Mul) "c ?f = Mul" "c ?g = Mul" by (cases "c ?f"; cases "c ?g", auto) thus ?thesis proof cases case Diff with False oFalse f_ge_g t_gr_s small_ss prc1 t_NS_s have "((f,ts),(g,ss)) \ ?rrr" unfolding rr_def by (cases "c ?f"; cases "c ?g", auto) with ind have "?ind (g,ss)" using Pair by auto with ss_S show ?thesis by simp next case Lex from False oFalse t_gr_s small_ss f_ge_g Lex have lex: "fst (lex_ext wpo n (map ((!) ts) (\ ?f)) (map ((!) ss) (\ ?g)))" by auto from False lex ts f_ge_g Lex have "((f,ts),(g,ss)) \ ?r" by (simp add: prc1) with t_NS_s have "((f,ts),(g,ss)) \ ?rrr" unfolding rr_def by auto with ind have "?ind (g,ss)" by auto with ss_S show ?thesis by auto next case Mul from False oFalse t_gr_s small_ss f_ge_g Mul have mul: "fst (mul_ext wpo (map ((!) ts) (\ ?f)) (map ((!) ss) (\ ?g)))" by auto from False mul ts f_ge_g Mul have "((f,ts),(g,ss)) \ ?r" by (simp add: prc1) with t_NS_s have "((f,ts),(g,ss)) \ ?rrr" unfolding rr_def by auto with ind have "?ind (g,ss)" by auto with ss_S show ?thesis by auto qed qed qed qed qed qed qed qed with Fun show ?case using \E[of _ f ts] by simp qed } from SN_I[OF this] show "SN {(s::('f, 'v)term, t). fst (wpo s t)}" . qed theorem WPO_SN_order_pair: "SN_order_pair WPO_S WPO_NS" proof show "refl WPO_NS" using wpo_ns_refl unfolding refl_on_def by auto show "trans WPO_NS" using wpo_compat unfolding trans_def by blast show "trans WPO_S" using wpo_compat wpo_s_imp_ns unfolding trans_def by blast show "WPO_NS O WPO_S \ WPO_S" using wpo_compat by blast show "WPO_S O WPO_NS \ WPO_S" using wpo_compat by blast show "SN WPO_S" using WPO_S_SN . qed theorem WPO_S_subst: "(s,t) \ WPO_S \ (s \ \, t \ \) \ WPO_S" for \ using wpo_stable by auto theorem WPO_NS_subst: "(s,t) \ WPO_NS \ (s \ \, t \ \) \ WPO_NS" for \ using wpo_stable by auto theorem WPO_NS_ctxt: "(s,t) \ WPO_NS \ (Fun f (bef @ s # aft), Fun f (bef @ t # aft)) \ WPO_NS" using wpo_ns_mono by blast theorem WPO_S_subset_WPO_NS: "WPO_S \ WPO_NS" using wpo_s_imp_ns by blast context (* if \ is the full status, then WPO_S is a simplification order *) assumes \_full: "\ f k. set (\ (f,k)) = {0 ..< k}" begin lemma subterm_wpo_s: "s \ t \ s \ t" proof (induct s t rule: supt.induct) case (arg s ss f) from arg[unfolded set_conv_nth] obtain i where i: "i < length ss" and s: "s = ss ! i" by auto from \_full[of f "length ss"] i have ii: "i \ set (\ (f,length ss))" by auto from subterm_wpo_s_arg[OF ii] s show ?case by auto next case (subt s ss t f) from subt wpo_s_imp_ns have "\ s \ set ss. wpo_ns s t" by blast from this[unfolded set_conv_nth] obtain i where ns: "ss ! i \ t" and i: "i < length ss" by auto from \_full[of f "length ss"] i have ii: "i \ set (\ (f,length ss))" by auto from subt have "Fun f ss \ t" by auto from NS_subterm[OF \_full this] ns ii show ?case by (auto simp: wpo.simps split: if_splits) qed (* Compatibility of the subterm relation with the order relation: a subterm is smaller *) lemma subterm_wpo_ns: assumes supteq: "s \ t" shows "s \ t" proof - from supteq have "s = t \ s \ t" by auto then show ?thesis proof assume "s = t" then show ?thesis using wpo_ns_refl by blast next assume "s \ t" from wpo_s_imp_ns[OF subterm_wpo_s[OF this]] show ?thesis . qed qed lemma wpo_s_mono: assumes rels: "s \ t" shows "Fun f (bef @ s # aft) \ Fun f (bef @ t # aft)" proof - let ?ss = "bef @ s # aft" let ?ts = "bef @ t # aft" let ?s = "Fun f ?ss" let ?t = "Fun f ?ts" let ?len = "Suc (length bef + length aft)" let ?f = "(f, ?len)" let ?\ = "\ ?f" from wpo_s_imp_ns[OF rels] have rel: "wpo_ns s t" . from wpo_ns_pre_mono[OF rel] have id: "(\j\set ?\. wpo_s ?s ((bef @ t # aft) ! j)) = True" "((?s,?t) \ NS) = True" "length ?ss = ?len" "length ?ts = ?len" by auto let ?lb = "length bef" from \_full[of f ?len] have lb_mem: "?lb \ set ?\" by auto then obtain i where \i: "?\ ! i = ?lb" and i: "i < length ?\" unfolding set_conv_nth by force let ?mss = "map ((!) ?ss) ?\" let ?mts = "map ((!) ?ts) ?\" have "fst (lex_ext wpo n ?mss ?mts)" unfolding lex_ext_iff fst_conv proof (intro conjI, force, rule disjI1, unfold length_map id, intro exI conjI, rule i, rule i, intro allI impI) show "wpo_s (?mss ! i) (?mts ! i)" using \i i rels by simp next fix j assume "j < i" with i have j: "j < length ?\" by auto with wpo_ns_pre_mono[OF rel] show "?mss ! j \ ?mts ! j" by blast qed moreover obtain lb nlb where part: "partition ((=) ?lb) ?\ = (lb, nlb)" by force hence mset_\: "mset ?\ = mset lb + mset nlb" by (induct ?\, auto) let ?mlbs = "map ((!) ?ss) lb" let ?mnlbs = "map ((!) ?ss) nlb" let ?mlbt = "map ((!) ?ts) lb" let ?mnlbt = "map ((!) ?ts) nlb" have id1: "mset ?mss = mset ?mnlbs + mset ?mlbs" using mset_\ by auto have id2: "mset ?mts = mset ?mnlbt + mset ?mlbt" using mset_\ by auto from part lb_mem have lb: "?lb \ set lb" by auto have "fst (mul_ext wpo ?mss ?mts)" unfolding mul_ext_def Let_def fst_conv proof (intro s_mul_extI_old, rule id1, rule id2) from lb show "mset ?mlbs \ {#}" by auto { fix i assume "i < length ?mnlbt" then obtain j where id: "?mnlbs ! i = ?ss ! j" "?mnlbt ! i = ?ts ! j" "j \ set nlb" by auto with part have "j \ ?lb" by auto hence "?ss ! j = ?ts ! j" by (auto simp: nth_append) thus "(?mnlbs ! i, ?mnlbt ! i) \ WPO_NS" unfolding id using wpo_ns_refl by auto } fix u assume "u \# mset ?mlbt" hence "u = t" using part by auto moreover have "s \# mset ?mlbs" using lb by force ultimately show "\ v. v \# mset ?mlbs \ (v,u) \ WPO_S" using rels by force qed auto ultimately show ?thesis unfolding wpo.simps[of ?s ?t] term.simps id prc_refl using order_tag.exhaust by (auto simp: Let_def) qed theorem WPO_S_ctxt: "(s,t) \ WPO_S \ (Fun f (bef @ s # aft), Fun f (bef @ t # aft)) \ WPO_S" using wpo_s_mono by blast theorem supt_subset_WPO_S: "{\} \ WPO_S" using subterm_wpo_s by blast theorem supteq_subset_WPO_NS: "{\} \ WPO_NS" using subterm_wpo_ns by blast end end end