(* 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⇩f::"'fun ⇒ nat" and arity⇩s::"'sets ⇒ nat" and public⇩f::"'fun ⇒ bool" and Ana⇩f::"'fun ⇒ ((('fun,'atom::finite,'sets,'lbl) prot_fun, nat) term list × nat list)" and Γ⇩f::"'fun ⇒ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" assumes Ana⇩f_assm1: "∀f. let (K, M) = Ana⇩f f in (∀k ∈ subterms⇩s⇩e⇩t (set K). is_Fun k ⟶ (is_Fu (the_Fun k)) ∧ length (args k) = arity⇩f (the_Fu (the_Fun k)))" and Ana⇩f_assm2: "∀f. let (K, M) = Ana⇩f f in ∀i ∈ fv⇩s⇩e⇩t (set K) ∪ set M. i < arity⇩f f" and public⇩f_assm: "∀f. arity⇩f f > (0::nat) ⟶ public⇩f f" and Γ⇩f_assm: "∀f. arity⇩f f = (0::nat) ⟶ Γ⇩f f ≠ None" and label_witness_assm: "label_witness1 ≠ label_witness2" begin lemma Ana⇩f_assm1_alt: assumes "Ana⇩f f = (K,M)" "k ∈ subterms⇩s⇩e⇩t (set K)" shows "(∃x. k = Var x) ∨ (∃h T. k = Fun (Fu h) T ∧ length T = arity⇩f h)" proof (cases k) case (Fun g T) let ?P = "λk. is_Fun k ⟶ is_Fu (the_Fun k) ∧ length (args k) = arity⇩f (the_Fu (the_Fun k))" let ?Q = "λK M. ∀k ∈ subterms⇩s⇩e⇩t (set K). ?P k" have "?Q (fst (Ana⇩f f)) (snd (Ana⇩f f))" using Ana⇩f_assm1 split_beta[of ?Q "Ana⇩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⇩f_assm2_alt: assumes "Ana⇩f f = (K,M)" "i ∈ fv⇩s⇩e⇩t (set K) ∪ set M" shows "i < arity⇩f f" using Ana⇩f_assm2 assms by fastforce subsection ‹Definitions› fun arity where "arity (Fu f) = arity⇩f f" | "arity (Set s) = arity⇩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⇩f f" | "public (Set s) = (arity⇩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⇩f f = length T ∧ arity⇩f f > 0 then let (K,M) = Ana⇩f f in (K ⋅⇩l⇩i⇩s⇩t (!) T, map ((!) T) M) else ([], []))" | "Ana _ = ([], [])" definition Γ⇩v where "Γ⇩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) = Γ⇩v v" | "Γ (Fun f T) = ( if arity f = 0 then case f of (Fu g) ⇒ TAtom (case Γ⇩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⇩f g = 0 ⟹ Γ (Fun (Fu g) []::('fun,'atom,'sets,'lbl) prot_term) = TAtom (case Γ⇩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⇩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⇩f f ≠ 0 ⟹ Γ (Fun (Fu f) T) = TComp (Fu f) (map Γ T)" (is "?A1 ⟹ ?A2") "arity⇩f f = 0 ⟹ Γ⇩f f = Some a ⟹ Γ (Fun (Fu f) T) = TAtom (Atom a)" (is "?B1 ⟹ ?B2 ⟹ ?B3") "arity⇩f f = 0 ⟹ Γ⇩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 "Γ⇩f f") simp_all show ?E by (cases "Γ⇩f f") simp_all show ?F by (cases "Γ⇩f f") simp_all qed lemma Γ_Set_simps[simp]: "arity⇩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⇩f g = (K',M')" shows "(K,M) = (if arity⇩f g = length T ∧ arity⇩f g > 0 then (K' ⋅⇩l⇩i⇩s⇩t (!) T, map ((!) T) M') else ([],[]))" (is ?A) and "(K,M) = (K' ⋅⇩l⇩i⇩s⇩t (!) T, map ((!) T) M') ∨ (K,M) = ([],[])" (is ?B) proof - show ?A using assms by (cases "arity⇩f g = length T ∧ arity⇩f g > 0") auto thus ?B by metis qed lemma Ana_Fu_intro: assumes "arity⇩f f = length T" "arity⇩f f > 0" and "Ana⇩f f = (K',M')" shows "Ana (Fun (Fu f) T) = (K' ⋅⇩l⇩i⇩s⇩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⇩f g = (K',M')" and "(K,M) ≠ ([],[])" shows "arity⇩f g = length T" (is ?A) and "(K,M) = (K' ⋅⇩l⇩i⇩s⇩t (!) T, map ((!) T) M')" (is ?B) proof - show ?A using assms by force moreover have "arity⇩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⇩f f = length T ∧ arity⇩f f > 0 ∧ (∃K M. Ana⇩f f = (K, M) ∧ Ana t = (K ⋅⇩l⇩i⇩s⇩t (!) T, map ((!) T) M))" using assms proof (induction t rule: Ana.induct) case (1 f T) hence *: "arity⇩f f = length T" "0 < arity⇩f f" "Ana (Fun (Fu f) T) = (case Ana⇩f f of (K, M) ⇒ (K ⋅⇩l⇩i⇩s⇩t (!) T, map ((!) T) M))" using Ana.simps(1)[of f T] unfolding Let_def by metis+ obtain K M where **: "Ana⇩f f = (K, M)" by (metis surj_pair) hence "Ana (Fun (Fu f) T) = (K ⋅⇩l⇩i⇩s⇩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⇩s⇩e⇩t (set K) ⊆ fv t" using assms proof (induction t rule: term.induct) case (Fun f T) have aux: "fv⇩s⇩e⇩t (set K ⋅⇩s⇩e⇩t (!) T) ⊆ fv⇩s⇩e⇩t (set T)" when K: "∀i ∈ fv⇩s⇩e⇩t (set K). i < length T" for K::"(('fun,'atom,'sets,'lbl) prot_fun, nat) term list" proof fix x assume "x ∈ fv⇩s⇩e⇩t (set K ⋅⇩s⇩e⇩t (!) T)" then obtain k where k: "k ∈ set K" "x ∈ fv (k ⋅ (!) T)" by auto have "∀i ∈ fv k. i < length T" using K k(1) by simp thus "x ∈ fv⇩s⇩e⇩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⇩f g = (K',M')" by force have "(K, M) ≠ ([], [])" using K by simp hence "(K, M) = (K' ⋅⇩l⇩i⇩s⇩t (!) T, map ((!) T) M')" "arity⇩f g = length T" using Ana_Fu_cases(1)[OF Fun.prems f *] by presburger+ hence ?case using aux[of K'] Ana⇩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⇩f h = (K',M')" by force have "(K,M) ≠ ([],[])" using Fun.prems(3) by auto hence "(K,M) = (K' ⋅⇩l⇩i⇩s⇩t (!) T, map ((!) T) M')" "⋀i. i ∈ fv⇩s⇩e⇩t (set K') ∪ set M' ⟹ i < length T" using Ana_Fu_cases(1)[OF Fun.prems(1) 2 1] Ana⇩f_assm2_alt[OF 1] by presburger+ hence "K = K' ⋅⇩l⇩i⇩s⇩t (!) T" and 3: "∀i∈fv⇩s⇩e⇩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 force hence 4: "Fun f T' ∈ subterms (k' ⋅ (!) T)" "fv k' ⊆ fv⇩s⇩e⇩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⇩s⇩e⇩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⇩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⇩f g = (K',M')" by force have "M = [] ∨ (arity⇩f g = length T ∧ M = map ((!) T) M')" using Ana_Fu_cases(1)[OF assms Fu *] by (meson prod.inject) thus ?thesis using Ana⇩f_assm2_alt[OF *] by auto qed auto lemma assm5: "Ana t = (K,M) ⟹ K ≠ [] ∨ M ≠ [] ⟹ Ana (t ⋅ δ) = (K ⋅⇩l⇩i⇩s⇩t δ, M ⋅⇩l⇩i⇩s⇩t δ)" proof (induction t rule: term.induct) case (Fun f T) thus ?case proof (cases f) case (Fu g) obtain K' M' where *: "Ana⇩f g = (K',M')" by force have **: "K = K' ⋅⇩l⇩i⇩s⇩t (!) T" "M = map ((!) T) M'" "arity⇩f g = length T" "∀i ∈ fv⇩s⇩e⇩t (set K') ∪ set M'. i < arity⇩f g" "0 < arity⇩f g" using Fun.prems(2) Ana_Fu_cases(1)[OF Fun.prems(1) Fu *] Ana⇩f_assm2_alt[OF *] by (meson prod.inject)+ have ***: "∀i ∈ fv⇩s⇩e⇩t (set K'). i < length T" "∀i ∈ set M'. i < length T" using **(3,4) by auto have "K ⋅⇩l⇩i⇩s⇩t δ = K' ⋅⇩l⇩i⇩s⇩t (!) (map (λt. t ⋅ δ) T)" "M ⋅⇩l⇩i⇩s⇩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 Γ⇩v_def by meson thus ?case using Var by (fastforce simp add: Γ⇩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⇩t⇩r⇩m (Γ (Var x))" unfolding wf⇩t⇩r⇩m_def by (auto simp add: Γ⇩v_def) lemma assm11: "arity f > 0 ⟹ public f" using public⇩f_assm by (cases f) auto lemma assm12: "Γ (Var (τ, n)) = Γ (Var (τ, m))" by (simp add: Γ⇩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 ⋅⇩l⇩i⇩s⇩t δ, M ⋅⇩l⇩i⇩s⇩t δ)" proof - show ?thesis proof (cases "(K, M) = ([],[])") case True { fix g assume f: "f = Fu g" obtain K' M' where "Ana⇩f g = (K',M')" by force 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⇩f g = (K',M')" by force have ***: "K = K' ⋅⇩l⇩i⇩s⇩t (!) T" "M = map ((!) T) M'" "arity⇩f g = length T" "∀i ∈ fv⇩s⇩e⇩t (set K') ∪ set M'. i < arity⇩f g" using Ana_Fu_cases(1)[OF assms ** *] False Ana⇩f_assm2_alt[OF *] by (meson prod.inject)+ have ****: "∀i∈fv⇩s⇩e⇩t (set K'). i < length T" "∀i∈set M'. i < length T" using ***(3,4) by auto have "K ⋅⇩l⇩i⇩s⇩t δ = K' ⋅⇩l⇩i⇩s⇩t (!) (map (λt. t ⋅ δ) T)" "M ⋅⇩l⇩i⇩s⇩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⇩s⇩u⇩b⇩s⇩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⇩s⇩e⇩t M) ∧ (∀t ∈ subst_range σ. t ∉ subterms⇩s⇩e⇩t (trms_transaction T)) ∧ wt⇩s⇩u⇩b⇩s⇩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⇩s⇩u⇩b⇩s⇩t ℐ ∧ wf⇩t⇩r⇩m⇩s (subst_range ℐ)" definition (in typed_model) welltyped_constraint_model where "welltyped_constraint_model ℐ 𝒜 ≡ wt⇩s⇩u⇩b⇩s⇩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⇩l⇩s⇩s⇩t 𝒜); transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t 𝒜) ⟧ ⟹ 𝒜@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α) ∈ reachable_constraints P" subsection ‹Minor Lemmata› lemma Γ⇩v_TAtom[simp]: "Γ⇩v (TAtom a, n) = TAtom a" unfolding Γ⇩v_def by simp lemma Γ⇩v_TAtom': assumes "a ≠ Bottom" shows "Γ⇩v (τ, n) = TAtom a ⟷ τ = TAtom a" proof assume "Γ⇩v (τ, n) = TAtom a" thus "τ = TAtom a" by (metis (no_types, lifting) assms Γ⇩v_def fst_conv term.inject(1)) qed simp lemma Γ⇩v_TAtom_inv: "Γ⇩v x = TAtom (Atom a) ⟹ ∃m. x = (TAtom (Atom a), m)" "Γ⇩v x = TAtom Value ⟹ ∃m. x = (TAtom Value, m)" "Γ⇩v x = TAtom SetType ⟹ ∃m. x = (TAtom SetType, m)" "Γ⇩v x = TAtom AttackType ⟹ ∃m. x = (TAtom AttackType, m)" "Γ⇩v x = TAtom OccursSecType ⟹ ∃m. x = (TAtom OccursSecType, m)" by (metis Γ⇩v_TAtom' surj_pair prot_atom.distinct(7), metis Γ⇩v_TAtom' surj_pair prot_atom.distinct(18), metis Γ⇩v_TAtom' surj_pair prot_atom.distinct(26), metis Γ⇩v_TAtom' surj_pair prot_atom.distinct(32), metis Γ⇩v_TAtom' surj_pair prot_atom.distinct(38)) lemma Γ⇩v_TAtom'': "(fst x = TAtom (Atom a)) = (Γ⇩v x = TAtom (Atom a))" (is "?A = ?A'") "(fst x = TAtom Value) = (Γ⇩v x = TAtom Value)" (is "?B = ?B'") "(fst x = TAtom SetType) = (Γ⇩v x = TAtom SetType)" (is "?C = ?C'") "(fst x = TAtom AttackType) = (Γ⇩v x = TAtom AttackType)" (is "?D = ?D'") "(fst x = TAtom OccursSecType) = (Γ⇩v x = TAtom OccursSecType)" (is "?E = ?E'") proof - have 1: "?A ⟹ ?A'" "?B ⟹ ?B'" "?C ⟹ ?C'" "?D ⟹ ?D'" "?E ⟹ ?E'" by (metis Γ⇩v_TAtom prod.collapse)+ have 2: "?A' ⟹ ?A" "?B' ⟹ ?B" "?C' ⟹ ?C" "?D' ⟹ ?D" "?E' ⟹ ?E" using Γ⇩v_TAtom Γ⇩v_TAtom_inv(1) apply fastforce using Γ⇩v_TAtom Γ⇩v_TAtom_inv(2) apply fastforce using Γ⇩v_TAtom Γ⇩v_TAtom_inv(3) apply fastforce using Γ⇩v_TAtom Γ⇩v_TAtom_inv(4) apply fastforce using Γ⇩v_TAtom Γ⇩v_TAtom_inv(5) by fastforce show "?A = ?A'" "?B = ?B'" "?C = ?C'" "?D = ?D'" "?E = ?E'" using 1 2 by metis+ qed lemma Γ⇩v_Var_image: "Γ⇩v ` X = Γ ` Var ` X" by force lemma Γ_Fu_const: assumes "arity⇩f g = 0" shows "∃a. Γ (Fun (Fu g) T) = TAtom (Atom a)" proof - have "Γ⇩f g ≠ None" using assms Γ⇩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⇩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⇩s s = 0" using * by simp hence False using Set assms by auto thus ?thesis by metis qed simp_all qed lemma Ana⇩f_keys_not_val_terms: assumes "Ana⇩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⇩s⇩e⇩t (set K)" using assms(2) funs_term_Fun_subterm[OF assms(3)] by (cases g) auto hence False using Ana⇩f_assm1_alt[OF assms(1) *] by simp } moreover { assume "is_PubConstValue g" then obtain n S where *: "Fun (PubConst Value n) S ∈ subterms⇩s⇩e⇩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⇩f_assm1_alt[OF assms(1) *] by simp } moreover { assume "is_Abs g" then obtain a S where *: "Fun (Abs a) S ∈ subterms⇩s⇩e⇩t (set K)" using assms(2) funs_term_Fun_subterm[OF assms(3)] by (cases g) auto hence False using Ana⇩f_assm1_alt[OF assms(1) *] by simp } ultimately show "¬is_Val g" "¬is_PubConstValue g" "¬is_Abs g" by metis+ qed lemma Ana⇩f_keys_not_pairs: assumes "Ana⇩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⇩s⇩e⇩t (set K)" using assms(2) funs_term_Fun_subterm[OF assms(3)] by (cases g) auto show False using Ana⇩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⇩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⇩f f = length S" "subterms k' ⊆ subterms⇩s⇩e⇩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⇩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⇩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⇩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⇩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⇩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⇩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⇩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⇩f f = length T ∧ arity⇩f f > 0") case True thus ?thesis using assms by (cases "Ana⇩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⇩f g = (K', R')" by fastforce have K: "K = K' ⋅⇩l⇩i⇩s⇩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⇩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⇩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⇩s⇩e⇩t M. OccursFact ∉ ⋃(funs_term ` set (snd (Ana s)))" "∀s ∈ subterms⇩s⇩e⇩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 ⋅⇩s⇩e⇩t θ ⊢ t ⋅ θ" and "∀x ∈ fv⇩s⇩e⇩t M ∪ fv t. (∃n. θ x = Fun (Val n) []) ∨ (∃n. θ x = Fun (PubConst Value n) [])" and "∀x ∈ fv⇩s⇩e⇩t M ∪ fv t. (∃n. σ x = Fun (Val n) [])" and "∀x ∈ fv⇩s⇩e⇩t M ∪ fv t. (∃n. θ x = Fun (PubConst Value n) []) ⟶ σ x ∈ M ∪ N" and "∀x ∈ fv⇩s⇩e⇩t M ∪ fv t. (∃n. θ x = Fun (Val n) []) ⟶ θ x = σ x" and "∀x ∈ fv⇩s⇩e⇩t M ∪ fv t. ∀y ∈ fv⇩s⇩e⇩t M ∪ fv t. θ x = θ y ⟷ σ x = σ y" and "∀n. ¬(Fun (PubConst Value n) [] ⊑⇩s⇩e⇩t insert t M)" shows "(M ⋅⇩s⇩e⇩t σ) ∪ N ⊢ t ⋅ σ" proof - obtain n where n: "intruder_deduct_num (M ⋅⇩s⇩e⇩t θ) n (t ⋅ θ)" using assms(1) deduct_num_if_deduct by blast hence "∃m ≤ n. intruder_deduct_num ((M ⋅⇩s⇩e⇩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 ⋅⇩s⇩e⇩t θ") case True note 2 = this have 3: "∀x ∈ fv⇩s⇩e⇩t M ∪ fv t. ∃c. θ x = Fun c []" "∀x ∈ fv⇩s⇩e⇩t M ∪ fv t. ∃c. σ x = Fun c []" using prems(2,3) by (blast, blast) have "t ⋅ σ ∈ M ⋅⇩s⇩e⇩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⇩s⇩e⇩t (M ⋅⇩s⇩e⇩t θ) = subterms⇩s⇩e⇩t M ⋅⇩s⇩e⇩t θ" "subterms⇩s⇩e⇩t (M ⋅⇩s⇩e⇩t σ) = subterms⇩s⇩e⇩t M ⋅⇩s⇩e⇩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 ⋅⇩s⇩e⇩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 ⋅⇩s⇩e⇩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 ⋅⇩s⇩e⇩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 ⋅⇩s⇩e⇩t σ" when "σ x ∈ M" using that subst_ground_ident[of "σ x" σ] by force ultimately have "σ x ∈ (M ⋅⇩s⇩e⇩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 ⋅⇩l⇩i⇩s⇩t θ" using t(1) by auto have ss: "∃l < n. intruder_deduct_num (M ⋅⇩s⇩e⇩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 ⋅⇩s⇩e⇩t σ) ∪ N) l (s ⋅ σ)" when s: "s ∈ set ss" for s proof - obtain l where l: "l < n" "intruder_deduct_num (M ⋅⇩s⇩e⇩t θ) l (s ⋅ θ)" using ss s by blast have *: "fv s ⊆ fv t" "subterms⇩s⇩e⇩t (insert s M) ⊆ subterms⇩s⇩e⇩t (insert t M)" using s unfolding Fun f ts by auto have "∃l' ≤ l. intruder_deduct_num ((M ⋅⇩s⇩e⇩t σ) ∪ N) l' (s ⋅ σ)" proof - have "∀x ∈ fv⇩s⇩e⇩t M ∪ fv s. (∃n. θ x = Fun (Val n) []) ∨ (∃n. θ x = Fun (PubConst Value n) [])" "∀x ∈ fv⇩s⇩e⇩t M ∪ fv s. ∃n. σ x = Fun (Val n) []" "∀x ∈ fv⇩s⇩e⇩t M ∪ fv s. (∃n. θ x = Fun (PubConst Value n) []) ⟶ σ x ∈ M ∪ N" "∀x ∈ fv⇩s⇩e⇩t M ∪ fv s. (∃n. θ x = Fun (Val n) []) ⟶ θ x = σ x" "∀x ∈ fv⇩s⇩e⇩t M ∪ fv s. ∀y ∈ fv⇩s⇩e⇩t M ∪ fv s. θ x = θ y ⟷ σ x = σ y" "∀n. Fun (PubConst Value n) [] ∉ subterms⇩s⇩e⇩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 ⋅⇩s⇩e⇩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 ⋅⇩l⇩i⇩s⇩t σ) = arity g" "public g" using t(2,3) unfolding f ts by auto let ?P = "λs l. l < n ∧ intruder_deduct_num ((M ⋅⇩s⇩e⇩t σ) ∪ N) l s" define steps where "steps ≡ λs. SOME l. ?P s l" have 2: "steps (s ⋅ σ) < n" "intruder_deduct_num ((M ⋅⇩s⇩e⇩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 ⋅⇩l⇩i⇩s⇩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 ⋅⇩l⇩i⇩s⇩t σ)" steps n] by (simp add: Suc_leI) qed show ?thesis using intruder_deduct_num.ComposeN[OF g, of "(M ⋅⇩s⇩e⇩t σ) ∪ N" steps] 2(2) 3 unfolding Fun by auto qed next assume "∃s ∈ subterms⇩s⇩e⇩t (M ⋅⇩s⇩e⇩t θ). (∃l < n. intruder_deduct_num (M ⋅⇩s⇩e⇩t θ) l s) ∧ (∀k ∈ set (fst (Ana s)). ∃l < n. intruder_deduct_num (M ⋅⇩s⇩e⇩t θ) l k) ∧ t ⋅ θ ∈ set (snd (Ana s))" then obtain s l where s: "s ∈ subterms⇩s⇩e⇩t M ⋅⇩s⇩e⇩t θ" "∀k ∈ set (fst (Ana s)). ∃l < n. intruder_deduct_num (M ⋅⇩s⇩e⇩t θ) l k" "t ⋅ θ ∈ set (snd (Ana s))" and l: "l < n" "intruder_deduct_num (M ⋅⇩s⇩e⇩t θ) l s" by (metis (no_types, lifting) M_subterms_eq(1)) obtain u where u: "u ⊑⇩s⇩e⇩t M" "s = u ⋅ θ" using s(1) by blast have u_fv: "fv u ⊆ fv⇩s⇩e⇩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 ⋅⇩l⇩i⇩s⇩t θ, R ⋅⇩l⇩i⇩s⇩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 ⋅⇩s⇩e⇩t θ) l (k ⋅ θ)" "t ⋅ θ ∈ set (R ⋅⇩l⇩i⇩s⇩t θ)" using s(2,3) by auto have IH1: "∃l < n. intruder_deduct_num ((M ⋅⇩s⇩e⇩t σ) ∪ N) l (u ⋅ σ)" proof - have "subterms u ⊆ subterms⇩s⇩e⇩t M" using u(1) subterms_subset by auto hence "subterms⇩s⇩e⇩t (insert u M) = subterms⇩s⇩e⇩t M" by blast hence *: "subterms⇩s⇩e⇩t (insert u M) ⊆ subterms⇩s⇩e⇩t (insert t M)" by auto have "∃l' ≤ l. intruder_deduct_num ((M ⋅⇩s⇩e⇩t σ) ∪ N) l' (u ⋅ σ)" proof - have "∀x ∈ fv⇩s⇩e⇩t M ∪ fv u. (∃n. θ x = Fun (Val n) []) ∨ (∃n. θ x = Fun (PubConst Value n) [])" "∀x ∈ fv⇩s⇩e⇩t M ∪ fv u. ∃n. σ x = Fun (Val n) []" "∀x ∈ fv⇩s⇩e⇩t M ∪ fv u. (∃n. θ x = Fun (PubConst Value n) []) ⟶ σ x ∈ M ∪ N" "∀x ∈ fv⇩s⇩e⇩t M ∪ fv u. (∃n. θ x = Fun (Val n) []) ⟶ θ x = σ x" "∀x ∈ fv⇩s⇩e⇩t M ∪ fv u. ∀y ∈ fv⇩s⇩e⇩t M ∪ fv u. θ x = θ y ⟷ σ x = σ y" "∀n. Fun (PubConst Value n) [] ∉ subterms⇩s⇩e⇩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 ⋅⇩s⇩e⇩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 ⋅⇩s⇩e⇩t σ) ∪ N) l (k ⋅ σ)" when k: "k ∈ set K" for k using k IH prems(2-) Ana⇩f_keys_not_val_terms s'(1) KR u(1) proof - have *: "fv k ⊆ fv⇩s⇩e⇩t M" using k KR Ana_keys_fv u(1) fv_subset_subterms by blast have **: "Fun (PubConst Value n) [] ⊑⇩s⇩e⇩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 ⋅⇩s⇩e⇩t θ) lk (k ⋅ θ)" using s'(1) k by fast have "∃l' ≤ lk. intruder_deduct_num ((M ⋅⇩s⇩e⇩t σ) ∪ N) l' (k ⋅ σ)" proof - have "∀x ∈ fv⇩s⇩e⇩t M ∪ fv k. (∃n. θ x = Fun (Val n) []) ∨ (∃n. θ x = Fun (PubConst Value n) [])" "∀x ∈ fv⇩s⇩e⇩t M ∪ fv k. ∃n. σ x = Fun (Val n) []" "∀x ∈ fv⇩s⇩e⇩t M ∪ fv k. (∃n. θ x = Fun (PubConst Value n) []) ⟶ σ x ∈ M ∪ N" "∀x ∈ fv⇩s⇩e⇩t M ∪ fv k. (∃n. θ x = Fun (Val n) []) ⟶ θ x = σ x" "∀x ∈ fv⇩s⇩e⇩t M ∪ fv k. ∀y ∈ fv⇩s⇩e⇩t M ∪ fv k. θ x = θ y ⟷ σ x = σ y" "∀n. Fun (PubConst Value n) [] ∉ subterms⇩s⇩e⇩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 ⋅⇩s⇩e⇩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 ⋅⇩l⇩i⇩s⇩t σ, R ⋅⇩l⇩i⇩s⇩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 ⋅⇩l⇩i⇩s⇩t σ)" proof - have r_subterm_u: "r ⊑ u" using r(1) KR Ana_subterm by blast have r_fv: "fv r ⊆ fv⇩s⇩e⇩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⇩s⇩e⇩t (insert t M)" by blast have r_subterm_M: "subterms r ⊆ subterms⇩s⇩e⇩t (insert t M)" using subterms_subset[OF r_subterm_u] in_subterms_subset_Union[OF u(1)] by (auto intro: subterms⇩s⇩e⇩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⇩s⇩e⇩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 ⋅⇩s⇩e⇩t σ) ∪ N) l1 (u ⋅ σ)" using IH1 by blast let ?P = "λs l. l < n ∧ intruder_deduct_num ((M ⋅⇩s⇩e⇩t σ) ∪ N) l s" define steps where "steps ≡ λs. SOME l. ?P s l" have 2: "steps (k ⋅ σ) < n" "intruder_deduct_num ((M ⋅⇩s⇩e⇩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 ⋅⇩l⇩i⇩s⇩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 ⋅⇩l⇩i⇩s⇩t σ)" steps n] by (simp add: Suc_leI) qed have IH2': "intruder_deduct_num ((M ⋅⇩s⇩e⇩t σ) ∪ N) (steps k) k" when k: "k ∈ set (K ⋅⇩l⇩i⇩s⇩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⇩s⇩u⇩b⇩s⇩t I" "wf⇩t⇩r⇩m⇩s (subst_range I)" shows "constraint_model I []" using I unfolding constraint_model_def by simp lemma welltyped_constraint_model_Nil: assumes I: "wt⇩s⇩u⇩b⇩s⇩t I" "interpretation⇩s⇩u⇩b⇩s⇩t I" "wf⇩t⇩r⇩m⇩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⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩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⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩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⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩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⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ ⊢ attack⟨n⟩" using rcv_attack in_ik⇩s⇩s⇩t_iff[of "attack⟨n⟩" "unlabel 𝒜"] ideduct_subst[OF intruder_deduct.Axiom[of "attack⟨n⟩" "ik⇩l⇩s⇩s⇩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⇩s⇩u⇩b⇩s⇩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 "Γ⇩v x = TAtom Value" using * Var by auto hence "τ = TAtom Value" using x Γ⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ))) ℐ" and s: "receive⟨ts⟩ ∈ set (unlabel (transaction_receive T ⋅⇩l⇩s⇩s⇩t θ))" shows "∀t ∈ set ts. IK ⊢ t ⋅ ℐ" proof - let ?R = "unlabel (dual⇩l⇩s⇩s⇩t (transaction_receive T ⋅⇩l⇩s⇩s⇩t θ))" let ?S = "λA. unlabel (dual⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩t θ))" let ?S' = "?S (transaction_receive T)" obtain l B s where B: "(l,send⟨ts⟩) = dual⇩l⇩s⇩s⇩t⇩p ((l,s) ⋅⇩l⇩s⇩s⇩t⇩p θ)" "prefix ((B ⋅⇩l⇩s⇩s⇩t θ)@[(l,s) ⋅⇩l⇩s⇩s⇩t⇩p θ]) (transaction_receive T ⋅⇩l⇩s⇩s⇩t θ)" using s dual⇩l⇩s⇩s⇩t_unlabel_steps_iff(2)[of ts "transaction_receive T ⋅⇩l⇩s⇩s⇩t θ"] dual⇩l⇩s⇩s⇩t_in_set_prefix_obtain_subst[of "send⟨ts⟩" "transaction_receive T" θ] by blast have 1: "unlabel (dual⇩l⇩s⇩s⇩t ((B ⋅⇩l⇩s⇩s⇩t θ)@[(l,s) ⋅⇩l⇩s⇩s⇩t⇩p θ])) = unlabel (dual⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩t θ))@[send⟨ts⟩]" using B(1) unlabel_append dual⇩l⇩s⇩s⇩t⇩p_subst dual⇩l⇩s⇩s⇩t_subst singleton_lst_proj(4) dual⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩t θ))@[send⟨ts⟩]) ℐ" using B 1 unfolding prefix_def unlabel_def by (metis dual⇩l⇩s⇩s⇩t_def map_append strand_sem_append_stateful) hence t_deduct: "∀t ∈ set ts. IK ∪ (ik⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩t θ)) ⋅⇩s⇩e⇩t ℐ) ⊢ t ⋅ ℐ" using strand_sem_append_stateful[of IK DB "unlabel (dual⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩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 ⋅⇩l⇩s⇩s⇩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 ⋅⇩l⇩s⇩s⇩t θ)). ∃t. s = receive⟨t⟩" by (simp add: list.pred_set is_Receive_def) hence "∀s ∈ set (unlabel (dual⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩t θ))). ∃t. s = send⟨t⟩" by (metis dual⇩l⇩s⇩s⇩t_memberD dual⇩l⇩s⇩s⇩t⇩p_inv(2) unlabel_in unlabel_mem_has_label) } ultimately have "∀s ∈ set ?R. ∃ts. s = send⟨ts⟩" by simp hence "ik⇩s⇩s⇩t ?R = {}" unfolding unlabel_def ik⇩s⇩s⇩t_def by fast hence "ik⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩t θ)) = {}" using B(2) 1 ik⇩s⇩s⇩t_append dual⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ))) ℐ" shows "⟨ac: t ∈ u⟩ ∈ set (unlabel (transaction_checks T ⋅⇩l⇩s⇩s⇩t θ)) ⟹ (t ⋅ ℐ, u ⋅ ℐ) ∈ DB" and "⟨ac: t ≐ u⟩ ∈ set (unlabel (transaction_checks T ⋅⇩l⇩s⇩s⇩t θ)) ⟹ t ⋅ ℐ = u ⋅ ℐ" proof - let ?s = "⟨ac: t ∈ u⟩" let ?s' = "⟨ac: t ≐ u⟩" let ?C = "set (unlabel (transaction_checks T ⋅⇩l⇩s⇩s⇩t θ))" let ?R = "transaction_receive T@transaction_checks T" let ?R' = "unlabel (dual⇩l⇩s⇩s⇩t (?R ⋅⇩l⇩s⇩s⇩t θ))" let ?S = "λA. unlabel (dual⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩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⇩s⇩s⇩t (unlabel (dual⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩t θ))) ℐ DB" have s_in: "?s ∈ ?C ⟹ ?s ∈ set (unlabel (?R ⋅⇩l⇩s⇩s⇩t θ))" "?s' ∈ ?C ⟹ ?s' ∈ set (unlabel (?R ⋅⇩l⇩s⇩s⇩t θ))" using subst_lsst_append[of "transaction_receive T"] unlabel_append[of "transaction_receive T"] by auto have 1: "unlabel (dual⇩l⇩s⇩s⇩t ((B ⋅⇩l⇩s⇩s⇩t θ)@[(l,s) ⋅⇩l⇩s⇩s⇩t⇩p θ])) = unlabel (dual⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩t θ))@[s']" when B: "(l,s') = dual⇩l⇩s⇩s⇩t⇩p ((l,s) ⋅⇩l⇩s⇩s⇩t⇩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⇩l⇩s⇩s⇩t⇩p_subst dual⇩l⇩s⇩s⇩t_subst singleton_lst_proj(4) dual⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩t θ))@[s']) ℐ" when B: "(l,s') = dual⇩l⇩s⇩s⇩t⇩p ((l,s) ⋅⇩l⇩s⇩s⇩t⇩p θ)" "prefix ((B ⋅⇩l⇩s⇩s⇩t θ)@[(l,s) ⋅⇩l⇩s⇩s⇩t⇩p θ]) (?R ⋅⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t_def by (metis (no_types) map_append) qed have s_sem: "?s ∈ ?C ⟹ (l,?s) = dual⇩l⇩s⇩s⇩t⇩p ((l,s) ⋅⇩l⇩s⇩s⇩t⇩p θ) ⟹ (t ⋅ ℐ, u ⋅ ℐ) ∈ ?dbupd B" "?s' ∈ ?C ⟹ (l,?s') = dual⇩l⇩s⇩s⇩t⇩p ((l,s) ⋅⇩l⇩s⇩s⇩t⇩p θ) ⟹ t ⋅ ℐ = u ⋅ ℐ" when B: "prefix ((B ⋅⇩l⇩s⇩s⇩t θ)@[(l,s) ⋅⇩l⇩s⇩s⇩t⇩p θ]) (?R ⋅⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩t θ))" _ ℐ] by (fastforce, fastforce) have 3: "∀a ∈ set (unlabel (dual⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩t θ))). ¬is_Insert a ∧ ¬is_Delete a" when B: "prefix ((B ⋅⇩l⇩s⇩s⇩t θ)@[(l,s) ⋅⇩l⇩s⇩s⇩t⇩p θ]) (?R ⋅⇩l⇩s⇩s⇩t θ)" for l s and B::"('fun,'atom,'sets,'lbl) prot_strand" proof - have "∀a ∈ set (unlabel (dual⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩t θ))). ?Q a" proof fix a assume a: "a ∈ set (unlabel (dual⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩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 ⋅⇩l⇩s⇩s⇩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 ⋅⇩l⇩s⇩s⇩t θ)). ?P a" using unlabel_mono[OF set_mono_prefix[OF append_prefixD[OF B]]] by blast obtain l where "(l,a) ∈ set (dual⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩t θ))" using a by (meson unlabel_mem_has_label) then obtain b where b: "(l,b) ∈ set (B ⋅⇩l⇩s⇩s⇩t θ)" "dual⇩l⇩s⇩s⇩t⇩p (l,b) = (l,a)" using dual⇩l⇩s⇩s⇩t_memberD by blast hence "?P b" using B_P unfolding unlabel_def by fastforce thus "?Q a" using dual⇩l⇩s⇩s⇩t⇩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⇩l⇩s⇩s⇩t⇩p ((l,s) ⋅⇩l⇩s⇩s⇩t⇩p θ)" "prefix ((B ⋅⇩l⇩s⇩s⇩t θ)@[(l,s) ⋅⇩l⇩s⇩s⇩t⇩p θ]) (?R ⋅⇩l⇩s⇩s⇩t θ)" using s_in(1)[OF s] dual⇩l⇩s⇩s⇩t_unlabel_steps_iff(6)[of _ t u] dual⇩l⇩s⇩s⇩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⇩s⇩s⇩t_no_upd[of "unlabel (dual⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩t θ))" ℐ DB] by simp qed show "t ⋅ ℐ = u ⋅ ℐ" when s: "?s' ∈ ?C" proof - obtain l B s where B: "(l,?s') = dual⇩l⇩s⇩s⇩t⇩p ((l,s) ⋅⇩l⇩s⇩s⇩t⇩p θ)" "prefix ((B ⋅⇩l⇩s⇩s⇩t θ)@[(l,s) ⋅⇩l⇩s⇩s⇩t⇩p θ]) (?R ⋅⇩l⇩s⇩s⇩t θ)" using s_in(2)[OF s] dual⇩l⇩s⇩s⇩t_unlabel_steps_iff(3)[of _ t u] dual⇩l⇩s⇩s⇩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⇩s⇩s⇩t_no_upd[of "unlabel (dual⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ))) ℐ" and "NegChecks X F G ∈ set (unlabel (transaction_checks T ⋅⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (?R ⋅⇩l⇩s⇩s⇩t θ))" let ?S = "λA. unlabel (dual⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩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 ⋅⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t⇩p ((l,s) ⋅⇩l⇩s⇩s⇩t⇩p θ)" "prefix ((B ⋅⇩l⇩s⇩s⇩t θ)@[(l,s) ⋅⇩l⇩s⇩s⇩t⇩p θ]) (?R ⋅⇩l⇩s⇩s⇩t θ)" using s dual⇩l⇩s⇩s⇩t_unlabel_steps_iff(7)[of X F G] dual⇩l⇩s⇩s⇩t_in_set_prefix_obtain_subst[of ?s ?R θ] by blast have 1: "unlabel (dual⇩l⇩s⇩s⇩t ((B ⋅⇩l⇩s⇩s⇩t θ)@[(l,s) ⋅⇩l⇩s⇩s⇩t⇩p θ])) = unlabel (dual⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩t θ))@[?s]" using B(1) unlabel_append dual⇩l⇩s⇩s⇩t⇩p_subst dual⇩l⇩s⇩s⇩t_subst singleton_lst_proj(4) dual⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩t θ))@[?s]) ℐ" using B 1 strand_sem_append_stateful subst_lsst_append unfolding prefix_def unlabel_def dual⇩l⇩s⇩s⇩t_def by (metis (no_types) map_append) hence s_sem: "negchecks_model ℐ (dbupd⇩s⇩s⇩t (unlabel (dual⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩t θ))) ℐ DB) X F G" using strand_sem_append_stateful[of IK DB "unlabel (dual⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩t θ))" "[?s]" ℐ] by fastforce have "∀a ∈ set (unlabel (dual⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩t θ))). ?Q a" proof fix a assume a: "a ∈ set (unlabel (dual⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩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 ⋅⇩l⇩s⇩s⇩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 ⋅⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩t θ))" using a by (meson unlabel_mem_has_label) then obtain b where b: "(l,b) ∈ set (B ⋅⇩l⇩s⇩s⇩t θ)" "dual⇩l⇩s⇩s⇩t⇩p (l,b) = (l,a)" using dual⇩l⇩s⇩s⇩t_memberD by blast hence "?P b" using B_P unfolding unlabel_def by fastforce thus "?Q a" using dual⇩l⇩s⇩s⇩t⇩p_inv[OF b(2)] by (cases b) auto qed hence "∀a ∈ set (unlabel (dual⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩t θ))). ¬is_Insert a ∧ ¬is_Delete a" by fastforce thus ?thesis using dbupd⇩s⇩s⇩t_no_upd[of "unlabel (dual⇩l⇩s⇩s⇩t (B ⋅⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ))) ℐ" and c: "NegChecks X [] [(t,u)] ∈ set (unlabel (transaction_checks T ⋅⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ))" and "rm ≡ λX. rm_vars (set X)" assumes T: "wellformed_transaction T" and I: " interpretation⇩s⇩u⇩b⇩s⇩t I" "wf⇩t⇩r⇩m⇩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⇩l⇩s⇩s⇩t (T ⋅⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t_list_all(1), metis (no_types) subst_sst_list_all(11) unlabel_subst dual⇩l⇩s⇩s⇩t_list_all(11), metis (no_types) subst_sst_list_all(10) unlabel_subst dual⇩l⇩s⇩s⇩t_list_all(10), metis (no_types) subst_sst_list_all(1) unlabel_subst dual⇩l⇩s⇩s⇩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⇩s⇩s⇩t (?C (transaction_receive T)) ⋅⇩s⇩e⇩t I) = M" "dbupd⇩s⇩s⇩t (?C (transaction_receive T)) I D = D" "M ∪ (ik⇩s⇩s⇩t (?C (transaction_checks T)) ⋅⇩s⇩e⇩t I) = M" "dbupd⇩s⇩s⇩t (?C (transaction_checks T)) I D = D" by (metis ik⇩s⇩s⇩t_snoc_no_receive_empty sup_bot.right_neutral, metis dbupd⇩s⇩s⇩t_no_upd, metis ik⇩s⇩s⇩t_snoc_no_receive_empty sup_bot.right_neutral, metis dbupd⇩s⇩s⇩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 ⋅⇩l⇩i⇩s⇩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 ⋅⇩l⇩i⇩s⇩t θ" by (metis ts dual⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩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' ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) θ" "G = G' ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) θ" using unlabel_mem_has_label[OF iffD2[OF dual⇩l⇩s⇩s⇩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 ⋅ δ ∘⇩s I ≠ s ⋅ δ ∘⇩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) ⋅⇩p δ ∘⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ))" and "rm ≡ λX. rm_vars (set X)" assumes T: "wellformed_transaction T" and I: " interpretation⇩s⇩u⇩b⇩s⇩t I" "wf⇩t⇩r⇩m⇩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⇩s⇩s⇩t (unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t δ))) ⋅⇩s⇩e⇩t ℐ) ⋅⇩α⇩s⇩e⇩t a = (trms⇩s⇩s⇩t (unlabel (transaction_send T)) ⋅⇩s⇩e⇩t δ ⋅⇩s⇩e⇩t ℐ) ⋅⇩α⇩s⇩e⇩t a" (is "?A = ?B") using dual_transaction_ik_is_transaction_send[OF assms] subst_lsst_unlabel[of "dual⇩l⇩s⇩s⇩t (transaction_strand T)" δ] ik⇩s⇩s⇩t_subst[of "unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand T))" δ] dual⇩l⇩s⇩s⇩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⇩s⇩e⇩t N. Γ⇩v x = TAtom Value" shows "SMP M ⊆ {a ⋅ δ | a δ. a ∈ N ∧ wt⇩s⇩u⇩b⇩s⇩t δ ∧ wf⇩t⇩r⇩m⇩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⇩s⇩u⇩b⇩s⇩t δ ∧ wf⇩t⇩r⇩m⇩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⇩s⇩u⇩b⇩s⇩t Var" "wf⇩t⇩r⇩m⇩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⇩s⇩u⇩b⇩s⇩t δ" "wf⇩t⇩r⇩m⇩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⇩s⇩u⇩b⇩s⇩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⇩s⇩u⇩b⇩s⇩t δ" "wf⇩t⇩r⇩m⇩s (subst_range δ)" by (auto simp add: S_def) have "wt⇩s⇩u⇩b⇩s⇩t (δ ∘⇩s θ)" "wf⇩t⇩r⇩m⇩s (subst_range (δ ∘⇩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 ⋅ δ ∘⇩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⇩s⇩u⇩b⇩s⇩t δ" "wf⇩t⇩r⇩m⇩s (subst_range δ)" by (auto simp add: S_def) obtain Ka Ta where a': "Ana a = (Ka,Ta)" by force have *: "K = Ka ⋅⇩l⇩i⇩s⇩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⇩s⇩s⇩t⇩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⇩t⇩r⇩m⇩s' arity (trms⇩l⇩s⇩s⇩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⇩s⇩s⇩t⇩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⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T) = {} ∧ set (transaction_fresh T) ⊆ fv⇩l⇩s⇩s⇩t (filter (is_Insert ∘ snd) (transaction_updates T)) ∪ fv⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_updates T) ∪ fv⇩l⇩s⇩s⇩t (transaction_send T) - set (transaction_fresh T) ⊆ fv⇩l⇩s⇩s⇩t (transaction_receive T) ∪ fv⇩l⇩s⇩s⇩t (transaction_checks T) ∧ (∀r ∈ set (unlabel (transaction_checks T)). is_Equality r ⟶ fv (the_rhs r) ⊆ fv⇩l⇩s⇩s⇩t (transaction_receive T)) ∧ fv⇩l⇩s⇩s⇩t (transaction_checks T) ⊆ fv⇩l⇩s⇩s⇩t (transaction_receive T) ∪ fv⇩l⇩s⇩s⇩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⟩⇩s) ∧ (is_NegChecks a ⟶ list_all (λ(_,t). t ≠ ⟨s⟩⇩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⇩s⇩e⇩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). Γ⇩v x = TAtom Value" (is ?B) and "∀x ∈ vars⇩l⇩s⇩s⇩t (transaction_strand T). Γ⇩v x = TAtom Value" (is ?C) and "bvars⇩l⇩s⇩s⇩t (transaction_strand T) = {}" (is ?D1) and "fv_transaction T ∩ bvars_transaction T = {}" (is ?D2) and "set (transaction_fresh T) ⊆ fv⇩l⇩s⇩s⇩t (filter (is_Insert ∘ snd) (transaction_updates T)) ∪ fv⇩l⇩s⇩s⇩t (transaction_send T)" (is ?E) and "set (transaction_fresh T) ⊆ fv⇩l⇩s⇩s⇩t (transaction_updates T) ∪ fv⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_checks T). x ∈ fv⇩l⇩s⇩s⇩t (transaction_receive T) ∨ (∃t s. select⟨t,s⟩ ∈ set (unlabel (transaction_checks T)) ∧ x ∈ fv t ∪ fv s)" (is ?H) and "fv⇩l⇩s⇩s⇩t (transaction_updates T) ∪ fv⇩l⇩s⇩s⇩t (transaction_send T) - set (transaction_fresh T) ⊆ fv⇩l⇩s⇩s⇩t (transaction_receive T) ∪ fv⇩l⇩s⇩s⇩t (transaction_checks T)" (is ?I) and "∀x ∈ set (unlabel (transaction_checks T)). is_Equality x ⟶ fv (the_rhs x) ⊆ fv⇩l⇩s⇩s⇩t (transaction_receive T)" (is ?J) (* TODO: why do we need this requirement? *) and "set (transaction_fresh T) ∩ fv⇩l⇩s⇩s⇩t (transaction_receive T) = {}" (is ?K1) and "set (transaction_fresh T) ∩ fv⇩l⇩s⇩s⇩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 ⊑ Γ⇩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 Γ⇩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⇩l⇩s⇩s⇩t (transaction_checks T)" hence "x ∈ fv⇩l⇩s⇩s⇩t (transaction_receive T) ∨ x ∈ fv⇩l⇩s⇩s⇩t ?selects" using T unfolding admissible_transaction'_def by blast thus "x ∈ fv⇩l⇩s⇩s⇩t (transaction_receive T) ∨ (∃t s. select⟨t,s⟩ ∈ set (unlabel (transaction_checks T)) ∧ x ∈ fv t ∪ fv s)" proof assume "x ∈ fv⇩l⇩s⇩s⇩t ?selects" then obtain r where r: "x ∈ fv⇩s⇩s⇩t⇩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⇩s⇩s⇩t⇩p ?x ⊆ fv⇩l⇩s⇩s⇩t (transaction_checks T)" using * by force hence ***: "fv⇩s⇩s⇩t⇩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⇩s⇩s⇩t⇩p ?x ⊆ fv⇩l⇩s⇩s⇩t (transaction_checks T)" using * by force hence ***: "fv⇩s⇩s⇩t⇩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⇩s⇩s⇩t⇩p ?x ⊆ fv⇩l⇩s⇩s⇩t (transaction_checks T)" "set (bvars⇩s⇩s⇩t⇩p ?x) ⊆ bvars⇩l⇩s⇩s⇩t (transaction_checks T)" using 0 by force+ moreover have "set (transaction_fresh T) ∩ fv⇩l⇩s⇩s⇩t (transaction_receive T) = {}" "set (transaction_fresh T) ∩ fv⇩l⇩s⇩s⇩t (transaction_checks T) = {}" using T_wf unfolding wellformed_transaction_def by fast+ ultimately have "fv⇩s⇩s⇩t⇩p ?x ∩ set (transaction_fresh T) = {}" "set (bvars⇩s⇩s⇩t⇩p ?x) ∩ set (transaction_fresh T) = {}" using admissible_transactionE(7)[OF T_adm] wellformed_transaction_wf⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_receive T)" (is ?A) and "x ∉ fv⇩l⇩s⇩s⇩t (transaction_checks T)" (is ?B) and "x ∉ vars⇩l⇩s⇩s⇩t (transaction_receive T)" (is ?C) and "x ∉ vars⇩l⇩s⇩s⇩t (transaction_checks T)" (is ?D) and "x ∉ bvars⇩l⇩s⇩s⇩t (transaction_receive T)" (is ?E) and "x ∉ bvars⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_updates T) ∪ fv⇩l⇩s⇩s⇩t (transaction_send T)" "set (transaction_fresh T) ∩ fv⇩l⇩s⇩s⇩t (transaction_receive T) = {}" "set (transaction_fresh T) ∩ fv⇩l⇩s⇩s⇩t (transaction_checks T) = {}" "fv_transaction T ∩ bvars_transaction T = {}" using admissible_transactionE[OF T] by argo+ have 1: "set (transaction_fresh T) ∩ bvars⇩l⇩s⇩s⇩t (transaction_checks T) = {}" using 0(1,4) fv_transaction_unfold[of T] bvars_transaction_unfold[of T] by blast have 2: "vars⇩l⇩s⇩s⇩t (transaction_receive T) = fv⇩l⇩s⇩s⇩t (transaction_receive T)" "bvars⇩l⇩s⇩s⇩t (transaction_receive T) = {}" using bvars_wellformed_transaction_unfold[OF T_wf] vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩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⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_receive T) ∨ (x ∈ fv⇩l⇩s⇩s⇩t (transaction_checks T) ∧ (∃t s. select⟨t,s⟩ ∈ set (unlabel (transaction_checks T)) ∧ x ∈ fv t ∪ fv s))" proof - have "x ∈ fv⇩l⇩s⇩s⇩t (transaction_receive T) ∪ fv⇩l⇩s⇩s⇩t (transaction_checks T) ∪ fv⇩l⇩s⇩s⇩t (transaction_updates T) ∪ fv⇩l⇩s⇩s⇩t (transaction_send T)" using x(1) fv⇩s⇩s⇩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⇩s⇩e⇩t (set ts)) ∨ (∃s. select⟨Var x, s⟩ ∈ set (unlabel (transaction_checks T)))" proof (cases "x ∈ fv⇩l⇩s⇩s⇩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 ⋅⇩l⇩s⇩s⇩t θ)) ∧ θ x ⊑⇩s⇩e⇩t set ts) ∨ (∃s. select⟨θ x, s⟩ ∈ set (unlabel (transaction_checks T ⋅⇩l⇩s⇩s⇩t θ)))" proof - note 0 = admissible_transaction_fv_in_receives_or_selects'[OF T x] have 1: "θ x ⊑⇩s⇩e⇩t set (ts ⋅⇩l⇩i⇩s⇩t θ)" when ts: "x ∈ fv⇩s⇩e⇩t (set ts)" for ts using that subst_mono_fv[of x _ θ] by auto have 2: "receive⟨ts ⋅⇩l⇩i⇩s⇩t θ⟩ ∈ set (A ⋅⇩s⇩s⇩t θ)" when "receive⟨ts⟩ ∈ set A" for ts A using that by fast have 3: "select⟨t ⋅ θ,s ⋅ θ⟩ ∈ set (A ⋅⇩s⇩s⇩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 eval_term.simps(1)) qed lemma admissible_transaction_fv_in_receives_or_selects_dual_subst: defines "f ≡ λS. unlabel (dual⇩l⇩s⇩s⇩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 ⋅⇩l⇩s⇩s⇩t θ)) ∧ θ x ⊑⇩s⇩e⇩t set ts) ∨ (∃s. select⟨θ x, s⟩ ∈ set (f (transaction_checks T ⋅⇩l⇩s⇩s⇩t θ)))" using admissible_transaction_fv_in_receives_or_selects_subst[OF T x, of θ] by (metis (no_types, lifting) f_def dual⇩l⇩s⇩s⇩t_unlabel_steps_iff(2) dual⇩l⇩s⇩s⇩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⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩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 Γ⇩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 "Γ⇩v x = TAtom Value" shows "x ∈ fv_transaction T" using assms(2,3) Γ⇩v_TAtom''(2)[of x] vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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 ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))" shows "∀t ∈ set ts. ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t ℐ ⊢ t ⋅ ℐ" proof - define θ where "θ ≡ ξ ∘⇩s σ ∘⇩s α" have t': "send⟨ts⟩ ∈ set (unlabel (dual⇩l⇩s⇩s⇩t (transaction_receive T ⋅⇩l⇩s⇩s⇩t θ)))" using t dual⇩l⇩s⇩s⇩t_unlabel_steps_iff(2) unfolding θ_def by blast then obtain T1 T2 where T: "unlabel (dual⇩l⇩s⇩s⇩t (transaction_receive T ⋅⇩l⇩s⇩s⇩t θ)) = T1@send⟨ts⟩#T2" using t' by (meson split_list) have "constr_sem_stateful ℐ (unlabel A@unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩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⇩s⇩s⇩t (unlabel A@T1) ⋅⇩s⇩e⇩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⇩l⇩s⇩s⇩t (transaction_receive T ⋅⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_receive T ⋅⇩l⇩s⇩s⇩t θ))" using x unfolding unlabel_def by fastforce then obtain ly where ly: "ly ∈ set (transaction_receive T ⋅⇩l⇩s⇩s⇩t θ)" "(l,x) = dual⇩l⇩s⇩s⇩t⇩p ly" unfolding dual⇩l⇩s⇩s⇩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 ⋅⇩l⇩s⇩s⇩t θ)" "(l,x) = dual⇩l⇩s⇩s⇩t⇩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) ⋅⇩l⇩s⇩s⇩t⇩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⇩s⇩s⇩t T1 = {}" unfolding ik⇩s⇩s⇩t_def is_Receive_def by fast hence "ik⇩s⇩s⇩t (unlabel A@T1) = ik⇩l⇩s⇩s⇩t A" using ik⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩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 ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α" let ?S = "λS. transaction_receive T@S" let ?S' = "λS. ?S S ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)))" using ℐ unfolding constraint_model_def by simp moreover have "dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t δ) = dual⇩l⇩s⇩s⇩t (?S (T1@[c]) ⋅⇩l⇩s⇩s⇩t δ)@ dual⇩l⇩s⇩s⇩t (T2@transaction_updates T@transaction_send T ⋅⇩l⇩s⇩s⇩t δ)" when "transaction_checks T = T1@c#T2" for T1 T2 c δ using that dual⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (?S' (T1@[(l,⟨Var (?x n) in ?s⟩)]))) = unlabel (A@dual⇩l⇩s⇩s⇩t (?S' T1))@[⟨α (?x n) in ?s⟩]" using T a σ dual⇩l⇩s⇩s⇩t_append subst_lsst_append unlabel_append ξ_empty by (fastforce simp add: transaction_fresh_subst_def unlabel_def dual⇩l⇩s⇩s⇩t_def subst_apply_labeled_stateful_strand_def subst_compose) moreover have "db⇩s⇩s⇩t (unlabel A) = db⇩s⇩s⇩t (unlabel (A@dual⇩l⇩s⇩s⇩t (?S' T1)))" by (simp add: T1 db⇩s⇩s⇩t_transaction_prefix_eq[OF T_wf] del: unlabel_append) ultimately have "∃M. strand_sem_stateful M (set (db⇩s⇩s⇩t (unlabel A) ℐ)) [⟨α (?x n) in ?s⟩] ℐ" using T'_model[OF T1] db⇩s⇩s⇩t_set_is_dbupd⇩s⇩s⇩t[of _ ℐ] strand_sem_append_stateful[of _ _ _ _ ℐ] by (simp add: db⇩s⇩s⇩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⇩s⇩s⇩t⇩p ⟨Var (?x n) not in ?s⟩" using vars⇩s⇩s⇩t⇩p_cases(9)[of "[]" "Var (?x n)" ?s] by auto hence "?x n ∈ vars⇩l⇩s⇩s⇩t (transaction_checks T)" using a unfolding vars⇩s⇩s⇩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⇩l⇩s⇩s⇩t (?S' (T1@[(l,⟨Var (?x n) not in ?s⟩)]))) = unlabel (A@dual⇩l⇩s⇩s⇩t (?S' T1))@[⟨α (?x n) not in ?s⟩]" using T a σ dual⇩l⇩s⇩s⇩t_append subst_lsst_append unlabel_append ξ_empty by (fastforce simp add: transaction_fresh_subst_def unlabel_def dual⇩l⇩s⇩s⇩t_def subst_apply_labeled_stateful_strand_def subst_compose) moreover have "db⇩s⇩s⇩t (unlabel A) = db⇩s⇩s⇩t (unlabel (A@dual⇩l⇩s⇩s⇩t (?S' T1)))" by (simp add: T1 db⇩s⇩s⇩t_transaction_prefix_eq[OF T_wf] del: unlabel_append) ultimately have "∃M. strand_sem_stateful M (set (db⇩s⇩s⇩t (unlabel A) ℐ)) [⟨α (?x n) not in ?s⟩] ℐ" using T'_model[OF T1] db⇩s⇩s⇩t_set_is_dbupd⇩s⇩s⇩t[of _ ℐ] strand_sem_append_stateful[of _ _ _ _ ℐ] by (simp add: db⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩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 ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α" let ?S = "λS. transaction_receive T@S" let ?S' = "λS. ?S S ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)))" using ℐ unfolding constraint_model_def by simp moreover have "dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t δ) = dual⇩l⇩s⇩s⇩t (?S (T1@[c]) ⋅⇩l⇩s⇩s⇩t δ)@ dual⇩l⇩s⇩s⇩t (T2@transaction_updates T@transaction_send T ⋅⇩l⇩s⇩s⇩t δ)" when "transaction_checks T = T1@c#T2" for T1 T2 c δ using that dual⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (?S' (T1@[(l,select⟨Var (?x n), ?s⟩)]))) = unlabel (A@dual⇩l⇩s⇩s⇩t (?S' T1))@[select⟨α (?x n), ?s⟩]" using T a σ dual⇩l⇩s⇩s⇩t_append subst_lsst_append unlabel_append ξ_empty by (fastforce simp add: transaction_fresh_subst_def unlabel_def dual⇩l⇩s⇩s⇩t_def subst_apply_labeled_stateful_strand_def subst_compose) moreover have "db⇩s⇩s⇩t (unlabel A) = db⇩s⇩s⇩t (unlabel (A@dual⇩l⇩s⇩s⇩t (?S' T1)))" by (simp add: T1 db⇩s⇩s⇩t_transaction_prefix_eq[OF T_wf] del: unlabel_append) ultimately have "∃M. strand_sem_stateful M (set (db⇩s⇩s⇩t (unlabel A) ℐ)) [⟨α (?x n) in ?s⟩] ℐ" using T'_model[OF T1] db⇩s⇩s⇩t_set_is_dbupd⇩s⇩s⇩t[of _ ℐ] strand_sem_append_stateful[of _ _ _ _ ℐ] by (simp add: db⇩s⇩s⇩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⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩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⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩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⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩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⇩s⇩s⇩t⇩p r ⊆ fv_transaction T" using r fv_transaction_unfold[of T] by auto have fv_r2: "fv⇩s⇩s⇩t⇩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⇩s⇩s⇩t⇩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. Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩v x = TAtom (Atom a))" and "∀x ∈ fv_transaction T. Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩v x = TAtom (Atom a))" and "∀x ∈ set (transaction_fresh T). Γ⇩v x = TAtom Value" proof - note T_wf = admissible_transaction_is_wellformed_transaction(1)[OF T_adm] show "∀x ∈ vars_transaction T. Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩v x = TAtom (Atom a))" using admissible_transactionE(3)[OF T_adm] by fast thus "∀x ∈ fv_transaction T. Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩v x = TAtom (Atom a))" using vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t by fast show "∀x ∈ set (transaction_fresh T). Γ⇩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⇩s⇩e⇩t (trms_transaction T)" and "Fun (PubConst Value n) S ∉ subterms⇩s⇩e⇩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⇩s⇩e⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ))) ∨ ⟨θ y != θ x⟩ ∈ set (unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ)))" using stateful_strand_step_subst_inI(8)[of _ _ "unlabel (transaction_strand T)" θ] subst_lsst_unlabel[of "transaction_strand T" θ] dual⇩l⇩s⇩s⇩t_unlabel_steps_iff(7)[of "[]" _ "[]"] by force then obtain B where B: "prefix (B@[⟨θ x != θ y⟩]) (unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ))) ∨ prefix (B@[⟨θ y != θ x⟩]) (unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ))" assumes T: "admissible_transaction' T" and I: " interpretation⇩s⇩u⇩b⇩s⇩t I" "wf⇩t⇩r⇩m⇩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⇩t⇩r⇩m⇩s: assumes "admissible_transaction_terms T" shows "wf⇩t⇩r⇩m⇩s (trms_transaction T)" by (rule conjunct1[OF assms[unfolded admissible_transaction_terms_def wf⇩t⇩r⇩m⇩s_code[symmetric]]]) lemma admissible_transactions_wf⇩t⇩r⇩m⇩s: assumes "admissible_transaction' T" shows "wf⇩t⇩r⇩m⇩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⇩t⇩r⇩m⇩s) qed lemma admissible_transaction_no_Ana_Attack: assumes "admissible_transaction_terms T" and "t ∈ subterms⇩s⇩e⇩t (trms_transaction T)" shows "attack⟨n⟩ ∉ set (snd (Ana t))" proof - obtain r where r: "r ∈ set (unlabel (transaction_strand T))" "t ∈ subterms⇩s⇩e⇩t (trms⇩s⇩s⇩t⇩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⇩s⇩e⇩t (trms⇩s⇩s⇩t⇩p r)" using Ana_subterm[OF t] r(2) subterms_subset by fast hence "∃f ∈ ⋃(funs_term ` trms⇩s⇩s⇩t⇩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 "Γ⇩v x = TAtom Value" proof - have "x ∈ vars_transaction T" using x vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t[of "unlabel (transaction_strand T)"] by blast thus "Γ⇩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)" "Γ⇩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. Γ⇩v x = TAtom Value ⟶ occurs (Var x) ∈ set (the_msgs (hd (unlabel ?rcvs)))" using T x unfolding admissible_transaction_occurs_checks_def Γ⇩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⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_send T)" using t(2) ts unfolding trms⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_send T)" proof - have "∃x ∈ set (transaction_fresh T). t = occurs (Var x)" when "t ∈ trms⇩l⇩s⇩s⇩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 ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩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 "Γ⇩v x = TAtom Value" "Var x ∈ set ts" "fv⇩s⇩e⇩t (set ts) = {x}" "∀n. ¬(Fun (Val n) [] ⊑⇩s⇩e⇩t set ts)" "T ∈ set P" "T = Transaction (λ(). []) [x] [] [] upds [(l, send⟨ts⟩)]" "upds = [] ∨ (upds = [(l', insert⟨Var x, ⟨s⟩⇩s⟩)] ∧ (∀T ∈ set P. ∀(l,a) ∈ set (transaction_strand T). ∀t. a ≠ select⟨t, ⟨s⟩⇩s⟩ ∧ a ≠ ⟨t in ⟨s⟩⇩s⟩ ∧ a ≠ ⟨t not in ⟨s⟩⇩s⟩ ∧ a ≠ delete⟨t, ⟨s⟩⇩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⟩⇩s) ∧ (is_NegChecks a ⟶ list_all (λ(_,t). t ≠ ⟨s⟩⇩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⇩s⇩e⇩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⇩s⇩e⇩t (set ts) = {x}" "upds = [] ∨ (upds = [(l', insert⟨Var x, ⟨s⟩⇩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⇩s⇩e⇩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⟩⇩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: "Γ⇩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) [] ⊑⇩s⇩e⇩t set ts)" using admissible_transactions_no_Value_consts[OF T_adm] by blast have "a ≠ select⟨t, ⟨s⟩⇩s⟩ ∧ a ≠ ⟨t in ⟨s⟩⇩s⟩ ∧ a ≠ ⟨t not in ⟨s⟩⇩s⟩ ∧ a ≠ delete⟨t, ⟨s⟩⇩s⟩" when upds: "upds = [(k, insert⟨Var x,⟨s⟩⇩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⟩⇩s⟩ ∧ a ≠ ⟨t in ⟨s⟩⇩s⟩ ∧ a ≠ ⟨t not in ⟨s⟩⇩s⟩ ∧ a ≠ delete⟨t, ⟨s⟩⇩s⟩" when upds: "upds = [(k, insert⟨Var x,⟨s⟩⇩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⟩⇩s) ∧ (is_NegChecks a ⟶ list_all (λ(_,t). t ≠ ⟨s⟩⇩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⇩s⇩e⇩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 force 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⇩s⇩u⇩b⇩s⇩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⇩t⇩r⇩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⇩t⇩r⇩m⇩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⇩l⇩s⇩s⇩t A)" shows "fv⇩s⇩e⇩t (α ` (⋃(vars_transaction ` set P))) ∩ (⋃(vars_transaction ` set P)) = {}" (is ?A) and "fv⇩s⇩e⇩t (α ` vars⇩l⇩s⇩s⇩t A) ∩ vars⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t A ∩ range_vars α = {}" (is ?D1) and "bvars⇩l⇩s⇩s⇩t A ∩ range_vars α = {}" (is ?D2) and "fv⇩l⇩s⇩s⇩t A ∩ range_vars α = {}" (is ?D3) proof - define X where "X ≡ ⋃(vars_transaction ` set P) ∪ vars⇩l⇩s⇩s⇩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 force 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⇩s⇩e⇩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⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t) show ?D1 using 3(3) unfolding X_def by auto thus ?D2 ?D3 by (simp_all add: disjoint_iff_not_equal vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t) qed lemma transaction_renaming_subst_wt: fixes α::"('fun,'atom,'sets,'lbl) prot_subst" assumes "transaction_renaming_subst α P X" shows "wt⇩s⇩u⇩b⇩s⇩t α" proof - { fix x::"('fun,'atom,'sets,'lbl) prot_var" obtain τ n where x: "x = (τ,n)" by force then obtain m where m: "α x = Var (τ,m)" using assms transaction_renaming_subst_is_renaming(1) by force hence "Γ (α x) = Γ⇩v x" using x by (simp add: Γ⇩v_def) } thus ?thesis by (simp add: wt⇩s⇩u⇩b⇩s⇩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⇩t⇩r⇩m (α v)" proof - obtain τ n where "v = (τ, n)" by force then obtain m where "α v = Var (τ, n + Suc m)" using transaction_renaming_subst_is_renaming(1)[OF assms] by force 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⇩t⇩r⇩m⇩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⇩l⇩s⇩s⇩t A)" shows "∃y. α x = Var y ∧ y ∉ ⋃(vars_transaction ` set P) ∪ vars⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩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⇩s⇩s⇩t (S ⋅⇩s⇩s⇩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⇩s⇩s⇩t S" "x ∈ fv (α y)" using fv⇩s⇩s⇩t_subst_obtain_var[OF x] by force 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 force 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⇩l⇩s⇩s⇩t A)" shows "fv⇩s⇩e⇩t (α ` fv⇩l⇩s⇩s⇩t A) ∩ fv⇩l⇩s⇩s⇩t A = {}" proof - have "fv⇩s⇩e⇩t (α ` vars⇩l⇩s⇩s⇩t A) ∩ vars⇩l⇩s⇩s⇩t A = {}" using assms transaction_renaming_subst_vars_disj(2) by blast moreover have "fv⇩l⇩s⇩s⇩t A ⊆ vars⇩l⇩s⇩s⇩t A" by (simp add: vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩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⇩t⇩r⇩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 force 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⇩s⇩u⇩b⇩s⇩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⇩t⇩r⇩m⇩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⇩s⇩e⇩t M" and "∀t ∈ subst_range σ. t ∉ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩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)" "Γ⇩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) Γ⇩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)" "Γ⇩v y = TAtom Value" obtains n where "(σ ∘⇩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 force 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 force 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. (ξ ∘⇩s σ ∘⇩s α) x = Fun c [] ∧ arity c = 0" and "x ∉ fst ` set (transaction_decl T ()) ⟹ x ∈ set (transaction_fresh T) ⟹ ∃c. (ξ ∘⇩s σ ∘⇩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. (ξ ∘⇩s σ ∘⇩s α) x = Fun (Val n) []" and "x ∉ fst ` set (transaction_decl T ()) ⟹ x ∉ set (transaction_fresh T) ⟹ ∃y. (ξ ∘⇩s σ ∘⇩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. (ξ ∘⇩s σ ∘⇩s α) x = Fun c [] ∧ arity c = 0" using subst_compose[of "ξ ∘⇩s σ" α x] subst_compose[of ξ σ x] by simp next assume x: "x ∉ fst ` set (transaction_decl T ())" "x ∈ set (transaction_fresh T)" have *: "(ξ ∘⇩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: "(ξ ∘⇩s σ) x = Fun c []" "¬public c" "arity c = 0" using σ x(2) unfolding transaction_fresh_subst_def by fastforce thus "∃c. (ξ ∘⇩s σ ∘⇩s α) x = Fun c [] ∧ ¬public c ∧ arity c = 0" using subst_compose[of "ξ ∘⇩s σ" α x] subst_compose[of ξ σ x] by simp assume "fst x = TAtom Value" hence "Γ ((ξ ∘⇩s σ) x) = TAtom Value" using * σ Γ⇩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. (ξ ∘⇩sσ ∘⇩s α) x = Fun (Val n) []" using c subst_compose[of "ξ ∘⇩s σ" α x] subst_compose[of ξ σ x] by simp next assume "x ∉ fst ` set (transaction_decl T ())" "x ∉ set (transaction_fresh T)" hence "(ξ ∘⇩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. (ξ ∘⇩s σ ∘⇩s α) x = Var y" using transaction_renaming_subst_is_renaming(1)[OF α] subst_compose[of "ξ ∘⇩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 (ξ ∘⇩s σ ∘⇩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). Γ⇩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 (ξ ∘⇩s σ ∘⇩s α)" "(ξ ∘⇩s σ ∘⇩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 Γ⇩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 Γ⇩v_TAtom'' by auto show "(∃n. t = Fun (Val n) []) ∨ (∃x. t = Var x)" when "∀x ∈ set (transaction_fresh T). Γ⇩v x = TAtom Value" using 0 2 that unfolding Γ⇩v_TAtom'' by auto show "t ∈ subst_range σ" when t': "is_Fun t" proof - obtain x where x: "(σ ∘⇩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⇩l⇩s⇩s⇩t 𝒜)" and α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t 𝒜)" and y: "y ∈ fv ((ξ ∘⇩s σ ∘⇩s α) x)" shows "ξ x = Var x" and "σ x = Var x" and "α x = Var y" and "(ξ ∘⇩s σ ∘⇩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 ((σ ∘⇩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" "(ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t 𝒜)" and α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t 𝒜)" shows "⋃(fv_transaction ` set P) ⊆ subst_domain (ξ ∘⇩s σ ∘⇩s α)" (is ?A) and "fv⇩l⇩s⇩s⇩t 𝒜 ⊆ subst_domain (ξ ∘⇩s σ ∘⇩s α)" (is ?B) and "T' ∈ set P ⟹ fv_transaction T' ⊆ subst_domain (ξ ∘⇩s σ ∘⇩s α)" (is "T' ∈ set P ⟹ ?C") and "T' ∈ set P ⟹ fv⇩l⇩s⇩s⇩t (transaction_strand T' ⋅⇩l⇩s⇩s⇩t (ξ ∘⇩s σ ∘⇩s α)) ⊆ range_vars (ξ ∘⇩s σ ∘⇩s α)" (is "T' ∈ set P ⟹ ?D") proof - have *: "x ∈ subst_domain (ξ ∘⇩s σ ∘⇩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: "(ξ ∘⇩s σ ∘⇩s α) x = (σ ∘⇩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 (σ ∘⇩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 "(σ ∘⇩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⇩s⇩s⇩t (unlabel (transaction_strand T') ⋅⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α) ⊆ range_vars (ξ ∘⇩s σ ∘⇩s α)" when T: "T' ∈ set P" using T fv⇩s⇩s⇩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⇩l⇩s⇩s⇩t 𝒜)" and α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t 𝒜)" shows "fv⇩s⇩e⇩t ((ξ ∘⇩s σ ∘⇩s α) ` (⋃(vars_transaction ` set P))) ∩ (⋃(vars_transaction ` set P)) = {}" (is ?A) and "x ∈ ⋃(vars_transaction ` set P) ⟹ fv ((ξ ∘⇩s σ ∘⇩s α) x) ∩ (⋃(vars_transaction ` set P)) = {}" (is "?B' ⟹ ?B") and "T' ∈ set P ⟹ vars_transaction T' ∩ range_vars (ξ ∘⇩s σ ∘⇩s α) = {}" (is "T' ∈ set P ⟹ ?C1") and "T' ∈ set P ⟹ bvars_transaction T' ∩ range_vars (ξ ∘⇩s σ ∘⇩s α) = {}" (is "T' ∈ set P ⟹ ?C2") and "T' ∈ set P ⟹ fv_transaction T' ∩ range_vars (ξ ∘⇩s σ ∘⇩s α) = {}" (is "T' ∈ set P ⟹ ?C3") and "vars⇩l⇩s⇩s⇩t 𝒜 ∩ range_vars (ξ ∘⇩s σ ∘⇩s α) = {}" (is ?D1) and "bvars⇩l⇩s⇩s⇩t 𝒜 ∩ range_vars (ξ ∘⇩s σ ∘⇩s α) = {}" (is ?D2) and "fv⇩l⇩s⇩s⇩t 𝒜 ∩ range_vars (ξ ∘⇩s σ ∘⇩s α) = {}" (is ?D3) and "range_vars ξ = {}" (is ?E1) and "range_vars σ = {}" (is ?E2) and "range_vars (ξ ∘⇩s σ ∘⇩s α) ⊆ range_vars α" (is ?E3) proof - note 0 = transaction_renaming_subst_vars_disj[OF α] define θ where "θ = ξ ∘⇩s σ ∘⇩s α" show ?A proof (cases "fv⇩s⇩e⇩t ((ξ ∘⇩s σ ∘⇩s α) ` (⋃(vars_transaction ` set P))) = {}") case False hence "∀x ∈ (⋃(vars_transaction ` set P)). (ξ ∘⇩s σ ∘⇩s α) x = α x ∨ fv ((ξ ∘⇩s σ ∘⇩s α) x) = {}" using transaction_decl_fresh_renaming_substs_range''[OF ξ σ α] by auto thus ?thesis using 0(1) unfolding θ_def[symmetric] 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 "ξ ∘⇩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⇩l⇩s⇩s⇩t 𝒜)" and α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t 𝒜)" and "bvars⇩l⇩s⇩s⇩t S ∩ subst_domain ξ = {}" and "bvars⇩l⇩s⇩s⇩t S ∩ subst_domain σ = {}" and "bvars⇩l⇩s⇩s⇩t S ∩ subst_domain α = {}" shows "subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (S ⋅⇩l⇩s⇩s⇩t (ξ ∘⇩s σ ∘⇩s α))) = subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t S) ⋅⇩s⇩e⇩t (ξ ∘⇩s σ ∘⇩s α)" proof - have 1: "∀x ∈ fv⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t S). (∃f. (ξ ∘⇩s σ ∘⇩s α) x = Fun f []) ∨ (∃y. (ξ ∘⇩s σ ∘⇩s α) x = Var y)" using transaction_decl_fresh_renaming_substs_range'[OF ξ σ α] by blast have 2: "bvars⇩l⇩s⇩s⇩t S ∩ subst_domain (ξ ∘⇩s σ ∘⇩s α) = {}" using assms(4-6) subst_domain_compose[of ξ σ] subst_domain_compose[of "ξ ∘⇩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⇩s⇩u⇩b⇩s⇩t (ξ ∘⇩s σ ∘⇩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⇩t⇩r⇩m⇩s (subst_range (ξ ∘⇩s σ ∘⇩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 "ξ ∘⇩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⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))" shows "∃y ∈ fv_transaction T - set (transaction_fresh T). (ξ ∘⇩s σ ∘⇩s α) y = Var x" proof - have "x ∈ fv⇩s⇩s⇩t (unlabel (transaction_strand T) ⋅⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" using x fv⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq[of "transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α"] unlabel_subst[of "transaction_strand T" "ξ ∘⇩s σ ∘⇩s α"] by argo then obtain y where "y ∈ fv_transaction T" "x ∈ fv ((ξ ∘⇩s σ ∘⇩s α) y)" by (metis fv⇩s⇩s⇩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). Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩v x = TAtom (Atom a))" and t: "t ∈ subst_range (ξ ∘⇩s σ ∘⇩s α)" shows "∄n. t = attack⟨n⟩" proof - note ξσα_wt = transaction_decl_fresh_renaming_substs_wt[OF ξ σ α] obtain x where x: "(ξ ∘⇩s σ ∘⇩s α) x = t" using t by auto have x_type: "Γ (Var x) = Γ (Var x ⋅ ξ)" "Γ (Var x) = Γ (Var x ⋅ ξ ∘⇩s σ ∘⇩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) 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⇩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 eval_term.simps(1)[of _ x ξ] eval_term.simps(1)[of _ x "ξ ∘⇩s σ ∘⇩s α"] by (cases "Γ⇩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 ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)) ⟹ ∃ts' s. send⟨ts'⟩ ∈ set (unlabel (transaction_send T)) ∧ occurs s ∈ set ts' ∧ t = s ⋅ ξ ∘⇩s σ ∘⇩s α" (is "?A ⟹ ?A'") and "receive⟨ts⟩ ∈ set (unlabel (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)) ⟹ ∃ts' s. receive⟨ts'⟩ ∈ set (unlabel (transaction_receive T)) ∧ occurs s ∈ set ts' ∧ t = s ⋅ ξ ∘⇩s σ ∘⇩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 ⋅ ξ ∘⇩s σ ∘⇩s α" using t stateful_strand_step_mem_substD(1)[ of ts "unlabel (transaction_strand T)" "ξ ∘⇩s σ ∘⇩s α"] unlabel_subst[of "transaction_strand T" "ξ ∘⇩s σ ∘⇩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 ⋅ ξ ∘⇩s σ ∘⇩s α = Fun c []) ∨ (∃y. s ⋅ ξ ∘⇩s σ ∘⇩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 ⋅ ξ ∘⇩s σ ∘⇩s α = Fun OccursSec []" "T ! 1 ⋅ ξ ∘⇩s σ ∘⇩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 force hence "t = u ⋅ ξ ∘⇩s σ ∘⇩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 ⋅ ξ ∘⇩s σ ∘⇩s α" using t stateful_strand_step_mem_substD(2)[ of ts "unlabel (transaction_strand T)" "ξ ∘⇩s σ ∘⇩s α"] unlabel_subst[of "transaction_strand T" "ξ ∘⇩s σ ∘⇩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 ⋅ ξ ∘⇩s σ ∘⇩s α = Fun c []) ∨ (∃y. s ⋅ ξ ∘⇩s σ ∘⇩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 ⋅ ξ ∘⇩s σ ∘⇩s α = Fun OccursSec []" "T ! 1 ⋅ ξ ∘⇩s σ ∘⇩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 force hence "t = u ⋅ ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t A)" shows "transaction_fresh_subst σ (transaction_proj n T) (trms⇩l⇩s⇩s⇩t (proj n A))" using assms transaction_proj_fresh_eq[of n T] contra_subsetD[OF subterms⇩s⇩e⇩t_mono[OF transaction_proj_trms_subset[of n T]]] contra_subsetD[OF subterms⇩s⇩e⇩t_mono[OF trms⇩s⇩s⇩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⇩l⇩s⇩s⇩t A)" shows "transaction_renaming_subst α (map (transaction_proj n) P) (vars⇩l⇩s⇩s⇩t (proj n A))" proof - let ?X = "λP A. ⋃(vars_transaction ` set P) ∪ vars⇩l⇩s⇩s⇩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'⇩s⇩s⇩t (fst ` set (transaction_decl T ()) ∪ set (transaction_fresh T)) (unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand T)))" and ξ: "transaction_decl_subst ξ T" and σ: "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t 𝒜)" and α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t 𝒜)" shows "wf'⇩s⇩s⇩t {} (unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)))" proof - have 0: "range_vars ξ ∩ bvars⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (transaction_strand T)) = {}" "range_vars σ ∩ bvars⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (transaction_strand T) ⋅⇩l⇩s⇩s⇩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⇩s⇩e⇩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⇩s⇩e⇩t (α ` fv⇩s⇩e⇩t (σ ` fv⇩s⇩e⇩t (ξ ` ?X))) = {}" using 0(3-5) 1 by auto hence "wf'⇩s⇩s⇩t {} (((unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand T)) ⋅⇩s⇩s⇩t ξ) ⋅⇩s⇩s⇩t σ) ⋅⇩s⇩s⇩t α)" by (metis wf⇩s⇩s⇩t_subst_apply[OF wf⇩s⇩s⇩t_subst_apply[OF wf⇩s⇩s⇩t_subst_apply[OF T]]]) thus ?thesis using dual⇩l⇩s⇩s⇩t_subst unlabel_subst labeled_stateful_strand_subst_comp[OF 0(1), of "σ ∘⇩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 "θ ≡ ξ ∘⇩s σ ∘⇩s α" assumes T_adm: "admissible_transaction' T" and ξσα: "transaction_decl_subst ξ T" "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t 𝒜)" "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t (g B)) ∧ transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (g B)) ∧ transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)) A" and "T ∈ set P" "transaction_decl_subst ξ T" "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t B)" "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t B)" and "b ∈ set (transaction_strand T)" "a = dual⇩l⇩s⇩s⇩t⇩p b ⋅⇩l⇩s⇩s⇩t⇩p ξ ∘⇩s σ ∘⇩s α" "fst a = fst b" proof - define f where "f ≡ λ(T,ξ,σ::('fun,'atom,'sets,'lbl) prot_subst,α). dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t (g B)) ∧ transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t⇩p b ⋅⇩l⇩s⇩s⇩t⇩p ξ ∘⇩s σ ∘⇩s α" "fst a = fst b" using T(2) dual⇩l⇩s⇩s⇩t_subst[of "transaction_strand T" "ξ ∘⇩s σ ∘⇩s α"] dual⇩l⇩s⇩s⇩t_memberD'[of a "transaction_strand T" "ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t (g B))" "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t 𝒜) ⟹ transaction_fresh_subst σ T2 (trms⇩l⇩s⇩s⇩t ℬ)" (is "?B1 ⟹ ?B2") "transaction_renaming_subst α P1 (vars⇩l⇩s⇩s⇩t 𝒜) ⟹ transaction_renaming_subst α P2 (vars⇩l⇩s⇩s⇩t ℬ)" (is "?C1 ⟹ ?C2") "transaction_renaming_subst α P2 (vars⇩l⇩s⇩s⇩t 𝒜) ⟹ transaction_renaming_subst α P1 (vars⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T1 ⋅⇩l⇩s⇩s⇩t θ)) = unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand T2 ⋅⇩l⇩s⇩s⇩t θ))" when "transaction_unlabel_eq T1 T2" for T1 T2 θ using 0[OF that] unlabel_subst[of _ θ] dual⇩l⇩s⇩s⇩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 "ξ ∘⇩s σ ∘⇩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 "ξ ∘⇩s σ ∘⇩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⇩t⇩r⇩m⇩s: assumes "∀T ∈ set P. wf⇩t⇩r⇩m⇩s (trms_transaction T)" and "𝒜 ∈ reachable_constraints P" shows "wf⇩t⇩r⇩m⇩s (trms⇩l⇩s⇩s⇩t 𝒜)" using assms(2) proof (induction 𝒜 rule: reachable_constraints.induct) case (step 𝒜 T ξ σ α) have "wf⇩t⇩r⇩m⇩s (trms_transaction T)" using assms(1) step.hyps(2) by blast hence "wf⇩t⇩r⇩m⇩s (trms_transaction T ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩s α)" using transaction_decl_fresh_renaming_substs_range_wf_trms[OF step.hyps(3-5)] by (metis wf_trms_subst) hence "wf⇩t⇩r⇩m⇩s (trms⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))" using wf⇩t⇩r⇩m⇩s_trms⇩s⇩s⇩t_subst unlabel_subst[of "transaction_strand T" "ξ ∘⇩s σ ∘⇩s α"] by metis hence "wf⇩t⇩r⇩m⇩s (trms⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)))" using trms⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq by blast thus ?case using step.IH unlabel_append[of 𝒜] trms⇩s⇩s⇩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). Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩v x = TAtom (Atom a))" shows "Γ⇩v ` fv⇩l⇩s⇩s⇩t 𝒜 ⊆ (⋃T ∈ set P. Γ⇩v ` fv_transaction T)" (is "?A 𝒜") and "Γ⇩v ` bvars⇩l⇩s⇩s⇩t 𝒜 ⊆ (⋃T ∈ set P. Γ⇩v ` bvars_transaction T)" (is "?B 𝒜") and "Γ⇩v ` vars⇩l⇩s⇩s⇩t 𝒜 ⊆ (⋃T ∈ set P. Γ⇩v ` vars_transaction T)" (is "?C 𝒜") using 𝒜 proof (induction 𝒜 rule: reachable_constraints.induct) case (step 𝒜 T ξ σ α) define T' where "T' ≡ dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" note 2 = transaction_decl_fresh_renaming_substs_wt[OF step.hyps(3-5)] have 3: "∀t ∈ subst_range (ξ ∘⇩s σ ∘⇩s α). fv t = {} ∨ (∃x. t = Var x)" using transaction_decl_fresh_renaming_substs_range'(1)[OF step.hyps(3-5)] by fastforce have "fv⇩l⇩s⇩s⇩t T' = fv⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" "bvars⇩l⇩s⇩s⇩t T' = bvars⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" "vars⇩l⇩s⇩s⇩t T' = vars⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" unfolding T'_def by (metis fv⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq, metis bvars⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq, metis vars⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq) hence "Γ ` Var ` fv⇩l⇩s⇩s⇩t T' ⊆ Γ ` Var ` fv_transaction T" "Γ ` Var ` bvars⇩l⇩s⇩s⇩t T' = Γ ` Var ` bvars_transaction T" "Γ ` Var ` vars⇩l⇩s⇩s⇩t T' ⊆ Γ ` Var ` vars_transaction T" using wt_subst_lsst_vars_type_subset[OF 2 3, of "transaction_strand T"] by argo+ hence "Γ⇩v ` fv⇩l⇩s⇩s⇩t T' ⊆ Γ⇩v ` fv_transaction T" "Γ⇩v ` bvars⇩l⇩s⇩s⇩t T' = Γ⇩v ` bvars_transaction T" "Γ⇩v ` vars⇩l⇩s⇩s⇩t T' ⊆ Γ⇩v ` vars_transaction T" by (metis Γ⇩v_Var_image)+ hence 4: "Γ⇩v ` fv⇩l⇩s⇩s⇩t T' ⊆ (⋃T ∈ set P. Γ⇩v ` fv_transaction T)" "Γ⇩v ` bvars⇩l⇩s⇩s⇩t T' ⊆ (⋃T ∈ set P. Γ⇩v ` bvars_transaction T)" "Γ⇩v ` vars⇩l⇩s⇩s⇩t T' ⊆ (⋃T ∈ set P. Γ⇩v ` vars_transaction T)" using step.hyps(2) by fast+ have 5: "Γ⇩v ` fv⇩l⇩s⇩s⇩t (𝒜 @ T') = (Γ⇩v ` fv⇩l⇩s⇩s⇩t 𝒜) ∪ (Γ⇩v ` fv⇩l⇩s⇩s⇩t T')" "Γ⇩v ` bvars⇩l⇩s⇩s⇩t (𝒜 @ T') = (Γ⇩v ` bvars⇩l⇩s⇩s⇩t 𝒜) ∪ (Γ⇩v ` bvars⇩l⇩s⇩s⇩t T')" "Γ⇩v ` vars⇩l⇩s⇩s⇩t (𝒜 @ T') = (Γ⇩v ` vars⇩l⇩s⇩s⇩t 𝒜) ∪ (Γ⇩v ` vars⇩l⇩s⇩s⇩t T')" using unlabel_append[of 𝒜 T'] fv⇩s⇩s⇩t_append[of "unlabel 𝒜" "unlabel T'"] bvars⇩s⇩s⇩t_append[of "unlabel 𝒜" "unlabel T'"] vars⇩s⇩s⇩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⇩s⇩s⇩t_def) } { case 2 thus ?case using step.IH(2) 4(2) 5(2) unfolding T'_def by (simp del: subst_subst_compose bvars⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T) = {}" shows "bvars⇩l⇩s⇩s⇩t 𝒜 = {}" using assms proof (induction) case init then show ?case unfolding unlabel_def by auto next case (step 𝒜 T ξ σ α) then have "bvars⇩l⇩s⇩s⇩t 𝒜 = {}" by metis moreover have "bvars⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)) = {}" using step by (metis bvars⇩l⇩s⇩s⇩t_subst bvars⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq) ultimately show ?case using bvars⇩s⇩s⇩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⇩l⇩s⇩s⇩t 𝒜 ∩ bvars⇩l⇩s⇩s⇩t 𝒜 = {}" proof - let ?X = "⋃T ∈ set P. bvars_transaction T" note 0 = admissible_transactions_fv_bvars_disj[OF P] have 1: "bvars⇩l⇩s⇩s⇩t 𝒜 ⊆ ?X" using 𝒜_reach proof (induction 𝒜 rule: reachable_constraints.induct) case (step 𝒜 T ξ σ α) have "bvars⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)) = bvars_transaction T" using bvars⇩s⇩s⇩t_subst[of "unlabel (transaction_strand T)" "ξ ∘⇩s σ ∘⇩s α"] bvars⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq[of "transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α"] dual⇩l⇩s⇩s⇩t_subst[of "transaction_strand T" "ξ ∘⇩s σ ∘⇩s α"] unlabel_subst[of "transaction_strand T" "ξ ∘⇩s σ ∘⇩s α"] by argo hence "bvars⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)) ⊆ ?X" using step.hyps(2) by blast thus ?case using step.IH bvars⇩s⇩s⇩t_append by auto qed (simp add: unlabel_def) have 2: "fv⇩l⇩s⇩s⇩t 𝒜 ∩ ?X = {}" using 𝒜_reach proof (induction 𝒜 rule: reachable_constraints.induct) case (step 𝒜 T ξ σ α) have "x ≠ y" when x: "x ∈ ?X" and y: "y ∈ fv⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" for x y proof - obtain y' where y': "y' ∈ fv_transaction T" "y ∈ fv ((ξ ∘⇩s σ ∘⇩s α) y')" using y unlabel_subst[of "transaction_strand T" "ξ ∘⇩s σ ∘⇩s α"] by (metis fv⇩s⇩s⇩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⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t by fast qed hence "fv⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α) ∩ ?X = {}" by blast thus ?case using step.IH fv⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq[of "transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α"] dual⇩l⇩s⇩s⇩t_subst[of "transaction_strand T" "ξ ∘⇩s σ ∘⇩s α"] unlabel_subst[of "transaction_strand T" "ξ ∘⇩s σ ∘⇩s α"] fv⇩s⇩s⇩t_append[of "unlabel 𝒜" "unlabel (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t 𝒜" shows "Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩v x = TAtom (Atom a))" proof - have 𝒜_wf⇩t⇩r⇩m⇩s: "wf⇩t⇩r⇩m⇩s (trms⇩l⇩s⇩s⇩t 𝒜)" by (metis reachable_constraints_wf⇩t⇩r⇩m⇩s admissible_transactions_wf⇩t⇩r⇩m⇩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). Γ⇩v x = TAtom Value" using protocol_transaction_vars_TAtom_typed(3) P by blast hence *: "Γ⇩v ` vars⇩l⇩s⇩s⇩t 𝒜 ⊆ (⋃T∈set P. Γ⇩v ` vars_transaction T)" using reachable_constraints_var_types_in_transactions[of 𝒜 P, OF 𝒜_reach] by auto have "Γ⇩v ` vars⇩l⇩s⇩s⇩t 𝒜 ⊆ TAtom ` insert Value (range Atom)" proof - have "Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩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. Γ⇩v ` vars_transaction T) ⊆ TAtom ` insert Value (range Atom)" using P by blast thus "Γ⇩v ` vars⇩l⇩s⇩s⇩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). Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩v x = TAtom (Atom a))" "∀T ∈ set P. ∀x ∈ vars_transaction T. ¬TAtom AttackType ⊑ Γ⇩v x" and x: "x ∈ vars⇩l⇩s⇩s⇩t 𝒜" shows "¬TAtom AttackType ⊑ Γ⇩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⇩l⇩s⇩s⇩t 𝒜" and "Γ⇩v x = TAtom Value" shows "x ∈ fv⇩l⇩s⇩s⇩t 𝒜" proof - have "∀T∈set P. bvars_transaction T = {}" using P admissible_transactionE(4) by metis hence 𝒜_no_bvars: "bvars⇩l⇩s⇩s⇩t 𝒜 = {}" using reachable_constraints_no_bvars[OF 𝒜_reach] by metis thus ?thesis using x vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩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⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (𝒜 ⋅⇩l⇩s⇩s⇩t ℐ)) = (subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜)) ⋅⇩s⇩e⇩t ℐ" proof - have 𝒜_wf⇩t⇩r⇩m⇩s: "wf⇩t⇩r⇩m⇩s (trms⇩l⇩s⇩s⇩t 𝒜)" by (metis reachable_constraints_wf⇩t⇩r⇩m⇩s admissible_transactions_wf⇩t⇩r⇩m⇩s P 𝒜_reach) from ℐ have ℐ': "welltyped_constraint_model ℐ 𝒜" using welltyped_constraint_model_prefix by auto have 1: "∀x ∈ fv⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜). (∃f. ℐ x = Fun f []) ∨ (∃y. ℐ x = Var y)" proof fix x assume xa: "x ∈ fv⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩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⇩t⇩r⇩m (Fun f T)" using ℐ unfolding welltyped_constraint_model_def constraint_model_def using wf_trm_subst_rangeD by metis moreover have "x ∈ vars⇩l⇩s⇩s⇩t 𝒜" using xa var_subterm_trms⇩s⇩s⇩t_is_vars⇩s⇩s⇩t[of x "unlabel 𝒜"] vars_iff_subtermeq[of x] by auto hence "∃a. Γ⇩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⇩s⇩u⇩b⇩s⇩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⇩l⇩s⇩s⇩t 𝒜 = {}" using reachable_constraints_no_bvars assms by metis then have 2: "bvars⇩l⇩s⇩s⇩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). Γ⇩v x = TAtom Value" and f: "f ∈ ⋃(funs_term ` trms⇩l⇩s⇩s⇩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) ⋅⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α" let ?T'' = "transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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 Γ⇩v_TAtom'' by fast show ?thesis using step proof (cases "f ∈ ⋃(funs_term ` trms⇩l⇩s⇩s⇩t 𝒜)") case False then obtain t where t: "t ∈ trms⇩s⇩s⇩t ?T'" "f ∈ funs_term t" using step.prems trms⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq[of ?T''] trms⇩s⇩s⇩t_append[of "unlabel 𝒜" "unlabel (dual⇩l⇩s⇩s⇩t ?T'')"] unlabel_append[of 𝒜 "dual⇩l⇩s⇩s⇩t ?T''"] unlabel_subst[of "transaction_strand T"] by fastforce show ?thesis using trms⇩s⇩s⇩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 ((ξ ∘⇩s σ ∘⇩s α) x)" then obtain x where "x ∈ fv_transaction T" "f ∈ funs_term ((ξ ∘⇩s σ ∘⇩s α) x)" by force thus ?thesis using transaction_decl_fresh_renaming_substs_range'(3)[ OF step.hyps(3-5) _ ξ_empty T_fresh[OF step.hyps(2), unfolded Γ⇩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⇩l⇩s⇩s⇩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 Γ⇩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⇩l⇩s⇩s⇩t A" shows "∃n. t = Fun (Val n) []" using 𝒜_reach occ proof (induction A rule: reachable_constraints.induct) case (step A T ξ σ α) define θ where "θ ≡ ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t A") case False hence "occurs t ∈ ik⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ))" using step.prems unfolding θ_def by simp hence "∃ts. occurs t ∈ set ts ∧ receive⟨ts⟩ ∈ set (unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ)))" unfolding ik⇩s⇩s⇩t_def by force hence "∃ts. occurs t ∈ set ts ∧ send⟨ts⟩ ∈ set (unlabel (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ))" using dual⇩l⇩s⇩s⇩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: "Γ⇩v x = TAtom Value" "x ∈ fv⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩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)" "(ξ ∘⇩s σ ∘⇩s α) y = Var x" using transaction_decl_fresh_renaming_substs_fv[OF step.hyps(3-5), of x] step.prems(1) fv⇩s⇩s⇩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" "Γ⇩v y = TAtom Value" using x(1) Γ⇩v_TAtom''[of x] Γ⇩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 ⋅⇩l⇩i⇩s⇩t ξ ∘⇩s σ ∘⇩s α⟩ ∈ set (unlabel (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))" using unlabel_subst[of "transaction_strand T" "ξ ∘⇩s σ ∘⇩s α"] stateful_strand_step_subst_inI(2)[of _ _ "ξ ∘⇩s σ ∘⇩s α"] by force have "occurs (Var y) ⋅ ξ ∘⇩s σ ∘⇩s α = occurs (Var x)" using y(2) by (auto simp del: subst_subst_compose) hence **: "occurs (Var x) ∈ set ts ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩s α" using ts(1) by force have "send⟨ts ⋅⇩l⇩i⇩s⇩t ξ ∘⇩s σ ∘⇩s α⟩ ∈ set (unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)))" using * dual⇩l⇩s⇩s⇩t_unlabel_steps_iff(2) by blast thus ?thesis using ** unlabel_append[of A] by force qed qed simp lemma reachable_constraints_db⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t 𝒜 ℐ)" shows "∃ss. s = Fun (Set ss) []" using 𝒜 d proof (induction) case (step 𝒜 TT ξ σ α) let ?TT = "transaction_strand TT ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α" let ?TTu = "unlabel ?TT" let ?TTd = "dual⇩l⇩s⇩s⇩t ?TT" let ?TTdu = "unlabel ?TTd" from step(6) have "(t, s) ∈ set (db'⇩s⇩s⇩t ?TTdu ℐ (db'⇩s⇩s⇩t (unlabel 𝒜) ℐ []))" by (metis db⇩s⇩s⇩t_append db⇩s⇩s⇩t_def step.prems unlabel_append) hence "(t, s) ∈ set (db'⇩s⇩s⇩t (unlabel 𝒜) ℐ []) ∨ (∃t' s'. insert⟨t',s'⟩ ∈ set ?TTdu ∧ t = t' ⋅ ℐ ∧ s = s' ⋅ ℐ)" using db⇩s⇩s⇩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 ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" using dual⇩l⇩s⇩s⇩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)) ⋅⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" by (simp add: subst_lsst_unlabel) hence "insert⟨t',s'⟩ ∈ (λx. x ⋅⇩s⇩s⇩t⇩p ξ ∘⇩s σ ∘⇩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 ⋅⇩s⇩s⇩t⇩p ξ ∘⇩s σ ∘⇩s α = insert⟨t',s'⟩" by auto hence "∃t'' s''. insert⟨t'',s''⟩ ∈ set (unlabel (transaction_strand TT)) ∧ t' = t'' ⋅ ξ ∘⇩s σ ∘⇩s α ∧ s' = s'' ⋅ ξ ∘⇩s σ ∘⇩s α" by (cases u) auto then obtain t'' s'' where t''s''_p: "insert⟨t'',s''⟩ ∈ set (unlabel (transaction_strand TT)) ∧ t' = t'' ⋅ ξ ∘⇩s σ ∘⇩s α ∧ s' = s'' ⋅ ξ ∘⇩s σ ∘⇩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'' ⋅ ξ ∘⇩s σ ∘⇩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⇩s⇩s⇩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⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (A@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ)) = (ik⇩l⇩s⇩s⇩t A) ∪ (trms⇩l⇩s⇩s⇩t (transaction_send T) ⋅⇩s⇩e⇩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⇩l⇩s⇩s⇩t A") case False hence "occurs t ∈ trms⇩l⇩s⇩s⇩t (transaction_send T) ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩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 ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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 ⋅ ξ ∘⇩s σ ∘⇩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 "ξ ∘⇩s σ ∘⇩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⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I). OccursFact ∉ ⋃(funs_term ` set (snd (Ana s)))" (is "?A A") and "∀s ∈ subterms⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I). OccursSec ∉ ⋃(funs_term ` set (snd (Ana s)))" (is "?B A") and "Fun OccursSec [] ∉ ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I" (is "?C A") and "∀x ∈ vars⇩l⇩s⇩s⇩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⇩s⇩u⇩b⇩s⇩t I" by (metis ℐ welltyped_constraint_model_def) have ℐ_wf⇩t⇩r⇩m⇩s: "wf⇩t⇩r⇩m⇩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⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T)) ⊆ vars_transaction T" "fv⇩s⇩e⇩t (subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T))) = fv⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T))" for T::"('fun,'atom,'sets,'lbl) prot_transaction" using fv_trms⇩s⇩s⇩t_subset(1)[of "unlabel (transaction_send T)"] vars_transaction_unfold fv_subterms_set[of "trms⇩l⇩s⇩s⇩t (transaction_send T)"] by blast+ have 0: "∀x ∈ fv⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T)). ∃a. Γ (Var x) = TAtom a" "∀x ∈ fv⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T)). Γ (Var x) ≠ TAtom OccursSecType" "∀x ∈ fv⇩s⇩e⇩t (subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T))). ∃a. Γ (Var x) = TAtom a" "∀x ∈ fv⇩s⇩e⇩t (subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (A@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ)) ⋅⇩s⇩e⇩t I = (ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I) ∪ (trms⇩l⇩s⇩s⇩t (transaction_send T) ⋅⇩s⇩e⇩t θ ⋅⇩s⇩e⇩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⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T) ⋅⇩s⇩e⇩t θ ⋅⇩s⇩e⇩t I) = subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T)) ⋅⇩s⇩e⇩t θ ⋅⇩s⇩e⇩t I" when "T ∈ set P" and θ: "wt⇩s⇩u⇩b⇩s⇩t θ" "wf⇩t⇩r⇩m⇩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⇩t⇩r⇩m⇩s]] by auto have 3: "wt⇩s⇩u⇩b⇩s⇩t (ξ ∘⇩s σ ∘⇩s α)" "wf⇩t⇩r⇩m⇩s (subst_range (ξ ∘⇩s σ ∘⇩s α))" when "T ∈ set P" "transaction_decl_subst ξ T" "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t A)" "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩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⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩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⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩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 force 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 (ξ ∘⇩s σ ∘⇩s α))" "OccursSec ∉ ⋃(funs_term ` subst_range (ξ ∘⇩s σ ∘⇩s α))" when ξσα: "transaction_decl_subst ξ T" "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t A)" "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩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 (ξ ∘⇩s σ ∘⇩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 (ξ ∘⇩s σ ∘⇩s α))" "OccursSec ∉ ⋃(funs_term ` subst_range (ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t A)" "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t A)" and x: "Var x ∈ trms⇩l⇩s⇩s⇩t (transaction_send T) ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t (transaction_send T)" "t ⋅ (ξ ∘⇩s σ ∘⇩s α) = Var x" using x by force 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 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⇩l⇩s⇩s⇩t A)" "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t A)" and x: "x ∈ fv⇩s⇩e⇩t ((ξ ∘⇩s σ ∘⇩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 ((ξ ∘⇩s σ ∘⇩s α) y)" using x by auto hence y': "(ξ ∘⇩s σ ∘⇩s α) y = Var x" using transaction_decl_fresh_renaming_substs_range'(3)[ OF ξσα _ ξ_empty[OF T ξσα(1)] T_fresh_type[OF T]] by (cases "(ξ ∘⇩s σ ∘⇩s α) y ∈ subst_range (ξ ∘⇩s σ ∘⇩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 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⇩l⇩s⇩s⇩t A)" "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t A)" and x: "Var x ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T)) ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩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⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T))" "t ⋅ (ξ ∘⇩s σ ∘⇩s α) = Var x" using x by force 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 eval_term.simps(1)) thus "I x ≠ Fun OccursSec []" "∄t. I x = occurs t" by auto qed have s_fv: "fv s ⊆ fv⇩s⇩e⇩t ((ξ ∘⇩s σ ∘⇩s α) ` vars_transaction T)" when s: "s ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T)) ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩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⇩s⇩e⇩t (subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T)) ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩s α)" using s by auto hence *: "x ∈ fv⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T) ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩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⇩s⇩e⇩t ((ξ ∘⇩s σ ∘⇩s α) ` vars⇩l⇩s⇩s⇩t (transaction_send T))" proof - obtain t where t: "t ∈ trms⇩l⇩s⇩s⇩t (transaction_send T)" "x ∈ fv (t ⋅ ξ ∘⇩s σ ∘⇩s α)" using * by fastforce hence "fv t ⊆ vars⇩l⇩s⇩s⇩t (transaction_send T)" using fv_trms⇩s⇩s⇩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⇩s⇩e⇩t ((ξ ∘⇩s σ ∘⇩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⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T)). OccursFact ∉ ⋃(funs_term ` set (snd (Ana s)))" using 4[OF step.hyps(2)] by blast have "∀s ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T)) ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩s α ⋅⇩s⇩e⇩t I. OccursFact ∉ ⋃(funs_term ` set (snd (Ana s)))" proof fix t assume t: "t ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T)) ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩s α ⋅⇩s⇩e⇩t I" then obtain s u where su: "s ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T)) ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩s α" "s ⋅ I = t" "u ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T))" "u ⋅ ξ ∘⇩s σ ∘⇩s α = s" by force obtain Ku Tu where KTu: "Ana u = (Ku,Tu)" by force have *: "OccursFact ∉ ⋃(funs_term ` set Tu)" "OccursFact ∉ ⋃(funs_term ` subst_range (ξ ∘⇩s σ ∘⇩s α))" "OccursFact ∉ ⋃(funs_term ` ⋃(((set ∘ snd ∘ Ana) ` subst_range (ξ ∘⇩s σ ∘⇩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 ⋅⇩l⇩i⇩s⇩t ξ ∘⇩s σ ∘⇩s α))" proof - { fix f assume f: "f ∈ ⋃(funs_term ` set (Tu ⋅⇩l⇩i⇩s⇩t ξ ∘⇩s σ ∘⇩s α))" then obtain tf where tf: "tf ∈ set Tu" "f ∈ funs_term (tf ⋅ ξ ∘⇩s σ ∘⇩s α)" by force hence "f ∈ funs_term tf ∨ f ∈ ⋃(funs_term ` subst_range (ξ ∘⇩s σ ∘⇩s α))" using funs_term_subst[of tf "ξ ∘⇩s σ ∘⇩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 = (ξ ∘⇩s σ ∘⇩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 "ξ ∘⇩s σ ∘⇩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) ⋅⇩l⇩i⇩s⇩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⇩s⇩e⇩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⇩s⇩e⇩t (set (snd (Ana s))). g ∈ funs_term (I x)" then obtain x where x: "x ∈ fv⇩s⇩e⇩t (set (snd (Ana s)))" "g ∈ funs_term (I x)" by force have "x ∈ fv s" using x(1) Ana_vars(2)[of s] by (cases "Ana s") auto hence "x ∈ fv⇩s⇩e⇩t ((ξ ∘⇩s σ ∘⇩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⇩t⇩r⇩m⇩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 "ξ ∘⇩s σ ∘⇩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⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T)) ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩s α ⋅⇩s⇩e⇩t I. OccursSec ∉ ⋃(funs_term ` set (snd (Ana s)))" proof fix t assume t: "t ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T)) ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩s α ⋅⇩s⇩e⇩t I" then obtain s u where su: "s ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T)) ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩s α" "s ⋅ I = t" "u ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T))" "u ⋅ ξ ∘⇩s σ ∘⇩s α = s" by force obtain Ku Tu where KTu: "Ana u = (Ku,Tu)" by force have *: "OccursSec ∉ ⋃(funs_term ` set Tu)" "OccursSec ∉ ⋃(funs_term ` subst_range (ξ ∘⇩s σ ∘⇩s α))" "OccursSec ∉ ⋃(funs_term ` ⋃(((set ∘ snd ∘ Ana) ` subst_range (ξ ∘⇩s σ ∘⇩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 ⋅⇩l⇩i⇩s⇩t ξ ∘⇩s σ ∘⇩s α))" proof - { fix f assume f: "f ∈ ⋃(funs_term ` set (Tu ⋅⇩l⇩i⇩s⇩t ξ ∘⇩s σ ∘⇩s α))" then obtain tf where tf: "tf ∈ set Tu" "f ∈ funs_term (tf ⋅ ξ ∘⇩s σ ∘⇩s α)" by force hence "f ∈ funs_term tf ∨ f ∈ ⋃(funs_term ` subst_range (ξ ∘⇩s σ ∘⇩s α))" using funs_term_subst[of tf "ξ ∘⇩s σ ∘⇩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 = (ξ ∘⇩s σ ∘⇩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 "ξ ∘⇩s σ ∘⇩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) ⋅⇩l⇩i⇩s⇩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⇩s⇩e⇩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⇩s⇩e⇩t (set (snd (Ana s))). g ∈ funs_term (I x)" then obtain x where x: "x ∈ fv⇩s⇩e⇩t (set (snd (Ana s)))" "g ∈ funs_term (I x)" by force have "x ∈ fv s" using x(1) Ana_vars(2)[of s] by (cases "Ana s") auto hence "x ∈ fv⇩s⇩e⇩t ((ξ ∘⇩s σ ∘⇩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⇩t⇩r⇩m⇩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 "ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t (transaction_send T)" using admissible_transaction_occurs_checksE5[OF T_occ[OF step.hyps(2)]] by blast have **: "Fun OccursSec [] ∉ subst_range (ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t (transaction_send T) ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩s α ⋅⇩s⇩e⇩t I" proof assume "Fun OccursSec [] ∈ trms⇩l⇩s⇩s⇩t (transaction_send T) ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩s α ⋅⇩s⇩e⇩t I" then obtain s where "s ∈ trms⇩l⇩s⇩s⇩t (transaction_send T) ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩s α" "s ⋅ I = Fun OccursSec []" by force moreover have "Fun OccursSec [] ∉ trms⇩l⇩s⇩s⇩t (transaction_send T) ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩s α" proof assume "Fun OccursSec [] ∈ trms⇩l⇩s⇩s⇩t (transaction_send T) ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩s α" then obtain u where "u ∈ trms⇩l⇩s⇩s⇩t (transaction_send T)" "u ⋅ ξ ∘⇩s σ ∘⇩s α = Fun OccursSec []" by force 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 "ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))" hence x': "x ∈ vars⇩s⇩s⇩t (unlabel (transaction_strand T) ⋅⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" by (metis vars⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq unlabel_subst) hence "x ∈ vars_transaction T ∨ x ∈ fv⇩s⇩e⇩t ((ξ ∘⇩s σ ∘⇩s α) ` vars_transaction T)" using vars⇩s⇩s⇩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⇩l⇩s⇩s⇩t A" "t ⋅ I = occurs s" shows "∃u. t = occurs u" proof - have "wt⇩s⇩u⇩b⇩s⇩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: "Γ⇩v ` fv⇩l⇩s⇩s⇩t A ⊆ (⋃T ∈ set P. Γ⇩v ` fv_transaction T)" "∀T ∈ set P. ∀x ∈ fv_transaction T. Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩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⇩s⇩s⇩t_is_fv⇩s⇩s⇩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⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t A)" using Fun t(1) by force hence "y ∈ vars⇩l⇩s⇩s⇩t A" using fv_ik_subset_fv_sst'[of "unlabel A"] vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩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⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I" shows "occurs t ∈ ik⇩l⇩s⇩s⇩t A" proof - have ℐ_wt: "wt⇩s⇩u⇩b⇩s⇩t I" using ℐ unfolding welltyped_constraint_model_def constraint_model_def by metis obtain s where s: "s ∈ ik⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t A" using 𝒜_reach ℐ x proof (induction A rule: reachable_constraints.induct) case (step A T ξ σ α) define θ where "θ ≡ ξ ∘⇩s σ ∘⇩s α" define T' where "T' ≡ dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I) (set (db⇩l⇩s⇩s⇩t A I)) (unlabel T') I" using step.prems unlabel_append[of A T'] db⇩s⇩s⇩t_set_is_dbupd⇩s⇩s⇩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⇩s⇩s⇩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 ⋅⇩l⇩s⇩s⇩t θ))" using dual⇩l⇩s⇩s⇩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 eval_term.simps(1)) hence "occurs (Var y) ⋅ θ ∈ set ts' ⋅⇩s⇩e⇩t θ" "receive⟨ts' ⋅⇩l⇩i⇩s⇩t θ⟩ ∈ set (unlabel (transaction_receive T ⋅⇩l⇩s⇩s⇩t θ))" using subst_lsst_unlabel_member[of "receive⟨ts'⟩" "transaction_receive T" θ] by fastforce+ hence "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I ⊢ occurs (Var y) ⋅ θ ⋅ I" using wellformed_transaction_sem_receives[ OF T_wf, of "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I" "set (db⇩l⇩s⇩s⇩t A I)" θ I "ts' ⋅⇩l⇩i⇩s⇩t θ"] ℐ_is_T_model unfolding T'_def list_all_iff by fastforce hence *: "ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I ⊢ occurs (θ y ⋅ I)" by auto have "occurs (θ y ⋅ I) ∈ ik⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I ⊢ occurs k" shows "occurs k ∈ ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I" and "occurs k ∈ ik⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩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 ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α" show ?case using step.IH step.hyps(2) bvars⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq[of ?T'] bvars⇩l⇩s⇩s⇩t_subst[of "transaction_strand T" "ξ ∘⇩s σ ∘⇩s α"] bvars⇩s⇩s⇩t_append[of "unlabel 𝒜" "unlabel (dual⇩l⇩s⇩s⇩t ?T')"] unlabel_append[of 𝒜 "dual⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩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 ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α" define X where "X ≡ ⋃T ∈ set P. bvars_transaction T" have "fv⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (𝒜@dual⇩l⇩s⇩s⇩t T') ∩ X = {}" using step.IH[unfolded X_def[symmetric]] fv⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t A ∩ bvars⇩l⇩s⇩s⇩t A = {}" using A proof (induction A rule: reachable_constraints.induct) case (step 𝒜 T ξ σ α) define T' where "T' ≡ dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t 𝒜 ∩ fv⇩l⇩s⇩s⇩t T' = {}" using 0(7) 1(4)[OF step.hyps(2)] fv⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq unfolding T'_def by (metis (no_types) disjoint_iff_not_equal subset_iff) have "bvars⇩l⇩s⇩s⇩t T' ⊆ ⋃(bvars_transaction ` set P)" "fv⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t 𝒜 ∩ bvars⇩l⇩s⇩s⇩t T' = {}" by blast have "fv⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α) ∩ bvars_transaction T = {}" using 0(4)[OF step.hyps(2)] 1(4)[OF step.hyps(2)] by blast hence 4: "fv⇩l⇩s⇩s⇩t T' ∩ bvars⇩l⇩s⇩s⇩t T' = {}" by (metis (no_types) T'_def fv⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq bvars⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq unlabel_subst bvars⇩s⇩s⇩t_subst) have "fv⇩l⇩s⇩s⇩t (𝒜@T') ∩ bvars⇩l⇩s⇩s⇩t (𝒜@T') = {}" using 2 3 4 step.IH unfolding unlabel_append[of 𝒜 T'] fv⇩s⇩s⇩t_append[of "unlabel 𝒜" "unlabel T'"] bvars⇩s⇩s⇩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⇩t⇩r⇩m⇩s' arity (trms_transaction T)" and A: "A ∈ reachable_constraints P" shows "wf⇩s⇩s⇩t (unlabel A)" and "wf⇩t⇩r⇩m⇩s (trms⇩l⇩s⇩s⇩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'⇩s⇩s⇩t (?X T) (unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand T)))" "fv⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (transaction_strand T)) ∩ bvars⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (transaction_strand T)) = {}" "wf⇩t⇩r⇩m⇩s (trms_transaction T)" when T: "T ∈ set P" for T unfolding admissible_transaction_terms_def by (metis T wellformed_transaction_wf⇩s⇩s⇩t(1), metis T wellformed_transaction_wf⇩s⇩s⇩t(2) fv⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq bvars⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq, metis T wf⇩t⇩r⇩m⇩s_code P(2)) from A have "wf⇩s⇩s⇩t (unlabel A) ∧ wf⇩t⇩r⇩m⇩s (trms⇩l⇩s⇩s⇩t A)" proof (induction A rule: reachable_constraints.induct) case (step A T ξ σ α) let ?T' = "dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" have IH: "wf'⇩s⇩s⇩t {} (unlabel A)" "fv⇩l⇩s⇩s⇩t A ∩ bvars⇩l⇩s⇩s⇩t A = {}" "wf⇩t⇩r⇩m⇩s (trms⇩l⇩s⇩s⇩t A)" using step.IH by metis+ have 1: "wf'⇩s⇩s⇩t {} (unlabel (A@?T'))" using transaction_decl_fresh_renaming_substs_wf_sst[OF 0(1)[OF step.hyps(2)] step.hyps(3-5)] wf⇩s⇩s⇩t_vars_mono[of "{}"] wf⇩s⇩s⇩t_append[OF IH(1)] by simp have 2: "fv⇩l⇩s⇩s⇩t (A@?T') ∩ bvars⇩l⇩s⇩s⇩t (A@?T') = {}" using reachable_constraints_fv_bvars_disj'[OF P(1)] reachable_constraints.step[OF step.hyps] by blast have "wf⇩t⇩r⇩m⇩s (trms⇩l⇩s⇩s⇩t ?T')" using trms⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq unlabel_subst wf_trms_subst[ OF transaction_decl_fresh_renaming_substs_range_wf_trms[OF step.hyps(3-5)], THEN wf⇩t⇩r⇩m⇩s_trms⇩s⇩s⇩t_subst, OF 0(3)[OF step.hyps(2)]] by metis hence 3: "wf⇩t⇩r⇩m⇩s (trms⇩l⇩s⇩s⇩t (A@?T'))" using IH(3) by auto show ?case using 1 2 3 by force qed simp thus "wf⇩s⇩s⇩t (unlabel A)" "wf⇩t⇩r⇩m⇩s (trms⇩l⇩s⇩s⇩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). Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩v x = TAtom (Atom a))" and t: "t ∈ subterms⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩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). Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩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⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t A)") case False hence "t ∈ subterms⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)))" using step.prems by simp hence "t ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T) ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩s α)" using dual_transaction_ik_is_transaction_send'[OF T_wf[OF step.hyps(2)]] by metis hence "t ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T)) ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩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⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (transaction_send T))" "t = s ⋅ ξ ∘⇩s σ ∘⇩s α" by force 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)) ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩s α" using Ana_subst'[of f S _ "snd (Ana s)" "ξ ∘⇩s σ ∘⇩s α"] s(2) s' n by (cases "Ana s") auto hence "attack⟨n⟩ ∈ set (snd (Ana s)) ∨ attack⟨n⟩ ∈ subst_range (ξ ∘⇩s σ ∘⇩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). Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩v x = TAtom (Atom a))" "∀T ∈ set P. ∀x ∈ vars_transaction T. ¬TAtom AttackType ⊑ Γ⇩v x" and ℐ: "welltyped_constraint_model ℐ 𝒜" and l: "ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ ⊢ attack⟨l⟩" shows "attack⟨l⟩ ∈ ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩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⇩s⇩u⇩b⇩s⇩t ℐ" "wf⇩t⇩r⇩m⇩s (subst_range ℐ)" "wt⇩s⇩u⇩b⇩s⇩t ℐ" using ℐ unfolding welltyped_constraint_model_def constraint_model_def by metis+ have 0: "wf⇩t⇩r⇩m⇩s (ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ)" when 𝒜: "𝒜 ∈ reachable_constraints P" for 𝒜 using reachable_constraints_wf⇩t⇩r⇩m⇩s[OF _ 𝒜] admissible_transaction_terms_wf⇩t⇩r⇩m⇩s P(2) ik⇩s⇩s⇩t_trms⇩s⇩s⇩t_subset[of "unlabel 𝒜"] wf_trms_subst[OF ℐ'(3)] by fast have 1: "∀x ∈ fv⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t 𝒜). ¬TAtom AttackType ⊑ Γ⇩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)) ⋅⇩s⇩e⇩t ℐ" when t: "t ∈ subterms⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t 𝒜)" for t proof assume "attack⟨l⟩ ∈ set (snd (Ana t)) ⋅⇩s⇩e⇩t ℐ" then obtain s where s: "s ∈ set (snd (Ana t))" "s ⋅ ℐ = attack⟨l⟩" by force 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⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t 𝒜)" using t fv_subterms by fastforce hence "Γ⇩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⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ)" for t proof assume "attack⟨l⟩ ∈ set (snd (Ana t))" then obtain s where s: "s ∈ subterms⇩s⇩e⇩t (ℐ ` fv⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t 𝒜))" "attack⟨l⟩ ∈ set (snd (Ana s))" using Ana_subst_subterms_cases[OF t] 2 by fast then obtain x where x: "x ∈ fv⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t 𝒜)" "s ⊑ ℐ x" by force hence "ℐ x ∈ subterms⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ)" using var_is_subterm[of x] subterms_subst_subset'[of ℐ "ik⇩l⇩s⇩s⇩t 𝒜"] by force hence *: "wf⇩t⇩r⇩m (ℐ x)" "wf⇩t⇩r⇩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 ⋅ ξ ∘⇩s σ ∘⇩s α = attack⟨n⟩" and hyps: "transaction_decl_subst ξ T" "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t 𝒜)" "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t 𝒜)" and T: "∀x ∈ set (transaction_fresh T). Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩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 (ξ ∘⇩s σ ∘⇩s α)" 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' ⋅⇩l⇩i⇩s⇩t θ ∧ (l,send⟨ts'⟩) ∈ set (transaction_strand T)" when ts: "(l,receive⟨ts⟩) ∈ set (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ))" for l ts θ and T::"('fun, 'atom, 'sets, 'lbl) prot_transaction" using subst_lsst_memD(2)[OF ts[unfolded dual⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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 ⋅⇩s⇩e⇩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⇩l⇩s⇩s⇩t (𝒜 @ dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))" using step.prems(1) in_ik⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))" using step.prems(1) unfolding unlabel_def by force then obtain ts' where ts': "ts = ts' ⋅⇩l⇩i⇩s⇩t ξ ∘⇩s σ ∘⇩s α" "(l', send⟨ts'⟩) ∈ set (transaction_strand T)" using 5 by meson then obtain t' where t': "t' ∈ set ts'" "t' ⋅ ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ" using private_const_deduct[OF _ l] 3 by simp then obtain ts where ts: "receive⟨ts⟩ ∈ set (unlabel 𝒜)" "attack⟨l⟩ ∈ set ts ⋅⇩s⇩e⇩t ℐ" using in_ik⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ ⊢ attack⟨n⟩" shows "attack⟨n⟩ ∈ ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩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). Γ⇩v x = TAtom Value" "∀T ∈ set P. ∀x ∈ vars_transaction T. ¬TAtom AttackType ⊑ Γ⇩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⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩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: "Γ⇩v x = TAtom Value" "x ∈ fv⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t 𝒜. Γ⇩v x = TAtom Value ⟶ (∃B. prefix B 𝒜 ∧ x ∉ fv⇩l⇩s⇩s⇩t B ∧ ℐ x ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ) (set (db⇩l⇩s⇩s⇩t 𝒜 ℐ)) (unlabel T') ℐ" using step.prems unlabel_append[of 𝒜 T'] db⇩s⇩s⇩t_set_is_dbupd⇩s⇩s⇩t[of "unlabel 𝒜" ℐ "[]"] strand_sem_append_stateful[of "{}" "{}" "unlabel 𝒜" "unlabel T'" ℐ] by (simp add: T'_def welltyped_constraint_model_def constraint_model_def db⇩s⇩s⇩t_def) have ℐ_interp: "interpretation⇩s⇩u⇩b⇩s⇩t ℐ" and ℐ_wt: "wt⇩s⇩u⇩b⇩s⇩t ℐ" and ℐ_wf⇩t⇩r⇩m⇩s: "wf⇩t⇩r⇩m⇩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⇩l⇩s⇩s⇩t T'" "Γ⇩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⇩s⇩s⇩t_append[of "unlabel 𝒜" "unlabel T'"] unlabel_append[of 𝒜 T'] unfolding T'_def by force have "x ∈ fv⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" using x(1) fv⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq unfolding T'_def by fastforce then obtain y where y: "y ∈ fv⇩l⇩s⇩s⇩t (transaction_strand T)" "x ∈ fv ((ξ ∘⇩s σ ∘⇩s α) y)" using fv⇩s⇩s⇩t_subst_obtain_var[of x "unlabel (transaction_strand T)" "ξ ∘⇩s σ ∘⇩s α"] unlabel_subst[of "transaction_strand T" "ξ ∘⇩s σ ∘⇩s α"] by auto have y_x: "(ξ ∘⇩s σ ∘⇩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 "Γ ((ξ ∘⇩s σ ∘⇩s α) y) = TAtom Value" using x(2) y_x by simp moreover have "wt⇩s⇩u⇩b⇩s⇩t (ξ ∘⇩s σ ∘⇩s α)" by (rule transaction_decl_fresh_renaming_substs_wt[OF step.hyps(3-5)]) ultimately have y_val: "Γ⇩v y = TAtom Value" by (metis wt⇩s⇩u⇩b⇩s⇩t_def Γ.simps(1)) have "Fun n [] = (ξ ∘⇩s σ ∘⇩s α) y ⋅ ℐ" using n y_x by simp hence y_n: "Fun n [] = (ξ ∘⇩s σ ∘⇩s α ∘⇩s ℐ) y" by (metis subst_subst_compose[of "Var y" "ξ ∘⇩s σ ∘⇩s α" ℐ] eval_term.simps(1)) have 𝒜_ik_ℐ_vals: "∀x ∈ fv⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t 𝒜). ∃f. ℐ x = Fun f []" proof - have "∃a. Γ (ℐ x) = Var a" when "x ∈ fv⇩l⇩s⇩s⇩t 𝒜" for x using that reachable_constraints_vars_TAtom_typed[OF step.hyps(1) P, of x] vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t[of "unlabel 𝒜"] wt_subst_trm''[OF ℐ_wt, of "Var x"] by force hence "∃f. ℐ x = Fun f []" when "x ∈ fv⇩l⇩s⇩s⇩t 𝒜" for x using that wf_trm_subst[OF ℐ_wf⇩t⇩r⇩m⇩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 eval_term.simps(1)[of _ x ℐ]) thus ?thesis using fv_ik_subset_fv_sst'[of "unlabel 𝒜"] vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t[of "unlabel 𝒜"] by blast qed hence 𝒜_subterms_subst_cong: "subterms⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ) = subterms⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t 𝒜) ⋅⇩s⇩e⇩t ℐ" by (metis ik⇩s⇩s⇩t_subst[of "unlabel 𝒜" ℐ] unlabel_subst[of 𝒜 ℐ] subterms_subst_lsst_ik[of 𝒜 ℐ]) have x_nin_𝒜: "x ∉ fv⇩l⇩s⇩s⇩t 𝒜" proof - have "x ∈ fv⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" using x(1) fv⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq unfolding T'_def by fast hence "x ∈ fv⇩s⇩s⇩t ((unlabel (transaction_strand T) ⋅⇩s⇩s⇩t σ) ⋅⇩s⇩s⇩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 ⋅⇩l⇩s⇩s⇩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⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩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 [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t 𝒜 ∨ (∃z ∈ fv⇩l⇩s⇩s⇩t 𝒜. Γ⇩v z = TAtom Value ∧ ℐ z = Fun n [])" proof assume y_in: "y ∈ fv⇩l⇩s⇩s⇩t (transaction_receive T)" then obtain ts where ts: "receive⟨ts⟩ ∈ set (unlabel (transaction_receive T))" "y ∈ fv⇩s⇩e⇩t (set ts)" using admissible_transaction_strand_step_cases(1)[OF T_adm] by force hence ts_deduct: "list_all (λt. ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ ⊢ t ⋅ ξ ∘⇩s σ ∘⇩s α ⋅ ℐ) ts" using wellformed_transaction_sem_receives[ OF T_wf, of "ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ" "set (db⇩l⇩s⇩s⇩t 𝒜 ℐ)" "ξ ∘⇩s σ ∘⇩s α" ℐ "ts ⋅⇩l⇩i⇩s⇩t ξ ∘⇩s σ ∘⇩s α"] ℐ_is_T_model subst_lsst_unlabel_member[of "receive⟨ts⟩" "transaction_receive T" "ξ ∘⇩s σ ∘⇩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 [] ⊑⇩s⇩e⇩t ik⇩l⇩s⇩s⇩t 𝒜 ∨ (∃z ∈ fv⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t 𝒜). Γ⇩v z = TAtom Value ∧ ℐ z = Fun n [])" proof - have "Fun n [] ⊑ ty ⋅ ξ ∘⇩s σ ∘⇩s α ⋅ ℐ" using imageI[of "Var y" "subterms ty" "λx. x ⋅ ξ ∘⇩s σ ∘⇩s α ∘⇩s ℐ"] var_is_subterm[OF ty(2)] subterms_subst_subset[of "ξ ∘⇩s σ ∘⇩s α ∘⇩s ℐ" ty] subst_subst_compose[of ty "ξ ∘⇩s σ ∘⇩s α" ℐ] y_n by (auto simp del: subst_subst_compose) hence "Fun n [] ⊑⇩s⇩e⇩t ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩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 [] ⊑⇩s⇩e⇩t ik⇩l⇩s⇩s⇩t 𝒜 ∨ (∃z ∈ fv⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t 𝒜). ℐ z = Fun n [])" using const_subterm_subst_cases[of n _ ℐ] 𝒜_ik_ℐ_vals by fastforce thus ?thesis using ℐ_wt n(2) unfolding wt⇩s⇩u⇩b⇩s⇩t_def is_Val_def is_Abs_def by force qed thus ?thesis using fv_ik_subset_fv_sst' ik⇩s⇩s⇩t_trms⇩s⇩s⇩t_subset[of "unlabel 𝒜"] 𝒜_subterms_subst_cong by fast next assume y_in: "y ∈ fv⇩l⇩s⇩s⇩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⟨(ξ ∘⇩s σ ∘⇩s α) y, Fun (Set s) []⟩ ∈ set (unlabel (transaction_checks T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))" using subst_lsst_unlabel_member by fastforce hence n_in_db: "(Fun n [], Fun (Set s) []) ∈ set (db'⇩s⇩s⇩t (unlabel 𝒜) ℐ [])" using wellformed_transaction_sem_pos_checks(1)[ OF T_wf, of "ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ" "set (db⇩l⇩s⇩s⇩t 𝒜 ℐ)" "ξ ∘⇩s σ ∘⇩s α" ℐ assign "(ξ ∘⇩s σ ∘⇩s α) y" "Fun (Set s) []"] ℐ_is_T_model n y_x unfolding T'_def db⇩s⇩s⇩t_def by fastforce obtain tn sn where tsn: "insert⟨tn,sn⟩ ∈ set (unlabel 𝒜)" "Fun n [] = tn ⋅ ℐ" using db⇩s⇩s⇩t_in_cases[OF n_in_db] by force have "Fun n [] = tn ∨ (∃z. Γ⇩v z = TAtom Value ∧ tn = Var z)" using ℐ_wt tsn(2) n(2) unfolding wt⇩s⇩u⇩b⇩s⇩t_def is_Val_def is_Abs_def by (cases tn) auto moreover have "tn ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜)" "fv tn ⊆ fv⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t 𝒜. Γ⇩v z = TAtom Value ∧ ℐ z = Fun n []" then obtain B where B: "prefix B 𝒜" "Fun n [] ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t B)" by (metis IH n(1)) thus ?thesis using n x_nin_𝒜 trms⇩s⇩s⇩t_unlabel_prefix_subset(1)[of B] by (metis (no_types, opaque_lifting) self_append_conv subset_iff subterms⇩s⇩e⇩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⇩l⇩s⇩s⇩t (𝒜@T')" "Γ⇩v x = TAtom Value" show "?Q x (𝒜@T')" proof (cases "x ∈ fv⇩l⇩s⇩s⇩t 𝒜") case False hence "x ∈ fv⇩l⇩s⇩s⇩t T'" using x(1) unlabel_append[of 𝒜] fv⇩s⇩s⇩t_append[of "unlabel 𝒜"] by simp then obtain B where B: "prefix B 𝒜" "x ∉ fv⇩l⇩s⇩s⇩t B" "ℐ x ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t B)" using x(2) 1 by blast 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) [] ⊑⇩s⇩e⇩t ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ" shows "Fun (Val n) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t 𝒜" proof - have *: "wf⇩s⇩s⇩t (unlabel 𝒜)" "constr_sem_stateful ℐ (unlabel 𝒜)" "interpretation⇩s⇩u⇩b⇩s⇩t ℐ" "wf⇩t⇩r⇩m⇩s (subst_range ℐ)" "wt⇩s⇩u⇩b⇩s⇩t ℐ" using reachable_constraints_wf(1)[OF P(1) _ 𝒜_reach] admissible_transaction_terms_wf⇩t⇩r⇩m⇩s ℐ P(2) n unfolding welltyped_constraint_model_def constraint_model_def wf⇩t⇩r⇩m⇩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) [] ⊑⇩s⇩e⇩t ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t ℐ" shows "Fun (Val n) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩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) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t A" obtains B T ξ σ α where "prefix (B@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)) A" and "B ∈ reachable_constraints P" "T ∈ set P" "transaction_decl_subst ξ T" "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t B)" "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t (g B)) ∧ transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩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) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" using n trms⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq unlabel_subst unfolding Ts(1) g_def f_def unlabel_def trms⇩s⇩s⇩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⇩l⇩s⇩s⇩t (g B))" "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩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⇩s⇩s⇩t (unlabel (transaction_strand T) ⋅⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α) = trms_transaction T ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩s α" using trms⇩s⇩s⇩t_subst[of _ "ξ ∘⇩s σ ∘⇩s α"] T_no_bvars by blast hence "Fun (Val n) [] ⊑⇩s⇩e⇩t trms_transaction T ⋅⇩s⇩e⇩t ξ ∘⇩s σ ∘⇩s α" by (metis T(2) unlabel_subst) hence "Fun (Val n) [] ⊑⇩s⇩e⇩t subst_range (ξ ∘⇩s σ ∘⇩s α)" using admissible_transaction_terms_no_Value_consts(1)[OF T_adm_terms] const_subterms_subst_cases'[of "Val n" "ξ ∘⇩s σ ∘⇩s α" "trms_transaction T"] by blast then obtain tn where tn: "tn ∈ subst_range (ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t A" obtains B T ξ σ α where "prefix (B@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)) A" and "B ∈ reachable_constraints P" "T ∈ set P" "transaction_decl_subst ξ T" "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t B)" "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩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) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t A" shows "occurs (Fun (Val n) []) ∈ ik⇩l⇩s⇩s⇩t A" proof - obtain B T ξ σ α where B: "prefix (B@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)) A" "B ∈ reachable_constraints P" "T ∈ set P" "transaction_decl_subst ξ T" "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t B)" "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t B)" "Fun (Val n) [] ∈ subst_range σ" using constraint_model_Value_in_constr_prefix_fresh_action[OF 𝒜_reach P val] by blast define θ where "θ ≡ ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t (transaction_send T)" using ts by force hence "occurs (Var x) ⋅ θ ∈ ik⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ))" using dual_transaction_ik_is_transaction_send'[OF T_wf(1), of θ] by fast hence "occurs (Fun (Val n) []) ∈ ik⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ))" using x(2) by simp thus ?thesis using B(1)[unfolded θ_def[symmetric]] unlabel_append[of B "dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ)"] ik⇩s⇩s⇩t_append[of "unlabel B" "unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t ℐ ⊢ t" shows "occurs (Fun (Val n) []) ∈ ik⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t 𝒜))" shows "¬is_PubConstValue f" and "¬is_Abs f" proof - obtain x where x: "x ∈ fv⇩l⇩s⇩s⇩t 𝒜" "f ∈ funs_term (ℐ x)" using f by force obtain T where T: "Fun f T ⊑ ℐ x" using funs_term_Fun_subterm[OF x(2)] by force have ℐ_interp: "interpretation⇩s⇩u⇩b⇩s⇩t ℐ" and ℐ_wt: "wt⇩s⇩u⇩b⇩s⇩t ℐ" and ℐ_wf⇩t⇩r⇩m⇩s: "wf⇩t⇩r⇩m⇩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⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩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⇩t⇩r⇩m⇩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 eval_term.simps(1)[of _ x ℐ]) hence 2: "ℐ x = Fun f []" using x(2) by force have 3: "Γ⇩v x ≠ TAtom AbsValue" using 0 by force have "¬is_PubConstValue f ∧ ¬is_Abs f" proof (cases "Γ⇩v x = TAtom Value") case True then obtain B where B: "prefix B 𝒜" "x ∉ fv⇩l⇩s⇩s⇩t B" "ℐ x ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t B)" using constraint_model_Value_var_in_constr_prefix[OF 𝒜_reach ℐ P P_occ] x(1) by fast have "ℐ x ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜)" using B(1,3) trms⇩s⇩s⇩t_append[of "unlabel B"] unlabel_append[of B] unfolding prefix_def by auto hence "f ∈ ⋃(funs_term ` trms⇩l⇩s⇩s⇩t 𝒜)" using x(2) funs_term_subterms_eq(2)[of "trms⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α) ∈ reachable_constraints P" and ℐ: "welltyped_constraint_model ℐ (𝒜@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))" and ξ: "transaction_decl_subst ξ T" and σ: "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t 𝒜)" and α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩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) [] = (ξ ∘⇩s σ ∘⇩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 ((ξ ∘⇩s σ ∘⇩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⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t[of "unlabel (transaction_strand T)"] by (blast, blast) have σx_type: "Γ (σ x) = TAtom Value" using σ x Γ⇩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: "(σ ∘⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" using x_not_bvar fv⇩s⇩s⇩t_subst_fv_subset[OF x(1), of "ξ ∘⇩s σ ∘⇩s α"] unlabel_subst[of "transaction_strand T" "ξ ∘⇩s σ ∘⇩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⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq[of "transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α"] fv⇩s⇩s⇩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⇩l⇩s⇩s⇩t 𝒜) ⊆ SMP (⋃T ∈ set P. trms_transaction T)" (is "?A 𝒜") and "SMP (pair`setops⇩s⇩s⇩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 ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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⇩s⇩u⇩b⇩s⇩t δ ∧ wf⇩t⇩r⇩m⇩s (subst_range δ) ∧ t = s ⋅ δ" let ?Q = "λt. ∃s δ. s ∈ N ∧ wt⇩s⇩u⇩b⇩s⇩t δ ∧ wf⇩t⇩r⇩m⇩s (subst_range δ) ∧ t = s ⋅ δ" have IH: "SMP (trms⇩l⇩s⇩s⇩t A) ⊆ SMP M" "SMP (pair ` setops⇩s⇩s⇩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⇩l⇩s⇩s⇩t (A@dual⇩l⇩s⇩s⇩t T')) = SMP (trms⇩l⇩s⇩s⇩t A) ∪ SMP (trms⇩l⇩s⇩s⇩t T')" "SMP (pair ` setops⇩s⇩s⇩t (unlabel (A@dual⇩l⇩s⇩s⇩t T'))) = SMP (pair ` setops⇩s⇩s⇩t (unlabel A)) ∪ SMP (pair ` setops⇩s⇩s⇩t (unlabel T'))" using trms⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq[of T'] setops⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq[of T'] trms⇩s⇩s⇩t_append[of "unlabel A" "unlabel (dual⇩l⇩s⇩s⇩t T')"] setops⇩s⇩s⇩t_append[of "unlabel A" "unlabel (dual⇩l⇩s⇩s⇩t T')"] unlabel_append[of A "dual⇩l⇩s⇩s⇩t T'"] image_Un[of pair "setops⇩s⇩s⇩t (unlabel A)" "setops⇩s⇩s⇩t (unlabel T')"] SMP_union[of "trms⇩l⇩s⇩s⇩t A" "trms⇩l⇩s⇩s⇩t T'"] SMP_union[of "pair ` setops⇩s⇩s⇩t (unlabel A)" "pair ` setops⇩s⇩s⇩t (unlabel T')"] by argo+ have 1: "SMP (trms⇩l⇩s⇩s⇩t T') ⊆ SMP M" proof (intro SMP_subset_I ballI) fix t show "t ∈ trms⇩l⇩s⇩s⇩t T' ⟹ ?P t" using trms⇩s⇩s⇩t_wt_subst_ex[OF ξσα_wt ξσα_wf, of t "unlabel (transaction_strand T)"] unlabel_subst[of "transaction_strand T" "ξ ∘⇩s σ ∘⇩s α"] step.hyps(2) unfolding T'_def M_def by auto qed have 2: "SMP (pair ` setops⇩s⇩s⇩t (unlabel T')) ⊆ SMP N" proof (intro SMP_subset_I ballI) fix t show "t ∈ pair ` setops⇩s⇩s⇩t (unlabel T') ⟹ ?Q t" using setops⇩s⇩s⇩t_wt_subst_ex[OF ξσα_wt ξσα_wf, of t "unlabel (transaction_strand T)"] unlabel_subst[of "transaction_strand T" "ξ ∘⇩s σ ∘⇩s α"] step.hyps(2) unfolding T'_def N_def by auto qed have "SMP (trms⇩l⇩s⇩s⇩t (A@dual⇩l⇩s⇩s⇩t T')) ⊆ SMP M" "SMP (pair ` setops⇩s⇩s⇩t (unlabel (A@dual⇩l⇩s⇩s⇩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⇩s⇩s⇩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). Γ⇩v x = TAtom Value" "∀T ∈ set P. transaction_decl T () = []" "∀T ∈ set P. admissible_transaction_terms T" "∀T ∈ set P. ∀x ∈ vars_transaction T. Γ⇩v x = TAtom Value ∨ (∃a. Γ⇩v x = ⟨a⟩⇩τ⇩a)" shows "Pair ∉ ⋃(funs_term ` SMP (trms⇩l⇩s⇩s⇩t A))" using A proof (induction A rule: reachable_constraints.induct) case (step A T ξ σ α) define T' where "T' ≡ dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t (A@T')) = SMP (trms⇩l⇩s⇩s⇩t A) ∪ SMP (trms⇩l⇩s⇩s⇩t T')" using SMP_union[of "trms⇩l⇩s⇩s⇩t A" "trms⇩l⇩s⇩s⇩t T'"] unlabel_append[of A T'] trms⇩s⇩s⇩t_append[of "unlabel A" "unlabel T'"] by simp have 1: "wf⇩t⇩r⇩m⇩s (trms⇩l⇩s⇩s⇩t T')" using reachable_constraints_wf⇩t⇩r⇩m⇩s[OF _ reachable_constraints.step[OF step.hyps]] admissible_transaction_terms_wf⇩t⇩r⇩m⇩s P(3) trms⇩s⇩s⇩t_append[of "unlabel A"] unlabel_append[of A] unfolding T'_def by force have 2: "Pair ∉ ⋃(funs_term ` (subst_range (ξ ∘⇩s σ ∘⇩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⇩s⇩s⇩t (unlabel (transaction_strand T) ⋅⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" for t using 2 trms⇩s⇩s⇩t_funs_term_cases[OF t] by force hence 3: "Pair ∉ funs_term t" when t: "t ∈ trms⇩l⇩s⇩s⇩t T'" for t using t unlabel_subst[of "transaction_strand T" "ξ ∘⇩s σ ∘⇩s α"] trms⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq[of "transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α"] unfolding T'_def by metis have "∃a. Γ⇩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. Γ⇩v x = TAtom a" when "x ∈ vars⇩s⇩s⇩t (unlabel (transaction_strand T) ⋅⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" for x using wt_subst_fv⇩s⇩e⇩t_termtype_subterm[OF _ ξσα_wt ξσα_wf, of x "vars_transaction T"] vars⇩s⇩s⇩t_subst_cases[OF that] by fastforce hence "∃a. Γ⇩v x = TAtom a" when "x ∈ vars⇩l⇩s⇩s⇩t T'" for x using that unlabel_subst[of "transaction_strand T" "ξ ∘⇩s σ ∘⇩s α"] vars⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq[of "transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α"] unfolding T'_def by simp hence "∃a. Γ⇩v x = TAtom a" when "x ∈ fv⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t T')" for x using that fv_trms⇩s⇩s⇩t_subset(1) by fast hence "Pair ∉ funs_term (Γ (Var x))" when "x ∈ fv⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩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⇩s⇩s⇩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⇩s⇩s⇩t (unlabel A)") case False hence "t ∈ pair ` setops⇩s⇩s⇩t (unlabel (transaction_strand T) ⋅⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" using step.prems setops⇩s⇩s⇩t_append unlabel_append setops⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq[of "transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α"] unlabel_subst[of "transaction_strand T" "ξ ∘⇩s σ ∘⇩s α"] by fastforce then obtain t' δ where t': "t' ∈ pair ` setops⇩s⇩s⇩t (unlabel (transaction_strand T))" "wt⇩s⇩u⇩b⇩s⇩t δ" "wf⇩t⇩r⇩m⇩s (subst_range δ)" "t = t' ⋅ δ" using setops⇩s⇩s⇩t_wt_subst_ex[OF ξσα_wt ξσα_wf] by blast then obtain s s' where s: "t' = pair (s,s')" using setops⇩s⇩s⇩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⇩s⇩s⇩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⇩t⇩r⇩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⇩s⇩s⇩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⇩s⇩s⇩t (unlabel A)" using t unfolding setops⇩s⇩s⇩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 ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩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 ⋅ ξ ∘⇩s σ ∘⇩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 (ξ ∘⇩s σ ∘⇩s α)") (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⇩s⇩s⇩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 force hence "(Fun (Set s) []::('fun,'atom,'sets,'lbl) prot_term) ∈ trms⇩l⇩s⇩s⇩t A" using t setops⇩s⇩s⇩t_member_iff[of c "Fun (Set s) []" "unlabel A"] by force hence "wf⇩t⇩r⇩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⇩t⇩r⇩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⇩s⇩s⇩t (unlabel A). ∀t ∈ pair ` setops⇩s⇩s⇩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⇩s⇩s⇩t (unlabel A). ∀t ∈ pair ` setops⇩s⇩s⇩t (unlabel A). (∃σ θ ρ. wt⇩s⇩u⇩b⇩s⇩t σ ∧ wt⇩s⇩u⇩b⇩s⇩t θ ∧ wf⇩t⇩r⇩m⇩s (subst_range σ) ∧ wf⇩t⇩r⇩m⇩s (subst_range θ) ∧ Unifier ρ (s ⋅ σ) (t ⋅ θ)) ⟶ (∃δ. Unifier δ s t)" proof (intro ballI impI) fix s t assume st: "s ∈ pair ` setops⇩s⇩s⇩t (unlabel A)" "t ∈ pair ` setops⇩s⇩s⇩t (unlabel A)" and "∃σ θ ρ. wt⇩s⇩u⇩b⇩s⇩t σ ∧ wt⇩s⇩u⇩b⇩s⇩t θ ∧ wf⇩t⇩r⇩m⇩s (subst_range σ) ∧ wf⇩t⇩r⇩m⇩s (subst_range θ) ∧ Unifier ρ (s ⋅ σ) (t ⋅ θ)" then obtain σ θ ρ where σ: "wt⇩s⇩u⇩b⇩s⇩t σ" "wt⇩s⇩u⇩b⇩s⇩t θ" "wf⇩t⇩r⇩m⇩s (subst_range σ)" "wf⇩t⇩r⇩m⇩s (subst_range θ)" "Unifier ρ (s ⋅ σ) (t ⋅ θ)" by force 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 force have "cs ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t A)" "ct ∈ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩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⇩t⇩r⇩m⇩s' arity (trms_transaction T)" using P admissible_transaction_is_wellformed_transaction(1,4) unfolding admissible_transaction_terms_def by fast+ ultimately have *: "wf⇩t⇩r⇩m cs" "wf⇩t⇩r⇩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⇩s⇩e⇩t N" "wf⇩t⇩r⇩m⇩s N" and P: "∀T ∈ set P. admissible_transaction T" "∀T ∈ set P. list_all tfr⇩s⇩s⇩t⇩p (unlabel (transaction_strand T))" and 𝒜: "𝒜 ∈ reachable_constraints P" shows "tfr⇩s⇩s⇩t (unlabel 𝒜)" using 𝒜 proof (induction 𝒜 rule: reachable_constraints.induct) case (step A T ξ σ α) define T' where "T' ≡ dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t (transaction_strand T) ∩ range_vars (ξ ∘⇩s σ ∘⇩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⇩t⇩r⇩m⇩s M" using admissible_transactions_wf⇩t⇩r⇩m⇩s P'(1) unfolding M(1) by blast have "tfr⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (A@T'))" using reachable_constraints_SMP_subset(1)[OF AT'_reach] tfr_subset(3)[OF M(4), of "trms⇩l⇩s⇩s⇩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⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (A@T') ∪ pair ` setops⇩s⇩s⇩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⇩s⇩s⇩t⇩p (unlabel (transaction_strand T))" using step.hyps(2) P(2) tfr⇩s⇩s⇩t⇩p_is_comp_tfr⇩s⇩s⇩t⇩p unfolding comp_tfr⇩s⇩s⇩t_def tfr⇩s⇩s⇩t_def by fastforce hence "list_all tfr⇩s⇩s⇩t⇩p (unlabel T')" using tfr⇩s⇩s⇩t⇩p_all_wt_subst_apply[OF _ ξσα_wt ξσα_wf ξσα_bvars_disj] dual⇩l⇩s⇩s⇩t_tfr⇩s⇩s⇩t⇩p[of "transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α"] unlabel_subst[of "transaction_strand T" "ξ ∘⇩s σ ∘⇩s α"] unfolding T'_def by argo hence 2: "list_all tfr⇩s⇩s⇩t⇩p (unlabel (A@T'))" using step.IH unlabel_append unfolding tfr⇩s⇩s⇩t_def by auto have "tfr⇩s⇩s⇩t (unlabel (A@T'))" using 1 2 by (metis tfr⇩s⇩s⇩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⇩s⇩e⇩t N" "wf⇩t⇩r⇩m⇩s N" and P: "∀T ∈ set P. wf⇩t⇩r⇩m⇩s' arity (trms_transaction T)" "∀T ∈ set P. list_all tfr⇩s⇩s⇩t⇩p (unlabel (transaction_strand T))" and 𝒜: "𝒜 ∈ reachable_constraints P" shows "tfr⇩s⇩s⇩t (unlabel 𝒜)" using 𝒜 proof (induction 𝒜 rule: reachable_constraints.induct) case (step A T ξ σ α) define T' where "T' ≡ dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t (transaction_strand T) ∩ range_vars (ξ ∘⇩s σ ∘⇩s α) = {}" by (rule transaction_decl_fresh_renaming_substs_vars_disj(4)[OF step.hyps(3,4,5,2)]) have wf_trms_M: "wf⇩t⇩r⇩m⇩s M" using P(1) setops⇩s⇩s⇩t_wf⇩t⇩r⇩m⇩s(2) unfolding M(1) pair_code wf⇩t⇩r⇩m⇩s_code[symmetric] by fast have "SMP (trms⇩l⇩s⇩s⇩t (A@T')) ⊆ SMP M" "SMP (pair ` setops⇩s⇩s⇩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⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (A@T') ∪ pair ` setops⇩s⇩s⇩t (unlabel (A@T')))" using tfr_subset(3)[OF M(4), of "trms⇩l⇩s⇩s⇩t (A@T') ∪ pair ` setops⇩s⇩s⇩t (unlabel (A@T'))"] SMP_union[of "trms⇩l⇩s⇩s⇩t (A@T')" "pair ` setops⇩s⇩s⇩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⇩s⇩s⇩t⇩p (unlabel (transaction_strand T))" using step.hyps(2) P(2) tfr⇩s⇩s⇩t⇩p_is_comp_tfr⇩s⇩s⇩t⇩p unfolding comp_tfr⇩s⇩s⇩t_def tfr⇩s⇩s⇩t_def by fastforce hence "list_all tfr⇩s⇩s⇩t⇩p (unlabel T')" using tfr⇩s⇩s⇩t⇩p_all_wt_subst_apply[OF _ ξσα_wt ξσα_wf ξσα_bvars_disj] dual⇩l⇩s⇩s⇩t_tfr⇩s⇩s⇩t⇩p[of "transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α"] unlabel_subst[of "transaction_strand T" "ξ ∘⇩s σ ∘⇩s α"] unfolding T'_def by argo hence 2: "list_all tfr⇩s⇩s⇩t⇩p (unlabel (A@T'))" using step.IH unlabel_append unfolding tfr⇩s⇩s⇩t_def by auto have "tfr⇩s⇩s⇩t (unlabel (A@T'))" using 1 2 by (metis tfr⇩s⇩s⇩t_def) thus ?case by (metis T'_def) qed simp lemma reachable_constraints_typing_cond⇩s⇩s⇩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⇩s⇩e⇩t N" "wf⇩t⇩r⇩m⇩s N" and P: "∀T ∈ set P. wellformed_transaction T" "∀T ∈ set P. wf⇩t⇩r⇩m⇩s' arity (trms_transaction T)" "∀T ∈ set P. list_all tfr⇩s⇩s⇩t⇩p (unlabel (transaction_strand T))" and 𝒜: "𝒜 ∈ reachable_constraints P" shows "typing_cond⇩s⇩s⇩t (unlabel 𝒜)" using reachable_constraints_wf[OF P(1,2) 𝒜] reachable_constraints_tfr'[OF M P(2,3) 𝒜] unfolding typing_cond⇩s⇩s⇩t_def by blast context begin private lemma reachable_constraints_typing_result_aux: assumes 0: "wf⇩s⇩s⇩t (unlabel 𝒜)" "tfr⇩s⇩s⇩t (unlabel 𝒜)" "wf⇩t⇩r⇩m⇩s (trms⇩l⇩s⇩s⇩t 𝒜)" shows "wf⇩s⇩s⇩t (unlabel (𝒜@[(l,send⟨[attack⟨n⟩]⟩)]))" "tfr⇩s⇩s⇩t (unlabel (𝒜@[(l,send⟨[attack⟨n⟩]⟩)]))" "wf⇩t⇩r⇩m⇩s (trms⇩l⇩s⇩s⇩t (𝒜@[(l,send⟨[attack⟨n⟩]⟩)]))" proof - let ?n = "[(l,send⟨[attack⟨n⟩]⟩)]" let ?A = "𝒜@?n" show "wf⇩s⇩s⇩t (unlabel ?A)" using 0(1) wf⇩s⇩s⇩t_append_suffix'[of "{}" "unlabel 𝒜" "unlabel ?n"] unlabel_append[of 𝒜 ?n] by simp show "wf⇩t⇩r⇩m⇩s (trms⇩l⇩s⇩s⇩t ?A)" using 0(3) trms⇩s⇩s⇩t_append[of "unlabel 𝒜" "unlabel ?n"] unlabel_append[of 𝒜 ?n] by fastforce have "∀t ∈ trms⇩l⇩s⇩s⇩t ?n ∪ pair ` setops⇩s⇩s⇩t (unlabel ?n). ∃c. t = Fun c []" "∀t ∈ trms⇩l⇩s⇩s⇩t ?n ∪ pair ` setops⇩s⇩s⇩t (unlabel ?n). Ana t = ([],[])" by (simp_all add: setops⇩s⇩s⇩t_def) hence "tfr⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t 𝒜 ∪ pair ` setops⇩s⇩s⇩t (unlabel 𝒜) ∪ (trms⇩l⇩s⇩s⇩t ?n ∪ pair ` setops⇩s⇩s⇩t (unlabel ?n)))" using 0(2) tfr_consts_mono unfolding tfr⇩s⇩s⇩t_def by blast hence "tfr⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (𝒜@?n) ∪ pair ` setops⇩s⇩s⇩t (unlabel (𝒜@?n)))" using unlabel_append[of 𝒜 ?n] trms⇩s⇩s⇩t_append[of "unlabel 𝒜" "unlabel ?n"] setops⇩s⇩s⇩t_append[of "unlabel 𝒜" "unlabel ?n"] by (simp add: setops⇩s⇩s⇩t_def) thus "tfr⇩s⇩s⇩t (unlabel ?A)" using 0(2) unlabel_append[of ?A ?n] unfolding tfr⇩s⇩s⇩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⇩s⇩e⇩t N" "wf⇩t⇩r⇩m⇩s N" and P: "∀T ∈ set P. admissible_transaction T" "∀T ∈ set P. list_all tfr⇩s⇩s⇩t⇩p (unlabel (transaction_strand T))" and A: "𝒜 ∈ reachable_constraints P" and ℐ: "constraint_model ℐ (𝒜@[(l, send⟨[attack⟨n⟩]⟩)])" shows "∃ℐ⇩τ. welltyped_constraint_model ℐ⇩τ (𝒜@[(l, send⟨[attack⟨n⟩]⟩)])" proof - have I: "interpretation⇩s⇩u⇩b⇩s⇩t ℐ" "wf⇩t⇩r⇩m⇩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⇩t⇩r⇩m⇩s' arity (trms_transaction T)" using P(1) 0(2) unfolding admissible_transaction_terms_def by blast have 3: "wf⇩s⇩s⇩t (unlabel 𝒜)" "tfr⇩s⇩s⇩t (unlabel 𝒜)" "wf⇩t⇩r⇩m⇩s (trms⇩l⇩s⇩s⇩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⇩s⇩e⇩t N" "wf⇩t⇩r⇩m⇩s N" and P: "∀T ∈ set P. wellformed_transaction T" "∀T ∈ set P. wf⇩t⇩r⇩m⇩s' arity (trms_transaction T)" "∀T ∈ set P. list_all tfr⇩s⇩s⇩t⇩p (unlabel (transaction_strand T))" and A: "𝒜 ∈ reachable_constraints P" and ℐ: "constraint_model ℐ (𝒜@[(l, send⟨[attack⟨n⟩]⟩)])" shows "∃ℐ⇩τ. welltyped_constraint_model ℐ⇩τ (𝒜@[(l, send⟨[attack⟨n⟩]⟩)])" proof - have I: "interpretation⇩s⇩u⇩b⇩s⇩t ℐ" "wf⇩t⇩r⇩m⇩s (subst_range ℐ)" "constr_sem_stateful ℐ (unlabel (𝒜@[(l, send⟨[attack⟨n⟩]⟩)]))" using ℐ unfolding constraint_model_def by metis+ have 0: "wf⇩s⇩s⇩t (unlabel 𝒜)" "tfr⇩s⇩s⇩t (unlabel 𝒜)" "wf⇩t⇩r⇩m⇩s (trms⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t proj_subst transaction_strand_proj) qed (simp add: reachable_constraints.init) context begin private lemma reachable_constraints_par_comp⇩l⇩s⇩s⇩t_aux: fixes P defines "Ts ≡ concat (map transaction_strand P)" assumes A: "A ∈ reachable_constraints P" shows "∀b ∈ set (dual⇩l⇩s⇩s⇩t A). ∃a ∈ set Ts. ∃δ. b = a ⋅⇩l⇩s⇩s⇩t⇩p δ ∧ wt⇩s⇩u⇩b⇩s⇩t δ ∧ wf⇩t⇩r⇩m⇩s (subst_range δ) ∧ (∀t ∈ subst_range δ. (∃x. t = Var x) ∨ (∃c. t = Fun c []))" (is "∀b ∈ set (dual⇩l⇩s⇩s⇩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 "θ ≡ ξ ∘⇩s σ ∘⇩s α" let ?R = "λA Ts. ∀b ∈ set A. ∃a ∈ set Ts. Q b a" have "wt⇩s⇩u⇩b⇩s⇩t θ" "wf⇩t⇩r⇩m⇩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⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (transaction_strand T)) ⋅⇩l⇩s⇩s⇩t θ) (transaction_strand T)" using dual⇩l⇩s⇩s⇩t_self_inverse[of "transaction_strand T"] by (auto simp add: Q_def subst_apply_labeled_stateful_strand_def) hence "?R (dual⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ))) (transaction_strand T)" by (metis dual⇩l⇩s⇩s⇩t_subst) hence "?R (dual⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ))) Ts" using step.hyps(2) unfolding Ts_def dual⇩l⇩s⇩s⇩t_def by fastforce thus ?case using step.IH unfolding Q_def θ_def by auto qed simp lemma reachable_constraints_par_comp⇩l⇩s⇩s⇩t: fixes P defines "f ≡ λM. {t ⋅ δ | t δ. t ∈ M ∧ wt⇩s⇩u⇩b⇩s⇩t δ ∧ wf⇩t⇩r⇩m⇩s (subst_range δ) ∧ fv (t ⋅ δ) = {}}" and "Ts ≡ concat (map transaction_strand P)" assumes P_pc: "comp_par_comp⇩l⇩s⇩s⇩t public arity Ana Γ Pair Ts M S" and A: "A ∈ reachable_constraints P" shows "par_comp⇩l⇩s⇩s⇩t A ((f S) - {m. intruder_synth {} m})" using par_comp⇩l⇩s⇩s⇩t_if_comp_par_comp⇩l⇩s⇩s⇩t'[OF P_pc, of "dual⇩l⇩s⇩s⇩t A", THEN par_comp⇩l⇩s⇩s⇩t_dual⇩l⇩s⇩s⇩t] reachable_constraints_par_comp⇩l⇩s⇩s⇩t_aux[OF A, unfolded Ts_def[symmetric]] unfolding f_def dual⇩l⇩s⇩s⇩t_self_inverse by fast end lemma reachable_constraints_par_comp_constr: fixes P f S defines "f ≡ λM. {t ⋅ δ | t δ. t ∈ M ∧ wt⇩s⇩u⇩b⇩s⇩t δ ∧ wf⇩t⇩r⇩m⇩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⇩s⇩e⇩t N" "wf⇩t⇩r⇩m⇩s N" and P: "∀T ∈ set P. wellformed_transaction T" "∀T ∈ set P. wf⇩t⇩r⇩m⇩s' arity (trms_transaction T)" "∀T ∈ set P. list_all tfr⇩s⇩s⇩t⇩p (unlabel (transaction_strand T))" "comp_par_comp⇩l⇩s⇩s⇩t public arity Ana Γ Pair Ts M_fun S" and 𝒜: "𝒜 ∈ reachable_constraints P" and ℐ: "constraint_model ℐ 𝒜" shows "∃ℐ⇩τ. welltyped_constraint_model ℐ⇩τ 𝒜 ∧ ((∀n. welltyped_constraint_model ℐ⇩τ (proj n 𝒜)) ∨ (∃𝒜' l t. prefix 𝒜' 𝒜 ∧ suffix [(l, receive⟨t⟩)] 𝒜' ∧ strand_leaks⇩l⇩s⇩s⇩t 𝒜' Sec ℐ⇩τ))" proof - have ℐ': "constr_sem_stateful ℐ (unlabel 𝒜)" "interpretation⇩s⇩u⇩b⇩s⇩t ℐ" using ℐ unfolding constraint_model_def by blast+ show ?thesis using reachable_constraints_par_comp⇩l⇩s⇩s⇩t[OF P(4)[unfolded Ts_def] 𝒜] reachable_constraints_typing_cond⇩s⇩s⇩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 ≡ λ𝒜. ∃ℐ⇩τ 𝒜'. Q ℐ⇩τ ∧ prefix 𝒜' 𝒜 ∧ (∃l' t. suffix [(l', receive⟨t⟩)] 𝒜') ∧ strand_leaks⇩l⇩s⇩s⇩t 𝒜' Sec ℐ⇩τ" assumes Sec: "∀s ∈ Sec. ¬{} ⊢⇩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 𝒜 ℐ⇩τ 𝒜' s n where 𝒜: "𝒜 ∈ reachable_constraints Ps" and 𝒜': "prefix 𝒜' 𝒜" "constr_sem_stateful ℐ⇩τ (proj_unl n 𝒜'@[send⟨[s]⟩])" and ℐ⇩τ: "Q ℐ⇩τ" and s: "s ∈ Sec - declassified⇩l⇩s⇩s⇩t 𝒜' ℐ⇩τ" unfolding leaks_def strand_leaks⇩l⇩s⇩s⇩t_def by fast have "¬{} ⊢⇩c s" "s ⋅ ℐ⇩τ = s" using s Sec by auto then obtain B k' u where "constr_sem_stateful ℐ⇩τ (proj_unl n B@[send⟨[s]⟩])" "prefix (proj n B) (proj n 𝒜)" "suffix [(k', receive⟨u⟩)] (proj n B)" "s ∈ Sec - declassified⇩l⇩s⇩s⇩t (proj n B) ℐ⇩τ" using constr_sem_stateful_proj_priv_term_prefix_obtain[OF 𝒜' s] unfolding welltyped_constraint_model_def constraint_model_def by metis thus ?thesis using ℐ⇩τ reachable_constraints_transaction_proj[OF 𝒜, of n] proj_idem[of n B] unfolding leaks_def strand_leaks⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)). ?P T a" proof fix a assume a: "a ∈ set (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))" then obtain b where b: "b ∈ set (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" "a = dual⇩l⇩s⇩s⇩t⇩p b" unfolding dual⇩l⇩s⇩s⇩t_def by auto then obtain c where c: "c ∈ set (transaction_strand T)" "b = c ⋅⇩l⇩s⇩s⇩t⇩p ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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: "¬{} ⊢⇩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⇩l⇩s⇩s⇩t B ℐ = declassified⇩l⇩s⇩s⇩t C ℐ ∧ constr_sem_stateful ℐ (unlabel C@[send⟨[s]⟩])" using A proof (induction A rule: reachable_constraints.induct) case (step A T ξ σ α) define θ where "θ ≡ ξ ∘⇩s σ ∘⇩s α" let ?T = "dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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} ⋅⇩s⇩e⇩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⇩l⇩s⇩s⇩t S ℐ = {t. decl_ik S ⊢ t}" for S unfolding declassified⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_receive T ⋅⇩l⇩s⇩s⇩t θ)))" "list_all is_Check_or_Assignment (unlabel (dual⇩l⇩s⇩s⇩t (transaction_checks T ⋅⇩l⇩s⇩s⇩t θ)))" "list_all is_Update (unlabel (dual⇩l⇩s⇩s⇩t (transaction_updates T ⋅⇩l⇩s⇩s⇩t θ)))" "list_all is_Receive (unlabel (dual⇩l⇩s⇩s⇩t (transaction_send T ⋅⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t_list_all(1)[of "transaction_receive T ⋅⇩l⇩s⇩s⇩t θ"] dual⇩l⇩s⇩s⇩t_list_all(11)[of "transaction_checks T ⋅⇩l⇩s⇩s⇩t θ"] dual⇩l⇩s⇩s⇩t_list_all(10)[of "transaction_updates T ⋅⇩l⇩s⇩s⇩t θ"] dual⇩l⇩s⇩s⇩t_list_all(2)[of "transaction_send T ⋅⇩l⇩s⇩s⇩t θ"] unfolding f_def by (metis unlabel_subst[of _ θ])+ hence "¬list_ex is_Receive (unlabel (dual⇩l⇩s⇩s⇩t (transaction_receive T ⋅⇩l⇩s⇩s⇩t θ)))" "¬list_ex is_Receive (unlabel (dual⇩l⇩s⇩s⇩t (transaction_checks T ⋅⇩l⇩s⇩s⇩t θ)))" "¬list_ex is_Receive (unlabel (dual⇩l⇩s⇩s⇩t (transaction_updates T ⋅⇩l⇩s⇩s⇩t θ)))" "list_all is_Receive (unlabel (dual⇩l⇩s⇩s⇩t (transaction_send T ⋅⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_send T ⋅⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t_map_fst_eq[of "tl (transaction_send T ⋅⇩l⇩s⇩s⇩t θ)"] unfolding T(4) dual⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t TB ℐ = declassified⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t B ⋅⇩s⇩e⇩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⇩l⇩s⇩s⇩t (A@?T) ⋅⇩s⇩e⇩t ℐ ⊢ s ⋅ ℐ" using ideduct_mono[OF 6(2)] subst_all_mono[of _ _ ℐ] ik⇩s⇩s⇩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⇩l⇩s⇩s⇩t B ℐ = declassified⇩l⇩s⇩s⇩t (A@?T) ℐ" proof - have "declassified⇩l⇩s⇩s⇩t TB ℐ = declassified⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t (constr B)) ∧ transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (constr' p) = {}" "⋀D. dbupd⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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' ⋅⇩l⇩i⇩s⇩t ξ ∘⇩s σ ∘⇩s α" using a dual⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (constr' p) = {}" using in_ik⇩l⇩s⇩s⇩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' ⋅ ξ ∘⇩s σ ∘⇩s α" "s = s' ⋅ ξ ∘⇩s σ ∘⇩s α" using a dual⇩l⇩s⇩s⇩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⇩s⇩s⇩t (unlabel (constr' p)) I D = D" using dbupd⇩s⇩s⇩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⇩l⇩s⇩s⇩t (constr' p) ⋅⇩s⇩e⇩t I)" let ?D' = "dbupd⇩s⇩s⇩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⇩l⇩s⇩s⇩t B = ik⇩l⇩s⇩s⇩t A" proof show "ik⇩l⇩s⇩s⇩t B ⊆ ik⇩l⇩s⇩s⇩t A" unfolding Ts(1) B_def constr_def by (induct Ts) (auto simp add: ik⇩s⇩s⇩t_def) show "ik⇩l⇩s⇩s⇩t A ⊆ ik⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (constr Ts) ⊆ ik⇩l⇩s⇩s⇩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⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" unfolding constr'_def p by fastforce have T_in: "T ∈ set P" and ξ: "transaction_decl_subst ξ T" and σ: "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t (constr Ts))" and α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (constr (filter h Ts)) ⊆ trms⇩l⇩s⇩s⇩t (constr Ts)" unfolding constr_def unlabel_def by fastforce hence σ': "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩t (constr (filter h Ts)))" using σ unfolding transaction_fresh_subst_def by fast have "vars⇩l⇩s⇩s⇩t (constr (filter h Ts)) ⊆ vars⇩l⇩s⇩s⇩t (constr Ts)" unfolding constr_def unlabel_def vars⇩s⇩s⇩t_def by auto hence α': "transaction_renaming_subst α (filter g P) (vars⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t 𝒜" shows "Γ⇩v x = TAtom Value" proof - obtain T where T: "T ∈ set P" "Γ⇩v x ∈ Γ⇩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⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩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⇩l⇩s⇩s⇩t 𝒜" shows "(∃n. ℐ x = Fun (Val n) []) ∨ (∃n. ℐ x = Fun (PubConst Value n) [])" proof - have x': "Γ (ℐ x) = TAtom Value" "fv (ℐ x) = {}" "wf⇩t⇩r⇩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⟩⇩s⟩ ∧ a ≠ ⟨t in ⟨k⟩⇩s⟩ ∧ a ≠ ⟨t not in ⟨k⟩⇩s⟩ ∧ a ≠ delete⟨t,⟨k⟩⇩s⟩" shows "(l,⟨ac: t ∈ s⟩) ∈ set A ⟹ (∃u. s = ⟨u⟩⇩s ∧ u ≠ k)" (is "?A A ⟹ ?Q A") and "(l,⟨t not in s⟩) ∈ set A ⟹ (∃u. s = ⟨u⟩⇩s ∧ u ≠ k)" (is "?B A ⟹ ?Q A") and "(l,delete⟨t,s⟩) ∈ set A ⟹ (∃u. s = ⟨u⟩⇩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 "θ ≡ ξ ∘⇩s σ ∘⇩s α" let ?T' = "dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t_steps_iff(6) subst_lsst_memD(6) by blast obtain u where u: "s' = ⟨u⟩⇩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⇩l⇩s⇩s⇩t_steps_iff(7) subst_lsst_memD(9) by blast obtain u where u: "s' = ⟨u⟩⇩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⇩l⇩s⇩s⇩t_steps_iff(5) subst_lsst_memD(5) by blast obtain u where u: "s' = ⟨u⟩⇩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⟩⇩s ∧ u ≠ k)) ∧ (∀l t s. (l,⟨t not in s⟩) ∈ set (A@B) ⟶ (∃u. s = ⟨u⟩⇩s ∧ u ≠ k)) ∧ (∀l t s. (l,delete⟨t,s⟩) ∈ set (A@B) ⟶ (∃u. s = ⟨u⟩⇩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) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t A)" and T: "T ∈ set P" "Var x ∈ set T_ts" "Γ⇩v x = TAtom Value" "fv⇩s⇩e⇩t (set T_ts) = {x}" "∀n. ¬(Fun (Val n) [] ⊑⇩s⇩e⇩t set T_ts)" "T = Transaction (λ(). []) [x] [] [] T_upds [(l1,send⟨T_ts⟩)]" "T_upds = [] ∨ (T_upds = [(l2,insert⟨Var x, ⟨k⟩⇩s⟩)] ∧ (∀T ∈ set P. ∀(l,a) ∈ set (transaction_strand T). ∀t. a ≠ select⟨t,⟨k⟩⇩s⟩ ∧ a ≠ ⟨t in ⟨k⟩⇩s⟩ ∧ a ≠ ⟨t not in ⟨k⟩⇩s⟩ ∧ a ≠ delete⟨t,⟨k⟩⇩s⟩))" shows "∃B. A@B ∈ reachable_constraints P ∧ B ∈ reachable_constraints P ∧ vars⇩l⇩s⇩s⇩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) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t A ⟶ Fun (Val n) [] ∉ ik⇩l⇩s⇩s⇩t B) ∧ (∀n. Fun (Val n) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t B ⟶ Fun (Val n) [] ∈ ik⇩l⇩s⇩s⇩t B) ∧ N = {n. Fun (Val n) [] ∈ ik⇩l⇩s⇩s⇩t B} ∧ checks_not_k B ∧ (∀l a. (l,a) ∈ set B ∧ is_Insert a ⟶ (l = l2 ∧ (∃n. a = insert⟨Fun (Val n) [],⟨k⟩⇩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⇩l⇩s⇩s⇩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⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩t (A@B))" using insert.prems insert.hyps(2) B(7,8) ik⇩s⇩s⇩t_trms⇩s⇩s⇩t_subset[of "unlabel B"] unfolding unlabel_append trms⇩s⇩s⇩t_append by blast ultimately show "∀t ∈ subst_range σ. t ∉ subterms⇩s⇩e⇩t (trms⇩l⇩s⇩s⇩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⇩s⇩e⇩t (trms_transaction T)" using T(5,7) unfolding σ_def T(6) by fastforce qed (force simp add: T(3) σ_def wt⇩s⇩u⇩b⇩s⇩t_def)+ hence σ': "transaction_fresh_subst σ T (trms⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (A@B))" unfolding transaction_renaming_subst_def by blast hence α': "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t B)" unfolding transaction_renaming_subst_def by auto define θ where "θ ≡ ξ ∘⇩s σ ∘⇩s α" define C where "C ≡ dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t θ)" have θx: "θ x = Fun (Val n) []" unfolding θ_def ξ_def σ_def subst_compose by force have "dual⇩l⇩s⇩s⇩t (transaction_receive T ⋅⇩l⇩s⇩s⇩t θ) = []" "dual⇩l⇩s⇩s⇩t (transaction_checks T ⋅⇩l⇩s⇩s⇩t θ) = []" using T(6) unfolding θ_def ξ_def σ_def by auto moreover have "(T_upds = [] ∧ dual⇩l⇩s⇩s⇩t (transaction_updates T ⋅⇩l⇩s⇩s⇩t θ) = []) ∨ (T_upds ≠ [] ∧ dual⇩l⇩s⇩s⇩t (transaction_updates T ⋅⇩l⇩s⇩s⇩t θ) = [(l2,insert⟨θ x, ⟨k⟩⇩s⟩)])" using subst_lsst_cons[of "(l2,insert⟨Var x,⟨k⟩⇩s⟩)" "[]" θ] T(6,7) unfolding dual⇩l⇩s⇩s⇩t_subst by auto hence "(T_upds = [] ∧ dual⇩l⇩s⇩s⇩t (transaction_updates T ⋅⇩l⇩s⇩s⇩t θ) = []) ∨ (T_upds ≠ [] ∧ dual⇩l⇩s⇩s⇩t (transaction_updates T ⋅⇩l⇩s⇩s⇩t θ) = [(l2,insert⟨Fun (Val n) [],⟨k⟩⇩s⟩)])" unfolding θ_def ξ_def σ_def by (auto simp: subst_compose) moreover have "dual⇩l⇩s⇩s⇩t (transaction_send T ⋅⇩l⇩s⇩s⇩t θ) = [(l1, receive⟨T_ts ⋅⇩l⇩i⇩s⇩t θ⟩)]" using subst_lsst_cons[of "(l1, receive⟨T_ts⟩)" "[]" θ] T(6) unfolding dual⇩l⇩s⇩s⇩t_subst by auto hence "dual⇩l⇩s⇩s⇩t (transaction_send T ⋅⇩l⇩s⇩s⇩t θ) = [(l1, receive⟨T_ts ⋅⇩l⇩i⇩s⇩t θ⟩)]" by auto ultimately have C: "(T_upds = [] ∧ C = [(l1, receive⟨T_ts ⋅⇩l⇩i⇩s⇩t θ⟩)]) ∨ (T_upds ≠ [] ∧ C = [(l2, insert⟨Fun (Val n) [],⟨k⟩⇩s⟩), (l1, receive⟨T_ts ⋅⇩l⇩i⇩s⇩t θ⟩)])" unfolding C_def transaction_dual_subst_unfold by force have C': "Fun (Val n) [] ∈ set (T_ts ⋅⇩l⇩i⇩s⇩t θ)" "Fun (Val n) [] ∈ ik⇩l⇩s⇩s⇩t C" using T(2) in_ik⇩s⇩s⇩t_iff[of _ "unlabel C"] C unfolding θ_def ξ_def σ_def by (force, force) have "fv (t ⋅ θ) = {}" when t: "t ⊑⇩s⇩e⇩t set T_ts" for t proof - have "fv t ⊆ {x}" using t T(4) fv_subset_subterms by blast hence "fv (t ⋅ ξ ∘⇩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 ⋅⇩l⇩i⇩s⇩t θ))" by auto have 2: "m = n" when m: "Fun (Val m) [] ⊑⇩s⇩e⇩t ik⇩l⇩s⇩s⇩t C" for m proof - have "Fun (Val m) [] ⊑⇩s⇩e⇩t set (T_ts ⋅⇩l⇩i⇩s⇩t θ)" using m C in_ik⇩s⇩s⇩t_iff[of _ "unlabel C"] by fastforce hence *: "Fun (Val m) [] ⊑⇩s⇩e⇩t set T_ts ⋅⇩s⇩e⇩t θ" by simp show ?thesis using const_subterms_subst_cases[OF *] T(4,5) θx by fastforce qed have C_trms: "trms⇩l⇩s⇩s⇩t C ⊆ {Fun (Val n) [],⟨k⟩⇩s} ∪ ik⇩l⇩s⇩s⇩t C" using C in_ik⇩s⇩s⇩t_iff[of _ "unlabel C"] by fastforce have 3: "m = n" when m: "Fun (Val m) [] ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩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⇩s⇩s⇩t_append by blast have Q5: "?Q5 (B@C)" using B(7) C' 3 unfolding unlabel_append ik⇩s⇩s⇩t_append trms⇩s⇩s⇩t_append by blast have Q6: "?Q6 (insert n N) (B@C)" using B(8) C' 2 unfolding unlabel_append ik⇩s⇩s⇩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 "priv_consts_of X = {t. t ⊑⇩s⇩e⇩t X ∧ (∃c. t = Fun c [] ∧ ¬public c ∧ arity c = 0)}" private fun mk_symb where "mk_symb (ξ, σ, I, T, α) = dual⇩l⇩s⇩s⇩t((transaction_strand T) ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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) ∘⇩s I) (fv⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s var_rename n))" private fun invαI where "invαI ns w = foldl (∘⇩s) Var (map2 mk_invαI ns w)" private fun mk_I where "mk_I (ξ, σ, I, T, α) = narrow I (fv⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))" private fun comb_I where "comb_I w = fold (∘⇩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⇩l⇩s⇩s⇩t s = {}" private fun ground_step :: "(_, _) stateful_strand_step ⇒ bool" where "ground_step s ⟷ fv⇩s⇩s⇩t⇩p s = {}" private fun ground_lstep :: "_ strand_label × (_, _) stateful_strand_step ⇒ bool" where "ground_lstep (l,s) ⟷ fv⇩s⇩s⇩t⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α); strand_sem_stateful IK DB (unlabel A) I; interpretation⇩s⇩u⇩b⇩s⇩t I; wf⇩t⇩r⇩m⇩s (subst_range I) ⟧ ⟹ (IK ∪ ((ik⇩l⇩s⇩s⇩t A) ⋅⇩s⇩e⇩t I), dbupd⇩s⇩s⇩t (unlabel A) I DB, trms ∪ trms⇩l⇩s⇩s⇩t A, vars ∪ vars⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t((transaction_strand T) ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" using T_symb_append'[of w "[(ξ, σ, I, T,α)]"] by auto private lemma ground_step_subst: assumes "ground_step a" shows "a = a ⋅⇩s⇩s⇩t⇩p σ" using assms proof (induction a) case (NegChecks Y F F') then have FY: "fv⇩p⇩a⇩i⇩r⇩s F - set Y = {}" unfolding ground_step.simps unfolding fv⇩s⇩s⇩t⇩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 ⋅⇩p (rm_vars (set Y) σ) = f" by auto then have "F = F ⋅⇩p⇩a⇩i⇩r⇩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⇩p⇩a⇩i⇩r⇩s F' - set Y = {}" unfolding ground_step.simps unfolding fv⇩s⇩s⇩t⇩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 ⋅⇩p (rm_vars (set Y) σ) = f" by auto then have "F' = F' ⋅⇩p⇩a⇩i⇩r⇩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 ⋅⇩l⇩s⇩s⇩t⇩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⇩p⇩a⇩i⇩r ` set ps) - set X. I x = I' x" shows "ps ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) I = ps ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) I'" proof - have "∀p ∈ set ps. p ⋅⇩p rm_vars (set X) I = p ⋅⇩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⇩p⇩a⇩i⇩r⇩s.elims fv⇩p⇩a⇩i⇩r⇩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⇩p⇩a⇩i⇩r⇩s.elims fv⇩p⇩a⇩i⇩r⇩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 ⋅⇩p rm_vars (set X) I = p ⋅⇩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⇩s⇩s⇩t⇩p T. I x = I' x" shows "T ⋅⇩s⇩s⇩t⇩p I = T ⋅⇩s⇩s⇩t⇩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⇩p⇩a⇩i⇩r ` set F) ∪ ⋃(fv⇩p⇩a⇩i⇩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⇩s⇩s⇩t⇩p (snd T). I x = I' x" shows "T ⋅⇩l⇩s⇩s⇩t⇩p I = T ⋅⇩l⇩s⇩s⇩t⇩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⇩l⇩s⇩s⇩t T. I x = I' x" shows "T ⋅⇩l⇩s⇩s⇩t I = T ⋅⇩l⇩s⇩s⇩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⇩s⇩s⇩t_Cons subst_lsst_cons unlabel_Cons(2)) qed private lemma transaction_renaming_subst_not_in_fv⇩l⇩s⇩s⇩t: fixes ξ σ α::"('fun,'atom,'sets,'lbl) prot_subst" and A::"('fun,'atom,'sets,'lbl) prot_constr" assumes "x ∈ fv⇩l⇩s⇩s⇩t A" assumes "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t A)" shows "x ∉ fv⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" proof - have 0: "x ∉ fv⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s α)" when x: "x ∈ vars⇩l⇩s⇩s⇩t A" and α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩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⇩s⇩s⇩t_subst_obtain_var subst_compose unlabel_subst by (metis (no_types, opaque_lifting) subset_iff top_greatest) qed have 1: "x ∉ fv⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s α)" when x: "x ∈ fv⇩l⇩s⇩s⇩t A" and α: "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩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⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t) show ?thesis using 1 assms by metis qed private lemma wf_comb_I_Nil: "wf⇩t⇩r⇩m⇩s (subst_range (comb_I []))" by auto private lemma comb_I_append: "comb_I (w @ [(ξ, σ, I, T, α)]) = (mk_I (ξ, σ, I, T, α) ∘⇩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⇩l⇩s⇩s⇩t ((T_symb w) ⋅⇩l⇩s⇩s⇩t (comb_I w)) ∧ DB = dbupd⇩s⇩s⇩t (unlabel ((T_symb w))) (comb_I w) {} ∧ trms = trms⇩l⇩s⇩s⇩t (T_symb w) ∧ vars = vars⇩l⇩s⇩s⇩t (T_symb w) ∧ interpretation⇩s⇩u⇩b⇩s⇩t (comb_I w) ∧ wf⇩t⇩r⇩m⇩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⇩l⇩s⇩s⇩t (T_symb w ⋅⇩l⇩s⇩s⇩t comb_I w)" "DB = dbupd⇩s⇩s⇩t (unlabel (T_symb w)) (comb_I w) {}" "trms = trms⇩l⇩s⇩s⇩t (T_symb w)" "vars = vars⇩l⇩s⇩s⇩t (T_symb w)" "interpretation⇩s⇩u⇩b⇩s⇩t (comb_I w)" "wf⇩t⇩r⇩m⇩s (subst_range (comb_I w))" by auto define w' where "w' = w @ [(ξ, σ, I, T, α)]" have interps_w: "∀x ∈ fv⇩l⇩s⇩s⇩t (T_symb w). (comb_I w) x = (comb_I w') x" proof fix x assume "x ∈ fv⇩l⇩s⇩s⇩t (T_symb w)" then have "x ∉ fv⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" using step(5) transaction_renaming_subst_not_in_fv⇩l⇩s⇩s⇩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, α) ∘⇩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⇩s⇩s⇩t (unlabel (mk_symb (ξ, σ, I, T, α))). I x = (comb_I w') x" proof fix x assume "x ∈ fv⇩s⇩s⇩t (unlabel (mk_symb (ξ, σ, I, T, α)))" then have a: "x ∈ (fv⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))" by (metis fv⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq mk_symb.simps) have "(comb_I w') x = (mk_I (ξ, σ, I, T, α) ∘⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t (T_symb w))" using step(4) step'(5) by auto moreover have "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (T_symb w) ⋅⇩s⇩e⇩t comb_I w') (dbupd⇩s⇩s⇩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⇩l⇩s⇩s⇩t (T_symb w ⋅⇩l⇩s⇩s⇩t comb_I w)) (dbupd⇩s⇩s⇩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⇩s⇩e⇩t (ik⇩l⇩s⇩s⇩t (T_symb w)). comb_I w x = comb_I w' x" using interps_w by (metis fv_ik⇩s⇩s⇩t_is_fv⇩s⇩s⇩t) then have "∀x∈fv t. comb_I w x = comb_I w' x" when t: "t ∈ ik⇩l⇩s⇩s⇩t (T_symb w)" for t using t by auto then have "t ⋅ comb_I w' = t ⋅ comb_I w" when t: "t ∈ ik⇩l⇩s⇩s⇩t (T_symb w)" for t using t term_subst_eq[of t "comb_I w'" "comb_I w"] by metis then have "ik⇩l⇩s⇩s⇩t (T_symb w) ⋅⇩s⇩e⇩t comb_I w' = ik⇩l⇩s⇩s⇩t (T_symb w) ⋅⇩s⇩e⇩t comb_I w" by auto also have "... = ik⇩l⇩s⇩s⇩t (T_symb w ⋅⇩l⇩s⇩s⇩t comb_I w)" by (metis ik⇩s⇩s⇩t_subst unlabel_subst) finally have "ik⇩l⇩s⇩s⇩t (T_symb w) ⋅⇩s⇩e⇩t comb_I w' = ik⇩l⇩s⇩s⇩t (T_symb w ⋅⇩l⇩s⇩s⇩t comb_I w)" by auto } moreover { have "dbupd⇩s⇩s⇩t (unlabel (T_symb w)) (comb_I w) {} = dbupd⇩s⇩s⇩t (unlabel (T_symb w)) (comb_I w') {}" by (metis db⇩s⇩s⇩t_subst_swap[OF interps_w] db⇩s⇩s⇩t_set_is_dbupd⇩s⇩s⇩t empty_set) } ultimately have "strand_sem_stateful (ik⇩l⇩s⇩s⇩t (T_symb w) ⋅⇩s⇩e⇩t comb_I w') (dbupd⇩s⇩s⇩t (unlabel (T_symb w)) (comb_I w') {}) (unlabel (mk_symb (ξ, σ, I, T, α))) I" by force then show "strand_sem_stateful (ik⇩l⇩s⇩s⇩t (T_symb w) ⋅⇩s⇩e⇩t comb_I w') (dbupd⇩s⇩s⇩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⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I) = ik⇩l⇩s⇩s⇩t (T_symb w' ⋅⇩l⇩s⇩s⇩t comb_I w')" proof - have AI: "ik⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩t I) = ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I" by (metis ik⇩s⇩s⇩t_subst unlabel_subst) have "ik⇩l⇩s⇩s⇩t (T_symb w' ⋅⇩l⇩s⇩s⇩t comb_I w') = ik⇩l⇩s⇩s⇩t (T_symb w ⋅⇩l⇩s⇩s⇩t comb_I w') ∪ ik⇩l⇩s⇩s⇩t (T_symb [(ξ, σ, I, T, α)] ⋅⇩l⇩s⇩s⇩t comb_I w')" unfolding w'_def by (simp add: subst_lsst_append) also have "... = ik⇩l⇩s⇩s⇩t (T_symb w ⋅⇩l⇩s⇩s⇩t comb_I w') ∪ ik⇩l⇩s⇩s⇩t (T_symb [(ξ, σ, I, T, α)] ⋅⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (T_symb w ⋅⇩l⇩s⇩s⇩t comb_I w) ∪ ik⇩l⇩s⇩s⇩t (T_symb [(ξ, σ, I, T, α)] ⋅⇩l⇩s⇩s⇩t I)" by (metis interps_w subst_apply_labeled_stateful_strand_swap) also have "... = IK ∪ ik⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩t I)" using step'(3) step.hyps(6) by auto also have "... = IK ∪ (ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I)" unfolding AI by auto finally show "IK ∪ (ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I) = ik⇩l⇩s⇩s⇩t (T_symb w' ⋅⇩l⇩s⇩s⇩t comb_I w')" using step'(3) step(6) T_symb.simps mk_symb.simps by auto qed moreover have "dbupd⇩s⇩s⇩t (unlabel A) I DB = dbupd⇩s⇩s⇩t (unlabel (T_symb w')) (comb_I w') {}" proof - have "dbupd⇩s⇩s⇩t (unlabel A) I DB = dbupd⇩s⇩s⇩t (unlabel A) I (dbupd⇩s⇩s⇩t (unlabel (T_symb w)) (comb_I w) {})" using step'(4) by auto moreover have "... = dbupd⇩s⇩s⇩t (unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))) I (dbupd⇩s⇩s⇩t (unlabel (T_symb w)) (comb_I w) {})" using step(6) by auto moreover have "... = dbupd⇩s⇩s⇩t (unlabel (mk_symb (ξ, σ, I, T, α))) I (dbupd⇩s⇩s⇩t (unlabel (T_symb w)) (comb_I w) {})" by auto moreover have "... = dbupd⇩s⇩s⇩t (unlabel (mk_symb (ξ, σ, I, T, α))) (comb_I w') (dbupd⇩s⇩s⇩t (unlabel (T_symb w)) (comb_I w) {})" by (metis (no_types, lifting) db⇩s⇩s⇩t_subst_swap[OF interps_T] db⇩s⇩s⇩t_set_is_dbupd⇩s⇩s⇩t empty_set) moreover have "... = dbupd⇩s⇩s⇩t (unlabel (mk_symb (ξ, σ, I, T, α))) (comb_I w') (dbupd⇩s⇩s⇩t (unlabel (T_symb w)) (comb_I w') {})" by (metis (no_types, lifting) db⇩s⇩s⇩t_subst_swap[OF interps_w] db⇩s⇩s⇩t_set_is_dbupd⇩s⇩s⇩t empty_set) moreover have "... = dbupd⇩s⇩s⇩t (unlabel (T_symb w) @ unlabel (mk_symb (ξ, σ, I, T, α))) (comb_I w') {}" using dbupd⇩s⇩s⇩t_append by metis moreover have "... = dbupd⇩s⇩s⇩t (unlabel ((T_symb w) @ mk_symb (ξ, σ, I, T, α))) (comb_I w') {}" by auto moreover have "... = dbupd⇩s⇩s⇩t (unlabel (T_symb w')) (comb_I w') {}" unfolding w'_def by auto ultimately show "dbupd⇩s⇩s⇩t (unlabel A) I DB = dbupd⇩s⇩s⇩t (unlabel (T_symb w')) (comb_I w') {}" by auto qed moreover have "trms ∪ trms⇩l⇩s⇩s⇩t A = trms⇩l⇩s⇩s⇩t (T_symb w')" proof - have "trms ∪ trms⇩l⇩s⇩s⇩t A = trms⇩l⇩s⇩s⇩t (T_symb w) ∪ trms⇩l⇩s⇩s⇩t A" using step'(5) by auto moreover have "... = trms⇩l⇩s⇩s⇩t (T_symb w) ∪ trms⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))" using step(6) by auto moreover have "... = trms⇩l⇩s⇩s⇩t (T_symb w) ∪ trms⇩l⇩s⇩s⇩t (T_symb [(ξ, σ, I, T, α)])" by auto moreover have "... = trms⇩l⇩s⇩s⇩t (T_symb w @ T_symb [(ξ, σ, I, T, α)])" by auto moreover have "... = trms⇩l⇩s⇩s⇩t (T_symb w')" unfolding w'_def by auto ultimately show "trms ∪ trms⇩l⇩s⇩s⇩t A = trms⇩l⇩s⇩s⇩t (T_symb w')" by auto qed moreover have "vars ∪ vars⇩l⇩s⇩s⇩t A = vars⇩l⇩s⇩s⇩t (T_symb w')" proof - have "vars ∪ vars⇩l⇩s⇩s⇩t A = vars⇩l⇩s⇩s⇩t (T_symb w) ∪ vars⇩l⇩s⇩s⇩t A" using step'(6) by fastforce moreover have "... = vars⇩l⇩s⇩s⇩t (T_symb w) ∪ vars⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))" using step(6) by auto moreover have "... = vars⇩l⇩s⇩s⇩t (T_symb w) ∪ vars⇩l⇩s⇩s⇩t (T_symb [(ξ, σ, I, T, α)])" by auto moreover have "... = vars⇩l⇩s⇩s⇩t (T_symb w @ T_symb [(ξ, σ, I, T, α)])" by auto moreover have "... = vars⇩l⇩s⇩s⇩t (T_symb w')" unfolding w'_def by auto ultimately show "vars ∪ vars⇩l⇩s⇩s⇩t A = vars⇩l⇩s⇩s⇩t (T_symb w')" using step(6) by auto qed moreover have interp_comb_I_w': "interpretation⇩s⇩u⇩b⇩s⇩t (comb_I w')" using interpretation_comp(1) step'(7) unfolding w'_def by auto moreover have "wf⇩t⇩r⇩m⇩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⇩t⇩r⇩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⇩s⇩u⇩b⇩s⇩t I" assumes "wf⇩t⇩r⇩m⇩s (subst_range I)" shows "∃w. (ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I, dbupd⇩s⇩s⇩t (unlabel A) I {}, trms⇩l⇩s⇩s⇩t A, vars⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t I, dbupd⇩s⇩s⇩t (unlabel 𝒜) I {}, trms⇩l⇩s⇩s⇩t 𝒜, vars⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t I, dbupd⇩s⇩s⇩t (unlabel 𝒜) I {}, trms⇩l⇩s⇩s⇩t 𝒜, vars⇩l⇩s⇩s⇩t 𝒜, w) ∈ ground_protocol_states_aux P" by auto define w' where "w' = w@[(ξ, σ, I, T, α)]" define 𝒜' where "𝒜' = 𝒜@dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" let ?T = "unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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⇩l⇩s⇩s⇩t 𝒜)" using step.hyps(4) . moreover have "transaction_renaming_subst α P (vars⇩l⇩s⇩s⇩t 𝒜)" using step.hyps(5) . moreover have "strand_sem_stateful (ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t I) (dbupd⇩s⇩s⇩t (unlabel 𝒜) I {}) ?T I" using step(7) strand_sem_append_stateful[of "{}" "{}" "unlabel 𝒜" ?T I] by auto moreover have "interpretation⇩s⇩u⇩b⇩s⇩t I" using assms(3) . moreover have "wf⇩t⇩r⇩m⇩s (subst_range I)" using step.prems(3) by fastforce ultimately have "((ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t I) ∪ (ik⇩s⇩s⇩t ?T ⋅⇩s⇩e⇩t I), dbupd⇩s⇩s⇩t ?T I (dbupd⇩s⇩s⇩t (unlabel 𝒜) I {}), trms⇩l⇩s⇩s⇩t 𝒜 ∪ trms⇩s⇩s⇩t ?T, vars⇩l⇩s⇩s⇩t 𝒜 ∪ vars⇩s⇩s⇩t ?T, w@[(ξ, σ, I, T, α)]) ∈ ground_protocol_states_aux P" using ground_protocol_states_aux.step[ OF w_p, of T ξ σ α "dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" I] by metis moreover have "ik⇩l⇩s⇩s⇩t 𝒜' ⋅⇩s⇩e⇩t I = (ik⇩l⇩s⇩s⇩t 𝒜 ⋅⇩s⇩e⇩t I) ∪ (ik⇩s⇩s⇩t ?T ⋅⇩s⇩e⇩t I)" unfolding 𝒜'_def by auto moreover have "dbupd⇩s⇩s⇩t (unlabel 𝒜') I {} = dbupd⇩s⇩s⇩t ?T I (dbupd⇩s⇩s⇩t (unlabel 𝒜) I {})" unfolding 𝒜'_def by (simp add: dbupd⇩s⇩s⇩t_append) moreover have "trms⇩l⇩s⇩s⇩t 𝒜' = trms⇩l⇩s⇩s⇩t 𝒜 ∪ trms⇩s⇩s⇩t ?T" unfolding 𝒜'_def by auto moreover have "vars⇩l⇩s⇩s⇩t 𝒜' = vars⇩l⇩s⇩s⇩t 𝒜 ∪ vars⇩s⇩s⇩t ?T" unfolding 𝒜'_def by auto ultimately have "(ik⇩l⇩s⇩s⇩t 𝒜' ⋅⇩s⇩e⇩t I, dbupd⇩s⇩s⇩t (unlabel 𝒜') I {}, trms⇩l⇩s⇩s⇩t 𝒜', vars⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩t I), dbupd⇩s⇩s⇩t (unlabel A) I {}) | A I. A ∈ reachable_constraints P ∧ strand_sem_stateful {} {} (unlabel A) I ∧ interpretation⇩s⇩u⇩b⇩s⇩t I ∧ wf⇩t⇩r⇩m⇩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⇩l⇩s⇩s⇩t (T_symb w ⋅⇩l⇩s⇩s⇩t comb_I w)" "DB = dbupd⇩s⇩s⇩t (unlabel (T_symb w)) (comb_I w) {}" "trms = trms⇩l⇩s⇩s⇩t (T_symb w)" "vars = vars⇩l⇩s⇩s⇩t (T_symb w)" "interpretation⇩s⇩u⇩b⇩s⇩t (comb_I w)" "wf⇩t⇩r⇩m⇩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⇩l⇩s⇩s⇩t (T_symb w ⋅⇩l⇩s⇩s⇩t (comb_I w))" "DB = dbupd⇩s⇩s⇩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⇩s⇩u⇩b⇩s⇩t (comb_I w) ∧ wf⇩t⇩r⇩m⇩s (subst_range (comb_I w))" by auto then show "∃A I. (IK, DB) = (ik⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩t I), dbupd⇩s⇩s⇩t (unlabel A) I {}) ∧ A ∈ reachable_constraints P ∧ strand_sem_stateful {} {} (unlabel A) I ∧ interpretation⇩s⇩u⇩b⇩s⇩t I ∧ wf⇩t⇩r⇩m⇩s (subst_range I)" by blast next fix IK DB assume "(IK, DB) ∈ {(ik⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩t I), dbupd⇩s⇩s⇩t (unlabel A) I {}) | A I. A ∈ reachable_constraints P ∧ strand_sem_stateful {} {} (unlabel A) I ∧ interpretation⇩s⇩u⇩b⇩s⇩t I ∧ wf⇩t⇩r⇩m⇩s (subst_range I)}" then obtain A I where A_I_p: "IK = ik⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩t I)" "DB = dbupd⇩s⇩s⇩t (unlabel A) I {}" "A ∈ reachable_constraints P" "strand_sem_stateful {} {} (unlabel A) I" "interpretation⇩s⇩u⇩b⇩s⇩t I" "wf⇩t⇩r⇩m⇩s (subst_range I)" by auto then have "∃w. (ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I, dbupd⇩s⇩s⇩t (unlabel A) I {}, trms⇩l⇩s⇩s⇩t A, vars⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I, DB, trms⇩l⇩s⇩s⇩t A, vars⇩l⇩s⇩s⇩t A, w) ∈ ground_protocol_states_aux P" using A_I_p by blast then have "∃w. (ik⇩s⇩s⇩t (unlabel A ⋅⇩s⇩s⇩t I), DB, trms⇩l⇩s⇩s⇩t A, vars⇩l⇩s⇩s⇩t A, w) ∈ ground_protocol_states_aux P" by (simp add: ik⇩s⇩s⇩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 "(δ ∘⇩s α) = (δ ∘⇩s (rm_vars (set X) α))" proof (rule ext) fix x show "(δ ∘⇩s α) x = (δ ∘⇩s rm_vars (set X) α) x" proof (cases "x ∈ set X") case True have gt: "ground_term (δ x)" using True assms by auto have "(δ ∘⇩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 "... = (δ ∘⇩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 "(δ ∘⇩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 "... = (δ ∘⇩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 "(δ ∘⇩s (rm_vars (set X) α)) = (rm_vars (set X) α ∘⇩s δ)" proof (rule ext) fix x show "(δ ∘⇩s rm_vars (set X) α) x = (rm_vars (set X) α ∘⇩s δ) x" proof (cases "x ∈ set X") case True then have gt: "ground_term (δ x)" using True assms(3,4) by auto have "(δ ∘⇩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) α ∘⇩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 (α ∘⇩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 ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α) (F' ⋅⇩p⇩a⇩i⇩r⇩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 ⋅ δ ∘⇩s (α ∘⇩s I) ≠ s ⋅ δ ∘⇩s (α ∘⇩s I)) ∨ (∃(t, s)∈set F'. (t, s) ⋅⇩p δ ∘⇩s (α ∘⇩s I) ∉ DB)" using a assms(1) unfolding negchecks_model_def by auto then show "(∃(t, s)∈set (F ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α). t ⋅ δ ∘⇩s I ≠ s ⋅ δ ∘⇩s I) ∨ (∃(t, s)∈set (F' ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α). (t, s) ⋅⇩p δ ∘⇩s I ∉ DB)" proof assume "∃(t, s)∈set F. t ⋅ δ ∘⇩s (α ∘⇩s I) ≠ s ⋅ δ ∘⇩s (α ∘⇩s I)" then obtain t s where t_s_p: "(t, s)∈set F" "t ⋅ δ ∘⇩s (α ∘⇩s I) ≠ s ⋅ δ ∘⇩s (α ∘⇩s I)" by auto from this(2) have "t ⋅ δ ∘⇩s ((rm_vars (set X) α) ∘⇩s I) ≠ s ⋅ δ ∘⇩s ((rm_vars (set X) α) ∘⇩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) α) ∘⇩s ( δ ∘⇩s I) ≠ s ⋅ (rm_vars (set X) α) ∘⇩s (δ ∘⇩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) α) ⋅ ( δ ∘⇩s I) ≠ s ⋅ (rm_vars (set X) α) ⋅ (δ ∘⇩s I)" by auto moreover have "(t ⋅ rm_vars (set X) α, s ⋅ rm_vars (set X) α) ∈ set (F ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α)" using subst_apply_pairs_pset_subst t_s_p(1) by fastforce ultimately have "∃(t, s)∈set (F ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α). t ⋅ δ ∘⇩s I ≠ s ⋅ δ ∘⇩s I" by auto then show ?thesis by auto next assume "∃(t, s)∈set F'. (t, s) ⋅⇩p δ ∘⇩s (α ∘⇩s I) ∉ DB" then obtain t s where t_s_p: "(t, s) ∈ set F'" "(t, s) ⋅⇩p δ ∘⇩s (α ∘⇩s I) ∉ DB" by auto from this(2) have "(t, s) ⋅⇩p δ ∘⇩s (rm_vars (set X) α ∘⇩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) ⋅⇩p rm_vars (set X) α ∘⇩s (δ ∘⇩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) α) ⋅⇩p δ ∘⇩s I ∉ DB" by auto moreover have "(t ⋅ rm_vars (set X) α, s ⋅ rm_vars (set X) α) ∈ set (F' ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α)" using t_s_p(1) subst_apply_pairs_pset_subst by fastforce ultimately have "(∃(t, s) ∈ set (F' ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α). (t, s) ⋅⇩p δ ∘⇩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 ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α) (F' ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α)" assumes "subst_range α ⊆ range Var" assumes "∀x∈set X. ∀y. α y ≠ Var x" shows "negchecks_model (α ∘⇩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 ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α). t ⋅ δ ∘⇩s I ≠ s ⋅ δ ∘⇩s (I)) ∨ (∃(t, s)∈set (F' ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α). (t, s) ⋅⇩p δ ∘⇩s I ∉ DB)" using a assms(1)unfolding negchecks_model_def by auto then show "(∃(t, s)∈set F. t ⋅ δ ∘⇩s (α ∘⇩s I) ≠ s ⋅ δ ∘⇩s (α ∘⇩s I)) ∨ (∃(t, s)∈set F'. (t, s) ⋅⇩p δ ∘⇩s (α ∘⇩s I) ∉ DB)" proof assume "∃(t, s)∈set (F ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α). t ⋅ δ ∘⇩s I ≠ s ⋅ δ ∘⇩s (I)" then obtain t s where t_s_p: "(t, s) ∈ set (F ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α)" "t ⋅ δ ∘⇩s I ≠ s ⋅ δ ∘⇩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) α ⋅ δ ∘⇩s I ≠ s' ⋅ rm_vars (set X) α ⋅ δ ∘⇩s I" using t_s_p by auto then have "t' ⋅ δ ⋅ rm_vars (set X) α ∘⇩s I ≠ s' ⋅ δ ⋅ rm_vars (set X) α ∘⇩s I" using a subst_comp_rm_vars_commute[OF assms(3,2)] by (metis (no_types, lifting) subst_subst) then have "t' ⋅ δ ⋅ α ∘⇩s I ≠ s' ⋅ δ ⋅ α ∘⇩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 ⋅ δ ∘⇩s (α ∘⇩s I) ≠ s ⋅ δ ∘⇩s (α ∘⇩s I)" by auto then show ?thesis by auto next assume "∃(t, s)∈set (F' ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α). (t, s) ⋅⇩p δ ∘⇩s I ∉ DB" then obtain t s where t_s_p: "(t, s) ∈ set (F' ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α)" "(t ⋅ δ ∘⇩s I, s ⋅ δ ∘⇩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) α ⋅ δ ∘⇩s I, s' ⋅ rm_vars (set X) α ⋅ δ ∘⇩s I) ∉ DB" using t_s_p by auto then have "(t' ⋅ δ ⋅ rm_vars (set X) α ∘⇩s I, s' ⋅ δ ⋅ rm_vars (set X) α ∘⇩s I) ∉ DB" using a subst_comp_rm_vars_commute[OF assms(3,2)] by (metis (no_types, lifting) subst_subst) then have "(t' ⋅ δ ⋅ α ∘⇩s I , s' ⋅ δ ⋅ α ∘⇩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) ⋅⇩p δ ∘⇩s (α ∘⇩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 (α ∘⇩s I) DB X F F' ⟷ negchecks_model I DB X (F ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α) (F' ⋅⇩p⇩a⇩i⇩r⇩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⇩s⇩s⇩t T. ∀y. α y ≠ Var x" shows "strand_sem_stateful IK DB (T ⋅⇩s⇩s⇩t α) I = strand_sem_stateful IK DB T (α ∘⇩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 ⋅⇩s⇩e⇩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⇩s⇩s⇩t T. ∀y. α y ≠ Var x" using NegChecks by auto have "∀x∈bvars⇩s⇩s⇩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 ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α) (F' ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α) ⟷ negchecks_model (α ∘⇩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 ⋅⇩s⇩s⇩t α) I ⟷ strand_sem_stateful IK DB T (α ∘⇩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) (ξσ ∘⇩s α)) = (rm_vars (set X) ξσ ∘⇩s rm_vars (set X) α)" proof (rule ext) fix x show "rm_vars (set X) (ξσ ∘⇩s α) x = (rm_vars (set X) ξσ ∘⇩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 ⋅⇩s⇩s⇩t ξσ ∘⇩s α = (T ⋅⇩s⇩s⇩t ξσ) ⋅⇩s⇩s⇩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 ⋅⇩l⇩s⇩s⇩t ξσ ∘⇩s α = (T ⋅⇩l⇩s⇩s⇩t ξσ) ⋅⇩l⇩s⇩s⇩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⇩s⇩e⇩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 (ξ ∘⇩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 (ξ ∘⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α))) I ⟷ strand_sem_stateful IK DB (unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ))) (α ∘⇩s I)" proof - have α_Var: "subst_range α ⊆ range Var" using assms(1) transaction_renaming_subst_is_renaming(5) by blast have "(∀x∈vars⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T). ∀y. α y ≠ Var x)" by (metis UnCI vars⇩s⇩s⇩t_is_fv⇩s⇩s⇩t_bvars⇩s⇩s⇩t) then have T_Vars: "(∀x∈bvars⇩l⇩s⇩s⇩t (dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ)). ∀y. α y ≠ Var x)" by (metis bvars⇩l⇩s⇩s⇩t_subst bvars⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq) have ground_ξ_σ: "ground (subst_range (ξ ∘⇩s σ))" using fresh_transaction_decl_subst_ground_subst_range using assms(2) assms(3) by blast from assms(1) ground_ξ_σ have "unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand T) ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α) = unlabel ((dual⇩l⇩s⇩s⇩t (transaction_strand T) ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ) ⋅⇩l⇩s⇩s⇩t α)" using stateful_strand_ground_subst_comp[of _ "unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand T))"] by (simp add: dual⇩l⇩s⇩s⇩t_subst unlabel_subst) then show ?thesis using strand_sem_stateful_substitution_lemma α_Var T_Vars by (metis dual⇩l⇩s⇩s⇩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⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ); strand_sem_stateful IK DB (unlabel A) I; interpretation⇩s⇩u⇩b⇩s⇩t I; wf⇩t⇩r⇩m⇩s (subst_range I) ⟧ ⟹ (IK ∪ ((ik⇩l⇩s⇩s⇩t A) ⋅⇩s⇩e⇩t I), dbupd⇩s⇩s⇩t (unlabel A) I DB, consts ∪ {t. t ⊑⇩s⇩e⇩t trms⇩l⇩s⇩s⇩t A ∧ (∃c. t = Fun c [] ∧ ¬public c ∧ arity c = 0)}) ∈ ground_protocol_states P" private lemma transaction_fresh_subst_priv_consts_of_iff: "transaction_fresh_subst σ T (priv_consts_of trms) ⟷ transaction_fresh_subst σ T trms" proof (cases "transaction_fresh_subst σ T (priv_consts_of trms) ∨ transaction_fresh_subst σ T trms") case True then have "∀t ∈ subst_range σ. ∃c. t = Fun c [] ∧ ¬ public c ∧ arity c = 0" unfolding transaction_fresh_subst_def by auto have "(∀t ∈ subst_range σ. t ∈ subterms⇩s⇩e⇩t (priv_consts_of trms) ⟷ t ∈ subterms⇩s⇩e⇩t trms)" proof fix t assume "t ∈ subst_range σ" then obtain c where c_p: "t = Fun c [] ∧ ¬ public c ∧ arity c = 0" using ‹∀t∈subst_range σ. ∃c. t = Fun c [] ∧ ¬ public c ∧ arity c = 0› by blast then have "Fun c [] ∈ subterms⇩s⇩e⇩t (priv_consts_of trms) ⟷ Fun c [] ∈ subterms⇩s⇩e⇩t trms" unfolding priv_consts_of_def by auto then show "t ∈ subterms⇩s⇩e⇩t (priv_consts_of trms) ⟷ t ∈ subterms⇩s⇩e⇩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. α ∘⇩s αinv = Var ∧ wf⇩t⇩r⇩m⇩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⇩t⇩r⇩m⇩s_range(2)) private lemma priv_consts_of_union_distr: "priv_consts_of (trms1 ∪ trms2) = priv_consts_of trms1 ∪ priv_consts_of trms2" unfolding priv_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⇩s⇩s⇩t_substitution_lemma: "dbupd⇩s⇩s⇩t T (α ∘⇩s I) DB = dbupd⇩s⇩s⇩t (T ⋅⇩s⇩s⇩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_priv_consts_of: assumes "subst_range α ⊆ range Var" shows "priv_consts_of T = priv_consts_of (T ⋅⇩s⇩e⇩t α)" proof (rule antisym; rule subsetI) fix x assume "x ∈ priv_consts_of T" then obtain t c where t_c_p: "t ∈ T ∧ x ⊑ t ∧ x = Fun c [] ∧ ¬ public c ∧ arity c = 0" unfolding priv_consts_of_def by auto moreover have "x ⊑ t ⋅ α" using t_c_p by (meson assms subst_Var_const_subterm_subst) ultimately show "x ∈ priv_consts_of (T ⋅⇩s⇩e⇩t α)" unfolding priv_consts_of_def by auto next fix x assume "x ∈ priv_consts_of (T ⋅⇩s⇩e⇩t α)" then obtain t c where t_c_p: "t ∈ T ∧ x ⊑ t ⋅ α ∧ x = Fun c [] ∧ ¬ public c ∧ arity c = 0" unfolding priv_consts_of_def by auto moreover have "x ⊑ t" using t_c_p by (meson assms subst_Var_const_subterm_subst) ultimately show "x ∈ priv_consts_of T" unfolding priv_consts_of_def by auto qed private lemma fst_set_subst_apply_set: "fst ` set F ⋅⇩s⇩e⇩t α = fst ` (set F ⋅⇩p⇩s⇩e⇩t α)" by (induction F) auto private lemma snd_set_subst_apply_set: "snd ` set F ⋅⇩s⇩e⇩t α = snd ` (set F ⋅⇩p⇩s⇩e⇩t α)" by (induction F) auto private lemma trms⇩p⇩a⇩i⇩r⇩s_fst_snd: "trms⇩p⇩a⇩i⇩r⇩s F = fst ` set F ∪ snd ` set F" by (auto simp add: rev_image_eqI) private lemma priv_consts_of_trms⇩s⇩s⇩t⇩p_range_Var: assumes "subst_range α ⊆ range Var" shows "priv_consts_of (trms⇩s⇩s⇩t⇩p a) = priv_consts_of (trms⇩s⇩s⇩t⇩p (a ⋅⇩s⇩s⇩t⇩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 "priv_consts_of (trms⇩p⇩a⇩i⇩r⇩s F) = priv_consts_of (fst ` set F ∪ snd ` set F)" using trms⇩p⇩a⇩i⇩r⇩s_fst_snd by metis also have "... = priv_consts_of (fst ` set F) ∪ priv_consts_of (snd ` set F)" using priv_consts_of_union_distr by blast also have "... = priv_consts_of ((fst ` set F) ⋅⇩s⇩e⇩t rm_vars (set X) α) ∪ priv_consts_of ((snd ` set F) ⋅⇩s⇩e⇩t rm_vars (set X) α)" using α_subs_rng_Var subst_Var_priv_consts_of[of "rm_vars (set X) α" "fst ` set F"] subst_Var_priv_consts_of[of "rm_vars (set X) α" "snd ` set F"] by auto also have "... = priv_consts_of (((fst ` set F) ⋅⇩s⇩e⇩t rm_vars (set X) α) ∪ ((snd ` set F) ⋅⇩s⇩e⇩t rm_vars (set X) α))" using priv_consts_of_union_distr by auto also have "... = priv_consts_of (fst ` set (F ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α) ∪ snd ` set (F ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α))" unfolding subst_apply_pairs_def fst_set_subst_apply_set snd_set_subst_apply_set by simp also have "... = priv_consts_of (trms⇩p⇩a⇩i⇩r⇩s (F ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α))" using trms⇩p⇩a⇩i⇩r⇩s_fst_snd[of "F ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α"] by metis finally have 1: "priv_consts_of (trms⇩p⇩a⇩i⇩r⇩s F) = priv_consts_of (trms⇩p⇩a⇩i⇩r⇩s (F ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α))" by auto have "priv_consts_of (trms⇩p⇩a⇩i⇩r⇩s F') = priv_consts_of (fst ` set F' ∪ snd ` set F')" using trms⇩p⇩a⇩i⇩r⇩s_fst_snd by metis also have "... = priv_consts_of (fst ` set F') ∪ priv_consts_of (snd ` set F')" using priv_consts_of_union_distr by blast also have "... = priv_consts_of ((fst ` set F') ⋅⇩s⇩e⇩t rm_vars (set X) α) ∪ priv_consts_of ((snd ` set F') ⋅⇩s⇩e⇩t rm_vars (set X) α)" using subst_Var_priv_consts_of[of "rm_vars (set X) α" "fst ` set F'"] α_subs_rng_Var subst_Var_priv_consts_of[of "rm_vars (set X) α" "snd ` set F'"] by auto also have "... = priv_consts_of ((fst ` set F' ⋅⇩s⇩e⇩t rm_vars (set X) α) ∪ (snd ` set F' ⋅⇩s⇩e⇩t rm_vars (set X) α))" using priv_consts_of_union_distr by auto also have "... = priv_consts_of (fst ` set (F' ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α) ∪ snd ` set (F' ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α))" unfolding subst_apply_pairs_def fst_set_subst_apply_set snd_set_subst_apply_set by simp also have "... = priv_consts_of (trms⇩p⇩a⇩i⇩r⇩s (F' ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α))" using trms⇩p⇩a⇩i⇩r⇩s_fst_snd[of "F' ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α"] by metis finally have 2: "priv_consts_of (trms⇩p⇩a⇩i⇩r⇩s F') = priv_consts_of (trms⇩p⇩a⇩i⇩r⇩s (F' ⋅⇩p⇩a⇩i⇩r⇩s rm_vars (set X) α))" by auto show ?case using 1 2 by (simp add: priv_consts_of_union_distr) qed (use subst_Var_priv_consts_of[of _ "{_, _}", OF assms] subst_Var_priv_consts_of[OF assms] in auto) private lemma priv_consts_of_trms⇩s⇩s⇩t_range_Var: assumes "subst_range α ⊆ range Var" shows "priv_consts_of (trms⇩s⇩s⇩t T) = priv_consts_of (trms⇩s⇩s⇩t (T ⋅⇩s⇩s⇩t α))" proof (induction T) case Nil then show ?case by auto next case (Cons a T) have "priv_consts_of (trms⇩s⇩s⇩t (a # T)) = priv_consts_of (trms⇩s⇩s⇩t [a] ∪ trms⇩s⇩s⇩t T)" by simp also have "... = priv_consts_of (trms⇩s⇩s⇩t [a]) ∪ priv_consts_of (trms⇩s⇩s⇩t T)" using priv_consts_of_union_distr by simp also have "... = priv_consts_of (trms⇩s⇩s⇩t⇩p a) ∪ priv_consts_of (trms⇩s⇩s⇩t T)" by simp also have "... = priv_consts_of (trms⇩s⇩s⇩t⇩p (a ⋅⇩s⇩s⇩t⇩p α)) ∪ priv_consts_of (trms⇩s⇩s⇩t T)" using priv_consts_of_trms⇩s⇩s⇩t⇩p_range_Var[OF assms] by simp also have "... = priv_consts_of (trms⇩s⇩s⇩t ([a] ⋅⇩s⇩s⇩t α)) ∪ priv_consts_of (trms⇩s⇩s⇩t T)" by (simp add: subst_apply_stateful_strand_def) also have "... = priv_consts_of (trms⇩s⇩s⇩t ([a] ⋅⇩s⇩s⇩t α)) ∪ priv_consts_of (trms⇩s⇩s⇩t (T ⋅⇩s⇩s⇩t α))" using local.Cons by simp also have "... = priv_consts_of (trms⇩s⇩s⇩t (a # T ⋅⇩s⇩s⇩t α))" by (simp add: priv_consts_of_union_distr subst_sst_cons) finally show ?case by simp qed private lemma priv_consts_of_trms⇩l⇩s⇩s⇩t_range_Var: assumes "subst_range α ⊆ range Var" shows "priv_consts_of (trms⇩l⇩s⇩s⇩t T) = priv_consts_of (trms⇩l⇩s⇩s⇩t (T ⋅⇩l⇩s⇩s⇩t α))" using priv_consts_of_trms⇩s⇩s⇩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, priv_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 priv_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 (ξ ∘⇩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' = α ∘⇩s I" define A' where "A' = dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ)" have "(IK, DB, priv_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 (priv_consts_of trms)" using step transaction_fresh_subst_priv_consts_of_iff by force moreover have "A' = dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩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⇩s⇩u⇩b⇩s⇩t I'" using step(8) unfolding I'_def by (meson interpretation_comp(1)) moreover have "wf⇩t⇩r⇩m⇩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⇩l⇩s⇩s⇩t A' ⋅⇩s⇩e⇩t I'), dbupd⇩s⇩s⇩t (unlabel A') I' DB, priv_consts_of trms ∪ priv_consts_of (trms⇩l⇩s⇩s⇩t A')) ∈ ground_protocol_states P" using ground_protocol_states.step[of IK DB "priv_consts_of trms" P T ξ σ A' I'] unfolding priv_consts_of_def by blast moreover have "ik⇩l⇩s⇩s⇩t A' ⋅⇩s⇩e⇩t I' = ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I" proof - have "ik⇩l⇩s⇩s⇩t A' ⋅⇩s⇩e⇩t I' = ik⇩l⇩s⇩s⇩t A' ⋅⇩s⇩e⇩t α ∘⇩s I" unfolding A'_def I'_def step(6) by auto also have "... = ik⇩l⇩s⇩s⇩t ((A' ⋅⇩l⇩s⇩s⇩t α) ⋅⇩l⇩s⇩s⇩t I)" unfolding unlabel_subst[symmetric] ik⇩s⇩s⇩t_subst by auto also have "... = ik⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩t I)" unfolding A'_def step(6) using labelled_stateful_strand_ground_subst_comp[of _ "transaction_strand T", OF ground_ξσ] by (simp add: dual⇩l⇩s⇩s⇩t_subst) also have "... = ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I" by (metis ik⇩s⇩s⇩t_subst unlabel_subst) finally show "ik⇩l⇩s⇩s⇩t A' ⋅⇩s⇩e⇩t I' = ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I" by auto qed moreover have "dbupd⇩s⇩s⇩t (unlabel A') I' DB = dbupd⇩s⇩s⇩t (unlabel A) I DB" proof - have "dbupd⇩s⇩s⇩t (unlabel A') I' DB = dbupd⇩s⇩s⇩t (unlabel A') (α ∘⇩s I) DB" unfolding A'_def I'_def step(6) using step by auto also have "... = dbupd⇩s⇩s⇩t (unlabel A' ⋅⇩s⇩s⇩t α) I DB" using dbupd⇩s⇩s⇩t_substitution_lemma by metis also have "... = dbupd⇩s⇩s⇩t (unlabel A) I DB" unfolding A'_def step(6) using stateful_strand_ground_subst_comp[of _ "unlabel (dual⇩l⇩s⇩s⇩t (transaction_strand T))"] ground_ξσ by (simp add: dual⇩l⇩s⇩s⇩t_subst_unlabel) finally show "dbupd⇩s⇩s⇩t (unlabel A') I' DB = dbupd⇩s⇩s⇩t (unlabel A) I DB" by auto qed moreover have "priv_consts_of (trms⇩l⇩s⇩s⇩t A') = priv_consts_of (trms⇩l⇩s⇩s⇩t A)" by (metis (no_types, lifting) A'_def α_Var priv_consts_of_trms⇩l⇩s⇩s⇩t_range_Var ground_ξσ labelled_stateful_strand_ground_subst_comp step.hyps(6) trms⇩s⇩s⇩t_unlabel_dual⇩l⇩s⇩s⇩t_eq) ultimately show ?case using priv_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 = priv_consts_of trms ∧ vars = vars⇩l⇩s⇩s⇩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 "{} = priv_consts_of {}" unfolding priv_consts_of_def by auto moreover have "{} = vars⇩l⇩s⇩s⇩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 = priv_consts_of trms" "vars = vars⇩l⇩s⇩s⇩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: "α ∘⇩s αinv = Var ∧ wf⇩t⇩r⇩m⇩s (subst_range αinv)" using transaction_renaming_subst_inv[of α P vars] by auto define A' where "A' = dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩s α)" define I' where "I' = αinv ∘⇩s I" define trms' where "trms' = trms ∪ trms⇩l⇩s⇩s⇩t A'" define vars' where "vars' = vars ∪ vars⇩l⇩s⇩s⇩t A'" define w' where "w' = w @ [(ξ, σ, I', T, α)]" define IK' where "IK' = IK ∪ (ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I)" define DB' where "DB' = dbupd⇩s⇩s⇩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_priv_consts_of_iff by auto moreover have "transaction_renaming_subst α P vars" using α_p . moreover have "A' = dual⇩l⇩s⇩s⇩t (transaction_strand T ⋅⇩l⇩s⇩s⇩t ξ ∘⇩s σ ∘⇩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⇩s⇩u⇩b⇩s⇩t I'" using step(7) by (simp add: I'_def interpretation_substI subst_compose) moreover have "wf⇩t⇩r⇩m⇩s (subst_range I')" using I'_def αinv_p step.hyps(8) wf_trms_subst_compose by blast moreover have "dbupd⇩s⇩s⇩t (unlabel A) I DB = dbupd⇩s⇩s⇩t (unlabel A') I' DB" proof - have "dbupd⇩s⇩s⇩t (unlabel A) I DB = dbupd⇩s⇩s⇩t (unlabel A) (α ∘⇩s αinv ∘⇩s I) DB" by (simp add: αinv_p) also have "... = dbupd⇩s⇩s⇩t (unlabel A') (αinv ∘⇩s I) DB" unfolding A'_def step(5) by (metis (no_types, lifting) dbupd⇩s⇩s⇩t_substitution_lemma dual⇩l⇩s⇩s⇩t_subst_unlabel subst_compose_assoc) also have "... = dbupd⇩s⇩s⇩t (unlabel A') I' DB" unfolding A'_def I'_def by auto finally show ?thesis by auto qed moreover have "IK ∪ (ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I) = IK ∪ (ik⇩l⇩s⇩s⇩t A' ⋅⇩s⇩e⇩t I')" proof - have "IK ∪ (ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t I) = IK ∪ (ik⇩l⇩s⇩s⇩t A ⋅⇩s⇩e⇩t (α ∘⇩s αinv) ∘⇩s I)" using αinv_p by auto also have "... = IK ∪ (ik⇩l⇩s⇩s⇩t A' ⋅⇩s⇩e⇩t I')" unfolding A'_def step(5) unlabel_subst[symmetric] ik⇩s⇩s⇩t_subst dual⇩l⇩s⇩s⇩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 ∪ priv_consts_of (trms⇩l⇩s⇩s⇩t A) = priv_consts_of trms'" proof - have α_Var: "subst_range α ⊆ range Var" using α_p transaction_renaming_subst_range by blast have ground_ξσ: "ground (subst_range (ξ ∘⇩s σ))" using fresh_transaction_decl_subst_ground_subst_range using step.hyps(3) step.hyps(4) by blast have "consts ∪ priv_consts_of (trms⇩l⇩s⇩s⇩t A) = (priv_consts_of trms) ∪ priv_consts_of (trms⇩l⇩s⇩s⇩t A)" using trms_w_vars_p(2) by blast also have "... = (priv_consts_of trms) ∪ priv_consts_of (trms⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩t α))" using priv_consts_of_trms⇩l⇩s⇩s⇩t_range_Var[of α, OF α_Var, of A] by auto also have "... = (priv_consts_of trms) ∪ priv_consts_of (trms⇩l⇩s⇩s⇩t A')" using step(5) A'_def ground_ξσ using labelled_stateful_strand_ground_subst_comp[of _ "(dual⇩l⇩s⇩s⇩t (transaction_strand T))"] by (simp add: dual⇩l⇩s⇩s⇩t_subst) also have "... = priv_consts_of (trms ∪ trms⇩l⇩s⇩s⇩t A')" using priv_consts_of_union_distr by blast also have "... = priv_consts_of trms'" unfolding trms'_def by auto finally show "?thesis" by auto qed moreover have "vars' = vars⇩l⇩s⇩s⇩t (T_symb w')" using P_state reachable_constraints_if_ground_protocol_states_aux by auto ultimately show ?case unfolding DB'_def IK'_def priv_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⇩l⇩s⇩s⇩t (A ⋅⇩l⇩s⇩s⇩t I), dbupd⇩s⇩s⇩t (unlabel A) I {}) | A I. A ∈ reachable_constraints P ∧ strand_sem_stateful {} {} (unlabel A) I ∧ interpretation⇩s⇩u⇩b⇩s⇩t I ∧ wf⇩t⇩r⇩m⇩s (subst_range I)}" using protocol_model_equivalence_aux2 protocol_model_equivalence_aux1 by auto end end end