diff --git a/thys/Automated_Stateful_Protocol_Verification/Stateful_Protocol_Model.thy b/thys/Automated_Stateful_Protocol_Verification/Stateful_Protocol_Model.thy --- a/thys/Automated_Stateful_Protocol_Verification/Stateful_Protocol_Model.thy +++ b/thys/Automated_Stateful_Protocol_Verification/Stateful_Protocol_Model.thy @@ -1,4410 +1,4410 @@ (* (C) Copyright Andreas Viktor Hess, DTU, 2020 (C) Copyright Sebastian A. Mödersheim, DTU, 2020 (C) Copyright Achim D. Brucker, University of Exeter, 2020 (C) Copyright Anders Schlichtkrull, DTU, 2020 All Rights Reserved. Redistribution and use in source and binary forms, with or without modification, are permitted provided that the following conditions are met: - Redistributions of source code must retain the above copyright notice, this list of conditions and the following disclaimer. - Redistributions in binary form must reproduce the above copyright notice, this list of conditions and the following disclaimer in the documentation and/or other materials provided with the distribution. - Neither the name of the copyright holder nor the names of its contributors may be used to endorse or promote products derived from this software without specific prior written permission. THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. *) (* Title: 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 *) section\Stateful Protocol Model\ theory Stateful_Protocol_Model imports Stateful_Protocol_Composition_and_Typing.Stateful_Compositionality Transactions Term_Abstraction begin subsection \Locale Setup\ locale stateful_protocol_model = fixes arity\<^sub>f::"'fun \ nat" and arity\<^sub>s::"'sets \ nat" and public\<^sub>f::"'fun \ bool" and Ana\<^sub>f::"'fun \ ((('fun,'atom::finite,'sets) prot_fun, nat) term list \ nat list)" and \\<^sub>f::"'fun \ 'atom option" and label_witness1::"'lbl" and label_witness2::"'lbl" assumes Ana\<^sub>f_assm1: "\f. let (K, M) = Ana\<^sub>f f in (\k \ subterms\<^sub>s\<^sub>e\<^sub>t (set K). is_Fun k \ (is_Fu (the_Fun k)) \ length (args k) = arity\<^sub>f (the_Fu (the_Fun k)))" and Ana\<^sub>f_assm2: "\f. let (K, M) = Ana\<^sub>f f in \i \ fv\<^sub>s\<^sub>e\<^sub>t (set K) \ set M. i < arity\<^sub>f f" and public\<^sub>f_assm: "\f. arity\<^sub>f f > (0::nat) \ public\<^sub>f f" and \\<^sub>f_assm: "\f. arity\<^sub>f f = (0::nat) \ \\<^sub>f f \ None" and label_witness_assm: "label_witness1 \ label_witness2" begin lemma Ana\<^sub>f_assm1_alt: assumes "Ana\<^sub>f f = (K,M)" "k \ subterms\<^sub>s\<^sub>e\<^sub>t (set K)" shows "(\x. k = Var x) \ (\h T. k = Fun (Fu h) T \ length T = arity\<^sub>f h)" proof (cases k) case (Fun g T) let ?P = "\k. is_Fun k \ is_Fu (the_Fun k) \ length (args k) = arity\<^sub>f (the_Fu (the_Fun k))" let ?Q = "\K M. \k \ subterms\<^sub>s\<^sub>e\<^sub>t (set K). ?P k" have "?Q (fst (Ana\<^sub>f f)) (snd (Ana\<^sub>f f))" using Ana\<^sub>f_assm1 split_beta[of ?Q "Ana\<^sub>f f"] by meson hence "?Q K M" using assms(1) by simp hence "?P k" using assms(2) by blast thus ?thesis using Fun by (cases g) auto qed simp lemma Ana\<^sub>f_assm2_alt: assumes "Ana\<^sub>f f = (K,M)" "i \ fv\<^sub>s\<^sub>e\<^sub>t (set K) \ set M" shows "i < arity\<^sub>f f" using Ana\<^sub>f_assm2 assms by fastforce subsection \Definitions\ fun arity where "arity (Fu f) = arity\<^sub>f f" | "arity (Set s) = arity\<^sub>s s" | "arity (Val _) = 0" | "arity (Abs _) = 0" | "arity Pair = 2" | "arity (Attack _) = 0" | "arity OccursFact = 2" | "arity OccursSec = 0" | "arity (PubConstAtom _ _) = 0" | "arity (PubConstSetType _) = 0" | "arity (PubConstAttackType _) = 0" | "arity (PubConstBottom _) = 0" | "arity (PubConstOccursSecType _) = 0" fun public where "public (Fu f) = public\<^sub>f f" | "public (Set s) = (arity\<^sub>s s > 0)" | "public (Val n) = snd n" | "public (Abs _) = False" | "public Pair = True" | "public (Attack _) = False" | "public OccursFact = True" | "public OccursSec = False" | "public (PubConstAtom _ _) = True" | "public (PubConstSetType _) = True" | "public (PubConstAttackType _) = True" | "public (PubConstBottom _) = True" | "public (PubConstOccursSecType _) = True" fun Ana where "Ana (Fun (Fu f) T) = ( if arity\<^sub>f f = length T \ arity\<^sub>f f > 0 then let (K,M) = Ana\<^sub>f f in (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M) else ([], []))" | "Ana _ = ([], [])" definition \\<^sub>v where "\\<^sub>v v \ ( if (\t \ subterms (fst v). case t of (TComp f T) \ arity f > 0 \ arity f = length T | _ \ True) then fst v else TAtom Bottom)" fun \ where "\ (Var v) = \\<^sub>v v" | "\ (Fun f T) = ( if arity f = 0 then case f of (Fu g) \ TAtom (case \\<^sub>f g of Some a \ Atom a | None \ Bottom) | (Val _) \ TAtom Value | (Abs _) \ TAtom Value | (Set _) \ TAtom SetType | (Attack _) \ TAtom AttackType | OccursSec \ TAtom OccursSecType | (PubConstAtom a _) \ TAtom (Atom a) | (PubConstSetType _) \ TAtom SetType | (PubConstAttackType _) \ TAtom AttackType | (PubConstBottom _) \ TAtom Bottom | (PubConstOccursSecType _) \ TAtom OccursSecType | _ \ TAtom Bottom else TComp f (map \ T))" lemma \_consts_simps[simp]: "arity\<^sub>f g = 0 \ \ (Fun (Fu g) []) = TAtom (case \\<^sub>f g of Some a \ Atom a | None \ Bottom)" "\ (Fun (Val n) []) = TAtom Value" "\ (Fun (Abs b) []) = TAtom Value" "arity\<^sub>s s = 0 \ \ (Fun (Set s) []) = TAtom SetType" "\ (Fun (Attack x) []) = TAtom AttackType" "\ (Fun OccursSec []) = TAtom OccursSecType" "\ (Fun (PubConstAtom a t) []) = TAtom (Atom a)" "\ (Fun (PubConstSetType t) []) = TAtom SetType" "\ (Fun (PubConstAttackType t) []) = TAtom AttackType" "\ (Fun (PubConstBottom t) []) = TAtom Bottom" "\ (Fun (PubConstOccursSecType t) []) = TAtom OccursSecType" by simp+ lemma \_Set_simps[simp]: "arity\<^sub>s s \ 0 \ \ (Fun (Set s) T) = TComp (Set s) (map \ T)" "\ (Fun (Set s) T) = TAtom SetType \ \ (Fun (Set s) T) = TComp (Set s) (map \ T)" "\ (Fun (Set s) T) \ TAtom Value" "\ (Fun (Set s) T) \ TAtom (Atom a)" "\ (Fun (Set s) T) \ TAtom AttackType" "\ (Fun (Set s) T) \ TAtom OccursSecType" "\ (Fun (Set s) T) \ TAtom Bottom" by auto subsection \Locale Interpretations\ lemma Ana_Fu_cases: assumes "Ana (Fun f T) = (K,M)" and "f = Fu g" and "Ana\<^sub>f g = (K',M')" shows "(K,M) = (if arity\<^sub>f g = length T \ arity\<^sub>f g > 0 then (K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M') else ([],[]))" (is ?A) and "(K,M) = (K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M') \ (K,M) = ([],[])" (is ?B) proof - show ?A using assms by (cases "arity\<^sub>f g = length T \ arity\<^sub>f g > 0") auto thus ?B by metis qed lemma Ana_Fu_intro: assumes "arity\<^sub>f f = length T" "arity\<^sub>f f > 0" and "Ana\<^sub>f f = (K',M')" shows "Ana (Fun (Fu f) T) = (K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M')" using assms by simp lemma Ana_Fu_elim: assumes "Ana (Fun f T) = (K,M)" and "f = Fu g" and "Ana\<^sub>f g = (K',M')" and "(K,M) \ ([],[])" shows "arity\<^sub>f g = length T" (is ?A) and "(K,M) = (K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M')" (is ?B) proof - show ?A using assms by force moreover have "arity\<^sub>f g > 0" using assms by force ultimately show ?B using assms by auto qed lemma Ana_nonempty_inv: assumes "Ana t \ ([],[])" shows "\f T. t = Fun (Fu f) T \ arity\<^sub>f f = length T \ arity\<^sub>f f > 0 \ (\K M. Ana\<^sub>f f = (K, M) \ Ana t = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M))" using assms proof (induction t rule: Ana.induct) case (1 f T) hence *: "arity\<^sub>f f = length T" "0 < arity\<^sub>f f" "Ana (Fun (Fu f) T) = (case Ana\<^sub>f f of (K, M) \ (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M))" using Ana.simps(1)[of f T] unfolding Let_def by metis+ obtain K M where **: "Ana\<^sub>f f = (K, M)" by (metis surj_pair) hence "Ana (Fun (Fu f) T) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M)" using *(3) by simp thus ?case using ** *(1,2) by blast qed simp_all lemma assm1: assumes "Ana t = (K,M)" shows "fv\<^sub>s\<^sub>e\<^sub>t (set K) \ fv t" using assms proof (induction t rule: term.induct) case (Fun f T) have aux: "fv\<^sub>s\<^sub>e\<^sub>t (set K \\<^sub>s\<^sub>e\<^sub>t (!) T) \ fv\<^sub>s\<^sub>e\<^sub>t (set T)" when K: "\i \ fv\<^sub>s\<^sub>e\<^sub>t (set K). i < length T" for K::"(('fun,'atom,'sets) prot_fun, nat) term list" proof fix x assume "x \ fv\<^sub>s\<^sub>e\<^sub>t (set K \\<^sub>s\<^sub>e\<^sub>t (!) T)" then obtain k where k: "k \ set K" "x \ fv (k \ (!) T)" by moura have "\i \ fv k. i < length T" using K k(1) by simp thus "x \ fv\<^sub>s\<^sub>e\<^sub>t (set T)" by (metis (no_types, lifting) k(2) contra_subsetD fv_set_mono image_subsetI nth_mem subst_apply_fv_unfold) qed { fix g assume f: "f = Fu g" and K: "K \ []" obtain K' M' where *: "Ana\<^sub>f g = (K',M')" by moura have "(K, M) \ ([], [])" using K by simp hence "(K, M) = (K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M')" "arity\<^sub>f g = length T" using Ana_Fu_cases(1)[OF Fun.prems f *] by presburger+ hence ?case using aux[of K'] Ana\<^sub>f_assm2_alt[OF *] by auto } thus ?case using Fun by (cases f) fastforce+ qed simp lemma assm2: assumes "Ana t = (K,M)" and "\g S'. Fun g S' \ t \ length S' = arity g" and "k \ set K" and "Fun f T' \ k" shows "length T' = arity f" using assms proof (induction t rule: term.induct) case (Fun g T) obtain h where 2: "g = Fu h" using Fun.prems(1,3) by (cases g) auto obtain K' M' where 1: "Ana\<^sub>f h = (K',M')" by moura have "(K,M) \ ([],[])" using Fun.prems(3) by auto hence "(K,M) = (K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T, map ((!) T) M')" "\i. i \ fv\<^sub>s\<^sub>e\<^sub>t (set K') \ set M' \ i < length T" using Ana_Fu_cases(1)[OF Fun.prems(1) 2 1] Ana\<^sub>f_assm2_alt[OF 1] by presburger+ hence "K = K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T" and 3: "\i\fv\<^sub>s\<^sub>e\<^sub>t (set K'). i < length T" by simp_all then obtain k' where k': "k' \ set K'" "k = k' \ (!) T" using Fun.prems(3) by moura hence 4: "Fun f T' \ subterms (k' \ (!) T)" "fv k' \ fv\<^sub>s\<^sub>e\<^sub>t (set K')" using Fun.prems(4) by auto show ?case proof (cases "\i \ fv k'. Fun f T' \ subterms (T ! i)") case True hence "Fun f T' \ subterms\<^sub>s\<^sub>e\<^sub>t (set T)" using k' Fun.prems(4) 3 by auto thus ?thesis using Fun.prems(2) by auto next case False then obtain S where "Fun f S \ subterms k'" "Fun f T' = Fun f S \ (!) T" using k'(2) Fun.prems(4) subterm_subst_not_img_subterm by force thus ?thesis using Ana\<^sub>f_assm1_alt[OF 1, of "Fun f S"] k'(1) by (cases f) auto qed qed simp lemma assm4: assumes "Ana (Fun f T) = (K, M)" shows "set M \ set T" using assms proof (cases f) case (Fu g) obtain K' M' where *: "Ana\<^sub>f g = (K',M')" by moura have "M = [] \ (arity\<^sub>f g = length T \ M = map ((!) T) M')" using Ana_Fu_cases(1)[OF assms Fu *] by (meson prod.inject) thus ?thesis using Ana\<^sub>f_assm2_alt[OF *] by auto qed auto lemma assm5: "Ana t = (K,M) \ K \ [] \ M \ [] \ Ana (t \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" proof (induction t rule: term.induct) case (Fun f T) thus ?case proof (cases f) case (Fu g) obtain K' M' where *: "Ana\<^sub>f g = (K',M')" by moura have **: "K = K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T" "M = map ((!) T) M'" "arity\<^sub>f g = length T" "\i \ fv\<^sub>s\<^sub>e\<^sub>t (set K') \ set M'. i < arity\<^sub>f g" "0 < arity\<^sub>f g" using Fun.prems(2) Ana_Fu_cases(1)[OF Fun.prems(1) Fu *] Ana\<^sub>f_assm2_alt[OF *] by (meson prod.inject)+ have ***: "\i \ fv\<^sub>s\<^sub>e\<^sub>t (set K'). i < length T" "\i \ set M'. i < length T" using **(3,4) by auto have "K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) (map (\t. t \ \) T)" "M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = map ((!) (map (\t. t \ \) T)) M'" using subst_idx_map[OF ***(2), of \] subst_idx_map'[OF ***(1), of \] **(1,2) by fast+ thus ?thesis using Fu * **(3,5) by auto qed auto qed simp sublocale intruder_model arity public Ana apply unfold_locales by (metis assm1, metis assm2, rule Ana.simps, metis assm4, metis assm5) adhoc_overloading INTRUDER_SYNTH intruder_synth adhoc_overloading INTRUDER_DEDUCT intruder_deduct lemma assm6: "arity c = 0 \ \a. \X. \ (Fun c X) = TAtom a" by (cases c) auto lemma assm7: "0 < arity f \ \ (Fun f T) = TComp f (map \ T)" by auto lemma assm8: "infinite {c. \ (Fun c []::('fun,'atom,'sets) prot_term) = TAtom a \ public c}" (is "?P a") proof - let ?T = "\f. (range f)::('fun,'atom,'sets) 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) prot_fun \ ?T f. \x \ UNIV. y = f x" let ?D = "\f b. ?T f \ {c. \ (Fun c []::('fun,'atom,'sets) 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) 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 "PubConstAtom b" a] by force next case Value thus ?thesis using sub_lmm[of "\n. Val (n,True)" a] by force next case SetType thus ?thesis using sub_lmm[of PubConstSetType a] by fastforce next case AttackType thus ?thesis using sub_lmm[of PubConstAttackType a] by fastforce next case Bottom thus ?thesis using sub_lmm[of PubConstBottom a] by fastforce next case OccursSecType thus ?thesis using sub_lmm[of PubConstOccursSecType a] by fastforce qed qed lemma assm9: "TComp f T \ \ t \ arity f > 0" proof (induction t rule: term.induct) case (Var x) hence "\ (Var x) \ TAtom Bottom" by force hence "\t \ subterms (fst x). case t of TComp f T \ arity f > 0 \ arity f = length T | _ \ True" using Var \.simps(1)[of x] unfolding \\<^sub>v_def by meson thus ?case using Var by (fastforce simp add: \\<^sub>v_def) next case (Fun g S) have "arity g \ 0" using Fun.prems Var_subtermeq assm6 by force thus ?case using Fun by (cases "TComp f T = TComp g (map \ S)") auto qed lemma assm10: "wf\<^sub>t\<^sub>r\<^sub>m (\ (Var x))" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by (auto simp add: \\<^sub>v_def) lemma assm11: "arity f > 0 \ public f" using public\<^sub>f_assm by (cases f) auto lemma assm12: "\ (Var (\, n)) = \ (Var (\, m))" by (simp add: \\<^sub>v_def) lemma assm13: "arity c = 0 \ Ana (Fun c T) = ([],[])" by (cases c) simp_all lemma assm14: assumes "Ana (Fun f T) = (K,M)" shows "Ana (Fun f T \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" proof - show ?thesis proof (cases "(K, M) = ([],[])") case True { fix g assume f: "f = Fu g" obtain K' M' where "Ana\<^sub>f g = (K',M')" by moura hence ?thesis using assms f True by auto } thus ?thesis using True assms by (cases f) auto next case False then obtain g where **: "f = Fu g" using assms by (cases f) auto obtain K' M' where *: "Ana\<^sub>f g = (K',M')" by moura have ***: "K = K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T" "M = map ((!) T) M'" "arity\<^sub>f g = length T" "\i \ fv\<^sub>s\<^sub>e\<^sub>t (set K') \ set M'. i < arity\<^sub>f g" using Ana_Fu_cases(1)[OF assms ** *] False Ana\<^sub>f_assm2_alt[OF *] by (meson prod.inject)+ have ****: "\i\fv\<^sub>s\<^sub>e\<^sub>t (set K'). i < length T" "\i\set M'. i < length T" using ***(3,4) by auto have "K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = K' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) (map (\t. t \ \) T)" "M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = map ((!) (map (\t. t \ \) T)) M'" using subst_idx_map[OF ****(2), of \] subst_idx_map'[OF ****(1), of \] ***(1,2) by auto thus ?thesis using assms * ** ***(3) by auto qed qed sublocale labeled_stateful_typed_model' arity public Ana \ Pair label_witness1 label_witness2 by unfold_locales (metis assm6, metis assm7, metis assm8, metis assm9, rule assm10, metis assm11, rule arity.simps(5), metis assm14, metis assm12, metis assm13, metis assm14, rule label_witness_assm) subsection \Minor Lemmata\ lemma \\<^sub>v_TAtom[simp]: "\\<^sub>v (TAtom a, n) = TAtom a" unfolding \\<^sub>v_def by simp lemma \\<^sub>v_TAtom': assumes "a \ Bottom" shows "\\<^sub>v (\, n) = TAtom a \ \ = TAtom a" proof assume "\\<^sub>v (\, n) = TAtom a" thus "\ = TAtom a" by (metis (no_types, lifting) assms \\<^sub>v_def fst_conv term.inject(1)) qed simp lemma \\<^sub>v_TAtom_inv: "\\<^sub>v x = TAtom (Atom a) \ \m. x = (TAtom (Atom a), m)" "\\<^sub>v x = TAtom Value \ \m. x = (TAtom Value, m)" "\\<^sub>v x = TAtom SetType \ \m. x = (TAtom SetType, m)" "\\<^sub>v x = TAtom AttackType \ \m. x = (TAtom AttackType, m)" "\\<^sub>v x = TAtom OccursSecType \ \m. x = (TAtom OccursSecType, m)" by (metis \\<^sub>v_TAtom' surj_pair prot_atom.distinct(7), metis \\<^sub>v_TAtom' surj_pair prot_atom.distinct(15), metis \\<^sub>v_TAtom' surj_pair prot_atom.distinct(21), metis \\<^sub>v_TAtom' surj_pair prot_atom.distinct(25), metis \\<^sub>v_TAtom' surj_pair prot_atom.distinct(30)) lemma \\<^sub>v_TAtom'': "(fst x = TAtom (Atom a)) = (\\<^sub>v x = TAtom (Atom a))" (is "?A = ?A'") "(fst x = TAtom Value) = (\\<^sub>v x = TAtom Value)" (is "?B = ?B'") "(fst x = TAtom SetType) = (\\<^sub>v x = TAtom SetType)" (is "?C = ?C'") "(fst x = TAtom AttackType) = (\\<^sub>v x = TAtom AttackType)" (is "?D = ?D'") "(fst x = TAtom OccursSecType) = (\\<^sub>v x = TAtom OccursSecType)" (is "?E = ?E'") proof - have 1: "?A \ ?A'" "?B \ ?B'" "?C \ ?C'" "?D \ ?D'" "?E \ ?E'" by (metis \\<^sub>v_TAtom prod.collapse)+ have 2: "?A' \ ?A" "?B' \ ?B" "?C' \ ?C" "?D' \ ?D" "?E' \ ?E" using \\<^sub>v_TAtom \\<^sub>v_TAtom_inv(1) apply fastforce using \\<^sub>v_TAtom \\<^sub>v_TAtom_inv(2) apply fastforce using \\<^sub>v_TAtom \\<^sub>v_TAtom_inv(3) apply fastforce using \\<^sub>v_TAtom \\<^sub>v_TAtom_inv(4) apply fastforce using \\<^sub>v_TAtom \\<^sub>v_TAtom_inv(5) by fastforce show "?A = ?A'" "?B = ?B'" "?C = ?C'" "?D = ?D'" "?E = ?E'" using 1 2 by metis+ qed lemma \\<^sub>v_Var_image: "\\<^sub>v ` X = \ ` Var ` X" by force lemma \_Fu_const: assumes "arity\<^sub>f g = 0" shows "\a. \ (Fun (Fu g) T) = TAtom (Atom a)" proof - have "\\<^sub>f g \ None" using assms \\<^sub>f_assm by blast thus ?thesis using assms by force qed lemma Fun_Value_type_inv: fixes T::"('fun,'atom,'sets) prot_term list" assumes "\ (Fun f T) = TAtom Value" shows "(\n. f = Val n) \ (\bs. f = Abs bs)" proof - have *: "arity f = 0" by (metis const_type_inv assms) show ?thesis using assms proof (cases f) case (Fu g) hence "arity\<^sub>f g = 0" using * by simp hence False using Fu \_Fu_const[of g T] assms by auto thus ?thesis by metis next case (Set s) hence "arity\<^sub>s s = 0" using * by simp hence False using Set assms by auto thus ?thesis by metis qed simp_all qed lemma abs_\: "\ t = \ (t \\<^sub>\ \)" by (induct t \ rule: abs_apply_term.induct) auto lemma Ana\<^sub>f_keys_not_pubval_terms: assumes "Ana\<^sub>f f = (K, T)" and "k \ set K" and "g \ funs_term k" shows "\is_Val g" proof assume "is_Val g" then obtain n S where *: "Fun (Val n) S \ subterms\<^sub>s\<^sub>e\<^sub>t (set K)" using assms(2) funs_term_Fun_subterm[OF assms(3)] by (cases g) auto show False using Ana\<^sub>f_assm1_alt[OF assms(1) *] by simp qed lemma Ana\<^sub>f_keys_not_abs_terms: assumes "Ana\<^sub>f f = (K, T)" and "k \ set K" and "g \ funs_term k" shows "\is_Abs g" proof assume "is_Abs g" then obtain a S where *: "Fun (Abs a) S \ subterms\<^sub>s\<^sub>e\<^sub>t (set K)" using assms(2) funs_term_Fun_subterm[OF assms(3)] by (cases g) auto show False using Ana\<^sub>f_assm1_alt[OF assms(1) *] by simp qed lemma Ana\<^sub>f_keys_not_pairs: assumes "Ana\<^sub>f f = (K, T)" and "k \ set K" and "g \ funs_term k" shows "g \ Pair" proof assume "g = Pair" then obtain S where *: "Fun Pair S \ subterms\<^sub>s\<^sub>e\<^sub>t (set K)" using assms(2) funs_term_Fun_subterm[OF assms(3)] by (cases g) auto show False using Ana\<^sub>f_assm1_alt[OF assms(1) *] by simp qed lemma Ana_Fu_keys_funs_term_subset: fixes K::"('fun,'atom,'sets) prot_term list" assumes "Ana (Fun (Fu f) S) = (K, T)" and "Ana\<^sub>f f = (K', T')" shows "\(funs_term ` set K) \ \(funs_term ` set K') \ funs_term (Fun (Fu f) S)" proof - { fix k assume k: "k \ set K" then obtain k' where k': "k' \ set K'" "k = k' \ (!) S" "arity\<^sub>f f = length S" "subterms k' \ subterms\<^sub>s\<^sub>e\<^sub>t (set K')" using assms Ana_Fu_elim[OF assms(1) _ assms(2)] by fastforce have 1: "funs_term k' \ \(funs_term ` set K')" using k'(1) by auto have "i < length S" when "i \ fv k'" for i using that Ana\<^sub>f_assm2_alt[OF assms(2), of i] k'(1,3) by auto hence 2: "funs_term (S ! i) \ funs_term (Fun (Fu f) S)" when "i \ fv k'" for i using that by force have "funs_term k \ \(funs_term ` set K') \ funs_term (Fun (Fu f) S)" using funs_term_subst[of k' "(!) S"] k'(2) 1 2 by fast } thus ?thesis by blast qed lemma Ana_Fu_keys_not_pubval_terms: fixes k::"('fun,'atom,'sets) prot_term" assumes "Ana (Fun (Fu f) S) = (K, T)" and "Ana\<^sub>f f = (K', T')" and "k \ set K" and "\g \ funs_term (Fun (Fu f) S). is_Val g \ \public g" shows "\g \ funs_term k. is_Val g \ \public g" using assms(3,4) Ana\<^sub>f_keys_not_pubval_terms[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) prot_term" assumes "Ana (Fun (Fu f) S) = (K, T)" and "Ana\<^sub>f f = (K', T')" and "k \ set K" and "\g \ funs_term (Fun (Fu f) S). \is_Abs g" shows "\g \ funs_term k. \is_Abs g" using assms(3,4) Ana\<^sub>f_keys_not_abs_terms[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) prot_term" assumes "Ana (Fun (Fu f) S) = (K, T)" and "Ana\<^sub>f f = (K', T')" and "k \ set K" and "\g \ funs_term (Fun (Fu f) S). g \ Pair" shows "\g \ funs_term k. g \ Pair" using assms(3,4) Ana\<^sub>f_keys_not_pairs[OF assms(2)] Ana_Fu_keys_funs_term_subset[OF assms(1,2)] by blast lemma deduct_occurs_in_ik: fixes t::"('fun,'atom,'sets) prot_term" assumes t: "M \ occurs t" and M: "\s \ subterms\<^sub>s\<^sub>e\<^sub>t M. OccursFact \ \(funs_term ` set (snd (Ana s)))" "\s \ subterms\<^sub>s\<^sub>e\<^sub>t M. OccursSec \ \(funs_term ` set (snd (Ana s)))" "Fun OccursSec [] \ M" shows "occurs t \ M" using private_fun_deduct_in_ik''[of M OccursFact "[Fun OccursSec [], t]" OccursSec] t M by fastforce lemma wellformed_transaction_sem_receives: fixes T::"('fun,'atom,'sets,'lbl) prot_transaction" assumes T_valid: "wellformed_transaction T" and \: "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \" and s: "receive\t\ \ set (unlabel (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" shows "IK \ t \ \" proof - let ?R = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?S = "\A. unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?S' = "?S (transaction_receive T)" obtain l B s where B: "(l,send\t\) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "prefix ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]) (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using s dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(2)[of t "transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_set_prefix_obtain_subst[of "send\t\" "transaction_receive T" \] by blast have 1: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \])) = unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))@[send\t\]" using B(1) unlabel_append dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst singleton_lst_proj(4) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_snoc subst_lsst_append subst_lsst_singleton by (metis (no_types, lifting) subst_apply_labeled_stateful_strand_step.simps ) have "strand_sem_stateful IK DB ?S' \" using \ strand_sem_append_stateful[of IK DB _ _ \] transaction_dual_subst_unfold[of T \] by fastforce hence "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))@[send\t\]) \" using B 1 unfolding prefix_def unlabel_def by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def map_append strand_sem_append_stateful) hence t_deduct: "IK \ (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \\<^sub>s\<^sub>e\<^sub>t \) \ t \ \" using strand_sem_append_stateful[of IK DB "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "[send\t\]" \] by simp have "\s \ set (unlabel (transaction_receive T)). \t. s = receive\t\" using T_valid wellformed_transaction_unlabel_cases(1)[OF T_valid] by auto moreover { fix A::"('fun,'atom,'sets,'lbl) prot_strand" and \ assume "\s \ set (unlabel A). \t. s = receive\t\" hence "\s \ set (unlabel (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)). \t. s = receive\t\" proof (induction A) case (Cons a A) thus ?case using subst_lsst_cons[of a A \] by (cases a) auto qed simp hence "\s \ set (unlabel (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)). \t. s = receive\t\" by (simp add: list.pred_set is_Receive_def) hence "\s \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))). \t. s = send\t\" by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_memberD dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_inv(2) unlabel_in unlabel_mem_has_label) } ultimately have "\s \ set ?R. \t. s = send\t\" by simp hence "ik\<^sub>s\<^sub>s\<^sub>t ?R = {}" unfolding unlabel_def ik\<^sub>s\<^sub>s\<^sub>t_def by fast hence "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = {}" using B(2) 1 ik\<^sub>s\<^sub>s\<^sub>t_append dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append by (metis (no_types, lifting) Un_empty map_append prefix_def unlabel_def) thus ?thesis using t_deduct by simp qed lemma wellformed_transaction_sem_selects: assumes T_valid: "wellformed_transaction T" and \: "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \" and "select\t,u\ \ set (unlabel (transaction_selects T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" shows "(t \ \, u \ \) \ DB" proof - let ?s = "select\t,u\" let ?R = "transaction_receive T@transaction_selects T" let ?R' = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?S = "\A. unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?S' = "?S (transaction_receive T)@?S (transaction_selects T)" let ?P = "\a. is_Receive a \ is_Assignment a" let ?Q = "\a. is_Send a \ is_Assignment a" have s: "?s \ set (unlabel (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using assms(3) subst_lsst_append[of "transaction_receive T"] unlabel_append[of "transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] by auto obtain l B s where B: "(l,?s) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "prefix ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]) (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using s dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(6)[of assign t u] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_set_prefix_obtain_subst[of ?s ?R \] by blast have 1: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \])) = unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))@[?s]" using B(1) unlabel_append dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst singleton_lst_proj(4) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_snoc subst_lsst_append subst_lsst_singleton by (metis (no_types, lifting) subst_apply_labeled_stateful_strand_step.simps) have "strand_sem_stateful IK DB ?S' \" using \ strand_sem_append_stateful[of IK DB _ _ \] transaction_dual_subst_unfold[of T \] by fastforce hence "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))@[?s]) \" using B 1 strand_sem_append_stateful subst_lsst_append unfolding prefix_def unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by (metis (no_types) map_append) hence in_db: "(t \ \, u \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \ DB" using strand_sem_append_stateful[of IK DB "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "[?s]" \] by simp have "\a \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))). ?Q a" proof fix a assume a: "a \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" have "\a \ set (unlabel ?R). ?P a" using wellformed_transaction_unlabel_cases(1)[OF T_valid] wellformed_transaction_unlabel_cases(2)[OF T_valid] unfolding unlabel_def by fastforce hence "\a \ set (unlabel (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)). ?P a" using stateful_strand_step_cases_subst(2,8)[of _ \] subst_lsst_unlabel[of ?R \] by (simp add: subst_apply_stateful_strand_def del: unlabel_append) hence B_P: "\a \ set (unlabel (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)). ?P a" using unlabel_mono[OF set_mono_prefix[OF append_prefixD[OF B(2)]]] by blast obtain l where "(l,a) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using a by (meson unlabel_mem_has_label) then obtain b where b: "(l,b) \ set (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (l,b) = (l,a)" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_memberD by blast hence "?P b" using B_P unfolding unlabel_def by fastforce thus "?Q a" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_inv[OF b(2)] by (cases b) auto qed hence "\a \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))). \is_Insert a \ \is_Delete a" by fastforce thus ?thesis using dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd[of "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" \ DB] in_db by simp qed lemma wellformed_transaction_sem_pos_checks: assumes T_valid: "wellformed_transaction T" and \: "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \" and "\t in u\ \ set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" shows "(t \ \, u \ \) \ DB" proof - let ?s = "\t in u\" let ?R = "transaction_receive T@transaction_selects T@transaction_checks T" let ?R' = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?S = "\A. unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?S' = "?S (transaction_receive T)@?S (transaction_selects T)@?S (transaction_checks T)" let ?P = "\a. is_Receive a \ is_Assignment a \ is_Check a" let ?Q = "\a. is_Send a \ is_Assignment a \ is_Check a" have s: "?s \ set (unlabel (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using assms(3) subst_lsst_append[of "transaction_receive T@transaction_selects T"] unlabel_append[of "transaction_receive T@transaction_selects T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] by auto obtain l B s where B: "(l,?s) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "prefix ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]) (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using s dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(6)[of check t u] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_set_prefix_obtain_subst[of ?s ?R \] by blast have 1: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \])) = unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))@[?s]" using B(1) unlabel_append dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst singleton_lst_proj(4) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_snoc subst_lsst_append subst_lsst_singleton by (metis (no_types, lifting) subst_apply_labeled_stateful_strand_step.simps ) have "strand_sem_stateful IK DB ?S' \" using \ strand_sem_append_stateful[of IK DB _ _ \] transaction_dual_subst_unfold[of T \] by fastforce hence "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))@[?s]) \" using B 1 strand_sem_append_stateful subst_lsst_append unfolding prefix_def unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by (metis (no_types) map_append) hence in_db: "(t \ \, u \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \ DB" using strand_sem_append_stateful[of IK DB "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "[?s]" \] by simp have "\a \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))). ?Q a" proof fix a assume a: "a \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" have "\a \ set (unlabel ?R). ?P a" using wellformed_transaction_unlabel_cases(1,2,3)[OF T_valid] unfolding unlabel_def by fastforce hence "\a \ set (unlabel (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)). ?P a" using stateful_strand_step_cases_subst(2,8,9)[of _ \] subst_lsst_unlabel[of ?R \] by (simp add: subst_apply_stateful_strand_def del: unlabel_append) hence B_P: "\a \ set (unlabel (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)). ?P a" using unlabel_mono[OF set_mono_prefix[OF append_prefixD[OF B(2)]]] by blast obtain l where "(l,a) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using a by (meson unlabel_mem_has_label) then obtain b where b: "(l,b) \ set (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (l,b) = (l,a)" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_memberD by blast hence "?P b" using B_P unfolding unlabel_def by fastforce thus "?Q a" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_inv[OF b(2)] by (cases b) auto qed hence "\a \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))). \is_Insert a \ \is_Delete a" by fastforce thus ?thesis using dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd[of "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" \ DB] in_db by simp qed lemma wellformed_transaction_sem_neg_checks: assumes T_valid: "wellformed_transaction T" and \: "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \" and "NegChecks X [] [(t,u)] \ set (unlabel (transaction_checks T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" shows "\\. subst_domain \ = set X \ ground (subst_range \) \ (t \ \ \ \, u \ \ \ \) \ DB" (is ?A) and "X = [] \ (t \ \, u \ \) \ DB" (is "?B \ ?B'") proof - let ?s = "NegChecks X [] [(t,u)]" let ?R = "transaction_receive T@transaction_selects T@transaction_checks T" let ?R' = "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?S = "\A. unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" let ?S' = "?S (transaction_receive T)@?S (transaction_selects T)@?S (transaction_checks T)" let ?P = "\a. is_Receive a \ is_Assignment a \ is_Check a" let ?Q = "\a. is_Send a \ is_Assignment a \ is_Check a" let ?U = "\\. subst_domain \ = set X \ ground (subst_range \)" have s: "?s \ set (unlabel (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using assms(3) subst_lsst_append[of "transaction_receive T@transaction_selects T"] unlabel_append[of "transaction_receive T@transaction_selects T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] by auto obtain l B s where B: "(l,?s) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ((l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "prefix ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \]) (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using s dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(7)[of X "[]" "[(t,u)]"] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_set_prefix_obtain_subst[of ?s ?R \] by blast have 1: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@[(l,s) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \])) = unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))@[?s]" using B(1) unlabel_append dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst singleton_lst_proj(4) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst_snoc subst_lsst_append subst_lsst_singleton by (metis (no_types, lifting) subst_apply_labeled_stateful_strand_step.simps) have "strand_sem_stateful IK DB ?S' \" using \ strand_sem_append_stateful[of IK DB _ _ \] transaction_dual_subst_unfold[of T \] by fastforce hence "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))@[?s]) \" using B 1 strand_sem_append_stateful subst_lsst_append unfolding prefix_def unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by (metis (no_types) map_append) hence "negchecks_model \ (dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \ DB) X [] [(t,u)]" using strand_sem_append_stateful[of IK DB "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" "[?s]" \] by fastforce hence in_db: "\\. ?U \ \ (t \ \ \ \, u \ \ \ \) \ dbupd\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \ DB" unfolding negchecks_model_def by simp have "\a \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))). ?Q a" proof fix a assume a: "a \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" have "\a \ set (unlabel ?R). ?P a" using wellformed_transaction_unlabel_cases(1,2,3)[OF T_valid] unfolding unlabel_def by fastforce hence "\a \ set (unlabel (?R \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)). ?P a" using stateful_strand_step_cases_subst(2,8,9)[of _ \] subst_lsst_unlabel[of ?R \] by (simp add: subst_apply_stateful_strand_def del: unlabel_append) hence B_P: "\a \ set (unlabel (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)). ?P a" using unlabel_mono[OF set_mono_prefix[OF append_prefixD[OF B(2)]]] by blast obtain l where "(l,a) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using a by (meson unlabel_mem_has_label) then obtain b where b: "(l,b) \ set (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (l,b) = (l,a)" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_memberD by blast hence "?P b" using B_P unfolding unlabel_def by fastforce thus "?Q a" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_inv[OF b(2)] by (cases b) auto qed hence "\a \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))). \is_Insert a \ \is_Delete a" by fastforce thus ?A using dbupd\<^sub>s\<^sub>s\<^sub>t_no_upd[of "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (B \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" \ DB] in_db by simp moreover have "\ = Var" "t \ \ = t" when "subst_domain \ = set []" for t and \::"('fun, 'atom, 'sets) 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_fv_in_receives_or_selects: assumes T: "wellformed_transaction T" and x: "x \ fv_transaction T" "x \ set (transaction_fresh T)" shows "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T)" proof - have "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_updates T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" using x(1) fv\<^sub>s\<^sub>s\<^sub>t_append unlabel_append by (metis transaction_strand_def append_assoc) thus ?thesis using T x(2) unfolding wellformed_transaction_def by blast qed lemma dual_transaction_ik_is_transaction_send'': fixes \ \::"('a,'b,'c) prot_subst" assumes "wellformed_transaction T" shows "(ik\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a = (trms\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>\\<^sub>s\<^sub>e\<^sub>t a" (is "?A = ?B") using dual_transaction_ik_is_transaction_send[OF assms] subst_lsst_unlabel[of "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)" \] ik\<^sub>s\<^sub>s\<^sub>t_subst[of "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))" \] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst[of "transaction_strand T" \] by (auto simp add: abs_apply_terms_def) lemma while_prot_terms_fun_mono: "mono (\M'. M \ \(subterms ` M') \ \((set \ fst \ Ana) ` M'))" unfolding mono_def by fast lemma while_prot_terms_SMP_overapprox: fixes M::"('fun,'atom,'sets) prot_terms" assumes N_supset: "M \ \(subterms ` N) \ \((set \ fst \ Ana) ` N) \ N" and Value_vars_only: "\x \ fv\<^sub>s\<^sub>e\<^sub>t N. \\<^sub>v x = TAtom Value" shows "SMP M \ {a \ \ | a \. a \ N \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)}" proof - define f where "f \ \M'. M \ \(subterms ` M') \ \((set \ fst \ Ana) ` M')" define S where "S \ {a \ \ | a \. a \ N \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)}" note 0 = Value_vars_only have "t \ S" when "t \ SMP M" for t using that proof (induction t rule: SMP.induct) case (MP t) hence "t \ N" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range Var)" using N_supset by auto hence "t \ Var \ S" unfolding S_def by blast thus ?case by simp next case (Subterm t t') then obtain \ a where a: "a \ \ = t" "a \ N" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (auto simp add: S_def) hence "\x \ fv a. \\. \ (Var x) = TAtom \" using 0 by auto hence *: "\x \ fv a. (\f. \ x = Fun f []) \ (\y. \ x = Var y)" using a(3) TAtom_term_cases[OF wf_trm_subst_rangeD[OF a(4)]] by (metis wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) obtain b where b: "b \ \ = t'" "b \ subterms a" using subterms_subst_subterm[OF *, of t'] Subterm.hyps(2) a(1) by fast hence "b \ N" using N_supset a(2) by blast thus ?case using a b(1) unfolding S_def by blast next case (Substitution t \) then obtain \ a where a: "a \ \ = t" "a \ N" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (auto simp add: S_def) have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" by (fact wt_subst_compose[OF a(3) Substitution.hyps(2)], fact wf_trms_subst_compose[OF a(4) Substitution.hyps(3)]) moreover have "t \ \ = a \ \ \\<^sub>s \" using a(1) subst_subst_compose[of a \ \] by simp ultimately show ?case using a(2) unfolding S_def by blast next case (Ana t K T k) then obtain \ a where a: "a \ \ = t" "a \ N" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (auto simp add: S_def) obtain Ka Ta where a': "Ana a = (Ka,Ta)" by moura have *: "K = Ka \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" proof (cases a) case (Var x) then obtain g U where gU: "t = Fun g U" using a(1) Ana.hyps(2,3) Ana_var by (cases t) simp_all have "\ (Var x) = TAtom Value" using Var a(2) 0 by auto hence "\ (Fun g U) = TAtom Value" using a(1,3) Var gU wt_subst_trm''[OF a(3), of a] by argo thus ?thesis using gU Fun_Value_type_inv Ana.hyps(2,3) by fastforce next case (Fun g U) thus ?thesis using a(1) a' Ana.hyps(2) Ana_subst'[of g U] by simp qed then obtain ka where ka: "k = ka \ \" "ka \ set Ka" using Ana.hyps(3) by auto have "ka \ set ((fst \ Ana) a)" using ka(2) a' by simp hence "ka \ N" using a(2) N_supset by auto thus ?case using ka a(3,4) unfolding S_def by blast qed thus ?thesis unfolding S_def by blast qed subsection \The Protocol Transition System, Defined in Terms of the Reachable Constraints\ definition transaction_fresh_subst where "transaction_fresh_subst \ T \ \ subst_domain \ = set (transaction_fresh T) \ (\t \ subst_range \. \n. t = Fun (Val (n,False)) []) \ (\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \ (\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction 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 \ \ \n \ max_var_set (\(vars_transaction ` set P) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \). \ = var_rename n" definition constraint_model where "constraint_model \ \ \ constr_sem_stateful \ (unlabel \) \ interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" definition welltyped_constraint_model where "welltyped_constraint_model \ \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ constraint_model \ \" 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 constraint_model_Val_is_Value_term: assumes "welltyped_constraint_model I A" and "t \ I = Fun (Val n) []" shows "t = Fun (Val n) [] \ (\m. t = Var (TAtom Value, m))" proof - have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" using assms(1) unfolding welltyped_constraint_model_def by simp moreover have "\ (Fun (Val n) []) = TAtom Value" by auto ultimately have *: "\ t = TAtom Value" by (metis (no_types) assms(2) wt_subst_trm'') show ?thesis proof (cases t) case (Var x) obtain \ m where x: "x = (\, m)" by (metis surj_pair) have "\\<^sub>v x = TAtom Value" using * Var by auto hence "\ = TAtom Value" using x \\<^sub>v_TAtom'[of Value \ m] by simp thus ?thesis using x Var by metis next case (Fun f T) thus ?thesis using assms(2) by auto qed qed text \ The set of symbolic constraints reachable in any symbolic run of the protocol \P\. \\\ 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: "[] \ reachable_constraints P" | step: "\\ \ reachable_constraints P; T \ set P; transaction_fresh_subst \ T \; transaction_renaming_subst \ P \ \ \ \@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \) \ reachable_constraints P" subsection \Admissible Transactions\ definition admissible_transaction_checks where "admissible_transaction_checks T \ \x \ set (unlabel (transaction_checks T)). is_Check x \ (is_InSet x \ is_Var (the_elem_term x) \ is_Fun_Set (the_set_term x) \ fst (the_Var (the_elem_term x)) = TAtom Value) \ (is_NegChecks x \ bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p x = [] \ ((the_eqs x = [] \ length (the_ins x) = 1) \ (the_ins x = [] \ length (the_eqs x) = 1))) \ (is_NegChecks x \ the_eqs x = [] \ (let h = hd (the_ins x) in is_Var (fst h) \ is_Fun_Set (snd h) \ fst (the_Var (fst h)) = TAtom Value))" definition admissible_transaction_selects where "admissible_transaction_selects T \ \x \ set (unlabel (transaction_selects T)). is_InSet x \ the_check x = Assign \ 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_updates where "admissible_transaction_updates T \ \x \ set (unlabel (transaction_updates T)). is_Update x \ is_Var (the_elem_term x) \ is_Fun_Set (the_set_term x) \ fst (the_Var (the_elem_term x)) = TAtom Value" definition admissible_transaction_terms where "admissible_transaction_terms T \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)) \ (\f \ \(funs_term ` trms_transaction T). \is_Val f \ \is_Abs f \ \is_PubConstSetType f \ f \ Pair \ \is_PubConstAttackType f \ \is_PubConstBottom f \ \is_PubConstOccursSecType f) \ (\r \ set (unlabel (transaction_strand T)). (\f \ \(funs_term ` (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r)). is_Attack f) \ (let t = the_msg r in is_Send r \ is_Fun t \ is_Attack (the_Fun t) \ args t = []))" definition admissible_transaction_occurs_checks where "admissible_transaction_occurs_checks T \ ( (\x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value \ receive\occurs (Var x)\ \ set (unlabel (transaction_receive T))) \ (\x \ set (transaction_fresh T). fst x = TAtom Value \ send\occurs (Var x)\ \ set (unlabel (transaction_send T))) \ (\r \ set (unlabel (transaction_receive T)). is_Receive r \ (OccursFact \ funs_term (the_msg r) \ OccursSec \ funs_term (the_msg r)) \ (\x \ fv_transaction T - set (transaction_fresh T). fst x = TAtom Value \ the_msg r = occurs (Var x))) \ (\r \ set (unlabel (transaction_send T)). is_Send r \ (OccursFact \ funs_term (the_msg r) \ OccursSec \ funs_term (the_msg r)) \ (\x \ set (transaction_fresh T). fst x = TAtom Value \ the_msg r = occurs (Var x))) )" definition admissible_transaction where "admissible_transaction T \ ( wellformed_transaction T \ distinct (transaction_fresh T) \ list_all (\x. fst x = TAtom Value) (transaction_fresh T) \ (\x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T). is_Var (fst x) \ (the_Var (fst x) = Value)) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand 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))) \ admissible_transaction_selects T \ admissible_transaction_checks T \ admissible_transaction_updates T \ admissible_transaction_terms T \ admissible_transaction_occurs_checks T )" lemma transaction_no_bvars: assumes "admissible_transaction T" shows "fv_transaction T = vars_transaction T" and "bvars_transaction T = {}" proof - have "wellformed_transaction T" "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) = {}" using assms unfolding admissible_transaction_def by blast+ thus "bvars_transaction T = {}" "fv_transaction T = vars_transaction T" using bvars_wellformed_transaction_unfold vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t by fast+ qed lemma 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 transaction_no_bvars(2) by fast lemma transaction_bvars_no_Value_type: assumes "admissible_transaction T" and "x \ bvars_transaction T" shows "\TAtom Value \ \\<^sub>v x" using assms transaction_no_bvars(2) by blast lemma transaction_receive_deduct: assumes T_adm: "admissible_transaction T" and \: "constraint_model \ (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \))" and \: "transaction_fresh_subst \ T A" and \: "transaction_renaming_subst \ P A" and t: "receive\t\ \ set (unlabel (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \))" shows "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" proof - define \ where "\ \ \ \\<^sub>s \" have t': "send\t\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" using t dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(2) unfolding \_def by blast then obtain T1 T2 where T: "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = T1@send\t\#T2" using t' by (meson split_list) have "constr_sem_stateful \ (unlabel A@unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" using \ unlabel_append[of A] unfolding constraint_model_def \_def by simp hence "constr_sem_stateful \ (unlabel A@T1@[send\t\])" using strand_sem_append_stateful[of "{}" "{}" "unlabel A@T1@[send\t\]" _ \] transaction_dual_subst_unfold[of T \] T by (metis append.assoc append_Cons append_Nil) hence "ik\<^sub>s\<^sub>s\<^sub>t (unlabel A@T1) \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" using strand_sem_append_stateful[of "{}" "{}" "unlabel A@T1" "[send\t\]" \] T by force moreover have "\is_Receive x" when x: "x \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" for x proof - have *: "is_Receive a" when "a \ set (unlabel (transaction_receive T))" for a using T_adm Ball_set[of "unlabel (transaction_receive T)" is_Receive] that unfolding admissible_transaction_def wellformed_transaction_def by blast obtain l where l: "(l,x) \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using x unfolding unlabel_def by fastforce then obtain ly where ly: "ly \ set (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "(l,x) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p ly" unfolding dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto obtain j y where j: "ly = (j,y)" by (metis surj_pair) hence "j = l" using ly(2) by (cases y) auto hence y: "(l,y) \ set (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "(l,x) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (l,y)" by (metis j ly(1), metis j ly(2)) obtain z where z: "z \ set (unlabel (transaction_receive T))" "(l,z) \ set (transaction_receive T)" "(l,y) = (l,z) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" using y(1) unfolding subst_apply_labeled_stateful_strand_def unlabel_def by force have "is_Receive y" using *[OF z(1)] z(3) by (cases z) auto thus "\is_Receive x" using l y by (cases y) auto qed hence "\is_Receive x" when "x \ set T1" for x using T that by simp hence "ik\<^sub>s\<^sub>s\<^sub>t T1 = {}" unfolding ik\<^sub>s\<^sub>s\<^sub>t_def is_Receive_def by fast hence "ik\<^sub>s\<^sub>s\<^sub>t (unlabel A@T1) = ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using ik\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel A" T1] by simp ultimately show ?thesis by (simp add: \_def) qed lemma transaction_checks_db: assumes T: "admissible_transaction T" and \: "constraint_model \ (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \))" and \: "transaction_fresh_subst \ T A" and \: "transaction_renaming_subst \ P A" shows "\Var (TAtom Value, n) in Fun (Set s) []\ \ set (unlabel (transaction_checks T)) \ (\ (TAtom Value, n) \ \, Fun (Set s) []) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \)" (is "?A \ ?B") and "\Var (TAtom Value, n) not in Fun (Set s) []\ \ set (unlabel (transaction_checks T)) \ (\ (TAtom Value, n) \ \, Fun (Set s) []) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \)" (is "?C \ ?D") proof - let ?x = "\n. (TAtom Value, n)" let ?s = "Fun (Set s) []" let ?T = "transaction_receive T@transaction_selects T@transaction_checks T" let ?T' = "?T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \" let ?S = "\S. transaction_receive T@transaction_selects T@S" let ?S' = "\S. ?S S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \" have T_valid: "wellformed_transaction T" using T by (simp add: admissible_transaction_def) have "constr_sem_stateful \ (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)))" using \ unfolding constraint_model_def by simp moreover have "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S (T1@[c]) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T2@transaction_updates T@transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" when "transaction_checks T = T1@c#T2" for T1 T2 c \ using that dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append subst_lsst_append unfolding transaction_strand_def by (metis append.assoc append_Cons append_Nil) ultimately have T'_model: "constr_sem_stateful \ (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' (T1@[(l,c)]))))" when "transaction_checks T = T1@(l,c)#T2" for T1 T2 l c using strand_sem_append_stateful[of _ _ _ _ \] by (simp add: that transaction_strand_def) show "?A \ ?B" proof - assume a: ?A hence *: "\Var (?x n) in ?s\ \ set (unlabel ?T)" unfolding transaction_strand_def unlabel_def by simp then obtain l T1 T2 where T1: "transaction_checks T = T1@(l,\Var (?x n) in ?s\)#T2" by (metis a split_list unlabel_mem_has_label) have "?x n \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using a by force hence "?x n \ set (transaction_fresh T)" using a transaction_fresh_vars_notin[OF T_valid] by fast hence "unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' (T1@[(l,\Var (?x n) in ?s\)]))) = unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' T1))@[\\ (?x n) in ?s\]" using T a \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append subst_lsst_append unlabel_append by (fastforce simp add: transaction_fresh_subst_def unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def subst_apply_labeled_stateful_strand_def) moreover have "db\<^sub>s\<^sub>s\<^sub>t (unlabel A) = db\<^sub>s\<^sub>s\<^sub>t (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' T1)))" by (simp add: T1 db\<^sub>s\<^sub>s\<^sub>t_transaction_prefix_eq[OF T_valid] del: unlabel_append) ultimately have "\M. strand_sem_stateful M (set (db\<^sub>s\<^sub>s\<^sub>t (unlabel A) \)) [\\ (?x n) in ?s\] \" using T'_model[OF T1] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of _ \] strand_sem_append_stateful[of _ _ _ _ \] by (simp add: db\<^sub>s\<^sub>s\<^sub>t_def del: unlabel_append) thus ?B by simp qed show "?C \ ?D" proof - assume a: ?C hence *: "\Var (?x n) not in ?s\ \ set (unlabel ?T)" unfolding transaction_strand_def unlabel_def by simp then obtain l T1 T2 where T1: "transaction_checks T = T1@(l,\Var (?x n) not in ?s\)#T2" by (metis a split_list unlabel_mem_has_label) have "?x n \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p \Var (?x n) not in ?s\" using vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_cases(9)[of "[]" "Var (?x n)" ?s] by auto hence "?x n \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using a unfolding vars\<^sub>s\<^sub>s\<^sub>t_def by force hence "?x n \ set (transaction_fresh T)" using a transaction_fresh_vars_notin[OF T_valid] by fast hence "unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' (T1@[(l,\Var (?x n) not in ?s\)]))) = unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' T1))@[\\ (?x n) not in ?s\]" using T a \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append subst_lsst_append unlabel_append by (fastforce simp add: transaction_fresh_subst_def unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def subst_apply_labeled_stateful_strand_def) moreover have "db\<^sub>s\<^sub>s\<^sub>t (unlabel A) = db\<^sub>s\<^sub>s\<^sub>t (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' T1)))" by (simp add: T1 db\<^sub>s\<^sub>s\<^sub>t_transaction_prefix_eq[OF T_valid] del: unlabel_append) ultimately have "\M. strand_sem_stateful M (set (db\<^sub>s\<^sub>s\<^sub>t (unlabel A) \)) [\\ (?x n) not in ?s\] \" using T'_model[OF T1] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of _ \] strand_sem_append_stateful[of _ _ _ _ \] by (simp add: db\<^sub>s\<^sub>s\<^sub>t_def del: unlabel_append) thus ?D using stateful_strand_sem_NegChecks_no_bvars(1)[of _ _ _ ?s \] by simp qed qed lemma transaction_selects_db: assumes T: "admissible_transaction T" and \: "constraint_model \ (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \))" and \: "transaction_fresh_subst \ T A" and \: "transaction_renaming_subst \ P A" shows "select\Var (TAtom Value, n), Fun (Set s) []\ \ set (unlabel (transaction_selects T)) \ (\ (TAtom Value, n) \ \, Fun (Set s) []) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \)" (is "?A \ ?B") proof - let ?x = "\n. (TAtom Value, n)" let ?s = "Fun (Set s) []" let ?T = "transaction_receive T@transaction_selects T@transaction_checks T" let ?T' = "?T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \" let ?S = "\S. transaction_receive T@S" let ?S' = "\S. ?S S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \" have T_valid: "wellformed_transaction T" using T by (simp add: admissible_transaction_def) have "constr_sem_stateful \ (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)))" using \ unfolding constraint_model_def by simp moreover have "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S (T1@[c]) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)@ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (T2@transaction_checks T @ transaction_updates T@transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" when "transaction_selects T = T1@c#T2" for T1 T2 c \ using that dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append subst_lsst_append unfolding transaction_strand_def by (metis append.assoc append_Cons append_Nil) ultimately have T'_model: "constr_sem_stateful \ (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' (T1@[(l,c)]))))" when "transaction_selects 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_selects T = T1@(l,select\Var (?x n), ?s\)#T2" by (metis a split_list unlabel_mem_has_label) have "?x n \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T)" using a by force hence "?x n \ set (transaction_fresh T)" using a transaction_fresh_vars_notin[OF T_valid] by fast hence "unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' (T1@[(l,select\Var (?x n), ?s\)]))) = unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' T1))@[select\\ (?x n), ?s\]" using T a \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_append subst_lsst_append unlabel_append by (fastforce simp add: transaction_fresh_subst_def unlabel_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def subst_apply_labeled_stateful_strand_def) moreover have "db\<^sub>s\<^sub>s\<^sub>t (unlabel A) = db\<^sub>s\<^sub>s\<^sub>t (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (?S' T1)))" by (simp add: T1 db\<^sub>s\<^sub>s\<^sub>t_transaction_prefix_eq[OF T_valid] del: unlabel_append) ultimately have "\M. strand_sem_stateful M (set (db\<^sub>s\<^sub>s\<^sub>t (unlabel A) \)) [\\ (?x n) in ?s\] \" using T'_model[OF T1] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of _ \] strand_sem_append_stateful[of _ _ _ _ \] by (simp add: db\<^sub>s\<^sub>s\<^sub>t_def del: unlabel_append) thus ?B by simp qed qed lemma transactions_have_no_Value_consts: assumes "admissible_transaction T" and "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))" shows "\a T. t = Fun (Val a) T" (is ?A) and "\a T. t = Fun (Abs a) T" (is ?B) proof - have "admissible_transaction_terms T" using assms(1) unfolding admissible_transaction_def by blast hence "\is_Val f" "\is_Abs f" when "f \ \(funs_term ` (trms_transaction T))" for f using that unfolding 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_Val f" "\is_Abs f" when "f \ funs_term t" for f using that by presburger+ show ?A using *(1) by force show ?B using *(2) by force qed lemma transactions_have_no_Value_consts': assumes "admissible_transaction T" and "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)" shows "\a T. Fun (Val a) T \ subterms t" and "\a T. Fun (Abs a) T \ subterms t" using transactions_have_no_Value_consts[OF assms(1)] assms(2) by fast+ lemma transactions_have_no_PubConsts: assumes "admissible_transaction T" and "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))" shows "\a T. t = Fun (PubConstSetType a) T" (is ?A) and "\a T. t = Fun (PubConstAttackType a) T" (is ?B) and "\a T. t = Fun (PubConstBottom a) T" (is ?C) and "\a T. t = Fun (PubConstOccursSecType a) T" (is ?D) proof - have "admissible_transaction_terms T" using assms(1) unfolding admissible_transaction_def by blast hence "\is_PubConstSetType f" "\is_PubConstAttackType f" "\is_PubConstBottom f" "\is_PubConstOccursSecType f" when "f \ \(funs_term ` (trms_transaction T))" for f using that unfolding 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_PubConstSetType f" "\is_PubConstAttackType f" "\is_PubConstBottom f" "\is_PubConstOccursSecType 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) by force show ?D using *(4) by force qed lemma transactions_have_no_PubConsts': assumes "admissible_transaction T" and "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)" shows "\a T. Fun (PubConstSetType a) T \ subterms t" and "\a T. Fun (PubConstAttackType a) T \ subterms t" and "\a T. Fun (PubConstBottom a) T \ subterms t" and "\a T. Fun (PubConstOccursSecType a) T \ subterms t" using transactions_have_no_PubConsts[OF assms(1)] assms(2) by fast+ 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_selects 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_selects 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_selects_def is_Fun_Set_def by fastforce+ have "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T)" using * by force hence ***: "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x \ set (transaction_fresh T) = {}" using T_valid unfolding wellformed_transaction_def by fast show ?A using **(1,2) *** by (cases t) auto show ?B using **(3,4,5) unfolding is_Set_def by (cases s) auto qed lemma transaction_inset_checks_are_Value_vars: assumes T_valid: "wellformed_transaction T" and "admissible_transaction_checks T" and "\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\" have *: "?x \ set (unlabel (transaction_checks T))" using assms(3) wellformed_transaction_unlabel_cases[OF T_valid, of ?x] by (auto simp add: transaction_strand_def unlabel_def) have **: "is_Var (the_elem_term ?x)" "fst (the_Var (the_elem_term ?x)) = TAtom Value" "is_Fun (the_set_term ?x)" "args (the_set_term ?x) = []" "is_Set (the_Fun (the_set_term ?x))" using * assms(2) unfolding admissible_transaction_checks_def is_Fun_Set_def by fastforce+ have "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using * by force hence ***: "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x \ set (transaction_fresh T) = {}" using T_valid unfolding wellformed_transaction_def by fast show ?A using **(1,2) *** by (cases t) auto show ?B using **(3,4,5) unfolding is_Set_def by (cases s) auto qed lemma transaction_notinset_checks_are_Value_vars: assumes T_valid: "wellformed_transaction T" and "admissible_transaction_checks T" and "\X\\\: F \\: G\ \ set (unlabel (transaction_strand T))" and "(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) proof - let ?x = "\X\\\: F \\: G\" have 0: "?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) hence 1: "F = [] \ length G = 1" using assms(2,4) unfolding admissible_transaction_checks_def by fastforce hence "hd G = (t,s)" using assms(4) 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 0 1 assms(2) unfolding admissible_transaction_checks_def Let_def is_Fun_Set_def by fastforce+ have "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T)" using 0 by force+ moreover have "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_checks T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T)" "set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) = {}" "set (transaction_fresh T) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T) = {}" using T_valid unfolding wellformed_transaction_def by fast+ ultimately have "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x \ set (transaction_fresh T) = {}" "set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p ?x) \ set (transaction_fresh T) = {}" using wellformed_transaction_wf\<^sub>s\<^sub>s\<^sub>t(2,3)[OF T_valid] fv_transaction_unfold[of T] bvars_transaction_unfold[of T] by blast+ hence ***: "fv t \ set (transaction_fresh T) = {}" using assms(4) 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 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_selects T)) \ \x s. r = select\Var x, Fun (Set s) []\ \ fst x = TAtom Value \ x \ fv_transaction T - set (transaction_fresh T)" (is "?B \ ?B'") and "r \ set (unlabel (transaction_checks T)) \ (\x s. (r = \Var x in 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\)" (is "?C \ ?C'") 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 "?D \ ?D'") and "r \ set (unlabel (transaction_send T)) \ \t. r = send\t\" (is "?E \ ?E'") proof - have T_valid: "wellformed_transaction T" using T_adm unfolding admissible_transaction_def by metis show "?A \ ?A'" using T_valid Ball_set[of "unlabel (transaction_receive T)" is_Receive] unfolding wellformed_transaction_def is_Receive_def by blast show "?E \ ?E'" using T_valid 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 have "admissible_transaction_selects T" using T_adm unfolding admissible_transaction_def by simp hence *: "is_InSet r" "the_check r = Assign" "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 unfolding admissible_transaction_selects_def is_Fun_Set_def by fast+ obtain rt rs where r': "r = select\rt,rs\" using *(1,2) by (cases r) auto obtain x where x: "rt = Var x" "fst x = TAtom Value" using *(3,7) r' by auto obtain f S where fS: "rs = Fun f S" using *(4) r' by auto obtain s where s: "f = Set s" using *(5) fS r' by (cases f) auto hence S: "S = []" using *(6) fS r' by (cases S) auto have fv_r1: "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p r \ fv_transaction T" using r fv_transaction_unfold[of T] by auto have fv_r2: "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p r \ set (transaction_fresh T) = {}" using r T_valid unfolding wellformed_transaction_def by fastforce show ?B' using r' x fS s S fv_r1 fv_r2 by simp qed show "?C \ ?C'" proof - assume r: ?C have adm_checks: "admissible_transaction_checks T" using assms unfolding admissible_transaction_def by simp have fv_r1: "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p r \ fv_transaction T" using r fv_transaction_unfold[of T] by auto have fv_r2: "fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p r \ set (transaction_fresh T) = {}" using r T_valid unfolding wellformed_transaction_def by fastforce have "(is_InSet r \ the_check r = Check) \ (is_Equality r \ the_check r = Check) \ is_NegChecks r" using r adm_checks unfolding admissible_transaction_checks_def by fast thus ?C' proof (elim disjE conjE) assume *: "is_InSet r" "the_check r = Check" 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 rt rs where r': "r = \rt in 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 ?C' using r' x fS s S fv_r1 fv_r2 by simp next assume *: "is_NegChecks r" hence **: "bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p r = []" "(the_eqs r = [] \ length (the_ins r) = 1) \ (the_ins r = [] \ length (the_eqs r) = 1)" using r adm_checks unfolding admissible_transaction_checks_def by fast+ show ?C' 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 r adm_checks * ***(1) unfolding admissible_transaction_checks_def is_Fun_Set_def by metis+ 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 ?C' 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 ?C' using *** **(1) * by (cases r) auto qed qed (auto simp add: is_Equality_def the_check_def) qed show "?D \ ?D'" proof - assume r: ?D have adm_upds: "admissible_transaction_updates T" using assms unfolding admissible_transaction_def by simp 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 ?D' using ts x fT ss T by blast qed qed lemma transaction_Value_vars_are_fv: assumes "admissible_transaction T" and "x \ vars_transaction T" and "\\<^sub>v x = TAtom Value" shows "x \ fv_transaction T" using assms \\<^sub>v_TAtom''(2)[of x] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] unfolding admissible_transaction_def by fast lemma protocol_transaction_vars_TAtom_typed: assumes P: "admissible_transaction T" shows "\x \ vars_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" and "\x \ fv_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" and "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" proof - have P': "wellformed_transaction T" using P unfolding admissible_transaction_def by fast show "\x \ vars_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" using P \\<^sub>v_TAtom'' unfolding admissible_transaction_def is_Var_def prot_atom.is_Atom_def the_Var_def by fastforce thus "\x \ fv_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" using vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t by fast have "list_all (\x. fst x = Var Value) (transaction_fresh T)" using P \\<^sub>v_TAtom'' unfolding admissible_transaction_def by fast thus "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" using \\<^sub>v_TAtom''(2) unfolding list_all_iff by fast qed lemma protocol_transactions_no_pubconsts: assumes "admissible_transaction T" shows "Fun (Val (n,True)) S \ subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction T)" using assms transactions_have_no_Value_consts(1) by fast lemma protocol_transactions_no_abss: assumes "admissible_transaction T" shows "Fun (Abs n) S \ subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction T)" using assms transactions_have_no_Value_consts(2) by fast lemma admissible_transaction_strand_sem_fv_ineq: assumes T_adm: "admissible_transaction T" and \: "strand_sem_stateful IK DB (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \" and x: "x \ fv_transaction T - set (transaction_fresh T)" and y: "y \ fv_transaction T - set (transaction_fresh T)" and x_not_y: "x \ y" shows "\ x \ \ \ \ y \ \" proof - have "\Var x != Var y\ \ set (unlabel (transaction_checks T)) \ \Var y != Var x\ \ set (unlabel (transaction_checks T))" using x y x_not_y T_adm unfolding admissible_transaction_def by auto hence "\Var x != Var y\ \ set (unlabel (transaction_strand T)) \ \Var y != Var x\ \ set (unlabel (transaction_strand T))" unfolding transaction_strand_def unlabel_def by auto hence "\\ x != \ y\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \ \\ y != \ x\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" using stateful_strand_step_subst_inI(8)[of _ _ "unlabel (transaction_strand T)" \] subst_lsst_unlabel[of "transaction_strand T" \] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(7)[of "[]" _ "[]"] by force then obtain B where B: "prefix (B@[\\ x != \ y\]) (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) \ prefix (B@[\\ y != \ x\]) (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" unfolding prefix_def by (metis (no_types, hide_lams) 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_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: assumes "admissible_transaction T" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms_transaction T)" by (metis wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code assms admissible_transaction_def admissible_transaction_terms_def) lemma admissible_transaction_no_Ana_Attack: assumes "admissible_transaction_terms T" and "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms_transaction T)" shows "attack\n\ \ set (snd (Ana t))" proof - obtain r where r: "r \ set (unlabel (transaction_strand T))" "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r)" using assms(2) by force obtain K M where t: "Ana t = (K, M)" by (metis surj_pair) show ?thesis proof assume n: "attack\n\ \ set (snd (Ana t))" hence "attack\n\ \ set M" using t by simp hence n': "attack\n\ \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r)" using Ana_subterm[OF t] r(2) subterms_subset by fast hence "\f \ \(funs_term ` trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p r). is_Attack f" using funs_term_Fun_subterm' unfolding is_Attack_def by fast hence "is_Send r" "is_Fun (the_msg r)" "is_Attack (the_Fun (the_msg r))" "args (the_msg r) = []" using assms(1) r(1) unfolding admissible_transaction_terms_def by metis+ hence "t = attack\n\" using n' r(2) unfolding is_Send_def is_Attack_def by auto thus False using n by fastforce qed qed lemma admissible_transaction_occurs_fv_types: assumes "admissible_transaction T" and "x \ vars_transaction T" shows "\a. \ (Var x) = TAtom a \ \ (Var x) \ TAtom OccursSecType" proof - have "is_Var (fst x)" "the_Var (fst x) = Value" using assms unfolding admissible_transaction_def by blast+ thus ?thesis using \\<^sub>v_TAtom''(2)[of x] by force qed lemma admissible_transaction_Value_vars: assumes T: "admissible_transaction T" and x: "x \ fv_transaction T" shows "\\<^sub>v x = TAtom Value" proof - have "x \ vars_transaction T" using x vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] by blast hence "is_Var (fst x)" "the_Var (fst x) = Value" using T assms unfolding admissible_transaction_def list_all_iff by fast+ thus "\\<^sub>v x = TAtom Value" using \\<^sub>v_TAtom''(2)[of x] by force qed subsection \Lemmata: Renaming and Fresh Substitutions\ lemma transaction_renaming_subst_is_renaming: fixes \::"('fun,'atom,'sets) prot_subst" assumes "transaction_renaming_subst \ P A" shows "\m. \ (\,n) = Var (\,n+Suc m)" using assms by (auto simp add: transaction_renaming_subst_def var_rename_def) lemma transaction_renaming_subst_is_renaming': fixes \::"('fun,'atom,'sets) prot_subst" assumes "transaction_renaming_subst \ P A" shows "\y. \ x = Var y" using assms by (auto simp add: transaction_renaming_subst_def var_rename_def) lemma transaction_renaming_subst_vars_disj: fixes \::"('fun,'atom,'sets) prot_subst" assumes "transaction_renaming_subst \ P A" shows "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (\(vars_transaction ` set P))) \ (\(vars_transaction ` set P)) = {}" (is ?A) and "fv\<^sub>s\<^sub>e\<^sub>t (\ ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = {}" (is ?B) and "T \ set P \ vars_transaction T \ range_vars \ = {}" (is "T \ set P \ ?C1") and "T \ set P \ bvars_transaction T \ range_vars \ = {}" (is "T \ set P \ ?C2") and "T \ set P \ fv_transaction T \ range_vars \ = {}" (is "T \ set P \ ?C3") and "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ range_vars \ = {}" (is ?D1) and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ range_vars \ = {}" (is ?D2) and "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ range_vars \ = {}" (is ?D3) proof - define X where "X \ \(vars_transaction ` set P) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" have 1: "finite X" by (simp add: X_def) obtain n where n: "n \ max_var_set X" "\ = var_rename n" using assms unfolding transaction_renaming_subst_def X_def by moura hence 2: "\x \ X. snd x < Suc n" using less_Suc_max_var_set[OF _ 1] unfolding var_rename_def by fastforce have 3: "x \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` X)" "fv (\ x) \ X = {}" "x \ range_vars \" when x: "x \ X" for x using 2 x n unfolding var_rename_def by force+ show ?A ?B using 3(1,2) unfolding X_def by auto show ?C1 when T: "T \ set P" using T 3(3) unfolding X_def by blast thus ?C2 ?C3 when T: "T \ set P" using T by (simp_all add: disjoint_iff_not_equal vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t) show ?D1 using 3(3) unfolding X_def by auto thus ?D2 ?D3 by (simp_all add: disjoint_iff_not_equal vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t) qed lemma transaction_renaming_subst_wt: fixes \::"('fun,'atom,'sets) prot_subst" assumes "transaction_renaming_subst \ P A" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" proof - { fix x::"('fun,'atom,'sets) prot_var" obtain \ n where x: "x = (\,n)" by moura then obtain m where m: "\ x = Var (\,m)" using assms transaction_renaming_subst_is_renaming by moura hence "\ (\ x) = \\<^sub>v x" using x by (simp add: \\<^sub>v_def) } thus ?thesis by (simp add: wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) qed lemma transaction_renaming_subst_is_wf_trm: fixes \::"('fun,'atom,'sets) prot_subst" assumes "transaction_renaming_subst \ P A" shows "wf\<^sub>t\<^sub>r\<^sub>m (\ v)" proof - obtain \ n where "v = (\, n)" by moura then obtain m where "\ v = Var (\, n + Suc m)" using transaction_renaming_subst_is_renaming[OF assms] by moura thus ?thesis by (metis wf_trm_Var) qed lemma transaction_renaming_subst_range_wf_trms: fixes \::"('fun,'atom,'sets) prot_subst" assumes "transaction_renaming_subst \ P A" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (metis transaction_renaming_subst_is_wf_trm[OF assms] wf_trm_subst_range_iff) lemma transaction_renaming_subst_range_notin_vars: fixes \::"('fun,'atom,'sets) prot_subst" assumes "transaction_renaming_subst \ P \" shows "\y. \ x = Var y \ y \ \(vars_transaction ` set P) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" proof - obtain \ n where x: "x = (\,n)" by (metis surj_pair) define y where "y \ \m. (\,n+Suc m)" have "\m \ max_var_set (\(vars_transaction ` set P) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \). \ x = Var (y m)" using assms x by (auto simp add: y_def transaction_renaming_subst_def var_rename_def) moreover have "finite (\(vars_transaction ` set P) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by auto ultimately show ?thesis using x unfolding y_def by force qed lemma transaction_renaming_subst_var_obtain: fixes \::"('fun,'atom,'sets) prot_subst" assumes x: "x \ fv\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>s\<^sub>t \)" and \: "transaction_renaming_subst \ P \" shows "\y. \ y = Var x" proof - obtain y where y: "y \ fv\<^sub>s\<^sub>s\<^sub>t S" "x \ fv (\ y)" using fv\<^sub>s\<^sub>s\<^sub>t_subst_obtain_var[OF x] by moura thus ?thesis using transaction_renaming_subst_is_renaming'[OF \, of y] by fastforce qed lemma transaction_fresh_subst_is_wf_trm: fixes \::"('fun,'atom,'sets) prot_subst" assumes "transaction_fresh_subst \ T A" shows "wf\<^sub>t\<^sub>r\<^sub>m (\ v)" proof (cases "v \ subst_domain \") case True then obtain n where "\ v = Fun (Val n) []" using assms unfolding transaction_fresh_subst_def by moura thus ?thesis by auto qed auto lemma transaction_fresh_subst_wt: fixes \::"('fun,'atom,'sets) prot_subst" assumes "transaction_fresh_subst \ T A" and "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" proof - have 1: "subst_domain \ = set (transaction_fresh T)" and 2: "\t \ subst_range \. \n. t = Fun (Val n) []" using assms(1) unfolding transaction_fresh_subst_def by metis+ { fix x::"('fun,'atom,'sets) prot_var" have "\ (Var x) = \ (\ x)" using assms(2) 1 2 by (cases "x \ subst_domain \") force+ } thus ?thesis by (simp add: wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) qed lemma transaction_fresh_subst_domain: fixes \::"('fun,'atom,'sets) prot_subst" assumes "transaction_fresh_subst \ T \" 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) prot_subst" assumes "transaction_fresh_subst \ T \" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (metis transaction_fresh_subst_is_wf_trm[OF assms] wf_trm_subst_range_iff) lemma transaction_fresh_subst_range_fresh: fixes \::"('fun,'atom,'sets) prot_subst" assumes "transaction_fresh_subst \ T \" shows "\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and "\t \ subst_range \. t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T))" using assms unfolding transaction_fresh_subst_def by meson+ lemma transaction_fresh_subst_sends_to_val: fixes \::"('fun,'atom,'sets) prot_subst" assumes "transaction_fresh_subst \ T \" and "y \ set (transaction_fresh T)" 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 thus ?thesis using assms that unfolding transaction_fresh_subst_def by fastforce qed lemma transaction_fresh_subst_sends_to_val': fixes \ \::"('fun,'atom,'sets) prot_subst" assumes "transaction_fresh_subst \ T \" and "y \ set (transaction_fresh T)" obtains n where "(\ \\<^sub>s \) y \ \ = Fun (Val n) []" "Fun (Val n) [] \ subst_range \" proof - obtain n where "\ y = Fun (Val n) []" "Fun (Val n) [] \ subst_range \" using transaction_fresh_subst_sends_to_val[OF assms] by moura thus ?thesis using that by (fastforce simp add: subst_compose_def) qed lemma transaction_fresh_subst_grounds_domain: fixes \::"('fun,'atom,'sets) prot_subst" assumes "transaction_fresh_subst \ T \" and "y \ set (transaction_fresh T)" shows "fv (\ y) = {}" proof - obtain n where "\ y = Fun (Val n) []" using transaction_fresh_subst_sends_to_val[OF assms] by moura thus ?thesis by simp qed lemma transaction_fresh_subst_transaction_renaming_subst_range: fixes \ \::"('fun,'atom,'sets) prot_subst" assumes "transaction_fresh_subst \ T \" "transaction_renaming_subst \ P \" shows "x \ set (transaction_fresh T) \ \n. (\ \\<^sub>s \) x = Fun (Val (n,False)) []" and "x \ set (transaction_fresh T) \ \y. (\ \\<^sub>s \) x = Var y" proof - assume "x \ set (transaction_fresh T)" then obtain n where "\ x = Fun (Val (n,False)) []" using assms(1) unfolding transaction_fresh_subst_def by fastforce thus "\n. (\ \\<^sub>s \) x = Fun (Val (n,False)) []" using subst_compose[of \ \ x] by simp next assume "x \ set (transaction_fresh T)" hence "\ x = Var x" using assms(1) unfolding transaction_fresh_subst_def by fastforce thus "\y. (\ \\<^sub>s \) x = Var y" using transaction_renaming_subst_is_renaming[OF assms(2)] subst_compose[of \ \ x] by (cases x) force qed lemma transaction_fresh_subst_transaction_renaming_subst_range': fixes \ \::"('fun,'atom,'sets) prot_subst" assumes "transaction_fresh_subst \ T \" "transaction_renaming_subst \ P \" and "t \ subst_range (\ \\<^sub>s \)" shows "(\n. t = Fun (Val (n,False)) []) \ (\x. t = Var x)" proof - obtain x where "x \ subst_domain (\ \\<^sub>s \)" "(\ \\<^sub>s \) x = t" using assms(3) by auto thus ?thesis using transaction_fresh_subst_transaction_renaming_subst_range[OF assms(1,2), of x] by auto qed lemma transaction_fresh_subst_transaction_renaming_subst_range'': fixes \ \::"('fun,'atom,'sets) prot_subst" assumes s: "transaction_fresh_subst \ T \" "transaction_renaming_subst \ P \" and y: "y \ fv ((\ \\<^sub>s \) x)" shows "\ x = Var x" and "\ x = Var y" and "(\ \\<^sub>s \) x = Var y" proof - have "\z. z \ fv (\ x)" using y subst_compose_fv' by fast hence x: "x \ subst_domain \" using y transaction_fresh_subst_domain[OF s(1)] transaction_fresh_subst_grounds_domain[OF s(1), of x] by blast thus "\ x = Var x" by blast thus "\ x = Var y" "(\ \\<^sub>s \) x = Var y" using y transaction_renaming_subst_is_renaming'[OF s(2), of x] unfolding subst_compose_def by fastforce+ qed lemma transaction_fresh_subst_transaction_renaming_subst_vars_subset: fixes \ \::"('fun,'atom,'sets) prot_subst" assumes \: "transaction_fresh_subst \ T \" and \: "transaction_renaming_subst \ P \" shows "\(fv_transaction ` set P) \ subst_domain (\ \\<^sub>s \)" (is ?A) and "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ subst_domain (\ \\<^sub>s \)" (is ?B) and "T' \ set P \ fv_transaction T' \ subst_domain (\ \\<^sub>s \)" (is "T' \ set P \ ?C") and "T' \ set P \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T' \\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ \\<^sub>s \)) \ range_vars (\ \\<^sub>s \)" (is "T' \ set P \ ?D") proof - have *: "x \ subst_domain (\ \\<^sub>s \)" for x 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 thus ?thesis using subst_domain_subst_compose[of \ \] by blast next case False hence "(\ \\<^sub>s \) x = \ x" unfolding subst_compose_def by fastforce moreover have "\ x \ Var x" using transaction_renaming_subst_is_renaming[OF \, of "fst x" "snd x"] by (cases x) auto ultimately show ?thesis by fastforce qed show ?A ?B using * by blast+ show ?C when T: "T' \ set P" using T * by blast hence "fv\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T') \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \) \ range_vars (\ \\<^sub>s \)" when T: "T' \ set P" using T fv\<^sub>s\<^sub>s\<^sub>t_subst_subset_range_vars_if_subset_domain by blast thus ?D when T: "T' \ set P" by (metis T unlabel_subst) qed lemma transaction_fresh_subst_transaction_renaming_subst_vars_disj: fixes \ \::"('fun,'atom,'sets) prot_subst" assumes \: "transaction_fresh_subst \ T \" and \: "transaction_renaming_subst \ P \" shows "fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \) ` (\(vars_transaction ` set P))) \ (\(vars_transaction ` set P)) = {}" (is ?A) and "x \ \(vars_transaction ` set P) \ fv ((\ \\<^sub>s \) x) \ (\(vars_transaction ` set P)) = {}" (is "?B' \ ?B") and "T' \ set P \ vars_transaction T' \ range_vars (\ \\<^sub>s \) = {}" (is "T' \ set P \ ?C1") and "T' \ set P \ bvars_transaction T' \ range_vars (\ \\<^sub>s \) = {}" (is "T' \ set P \ ?C2") and "T' \ set P \ fv_transaction T' \ range_vars (\ \\<^sub>s \) = {}" (is "T' \ set P \ ?C3") and "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ range_vars (\ \\<^sub>s \) = {}" (is ?D1) and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ range_vars (\ \\<^sub>s \) = {}" (is ?D2) and "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ range_vars (\ \\<^sub>s \) = {}" (is ?D3) proof - note 0 = transaction_renaming_subst_vars_disj[OF \] show ?A proof (cases "fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \) ` (\(vars_transaction ` set P))) = {}") case False hence "\x \ (\(vars_transaction ` set P)). (\ \\<^sub>s \) x = \ x \ fv ((\ \\<^sub>s \) x) = {}" using transaction_fresh_subst_transaction_renaming_subst_range''[OF \ \] by auto thus ?thesis using 0(1) by force qed blast thus "?B' \ ?B" by auto have 1: "range_vars (\ \\<^sub>s \) \ range_vars \" using range_vars_subst_compose_subset[of \ \] transaction_fresh_subst_domain[OF \] transaction_fresh_subst_grounds_domain[OF \] by force 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_fresh_subst_transaction_renaming_subst_trms: fixes \ \::"('fun,'atom,'sets) prot_subst" assumes "transaction_fresh_subst \ T \" "transaction_renaming_subst \ P \" and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ subst_domain \ = {}" and "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ subst_domain \ = {}" shows "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (S \\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ \\<^sub>s \))) = subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S) \\<^sub>s\<^sub>e\<^sub>t (\ \\<^sub>s \)" proof - have 1: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t S). (\f. (\ \\<^sub>s \) x = Fun f []) \ (\y. (\ \\<^sub>s \) x = Var y)" using transaction_fresh_subst_transaction_renaming_subst_range[OF assms(1,2)] by blast have 2: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ subst_domain (\ \\<^sub>s \) = {}" using assms(3,4) subst_domain_compose[of \ \] by blast show ?thesis using subterms_subst_lsst[OF 1 2] by simp qed lemma transaction_fresh_subst_transaction_renaming_wt: fixes \ \::"('fun,'atom,'sets) prot_subst" assumes "transaction_fresh_subst \ T \" "transaction_renaming_subst \ P \" and "\x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" using transaction_renaming_subst_wt[OF assms(2)] transaction_fresh_subst_wt[OF assms(1,3)] by (metis wt_subst_compose) lemma transaction_fresh_subst_transaction_renaming_fv: fixes \ \::"('fun,'atom,'sets) prot_subst" assumes \: "transaction_fresh_subst \ T A" and \: "transaction_renaming_subst \ P A" and x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \))" shows "\y \ fv_transaction T - set (transaction_fresh T). (\ \\<^sub>s \) y = Var x" proof - have "x \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T) \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" using x fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \"] by argo then obtain y where "y \ fv_transaction T" "x \ fv ((\ \\<^sub>s \) y)" by (metis fv\<^sub>s\<^sub>s\<^sub>t_subst_obtain_var) thus ?thesis using transaction_fresh_subst_transaction_renaming_subst_range[OF \ \, of y] by (cases "y \ set (transaction_fresh T)") force+ qed lemma transaction_fresh_subst_transaction_renaming_subst_occurs_fact_send_receive: fixes t::"('fun,'atom,'sets) prot_term" assumes \: "transaction_fresh_subst \ T \" and \: "transaction_renaming_subst \ P \" and T: "wellformed_transaction T" shows "send\occurs t\ \ set (unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)) \ \s. send\occurs s\ \ set (unlabel (transaction_send T)) \ t = s \ \ \\<^sub>s \" (is "?A \ ?A'") and "receive\occurs t\ \ set (unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)) \ \s. receive\occurs s\ \ set (unlabel (transaction_receive T)) \ t = s \ \ \\<^sub>s \" (is "?B \ ?B'") proof - assume ?A then obtain s where s: "send\s\ \ set (unlabel (transaction_strand T))" "occurs t = s \ \ \\<^sub>s \" using stateful_strand_step_subst_inv_cases(1)[ of "occurs t" "unlabel (transaction_strand T)" "\ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \"] by auto note 0 = s(2) transaction_fresh_subst_transaction_renaming_subst_range[OF \ \] have "\u. s = occurs u" proof (cases s) case (Var x) hence "(\n. s \ \ \\<^sub>s \ = Fun (Val (n, False)) []) \ (\y. s \ \ \\<^sub>s \ = Var y)" using 0(2,3)[of x] by (auto simp del: subst_subst_compose) thus ?thesis using 0(1) by simp next case (Fun f T) hence 1: "f = OccursFact" "length T = 2" "T ! 0 \ \ \\<^sub>s \ = Fun OccursSec []" "T ! 1 \ \ \\<^sub>s \ = t" using 0(1) by auto have "T ! 0 = Fun OccursSec []" proof (cases "T ! 0") case (Var x) thus ?thesis using 0(2,3)[of x] 1(3) by (auto simp del: subst_subst_compose) qed (use 1(3) in simp) thus ?thesis using Fun 1 0(1) by (auto simp del: subst_subst_compose) qed then obtain u where u: "s = occurs u" by moura hence "t = u \ \ \\<^sub>s \" using s(2) by fastforce thus ?A' using s u wellformed_transaction_strand_unlabel_memberD(8)[OF T] by metis next assume ?B then obtain s where s: "receive\s\ \ set (unlabel (transaction_strand T))" "occurs t = s \ \ \\<^sub>s \" using stateful_strand_step_subst_inv_cases(2)[ of "occurs t" "unlabel (transaction_strand T)" "\ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \"] by auto note 0 = s(2) transaction_fresh_subst_transaction_renaming_subst_range[OF \ \] have "\u. s = occurs u" proof (cases s) case (Var x) hence "(\n. s \ \ \\<^sub>s \ = Fun (Val (n, False)) []) \ (\y. s \ \ \\<^sub>s \ = Var y)" using 0(2,3)[of x] by (auto simp del: subst_subst_compose) thus ?thesis using 0(1) by simp next case (Fun f T) hence 1: "f = OccursFact" "length T = 2" "T ! 0 \ \ \\<^sub>s \ = Fun OccursSec []" "T ! 1 \ \ \\<^sub>s \ = t" using 0(1) by auto have "T ! 0 = Fun OccursSec []" proof (cases "T ! 0") case (Var x) thus ?thesis using 0(2,3)[of x] 1(3) by (auto simp del: subst_subst_compose) qed (use 1(3) in simp) thus ?thesis using Fun 1 0(1) by (auto simp del: subst_subst_compose) qed then obtain u where u: "s = occurs u" by moura hence "t = u \ \ \\<^sub>s \" using s(2) by fastforce thus ?B' using s u wellformed_transaction_strand_unlabel_memberD(1)[OF T] by metis qed lemma transaction_fresh_subst_proj: assumes "transaction_fresh_subst \ T A" shows "transaction_fresh_subst \ (transaction_proj n T) (proj n A)" using assms transaction_proj_fresh_eq[of n T] contra_subsetD[OF subterms\<^sub>s\<^sub>e\<^sub>t_mono[OF transaction_proj_trms_subset[of n T]]] contra_subsetD[OF subterms\<^sub>s\<^sub>e\<^sub>t_mono[OF trms\<^sub>s\<^sub>s\<^sub>t_proj_subset(1)[of n A]]] unfolding transaction_fresh_subst_def by metis lemma transaction_renaming_subst_proj: assumes "transaction_renaming_subst \ P A" shows "transaction_renaming_subst \ (map (transaction_proj n) P) (proj n A)" proof - let ?X = "\P A. \(vars_transaction ` set P) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" define Y where "Y \ ?X (map (transaction_proj n) P) (proj n A)" define Z where "Z \ ?X P A" have "Y \ Z" using sst_vars_proj_subset(3)[of n A] transaction_proj_vars_subset[of n] unfolding Y_def Z_def by fastforce hence "insert 0 (snd ` Y) \ insert 0 (snd ` Z)" by blast moreover have "finite (insert 0 (snd ` Z))" "finite (insert 0 (snd ` Y))" unfolding Y_def Z_def by auto ultimately have 0: "max_var_set Y \ max_var_set Z" using Max_mono by blast have "\n\max_var_set Z. \ = var_rename n" using assms unfolding transaction_renaming_subst_def Z_def by blast hence "\n\max_var_set Y. \ = var_rename n" using 0 le_trans by fast thus ?thesis unfolding transaction_renaming_subst_def Y_def by blast qed lemma protocol_transaction_wf_subst: fixes \ \::"('fun,'atom,'sets) prot_subst" assumes T: "wf'\<^sub>s\<^sub>s\<^sub>t (set (transaction_fresh T)) (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)))" and \: "transaction_fresh_subst \ T \" and \: "transaction_renaming_subst \ P \" shows "wf'\<^sub>s\<^sub>s\<^sub>t {} (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)))" proof - have 0: "range_vars \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)) = {}" "ground (\ ` set (transaction_fresh T))" "ground (\ ` {})" using transaction_fresh_subst_domain[OF \] transaction_fresh_subst_grounds_domain[OF \] by fastforce+ have "wf'\<^sub>s\<^sub>s\<^sub>t {} ((unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)) \\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>s\<^sub>t \)" by (metis wf\<^sub>s\<^sub>s\<^sub>t_subst_apply[OF wf\<^sub>s\<^sub>s\<^sub>t_subst_apply[OF T]] 0(2,3)) thus ?thesis by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst labeled_stateful_strand_subst_comp[OF 0(1)]) qed subsection \Lemmata: Reachable Constraints\ lemma reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: assumes "\T \ set P. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms_transaction T)" and "\ \ reachable_constraints P" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using assms(2) proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \) have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms_transaction T)" using assms(1) step.hyps(2) by blast moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using wf_trms_subst_compose[of \ \] transaction_renaming_subst_range_wf_trms[OF step.hyps(4)] transaction_fresh_subst_range_wf_trms[OF step.hyps(3)] by fastforce ultimately have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms_transaction T \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \)" by (metis wf_trms_subst) hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \))" using wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_trms\<^sub>s\<^sub>s\<^sub>t_subst unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \"] by metis hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)))" using trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq by blast thus ?case using step.IH unlabel_append[of \] trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \"] by auto qed simp lemma reachable_constraints_TAtom_types: assumes "\ \ reachable_constraints P" and "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" shows "\\<^sub>v ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ (\T \ set P. \\<^sub>v ` fv_transaction T)" (is "?A \") and "\\<^sub>v ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ (\T \ set P. \\<^sub>v ` bvars_transaction T)" (is "?B \") and "\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ (\T \ set P. \\<^sub>v ` vars_transaction T)" (is "?C \") using assms(1) proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \) define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" have 2: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" using transaction_renaming_subst_wt[OF step.hyps(4)] transaction_fresh_subst_wt[OF step.hyps(3)] by (metis step.hyps(2) assms(2) wt_subst_compose) have 3: "\t \ subst_range (\ \\<^sub>s \). fv t = {} \ (\x. t = Var x)" using transaction_fresh_subst_transaction_renaming_subst_range'[OF step.hyps(3,4)] by fastforce have "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" unfolding T'_def by (metis fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq, metis bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq, metis vars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq) hence "\ ` Var ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ \ ` Var ` fv_transaction T" "\ ` Var ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = \ ` Var ` bvars_transaction T" "\ ` Var ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ \ ` Var ` vars_transaction T" using wt_subst_lsst_vars_type_subset[OF 2 3, of "transaction_strand T"] by argo+ hence "\\<^sub>v ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ \\<^sub>v ` fv_transaction T" "\\<^sub>v ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = \\<^sub>v ` bvars_transaction T" "\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ \\<^sub>v ` vars_transaction T" by (metis \\<^sub>v_Var_image)+ hence 4: "\\<^sub>v ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ (\T \ set P. \\<^sub>v ` fv_transaction T)" "\\<^sub>v ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ (\T \ set P. \\<^sub>v ` bvars_transaction T)" "\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ (\T \ set P. \\<^sub>v ` vars_transaction T)" using step.hyps(2) by fast+ have 5: "\\<^sub>v ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ @ T') = (\\<^sub>v ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ (\\<^sub>v ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" "\\<^sub>v ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ @ T') = (\\<^sub>v ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ (\\<^sub>v ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" "\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ @ T') = (\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ (\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" using unlabel_append[of \ T'] fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel T'"] bvars\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel T'"] vars\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel T'"] by auto { case 1 thus ?case using step.IH(1) 4(1) 5(1) unfolding T'_def by (simp del: subst_subst_compose fv\<^sub>s\<^sub>s\<^sub>t_def) } { case 2 thus ?case using step.IH(2) 4(2) 5(2) unfolding T'_def by (simp del: subst_subst_compose bvars\<^sub>s\<^sub>s\<^sub>t_def) } { case 3 thus ?case using step.IH(3) 4(3) 5(3) unfolding T'_def by (simp del: subst_subst_compose) } qed simp_all lemma reachable_constraints_no_bvars: assumes \: "\ \ reachable_constraints P" and P: "\T \ set P. bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) = {}" shows "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" using assms proof (induction) case init then show ?case unfolding unlabel_def by auto next case (step \ T \ \) then have "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" by metis moreover have "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)) = {}" using step by (metis bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq) ultimately show ?case using bvars\<^sub>s\<^sub>s\<^sub>t_append unlabel_append by (metis sup_bot.left_neutral) qed lemma reachable_constraints_fv_bvars_disj: assumes \_reach: "\ \ reachable_constraints P" and P: "\S \ set P. admissible_transaction S" shows "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" proof - let ?X = "\T \ set P. bvars_transaction T" note 0 = transactions_fv_bvars_disj[OF P] have 1: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ ?X" using \_reach proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \) have "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)) = bvars_transaction T" using bvars\<^sub>s\<^sub>s\<^sub>t_subst[of "unlabel (transaction_strand T)" "\ \\<^sub>s \"] bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \"] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst[of "transaction_strand T" "\ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \"] by argo hence "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)) \ ?X" using step.hyps(2) by blast thus ?case using step.IH bvars\<^sub>s\<^sub>s\<^sub>t_append by auto qed (simp add: unlabel_def bvars\<^sub>s\<^sub>s\<^sub>t_def) have 2: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ ?X = {}" using \_reach proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \) have "x \ y" when x: "x \ ?X" and y: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" for x y proof - obtain y' where y': "y' \ fv_transaction T" "y \ fv ((\ \\<^sub>s \) y')" using y unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \"] by (metis fv\<^sub>s\<^sub>s\<^sub>t_subst_obtain_var) have "y \ \(vars_transaction ` set P)" using transaction_fresh_subst_transaction_renaming_subst_range''[OF step.hyps(3,4) y'(2)] transaction_renaming_subst_range_notin_vars[OF step.hyps(4), of y'] by auto thus ?thesis using x vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t by fast qed hence "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \) \ ?X = {}" by blast thus ?case using step.IH fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \"] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst[of "transaction_strand T" "\ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \"] fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)"] unlabel_append[of \ "transaction_strand T"] by force qed (simp add: unlabel_def fv\<^sub>s\<^sub>s\<^sub>t_def) show ?thesis using 0 1 2 by blast qed lemma reachable_constraints_vars_TAtom_typed: assumes \_reach: "\ \ reachable_constraints P" and P: "\T \ set P. admissible_transaction T" and x: "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" shows "\\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" proof - have \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (metis reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s admissible_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P \_reach) have T_adm: "admissible_transaction T" when "T \ set P" for T by (meson that Ball_set P) have "\T\set P. \x\set (transaction_fresh T). \\<^sub>v x = TAtom Value" using protocol_transaction_vars_TAtom_typed(3) P by blast hence *: "\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ (\T\set P. \\<^sub>v ` vars_transaction T)" using reachable_constraints_TAtom_types[of \ P, OF \_reach] by auto have "\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ TAtom ` insert Value (range Atom)" proof - have "\\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" when "T \ set P" "x \ vars_transaction T" for T x using that protocol_transaction_vars_TAtom_typed(1)[of T] P unfolding admissible_transaction_def by blast hence "(\T\set P. \\<^sub>v ` vars_transaction T) \ TAtom ` insert Value (range Atom)" using P by blast thus "\\<^sub>v ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ TAtom ` insert Value (range Atom)" using * by auto qed thus ?thesis using x by auto qed lemma reachable_constraints_Value_vars_are_fv: assumes \_reach: "\ \ reachable_constraints P" and P: "\T \ set P. admissible_transaction T" and x: "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" and "\\<^sub>v x = TAtom Value" shows "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" proof - have "\T\set P. bvars_transaction T = {}" using P unfolding list_all_iff admissible_transaction_def by metis hence \_no_bvars: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" using reachable_constraints_no_bvars[OF \_reach] by metis thus ?thesis using x vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] by blast qed lemma reachable_constraints_subterms_subst: assumes \_reach: "\ \ reachable_constraints P" and \: "welltyped_constraint_model \ \" and P: "\T \ set P. admissible_transaction T" shows "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\ \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = (subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \\<^sub>s\<^sub>e\<^sub>t \" proof - have \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" by (metis reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s admissible_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P \_reach) from \ have \': "welltyped_constraint_model \ \" using welltyped_constraint_model_prefix by auto have 1: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \). (\f. \ x = Fun f []) \ (\y. \ x = Var y)" proof fix x assume xa: "x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" have "\f T. \ x = Fun f T" using \ interpretation_grounds[of \ "Var x"] unfolding welltyped_constraint_model_def constraint_model_def by (cases "\ x") auto then obtain f T where fT_p: "\ x = Fun f T" by auto hence "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" using \ unfolding welltyped_constraint_model_def constraint_model_def using wf_trm_subst_rangeD by metis moreover have "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using xa var_subterm_trms\<^sub>s\<^sub>s\<^sub>t_is_vars\<^sub>s\<^sub>s\<^sub>t[of x "unlabel \"] vars_iff_subtermeq[of x] by auto hence "\a. \\<^sub>v x = TAtom a" using reachable_constraints_vars_TAtom_typed[OF \_reach P] by blast hence "\a. \ (Var x) = TAtom a" by simp hence "\a. \ (Fun f T) = TAtom a" by (metis (no_types, hide_lams) \' welltyped_constraint_model_def fT_p wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) ultimately show "(\f. \ x = Fun f []) \ (\y. \ x = Var y)" using TAtom_term_cases fT_p by metis qed have "\T\set P. bvars_transaction T = {}" using assms unfolding list_all_iff admissible_transaction_def by metis then have "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}" using reachable_constraints_no_bvars assms by metis then have 2: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ subst_domain \ = {}" by auto show ?thesis using subterms_subst_lsst[OF _ 2] 1 by simp qed lemma reachable_constraints_val_funs_private: assumes \_reach: "\ \ reachable_constraints P" and P: "\T \ set P. admissible_transaction T" and f: "f \ \(funs_term ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "is_Val f \ \public f" and "\is_Abs f" proof - have "(is_Val f \ \public f) \ \is_Abs f" using \_reach f proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \) let ?T' = "unlabel (transaction_strand T) \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \" let ?T'' = "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \" have T: "admissible_transaction_terms T" using P step.hyps(2) unfolding admissible_transaction_def by metis show ?thesis using step proof (cases "f \ \(funs_term ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)") case False then obtain t where t: "t \ trms\<^sub>s\<^sub>s\<^sub>t ?T'" "f \ funs_term t" using step.prems trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of ?T''] trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?T'')"] unlabel_append[of \ "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?T''"] unlabel_subst[of "transaction_strand T"] by fastforce show ?thesis using trms\<^sub>s\<^sub>s\<^sub>t_funs_term_cases[OF t] proof assume "\u \ trms_transaction T. f \ funs_term u" thus ?thesis using T unfolding admissible_transaction_terms_def by blast next assume "\x \ fv_transaction T. f \ funs_term ((\ \\<^sub>s \) x)" then obtain x where "x \ fv_transaction T" "f \ funs_term ((\ \\<^sub>s \) x)" by moura thus ?thesis using transaction_fresh_subst_transaction_renaming_subst_range[OF step.hyps(3,4), of x] by (force simp del: subst_subst_compose) qed qed simp qed simp thus "is_Val f \ \public f" "\is_Abs f" by simp_all qed lemma reachable_constraints_occurs_fact_ik_case: assumes \_reach: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction T" and occ: "occurs t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" shows "\n. t = Fun (Val (n,False)) []" using \_reach occ proof (induction A rule: reachable_constraints.induct) case (step A T \ \) define \ where "\ \ \ \\<^sub>s \" have T: "wellformed_transaction T" "admissible_transaction_occurs_checks T" using P step.hyps(2) unfolding list_all_iff admissible_transaction_def by blast+ show ?case proof (cases "occurs t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A") case False hence "occurs t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using step.prems unfolding \_def by simp hence "receive\occurs t\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)))" unfolding ik\<^sub>s\<^sub>s\<^sub>t_def by force hence "send\occurs t\ \ set (unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(1) by blast then obtain s where s: "send\s\ \ set (unlabel (transaction_strand T))" "s \ \ = occurs t" by (metis (no_types) stateful_strand_step_subst_inv_cases(1) unlabel_subst) note 0 = transaction_fresh_subst_transaction_renaming_subst_range[OF step.hyps(3,4)] have 1: "send\s\ \ set (unlabel (transaction_send T))" using s(1) wellformed_transaction_strand_unlabel_memberD(8)[OF T(1)] by blast have 2: "is_Send (send\s\)" unfolding is_Send_def by simp have 3: "\u. s = occurs u" proof - { fix z have "(\n. \ z = Fun (Val (n, False)) []) \ (\y. \ z = Var y)" using 0 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(2) by (cases s) auto thus ?thesis using *(2) s(2) by (cases u) auto qed obtain x where x: "x \ set (transaction_fresh T)" "s = occurs (Var x)" using T(2) 1 2 3 unfolding admissible_transaction_occurs_checks_def by fastforce have "t = \ x" using s(2) x(2) by auto thus ?thesis using 0(1)[OF x(1)] unfolding \_def by fast qed (simp add: step.IH) qed simp lemma reachable_constraints_occurs_fact_send_ex: assumes \_reach: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction T" and x: "\\<^sub>v x = TAtom Value" "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" (* shows "\B. prefix B A \ x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \ send\occurs (Var x)\ \ set (unlabel A)" *) shows "send\occurs (Var x)\ \ set (unlabel A)" using \_reach x(2) proof (induction A rule: reachable_constraints.induct) case (step A T \ \) have T: "admissible_transaction_occurs_checks T" using P step.hyps(2) unfolding list_all_iff admissible_transaction_def by blast show ?case proof (cases "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A") case True show ?thesis using step.IH[OF True] unlabel_append[of A] by auto next case False then obtain y where y: "y \ fv_transaction T - set (transaction_fresh T)" "(\ \\<^sub>s \) y = Var x" using transaction_fresh_subst_transaction_renaming_fv[OF step.hyps(3,4), of x] step.prems(1) fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel A"] unlabel_append[of A] by auto have "\ y = Var y" using y(1) step.hyps(3) unfolding transaction_fresh_subst_def by auto hence "\ y = Var x" using y(2) unfolding subst_compose_def by simp hence y_val: "fst y = TAtom Value" using x(1) \\<^sub>v_TAtom''[of x] \\<^sub>v_TAtom''[of y] wt_subst_trm''[OF transaction_renaming_subst_wt[OF step.hyps(4)], of "Var y"] by force hence "receive\occurs (Var y)\ \ set (unlabel (transaction_receive T))" using y(1) T unfolding admissible_transaction_occurs_checks_def by fast hence *: "receive\occurs (Var y)\ \ set (unlabel (transaction_strand T))" using transaction_strand_subsets(6) by blast have "receive\occurs (Var x)\ \ set (unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \))" using y(2) unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \"] stateful_strand_step_subst_inI(2)[OF *, of "\ \\<^sub>s \"] by (auto simp del: subst_subst_compose) hence "send\occurs (Var x)\ \ set (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)))" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(2) by blast thus ?thesis using unlabel_append[of A] by fastforce qed qed simp lemma reachable_constraints_db\<^sub>l\<^sub>s\<^sub>s\<^sub>t_set_args_empty: assumes \: "\ \ reachable_constraints P" and PP: "list_all wellformed_transaction P" and admissible_transaction_updates: "let f = (\T. \x \ set (unlabel (transaction_updates T)). is_Update x \ is_Var (the_elem_term x) \ is_Fun_Set (the_set_term x) \ fst (the_Var (the_elem_term x)) = TAtom Value) in list_all f P" and d: "(t, s) \ set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" shows "\ss. s = Fun (Set ss) []" using \ d proof (induction) case (step \ TT \ \) let ?TT = "transaction_strand TT \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \" let ?TTu = "unlabel ?TT" let ?TTd = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?TT" let ?TTdu = "unlabel ?TTd" from step(6) have "(t, s) \ set (db'\<^sub>s\<^sub>s\<^sub>t ?TTdu \ (db'\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ []))" unfolding db\<^sub>s\<^sub>s\<^sub>t_def by (simp add: db\<^sub>s\<^sub>s\<^sub>t_append) hence "(t, s) \ set (db'\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ []) \ (\t' s'. insert\t',s'\ \ set ?TTdu \ t = t' \ \ \ s = s' \ \)" using db\<^sub>s\<^sub>s\<^sub>t_in_cases[of t "s" ?TTdu \] by metis thus ?case proof assume "\t' s'. insert\t',s'\ \ set ?TTdu \ t = t' \ \ \ s = s' \ \" then obtain t' s' where t's'_p: "insert\t',s'\ \ set ?TTdu" "t = t' \ \" "s = s' \ \" by metis then obtain lll where "(lll, insert\t',s'\) \ set ?TTd" by (meson unlabel_mem_has_label) hence "(lll, insert\t',s'\) \ set (transaction_strand TT \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_steps_iff(4) by blast hence "insert\t',s'\ \ set ?TTu" by (meson unlabel_in) hence "insert\t',s'\ \ set ((unlabel (transaction_strand TT)) \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" by (simp add: subst_lsst_unlabel) hence "insert\t',s'\ \ (\x. x \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \) ` set (unlabel (transaction_strand TT))" unfolding subst_apply_stateful_strand_def by auto then obtain u where "u \ set (unlabel (transaction_strand TT)) \ u \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \ = insert\t',s'\" by auto hence "\t'' s''. insert\t'',s''\ \ set (unlabel (transaction_strand TT)) \ t' = t'' \ \ \\<^sub>s \ \ s' = s'' \ \ \\<^sub>s \" by (cases u) auto then obtain t'' s'' where t''s''_p: "insert\t'',s''\ \ set (unlabel (transaction_strand TT)) \ t' = t'' \ \ \\<^sub>s \ \ s' = s'' \ \ \\<^sub>s \" by auto hence "insert\t'',s''\ \ set (unlabel (transaction_updates TT))" using is_Update_in_transaction_updates[of "insert\t'',s''\" TT] using PP step(2) unfolding list_all_iff by auto moreover have "\x\set (unlabel (transaction_updates TT)). is_Fun_Set (the_set_term x)" using step(2) admissible_transaction_updates unfolding is_Fun_Set_def list_all_iff by auto ultimately have "is_Fun_Set (the_set_term (insert\t'',s''\))" by auto moreover have "s' = s'' \ \ \\<^sub>s \" using t''s''_p by blast ultimately have "is_Fun_Set (the_set_term (insert\t',s'\))" by (auto simp add: is_Fun_Set_subst) hence "is_Fun_Set s" by (simp add: t's'_p(3) is_Fun_Set_subst) thus ?case using is_Fun_Set_exi by auto qed (auto simp add: step db\<^sub>s\<^sub>s\<^sub>t_def) qed auto lemma reachable_constraints_occurs_fact_ik_ground: assumes \_reach: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction T" and t: "occurs t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" shows "fv (occurs t) = {}" proof - have 0: "admissible_transaction T" when "T \ set P" for T using P that unfolding list_all_iff by simp have 1: "wellformed_transaction T" when "T \ set P" for T using 0[OF that] unfolding admissible_transaction_def by simp have 2: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) = (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \)" when "T \ set P" for T \ and A::"('fun,'atom,'sets,'lbl) prot_constr" using dual_transaction_ik_is_transaction_send'[OF 1[OF that]] by fastforce have 3: "admissible_transaction_occurs_checks T" when "T \ set P" for T using 0[OF that] unfolding admissible_transaction_def by simp show ?thesis using \_reach t proof (induction A rule: reachable_constraints.induct) case (step A T \ \) thus ?case proof (cases "occurs t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A") case False hence "occurs t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \" using 2[OF step.hyps(2)] step.prems by blast hence "send\occurs t\ \ set (unlabel (transaction_send T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \))" using wellformed_transaction_send_receive_subst_trm_cases(2)[OF 1[OF step.hyps(2)]] by blast then obtain s where s: "send\occurs s\ \ set (unlabel (transaction_send T))" "t = s \ \ \\<^sub>s \" using transaction_fresh_subst_transaction_renaming_subst_occurs_fact_send_receive(1)[ OF step.hyps(3,4) 1[OF step.hyps(2)]] transaction_strand_subst_subsets(10) by blast obtain x where x: "x \ set (transaction_fresh T)" "s = Var x" using s(1) 3[OF step.hyps(2)] unfolding admissible_transaction_occurs_checks_def by fastforce have "fv t = {}" using transaction_fresh_subst_transaction_renaming_subst_range(1)[OF step.hyps(3,4) x(1)] s(2) x(2) 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" shows "\s \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I). OccursFact \ \(funs_term ` set (snd (Ana s)))" (is "?A A") and "\s \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I). OccursSec \ \(funs_term ` set (snd (Ana s)))" (is "?B A") and "Fun OccursSec [] \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" (is "?C A") and "\x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. I x \ Fun OccursSec []" (is "?D A") proof - have T_adm: "admissible_transaction T" when "T \ set P" for T using P that unfolding list_all_iff by simp have T_valid: "wellformed_transaction T" when "T \ set P" for T using T_adm[OF that] unfolding admissible_transaction_def by blast have T_occ: "admissible_transaction_occurs_checks T" when "T \ set P" for T using T_adm[OF that] unfolding admissible_transaction_def by blast have \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" by (metis \ welltyped_constraint_model_def) have \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range I)" by (metis \ welltyped_constraint_model_def constraint_model_def) have \_grounds: "fv (I x) = {}" "\f T. I x = Fun f T" for x using \ interpretation_grounds[of I, of "Var x"] empty_fv_exists_fun[of "I x"] unfolding welltyped_constraint_model_def constraint_model_def by auto have 00: "fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \ vars_transaction T" "fv\<^sub>s\<^sub>e\<^sub>t (subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))) = fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" for T::"('fun,'atom,'sets,'lbl) prot_transaction" using fv_trms\<^sub>s\<^sub>s\<^sub>t_subset(1)[of "unlabel (transaction_send T)"] vars_transaction_unfold fv_subterms_set[of "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)"] by blast+ have 0: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). \a. \ (Var x) = TAtom a" "\x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). \ (Var x) \ TAtom OccursSecType" "\x \ fv\<^sub>s\<^sub>e\<^sub>t (subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))). \a. \ (Var x) = TAtom a" "\x \ fv\<^sub>s\<^sub>e\<^sub>t (subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))). \ (Var x) \ TAtom OccursSecType" "\x \ vars_transaction T. \a. \ (Var x) = TAtom a" "\x \ vars_transaction T. \ (Var x) \ TAtom OccursSecType" when "T \ set P" for T using admissible_transaction_occurs_fv_types[OF T_adm[OF that]] 00 by blast+ have 1: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) \\<^sub>s\<^sub>e\<^sub>t I = (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I) \ (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t I)" when "T \ set P" for T \ and A::"('fun,'atom,'sets,'lbl) prot_constr" using dual_transaction_ik_is_transaction_send'[OF T_valid[OF that]] by fastforce have 2: "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t I) = subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t I" when "T \ set P" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" for T \ using wt_subst_TAtom_subterms_set_subst[OF wt_subst_compose[OF \(1) \_wt] 0(1)[OF that(1)]] wf_trm_subst_rangeD[OF wf_trms_subst_compose[OF \(2) \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s]] by auto have 3: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" when "T \ set P" "transaction_fresh_subst \ T A" "transaction_renaming_subst \ P 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_fresh_subst_transaction_renaming_wt[OF that(2,3)] transaction_fresh_subst_range_wf_trms[OF that(2)] transaction_renaming_subst_range_wf_trms[OF that(3)] wf_trms_subst_compose by simp_all have 4: "\s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). OccursFact \ \(funs_term ` set (snd (Ana s))) \ OccursSec \ \(funs_term ` set (snd (Ana s)))" when T: "T \ set P" for T proof fix t assume t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" then obtain s where s: "send\s\ \ set (unlabel (transaction_send T))" "t \ subterms s" using wellformed_transaction_unlabel_cases(5)[OF T_valid[OF T]] by fastforce have s_occ: "\x. s = occurs (Var x)" when "OccursFact \ funs_term t \ OccursSec \ funs_term t" proof - have "OccursFact \ funs_term s \ OccursSec \ funs_term s" using that subtermeq_imp_funs_term_subset[OF s(2)] by blast thus ?thesis using s T_occ[OF T] unfolding admissible_transaction_occurs_checks_def by fastforce qed obtain K T' where K: "Ana t = (K,T')" by moura show "OccursFact \ \(funs_term ` set (snd (Ana t))) \ OccursSec \ \(funs_term ` set (snd (Ana t)))" proof (rule ccontr) assume "\(OccursFact \ \(funs_term ` set (snd (Ana t))) \ OccursSec \ \(funs_term ` set (snd (Ana t))))" hence a: "OccursFact \ \(funs_term ` set (snd (Ana t))) \ OccursSec \ \(funs_term ` set (snd (Ana t)))" by simp hence "OccursFact \ \(funs_term ` set T') \ OccursSec \ \(funs_term ` set T')" using K by simp hence "OccursFact \ funs_term t \ OccursSec \ funs_term t" using Ana_subterm[OF K] funs_term_subterms_eq(1)[of t] by blast then obtain x where x: "t \ subterms (occurs (Var x))" using s(2) s_occ by blast thus False using a by fastforce qed qed have 5: "OccursFact \ \(funs_term ` subst_range (\ \\<^sub>s \))" "OccursSec \ \(funs_term ` subst_range (\ \\<^sub>s \))" when \\: "transaction_fresh_subst \ T A" "transaction_renaming_subst \ P A" for \ \ and T::"('fun,'atom,'sets,'lbl) prot_transaction" and A::"('fun,'atom,'sets,'lbl) prot_constr" proof - have "OccursFact \ funs_term t" "OccursSec \ funs_term t" when "t \ subst_range (\ \\<^sub>s \)" for t using transaction_fresh_subst_transaction_renaming_subst_range'[OF \\ that] by auto thus "OccursFact \ \(funs_term ` subst_range (\ \\<^sub>s \))" "OccursSec \ \(funs_term ` subst_range (\ \\<^sub>s \))" by blast+ qed have 6: "I x \ Fun OccursSec []" "\t. I x = occurs t" "\a. \ (I x) = TAtom a \ a \ OccursSecType" when T: "T \ set P" and \\: "transaction_fresh_subst \ T A" "transaction_renaming_subst \ P A" and x: "Var x \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \" for x \ \ and T::"('fun,'atom,'sets,'lbl) prot_transaction" and A::"('fun,'atom,'sets,'lbl) prot_constr" proof - obtain t where t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "t \ (\ \\<^sub>s \) = Var x" using x by moura then obtain y where y: "t = Var y" by (cases t) auto have "\a. \ t = TAtom a \ a \ OccursSecType" using 0(1,2)[OF T] t(1) y by force thus "\a. \ (I x) = TAtom a \ a \ OccursSecType" using wt_subst_trm''[OF 3(1)[OF T \\]] wt_subst_trm''[OF \_wt] t(2) by (metis subst_apply_term.simps(1)) 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_fresh_subst \ T A" "transaction_renaming_subst \ P A" and x: "x \ fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \) ` vars_transaction T)" for x \ \ and T::"('fun,'atom,'sets,'lbl) prot_transaction" and A::"('fun,'atom,'sets,'lbl) prot_constr" proof - obtain y where y: "y \ vars_transaction T" "x \ fv ((\ \\<^sub>s \) y)" using x by auto hence y': "(\ \\<^sub>s \) y = Var x" using transaction_fresh_subst_transaction_renaming_subst_range'[OF \\] by (cases "(\ \\<^sub>s \) y \ subst_range (\ \\<^sub>s \)") force+ have "\a. \ (Var y) = TAtom a \ a \ OccursSecType" using 0(5,6)[OF T] y by force thus "\a. \ (I x) = TAtom a \ a \ OccursSecType" using wt_subst_trm''[OF 3(1)[OF T \\]] wt_subst_trm''[OF \_wt] y' by (metis subst_apply_term.simps(1)) 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_fresh_subst \ T A" "transaction_renaming_subst \ P A" and x: "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \" for x \ \ and T::"('fun,'atom,'sets,'lbl) prot_transaction" and A::"('fun,'atom,'sets,'lbl) prot_constr" proof - obtain t where t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" "t \ (\ \\<^sub>s \) = Var x" using x by moura then obtain y where y: "t = Var y" by (cases t) auto have "\a. \ t = TAtom a \ a \ OccursSecType" using 0(3,4)[OF T] t(1) y by force thus "\a. \ (I x) = TAtom a \ a \ OccursSecType" using wt_subst_trm''[OF 3(1)[OF T \\]] wt_subst_trm''[OF \_wt] t(2) by (metis subst_apply_term.simps(1)) thus "I x \ Fun OccursSec []" "\t. I x = occurs t" by auto qed have s_fv: "fv s \ fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \) ` vars_transaction T)" when s: "s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \" and T: "T \ set P" for s and \ \::"('fun,'atom,'sets) prot_subst" and T::"('fun,'atom,'sets,'lbl) prot_transaction" proof fix x assume "x \ fv s" hence "x \ fv\<^sub>s\<^sub>e\<^sub>t (subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \)" using s by auto hence *: "x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \)" using fv_subterms_set_subst' by fast have **: "list_all is_Send (unlabel (transaction_send T))" using T_valid[OF T] unfolding wellformed_transaction_def by blast have "x \ fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \) ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" proof - obtain t where t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "x \ fv (t \ \ \\<^sub>s \)" using * by fastforce hence "fv t \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" using fv_trms\<^sub>s\<^sub>s\<^sub>t_subset(1)[of "unlabel (transaction_send T)"] by auto thus ?thesis using t(2) subst_apply_fv_subset by fast qed thus "x \ fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \) ` vars_transaction T)" using vars_transaction_unfold[of T] by fastforce qed show "?A A" using \_reach proof (induction A rule: reachable_constraints.induct) case (step A T \ \) have *: "\s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)). OccursFact \ \(funs_term ` set (snd (Ana s)))" using 4[OF step.hyps(2)] by blast have "\s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s\<^sub>e\<^sub>t I. OccursFact \ \(funs_term ` set (snd (Ana s)))" proof fix t assume t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s\<^sub>e\<^sub>t I" then obtain s u where su: "s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \" "s \ I = t" "u \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" "u \ \ \\<^sub>s \ = s" by force obtain Ku Tu where KTu: "Ana u = (Ku,Tu)" by moura have *: "OccursFact \ \(funs_term ` set Tu)" "OccursFact \ \(funs_term ` subst_range (\ \\<^sub>s \))" "OccursFact \ \(funs_term ` \(((set \ snd \ Ana) ` subst_range (\ \\<^sub>s \))))" using transaction_fresh_subst_transaction_renaming_subst_range'[OF step.hyps(3,4)] 4[OF step.hyps(2)] su(3) KTu by fastforce+ have "OccursFact \ \(funs_term ` set (Tu \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \))" proof - { fix f assume f: "f \ \(funs_term ` set (Tu \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \))" then obtain tf where tf: "tf \ set Tu" "f \ funs_term (tf \ \ \\<^sub>s \)" by moura hence "f \ funs_term tf \ f \ \(funs_term ` subst_range (\ \\<^sub>s \))" using funs_term_subst[of tf "\ \\<^sub>s \"] by force hence "f \ OccursFact" using *(1,2) tf(1) by blast } thus ?thesis by metis qed hence **: "OccursFact \ \(funs_term ` set (snd (Ana s)))" proof (cases u) case (Var xu) hence "s = (\ \\<^sub>s \) xu" using su(4) by (metis subst_apply_term.simps(1)) thus ?thesis using *(3) by fastforce qed (use su(4) KTu Ana_subst'[of _ _ Ku Tu "\ \\<^sub>s \"] in simp) show "OccursFact \ \(funs_term ` set (snd (Ana t)))" proof (cases s) case (Var sx) then obtain a where a: "\ (I sx) = Var a" using su(1) 8(3)[OF step.hyps(2,3,4), of sx] by fast hence "Ana (I sx) = ([],[])" by (metis \_grounds(2) const_type_inv[THEN Ana_const]) thus ?thesis using Var su(2) by simp next case (Fun f S) hence snd_Ana_t: "snd (Ana t) = snd (Ana s) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t I" using su(2) Ana_subst'[of f S _ "snd (Ana s)" I] by (cases "Ana s") simp_all { fix g assume "g \ \(funs_term ` set (snd (Ana t)))" hence "g \ \(funs_term ` set (snd (Ana s))) \ (\x \ fv\<^sub>s\<^sub>e\<^sub>t (set (snd (Ana s))). g \ funs_term (I x))" using snd_Ana_t funs_term_subst[of _ I] by auto hence "g \ OccursFact" proof assume "\x \ fv\<^sub>s\<^sub>e\<^sub>t (set (snd (Ana s))). g \ funs_term (I x)" then obtain x where x: "x \ fv\<^sub>s\<^sub>e\<^sub>t (set (snd (Ana s)))" "g \ funs_term (I x)" by moura have "x \ fv s" using x(1) Ana_vars(2)[of s] by (cases "Ana s") auto hence "x \ fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \) ` 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,3,4)] const_type_inv by metis hence "h \ OccursFact" by auto moreover have "U = []" using h(1,2,4) const_type_inv_wf[of h U a] \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s by fastforce ultimately show ?thesis using h(1) x(2) by auto qed (use ** in blast) } thus ?thesis by blast qed qed thus ?case using step.IH step.prems 1[OF step.hyps(2), of A "\ \\<^sub>s \"] 2[OF step.hyps(2) 3[OF step.hyps(2,3,4)]] by auto qed simp show "?B A" using \_reach proof (induction A rule: reachable_constraints.induct) case (step A T \ \) have "\s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s\<^sub>e\<^sub>t I. OccursSec \ \(funs_term ` set (snd (Ana s)))" proof fix t assume t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s\<^sub>e\<^sub>t I" then obtain s u where su: "s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \" "s \ I = t" "u \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" "u \ \ \\<^sub>s \ = s" by force obtain Ku Tu where KTu: "Ana u = (Ku,Tu)" by moura have *: "OccursSec \ \(funs_term ` set Tu)" "OccursSec \ \(funs_term ` subst_range (\ \\<^sub>s \))" "OccursSec \ \(funs_term ` \(((set \ snd \ Ana) ` subst_range (\ \\<^sub>s \))))" using transaction_fresh_subst_transaction_renaming_subst_range'[OF step.hyps(3,4)] 4[OF step.hyps(2)] su(3) KTu by fastforce+ have "OccursSec \ \(funs_term ` set (Tu \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \))" proof - { fix f assume f: "f \ \(funs_term ` set (Tu \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \\<^sub>s \))" then obtain tf where tf: "tf \ set Tu" "f \ funs_term (tf \ \ \\<^sub>s \)" by moura hence "f \ funs_term tf \ f \ \(funs_term ` subst_range (\ \\<^sub>s \))" using funs_term_subst[of tf "\ \\<^sub>s \"] by force hence "f \ OccursSec" using *(1,2) tf(1) by blast } thus ?thesis by metis qed hence **: "OccursSec \ \(funs_term ` set (snd (Ana s)))" proof (cases u) case (Var xu) hence "s = (\ \\<^sub>s \) xu" using su(4) by (metis subst_apply_term.simps(1)) thus ?thesis using *(3) by fastforce qed (use su(4) KTu Ana_subst'[of _ _ Ku Tu "\ \\<^sub>s \"] in simp) show "OccursSec \ \(funs_term ` set (snd (Ana t)))" proof (cases s) case (Var sx) then obtain a where a: "\ (I sx) = Var a" using su(1) 8(3)[OF step.hyps(2,3,4), of sx] by fast hence "Ana (I sx) = ([],[])" by (metis \_grounds(2) const_type_inv[THEN Ana_const]) thus ?thesis using Var su(2) by simp next case (Fun f S) hence snd_Ana_t: "snd (Ana t) = snd (Ana s) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t I" using su(2) Ana_subst'[of f S _ "snd (Ana s)" I] by (cases "Ana s") simp_all { fix g assume "g \ \(funs_term ` set (snd (Ana t)))" hence "g \ \(funs_term ` set (snd (Ana s))) \ (\x \ fv\<^sub>s\<^sub>e\<^sub>t (set (snd (Ana s))). g \ funs_term (I x))" using snd_Ana_t funs_term_subst[of _ I] by auto hence "g \ OccursSec" proof assume "\x \ fv\<^sub>s\<^sub>e\<^sub>t (set (snd (Ana s))). g \ funs_term (I x)" then obtain x where x: "x \ fv\<^sub>s\<^sub>e\<^sub>t (set (snd (Ana s)))" "g \ funs_term (I x)" by moura have "x \ fv s" using x(1) Ana_vars(2)[of s] by (cases "Ana s") auto hence "x \ fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \) ` 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,3,4)] const_type_inv by metis hence "h \ OccursSec" by auto moreover have "U = []" using h(1,2,4) const_type_inv_wf[of h U a] \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s by fastforce ultimately show ?thesis using h(1) x(2) by auto qed (use ** in blast) } thus ?thesis by blast qed qed thus ?case using step.IH step.prems 1[OF step.hyps(2), of A "\ \\<^sub>s \"] 2[OF step.hyps(2) 3[OF step.hyps(2,3,4)]] by auto qed simp show "?C A" using \_reach proof (induction A rule: reachable_constraints.induct) case (step A T \ \) have *: "Fun OccursSec [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" using wellformed_transaction_unlabel_cases(5)[OF T_valid[OF step.hyps(2)]] T_occ[OF step.hyps(2)] unfolding admissible_transaction_occurs_checks_def by fastforce have **: "Fun OccursSec [] \ subst_range (\ \\<^sub>s \)" using transaction_fresh_subst_transaction_renaming_subst_range'[OF step.hyps(3,4)] by auto have "Fun OccursSec [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s\<^sub>e\<^sub>t I" proof assume "Fun OccursSec [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \ \\<^sub>s\<^sub>e\<^sub>t I" then obtain s where "s \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \" "s \ I = Fun OccursSec []" by moura moreover have "Fun OccursSec [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \" proof assume "Fun OccursSec [] \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \" then obtain u where "u \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)" "u \ \ \\<^sub>s \ = Fun OccursSec []" by moura thus False using * ** by (cases u) (force simp del: subst_subst_compose)+ qed ultimately show False using 6[OF step.hyps(2,3,4)] by (cases s) auto qed thus ?case using step.IH step.prems 1[OF step.hyps(2), of A "\ \\<^sub>s \"] by fast qed simp show "?D A" using \_reach proof (induction A rule: reachable_constraints.induct) case (step A T \ \) { fix x assume x: "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \))" hence x': "x \ vars\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T) \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" by (metis vars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unlabel_subst) hence "x \ vars_transaction T \ x \ fv\<^sub>s\<^sub>e\<^sub>t ((\ \\<^sub>s \) ` vars_transaction T)" using vars\<^sub>s\<^sub>s\<^sub>t_subst_cases[OF x'] by metis moreover have "I x \ Fun OccursSec []" when "x \ vars_transaction T" using that 0(5,6)[OF step.hyps(2)] wt_subst_trm''[OF \_wt, of "Var x"] by fastforce ultimately have "I x \ Fun OccursSec []" using 7(1)[OF step.hyps(2,3,4), 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 t: "t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "t \ I = occurs s" shows "\u. t = occurs u" proof - have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" using \ unfolding welltyped_constraint_model_def constraint_model_def by metis hence 0: "\ t = \ (occurs s)" using t(2) wt_subst_trm'' by metis have 1: "\\<^sub>v ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ (\T \ set P. \\<^sub>v ` fv_transaction T)" "\T \ set P. \x \ fv_transaction T. \\<^sub>v x = TAtom Value \ (\a. \\<^sub>v x = TAtom (Atom a))" using reachable_constraints_TAtom_types(1)[OF \_reach] protocol_transaction_vars_TAtom_typed(2,3) P by fast+ show ?thesis proof (cases t) case (Var x) thus ?thesis using 0 1 t(1) var_subterm_ik\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t[of x "unlabel A"] by fastforce next case (Fun f T) hence 2: "f = OccursFact" "length T = Suc (Suc 0)" "T ! 0 \ I = Fun OccursSec []" using t(2) by auto have "T ! 0 = Fun OccursSec []" proof (cases "T ! 0") case (Var y) hence "I y = Fun OccursSec []" using Fun 2(3) by simp moreover have "Var y \ set T" using Var 2(2) length_Suc_conv[of T 1] by auto hence "y \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" using Fun t(1) by force hence "y \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using fv_ik_subset_fv_sst'[of "unlabel A"] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel A"] by blast ultimately have False using reachable_constraints_occurs_fact_ik_funs_terms(4)[OF \_reach \ P] by blast thus ?thesis by simp qed (use 2(3) in simp) moreover have "\u u'. T = [u,u']" - using 2(2) by (metis (no_types) length_0_conv length_Suc_conv) + using 2(2) by (metis (no_types, lifting) Suc_length_conv length_0_conv) 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 t: "occurs t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" shows "occurs t \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" proof - have \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t I" using \ unfolding welltyped_constraint_model_def constraint_model_def by metis obtain s where s: "s \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "s \ I = occurs t" using t by auto hence u: "\u. s = occurs u" using \_wt reachable_constraints_occurs_fact_ik_subst_aux[OF \_reach \ P] by blast hence "fv s = {}" using reachable_constraints_occurs_fact_ik_ground[OF \_reach P] 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 x: "send\occurs (Var x)\ \ set (unlabel A)" shows "occurs (I x) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using \_reach \ x proof (induction A rule: reachable_constraints.induct) case (step A T \ \) define \ where "\ \ \ \\<^sub>s \" define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" have T_adm: "admissible_transaction T" using P step.hyps(2) unfolding list_all_iff by blast have T_valid: "wellformed_transaction T" using T_adm unfolding admissible_transaction_def by blast have T_adm_occ: "admissible_transaction_occurs_checks T" using T_adm unfolding admissible_transaction_def by blast have \_is_T_model: "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I) (set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I)) (unlabel T') I" using step.prems unlabel_append[of A T'] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of "unlabel A" I "[]"] strand_sem_append_stateful[of "{}" "{}" "unlabel A" "unlabel T'" I] by (simp add: T'_def \_def welltyped_constraint_model_def constraint_model_def db\<^sub>s\<^sub>s\<^sub>t_def) show ?case proof (cases "send\occurs (Var x)\ \ set (unlabel A)") case False hence "send\occurs (Var x)\ \ set (unlabel T')" using step.prems(2) unfolding T'_def \_def by simp hence "receive\occurs (Var x)\ \ set (unlabel (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_unlabel_steps_iff(2) unfolding T'_def by blast then obtain y where y: "receive\occurs (Var y)\ \ set (unlabel (transaction_receive T))" "\ y = Var x" using transaction_fresh_subst_transaction_renaming_subst_occurs_fact_send_receive(2)[ OF step.hyps(3,4) T_valid] subst_to_var_is_var[of _ \ x] unfolding \_def by (force simp del: subst_subst_compose) hence "receive\occurs (Var y) \ \\ \ set (unlabel (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" using subst_lsst_unlabel_member[of "receive\occurs (Var y)\" "transaction_receive T" \] by fastforce hence "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ occurs (Var y) \ \ \ I" using wellformed_transaction_sem_receives[ OF T_valid, of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I" "set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I)" \ I "occurs (Var y) \ \"] \_is_T_model by (metis T'_def) hence *: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t I \ occurs (\ y \ I)" by auto have "occurs (\ y \ I) \ ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using deduct_occurs_in_ik[OF *] reachable_constraints_occurs_fact_ik_subst[ OF step.hyps(1) welltyped_constraint_model_prefix[OF step.prems(1)] P, of "\ y \ I"] reachable_constraints_occurs_fact_ik_funs_terms[ OF step.hyps(1) welltyped_constraint_model_prefix[OF step.prems(1)] P] by blast thus ?thesis using y(2) by simp qed (simp add: step.IH[OF welltyped_constraint_model_prefix[OF step.prems(1)]]) qed simp lemma reachable_contraints_fv_bvars_subset: assumes A: "A \ reachable_constraints P" shows "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ (\T \ set P. bvars_transaction T)" using assms proof (induction A rule: reachable_constraints.induct) case (step \ T \ \) let ?T' = "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \" show ?case using step.IH step.hyps(2) bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of ?T'] bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst[of "transaction_strand T" "\ \\<^sub>s \"] bvars\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?T')"] unlabel_append[of \ "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?T'"] by (metis (no_types, lifting) SUP_upper Un_subset_iff) qed simp lemma reachable_contraints_fv_disj: assumes A: "A \ reachable_constraints P" shows "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ (\T \ set P. bvars_transaction T) = {}" using A proof (induction A rule: reachable_constraints.induct) case (step \ T \ \) define T' where "T' \ transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \" define X where "X \ \T \ set P. bvars_transaction T" have "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ X = {}" using transaction_fresh_subst_transaction_renaming_subst_vars_disj(4)[OF step.hyps(3,4)] transaction_fresh_subst_transaction_renaming_subst_vars_subset(4)[OF step.hyps(3,4,2)] unfolding T'_def X_def by blast hence "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T') \ X = {}" using step.IH[unfolded X_def[symmetric]] fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of T'] by auto thus ?case unfolding T'_def X_def by blast qed simp lemma reachable_contraints_fv_bvars_disj: assumes P: "\T \ set P. wellformed_transaction T" and A: "A \ reachable_constraints P" shows "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = {}" using A proof (induction A rule: reachable_constraints.induct) case (step \ T \ \) define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" note 0 = transaction_fresh_subst_transaction_renaming_subst_vars_disj[OF step.hyps(3,4)] note 1 = transaction_fresh_subst_transaction_renaming_subst_vars_subset[OF step.hyps(3,4)] have 2: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = {}" using 0(7) 1(4)[OF step.hyps(2)] fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unfolding T'_def by (metis (no_types) disjoint_iff_not_equal subset_iff) have "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ \(bvars_transaction ` set P)" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ \(bvars_transaction ` set P) = {}" using reachable_contraints_fv_bvars_subset[OF reachable_constraints.step[OF step.hyps]] reachable_contraints_fv_disj[OF reachable_constraints.step[OF step.hyps]] unfolding T'_def by auto hence 3: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = {}" by blast have "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \) \ bvars_transaction T = {}" using 0(4)[OF step.hyps(2)] 1(4)[OF step.hyps(2)] by blast hence 4: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' = {}" by (metis (no_types) T'_def fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unlabel_subst bvars\<^sub>s\<^sub>s\<^sub>t_subst) have "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T') = {}" using 2 3 4 step.IH unfolding unlabel_append[of \ T'] fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel T'"] bvars\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel T'"] by fast thus ?case unfolding T'_def by blast qed simp lemma reachable_constraints_wf: assumes P: "\T \ set P. wellformed_transaction T" "\T \ set P. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms_transaction T)" and A: "A \ reachable_constraints P" shows "wf\<^sub>s\<^sub>s\<^sub>t (unlabel A)" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" proof - have "wellformed_transaction T" when "T \ set P" for T using P(1) that by fast+ hence 0: "wf'\<^sub>s\<^sub>s\<^sub>t (set (transaction_fresh T)) (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)))" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)) = {}" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms_transaction T)" when T: "T \ set P" for T unfolding admissible_transaction_terms_def by (metis T wellformed_transaction_wf\<^sub>s\<^sub>s\<^sub>t(1), metis T wellformed_transaction_wf\<^sub>s\<^sub>s\<^sub>t(2) fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq, metis T wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code P(2)) from A have "wf\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" proof (induction A rule: reachable_constraints.induct) case (step A T \ \) let ?T' = "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" have IH: "wf'\<^sub>s\<^sub>s\<^sub>t {} (unlabel A)" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = {}" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" using step.IH by metis+ have 1: "wf'\<^sub>s\<^sub>s\<^sub>t {} (unlabel (A@?T'))" using protocol_transaction_wf_subst[OF 0(1)[OF step.hyps(2)] step.hyps(3,4)] wf\<^sub>s\<^sub>s\<^sub>t_vars_mono[of "{}"] wf\<^sub>s\<^sub>s\<^sub>t_append[OF IH(1)] by simp have 2: "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@?T') \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@?T') = {}" using reachable_contraints_fv_bvars_disj[OF P(1)] reachable_constraints.step[OF step.hyps] by blast have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t ?T')" using trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unlabel_subst wf_trms_subst[ OF wf_trms_subst_compose[ OF transaction_fresh_subst_range_wf_trms[OF step.hyps(3)] transaction_renaming_subst_range_wf_trms[OF step.hyps(4)]], THEN wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_trms\<^sub>s\<^sub>s\<^sub>t_subst, OF 0(3)[OF step.hyps(2)]] by metis hence 3: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@?T'))" using IH(3) by auto show ?case using 1 2 3 by force qed simp thus "wf\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" by metis+ qed lemma reachable_constraints_no_Ana_Attack: assumes \: "\ \ reachable_constraints P" and P: "\T \ set P. admissible_transaction T" and t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" shows "attack\n\ \ set (snd (Ana t))" proof - have T_adm: "admissible_transaction T" when "T \ set P" for T using P that by blast have T_adm_term: "admissible_transaction_terms T" when "T \ set P" for T using T_adm[OF that] unfolding admissible_transaction_def by blast have T_valid: "wellformed_transaction T" when "T \ set P" for T using T_adm[OF that] unfolding admissible_transaction_def by blast show ?thesis using \ t proof (induction \ rule: reachable_constraints.induct) case (step A T \ \) thus ?case proof (cases "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)") case False hence "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)))" using step.prems by simp hence "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \)" using dual_transaction_ik_is_transaction_send'[OF T_valid[OF step.hyps(2)]] by metis hence "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \" using transaction_fresh_subst_transaction_renaming_subst_trms[ OF step.hyps(3,4), of "transaction_send T"] wellformed_transaction_unlabel_cases(5)[OF T_valid[OF step.hyps(2)]] by fastforce then obtain s where s: "s \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_send T))" "t = s \ \ \\<^sub>s \" by moura hence s': "attack\n\ \ set (snd (Ana s))" using admissible_transaction_no_Ana_Attack[OF T_adm_term[OF step.hyps(2)]] trms_transaction_unfold[of T] by blast note * = transaction_fresh_subst_transaction_renaming_subst_range'[OF step.hyps(3,4)] show ?thesis proof assume n: "attack\n\ \ set (snd (Ana t))" thus False proof (cases s) case (Var x) thus ?thesis using Var * n s(2) by (force simp del: subst_subst_compose) next case (Fun f T) hence "attack\n\ \ set (snd (Ana s)) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>s \" using Ana_subst'[of f T _ "snd (Ana s)" "\ \\<^sub>s \"] s(2) s' n by (cases "Ana s") auto hence "attack\n\ \ set (snd (Ana s)) \ attack\n\ \ subst_range (\ \\<^sub>s \)" using const_mem_subst_cases' by fast thus ?thesis using * s' by blast qed qed qed simp qed simp 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 x: "\\<^sub>v x = TAtom Value" "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" shows "\n. I x = Fun (Val (n,False)) []" using reachable_constraints_occurs_fact_send_ex[OF \_reach P x] reachable_constraints_occurs_fact_send_in_ik[OF \_reach \ P] reachable_constraints_occurs_fact_ik_case[OF \_reach P] 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 x: "(TAtom Value, m) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" shows "\n. I (TAtom Value, m) = Fun (Val (n,False)) []" using constraint_model_Value_term_is_Val[OF \_reach \ P _ 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" shows "\x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \\<^sub>v x = TAtom Value \ (\B. prefix B \ \ x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \ \ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B))" (is "?P \") using \_reach \ proof (induction \ rule: reachable_constraints.induct) case (step \ T \ \) have IH: "?P \" using step welltyped_constraint_model_prefix by fast define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" have T_adm: "admissible_transaction T" by (metis P step.hyps(2)) have T_valid: "wellformed_transaction T" by (metis T_adm admissible_transaction_def) have \_is_T_model: "strand_sem_stateful (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) (set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)) (unlabel T') \" using step.prems unlabel_append[of \ T'] db\<^sub>s\<^sub>s\<^sub>t_set_is_dbupd\<^sub>s\<^sub>s\<^sub>t[of "unlabel \" \ "[]"] strand_sem_append_stateful[of "{}" "{}" "unlabel \" "unlabel T'" \] by (simp add: T'_def welltyped_constraint_model_def constraint_model_def db\<^sub>s\<^sub>s\<^sub>t_def) have \_interp: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (metis \ welltyped_constraint_model_def constraint_model_def, metis \ welltyped_constraint_model_def, metis \ welltyped_constraint_model_def constraint_model_def) have 1: "\B. prefix B \ \ x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \ \ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" when x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" "\\<^sub>v x = TAtom Value" for x proof - obtain n where n: "\ x = Fun n []" "is_Val n \ is_Abs n" "\public n" using constraint_model_Value_term_is_Val[ OF reachable_constraints.step[OF step.hyps] step.prems P x(2)] x(1) fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \" "unlabel T'"] unlabel_append[of \ T'] unfolding T'_def by moura have "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" using x(1) fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unfolding T'_def by fastforce then obtain y where y: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)" "x \ fv ((\ \\<^sub>s \) y)" using fv\<^sub>s\<^sub>s\<^sub>t_subst_obtain_var[of x "unlabel (transaction_strand T)" "\ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \"] by auto have y_x: "(\ \\<^sub>s \) y = Var x" using y(2) transaction_fresh_subst_transaction_renaming_subst_range[OF step.hyps(3,4), of y] by force have "\ ((\ \\<^sub>s \) y) = TAtom Value" using x(2) y_x by simp moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" using protocol_transaction_vars_TAtom_typed(3) P(1) step.hyps(2) transaction_fresh_subst_transaction_renaming_wt[OF step.hyps(3,4)] by fast ultimately have y_val: "\\<^sub>v y = TAtom Value" by (metis wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def \.simps(1)) have y_not_fresh: "y \ set (transaction_fresh T)" using y(2) transaction_fresh_subst_transaction_renaming_subst_range(1)[OF step.hyps(3,4)] by fastforce have y_n: "Fun n [] = (\ \\<^sub>s \) y \ \" using n y_x by simp hence y_n': "Fun n [] = (\ \\<^sub>s \ \\<^sub>s \) y" by (metis subst_subst_compose[of "Var y" "\ \\<^sub>s \" \] subst_apply_term.simps(1)) have "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T) \ y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T)" using wellformed_transaction_fv_in_receives_or_selects[OF T_valid] y(1) y_not_fresh by blast hence n_cases: "Fun n [] \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ (\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \\<^sub>v z = TAtom Value \ \ z = Fun n [])" proof assume y_in: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_receive T)" then obtain t where t: "receive\t\ \ set (unlabel (transaction_receive T))" "y \ fv t" using admissible_transaction_strand_step_cases(1)[OF T_adm] by force hence "receive\t \ \ \\<^sub>s \\ \ set (unlabel (transaction_receive T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \))" using subst_lsst_unlabel_member[of "receive\t\" "transaction_receive T" "\ \\<^sub>s \"] by fastforce hence *: "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \ \\<^sub>s \ \ \" using wellformed_transaction_sem_receives[ OF T_valid, of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" "set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" "\ \\<^sub>s \" \ "t \ \ \\<^sub>s \"] \_is_T_model by (metis T'_def) have "\a. \ (\ x) = Var a" when "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" for x using that reachable_constraints_vars_TAtom_typed[OF step.hyps(1) P, of x] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] wt_subst_trm''[OF \_wt, of "Var x"] by force hence "\f. \ x = Fun f []" when "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" for x using that wf_trm_subst[OF \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s, of "Var x"] wf_trm_Var[of x] const_type_inv_wf empty_fv_exists_fun[OF interpretation_grounds[OF \_interp], of "Var x"] by (metis subst_apply_term.simps(1)[of x \]) hence \_ik_\_vals: "\x \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \). \f. \ x = Fun f []" using fv_ik_subset_fv_sst'[of "unlabel \"] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] by blast hence "subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \) = subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \" using ik\<^sub>s\<^sub>s\<^sub>t_subst[of "unlabel \" \] unlabel_subst[of \ \] subterms_subst_lsst_ik[of \ \] by metis moreover have "v \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" when "v \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" for v by (meson contra_subsetD fv_ik_subset_fv_sst' that) moreover have "Fun n [] \ subterms (t \ \ \\<^sub>s \ \ \)" using imageI[of "Var y" "subterms t" "\x. x \ \ \\<^sub>s \ \\<^sub>s \"] var_is_subterm[OF t(2)] subterms_subst_subset[of "\ \\<^sub>s \ \\<^sub>s \" t] subst_subst_compose[of t "\ \\<^sub>s \" \] y_n' by (auto simp del: subst_subst_compose) hence "Fun n [] \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \)" using private_fun_deduct_in_ik[OF *, of n "[]"] n(2,3) unfolding is_Val_def is_Abs_def by auto hence "Fun n [] \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ (\z \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \). Fun n [] \ subterms (\ z))" using const_subterm_subst_cases[of n _ \] by auto hence "Fun n [] \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ (\z \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \). \ z = Fun n [])" using \_ik_\_vals by fastforce hence "Fun n [] \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ (\z \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \). \\<^sub>v z = TAtom Value \ \ z = Fun n [])" using \_wt n(2) unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def is_Val_def is_Abs_def by force ultimately show ?thesis using ik\<^sub>s\<^sub>s\<^sub>t_trms\<^sub>s\<^sub>s\<^sub>t_subset[of "unlabel \"] by fast next assume y_in: "y \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_selects T)" then obtain s where s: "select\Var y,Fun (Set s) []\ \ set (unlabel (transaction_selects T))" using admissible_transaction_strand_step_cases(2)[OF T_adm] by force hence "select\(\ \\<^sub>s \) y, Fun (Set s) []\ \ set (unlabel (transaction_selects T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \))" using subst_lsst_unlabel_member by fastforce hence n_in_db: "(Fun n [], Fun (Set s) []) \ set (db'\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ [])" using wellformed_transaction_sem_selects[ OF T_valid, of "ik\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s\<^sub>e\<^sub>t \" "set (db\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \)" "\ \\<^sub>s \" \ "(\ \\<^sub>s \) y" "Fun (Set s) []"] \_is_T_model n y_x unfolding T'_def db\<^sub>s\<^sub>s\<^sub>t_def by fastforce obtain tn sn where tsn: "insert\tn,sn\ \ set (unlabel \)" "Fun n [] = tn \ \" using db\<^sub>s\<^sub>s\<^sub>t_in_cases[OF n_in_db] by force have "Fun n [] = tn \ (\z. \\<^sub>v z = TAtom Value \ tn = Var z)" using \_wt tsn(2) n(2) unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def is_Val_def is_Abs_def by (cases tn) auto moreover have "tn \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "fv tn \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" using tsn(1) in_subterms_Union by force+ ultimately show ?thesis using tsn(2) by auto qed have x_nin_\: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" proof - have "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" using x(1) fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq unfolding T'_def by fast hence "x \ fv\<^sub>s\<^sub>s\<^sub>t ((unlabel (transaction_strand T) \\<^sub>s\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>s\<^sub>t \)" using transaction_fresh_subst_grounds_domain[OF step.hyps(3)] step.hyps(3) labeled_stateful_strand_subst_comp[of \ "transaction_strand T" \] unlabel_subst[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" \] unlabel_subst[of "transaction_strand T" \] by (simp add: transaction_fresh_subst_def range_vars_alt_def) then obtain y where y: "\ y = Var x" using transaction_renaming_subst_var_obtain[OF _ step.hyps(4)] by blast thus ?thesis using transaction_renaming_subst_range_notin_vars[OF step.hyps(4), of y] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] by auto qed from n_cases show ?thesis proof assume "\z \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \\<^sub>v z = TAtom Value \ \ z = Fun n []" then obtain B where B: "prefix B \" "Fun n [] \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" by (metis IH n(1)) thus ?thesis using n x_nin_\ trms\<^sub>s\<^sub>s\<^sub>t_unlabel_prefix_subset(1)[of B] by (metis (no_types, hide_lams) self_append_conv subset_iff subterms\<^sub>s\<^sub>e\<^sub>t_mono prefix_def) qed (use n x_nin_\ in fastforce) qed have "?P (\@T')" proof (intro ballI impI) fix x assume x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\@T')" "\\<^sub>v x = TAtom Value" show "\B. prefix B (\@T') \ x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t B \ \ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" proof (cases "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \") case False hence x': "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" using x(1) unlabel_append[of \] fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \"] by simp then obtain B where B: "prefix B \" "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t B" "\ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" using x(2) 1 by moura thus ?thesis using prefix_prefix by fast qed (use x(2) IH prefix_prefix in fast) qed thus ?case unfolding T'_def by blast qed simp lemma admissible_transaction_occurs_checks_prop: assumes \_reach: "\ \ reachable_constraints P" and \: "welltyped_constraint_model \ \" and P: "\T \ set P. admissible_transaction T" and f: "f \ \(funs_term ` (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" shows "is_Val f \ \public f" and "\is_Abs f" proof - obtain x where x: "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \" "f \ funs_term (\ x)" using f by moura obtain T where T: "Fun f T \ \ x" using funs_term_Fun_subterm[OF x(2)] by moura have \_interp: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (metis \ welltyped_constraint_model_def constraint_model_def, metis \ welltyped_constraint_model_def, metis \ welltyped_constraint_model_def constraint_model_def) have 1: "\ (Var x) = \ (\ x)" using wt_subst_trm''[OF \_wt, of "Var x"] by simp hence "\a. \ (\ x) = Var a" using x(1) reachable_constraints_vars_TAtom_typed[OF \_reach P, of x] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel \"] by force hence "\f. \ x = Fun f []" using x(1) wf_trm_subst[OF \_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s, of "Var x"] wf_trm_Var[of x] const_type_inv_wf empty_fv_exists_fun[OF interpretation_grounds[OF \_interp], of "Var x"] by (metis subst_apply_term.simps(1)[of x \]) hence 2: "\ x = Fun f []" using x(2) by force have "(is_Val f \ \public f) \ \is_Abs f" proof (cases "\\<^sub>v x = TAtom Value") case True then obtain B where B: "prefix B \" "x \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t B" "\ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t B)" using constraint_model_Value_var_in_constr_prefix[OF \_reach \ P] x(1) by fast have "\ x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using B(1,3) trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel B"] unlabel_append[of B] unfolding prefix_def by auto hence "f \ \(funs_term ` trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using x(2) funs_term_subterms_eq(2)[of "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] by blast thus ?thesis using reachable_constraints_val_funs_private[OF \_reach P] by blast+ next case False thus ?thesis using x 1 2 by (cases f) auto qed thus "is_Val f \ \public 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 f: "f \ \(funs_term ` (\ ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))" shows "\n. f = Val (n,True)" and "\n. f = Abs n" using admissible_transaction_occurs_checks_prop[OF \_reach \ P f] by auto lemma transaction_var_becomes_Val: assumes \_reach: "\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \) \ reachable_constraints P" and \: "welltyped_constraint_model \ (\@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \))" and \: "transaction_fresh_subst \ T \" and \: "transaction_renaming_subst \ P \" and P: "\T \ set P. admissible_transaction T" and T: "T \ set P" and x: "x \ fv_transaction T" "fst x = TAtom Value" shows "\n. Fun (Val (n,False)) [] = (\ \\<^sub>s \) x \ \" proof - obtain m where m: "x = (TAtom Value, m)" by (metis x(2) eq_fst_iff) have x_not_bvar: "x \ bvars_transaction T" "fv ((\ \\<^sub>s \) x) \ bvars_transaction T = {}" using x(1) transactions_fv_bvars_disj[OF P] T transaction_fresh_subst_transaction_renaming_subst_vars_disj(2)[OF \ \, of x] vars\<^sub>s\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>s\<^sub>t[of "unlabel (transaction_strand T)"] by blast+ show ?thesis proof (cases "x \ subst_domain \") case True then obtain n where "\ x = Fun (Val (n, False)) []" using \ unfolding transaction_fresh_subst_def by fastforce thus ?thesis using subst_compose[of \ \ x] by simp next case False hence "\ x = Var x" by auto then obtain n where n: "(\ \\<^sub>s \) x = Var (TAtom Value, n)" using m transaction_renaming_subst_is_renaming[OF \] subst_compose[of \ \ x] by force hence "(TAtom Value, n) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" using x_not_bvar fv\<^sub>s\<^sub>s\<^sub>t_subst_fv_subset[OF x(1), of "\ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \"] by force hence "\n'. \ (TAtom Value, n) = Fun (Val (n',False)) []" using constraint_model_Value_term_is_Val'[OF \_reach \ P, of n] x fv\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \"] fv\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel \"] unlabel_append[of \] by fastforce thus ?thesis using n by simp qed qed lemma reachable_constraints_SMP_subset: assumes \: "\ \ reachable_constraints P" and P: "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" shows "SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ SMP (\T \ set P. trms_transaction T)" (is "?A \") and "SMP (pair`setops\<^sub>s\<^sub>s\<^sub>t (unlabel \)) \ SMP (\T\set P. pair`setops_transaction T)" (is "?B \") proof - have "?A \ \ ?B \" using \ proof (induction \ rule: reachable_constraints.induct) case (step A T \ \) define T' where "T' \ transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \" define M where "M \ \T \ set P. trms_transaction T" define N where "N \ \T \ set P. pair ` setops_transaction T" let ?P = "\t. \s \. s \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = s \ \" let ?Q = "\t. \s \. s \ N \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = s \ \" have IH: "SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ SMP M" "SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)) \ SMP N" using step.IH by (metis M_def, metis N_def) have \\_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" using P(1) step.hyps(2) transaction_fresh_subst_transaction_renaming_wt[OF step.hyps(3,4)] by fast have \\_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using transaction_fresh_subst_range_wf_trms[OF step.hyps(3)] transaction_renaming_subst_range_wf_trms[OF step.hyps(4)] by (metis wf_trms_subst_compose) have 0: "SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')) = SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" "SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'))) = SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)) \ SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel T'))" using trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of T'] setops\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of T'] trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel A" "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')"] setops\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel A" "unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')"] unlabel_append[of A "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'"] image_Un[of pair "setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "setops\<^sub>s\<^sub>s\<^sub>t (unlabel T')"] SMP_union[of "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'"] SMP_union[of "pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel T')"] by argo+ have 1: "SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T') \ SMP M" proof (intro SMP_subset_I ballI) fix t show "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T' \ ?P t" using trms\<^sub>s\<^sub>s\<^sub>t_wt_subst_ex[OF \\_wt \\_wf, of t "unlabel (transaction_strand T)"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \"] step.hyps(2) unfolding T'_def M_def by auto qed have 2: "SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel T')) \ SMP N" proof (intro SMP_subset_I ballI) fix t show "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel T') \ ?Q t" using setops\<^sub>s\<^sub>s\<^sub>t_wt_subst_ex[OF \\_wt \\_wf, of t "unlabel (transaction_strand T)"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \"] step.hyps(2) unfolding T'_def N_def by auto qed have "SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')) \ SMP M" "SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'))) \ SMP N" using 0 1 2 IH by blast+ thus ?case unfolding T'_def M_def N_def by blast qed (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) thus "?A \" "?B \" by metis+ qed lemma reachable_constraints_no_Pair_fun: assumes A: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction T" shows "Pair \ \(funs_term ` SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A))" using A proof (induction A rule: reachable_constraints.induct) case (step A T \ \) define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" have T_adm: "admissible_transaction T" using step.hyps(2) P unfolding list_all_iff by blast have \\_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" using protocol_transaction_vars_TAtom_typed(3) P(1) step.hyps(2) transaction_fresh_subst_transaction_renaming_wt[OF step.hyps(3,4)] by fast have \\_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using transaction_fresh_subst_range_wf_trms[OF step.hyps(3)] transaction_renaming_subst_range_wf_trms[OF step.hyps(4)] by (metis wf_trms_subst_compose) have 0: "SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T')) = SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) \ SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" using SMP_union[of "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'"] unlabel_append[of A T'] trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel A" "unlabel T'"] by simp have 1: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" using reachable_constraints_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s[OF _ reachable_constraints.step[OF step.hyps]] admissible_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P trms\<^sub>s\<^sub>s\<^sub>t_append[of "unlabel A"] unlabel_append[of A] unfolding T'_def by force have 2: "Pair \ \(funs_term ` (subst_range (\ \\<^sub>s \)))" using transaction_fresh_subst_transaction_renaming_subst_range'[OF step.hyps(3,4)] by force have "Pair \ \(funs_term ` (trms_transaction T))" using T_adm unfolding admissible_transaction_def admissible_transaction_terms_def by blast hence "Pair \ funs_term t" when t: "t \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T) \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" for t using 2 trms\<^sub>s\<^sub>s\<^sub>t_funs_term_cases[OF t] by force hence 3: "Pair \ funs_term t" when t: "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" for t using t unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \"] trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \"] unfolding T'_def by metis have "\a. \\<^sub>v x = TAtom a" when "x \ vars_transaction T" for x using that protocol_transaction_vars_TAtom_typed(1) P step.hyps(2) by fast hence "\a. \\<^sub>v x = TAtom a" when "x \ vars\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T) \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" for x using wt_subst_fv\<^sub>s\<^sub>e\<^sub>t_termtype_subterm[OF _ \\_wt \\_wf, of x "vars_transaction T"] vars\<^sub>s\<^sub>s\<^sub>t_subst_cases[OF that] by fastforce hence "\a. \\<^sub>v x = TAtom a" when "x \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t T'" for x using that unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \"] vars\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \"] unfolding T'_def by simp hence "\a. \\<^sub>v x = TAtom a" when "x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" for x using that fv_trms\<^sub>s\<^sub>s\<^sub>t_subset(1) by fast hence "Pair \ funs_term (\ (Var x))" when "x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" for x using that by fastforce moreover have "Pair \ funs_term s" when s: "Ana s = (K, M)" "Pair \ \(funs_term ` set K)" for s::"('fun,'atom,'sets) prot_term" and K M proof (cases s) case (Fun f S) thus ?thesis using s Ana_Fu_keys_not_pairs[of _ S K M] by (cases f) force+ qed (use s in simp) ultimately have "Pair \ funs_term t" when t: "t \ SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t T')" for t using t 3 SMP_funs_term[OF t _ _ 1, of Pair] funs_term_type_iff by fastforce thus ?case using 0 step.IH(1) unfolding T'_def by blast qed simp lemma reachable_constraints_setops_form: assumes A: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction T" and t: "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" shows "\c s. t = pair (c, Fun (Set s) []) \ \ c = TAtom Value" using A t proof (induction A rule: reachable_constraints.induct) case (step A T \ \) have T_adm: "admissible_transaction T" when "T \ set P" for T using P that unfolding list_all_iff by simp have T_adm': "admissible_transaction_selects T" "admissible_transaction_checks T" "admissible_transaction_updates T" when "T \ set P" for T using T_adm[OF that] unfolding admissible_transaction_def by simp_all have T_valid: "wellformed_transaction T" when "T \ set P" for T using T_adm[OF that] unfolding admissible_transaction_def by blast have \\_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" using protocol_transaction_vars_TAtom_typed(3) P(1) step.hyps(2) transaction_fresh_subst_transaction_renaming_wt[OF step.hyps(3,4)] by fast have \\_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using transaction_fresh_subst_range_wf_trms[OF step.hyps(3)] transaction_renaming_subst_range_wf_trms[OF step.hyps(4)] by (metis wf_trms_subst_compose) show ?case using step.IH proof (cases "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)") case False hence "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T) \\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" using step.prems setops\<^sub>s\<^sub>s\<^sub>t_append unlabel_append setops\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \"] by fastforce then obtain t' \ where t': "t' \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (transaction_strand T))" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t = t' \ \" using setops\<^sub>s\<^sub>s\<^sub>t_wt_subst_ex[OF \\_wt \\_wf] by blast then obtain s s' where s: "t' = pair (s,s')" using setops\<^sub>s\<^sub>s\<^sub>t_are_pairs by fastforce moreover have "InSet ac s s' = InSet Assign s s' \ InSet ac s s' = InSet Check s s'" for ac by (cases ac) simp_all ultimately have "\n. s = Var (Var Value, n)" "\u. s' = Fun (Set u) []" using t'(1) setops\<^sub>s\<^sub>s\<^sub>t_member_iff[of s s' "unlabel (transaction_strand T)"] pair_in_pair_image_iff[of s s'] transaction_inserts_are_Value_vars[ OF T_valid[OF step.hyps(2)] T_adm'(3)[OF step.hyps(2)], of s s'] transaction_deletes_are_Value_vars[ OF T_valid[OF step.hyps(2)] T_adm'(3)[OF step.hyps(2)], of s s'] transaction_selects_are_Value_vars[ OF T_valid[OF step.hyps(2)] T_adm'(1)[OF step.hyps(2)], of s s'] transaction_inset_checks_are_Value_vars[ OF T_valid[OF step.hyps(2)] T_adm'(2)[OF step.hyps(2)], of s s'] transaction_notinset_checks_are_Value_vars[ OF T_valid[OF step.hyps(2)] T_adm'(2)[OF step.hyps(2)], of _ _ _ s s'] by metis+ then obtain ss n where ss: "t = pair (\ (Var Value, n), Fun (Set ss) [])" using t'(4) s unfolding pair_def by force have "\ (\ (Var Value, n)) = TAtom Value" "wf\<^sub>t\<^sub>r\<^sub>m (\ (Var Value, n))" using t'(2) wt_subst_trm''[OF t'(2), of "Var (Var Value, n)"] apply simp using t'(3) by (cases "(Var Value, n) \ subst_domain \") auto thus ?thesis using ss by blast qed simp qed (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) lemma reachable_constraints_setops_type: fixes t::"('fun,'atom,'sets) prot_term" assumes A: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction T" and t: "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" shows "\ t = TComp Pair [TAtom Value, TAtom SetType]" proof - obtain s c where s: "t = pair (c, Fun (Set s) [])" "\ c = TAtom Value" using reachable_constraints_setops_form[OF A P t] by moura hence "(Fun (Set s) []::('fun,'atom,'sets) prot_term) \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" using t setops\<^sub>s\<^sub>s\<^sub>t_member_iff[of c "Fun (Set s) []" "unlabel A"] by force hence "wf\<^sub>t\<^sub>r\<^sub>m (Fun (Set s) []::('fun,'atom,'sets) prot_term)" using reachable_constraints_wf(2) P A unfolding admissible_transaction_def admissible_transaction_terms_def by blast hence "arity (Set s) = 0" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by simp thus ?thesis using s unfolding pair_def by fastforce qed lemma reachable_constraints_setops_same_type_if_unifiable: assumes A: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction T" shows "\s \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A). \t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A). (\\. Unifier \ s t) \ \ s = \ t" (is "?P A") using reachable_constraints_setops_type[OF A P] by simp lemma reachable_constraints_setops_unfiable_if_wt_instance_unifiable: assumes A: "A \ reachable_constraints P" and P: "\T \ set P. admissible_transaction T" shows "\s \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A). \t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A). (\\ \ \. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ Unifier \ (s \ \) (t \ \)) \ (\\. Unifier \ s t)" proof (intro ballI impI) fix s t assume st: "s \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" and "\\ \ \. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ Unifier \ (s \ \) (t \ \)" then obtain \ \ \ where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "Unifier \ (s \ \) (t \ \)" by moura obtain fs ft cs ct where c: "s = pair (cs, Fun (Set fs) [])" "t = pair (ct, Fun (Set ft) [])" "\ cs = TAtom Value" "\ ct = TAtom Value" using reachable_constraints_setops_form[OF A P st(1)] reachable_constraints_setops_form[OF A P st(2)] by moura have "cs \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "ct \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" using c(1,2) setops_subterm_trms[OF st(1), of cs] setops_subterm_trms[OF st(2), of ct] Fun_param_is_subterm[of cs "args s"] Fun_param_is_subterm[of ct "args t"] unfolding pair_def by simp_all moreover have "\T \ set P. wellformed_transaction T" "\T \ set P. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms_transaction T)" using P unfolding admissible_transaction_def admissible_transaction_terms_def by fast+ ultimately have *: "wf\<^sub>t\<^sub>r\<^sub>m cs" "wf\<^sub>t\<^sub>r\<^sub>m ct" using reachable_constraints_wf(2)[OF _ _ A] wf_trms_subterms by blast+ have "(\x. cs = Var x) \ (\c d. cs = Fun c [])" using const_type_inv_wf c(3) *(1) by (cases cs) auto moreover have "(\x. ct = Var x) \ (\c d. ct = Fun c [])" using const_type_inv_wf c(4) *(2) by (cases ct) auto ultimately show "\\. Unifier \ s t" using reachable_constraints_setops_form[OF A P] reachable_constraints_setops_type[OF A P] st \ c unfolding pair_def by auto qed lemma reachable_constraints_tfr: assumes M: "M \ \T \ set P. trms_transaction T" "has_all_wt_instances_of \ M N" "finite N" "tfr\<^sub>s\<^sub>e\<^sub>t N" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and P: "\T \ set P. admissible_transaction T" "\T \ set P. list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" and \: "\ \ reachable_constraints P" shows "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel \)" using \ proof (induction \ rule: reachable_constraints.induct) case (step A T \ \) define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" have P': "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" "\T \ set P. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms_transaction T)" using P(1) protocol_transaction_vars_TAtom_typed(3) admissible_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s by blast+ have AT'_reach: "A@T' \ reachable_constraints P" using reachable_constraints.step[OF step.hyps] unfolding T'_def by metis have \\_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" using P'(1) step.hyps(2) transaction_fresh_subst_transaction_renaming_wt[OF step.hyps(3,4)] by fast have \\_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using transaction_fresh_subst_range_wf_trms[OF step.hyps(3)] transaction_renaming_subst_range_wf_trms[OF step.hyps(4)] by (metis wf_trms_subst_compose) have \\_bvars_disj: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) \ range_vars (\ \\<^sub>s \) = {}" by (rule transaction_fresh_subst_transaction_renaming_subst_vars_disj(4)[OF step.hyps(3,4,2)]) have wf_trms_M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" using admissible_transactions_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s P(1) unfolding M(1) by blast have "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T'))" using reachable_constraints_SMP_subset(1)[OF AT'_reach P'(1)] tfr_subset(3)[OF M(4), of "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T')"] SMP_SMP_subset[of M N] SMP_I'[OF wf_trms_M M(5,2)] unfolding M(1) by blast moreover have "\p. Ana (pair p) = ([],[])" unfolding pair_def by auto ultimately have 1: "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T') \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@T')))" using tfr_setops_if_tfr_trms[of "unlabel (A@T')"] reachable_constraints_no_Pair_fun[OF AT'_reach P(1)] reachable_constraints_setops_same_type_if_unifiable[OF AT'_reach P(1)] reachable_constraints_setops_unfiable_if_wt_instance_unifiable[OF AT'_reach P(1)] by blast have "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" using step.hyps(2) P(2) tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p unfolding comp_tfr\<^sub>s\<^sub>s\<^sub>t_def tfr\<^sub>s\<^sub>s\<^sub>t_def by fastforce hence "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel T')" using tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_all_wt_subst_apply[OF _ \\_wt \\_wf \\_bvars_disj] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \"] unfolding T'_def by argo hence 2: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (A@T'))" using step.IH unlabel_append unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by auto have "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel (A@T'))" using 1 2 by (metis tfr\<^sub>s\<^sub>s\<^sub>t_def) thus ?case by (metis T'_def) qed simp lemma reachable_constraints_tfr': assumes M: "M \ \T \ set P. trms_transaction T \ pair' Pair ` setops_transaction T" "has_all_wt_instances_of \ M N" "finite N" "tfr\<^sub>s\<^sub>e\<^sub>t N" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and P: "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" "\T \ set P. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms_transaction T)" "\T \ set P. list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" and \: "\ \ reachable_constraints P" shows "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel \)" using \ proof (induction \ rule: reachable_constraints.induct) case (step A T \ \) define T' where "T' \ dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \)" have AT'_reach: "A@T' \ reachable_constraints P" using reachable_constraints.step[OF step.hyps] unfolding T'_def by metis have \\_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" using P(1) step.hyps(2) transaction_fresh_subst_transaction_renaming_wt[OF step.hyps(3,4)] by fast have \\_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using transaction_fresh_subst_range_wf_trms[OF step.hyps(3)] transaction_renaming_subst_range_wf_trms[OF step.hyps(4)] by (metis wf_trms_subst_compose) have \\_bvars_disj: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T) \ range_vars (\ \\<^sub>s \) = {}" by (rule transaction_fresh_subst_transaction_renaming_subst_vars_disj(4)[OF step.hyps(3,4,2)]) have wf_trms_M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" using P(2) setops\<^sub>s\<^sub>s\<^sub>t_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s(2) unfolding M(1) pair_code wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[symmetric] by fast have "SMP (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T')) \ SMP M" "SMP (pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@T'))) \ SMP M" using reachable_constraints_SMP_subset[OF AT'_reach P(1)] SMP_mono[of "\T \ set P. trms_transaction T" M] SMP_mono[of "\T \ set P. pair ` setops_transaction T" M] unfolding M(1) pair_code[symmetric] by blast+ hence 1: "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T') \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@T')))" using tfr_subset(3)[OF M(4), of "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T') \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@T'))"] SMP_union[of "trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@T')" "pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel (A@T'))"] SMP_SMP_subset[of M N] SMP_I'[OF wf_trms_M M(5,2)] by blast have "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" using step.hyps(2) P(3) tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_is_comp_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p unfolding comp_tfr\<^sub>s\<^sub>s\<^sub>t_def tfr\<^sub>s\<^sub>s\<^sub>t_def by fastforce hence "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel T')" using tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p_all_wt_subst_apply[OF _ \\_wt \\_wf \\_bvars_disj] dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p[of "transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \\<^sub>s \"] unlabel_subst[of "transaction_strand T" "\ \\<^sub>s \"] unfolding T'_def by argo hence 2: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (A@T'))" using step.IH unlabel_append unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by auto have "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel (A@T'))" using 1 2 by (metis tfr\<^sub>s\<^sub>s\<^sub>t_def) thus ?case by (metis T'_def) qed simp lemma reachable_constraints_typing_cond\<^sub>s\<^sub>s\<^sub>t: assumes M: "M \ \T \ set P. trms_transaction T \ pair' Pair ` setops_transaction T" "has_all_wt_instances_of \ M N" "finite N" "tfr\<^sub>s\<^sub>e\<^sub>t N" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and P: "\T \ set P. wellformed_transaction T" "\T \ set P. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms_transaction T)" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" "\T \ set P. list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" and \: "\ \ reachable_constraints P" shows "typing_cond\<^sub>s\<^sub>s\<^sub>t (unlabel \)" using reachable_constraints_wf[OF P(1,2) \] reachable_constraints_tfr'[OF M P(3,2,4) \] unfolding typing_cond\<^sub>s\<^sub>s\<^sub>t_def by blast context begin private lemma reachable_constraints_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_aux: fixes P defines "Ts \ concat (map transaction_strand P)" assumes P_fresh_wf: "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" (is "\T \ set P. ?fresh_wf T") and A: "A \ reachable_constraints P" shows "\b \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A). \a \ set Ts. \\. b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ (\t \ subst_range \. (\x. t = Var x) \ (\c. t = Fun c []))" (is "\b \ set (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A). \a \ set Ts. ?P b a") using A proof (induction A rule: reachable_constraints.induct) case (step \ T \ \) define Q where "Q \ ?P" define \ where "\ \ \ \\<^sub>s \" let ?R = "\A Ts. \b \ set A. \a \ set Ts. Q b a" have T_fresh_wf: "?fresh_wf T" using step.hyps(2) P_fresh_wf by blast have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "\t \ subst_range \. (\x. t = Var x) \ (\c. t = Fun c [])" using wt_subst_compose[ OF transaction_fresh_subst_wt[OF step.hyps(3) T_fresh_wf] transaction_renaming_subst_wt[OF step.hyps(4)]] wf_trms_subst_compose[ OF transaction_fresh_subst_range_wf_trms[OF step.hyps(3)] transaction_renaming_subst_range_wf_trms[OF step.hyps(4)]] transaction_fresh_subst_transaction_renaming_subst_range'[OF step.hyps(3,4)] unfolding \_def by metis+ hence "?R (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T)) \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) (transaction_strand T)" using dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_self_inverse[of "transaction_strand T"] by (auto simp add: Q_def subst_apply_labeled_stateful_strand_def) hence "?R (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) (transaction_strand T)" by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst) hence "?R (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (transaction_strand T \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \))) Ts" using step.hyps(2) unfolding Ts_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by fastforce thus ?case using step.IH unfolding Q_def \_def by auto qed simp lemma reachable_constraints_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t: fixes P defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and "Ts \ concat (map transaction_strand P)" assumes P_pc: "comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ Pair Ts M S" and P_wf: "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" and A: "A \ reachable_constraints P" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A ((f (set S)) - {m. intruder_synth {} m})" using par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_if_comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t'[OF P_pc, of "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A", THEN par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t] reachable_constraints_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_aux[OF P_wf A, unfolded Ts_def[symmetric]] unfolding f_def dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_self_inverse by fast end lemma reachable_constraints_par_comp_constr: fixes P f S defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and "Ts \ concat (map transaction_strand P)" and "Sec \ (f (set S)) - {m. intruder_synth {} m}" and "M \ \T \ set P. trms_transaction T \ pair' Pair ` setops_transaction T" assumes M: "has_all_wt_instances_of \ M N" "finite N" "tfr\<^sub>s\<^sub>e\<^sub>t N" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and P: "\T \ set P. wellformed_transaction T" "\T \ set P. wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity (trms_transaction T)" "\T \ set P. \x \ set (transaction_fresh T). \\<^sub>v x = TAtom Value" "\T \ set P. list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (transaction_strand T))" "comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ Pair Ts M_fun S" and \: "\ \ reachable_constraints P" and \: "constraint_model \ \" shows "\\\<^sub>\. welltyped_constraint_model \\<^sub>\ \ \ ((\n. welltyped_constraint_model \\<^sub>\ (proj n \)) \ (\\'. prefix \' \ \ strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' Sec \\<^sub>\))" proof - have \': "constr_sem_stateful \ (unlabel \)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using \ unfolding constraint_model_def by blast+ show ?thesis using reachable_constraints_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t[OF P(5,3)[unfolded Ts_def] \] reachable_constraints_typing_cond\<^sub>s\<^sub>s\<^sub>t[OF M_def M P(1,2,3,4) \] par_comp_constr_stateful[OF _ _ \', of Sec] unfolding f_def Sec_def welltyped_constraint_model_def constraint_model_def by blast qed end end diff --git a/thys/Smith_Normal_Form/Diagonal_To_Smith_JNF.thy b/thys/Smith_Normal_Form/Diagonal_To_Smith_JNF.thy --- a/thys/Smith_Normal_Form/Diagonal_To_Smith_JNF.thy +++ b/thys/Smith_Normal_Form/Diagonal_To_Smith_JNF.thy @@ -1,657 +1,659 @@ (* Author: Jose Divasón Email: jose.divason@unirioja.es *) section \Algorithm to transform a diagonal matrix into its Smith normal form in JNF\ theory Diagonal_To_Smith_JNF imports Admits_SNF_From_Diagonal_Iff_Bezout_Ring begin text \In this file, we implement an algorithm to transform a diagonal matrix into its Smith normal form, using the JNF library. There are, at least, three possible options: \begin{enumerate} \item Implement and prove the soundness of the algorithm from scratch in JNF \item Implement it in JNF and connect it to the HOL Analysis version by means of transfer rules. Thus, we could obtain the soundness lemma in JNF. \item Implement it in JNF, with calls to the HOL Analysis version by means of the functions @{text " from_hma\<^sub>m"} and @{text "to_hma\<^sub>m"}. That is, transform the matrix to HOL Analysis, apply the existing algorith in HOL Analysis to get the Smith normal form and then transform the output to JNF. Then, we could try to get the soundness theorem in JNF by means of transfer rules and local type definitions. \end{enumerate} The first option requires much effort. As we will see, the third option is not possible. \ subsection \Attempt with the third option: definitions and conditional transfer rules\ context fixes A::"'a::bezout_ring mat" assumes "A \ carrier_mat CARD('nr::mod_type) CARD('nc::mod_type)" begin private definition "diagonal_to_Smith_PQ_JNF' bezout = ( let A' = Mod_Type_Connect.to_hma\<^sub>m A::'a^'nc::mod_type^'nr::mod_type; (P,S,Q) = (diagonal_to_Smith_PQ A' bezout) in (Mod_Type_Connect.from_hma\<^sub>m P, Mod_Type_Connect.from_hma\<^sub>m S, Mod_Type_Connect.from_hma\<^sub>m Q))" end text \This approach will not work. The type is necessary in the definition of the function. That is, outside the context, the function will be: @{text "diagonal_to_Smith_PQ_JNF' TYPE('nc) TYPE('nr) A bezout"} And we cannot get rid of such @{text "TYPE('nc)"}. That is, we could get a lemma like: @{theory_text " lemma assumes A \ carrier_mat m n and (P,S,Q) = diagonal_to_Smith_PQ_JNF' TYPE('nr::mod_type) TYPE('nc::mod_type) A bezout shows invertible_mat P \ invertible_mat Q \ S = P * A * Q \ Smith_normal_form_mat S "} But we wouldn't be able to get rid of such types. \ subsection \Attempt with the second option: implementation and soundness in JNF\ definition "diagonal_step_JNF A i j d v = Matrix.mat (dim_row A) (dim_col A) (\ (a,b). if a = i \ b = i then d else if a = j \ b = j then v * (A $$ (j,j)) else A $$ (a,b))" text \Conditional transfer rules are required, so I prove them within context with assumptions.\ context includes lifting_syntax fixes i and j::nat assumes i: "i 'a :: comm_ring_1 ^ 'nc :: mod_type ^ 'nr :: mod_type \ _) ===> (=) ===> (=) ===> Mod_Type_Connect.HMA_M) (\A. diagonal_step_JNF A i j) (\B. diagonal_step B i j)" by (intro rel_funI, goal_cases, auto simp add: Mod_Type_Connect.HMA_M_def diagonal_step_JNF_def diagonal_step_def) (rule eq_matI, auto simp add: Mod_Type_Connect.from_hma\<^sub>m_def, insert from_nat_eq_imp_eq i j, auto) end definition diagonal_step_PQ_JNF :: "'a::{bezout_ring} mat \ nat \ nat \ 'a bezout \ ('a mat \ ('a mat))" where "diagonal_step_PQ_JNF A i k bezout = (let m = dim_row A; n = dim_col A; (p, q, u, v, d) = bezout (A $$ (i,i)) (A $$ (k,k)); P = addrow (-v) k i (swaprows i k (addrow p k i (1\<^sub>m m))); Q = multcol k (-1) (addcol u k i (addcol q i k (1\<^sub>m n))) in (P,Q) )" context includes lifting_syntax fixes i and k::nat assumes i: "i < min (CARD('nr::mod_type)) (CARD('nc::mod_type))" and k: "k < min (CARD('nr::mod_type)) (CARD('nc::mod_type))" begin lemma HMA_diagonal_step_PQ[transfer_rule]: "((Mod_Type_Connect.HMA_M :: _ \ 'a :: bezout_ring ^ 'nc :: mod_type ^ 'nr :: mod_type \ _) ===> (=) ===> rel_prod Mod_Type_Connect.HMA_M Mod_Type_Connect.HMA_M) (\A bezout. diagonal_step_PQ_JNF A i k bezout) (\A bezout. diagonal_step_PQ A i k bezout)" proof (intro rel_funI, goal_cases) case (1 A A' bezout bezout') note HMA_M_AA'[transfer_rule] = 1(1) let ?d_JNF = "(diagonal_step_PQ_JNF A i k bezout)" let ?d_HA = "(diagonal_step_PQ A' i k bezout)" have [transfer_rule]: "Mod_Type_Connect.HMA_I k (from_nat k::'nc)" and [transfer_rule]: "Mod_Type_Connect.HMA_I k (from_nat k::'nr)" by (metis Mod_Type_Connect.HMA_I_def k min.strict_boundedE to_nat_from_nat_id)+ have [transfer_rule]: "Mod_Type_Connect.HMA_I i (from_nat i::'nc)" and [transfer_rule]: "Mod_Type_Connect.HMA_I i (from_nat i::'nr)" by (metis Mod_Type_Connect.HMA_I_def i min.strict_boundedE to_nat_from_nat_id)+ have [transfer_rule]: "A $$ (i,i) = A' $h from_nat i $h from_nat i" proof - have "A $$ (i,i) = index_hma A' (from_nat i) (from_nat i)" by (transfer, simp) also have "... = A' $h from_nat i $h from_nat i" unfolding index_hma_def by auto finally show ?thesis . qed have [transfer_rule]: "A $$ (k,k) = A' $h from_nat k $h from_nat k" proof - have "A $$ (k,k) = index_hma A' (from_nat k) (from_nat k)" by (transfer, simp) also have "... = A' $h from_nat k $h from_nat k" unfolding index_hma_def by auto finally show ?thesis . qed have dim_row_CARD: "dim_row A = CARD('nr)" using HMA_M_AA' Mod_Type_Connect.dim_row_transfer_rule by blast have dim_col_CARD: "dim_col A = CARD('nc)" using HMA_M_AA' Mod_Type_Connect.dim_col_transfer_rule by blast let ?p = "fst (bezout (A' $h from_nat i $h from_nat i) (A' $h from_nat k $h from_nat k))" let ?v = "fst (snd (snd (snd (bezout (A $$ (i, i)) (A $$ (k, k))))))" have "Mod_Type_Connect.HMA_M (fst ?d_JNF) (fst ?d_HA)" unfolding diagonal_step_PQ_JNF_def diagonal_step_PQ_def Mod_Type_Connect.HMA_M_def unfolding Let_def split_beta dim_row_CARD by (auto, transfer, auto simp add: Mod_Type_Connect.HMA_M_def Rel_def rel_funI) moreover have "Mod_Type_Connect.HMA_M (snd ?d_JNF) (snd ?d_HA)" unfolding diagonal_step_PQ_JNF_def diagonal_step_PQ_def Mod_Type_Connect.HMA_M_def unfolding Let_def split_beta dim_col_CARD by (auto, transfer, auto simp add: Mod_Type_Connect.HMA_M_def Rel_def rel_funI) ultimately show ?case unfolding rel_prod_conv using 1 by (simp add: split_beta) qed end fun diagonal_to_Smith_i_PQ_JNF :: "nat list \ nat \ ('a::{bezout_ring} bezout) \ ('a mat \ 'a mat \ 'a mat) \ ('a mat \ 'a mat \ 'a mat)" where "diagonal_to_Smith_i_PQ_JNF [] i bezout (P,A,Q) = (P,A,Q)" | "diagonal_to_Smith_i_PQ_JNF (j#xs) i bezout (P,A,Q) = ( if A $$ (i,i) dvd A $$ (j,j) then diagonal_to_Smith_i_PQ_JNF xs i bezout (P,A,Q) else let (p, q, u, v, d) = bezout (A $$ (i,i)) (A $$ (j,j)); A' = diagonal_step_JNF A i j d v; (P',Q') = diagonal_step_PQ_JNF A i j bezout in diagonal_to_Smith_i_PQ_JNF xs i bezout (P'*P,A',Q*Q') \ \Apply the step\ ) " context includes lifting_syntax fixes i and xs assumes i: "i < min (CARD('nr::mod_type)) (CARD('nc::mod_type))" and xs: "\j\set xs. j < min (CARD('nr::mod_type)) (CARD('nc::mod_type))" begin declare diagonal_step_PQ.simps[simp del] lemma HMA_diagonal_to_Smith_i_PQ_aux: "HMA_M3 (P,A,Q) (P' :: 'a :: bezout_ring ^ 'nr :: mod_type ^ 'nr :: mod_type, A' :: 'a :: bezout_ring ^ 'nc :: mod_type ^ 'nr :: mod_type, Q' :: 'a :: bezout_ring ^ 'nc :: mod_type ^ 'nc :: mod_type) \ HMA_M3 (diagonal_to_Smith_i_PQ_JNF xs i bezout (P,A,Q)) (diagonal_to_Smith_i_PQ xs i bezout (P',A',Q'))" using i xs proof (induct xs i bezout "(P',A',Q')" arbitrary: P' A' Q' P A Q rule: diagonal_to_Smith_i_PQ.induct) case (1 i bezout P' A' Q') then show ?case by auto next case (2 j xs i bezout P' A' Q') note HMA_M3[transfer_rule] = "2.prems"(1) note i = 2(4) note j = 2(5) note IH1="2.hyps"(1) note IH2="2.hyps"(2) have j_min: "j < min CARD('nr) CARD('nc)" using j by auto have HMA_M_AA'[transfer_rule]: "Mod_Type_Connect.HMA_M A A'" using HMA_M3 by auto have [transfer_rule]: "Mod_Type_Connect.HMA_I j (from_nat j::'nc)" and [transfer_rule]: "Mod_Type_Connect.HMA_I j (from_nat j::'nr)" by (metis Mod_Type_Connect.HMA_I_def j_min min.strict_boundedE to_nat_from_nat_id)+ have [transfer_rule]: "Mod_Type_Connect.HMA_I i (from_nat i::'nc)" and [transfer_rule]: "Mod_Type_Connect.HMA_I i (from_nat i::'nr)" by (metis Mod_Type_Connect.HMA_I_def i min.strict_boundedE to_nat_from_nat_id)+ have [transfer_rule]: "A $$ (i, i) = A' $h from_nat i $h from_nat i" proof - have "A $$ (i,i) = index_hma A' (from_nat i) (from_nat i)" by (transfer, simp) also have "... = A' $h from_nat i $h from_nat i" unfolding index_hma_def by auto finally show ?thesis . qed have [transfer_rule]: "A $$ (j, j) = A' $h from_nat j $h from_nat j" proof - have "A $$ (j,j) = index_hma A' (from_nat j) (from_nat j)" by (transfer, simp) also have "... = A' $h from_nat j $h from_nat j" unfolding index_hma_def by auto finally show ?thesis . qed show ?case proof (cases "A $$ (i, i) dvd A $$ (j, j)") case True hence "A' $h from_nat i $h from_nat i dvd A' $h from_nat j $h from_nat j" by transfer then show ?thesis using True IH1 HMA_M3 i j by auto next case False obtain p q u v d where b: "(p, q, u, v, d) = bezout (A $$ (i,i)) (A $$ (j,j))" by (metis prod_cases5) let ?A'_JNF = "diagonal_step_JNF A i j d v" obtain P''_JNF Q''_JNF where P''Q''_JNF: "(P''_JNF,Q''_JNF) = diagonal_step_PQ_JNF A i j bezout" by (metis surjective_pairing) have not_dvd: "\ A' $h from_nat i $h from_nat i dvd A' $h from_nat j $h from_nat j" using False by transfer let ?A' = "diagonal_step A' i j d v" obtain P'' Q'' where P''Q'': "(P'',Q'') = diagonal_step_PQ A' i j bezout" by (metis surjective_pairing) have b2: "(p, q, u, v, d) = bezout (A' $h from_nat i $h from_nat i) (A' $h from_nat j $h from_nat j)" using b by (transfer,auto) let ?D_HA = "diagonal_to_Smith_i_PQ xs i bezout (P''**P',?A',Q'**Q'')" let ?D_JNF = "diagonal_to_Smith_i_PQ_JNF xs i bezout (P''_JNF*P,?A'_JNF,Q*Q''_JNF)" have rw_1: "diagonal_to_Smith_i_PQ_JNF (j # xs) i bezout (P, A, Q) = ?D_JNF" using False b P''Q''_JNF by (auto, unfold split_beta, metis fst_conv snd_conv) have rw_2: "diagonal_to_Smith_i_PQ (j # xs) i bezout (P', A', Q') = ?D_HA" using not_dvd b2 P''Q'' by (auto, unfold split_beta, metis fst_conv snd_conv) have "HMA_M3 ?D_JNF ?D_HA" proof (rule IH2[OF not_dvd b2], auto) have j: "j < min CARD('nr) CARD('nc)" using j by auto have [transfer_rule]: "rel_prod Mod_Type_Connect.HMA_M Mod_Type_Connect.HMA_M (diagonal_step_PQ_JNF A i j bezout) (diagonal_step_PQ A' i j bezout)" using HMA_diagonal_step_PQ[OF i j] HMA_M_AA' unfolding rel_fun_def by auto hence [transfer_rule]: "Mod_Type_Connect.HMA_M P''_JNF P''" and [transfer_rule]: "Mod_Type_Connect.HMA_M Q''_JNF Q''" using P''Q'' P''Q''_JNF unfolding rel_prod_conv split_beta by (metis fst_conv, metis snd_conv) have [transfer_rule]: "Mod_Type_Connect.HMA_M P P'" using HMA_M3 by auto show "Mod_Type_Connect.HMA_M (P''_JNF * P) (P'' ** P')" (* apply (transfer, auto) does not finish the goal*) by (transfer_prover_start, transfer_step+, auto) (* note HMA_diagonal_step[OF i j,transfer_rule]*) (*transfer does not work for the following goal*) show "Mod_Type_Connect.HMA_M (diagonal_step_JNF A i j d v) (diagonal_step A' i j d v)" using HMA_diagonal_step[OF i j] HMA_M_AA' unfolding rel_fun_def by auto have [transfer_rule]: "Mod_Type_Connect.HMA_M Q Q'" using HMA_M3 by auto show "Mod_Type_Connect.HMA_M (Q * Q''_JNF) (Q' ** Q'')" by (transfer_prover_start, transfer_step+, auto) qed (insert i j P''Q'', auto) then show ?thesis using rw_1 rw_2 by auto qed qed lemma HMA_diagonal_to_Smith_i_PQ[transfer_rule]: "((=) ===> (HMA_M3 :: (_ \ (_\('a :: bezout_ring ^ 'nc :: mod_type ^ 'nr :: mod_type) \ _) \_)) ===> HMA_M3) (diagonal_to_Smith_i_PQ_JNF xs i) (diagonal_to_Smith_i_PQ xs i)" proof (intro rel_funI, goal_cases) case (1 x y bezout bezout') then show ?case using HMA_diagonal_to_Smith_i_PQ_aux by (auto, smt HMA_M3.elims(2)) qed end fun Diagonal_to_Smith_row_i_PQ_JNF where "Diagonal_to_Smith_row_i_PQ_JNF i bezout (P,A,Q) = diagonal_to_Smith_i_PQ_JNF [i + 1.. (HMA_M3 :: (_ \ (_ \ ('a::bezout_ring^'nc::mod_type^'nr::mod_type) \ _) \ _)) ===> HMA_M3) (Diagonal_to_Smith_row_i_PQ_JNF i) (Diagonal_to_Smith_row_i_PQ i)" proof (intro rel_funI, clarify, goal_cases) case (1 _ bezout P A Q P' A' Q') note HMA_M3[transfer_rule] = 1 let ?xs1="[i + 1..j\set ?xs1. j < min CARD('nr) CARD('nc)" using i by (metis atLeastLessThan_iff ncols_def nrows_def set_upt xs_eq) have rel: "HMA_M3 (diagonal_to_Smith_i_PQ_JNF ?xs1 i bezout (P,A,Q)) (diagonal_to_Smith_i_PQ ?xs1 i bezout (P',A',Q'))" using HMA_diagonal_to_Smith_i_PQ[OF i j_xs] HMA_M3 unfolding rel_fun_def by blast then show ?case unfolding Diagonal_to_Smith_row_i_PQ_JNF_def Diagonal_to_Smith_row_i_PQ_def by (metis Suc_eq_plus1 xs_eq) qed end fun diagonal_to_Smith_aux_PQ_JNF where "diagonal_to_Smith_aux_PQ_JNF [] bezout (P,A,Q) = (P,A,Q)" | "diagonal_to_Smith_aux_PQ_JNF (i#xs) bezout (P,A,Q) = diagonal_to_Smith_aux_PQ_JNF xs bezout (Diagonal_to_Smith_row_i_PQ_JNF i bezout (P,A,Q))" context includes lifting_syntax fixes xs assumes xs: "\j\set xs. j < min (CARD('nr::mod_type)) (CARD('nc::mod_type))" begin lemma HMA_diagonal_to_Smith_aux_PQ_JNF[transfer_rule]: "((=) ===> (HMA_M3 :: (_ \ (_ \ ('a::bezout_ring^'nc::mod_type^'nr::mod_type) \ _) \ _)) ===> HMA_M3) (diagonal_to_Smith_aux_PQ_JNF xs) (diagonal_to_Smith_aux_PQ xs)" proof (intro rel_funI, clarify, goal_cases) case (1 _ bezout P A Q P' A' Q') note HMA_M3[transfer_rule] = 1 show ?case using xs HMA_M3 proof (induct xs arbitrary: P' A' Q' P A Q) case Nil then show ?case by auto next case (Cons i xs) note IH = Cons(1) note HMA_M3 = Cons.prems(2) have i: "i < min CARD('nr) CARD('nc)" using Cons.prems by auto let ?D_JNF = "(Diagonal_to_Smith_row_i_PQ_JNF i bezout (P, A, Q))" let ?D_HA = "(Diagonal_to_Smith_row_i_PQ i bezout (P', A', Q'))" have rw_1: "diagonal_to_Smith_aux_PQ_JNF (i # xs) bezout (P, A, Q) = diagonal_to_Smith_aux_PQ_JNF xs bezout ?D_JNF" by auto have rw_2: "diagonal_to_Smith_aux_PQ (i # xs) bezout (P', A', Q') = diagonal_to_Smith_aux_PQ xs bezout ?D_HA" by auto have "HMA_M3 ?D_JNF ?D_HA" using HMA_Diagonal_to_Smith_row_i_PQ[OF i] HMA_M3 unfolding rel_fun_def by blast then show ?case by (auto, smt Cons.hyps HMA_M3.elims(2) list.set_intros(2) local.Cons(2)) qed qed end fun diagonal_to_Smith_PQ_JNF where "diagonal_to_Smith_PQ_JNF A bezout = diagonal_to_Smith_aux_PQ_JNF [0..m (dim_row A),A,1\<^sub>m (dim_col A))" declare diagonal_to_Smith_PQ_JNF.simps[simp del] lemmas diagonal_to_Smith_PQ_JNF_def = diagonal_to_Smith_PQ_JNF.simps lemma diagonal_step_PQ_JNF_dim: assumes A: "A \ carrier_mat m n" and d: "diagonal_step_PQ_JNF A i j bezout = (P,Q)" shows "P \ carrier_mat m m \ Q \ carrier_mat n n" using A d unfolding diagonal_step_PQ_JNF_def split_beta Let_def by auto lemma diagonal_step_JNF_dim: assumes A: "A \ carrier_mat m n" shows "diagonal_step_JNF A i j d v \ carrier_mat m n" using A unfolding diagonal_step_JNF_def by auto lemma diagonal_to_Smith_i_PQ_JNF_dim: assumes "P' \ carrier_mat m m \ A' \ carrier_mat m n \ Q' \ carrier_mat n n" and "diagonal_to_Smith_i_PQ_JNF xs i bezout (P',A',Q') = (P,A,Q)" shows "P \ carrier_mat m m \ A \ carrier_mat m n \ Q \ carrier_mat n n" using assms proof (induct xs i bezout "(P',A',Q')" arbitrary: P A Q P' A' Q' rule: diagonal_to_Smith_i_PQ_JNF.induct) case (1 i bezout P A Q) then show ?case by auto next case (2 j xs i bezout P' A' Q') show ?case proof (cases "A' $$ (i, i) dvd A' $$ (j, j)") case True then show ?thesis using 2 by auto next case False obtain p q u v d where b: "(p, q, u, v, d) = bezout (A' $$ (i,i)) (A' $$ (j,j))" by (metis prod_cases5) let ?A' = "diagonal_step_JNF A' i j d v" obtain P'' Q'' where P''Q'': "(P'',Q'') = diagonal_step_PQ_JNF A' i j bezout" by (metis surjective_pairing) let ?A' = "diagonal_step_JNF A' i j d v" let ?D_JNF = "diagonal_to_Smith_i_PQ_JNF xs i bezout (P''*P',?A',Q'*Q'')" have rw_1: "diagonal_to_Smith_i_PQ_JNF (j # xs) i bezout (P', A', Q') = ?D_JNF" using False b P''Q'' by (auto, unfold split_beta, metis fst_conv snd_conv) show ?thesis proof (rule "2.hyps"(2)[OF False b]) show "?D_JNF = (P,A,Q)" using rw_1 2 by auto have "P'' \ carrier_mat m m" and "Q'' \ carrier_mat n n" using diagonal_step_PQ_JNF_dim[OF _ P''Q''[symmetric]] "2.prems" by auto thus "P'' * P' \ carrier_mat m m \ ?A' \ carrier_mat m n \ Q' * Q'' \ carrier_mat n n" using diagonal_step_JNF_dim 2 by (metis mult_carrier_mat) qed (insert P''Q'', auto) qed qed lemma Diagonal_to_Smith_row_i_PQ_JNF_dim: assumes "P' \ carrier_mat m m \ A' \ carrier_mat m n \ Q' \ carrier_mat n n" and "Diagonal_to_Smith_row_i_PQ_JNF i bezout (P',A',Q') = (P,A,Q)" shows "P \ carrier_mat m m \ A \ carrier_mat m n \ Q \ carrier_mat n n" by (rule diagonal_to_Smith_i_PQ_JNF_dim, insert assms, auto simp add: Diagonal_to_Smith_row_i_PQ_JNF_def) lemma diagonal_to_Smith_aux_PQ_JNF_dim: assumes "P' \ carrier_mat m m \ A' \ carrier_mat m n \ Q' \ carrier_mat n n" and "diagonal_to_Smith_aux_PQ_JNF xs bezout (P',A',Q') = (P,A,Q)" shows "P \ carrier_mat m m \ A \ carrier_mat m n \ Q \ carrier_mat n n" using assms proof (induct xs bezout "(P',A',Q')" arbitrary: P A Q P' A' Q' rule: diagonal_to_Smith_aux_PQ_JNF.induct) case (1 bezout P A Q) then show ?case by simp next case (2 i xs bezout P' A' Q') let ?D="(Diagonal_to_Smith_row_i_PQ_JNF i bezout (P', A', Q'))" have "diagonal_to_Smith_aux_PQ_JNF (i # xs) bezout (P', A', Q') = diagonal_to_Smith_aux_PQ_JNF xs bezout ?D" by auto hence *: "... = (P,A,Q)" using 2 by auto let ?P="fst ?D" let ?S="fst (snd ?D)" let ?Q="snd (snd ?D)" show ?case proof (rule "2.hyps") show "Diagonal_to_Smith_row_i_PQ_JNF i bezout (P', A', Q') = (?P,?S,?Q)" by auto show "diagonal_to_Smith_aux_PQ_JNF xs bezout (?P, ?S, ?Q) = (P, A, Q)" using * by simp show "?P \ carrier_mat m m \ ?S \ carrier_mat m n \ ?Q \ carrier_mat n n" by (rule Diagonal_to_Smith_row_i_PQ_JNF_dim, insert 2, auto) qed qed lemma diagonal_to_Smith_PQ_JNF_dim: assumes "A \ carrier_mat m n" and PSQ: "diagonal_to_Smith_PQ_JNF A bezout = (P,S,Q)" shows "P \ carrier_mat m m \ S \ carrier_mat m n \ Q \ carrier_mat n n" by (rule diagonal_to_Smith_aux_PQ_JNF_dim, insert assms, auto simp add: diagonal_to_Smith_PQ_JNF_def) context includes lifting_syntax begin lemma HMA_diagonal_to_Smith_PQ_JNF[transfer_rule]: "((Mod_Type_Connect.HMA_M) ===> (=) ===> HMA_M3) (diagonal_to_Smith_PQ_JNF) (diagonal_to_Smith_PQ)" proof (intro rel_funI, clarify, goal_cases) case (1 A A' _ bezout) let ?xs1 = "[0..j\set ?xs1. j < min CARD('c) CARD('b)" using dc dr less_imp_diff_less by auto let ?D_JNF = "diagonal_to_Smith_aux_PQ_JNF ?xs1 bezout ?PAQ" let ?D_HA = "diagonal_to_Smith_aux_PQ ?xs1 bezout (mat 1, A', mat 1)" have mat_rel_init: "HMA_M3 ?PAQ (mat 1, A', mat 1)" proof - have "Mod_Type_Connect.HMA_M (1\<^sub>m (dim_row A)) (mat 1::'a^'c::mod_type^'c::mod_type)" unfolding dr by (transfer_prover_start,transfer_step, auto) moreover have "Mod_Type_Connect.HMA_M (1\<^sub>m (dim_col A)) (mat 1::'a^'b::mod_type^'b::mod_type)" unfolding dc by (transfer_prover_start,transfer_step, auto) ultimately show ?thesis using 1 by auto qed have "HMA_M3 ?D_JNF ?D_HA" using HMA_diagonal_to_Smith_aux_PQ_JNF[OF j_xs] mat_rel_init unfolding rel_fun_def by blast then show ?case using xs_eq unfolding diagonal_to_Smith_PQ_JNF_def diagonal_to_Smith_PQ_def by auto qed end subsection \Applying local type definitions\ text \Now we get the soundness lemma in JNF, via the one in HOL Analysis. I need transfer rules and local type definitions.\ context includes lifting_syntax begin private lemma diagonal_to_Smith_PQ_JNF_with_types: assumes A: "A \ carrier_mat CARD('nr::mod_type) CARD('nc::mod_type)" and S: "S \ carrier_mat CARD('nr) CARD('nc)" and P: "P \ carrier_mat CARD('nr) CARD('nr)" and Q: "Q \ carrier_mat CARD('nc) CARD('nc)" and PSQ: "diagonal_to_Smith_PQ_JNF A bezout = (P, S, Q)" and d:"isDiagonal_mat A" and ib: "is_bezout_ext bezout" shows "S = P * A * Q \ invertible_mat P \ invertible_mat Q \ Smith_normal_form_mat S" proof - let ?P = "Mod_Type_Connect.to_hma\<^sub>m P::'a^'nr::mod_type^'nr::mod_type" let ?A = "Mod_Type_Connect.to_hma\<^sub>m A::'a^'nc::mod_type^'nr::mod_type" let ?Q = "Mod_Type_Connect.to_hma\<^sub>m Q::'a^'nc::mod_type^'nc::mod_type" let ?S = "Mod_Type_Connect.to_hma\<^sub>m S::'a^'nc::mod_type^'nr::mod_type" have [transfer_rule]: "Mod_Type_Connect.HMA_M A ?A" by (simp add: Mod_Type_Connect.HMA_M_def A) moreover have [transfer_rule]: "Mod_Type_Connect.HMA_M P ?P" by (simp add: Mod_Type_Connect.HMA_M_def P) moreover have [transfer_rule]: "Mod_Type_Connect.HMA_M Q ?Q" by (simp add: Mod_Type_Connect.HMA_M_def Q) moreover have [transfer_rule]: "Mod_Type_Connect.HMA_M S ?S" by (simp add: Mod_Type_Connect.HMA_M_def S) ultimately have [transfer_rule]: "HMA_M3 (P,S,Q) (?P,?S,?Q)" by simp have [transfer_rule]: "bezout = bezout" .. have PSQ2: "(?P,?S,?Q) = diagonal_to_Smith_PQ ?A bezout" by (transfer, insert PSQ, auto) have "?S = ?P**?A**?Q \ invertible ?P \ invertible ?Q \ Smith_normal_form ?S" by (rule diagonal_to_Smith_PQ'[OF _ ib PSQ2], transfer, auto simp add: d) with this[untransferred] show ?thesis by auto qed private lemma diagonal_to_Smith_PQ_JNF_mod_ring_with_types: assumes A: "A \ carrier_mat CARD('nr::nontriv mod_ring) CARD('nc::nontriv mod_ring)" and S: "S \ carrier_mat CARD('nr mod_ring) CARD('nc mod_ring)" and P: "P \ carrier_mat CARD('nr mod_ring) CARD('nr mod_ring)" and Q: "Q \ carrier_mat CARD('nc mod_ring) CARD('nc mod_ring)" and PSQ: "diagonal_to_Smith_PQ_JNF A bezout = (P, S, Q)" and d:"isDiagonal_mat A" and ib: "is_bezout_ext bezout" shows "S = P * A * Q \ invertible_mat P \ invertible_mat Q \ Smith_normal_form_mat S" by (rule diagonal_to_Smith_PQ_JNF_with_types[OF assms]) (*I don't know how to internalize the sort constraint of 'nr and 'nc at once, so I do it in two steps.*) thm diagonal_to_Smith_PQ_JNF_mod_ring_with_types[unfolded CARD_mod_ring, internalize_sort "'nr::nontriv"] private lemma diagonal_to_Smith_PQ_JNF_internalized_first: "class.nontriv TYPE('a::type) \ A \ carrier_mat CARD('a) CARD('nc::nontriv) \ S \ carrier_mat CARD('a) CARD('nc) \ P \ carrier_mat CARD('a) CARD('a) \ Q \ carrier_mat CARD('nc) CARD('nc) \ diagonal_to_Smith_PQ_JNF A bezout = (P, S, Q) \ isDiagonal_mat A \ is_bezout_ext bezout \ S = P * A * Q \ invertible_mat P \ invertible_mat Q \ Smith_normal_form_mat S" using diagonal_to_Smith_PQ_JNF_mod_ring_with_types[unfolded CARD_mod_ring, internalize_sort "'nr::nontriv"] by blast private lemma diagonal_to_Smith_PQ_JNF_internalized: "class.nontriv TYPE('c::type) \ class.nontriv TYPE('a::type) \ A \ carrier_mat CARD('a) CARD('c) \ S \ carrier_mat CARD('a) CARD('c) \ P \ carrier_mat CARD('a) CARD('a) \ Q \ carrier_mat CARD('c) CARD('c) \ diagonal_to_Smith_PQ_JNF A bezout = (P, S, Q) \ isDiagonal_mat A \ is_bezout_ext bezout \ S = P * A * Q \ invertible_mat P \ invertible_mat Q \ Smith_normal_form_mat S" using diagonal_to_Smith_PQ_JNF_internalized_first[internalize_sort "'nc::nontriv"] by blast context fixes m::nat and n::nat assumes local_typedef1: "\(Rep :: ('b \ int)) Abs. type_definition Rep Abs {0..(Rep :: ('c \ int)) Abs. type_definition Rep Abs {0..1" and n: "n>1" begin lemma type_to_set1: shows "class.nontriv TYPE('b)" (is ?a) and "m=CARD('b)" (is ?b) proof - from local_typedef1 obtain Rep::"('b \ int)" and Abs where t: "type_definition Rep Abs {0.. int)" and Abs where t: "type_definition Rep Abs {0.. carrier_mat m n" assumes PSQ: "(P,S,Q) = diagonal_to_Smith_PQ_JNF A bezout" shows "S = P*A*Q \ invertible_mat P \ invertible_mat Q \ Smith_normal_form_mat S \ P \ carrier_mat m m \ S \ carrier_mat m n \ Q \ carrier_mat n n" proof - have dim_matrices: "P \ carrier_mat m m \ S \ carrier_mat m n \ Q \ carrier_mat n n" by (rule diagonal_to_Smith_PQ_JNF_dim[OF A_dim PSQ[symmetric]]) show ?thesis using diagonal_to_Smith_PQ_JNF_internalized[where ?'c='c, where ?'a='b, OF type_to_set2(1) type_to_set(1), of m A S P Q] unfolding type_to_set1(2)[symmetric] type_to_set2(2)[symmetric] using assms m dim_matrices local_typedef1 by auto qed end end (*Canceling the first local type definitions (I was not able to cancel both in one step)*) context begin private lemma diagonal_to_Smith_PQ_JNF_canceled_first: "\Rep Abs. type_definition Rep Abs {0.. {0.. {} \ 1 < m \ 1 < n \ isDiagonal_mat A \ is_bezout_ext bezout \ A \ carrier_mat m n \ (P, S, Q) = diagonal_to_Smith_PQ_JNF A bezout \ S = P * A * Q \ invertible_mat P \ invertible_mat Q \ Smith_normal_form_mat S \ P \ carrier_mat m m \ S \ carrier_mat m n \ Q \ carrier_mat n n" using diagonal_to_Smith_PQ_JNF_local_typedef[cancel_type_definition] by blast (*Canceling the second*) private lemma diagonal_to_Smith_PQ_JNF_canceled_both: "{0.. {} \ {0.. {} \ 1 < m \ 1 < n \ isDiagonal_mat A \ is_bezout_ext bezout \ A \ carrier_mat m n \ (P, S, Q) = diagonal_to_Smith_PQ_JNF A bezout \ S = P * A * Q \ invertible_mat P \ invertible_mat Q \ Smith_normal_form_mat S \ P \ carrier_mat m m \ S \ carrier_mat m n \ Q \ carrier_mat n n" using diagonal_to_Smith_PQ_JNF_canceled_first[cancel_type_definition] by blast subsection \The final result\ lemma diagonal_to_Smith_PQ_JNF: assumes A: "isDiagonal_mat A" and ib: "is_bezout_ext bezout" and "A \ carrier_mat m n" and PBQ: "(P,S,Q) = diagonal_to_Smith_PQ_JNF A bezout" (*The following two assumptions appear since mod_type requires 11" and m: "m>1" shows "S = P*A*Q \ invertible_mat P \ invertible_mat Q \ Smith_normal_form_mat S \ P \ carrier_mat m m \ S \ carrier_mat m n \ Q \ carrier_mat n n" using diagonal_to_Smith_PQ_JNF_canceled_both[OF _ _ m n] using assms by force end -end \ No newline at end of file +end