(* (C) Copyright Andreas Viktor Hess, DTU, 2015-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: Strands_and_Constraints.thy Author: Andreas Viktor Hess, DTU *) section \Strands and Symbolic Intruder Constraints\ theory Strands_and_Constraints imports Messages More_Unification Intruder_Deduction begin subsection \Constraints, Strands and Related Definitions\ datatype poscheckvariant = Assign ("assign") | Check ("check") text \ A strand (or constraint) step is either a message transmission (either a message being sent \Send\ or being received \Receive\) or a check on messages (a positive check \Equality\---which can be either an "assignment" or just a check---or a negative check \Inequality\) \ datatype (funs\<^sub>s\<^sub>t\<^sub>p: 'a, vars\<^sub>s\<^sub>t\<^sub>p: 'b) strand_step = Send "('a,'b) term" ("send\_\\<^sub>s\<^sub>t" 80) | Receive "('a,'b) term" ("receive\_\\<^sub>s\<^sub>t" 80) | Equality poscheckvariant "('a,'b) term" "('a,'b) term" ("\_: _ \ _\\<^sub>s\<^sub>t" [80,80]) | Inequality (bvars\<^sub>s\<^sub>t\<^sub>p: "'b list") "(('a,'b) term \ ('a,'b) term) list" ("\_\\\: _\\<^sub>s\<^sub>t" [80,80]) where "bvars\<^sub>s\<^sub>t\<^sub>p (Send _) = []" | "bvars\<^sub>s\<^sub>t\<^sub>p (Receive _) = []" | "bvars\<^sub>s\<^sub>t\<^sub>p (Equality _ _ _) = []" text \ A strand is a finite sequence of strand steps (constraints and strands share the same datatype) \ type_synonym ('a,'b) strand = "('a,'b) strand_step list" type_synonym ('a,'b) strands = "('a,'b) strand set" abbreviation "trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ \(t,t') \ set F. {t,t'}" fun trms\<^sub>s\<^sub>t\<^sub>p::"('a,'b) strand_step \ ('a,'b) terms" where "trms\<^sub>s\<^sub>t\<^sub>p (Send t) = {t}" | "trms\<^sub>s\<^sub>t\<^sub>p (Receive t) = {t}" | "trms\<^sub>s\<^sub>t\<^sub>p (Equality _ t t') = {t,t'}" | "trms\<^sub>s\<^sub>t\<^sub>p (Inequality _ F) = trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" lemma vars\<^sub>s\<^sub>t\<^sub>p_unfold[simp]: "vars\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t\<^sub>p x) \ set (bvars\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) auto text \The set of terms occurring in a strand\ definition trms\<^sub>s\<^sub>t where "trms\<^sub>s\<^sub>t S \ \(trms\<^sub>s\<^sub>t\<^sub>p ` set S)" fun trms_list\<^sub>s\<^sub>t\<^sub>p::"('a,'b) strand_step \ ('a,'b) term list" where "trms_list\<^sub>s\<^sub>t\<^sub>p (Send t) = [t]" | "trms_list\<^sub>s\<^sub>t\<^sub>p (Receive t) = [t]" | "trms_list\<^sub>s\<^sub>t\<^sub>p (Equality _ t t') = [t,t']" | "trms_list\<^sub>s\<^sub>t\<^sub>p (Inequality _ F) = concat (map (\(t,t'). [t,t']) F)" text \The set of terms occurring in a strand (list variant)\ definition trms_list\<^sub>s\<^sub>t where "trms_list\<^sub>s\<^sub>t S \ remdups (concat (map trms_list\<^sub>s\<^sub>t\<^sub>p S))" text \The set of variables occurring in a sent message\ definition fv\<^sub>s\<^sub>n\<^sub>d::"('a,'b) strand_step \ 'b set" where "fv\<^sub>s\<^sub>n\<^sub>d x \ case x of Send t \ fv t | _ \ {}" text \The set of variables occurring in a received message\ definition fv\<^sub>r\<^sub>c\<^sub>v::"('a,'b) strand_step \ 'b set" where "fv\<^sub>r\<^sub>c\<^sub>v x \ case x of Receive t \ fv t | _ \ {}" text \The set of variables occurring in an equality constraint\ definition fv\<^sub>e\<^sub>q::"poscheckvariant \ ('a,'b) strand_step \ 'b set" where "fv\<^sub>e\<^sub>q ac x \ case x of Equality ac' s t \ if ac = ac' then fv s \ fv t else {} | _ \ {}" text \The set of variables occurring at the left-hand side of an equality constraint\ definition fv_l\<^sub>e\<^sub>q::"poscheckvariant \ ('a,'b) strand_step \ 'b set" where "fv_l\<^sub>e\<^sub>q ac x \ case x of Equality ac' s t \ if ac = ac' then fv s else {} | _ \ {}" text \The set of variables occurring at the right-hand side of an equality constraint\ definition fv_r\<^sub>e\<^sub>q::"poscheckvariant \ ('a,'b) strand_step \ 'b set" where "fv_r\<^sub>e\<^sub>q ac x \ case x of Equality ac' s t \ if ac = ac' then fv t else {} | _ \ {}" text \The free variables of inequality constraints\ definition fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q::"('a,'b) strand_step \ 'b set" where "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x \ case x of Inequality X F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X | _ \ {}" fun fv\<^sub>s\<^sub>t\<^sub>p::"('a,'b) strand_step \ 'b set" where "fv\<^sub>s\<^sub>t\<^sub>p (Send t) = fv t" | "fv\<^sub>s\<^sub>t\<^sub>p (Receive t) = fv t" | "fv\<^sub>s\<^sub>t\<^sub>p (Equality _ t t') = fv t \ fv t'" | "fv\<^sub>s\<^sub>t\<^sub>p (Inequality X F) = (\(t,t') \ set F. fv t \ fv t') - set X" text \The set of free variables of a strand\ definition fv\<^sub>s\<^sub>t::"('a,'b) strand \ 'b set" where "fv\<^sub>s\<^sub>t S \ \(set (map fv\<^sub>s\<^sub>t\<^sub>p S))" text \The set of bound variables of a strand\ definition bvars\<^sub>s\<^sub>t::"('a,'b) strand \ 'b set" where "bvars\<^sub>s\<^sub>t S \ \(set (map (set \ bvars\<^sub>s\<^sub>t\<^sub>p) S))" text \The set of all variables occurring in a strand\ definition vars\<^sub>s\<^sub>t::"('a,'b) strand \ 'b set" where "vars\<^sub>s\<^sub>t S \ \(set (map vars\<^sub>s\<^sub>t\<^sub>p S))" abbreviation wfrestrictedvars\<^sub>s\<^sub>t\<^sub>p::"('a,'b) strand_step \ 'b set" where "wfrestrictedvars\<^sub>s\<^sub>t\<^sub>p x \ case x of Inequality _ _ \ {} | Equality Check _ _ \ {} | _ \ vars\<^sub>s\<^sub>t\<^sub>p x" text \The variables of a strand whose occurrences might be restricted by well-formedness constraints\ definition wfrestrictedvars\<^sub>s\<^sub>t::"('a,'b) strand \ 'b set" where "wfrestrictedvars\<^sub>s\<^sub>t S \ \(set (map wfrestrictedvars\<^sub>s\<^sub>t\<^sub>p S))" abbreviation wfvarsoccs\<^sub>s\<^sub>t\<^sub>p where "wfvarsoccs\<^sub>s\<^sub>t\<^sub>p x \ case x of Send t \ fv t | Equality Assign s t \ fv s | _ \ {}" text \The variables of a strand that occur in sent messages or as variables in assignments\ definition wfvarsoccs\<^sub>s\<^sub>t where "wfvarsoccs\<^sub>s\<^sub>t S \ \(set (map wfvarsoccs\<^sub>s\<^sub>t\<^sub>p S))" text \The variables occurring at the right-hand side of assignment steps\ fun assignment_rhs\<^sub>s\<^sub>t where "assignment_rhs\<^sub>s\<^sub>t [] = {}" | "assignment_rhs\<^sub>s\<^sub>t (Equality Assign t t'#S) = insert t' (assignment_rhs\<^sub>s\<^sub>t S)" | "assignment_rhs\<^sub>s\<^sub>t (x#S) = assignment_rhs\<^sub>s\<^sub>t S" text \The set function symbols occurring in a strand\ definition funs\<^sub>s\<^sub>t::"('a,'b) strand \ 'a set" where "funs\<^sub>s\<^sub>t S \ \(set (map funs\<^sub>s\<^sub>t\<^sub>p S))" fun subst_apply_strand_step::"('a,'b) strand_step \ ('a,'b) subst \ ('a,'b) strand_step" (infix "\\<^sub>s\<^sub>t\<^sub>p" 51) where "Send t \\<^sub>s\<^sub>t\<^sub>p \ = Send (t \ \)" | "Receive t \\<^sub>s\<^sub>t\<^sub>p \ = Receive (t \ \)" | "Equality a t t' \\<^sub>s\<^sub>t\<^sub>p \ = Equality a (t \ \) (t' \ \)" | "Inequality X F \\<^sub>s\<^sub>t\<^sub>p \ = Inequality X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" text \Substitution application for strands\ definition subst_apply_strand::"('a,'b) strand \ ('a,'b) subst \ ('a,'b) strand" (infix "\\<^sub>s\<^sub>t" 51) where "S \\<^sub>s\<^sub>t \ \ map (\x. x \\<^sub>s\<^sub>t\<^sub>p \) S" text \The semantics of inequality constraints\ definition "ineq_model (\::('a,'b) subst) X F \ (\\. subst_domain \ = set X \ ground (subst_range \) \ list_ex (\f. fst f \ (\ \\<^sub>s \) \ snd f \ (\ \\<^sub>s \)) F)" fun simple\<^sub>s\<^sub>t\<^sub>p where "simple\<^sub>s\<^sub>t\<^sub>p (Receive t) = True" | "simple\<^sub>s\<^sub>t\<^sub>p (Send (Var v)) = True" | "simple\<^sub>s\<^sub>t\<^sub>p (Inequality X F) = (\\. ineq_model \ X F)" | "simple\<^sub>s\<^sub>t\<^sub>p _ = False" text \Simple constraints\ definition simple where "simple S \ list_all simple\<^sub>s\<^sub>t\<^sub>p S" text \The intruder knowledge of a constraint\ fun ik\<^sub>s\<^sub>t::"('a,'b) strand \ ('a,'b) terms" where "ik\<^sub>s\<^sub>t [] = {}" | "ik\<^sub>s\<^sub>t (Receive t#S) = insert t (ik\<^sub>s\<^sub>t S)" | "ik\<^sub>s\<^sub>t (_#S) = ik\<^sub>s\<^sub>t S" text \Strand well-formedness\ fun wf\<^sub>s\<^sub>t::"'b set \ ('a,'b) strand \ bool" where "wf\<^sub>s\<^sub>t V [] = True" | "wf\<^sub>s\<^sub>t V (Receive t#S) = (fv t \ V \ wf\<^sub>s\<^sub>t V S)" | "wf\<^sub>s\<^sub>t V (Send t#S) = wf\<^sub>s\<^sub>t (V \ fv t) S" | "wf\<^sub>s\<^sub>t V (Equality Assign s t#S) = (fv t \ V \ wf\<^sub>s\<^sub>t (V \ fv s) S)" | "wf\<^sub>s\<^sub>t V (Equality Check s t#S) = wf\<^sub>s\<^sub>t V S" | "wf\<^sub>s\<^sub>t V (Inequality _ _#S) = wf\<^sub>s\<^sub>t V S" text \Well-formedness of constraint states\ definition wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r::"('a,'b) strand \ ('a,'b) subst \ bool" where "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \ \ (wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>s\<^sub>t {} S \ subst_domain \ \ vars\<^sub>s\<^sub>t S = {} \ range_vars \ \ bvars\<^sub>s\<^sub>t S = {} \ fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {})" declare trms\<^sub>s\<^sub>t_def[simp] declare fv\<^sub>s\<^sub>n\<^sub>d_def[simp] declare fv\<^sub>r\<^sub>c\<^sub>v_def[simp] declare fv\<^sub>e\<^sub>q_def[simp] declare fv_l\<^sub>e\<^sub>q_def[simp] declare fv_r\<^sub>e\<^sub>q_def[simp] declare fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q_def[simp] declare fv\<^sub>s\<^sub>t_def[simp] declare vars\<^sub>s\<^sub>t_def[simp] declare bvars\<^sub>s\<^sub>t_def[simp] declare wfrestrictedvars\<^sub>s\<^sub>t_def[simp] declare wfvarsoccs\<^sub>s\<^sub>t_def[simp] lemmas wf\<^sub>s\<^sub>t_induct = wf\<^sub>s\<^sub>t.induct[case_names Nil ConsRcv ConsSnd ConsEq ConsEq2 ConsIneq] lemmas ik\<^sub>s\<^sub>t_induct = ik\<^sub>s\<^sub>t.induct[case_names Nil ConsRcv ConsSnd ConsEq ConsIneq] lemmas assignment_rhs\<^sub>s\<^sub>t_induct = assignment_rhs\<^sub>s\<^sub>t.induct[case_names Nil ConsEq2 ConsSnd ConsRcv ConsEq ConsIneq] subsubsection \Lexicographical measure on strands\ definition size\<^sub>s\<^sub>t::"('a,'b) strand \ nat" where "size\<^sub>s\<^sub>t S \ size_list (\x. Max (insert 0 (size ` trms\<^sub>s\<^sub>t\<^sub>p x))) S" definition measure\<^sub>s\<^sub>t::"((('a, 'b) strand \ ('a,'b) subst) \ ('a, 'b) strand \ ('a,'b) subst) set" where "measure\<^sub>s\<^sub>t \ measures [\(S,\). card (fv\<^sub>s\<^sub>t S), \(S,\). size\<^sub>s\<^sub>t S]" lemma measure\<^sub>s\<^sub>t_alt_def: "((s,x),(t,y)) \ measure\<^sub>s\<^sub>t = (card (fv\<^sub>s\<^sub>t s) < card (fv\<^sub>s\<^sub>t t) \ (card (fv\<^sub>s\<^sub>t s) = card (fv\<^sub>s\<^sub>t t) \ size\<^sub>s\<^sub>t s < size\<^sub>s\<^sub>t t))" by (simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) lemma measure\<^sub>s\<^sub>t_trans: "trans measure\<^sub>s\<^sub>t" by (simp add: trans_def measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) subsubsection \Some lemmata\ lemma trms_list\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>t: "trms\<^sub>s\<^sub>t S = set (trms_list\<^sub>s\<^sub>t S)" unfolding trms\<^sub>s\<^sub>t_def trms_list\<^sub>s\<^sub>t_def proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp lemma subst_apply_strand_step_def: "s \\<^sub>s\<^sub>t\<^sub>p \ = (case s of Send t \ Send (t \ \) | Receive t \ Receive (t \ \) | Equality a t t' \ Equality a (t \ \) (t' \ \) | Inequality X F \ Inequality X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \))" by (cases s) simp_all lemma subst_apply_strand_nil[simp]: "[] \\<^sub>s\<^sub>t \ = []" unfolding subst_apply_strand_def by simp lemma finite_funs\<^sub>s\<^sub>t\<^sub>p[simp]: "finite (funs\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) auto lemma finite_funs\<^sub>s\<^sub>t[simp]: "finite (funs\<^sub>s\<^sub>t S)" unfolding funs\<^sub>s\<^sub>t_def by simp lemma finite_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s[simp]: "finite (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s x)" by (induct x) auto lemma finite_trms\<^sub>s\<^sub>t\<^sub>p[simp]: "finite (trms\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) auto lemma finite_vars\<^sub>s\<^sub>t\<^sub>p[simp]: "finite (vars\<^sub>s\<^sub>t\<^sub>p x)" by auto lemma finite_bvars\<^sub>s\<^sub>t\<^sub>p[simp]: "finite (set (bvars\<^sub>s\<^sub>t\<^sub>p x))" by rule lemma finite_fv\<^sub>s\<^sub>n\<^sub>d[simp]: "finite (fv\<^sub>s\<^sub>n\<^sub>d x)" by (cases x) auto lemma finite_fv\<^sub>r\<^sub>c\<^sub>v[simp]: "finite (fv\<^sub>r\<^sub>c\<^sub>v x)" by (cases x) auto lemma finite_fv\<^sub>s\<^sub>t\<^sub>p[simp]: "finite (fv\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) auto lemma finite_vars\<^sub>s\<^sub>t[simp]: "finite (vars\<^sub>s\<^sub>t S)" by simp lemma finite_bvars\<^sub>s\<^sub>t[simp]: "finite (bvars\<^sub>s\<^sub>t S)" by simp lemma finite_fv\<^sub>s\<^sub>t[simp]: "finite (fv\<^sub>s\<^sub>t S)" by simp lemma finite_wfrestrictedvars\<^sub>s\<^sub>t\<^sub>p[simp]: "finite (wfrestrictedvars\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) (auto split: poscheckvariant.splits) lemma finite_wfrestrictedvars\<^sub>s\<^sub>t[simp]: "finite (wfrestrictedvars\<^sub>s\<^sub>t S)" using finite_wfrestrictedvars\<^sub>s\<^sub>t\<^sub>p by auto lemma finite_wfvarsoccs\<^sub>s\<^sub>t\<^sub>p[simp]: "finite (wfvarsoccs\<^sub>s\<^sub>t\<^sub>p x)" by (cases x) (auto split: poscheckvariant.splits) lemma finite_wfvarsoccs\<^sub>s\<^sub>t[simp]: "finite (wfvarsoccs\<^sub>s\<^sub>t S)" using finite_wfvarsoccs\<^sub>s\<^sub>t\<^sub>p by auto lemma finite_ik\<^sub>s\<^sub>t[simp]: "finite (ik\<^sub>s\<^sub>t S)" by (induct S rule: ik\<^sub>s\<^sub>t.induct) simp_all lemma finite_assignment_rhs\<^sub>s\<^sub>t[simp]: "finite (assignment_rhs\<^sub>s\<^sub>t S)" by (induct S rule: assignment_rhs\<^sub>s\<^sub>t.induct) simp_all lemma ik\<^sub>s\<^sub>t_is_rcv_set: "ik\<^sub>s\<^sub>t A = {t. Receive t \ set A}" by (induct A rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik\<^sub>s\<^sub>tD[dest]: "t \ ik\<^sub>s\<^sub>t S \ Receive t \ set S" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik\<^sub>s\<^sub>tD'[dest]: "t \ ik\<^sub>s\<^sub>t S \ t \ trms\<^sub>s\<^sub>t S" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik\<^sub>s\<^sub>tD''[dest]: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t S) \ t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik\<^sub>s\<^sub>t_subterm_exD: assumes "t \ ik\<^sub>s\<^sub>t S" shows "\x \ set S. t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t\<^sub>p x)" using assms ik\<^sub>s\<^sub>tD by force lemma assignment_rhs\<^sub>s\<^sub>tD[dest]: "t \ assignment_rhs\<^sub>s\<^sub>t S \ \t'. Equality Assign t' t \ set S" by (induct S rule: assignment_rhs\<^sub>s\<^sub>t.induct) auto lemma assignment_rhs\<^sub>s\<^sub>tD'[dest]: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (assignment_rhs\<^sub>s\<^sub>t S) \ t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" by (induct S rule: assignment_rhs\<^sub>s\<^sub>t.induct) auto lemma bvars\<^sub>s\<^sub>t_split: "bvars\<^sub>s\<^sub>t (S@S') = bvars\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S'" unfolding bvars\<^sub>s\<^sub>t_def by auto lemma bvars\<^sub>s\<^sub>t_singleton: "bvars\<^sub>s\<^sub>t [x] = set (bvars\<^sub>s\<^sub>t\<^sub>p x)" unfolding bvars\<^sub>s\<^sub>t_def by auto lemma strand_fv_bvars_disjointD: assumes "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" "Inequality X F \ set S" shows "set X \ bvars\<^sub>s\<^sub>t S" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ fv\<^sub>s\<^sub>t S" using assms by (induct S) fastforce+ lemma strand_fv_bvars_disjoint_unfold: assumes "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" "Inequality X F \ set S" "Inequality Y G \ set S" shows "set Y \ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X) = {}" proof - have "set X \ bvars\<^sub>s\<^sub>t S" "set Y \ bvars\<^sub>s\<^sub>t S" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ fv\<^sub>s\<^sub>t S" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G - set Y \ fv\<^sub>s\<^sub>t S" using strand_fv_bvars_disjointD[OF assms(1)] assms(2,3) by auto thus ?thesis using assms(1) by fastforce qed lemma strand_subst_hom[iff]: "(S@S') \\<^sub>s\<^sub>t \ = (S \\<^sub>s\<^sub>t \)@(S' \\<^sub>s\<^sub>t \)" "(x#S) \\<^sub>s\<^sub>t \ = (x \\<^sub>s\<^sub>t\<^sub>p \)#(S \\<^sub>s\<^sub>t \)" unfolding subst_apply_strand_def by auto lemma strand_subst_comp: "range_vars \ \ bvars\<^sub>s\<^sub>t S = {} \ S \\<^sub>s\<^sub>t \ \\<^sub>s \ = ((S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \)" proof (induction S) case (Cons x S) have *: "range_vars \ \ bvars\<^sub>s\<^sub>t S = {}" "range_vars \ \ (set (bvars\<^sub>s\<^sub>t\<^sub>p x)) = {}" using Cons bvars\<^sub>s\<^sub>t_split[of "[x]" S] append_Cons inf_sup_absorb by (metis (no_types, lifting) Int_iff Un_commute disjoint_iff_not_equal self_append_conv2, metis append_self_conv2 bvars\<^sub>s\<^sub>t_singleton inf_bot_right inf_left_commute) hence IH: "S \\<^sub>s\<^sub>t \ \\<^sub>s \ = (S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \" using Cons.IH by auto have "(x#S \\<^sub>s\<^sub>t \ \\<^sub>s \) = (x \\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \)#(S \\<^sub>s\<^sub>t \ \\<^sub>s \)" by (metis strand_subst_hom(2)) hence "... = (x \\<^sub>s\<^sub>t\<^sub>p \ \\<^sub>s \)#((S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \)" by (metis IH) hence "... = ((x \\<^sub>s\<^sub>t\<^sub>p \) \\<^sub>s\<^sub>t\<^sub>p \)#((S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \)" using rm_vars_comp[OF *(2)] proof (induction x) case (Inequality X F) thus ?case by (induct F) (auto simp add: subst_apply_pairs_def subst_apply_strand_step_def) qed (simp_all add: subst_apply_strand_step_def) thus ?case using IH by auto qed (simp add: subst_apply_strand_def) lemma strand_substI[intro]: "subst_domain \ \ fv\<^sub>s\<^sub>t S = {} \ S \\<^sub>s\<^sub>t \ = S" "subst_domain \ \ vars\<^sub>s\<^sub>t S = {} \ S \\<^sub>s\<^sub>t \ = S" proof - show "subst_domain \ \ vars\<^sub>s\<^sub>t S = {} \ S \\<^sub>s\<^sub>t \ = S" proof (induction S) case (Cons x S) hence "S \\<^sub>s\<^sub>t \ = S" by auto moreover have "vars\<^sub>s\<^sub>t\<^sub>p x \ subst_domain \ = {}" using Cons.prems by auto hence "x \\<^sub>s\<^sub>t\<^sub>p \ = x" proof (induction x) case (Inequality X F) thus ?case by (induct F) (force simp add: subst_apply_pairs_def)+ qed auto ultimately show ?case by simp qed (simp add: subst_apply_strand_def) show "subst_domain \ \ fv\<^sub>s\<^sub>t S = {} \ S \\<^sub>s\<^sub>t \ = S" proof (induction S) case (Cons x S) hence "S \\<^sub>s\<^sub>t \ = S" by auto moreover have "fv\<^sub>s\<^sub>t\<^sub>p x \ subst_domain \ = {}" using Cons.prems by auto hence "x \\<^sub>s\<^sub>t\<^sub>p \ = x" proof (induction x) case (Inequality X F) thus ?case by (induct F) (force simp add: subst_apply_pairs_def)+ qed auto ultimately show ?case by simp qed (simp add: subst_apply_strand_def) qed lemma strand_substI': "fv\<^sub>s\<^sub>t S = {} \ S \\<^sub>s\<^sub>t \ = S" "vars\<^sub>s\<^sub>t S = {} \ S \\<^sub>s\<^sub>t \ = S" by (metis inf_bot_right strand_substI(1), metis inf_bot_right strand_substI(2)) lemma strand_subst_set: "(set (S \\<^sub>s\<^sub>t \)) = ((\x. x \\<^sub>s\<^sub>t\<^sub>p \) ` (set S))" by (auto simp add: subst_apply_strand_def) lemma strand_map_inv_set_snd_rcv_subst: assumes "finite (M::('a,'b) terms)" shows "set ((map Send (inv set M)) \\<^sub>s\<^sub>t \) = Send ` (M \\<^sub>s\<^sub>e\<^sub>t \)" (is ?A) "set ((map Receive (inv set M)) \\<^sub>s\<^sub>t \) = Receive ` (M \\<^sub>s\<^sub>e\<^sub>t \)" (is ?B) proof - { fix f::"('a,'b) term \ ('a,'b) strand_step" assume f: "f = Send \ f = Receive" from assms have "set ((map f (inv set M)) \\<^sub>s\<^sub>t \) = f ` (M \\<^sub>s\<^sub>e\<^sub>t \)" proof (induction rule: finite_induct) case empty thus ?case unfolding inv_def by auto next case (insert m M) have "set (map f (inv set (insert m M)) \\<^sub>s\<^sub>t \) = insert (f m \\<^sub>s\<^sub>t\<^sub>p \) (set (map f (inv set M) \\<^sub>s\<^sub>t \))" by (simp add: insert.hyps(1) inv_set_fset subst_apply_strand_def) thus ?case using f insert.IH by auto qed } thus "?A" "?B" by auto qed lemma strand_ground_subst_vars_subset: assumes "ground (subst_range \)" shows "vars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ vars\<^sub>s\<^sub>t S" proof (induction S) case (Cons x S) have "vars\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ vars\<^sub>s\<^sub>t\<^sub>p x" using ground_subst_fv_subset[OF assms] proof (cases x) case (Inequality X F) let ?\ = "rm_vars (set X) \" have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" proof (induction F) case (Cons f F) obtain t t' where f: "f = (t,t')" by (metis surj_pair) hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (f#F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\) = fv (t \ ?\) \ fv (t' \ ?\) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (f#F) = fv t \ fv t' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" by (auto simp add: subst_apply_pairs_def) thus ?case using ground_subst_fv_subset[OF ground_subset[OF rm_vars_img_subset assms, of "set X"]] Cons.IH by (metis (no_types, lifting) Un_mono) qed (simp add: subst_apply_pairs_def) moreover have "vars\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ set X" "vars\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ set X" using Inequality by (auto simp add: subst_apply_pairs_def) ultimately show ?thesis by auto qed auto thus ?case using Cons.IH by auto qed (simp add: subst_apply_strand_def) lemma ik_union_subset: "\(P ` ik\<^sub>s\<^sub>t S) \ (\x \ (set S). \(P ` trms\<^sub>s\<^sub>t\<^sub>p x))" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik_snd_empty[simp]: "ik\<^sub>s\<^sub>t (map Send X) = {}" by (induct "map Send X" arbitrary: X rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik_snd_empty'[simp]: "ik\<^sub>s\<^sub>t [Send t] = {}" by simp lemma ik_append[iff]: "ik\<^sub>s\<^sub>t (S@S') = ik\<^sub>s\<^sub>t S \ ik\<^sub>s\<^sub>t S'" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik_cons: "ik\<^sub>s\<^sub>t (x#S) = ik\<^sub>s\<^sub>t [x] \ ik\<^sub>s\<^sub>t S" using ik_append[of "[x]" S] by simp lemma assignment_rhs_append[iff]: "assignment_rhs\<^sub>s\<^sub>t (S@S') = assignment_rhs\<^sub>s\<^sub>t S \ assignment_rhs\<^sub>s\<^sub>t S'" by (induct S rule: assignment_rhs\<^sub>s\<^sub>t.induct) auto lemma eqs_rcv_map_empty: "assignment_rhs\<^sub>s\<^sub>t (map Receive M) = {}" by auto lemma ik_rcv_map: assumes "t \ set L" shows "t \ ik\<^sub>s\<^sub>t (map Receive L)" proof - { fix L L' have "t \ ik\<^sub>s\<^sub>t [Receive t]" by auto hence "t \ ik\<^sub>s\<^sub>t (map Receive L@Receive t#map Receive L')" using ik_append by auto hence "t \ ik\<^sub>s\<^sub>t (map Receive (L@t#L'))" by auto } thus ?thesis using assms split_list_last by force qed lemma ik_subst: "ik\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) = ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \" by (induct rule: ik\<^sub>s\<^sub>t_induct) auto lemma ik_rcv_map': assumes "t \ ik\<^sub>s\<^sub>t (map Receive L)" shows "t \ set L" using assms by force lemma ik_append_subset[simp]: "ik\<^sub>s\<^sub>t S \ ik\<^sub>s\<^sub>t (S@S')" "ik\<^sub>s\<^sub>t S' \ ik\<^sub>s\<^sub>t (S@S')" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma assignment_rhs_append_subset[simp]: "assignment_rhs\<^sub>s\<^sub>t S \ assignment_rhs\<^sub>s\<^sub>t (S@S')" "assignment_rhs\<^sub>s\<^sub>t S' \ assignment_rhs\<^sub>s\<^sub>t (S@S')" by (induct S rule: assignment_rhs\<^sub>s\<^sub>t.induct) auto lemma trms\<^sub>s\<^sub>t_cons: "trms\<^sub>s\<^sub>t (x#S) = trms\<^sub>s\<^sub>t\<^sub>p x \ trms\<^sub>s\<^sub>t S" by simp lemma trm_strand_subst_cong: "t \ trms\<^sub>s\<^sub>t S \ t \ \ \ trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ (\X F. Inequality X F \ set S \ t \ rm_vars (set X) \ \ trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \))" (is "t \ trms\<^sub>s\<^sub>t S \ ?P t \ S") "t \ trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ (\t'. t = t' \ \ \ t' \ trms\<^sub>s\<^sub>t S) \ (\X F. Inequality X F \ set S \ (\t' \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F. t = t' \ rm_vars (set X) \))" (is "t \ trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ ?Q t \ S") proof - show "t \ trms\<^sub>s\<^sub>t S \ ?P t \ S" proof (induction S) case (Cons x S) show ?case proof (cases "t \ trms\<^sub>s\<^sub>t S") case True hence "?P t \ S" using Cons by simp thus ?thesis by (cases x) (metis (no_types, lifting) Un_iff list.set_intros(2) strand_subst_hom(2) trms\<^sub>s\<^sub>t_cons)+ next case False hence "t \ trms\<^sub>s\<^sub>t\<^sub>p x" using Cons.prems by auto thus ?thesis proof (induction x) case (Inequality X F) hence "t \ rm_vars (set X) \ \ trms\<^sub>s\<^sub>t\<^sub>p (Inequality X F \\<^sub>s\<^sub>t\<^sub>p \)" by (induct F) (auto simp add: subst_apply_pairs_def subst_apply_strand_step_def) thus ?case by fastforce qed (auto simp add: subst_apply_strand_step_def) qed qed simp show "t \ trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ ?Q t \ S" proof (induction S) case (Cons x S) show ?case proof (cases "t \ trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)") case True hence "?Q t \ S" using Cons by simp thus ?thesis by (cases x) force+ next case False hence "t \ trms\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \)" using Cons.prems by auto thus ?thesis proof (induction x) case (Inequality X F) hence "t \ trms\<^sub>s\<^sub>t\<^sub>p (Inequality X F) \\<^sub>s\<^sub>e\<^sub>t rm_vars (set X) \" by (induct F) (force simp add: subst_apply_pairs_def)+ thus ?case by fastforce qed (auto simp add: subst_apply_strand_step_def) qed qed simp qed subsection \Lemmata: Free Variables of Strands\ lemma fv_trm_snd_rcv[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t\<^sub>p (Send t)) = fv t" "fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t\<^sub>p (Receive t)) = fv t" by simp_all lemma in_strand_fv_subset: "x \ set S \ vars\<^sub>s\<^sub>t\<^sub>p x \ vars\<^sub>s\<^sub>t S" by fastforce lemma in_strand_fv_subset_snd: "Send t \ set S \ fv t \ \(set (map fv\<^sub>s\<^sub>n\<^sub>d S))" by auto lemma in_strand_fv_subset_rcv: "Receive t \ set S \ fv t \ \(set (map fv\<^sub>r\<^sub>c\<^sub>v S))" by auto lemma fv\<^sub>s\<^sub>n\<^sub>dE: assumes "v \ \(set (map fv\<^sub>s\<^sub>n\<^sub>d S))" obtains t where "send\t\\<^sub>s\<^sub>t \ set S" "v \ fv t" proof - have "\t. send\t\\<^sub>s\<^sub>t \ set S \ v \ fv t" by (metis (no_types, lifting) assms UN_E empty_iff set_map strand_step.case_eq_if fv\<^sub>s\<^sub>n\<^sub>d_def strand_step.collapse(1)) thus ?thesis by (metis that) qed lemma fv\<^sub>r\<^sub>c\<^sub>vE: assumes "v \ \(set (map fv\<^sub>r\<^sub>c\<^sub>v S))" obtains t where "receive\t\\<^sub>s\<^sub>t \ set S" "v \ fv t" proof - have "\t. receive\t\\<^sub>s\<^sub>t \ set S \ v \ fv t" by (metis (no_types, lifting) assms UN_E empty_iff set_map strand_step.case_eq_if fv\<^sub>r\<^sub>c\<^sub>v_def strand_step.collapse(2)) thus ?thesis by (metis that) qed lemma vars\<^sub>s\<^sub>t\<^sub>pI[intro]: "x \ fv\<^sub>s\<^sub>t\<^sub>p s \ x \ vars\<^sub>s\<^sub>t\<^sub>p s" by (induct s rule: fv\<^sub>s\<^sub>t\<^sub>p.induct) auto lemma vars\<^sub>s\<^sub>tI[intro]: "x \ fv\<^sub>s\<^sub>t S \ x \ vars\<^sub>s\<^sub>t S" using vars\<^sub>s\<^sub>t\<^sub>pI by fastforce lemma fv\<^sub>s\<^sub>t_subset_vars\<^sub>s\<^sub>t[simp]: "fv\<^sub>s\<^sub>t S \ vars\<^sub>s\<^sub>t S" using vars\<^sub>s\<^sub>tI by force lemma vars\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>t: "vars\<^sub>s\<^sub>t S = fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S" proof (induction S) case (Cons x S) thus ?case proof (induction x) case (Inequality X F) thus ?case by (induct F) auto qed auto qed simp lemma fv\<^sub>s\<^sub>t\<^sub>p_is_subterm_trms\<^sub>s\<^sub>t\<^sub>p: "x \ fv\<^sub>s\<^sub>t\<^sub>p a \ Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t\<^sub>p a)" using var_is_subterm by (cases a) force+ lemma fv\<^sub>s\<^sub>t_is_subterm_trms\<^sub>s\<^sub>t: "x \ fv\<^sub>s\<^sub>t A \ Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t A)" proof (induction A) case (Cons a A) thus ?case using fv\<^sub>s\<^sub>t\<^sub>p_is_subterm_trms\<^sub>s\<^sub>t\<^sub>p by (cases "x \ fv\<^sub>s\<^sub>t A") auto qed simp lemma vars_st_snd_map: "vars\<^sub>s\<^sub>t (map Send X) = fv (Fun f X)" by auto lemma vars_st_rcv_map: "vars\<^sub>s\<^sub>t (map Receive X) = fv (Fun f X)" by auto lemma vars_snd_rcv_union: "vars\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>s\<^sub>n\<^sub>d x \ fv\<^sub>r\<^sub>c\<^sub>v x \ fv\<^sub>e\<^sub>q assign x \ fv\<^sub>e\<^sub>q check x \ fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x \ set (bvars\<^sub>s\<^sub>t\<^sub>p x)" proof (cases x) case (Equality ac t t') thus ?thesis by (cases ac) auto qed auto lemma fv_snd_rcv_union: "fv\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>s\<^sub>n\<^sub>d x \ fv\<^sub>r\<^sub>c\<^sub>v x \ fv\<^sub>e\<^sub>q assign x \ fv\<^sub>e\<^sub>q check x \ fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x" proof (cases x) case (Equality ac t t') thus ?thesis by (cases ac) auto qed auto lemma fv_snd_rcv_empty[simp]: "fv\<^sub>s\<^sub>n\<^sub>d x = {} \ fv\<^sub>r\<^sub>c\<^sub>v x = {}" by (cases x) simp_all lemma vars_snd_rcv_strand[iff]: "vars\<^sub>s\<^sub>t (S::('a,'b) strand) = (\(set (map fv\<^sub>s\<^sub>n\<^sub>d S))) \ (\(set (map fv\<^sub>r\<^sub>c\<^sub>v S))) \ (\(set (map (fv\<^sub>e\<^sub>q assign) S))) \ (\(set (map (fv\<^sub>e\<^sub>q check) S))) \ (\(set (map fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q S))) \ bvars\<^sub>s\<^sub>t S" unfolding bvars\<^sub>s\<^sub>t_def proof (induction S) case (Cons x S) have "\s V. vars\<^sub>s\<^sub>t\<^sub>p (s::('a,'b) strand_step) \ V = fv\<^sub>s\<^sub>n\<^sub>d s \ fv\<^sub>r\<^sub>c\<^sub>v s \ fv\<^sub>e\<^sub>q assign s \ fv\<^sub>e\<^sub>q check s \ fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q s \ set (bvars\<^sub>s\<^sub>t\<^sub>p s) \ V" by (metis vars_snd_rcv_union) thus ?case using Cons.IH by (auto simp add: sup_assoc sup_left_commute) qed simp lemma fv_snd_rcv_strand[iff]: "fv\<^sub>s\<^sub>t (S::('a,'b) strand) = (\(set (map fv\<^sub>s\<^sub>n\<^sub>d S))) \ (\(set (map fv\<^sub>r\<^sub>c\<^sub>v S))) \ (\(set (map (fv\<^sub>e\<^sub>q assign) S))) \ (\(set (map (fv\<^sub>e\<^sub>q check) S))) \ (\(set (map fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q S)))" unfolding bvars\<^sub>s\<^sub>t_def proof (induction S) case (Cons x S) have "\s V. fv\<^sub>s\<^sub>t\<^sub>p (s::('a,'b) strand_step) \ V = fv\<^sub>s\<^sub>n\<^sub>d s \ fv\<^sub>r\<^sub>c\<^sub>v s \ fv\<^sub>e\<^sub>q assign s \ fv\<^sub>e\<^sub>q check s \ fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q s \ V" by (metis fv_snd_rcv_union) thus ?case using Cons.IH by (auto simp add: sup_assoc sup_left_commute) qed simp lemma vars_snd_rcv_strand2[iff]: "wfrestrictedvars\<^sub>s\<^sub>t (S::('a,'b) strand) = (\(set (map fv\<^sub>s\<^sub>n\<^sub>d S))) \ (\(set (map fv\<^sub>r\<^sub>c\<^sub>v S))) \ (\(set (map (fv\<^sub>e\<^sub>q assign) S)))" by (induct S) (auto simp add: split: strand_step.split poscheckvariant.split) lemma fv_snd_rcv_strand_subset[simp]: "\(set (map fv\<^sub>s\<^sub>n\<^sub>d S)) \ fv\<^sub>s\<^sub>t S" "\(set (map fv\<^sub>r\<^sub>c\<^sub>v S)) \ fv\<^sub>s\<^sub>t S" "\(set (map (fv\<^sub>e\<^sub>q ac) S)) \ fv\<^sub>s\<^sub>t S" "\(set (map fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q S)) \ fv\<^sub>s\<^sub>t S" "wfvarsoccs\<^sub>s\<^sub>t S \ fv\<^sub>s\<^sub>t S" proof - show "\(set (map fv\<^sub>s\<^sub>n\<^sub>d S)) \ fv\<^sub>s\<^sub>t S" "\(set (map fv\<^sub>r\<^sub>c\<^sub>v S)) \ fv\<^sub>s\<^sub>t S" "\(set (map fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q S)) \ fv\<^sub>s\<^sub>t S" using fv_snd_rcv_strand[of S] by auto show "\(set (map (fv\<^sub>e\<^sub>q ac) S)) \ fv\<^sub>s\<^sub>t S" by (induct S) (auto split: strand_step.split poscheckvariant.split) show "wfvarsoccs\<^sub>s\<^sub>t S \ fv\<^sub>s\<^sub>t S" by (induct S) (auto split: strand_step.split poscheckvariant.split) qed lemma vars_snd_rcv_strand_subset2[simp]: "\(set (map fv\<^sub>s\<^sub>n\<^sub>d S)) \ wfrestrictedvars\<^sub>s\<^sub>t S" "\(set (map fv\<^sub>r\<^sub>c\<^sub>v S)) \ wfrestrictedvars\<^sub>s\<^sub>t S" "\(set (map (fv\<^sub>e\<^sub>q assign) S)) \ wfrestrictedvars\<^sub>s\<^sub>t S" "wfvarsoccs\<^sub>s\<^sub>t S \ wfrestrictedvars\<^sub>s\<^sub>t S" by (induction S) (auto split: strand_step.split poscheckvariant.split) lemma wfrestrictedvars\<^sub>s\<^sub>t_subset_vars\<^sub>s\<^sub>t: "wfrestrictedvars\<^sub>s\<^sub>t S \ vars\<^sub>s\<^sub>t S" by (induction S) (auto split: strand_step.split poscheckvariant.split) lemma subst_sends_strand_step_fv_to_img: "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>s\<^sub>t\<^sub>p x \ range_vars \" using subst_sends_fv_to_img[of _ \] proof (cases x) case (Inequality X F) let ?\ = "rm_vars (set X) \" have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ range_vars ?\" proof (induction F) case (Cons f F) thus ?case using subst_sends_fv_to_img[of _ ?\] by (auto simp add: subst_apply_pairs_def) qed (auto simp add: subst_apply_pairs_def) hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ range_vars \" using rm_vars_img_subset[of "set X" \] fv_set_mono unfolding range_vars_alt_def by blast+ thus ?thesis using Inequality by (auto simp add: subst_apply_strand_step_def) qed (auto simp add: subst_apply_strand_step_def) lemma subst_sends_strand_fv_to_img: "fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t S \ range_vars \" proof (induction S) case (Cons x S) have *: "fv\<^sub>s\<^sub>t (x#S \\<^sub>s\<^sub>t \) = fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" "fv\<^sub>s\<^sub>t (x#S) \ range_vars \ = fv\<^sub>s\<^sub>t\<^sub>p x \ fv\<^sub>s\<^sub>t S \ range_vars \" by auto thus ?case using Cons.IH subst_sends_strand_step_fv_to_img[of x \] by auto qed simp lemma ineq_apply_subst: assumes "subst_domain \ \ set X = {}" shows "(Inequality X F) \\<^sub>s\<^sub>t\<^sub>p \ = Inequality X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using rm_vars_apply'[OF assms] by (simp add: subst_apply_strand_step_def) lemma fv_strand_step_subst: assumes "P = fv\<^sub>s\<^sub>t\<^sub>p \ P = fv\<^sub>r\<^sub>c\<^sub>v \ P = fv\<^sub>s\<^sub>n\<^sub>d \ P = fv\<^sub>e\<^sub>q ac \ P = fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q" and "set (bvars\<^sub>s\<^sub>t\<^sub>p x) \ (subst_domain \ \ range_vars \) = {}" shows "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (P x)) = P (x \\<^sub>s\<^sub>t\<^sub>p \)" proof (cases x) case (Send t) hence "vars\<^sub>s\<^sub>t\<^sub>p x = fv t" "fv\<^sub>s\<^sub>n\<^sub>d x = fv t" by auto thus ?thesis using assms Send subst_apply_fv_unfold[of _ \] by auto next case (Receive t) hence "vars\<^sub>s\<^sub>t\<^sub>p x = fv t" "fv\<^sub>r\<^sub>c\<^sub>v x = fv t" by auto thus ?thesis using assms Receive subst_apply_fv_unfold[of _ \] by auto next case (Equality ac' t t') show ?thesis proof (cases "ac = ac'") case True hence "vars\<^sub>s\<^sub>t\<^sub>p x = fv t \ fv t'" "fv\<^sub>e\<^sub>q ac x = fv t \ fv t'" using Equality by auto thus ?thesis using assms Equality subst_apply_fv_unfold[of _ \] True by auto next case False hence "vars\<^sub>s\<^sub>t\<^sub>p x = fv t \ fv t'" "fv\<^sub>e\<^sub>q ac x = {}" using Equality by auto thus ?thesis using assms Equality subst_apply_fv_unfold[of _ \] False by auto qed next case (Inequality X F) hence 1: "set X \ (subst_domain \ \ range_vars \) = {}" "x \\<^sub>s\<^sub>t\<^sub>p \ = Inequality X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" "rm_vars (set X) \ = \" using assms ineq_apply_subst[of \ X F] rm_vars_apply'[of \ "set X"] unfolding range_vars_alt_def by force+ have 2: "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" using Inequality by auto hence "fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) - set X" using fv\<^sub>s\<^sub>e\<^sub>t_subst_img_eq[OF 1(1), of "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F"] by simp hence 3: "fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" by (metis fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_step_subst) have 4: "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" using 1(2) by auto show ?thesis using assms(1) Inequality subst_apply_fv_unfold[of _ \] 1(2) 2 3 4 unfolding fv\<^sub>e\<^sub>q_def fv\<^sub>r\<^sub>c\<^sub>v_def fv\<^sub>s\<^sub>n\<^sub>d_def by (metis (no_types) Sup_empty image_empty fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.simps fv\<^sub>s\<^sub>e\<^sub>t.simps fv\<^sub>s\<^sub>t\<^sub>p.simps(4) strand_step.simps(20)) qed lemma fv_strand_subst: assumes "P = fv\<^sub>s\<^sub>t\<^sub>p \ P = fv\<^sub>r\<^sub>c\<^sub>v \ P = fv\<^sub>s\<^sub>n\<^sub>d \ P = fv\<^sub>e\<^sub>q ac \ P = fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q" and "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (\(set (map P S)))) = \(set (map P (S \\<^sub>s\<^sub>t \)))" using assms(2) proof (induction S) case (Cons x S) hence *: "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" "set (bvars\<^sub>s\<^sub>t\<^sub>p x) \ (subst_domain \ \ range_vars \) = {}" unfolding bvars\<^sub>s\<^sub>t_def by force+ hence **: "fv\<^sub>s\<^sub>e\<^sub>t (\ ` P x) = P (x \\<^sub>s\<^sub>t\<^sub>p \)" using fv_strand_step_subst[OF assms(1), of x \] by auto have "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (\(set (map P (x#S))))) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` P x) \ (\(set (map P ((S \\<^sub>s\<^sub>t \)))))" using Cons unfolding range_vars_alt_def bvars\<^sub>s\<^sub>t_def by force hence "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (\(set (map P (x#S))))) = P (x \\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (\(set (map P S))))" using ** by simp thus ?case using Cons.IH[OF *(1)] unfolding bvars\<^sub>s\<^sub>t_def by simp qed simp lemma fv_strand_subst2: assumes "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (wfrestrictedvars\<^sub>s\<^sub>t S)) = wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" by (metis (no_types, lifting) assms fv\<^sub>s\<^sub>e\<^sub>t.simps vars_snd_rcv_strand2 fv_strand_subst UN_Un image_Un) lemma fv_strand_subst': assumes "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (fv\<^sub>s\<^sub>t S)) = fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" by (metis assms fv_strand_subst fv\<^sub>s\<^sub>t_def) lemma fv_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s: "fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" by auto lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_in_fv_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s: "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ x \ fv\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F)" using fv_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s[of F] by blast lemma trms\<^sub>s\<^sub>t_append: "trms\<^sub>s\<^sub>t (A@B) = trms\<^sub>s\<^sub>t A \ trms\<^sub>s\<^sub>t B" by auto lemma trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst: "trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (a \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) = trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s a \\<^sub>s\<^sub>e\<^sub>t \" by (auto simp add: subst_apply_pairs_def) lemma trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_fv_subst_subset: "t \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv (t \ \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" by (force simp add: subst_apply_pairs_def) lemma trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_fv_subst_subset': fixes t::"('a,'b) term" and \::"('a,'b) subst" assumes "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F)" shows "fv (t \ \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" proof - { fix x assume "x \ fv t" hence "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" using fv_subset[OF assms] fv_subterms_set[of "trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F"] fv_trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_is_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s[of F] by blast hence "fv (\ x) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_fv_subset by fast } thus ?thesis by (meson fv_subst_obtain_var subset_iff) qed lemma trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_funs_term_cases: assumes "t \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" "f \ funs_term t" shows "(\u \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F. f \ funs_term u) \ (\x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F. f \ funs_term (\ x))" using assms(1) proof (induction F) case (Cons g F) obtain s u where g: "g = (s,u)" by (metis surj_pair) show ?case proof (cases "t \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)") case False thus ?thesis using assms(2) Cons.prems g funs_term_subst[of _ \] by (auto simp add: subst_apply_pairs_def) qed (use Cons.IH in fastforce) qed simp lemma trm\<^sub>s\<^sub>t\<^sub>p_subst: assumes "subst_domain \ \ set (bvars\<^sub>s\<^sub>t\<^sub>p a) = {}" shows "trms\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>t\<^sub>p \) = trms\<^sub>s\<^sub>t\<^sub>p a \\<^sub>s\<^sub>e\<^sub>t \" proof - have "rm_vars (set (bvars\<^sub>s\<^sub>t\<^sub>p a)) \ = \" using assms by force thus ?thesis using assms by (auto simp add: subst_apply_pairs_def subst_apply_strand_step_def split: strand_step.splits) qed lemma trms\<^sub>s\<^sub>t_subst: assumes "subst_domain \ \ bvars\<^sub>s\<^sub>t A = {}" shows "trms\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>t \) = trms\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" using assms proof (induction A) case (Cons a A) have 1: "subst_domain \ \ bvars\<^sub>s\<^sub>t A = {}" "subst_domain \ \ set (bvars\<^sub>s\<^sub>t\<^sub>p a) = {}" using Cons.prems by auto hence IH: "trms\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ = trms\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>t \)" using Cons.IH by simp have "trms\<^sub>s\<^sub>t (a#A) = trms\<^sub>s\<^sub>t\<^sub>p a \ trms\<^sub>s\<^sub>t A" by auto hence 2: "trms\<^sub>s\<^sub>t (a#A) \\<^sub>s\<^sub>e\<^sub>t \ = (trms\<^sub>s\<^sub>t\<^sub>p a \\<^sub>s\<^sub>e\<^sub>t \) \ (trms\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \)" by (metis image_Un) have "trms\<^sub>s\<^sub>t (a#A \\<^sub>s\<^sub>t \) = (trms\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>t\<^sub>p \)) \ trms\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>t \)" by (auto simp add: subst_apply_strand_def) hence 3: "trms\<^sub>s\<^sub>t (a#A \\<^sub>s\<^sub>t \) = (trms\<^sub>s\<^sub>t\<^sub>p a \\<^sub>s\<^sub>e\<^sub>t \) \ trms\<^sub>s\<^sub>t (A \\<^sub>s\<^sub>t \)" using trm\<^sub>s\<^sub>t\<^sub>p_subst[OF 1(2)] by auto show ?case using IH 2 3 by metis qed (simp add: subst_apply_strand_def) lemma strand_map_set_subst: assumes \: "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "\(set (map trms\<^sub>s\<^sub>t\<^sub>p (S \\<^sub>s\<^sub>t \))) = (\(set (map trms\<^sub>s\<^sub>t\<^sub>p S))) \\<^sub>s\<^sub>e\<^sub>t \" using assms proof (induction S) case (Cons x S) hence "bvars\<^sub>s\<^sub>t [x] \ subst_domain \ = {}" "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" unfolding bvars\<^sub>s\<^sub>t_def by force+ hence *: "subst_domain \ \ set (bvars\<^sub>s\<^sub>t\<^sub>p x) = {}" "\(set (map trms\<^sub>s\<^sub>t\<^sub>p (S \\<^sub>s\<^sub>t \))) = \(set (map trms\<^sub>s\<^sub>t\<^sub>p S)) \\<^sub>s\<^sub>e\<^sub>t \" using Cons.IH(1) bvars\<^sub>s\<^sub>t_singleton[of x] by auto hence "trms\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = (trms\<^sub>s\<^sub>t\<^sub>p x) \\<^sub>s\<^sub>e\<^sub>t \" proof (cases x) case (Inequality X F) thus ?thesis using rm_vars_apply'[of \ "set X"] * by (metis (no_types, lifting) image_cong trm\<^sub>s\<^sub>t\<^sub>p_subst) qed simp_all thus ?case using * subst_all_insert by auto qed simp lemma subst_apply_fv_subset_strand_trm: assumes P: "P = fv\<^sub>s\<^sub>t\<^sub>p \ P = fv\<^sub>r\<^sub>c\<^sub>v \ P = fv\<^sub>s\<^sub>n\<^sub>d \ P = fv\<^sub>e\<^sub>q ac \ P = fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q" and fv_sub: "fv t \ \(set (map P S)) \ V" and \: "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "fv (t \ \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using fv_strand_subst[OF P \] subst_apply_fv_subset[OF fv_sub, of \] by force lemma subst_apply_fv_subset_strand_trm2: assumes fv_sub: "fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" and \: "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "fv (t \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using fv_strand_subst2[OF \] subst_apply_fv_subset[OF fv_sub, of \] by force lemma subst_apply_fv_subset_strand: assumes P: "P = fv\<^sub>s\<^sub>t\<^sub>p \ P = fv\<^sub>r\<^sub>c\<^sub>v \ P = fv\<^sub>s\<^sub>n\<^sub>d \ P = fv\<^sub>e\<^sub>q ac \ P = fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q" and P_subset: "P x \ \(set (map P S)) \ V" and \: "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" "set (bvars\<^sub>s\<^sub>t\<^sub>p x) \ (subst_domain \ \ range_vars \) = {}" shows "P (x \\<^sub>s\<^sub>t\<^sub>p \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" proof (cases x) case (Send t) hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv t" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \)" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = fv t" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \)" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" by auto hence **: "(P x = fv t \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv t" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \)" hence "fv t \ \(set (map P S)) \ V" using P_subset by auto hence "fv (t \ \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" unfolding vars\<^sub>s\<^sub>t_def using P subst_apply_fv_subset_strand_trm assms by blast hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \)\ by force } ultimately show ?thesis by metis next case (Receive t) hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv t" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \)" "fv\<^sub>r\<^sub>c\<^sub>v x = fv t" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \)" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" by auto hence **: "(P x = fv t \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv t" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \)" hence "fv t \ \(set (map P S)) \ V" using P_subset by auto hence "fv (t \ \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" unfolding vars\<^sub>s\<^sub>t_def using P subst_apply_fv_subset_strand_trm assms by blast hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \)\ by blast } ultimately show ?thesis by metis next case (Equality ac' t t') show ?thesis proof (cases "ac' = ac") case True hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv t \ fv t'" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = fv t \ fv t'" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" using Equality by auto hence **: "(P x = fv t \ fv t' \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv t \ fv t'" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" hence "fv t \ \(set (map P S)) \ V" "fv t' \ \(set (map P S)) \ V" using P_subset by auto hence "fv (t \ \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" "fv (t' \ \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" unfolding vars\<^sub>s\<^sub>t_def using P subst_apply_fv_subset_strand_trm assms by metis+ hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)\ by blast } ultimately show ?thesis by metis next case False hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv t \ fv t'" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" using Equality by auto hence **: "(P x = fv t \ fv t' \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv t \ fv t'" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" hence "fv t \ \(set (map P S)) \ V" "fv t' \ \(set (map P S)) \ V" using P_subset by auto hence "fv (t \ \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" "fv (t' \ \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" unfolding vars\<^sub>s\<^sub>t_def using P subst_apply_fv_subset_strand_trm assms by metis+ hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)\ by blast } ultimately show ?thesis by metis qed next case (Inequality X F) hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" using \(2) ineq_apply_subst[of \ X F] by force+ hence **: "(P x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ \(set (map P S)) \ V" using P_subset by auto hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" proof (induction F) case (Cons f G) hence IH: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" by (metis (no_types, lifting) Diff_subset_conv UN_insert le_sup_iff list.simps(15) fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.simps) obtain t t' where f: "f = (t,t')" by (metis surj_pair) hence "fv t \ \(set (map P S)) \ (V \ set X)" "fv t' \ \(set (map P S)) \ (V \ set X)" using Cons.prems by auto hence "fv (t \ \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" "fv (t' \ \) \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" using subst_apply_fv_subset_strand_trm[OF P _ assms(3)] by blast+ thus ?case using f IH by (auto simp add: subst_apply_pairs_def) qed (simp add: subst_apply_pairs_def) moreover have "fv\<^sub>s\<^sub>e\<^sub>t (\ ` set X) = set X" using assms(4) Inequality by force ultimately have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X \ \(set (map P (S \\<^sub>s\<^sub>t \))) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" by auto hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X\ by blast } ultimately show ?thesis by metis qed lemma subst_apply_fv_subset_strand2: assumes P: "P = fv\<^sub>s\<^sub>t\<^sub>p \ P = fv\<^sub>r\<^sub>c\<^sub>v \ P = fv\<^sub>s\<^sub>n\<^sub>d \ P = fv\<^sub>e\<^sub>q ac \ P = fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q \ P = fv_r\<^sub>e\<^sub>q ac" and P_subset: "P x \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" and \: "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" "set (bvars\<^sub>s\<^sub>t\<^sub>p x) \ (subst_domain \ \ range_vars \) = {}" shows "P (x \\<^sub>s\<^sub>t\<^sub>p \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" proof (cases x) case (Send t) hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv t" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \)" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = fv t" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \)" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv_r\<^sub>e\<^sub>q ac x = {}" "fv_r\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" by auto hence **: "(P x = fv t \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv t" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \)" hence "fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using P_subset by auto hence "fv (t \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using P subst_apply_fv_subset_strand_trm2 assms by blast hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \)\ by blast } ultimately show ?thesis by metis next case (Receive t) hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv t" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \)" "fv\<^sub>r\<^sub>c\<^sub>v x = fv t" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \)" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv_r\<^sub>e\<^sub>q ac x = {}" "fv_r\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" by auto hence **: "(P x = fv t \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv t" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \)" hence "fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using P_subset by auto hence "fv (t \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using P subst_apply_fv_subset_strand_trm2 assms by blast hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \)\ by blast } ultimately show ?thesis by metis next case (Equality ac' t t') show ?thesis proof (cases "ac' = ac") case True hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv t \ fv t'" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = fv t \ fv t'" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv_r\<^sub>e\<^sub>q ac x = fv t'" "fv_r\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t' \ \)" using Equality by auto hence **: "(P x = fv t \ fv t' \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}) \ (P x = fv t' \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t' \ \))" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv t \ fv t'" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" hence "fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using P_subset by auto hence "fv (t \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" "fv (t' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using P subst_apply_fv_subset_strand_trm2 assms by blast+ hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)\ by blast } moreover { assume "P x = fv t'" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t' \ \)" hence "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using P_subset by auto hence "fv (t' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using P subst_apply_fv_subset_strand_trm2 assms by blast+ hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t' \ \)\ by blast } ultimately show ?thesis by metis next case False hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv t \ fv t'" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv_r\<^sub>e\<^sub>q ac x = {}" "fv_r\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" using Equality by auto hence **: "(P x = fv t \ fv t' \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}) \ (P x = fv t' \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t' \ \))" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv t \ fv t'" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" hence "fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using P_subset by auto hence "fv (t \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" "fv (t' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using P subst_apply_fv_subset_strand_trm2 assms by blast+ hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)\ by blast } moreover { assume "P x = fv t'" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t' \ \)" hence "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using P_subset by auto hence "fv (t' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using P subst_apply_fv_subset_strand_trm2 assms by blast+ hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv (t' \ \)\ by blast } ultimately show ?thesis by metis qed next case (Inequality X F) hence *: "fv\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" "fv\<^sub>r\<^sub>c\<^sub>v x = {}" "fv\<^sub>r\<^sub>c\<^sub>v (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>s\<^sub>n\<^sub>d x = {}" "fv\<^sub>s\<^sub>n\<^sub>d (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>e\<^sub>q ac x = {}" "fv\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" "fv\<^sub>i\<^sub>n\<^sub>e\<^sub>q (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" "fv_r\<^sub>e\<^sub>q ac x = {}" "fv_r\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" using \(2) ineq_apply_subst[of \ X F] by force+ hence **: "(P x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X) \ (P x = {} \ P (x \\<^sub>s\<^sub>t\<^sub>p \) = {})" by (metis P) moreover { assume "P x = {}" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = {}" hence ?thesis by simp } moreover { assume "P x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" "P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X" hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using P_subset by auto hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" proof (induction F) case (Cons f G) hence IH: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" by (metis (no_types, lifting) Diff_subset_conv UN_insert le_sup_iff list.simps(15) fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s.simps) obtain t t' where f: "f = (t,t')" by (metis surj_pair) hence "fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ (V \ set X)" "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ (V \ set X)" using Cons.prems by auto hence "fv (t \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" "fv (t' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ set X))" using subst_apply_fv_subset_strand_trm2[OF _ assms(3)] P by blast+ thus ?case using f IH by (auto simp add: subst_apply_pairs_def) qed (simp add: subst_apply_pairs_def) moreover have "fv\<^sub>s\<^sub>e\<^sub>t (\ ` set X) = set X" using assms(4) Inequality by force ultimately have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" by fastforce hence ?thesis using \P (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) - set X\ by blast } ultimately show ?thesis by metis qed lemma strand_subst_fv_bounded_if_img_bounded: assumes "range_vars \ \ fv\<^sub>s\<^sub>t S" shows "fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t S" using subst_sends_strand_fv_to_img[of S \] assms by blast lemma strand_fv_subst_subset_if_subst_elim: assumes "subst_elim \ v" and "v \ fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "v \ fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" proof (cases "v \ fv\<^sub>s\<^sub>t S") case True thus ?thesis proof (induction S) case (Cons x S) have *: "v \ fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \)" using assms(1) proof (cases x) case (Inequality X F) hence "subst_elim (rm_vars (set X) \) v \ v \ set X" using assms(1) by blast moreover have "fv\<^sub>s\<^sub>t\<^sub>p (Inequality X F \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) - set X" using Inequality by auto ultimately have "v \ fv\<^sub>s\<^sub>t\<^sub>p (Inequality X F \\<^sub>s\<^sub>t\<^sub>p \)" by (induct F) (auto simp add: subst_elim_def subst_apply_pairs_def) thus ?thesis using Inequality by simp qed (simp_all add: subst_elim_def) moreover have "v \ fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" using Cons.IH proof (cases "v \ fv\<^sub>s\<^sub>t S") case False moreover have "v \ range_vars \" by (simp add: subst_elimD''[OF assms(1)] range_vars_alt_def) ultimately show ?thesis by (meson UnE subsetCE subst_sends_strand_fv_to_img) qed simp ultimately show ?case by auto qed simp next case False thus ?thesis using assms fv_strand_subst' unfolding subst_elim_def by (metis (mono_tags, hide_lams) fv\<^sub>s\<^sub>e\<^sub>t.simps imageE mem_simps(8) subst_apply_term.simps(1)) qed lemma strand_fv_subst_subset_if_subst_elim': assumes "subst_elim \ v" "v \ fv\<^sub>s\<^sub>t S" "range_vars \ \ fv\<^sub>s\<^sub>t S" shows "fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t S" using strand_fv_subst_subset_if_subst_elim[OF assms(1)] assms(2) strand_subst_fv_bounded_if_img_bounded[OF assms(3)] by blast lemma fv_ik_is_fv_rcv: "fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t S) = \(set (map fv\<^sub>r\<^sub>c\<^sub>v S))" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma fv_ik_subset_fv_st[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t S) \ wfrestrictedvars\<^sub>s\<^sub>t S" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma fv_assignment_rhs_subset_fv_st[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (assignment_rhs\<^sub>s\<^sub>t S) \ wfrestrictedvars\<^sub>s\<^sub>t S" by (induct S rule: assignment_rhs\<^sub>s\<^sub>t.induct) force+ lemma fv_ik_subset_fv_st'[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t S) \ fv\<^sub>s\<^sub>t S" by (induct S rule: ik\<^sub>s\<^sub>t.induct) auto lemma ik\<^sub>s\<^sub>t_var_is_fv: "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t A) \ x \ fv\<^sub>s\<^sub>t A" by (meson fv_ik_subset_fv_st'[of A] fv_subset_subterms subsetCE term.set_intros(3)) lemma fv_assignment_rhs_subset_fv_st'[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (assignment_rhs\<^sub>s\<^sub>t S) \ fv\<^sub>s\<^sub>t S" by (induct S rule: assignment_rhs\<^sub>s\<^sub>t.induct) auto lemma ik\<^sub>s\<^sub>t_assignment_rhs\<^sub>s\<^sub>t_wfrestrictedvars_subset: "fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>s\<^sub>t A) \ wfrestrictedvars\<^sub>s\<^sub>t A" using fv_ik_subset_fv_st[of A] fv_assignment_rhs_subset_fv_st[of A] by simp+ lemma strand_step_id_subst[iff]: "x \\<^sub>s\<^sub>t\<^sub>p Var = x" by (cases x) auto lemma strand_id_subst[iff]: "S \\<^sub>s\<^sub>t Var = S" using strand_step_id_subst by (induct S) auto lemma strand_subst_vars_union_bound[simp]: "vars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ vars\<^sub>s\<^sub>t S \ range_vars \" proof (induction S) case (Cons x S) moreover have "vars\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ vars\<^sub>s\<^sub>t\<^sub>p x \ range_vars \" using subst_sends_fv_to_img[of _ \] proof (cases x) case (Inequality X F) define \' where "\' \ rm_vars (set X) \" have 0: "range_vars \' \ range_vars \" using rm_vars_img[of "set X" \] by (auto simp add: \'_def subst_domain_def range_vars_alt_def) have "vars\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \') \ set X" "vars\<^sub>s\<^sub>t\<^sub>p x = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ set X" using Inequality by (auto simp add: \'_def) moreover have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \') \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ range_vars \" proof (induction F) case (Cons f G) obtain t t' where f: "f = (t,t')" by moura hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (f#G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \') = fv (t \ \') \ fv (t' \ \') \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \')" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (f#G) = fv t \ fv t' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" by (auto simp add: subst_apply_pairs_def) thus ?case using 0 Cons.IH subst_sends_fv_to_img[of t \'] subst_sends_fv_to_img[of t' \'] unfolding f by auto qed (simp add: subst_apply_pairs_def) ultimately show ?thesis by auto qed auto ultimately show ?case by auto qed simp lemma strand_vars_split: "vars\<^sub>s\<^sub>t (S@S') = vars\<^sub>s\<^sub>t S \ vars\<^sub>s\<^sub>t S'" "wfrestrictedvars\<^sub>s\<^sub>t (S@S') = wfrestrictedvars\<^sub>s\<^sub>t S \ wfrestrictedvars\<^sub>s\<^sub>t S'" "fv\<^sub>s\<^sub>t (S@S') = fv\<^sub>s\<^sub>t S \ fv\<^sub>s\<^sub>t S'" by auto lemma bvars_subst_ident: "bvars\<^sub>s\<^sub>t S = bvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" unfolding bvars\<^sub>s\<^sub>t_def by (induct S) (simp_all add: subst_apply_strand_step_def split: strand_step.splits) lemma strand_subst_subst_idem: assumes "subst_idem \" "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t S" "subst_domain \ \ fv\<^sub>s\<^sub>t S = {}" "range_vars \ \ bvars\<^sub>s\<^sub>t S = {}" "range_vars \ \ bvars\<^sub>s\<^sub>t S = {}" shows "(S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \ = (S \\<^sub>s\<^sub>t \)" and "(S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t (\ \\<^sub>s \) = (S \\<^sub>s\<^sub>t \)" proof - from assms(2,3) have "fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ subst_domain \ = {}" using subst_sends_strand_fv_to_img[of S \] by blast thus "(S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \ = (S \\<^sub>s\<^sub>t \)" by blast thus "(S \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t (\ \\<^sub>s \) = (S \\<^sub>s\<^sub>t \)" by (metis assms(1,4,5) bvars_subst_ident strand_subst_comp subst_idem_def) qed lemma strand_subst_img_bound: assumes "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t S" and "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "range_vars \ \ fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" proof - have "subst_domain \ \ \(set (map fv\<^sub>s\<^sub>t\<^sub>p S))" by (metis (no_types) fv\<^sub>s\<^sub>t_def Un_subset_iff assms(1)) thus ?thesis unfolding range_vars_alt_def fv\<^sub>s\<^sub>t_def by (metis subst_range.simps fv_set_mono fv_strand_subst Int_commute assms(2) image_Un le_iff_sup) qed lemma strand_subst_img_bound': assumes "subst_domain \ \ range_vars \ \ vars\<^sub>s\<^sub>t S" and "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "range_vars \ \ vars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" proof - have "(subst_domain \ \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` subst_domain \)) \ vars\<^sub>s\<^sub>t S = subst_domain \ \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` subst_domain \)" using assms(1) by (metis inf.absorb_iff1 range_vars_alt_def subst_range.simps) hence "range_vars \ \ fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" using vars_snd_rcv_strand fv_snd_rcv_strand assms(2) strand_subst_img_bound unfolding range_vars_alt_def by (metis (no_types) inf_le2 inf_sup_distrib1 subst_range.simps sup_bot.right_neutral) thus "range_vars \ \ vars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" by (metis fv_snd_rcv_strand le_supI1 vars_snd_rcv_strand) qed lemma strand_subst_all_fv_subset: assumes "fv t \ fv\<^sub>s\<^sub>t S" "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "fv (t \ \) \ fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" using assms by (metis fv_strand_subst' Int_commute subst_apply_fv_subset) lemma strand_subst_not_dom_fixed: assumes "v \ fv\<^sub>s\<^sub>t S" and "v \ subst_domain \" shows "v \ fv\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" using assms proof (induction S) case (Cons x S') have 1: "\X. v \ subst_domain (rm_vars (set X) \)" using Cons.prems(2) rm_vars_dom_subset by force show ?case proof (cases "v \ fv\<^sub>s\<^sub>t S'") case True thus ?thesis using Cons.IH[OF _ Cons.prems(2)] by auto next case False hence 2: "v \ fv\<^sub>s\<^sub>t\<^sub>p x" using Cons.prems(1) by simp hence "v \ fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \)" using Cons.prems(2) subst_not_dom_fixed proof (cases x) case (Inequality X F) hence "v \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" using 2 by simp hence "v \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" using subst_not_dom_fixed[OF _ 1] by (induct F) (auto simp add: subst_apply_pairs_def) thus ?thesis using Inequality 2 by auto qed (force simp add: subst_domain_def)+ thus ?thesis by auto qed qed simp lemma strand_vars_unfold: "v \ vars\<^sub>s\<^sub>t S \ \S' x S''. S = S'@x#S'' \ v \ vars\<^sub>s\<^sub>t\<^sub>p x" proof (induction S) case (Cons x S) thus ?case proof (cases "v \ vars\<^sub>s\<^sub>t\<^sub>p x") case True thus ?thesis by blast next case False hence "v \ vars\<^sub>s\<^sub>t S" using Cons.prems by auto thus ?thesis using Cons.IH by (metis append_Cons) qed qed simp lemma strand_fv_unfold: "v \ fv\<^sub>s\<^sub>t S \ \S' x S''. S = S'@x#S'' \ v \ fv\<^sub>s\<^sub>t\<^sub>p x" proof (induction S) case (Cons x S) thus ?case proof (cases "v \ fv\<^sub>s\<^sub>t\<^sub>p x") case True thus ?thesis by blast next case False hence "v \ fv\<^sub>s\<^sub>t S" using Cons.prems by auto thus ?thesis using Cons.IH by (metis append_Cons) qed qed simp lemma subterm_if_in_strand_ik: "t \ ik\<^sub>s\<^sub>t S \ \t'. Receive t' \ set S \ t \ t'" by (induct S rule: ik\<^sub>s\<^sub>t_induct) auto lemma fv_subset_if_in_strand_ik: "t \ ik\<^sub>s\<^sub>t S \ fv t \ \(set (map fv\<^sub>r\<^sub>c\<^sub>v S))" proof - assume "t \ ik\<^sub>s\<^sub>t S" then obtain t' where "Receive t' \ set S" "t \ t'" by (metis subterm_if_in_strand_ik) hence "fv t \ fv t'" by (simp add: subtermeq_vars_subset) thus ?thesis using in_strand_fv_subset_rcv[OF \Receive t' \ set S\] by auto qed lemma fv_subset_if_in_strand_ik': "t \ ik\<^sub>s\<^sub>t S \ fv t \ fv\<^sub>s\<^sub>t S" using fv_subset_if_in_strand_ik[of t S] fv_snd_rcv_strand_subset(2)[of S] by blast lemma vars_subset_if_in_strand_ik2: "t \ ik\<^sub>s\<^sub>t S \ fv t \ wfrestrictedvars\<^sub>s\<^sub>t S" using fv_subset_if_in_strand_ik[of t S] vars_snd_rcv_strand_subset2(2)[of S] by blast subsection \Lemmata: Simple Strands\ lemma simple_Cons[dest]: "simple (s#S) \ simple S" unfolding simple_def by auto lemma simple_split[dest]: assumes "simple (S@S')" shows "simple S" "simple S'" using assms unfolding simple_def by auto lemma simple_append[intro]: "\simple S; simple S'\ \ simple (S@S')" unfolding simple_def by auto lemma simple_append_sym[sym]: "simple (S@S') \ simple (S'@S)" by auto lemma not_simple_if_snd_fun: "(\S' S'' f X. S = S'@Send (Fun f X)#S'') \ \simple S" unfolding simple_def by auto lemma not_list_all_elim: "\list_all P A \ \B x C. A = B@x#C \ \P x \ list_all P B" proof (induction A rule: List.rev_induct) case (snoc a A) show ?case proof (cases "list_all P A") case True thus ?thesis using snoc.prems by auto next case False then obtain B x C where "A = B@x#C" "\P x" "list_all P B" using snoc.IH[OF False] by auto thus ?thesis by auto qed qed simp lemma not_simple\<^sub>s\<^sub>t\<^sub>p_elim: assumes "\simple\<^sub>s\<^sub>t\<^sub>p x" shows "(\f T. x = Send (Fun f T)) \ (\a t t'. x = Equality a t t') \ (\X F. x = Inequality X F \ \(\\. ineq_model \ X F))" using assms by (cases x) (fastforce elim: simple\<^sub>s\<^sub>t\<^sub>p.elims)+ lemma not_simple_elim: assumes "\simple S" shows "(\A B f T. S = A@Send (Fun f T)#B \ simple A) \ (\A B a t t'. S = A@Equality a t t'#B \ simple A) \ (\A B X F. S = A@Inequality X F#B \ \(\\. ineq_model \ X F))" by (metis assms not_list_all_elim not_simple\<^sub>s\<^sub>t\<^sub>p_elim simple_def) lemma simple_fun_prefix_unique: assumes "A = S@Send (Fun f X)#S'" "simple S" shows "\T g Y T'. A = T@Send (Fun g Y)#T' \ simple T \ S = T \ f = g \ X = Y \ S' = T'" proof - { fix T g Y T' assume *: "A = T@Send (Fun g Y)#T'" "simple T" { assume "length S < length T" hence False using assms * by (metis id_take_nth_drop not_simple_if_snd_fun nth_append nth_append_length) } moreover { assume "length S > length T" hence False using assms * by (metis id_take_nth_drop not_simple_if_snd_fun nth_append nth_append_length) } ultimately have "S = T" using assms * by (meson List.append_eq_append_conv linorder_neqE_nat) } thus ?thesis using assms(1) by blast qed lemma simple_snd_is_var: "\Send t \ set S; simple S\ \ \v. t = Var v" unfolding simple_def by (metis list_all_append list_all_simps(1) simple\<^sub>s\<^sub>t\<^sub>p.elims(2) split_list_first strand_step.distinct(1) strand_step.distinct(5) strand_step.inject(1)) subsection \Lemmata: Strand Measure\ lemma measure\<^sub>s\<^sub>t_wellfounded: "wf measure\<^sub>s\<^sub>t" unfolding measure\<^sub>s\<^sub>t_def by simp lemma strand_size_append[iff]: "size\<^sub>s\<^sub>t (S@S') = size\<^sub>s\<^sub>t S + size\<^sub>s\<^sub>t S'" by (induct S) (auto simp add: size\<^sub>s\<^sub>t_def) lemma strand_size_map_fun_lt[simp]: "size\<^sub>s\<^sub>t (map Send X) < size (Fun f X)" "size\<^sub>s\<^sub>t (map Send X) < size\<^sub>s\<^sub>t [Send (Fun f X)]" "size\<^sub>s\<^sub>t (map Send X) < size\<^sub>s\<^sub>t [Receive (Fun f X)]" by (induct X) (auto simp add: size\<^sub>s\<^sub>t_def) lemma strand_size_rm_fun_lt[simp]: "size\<^sub>s\<^sub>t (S@S') < size\<^sub>s\<^sub>t (S@Send (Fun f X)#S')" "size\<^sub>s\<^sub>t (S@S') < size\<^sub>s\<^sub>t (S@Receive (Fun f X)#S')" by (induct S) (auto simp add: size\<^sub>s\<^sub>t_def) lemma strand_fv_card_map_fun_eq: "card (fv\<^sub>s\<^sub>t (S@Send (Fun f X)#S')) = card (fv\<^sub>s\<^sub>t (S@(map Send X)@S'))" proof - have "fv\<^sub>s\<^sub>t (S@Send (Fun f X)#S') = fv\<^sub>s\<^sub>t (S@(map Send X)@S')" by auto thus ?thesis by simp qed lemma strand_fv_card_rm_fun_le[simp]: "card (fv\<^sub>s\<^sub>t (S@S')) \ card (fv\<^sub>s\<^sub>t (S@Send (Fun f X)#S'))" by (force intro: card_mono) lemma strand_fv_card_rm_eq_le[simp]: "card (fv\<^sub>s\<^sub>t (S@S')) \ card (fv\<^sub>s\<^sub>t (S@Equality a t t'#S'))" by (force intro: card_mono) subsection \Lemmata: Well-formed Strands\ lemma wf_prefix[dest]: "wf\<^sub>s\<^sub>t V (S@S') \ wf\<^sub>s\<^sub>t V S" by (induct S rule: wf\<^sub>s\<^sub>t.induct) auto lemma wf_vars_mono[simp]: "wf\<^sub>s\<^sub>t V S \ wf\<^sub>s\<^sub>t (V \ W) S" proof (induction S arbitrary: V) case (Cons x S) thus ?case proof (cases x) case (Send t) hence "wf\<^sub>s\<^sub>t (V \ fv t \ W) S" using Cons.prems(1) Cons.IH by simp thus ?thesis using Send by (simp add: sup_commute sup_left_commute) next case (Equality a t t') show ?thesis proof (cases a) case Assign hence "wf\<^sub>s\<^sub>t (V \ fv t \ W) S" "fv t' \ V \ W" using Equality Cons.prems(1) Cons.IH by auto thus ?thesis using Equality Assign by (simp add: sup_commute sup_left_commute) next case Check thus ?thesis using Equality Cons by auto qed qed auto qed simp lemma wf\<^sub>s\<^sub>tI[intro]: "wfrestrictedvars\<^sub>s\<^sub>t S \ V \ wf\<^sub>s\<^sub>t V S" proof (induction S) case (Cons x S) thus ?case proof (cases x) case (Send t) hence "wf\<^sub>s\<^sub>t V S" "V \ fv t = V" using Cons by auto thus ?thesis using Send by simp next case (Equality a t t') show ?thesis proof (cases a) case Assign hence "wf\<^sub>s\<^sub>t V S" "fv t' \ V" using Equality Cons by auto thus ?thesis using wf_vars_mono Equality Assign by simp next case Check thus ?thesis using Equality Cons by auto qed qed simp_all qed simp lemma wf\<^sub>s\<^sub>tI'[intro]: "\(fv\<^sub>r\<^sub>c\<^sub>v ` set S) \ \(fv_r\<^sub>e\<^sub>q assign ` set S) \ V \ wf\<^sub>s\<^sub>t V S" proof (induction S) case (Cons x S) thus ?case proof (cases x) case (Equality a t t') thus ?thesis using Cons by (cases a) auto qed simp_all qed simp lemma wf_append_exec: "wf\<^sub>s\<^sub>t V (S@S') \ wf\<^sub>s\<^sub>t (V \ wfvarsoccs\<^sub>s\<^sub>t S) S'" proof (induction S arbitrary: V) case (Cons x S V) thus ?case proof (cases x) case (Send t) hence "wf\<^sub>s\<^sub>t (V \ fv t \ wfvarsoccs\<^sub>s\<^sub>t S) S'" using Cons.prems Cons.IH by simp thus ?thesis using Send by (auto simp add: sup_assoc) next case (Equality a t t') show ?thesis proof (cases a) case Assign hence "wf\<^sub>s\<^sub>t (V \ fv t \ wfvarsoccs\<^sub>s\<^sub>t S) S'" using Equality Cons.prems Cons.IH by auto thus ?thesis using Equality Assign by (auto simp add: sup_assoc) next case Check hence "wf\<^sub>s\<^sub>t (V \ wfvarsoccs\<^sub>s\<^sub>t S) S'" using Equality Cons.prems Cons.IH by auto thus ?thesis using Equality Check by (auto simp add: sup_assoc) qed qed auto qed simp lemma wf_append_suffix: "wf\<^sub>s\<^sub>t V S \ wfrestrictedvars\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V \ wf\<^sub>s\<^sub>t V (S@S')" proof (induction V S rule: wf\<^sub>s\<^sub>t_induct) case (ConsSnd V t S) hence *: "wf\<^sub>s\<^sub>t (V \ fv t) S" by simp_all hence "wfrestrictedvars\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>t S \ (V \ fv t)" using ConsSnd.prems(2) by fastforce thus ?case using ConsSnd.IH * by simp next case (ConsRcv V t S) hence *: "fv t \ V" "wf\<^sub>s\<^sub>t V S" by simp_all hence "wfrestrictedvars\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" using ConsRcv.prems(2) by fastforce thus ?case using ConsRcv.IH * by simp next case (ConsEq V t t' S) hence *: "fv t' \ V" "wf\<^sub>s\<^sub>t (V \ fv t) S" by simp_all moreover have "vars\<^sub>s\<^sub>t\<^sub>p (Equality Assign t t') = fv t \ fv t'" by simp moreover have "wfrestrictedvars\<^sub>s\<^sub>t (Equality Assign t t'#S) = fv t \ fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S" by auto ultimately have "wfrestrictedvars\<^sub>s\<^sub>t S' \ wfrestrictedvars\<^sub>s\<^sub>t S \ (V \ fv t)" using ConsEq.prems(2) by blast thus ?case using ConsEq.IH * by simp qed (simp_all add: wf\<^sub>s\<^sub>tI) lemma wf_append_suffix': assumes "wf\<^sub>s\<^sub>t V S" and "\(fv\<^sub>r\<^sub>c\<^sub>v ` set S') \ \(fv_r\<^sub>e\<^sub>q assign ` set S') \ wfvarsoccs\<^sub>s\<^sub>t S \ V" shows "wf\<^sub>s\<^sub>t V (S@S')" using assms proof (induction V S rule: wf\<^sub>s\<^sub>t_induct) case (ConsSnd V t S) hence *: "wf\<^sub>s\<^sub>t (V \ fv t) S" by simp_all have "wfvarsoccs\<^sub>s\<^sub>t (send\t\\<^sub>s\<^sub>t#S) = fv t \ wfvarsoccs\<^sub>s\<^sub>t S" unfolding wfvarsoccs\<^sub>s\<^sub>t_def by simp hence "(\a\set S'. fv\<^sub>r\<^sub>c\<^sub>v a) \ (\a\set S'. fv_r\<^sub>e\<^sub>q assign a) \ wfvarsoccs\<^sub>s\<^sub>t S \ (V \ fv t)" using ConsSnd.prems(2) unfolding wfvarsoccs\<^sub>s\<^sub>t_def by auto thus ?case using ConsSnd.IH[OF *] by auto next case (ConsEq V t t' S) hence *: "fv t' \ V" "wf\<^sub>s\<^sub>t (V \ fv t) S" by simp_all have "wfvarsoccs\<^sub>s\<^sub>t (\assign: t \ t'\\<^sub>s\<^sub>t#S) = fv t \ wfvarsoccs\<^sub>s\<^sub>t S" unfolding wfvarsoccs\<^sub>s\<^sub>t_def by simp hence "(\a\set S'. fv\<^sub>r\<^sub>c\<^sub>v a) \ (\a\set S'. fv_r\<^sub>e\<^sub>q assign a) \ wfvarsoccs\<^sub>s\<^sub>t S \ (V \ fv t)" using ConsEq.prems(2) unfolding wfvarsoccs\<^sub>s\<^sub>t_def by auto thus ?case using ConsEq.IH[OF *(2)] *(1) by auto qed (auto simp add: wf\<^sub>s\<^sub>tI') lemma wf_send_compose: "wf\<^sub>s\<^sub>t V (S@(map Send X)@S') = wf\<^sub>s\<^sub>t V (S@Send (Fun f X)#S')" proof (induction S arbitrary: V) case Nil thus ?case proof (induction X arbitrary: V) case (Cons y Y) thus ?case by (simp add: sup_assoc) qed simp next case (Cons s S) thus ?case proof (cases s) case (Equality ac t t') thus ?thesis using Cons by (cases ac) auto qed auto qed lemma wf_snd_append[iff]: "wf\<^sub>s\<^sub>t V (S@[Send t]) = wf\<^sub>s\<^sub>t V S" by (induct S rule: wf\<^sub>s\<^sub>t.induct) simp_all lemma wf_snd_append': "wf\<^sub>s\<^sub>t V S \ wf\<^sub>s\<^sub>t V (Send t#S)" by simp lemma wf_rcv_append[dest]: "wf\<^sub>s\<^sub>t V (S@Receive t#S') \ wf\<^sub>s\<^sub>t V (S@S')" by (induct S rule: wf\<^sub>s\<^sub>t.induct) simp_all lemma wf_rcv_append'[intro]: "\wf\<^sub>s\<^sub>t V (S@S'); fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ V\ \ wf\<^sub>s\<^sub>t V (S@Receive t#S')" proof (induction S rule: wf\<^sub>s\<^sub>t_induct) case (ConsRcv V t' S) hence "wf\<^sub>s\<^sub>t V (S@S')" "fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" by auto+ thus ?case using ConsRcv by auto next case (ConsEq V t' t'' S) hence "fv t'' \ V" by simp moreover have "wfrestrictedvars\<^sub>s\<^sub>t (Equality Assign t' t''#S) = fv t' \ fv t'' \ wfrestrictedvars\<^sub>s\<^sub>t S" by auto ultimately have "fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ (V \ fv t')" using ConsEq.prems(2) by blast thus ?case using ConsEq by auto qed auto lemma wf_rcv_append''[intro]: "\wf\<^sub>s\<^sub>t V S; fv t \ \(set (map fv\<^sub>s\<^sub>n\<^sub>d S))\ \ wf\<^sub>s\<^sub>t V (S@[Receive t])" by (induct S) (simp, metis vars_snd_rcv_strand_subset2(1) append_Nil2 le_supI1 order_trans wf_rcv_append') lemma wf_rcv_append'''[intro]: "\wf\<^sub>s\<^sub>t V S; fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ V\ \ wf\<^sub>s\<^sub>t V (S@[Receive t])" by (simp add: wf_rcv_append'[of _ _ "[]"]) lemma wf_eq_append[dest]: "wf\<^sub>s\<^sub>t V (S@Equality a t t'#S') \ fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ V \ wf\<^sub>s\<^sub>t V (S@S')" proof (induction S rule: wf\<^sub>s\<^sub>t_induct) case (Nil V) hence "wf\<^sub>s\<^sub>t (V \ fv t) S'" by (cases a) auto moreover have "V \ fv t = V" using Nil by auto ultimately show ?case by simp next case (ConsRcv V u S) hence "wf\<^sub>s\<^sub>t V (S @ Equality a t t' # S')" "fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" "fv u \ V" by fastforce+ hence "wf\<^sub>s\<^sub>t V (S@S')" using ConsRcv.IH by auto thus ?case using \fv u \ V\ by simp next case (ConsEq V u u' S) hence "wf\<^sub>s\<^sub>t (V \ fv u) (S@Equality a t t'#S')" "fv t \ wfrestrictedvars\<^sub>s\<^sub>t S \ (V \ fv u)" "fv u' \ V" by auto hence "wf\<^sub>s\<^sub>t (V \ fv u) (S@S')" using ConsEq.IH by auto thus ?case using \fv u' \ V\ by simp qed auto lemma wf_eq_append'[intro]: "\wf\<^sub>s\<^sub>t V (S@S'); fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V\ \ wf\<^sub>s\<^sub>t V (S@Equality a t t'#S')" proof (induction S rule: wf\<^sub>s\<^sub>t_induct) case Nil thus ?case by (cases a) auto next case (ConsEq V u u' S) hence "wf\<^sub>s\<^sub>t (V \ fv u) (S@S')" "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V \ fv u" by fastforce+ thus ?case using ConsEq by auto next case (ConsEq2 V u u' S) hence "wf\<^sub>s\<^sub>t V (S@S')" by auto thus ?case using ConsEq2 by auto next case (ConsRcv V u S) hence "wf\<^sub>s\<^sub>t V (S@S')" "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V" by fastforce+ thus ?case using ConsRcv by auto next case (ConsSnd V u S) hence "wf\<^sub>s\<^sub>t (V \ fv u) (S@S')" "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ (V \ fv u)" by fastforce+ thus ?case using ConsSnd by auto qed auto lemma wf_eq_append''[intro]: "\wf\<^sub>s\<^sub>t V (S@S'); fv t' \ wfvarsoccs\<^sub>s\<^sub>t S \ V\ \ wf\<^sub>s\<^sub>t V (S@[Equality a t t']@S')" proof (induction S rule: wf\<^sub>s\<^sub>t_induct) case Nil thus ?case by (cases a) auto next case (ConsEq V u u' S) hence "wf\<^sub>s\<^sub>t (V \ fv u) (S@S')" "fv t' \ wfvarsoccs\<^sub>s\<^sub>t S \ V \ fv u" by fastforce+ thus ?case using ConsEq by auto next case (ConsEq2 V u u' S) hence "wf\<^sub>s\<^sub>t (V \ fv u) (S@S')" "fv t' \ wfvarsoccs\<^sub>s\<^sub>t S \ V \ fv u" by fastforce+ thus ?case using ConsEq2 by auto next case (ConsRcv V u S) hence "wf\<^sub>s\<^sub>t V (S@S')" "fv t' \ wfvarsoccs\<^sub>s\<^sub>t S \ V" by fastforce+ thus ?case using ConsRcv by auto next case (ConsSnd V u S) hence "wf\<^sub>s\<^sub>t (V \ fv u) (S@S')" "fv t' \ wfvarsoccs\<^sub>s\<^sub>t S \ (V \ fv u)" by auto thus ?case using ConsSnd by auto qed auto lemma wf_eq_append'''[intro]: "\wf\<^sub>s\<^sub>t V S; fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S \ V\ \ wf\<^sub>s\<^sub>t V (S@[Equality a t t'])" by (simp add: wf_eq_append'[of _ _ "[]"]) lemma wf_eq_check_append[dest]: "wf\<^sub>s\<^sub>t V (S@Equality Check t t'#S') \ wf\<^sub>s\<^sub>t V (S@S')" by (induct S rule: wf\<^sub>s\<^sub>t.induct) simp_all lemma wf_eq_check_append'[intro]: "wf\<^sub>s\<^sub>t V (S@S') \ wf\<^sub>s\<^sub>t V (S@Equality Check t t'#S')" by (induct S rule: wf\<^sub>s\<^sub>t.induct) auto lemma wf_eq_check_append''[intro]: "wf\<^sub>s\<^sub>t V S \ wf\<^sub>s\<^sub>t V (S@[Equality Check t t'])" by (induct S rule: wf\<^sub>s\<^sub>t.induct) auto lemma wf_ineq_append[dest]: "wf\<^sub>s\<^sub>t V (S@Inequality X F#S') \ wf\<^sub>s\<^sub>t V (S@S')" by (induct S rule: wf\<^sub>s\<^sub>t.induct) simp_all lemma wf_ineq_append'[intro]: "wf\<^sub>s\<^sub>t V (S@S') \ wf\<^sub>s\<^sub>t V (S@Inequality X F#S')" by (induct S rule: wf\<^sub>s\<^sub>t.induct) auto lemma wf_ineq_append''[intro]: "wf\<^sub>s\<^sub>t V S \ wf\<^sub>s\<^sub>t V (S@[Inequality X F])" by (induct S rule: wf\<^sub>s\<^sub>t.induct) auto lemma wf_rcv_fv_single[elim]: "wf\<^sub>s\<^sub>t V (Receive t#S') \ fv t \ V" by simp lemma wf_rcv_fv: "wf\<^sub>s\<^sub>t V (S@Receive t#S') \ fv t \ wfvarsoccs\<^sub>s\<^sub>t S \ V" by (induct S arbitrary: V) (auto split!: strand_step.split poscheckvariant.split) lemma wf_eq_fv: "wf\<^sub>s\<^sub>t V (S@Equality Assign t t'#S') \ fv t' \ wfvarsoccs\<^sub>s\<^sub>t S \ V" by (induct S arbitrary: V) (auto split!: strand_step.split poscheckvariant.split) lemma wf_simple_fv_occurrence: assumes "wf\<^sub>s\<^sub>t {} S" "simple S" "v \ wfrestrictedvars\<^sub>s\<^sub>t S" shows "\S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f. S = S\<^sub>p\<^sub>r\<^sub>e@Send (Var v)#S\<^sub>s\<^sub>u\<^sub>f \ v \ wfrestrictedvars\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e" using assms proof (induction S rule: List.rev_induct) case (snoc x S) from \wf\<^sub>s\<^sub>t {} (S@[x])\ have "wf\<^sub>s\<^sub>t {} S" "wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>s\<^sub>t S) [x]" using wf_append_exec[THEN wf_vars_mono, of "{}" S "[x]" "wfrestrictedvars\<^sub>s\<^sub>t S - wfvarsoccs\<^sub>s\<^sub>t S"] vars_snd_rcv_strand_subset2(4)[of S] Diff_partition[of "wfvarsoccs\<^sub>s\<^sub>t S" "wfrestrictedvars\<^sub>s\<^sub>t S"] by auto from \simple (S@[x])\ have "simple S" "simple\<^sub>s\<^sub>t\<^sub>p x" unfolding simple_def by auto show ?case proof (cases "v \ wfrestrictedvars\<^sub>s\<^sub>t S") case False show ?thesis proof (cases x) case (Receive t) hence "fv t \ wfrestrictedvars\<^sub>s\<^sub>t S" using \wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>s\<^sub>t S) [x]\ by simp hence "v \ wfrestrictedvars\<^sub>s\<^sub>t S" using \v \ wfrestrictedvars\<^sub>s\<^sub>t (S@[x])\ \x = Receive t\ by auto thus ?thesis using \x = Receive t\ snoc.IH[OF \wf\<^sub>s\<^sub>t {} S\ \simple S\] by fastforce next case (Send t) hence "v \ vars\<^sub>s\<^sub>t\<^sub>p x" using \v \ wfrestrictedvars\<^sub>s\<^sub>t (S@[x])\ False by auto from Send obtain w where "t = Var w" using \simple\<^sub>s\<^sub>t\<^sub>p x\ by (cases t) simp_all hence "v = w" using \x = Send t\ \v \ vars\<^sub>s\<^sub>t\<^sub>p x\ by simp thus ?thesis using \x = Send t\ \v \ wfrestrictedvars\<^sub>s\<^sub>t S\ \t = Var w\ by auto next case (Equality ac t t') thus ?thesis using snoc.prems(2) unfolding simple_def by auto next case (Inequality t t') thus ?thesis using False snoc.prems(3) by auto qed qed (use snoc.IH[OF \wf\<^sub>s\<^sub>t {} S\ \simple S\] in fastforce) qed simp lemma Unifier_strand_fv_subset: assumes g_in_ik: "t \ ik\<^sub>s\<^sub>t S" and \: "Unifier \ (Fun f X) t" and disj: "bvars\<^sub>s\<^sub>t S \ (subst_domain \ \ range_vars \) = {}" shows "fv (Fun f X \ \) \ \(set (map fv\<^sub>r\<^sub>c\<^sub>v (S \\<^sub>s\<^sub>t \)))" by (metis (no_types) fv_subset_if_in_strand_ik[OF g_in_ik] disj \ fv_strand_subst subst_apply_fv_subset) lemma wf\<^sub>s\<^sub>t_induct'[consumes 1, case_names Nil ConsSnd ConsRcv ConsEq ConsEq2 ConsIneq]: fixes S::"('a,'b) strand" assumes "wf\<^sub>s\<^sub>t V S" "P []" "\t S. \wf\<^sub>s\<^sub>t V S; P S\ \ P (S@[Send t])" "\t S. \wf\<^sub>s\<^sub>t V S; P S; fv t \ V \ wfvarsoccs\<^sub>s\<^sub>t S\ \ P (S@[Receive t])" "\t t' S. \wf\<^sub>s\<^sub>t V S; P S; fv t' \ V \ wfvarsoccs\<^sub>s\<^sub>t S\ \ P (S@[Equality Assign t t'])" "\t t' S. \wf\<^sub>s\<^sub>t V S; P S\ \ P (S@[Equality Check t t'])" "\X F S. \wf\<^sub>s\<^sub>t V S; P S\ \ P (S@[Inequality X F])" shows "P S" using assms proof (induction S rule: List.rev_induct) case (snoc x S) hence *: "wf\<^sub>s\<^sub>t V S" "wf\<^sub>s\<^sub>t (V \ wfvarsoccs\<^sub>s\<^sub>t S) [x]" by (metis wf_prefix, metis wf_append_exec) have IH: "P S" using snoc.IH[OF *(1)] snoc.prems by auto note ** = snoc.prems(3,4,5,6,7)[OF *(1) IH] *(2) show ?case using **(1,2,4,5,6) proof (cases x) case (Equality ac t t') then show ?thesis using **(3,4,6) by (cases ac) auto qed auto qed simp lemma wf_subst_apply: "wf\<^sub>s\<^sub>t V S \ wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) (S \\<^sub>s\<^sub>t \)" proof (induction S arbitrary: V rule: wf\<^sub>s\<^sub>t_induct) case (ConsRcv V t S) hence "wf\<^sub>s\<^sub>t V S" "fv t \ V" by simp_all hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) (S \\<^sub>s\<^sub>t \)" "fv (t \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using ConsRcv.IH subst_apply_fv_subset by simp_all thus ?case by simp next case (ConsSnd V t S) hence "wf\<^sub>s\<^sub>t (V \ fv t) S" by simp hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ fv t))) (S \\<^sub>s\<^sub>t \)" using ConsSnd.IH by metis hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv (t \ \)) (S \\<^sub>s\<^sub>t \)" using subst_apply_fv_union by metis thus ?case by simp next case (ConsEq V t t' S) hence "wf\<^sub>s\<^sub>t (V \ fv t) S" "fv t' \ V" by auto hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ fv t))) (S \\<^sub>s\<^sub>t \)" and *: "fv (t' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using ConsEq.IH subst_apply_fv_subset by force+ hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv (t \ \)) (S \\<^sub>s\<^sub>t \)" using subst_apply_fv_union by metis thus ?case using * by simp qed simp_all lemma wf_unify: assumes wf: "wf\<^sub>s\<^sub>t V (S@Send (Fun f X)#S')" and g_in_ik: "t \ ik\<^sub>s\<^sub>t S" and \: "Unifier \ (Fun f X) t" and disj: "bvars\<^sub>s\<^sub>t (S@Send (Fun f X)#S') \ (subst_domain \ \ range_vars \) = {}" shows "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) ((S@S') \\<^sub>s\<^sub>t \)" using assms proof (induction S' arbitrary: V rule: List.rev_induct) case (snoc x S' V) have fun_fv_bound: "fv (Fun f X \ \) \ \(set (map fv\<^sub>r\<^sub>c\<^sub>v (S \\<^sub>s\<^sub>t \)))" using snoc.prems(4) bvars\<^sub>s\<^sub>t_split Unifier_strand_fv_subset[OF g_in_ik \] by auto hence "fv (Fun f X \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \))" using fv_ik_is_fv_rcv by metis hence "fv (Fun f X \ \) \ wfrestrictedvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" using fv_ik_subset_fv_st[of "S \\<^sub>s\<^sub>t \"] by blast hence *: "fv ((Fun f X) \ \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" by fastforce from snoc.prems(1) have "wf\<^sub>s\<^sub>t V (S@Send (Fun f X)#S')" using wf_prefix[of V "S@Send (Fun f X)#S'" "[x]"] by simp hence **: "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) ((S@S') \\<^sub>s\<^sub>t \)" using snoc.IH[OF _ snoc.prems(2,3)] snoc.prems(4) by auto from snoc.prems(1) have ***: "wf\<^sub>s\<^sub>t (V \ wfvarsoccs\<^sub>s\<^sub>t (S@Send (Fun f X)#S')) [x]" using wf_append_exec[of V "(S@Send (Fun f X)#S')" "[x]"] by simp from snoc.prems(4) have disj': "bvars\<^sub>s\<^sub>t (S@S') \ (subst_domain \ \ range_vars \) = {}" "set (bvars\<^sub>s\<^sub>t\<^sub>p x) \ (subst_domain \ \ range_vars \) = {}" by auto show ?case proof (cases x) case (Send t) thus ?thesis using wf_snd_append[of "fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" "(S@S') \\<^sub>s\<^sub>t \"] ** by auto next case (Receive t) hence "fv\<^sub>s\<^sub>t\<^sub>p x \ V \ wfvarsoccs\<^sub>s\<^sub>t (S@Send (Fun f X)#S')" using *** by auto hence "fv\<^sub>s\<^sub>t\<^sub>p x \ V \ wfrestrictedvars\<^sub>s\<^sub>t (S@Send (Fun f X)#S')" using vars_snd_rcv_strand_subset2(4)[of "S@Send (Fun f X)#S'"] by blast hence "fv\<^sub>s\<^sub>t\<^sub>p x \ V \ fv (Fun f X) \ wfrestrictedvars\<^sub>s\<^sub>t (S@S')" by auto hence "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv ((Fun f X) \ \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" by (metis (no_types) inf_sup_aci(5) subst_apply_fv_subset_strand2 subst_apply_fv_union disj') hence "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" using * by blast hence "fv (t \ \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) " using \x = Receive t\ by auto hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) (((S@S') \\<^sub>s\<^sub>t \)@[Receive (t \ \)])" using wf_rcv_append'''[OF **, of "t \ \"] by metis thus ?thesis using \x = Receive t\ by auto next case (Equality ac s s') show ?thesis proof (cases ac) case Assign hence "fv s' \ V \ wfvarsoccs\<^sub>s\<^sub>t (S@Send (Fun f X)#S')" using Equality *** by auto hence "fv s' \ V \ wfrestrictedvars\<^sub>s\<^sub>t (S@Send (Fun f X)#S')" using vars_snd_rcv_strand_subset2(4)[of "S@Send (Fun f X)#S'"] by blast hence "fv s' \ V \ fv (Fun f X) \ wfrestrictedvars\<^sub>s\<^sub>t (S@S')" by auto moreover have "fv s' = fv_r\<^sub>e\<^sub>q ac x" "fv (s' \ \) = fv_r\<^sub>e\<^sub>q ac (x \\<^sub>s\<^sub>t\<^sub>p \)" using Equality by simp_all ultimately have "fv (s' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv (Fun f X \ \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" using subst_apply_fv_subset_strand2[of "fv\<^sub>e\<^sub>q ac" ac x] by (metis disj'(1) subst_apply_fv_subset_strand_trm2 subst_apply_fv_union sup_commute) hence "fv (s' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" using * by blast hence "fv (s' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using \x = Equality ac s s'\ by auto hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) (((S@S') \\<^sub>s\<^sub>t \)@[Equality ac (s \ \) (s' \ \)])" using wf_eq_append'''[OF **] by metis thus ?thesis using \x = Equality ac s s'\ by auto next case Check thus ?thesis using wf_eq_check_append''[OF **] Equality by simp qed next case (Inequality t t') thus ?thesis using wf_ineq_append''[OF **] by simp qed qed (auto dest: wf_subst_apply) lemma wf_equality: assumes wf: "wf\<^sub>s\<^sub>t V (S@Equality ac t t'#S')" and \: "mgu t t' = Some \" and disj: "bvars\<^sub>s\<^sub>t (S@Equality ac t t'#S') \ (subst_domain \ \ range_vars \) = {}" shows "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) ((S@S') \\<^sub>s\<^sub>t \)" using assms proof (induction S' arbitrary: V rule: List.rev_induct) case Nil thus ?case using wf_prefix[of V S "[Equality ac t t']"] wf_subst_apply[of V S \] by auto next case (snoc x S' V) show ?case proof (cases ac) case Assign hence "fv t' \ V \ wfvarsoccs\<^sub>s\<^sub>t S" using wf_eq_fv[of V, of S t t' "S'@[x]"] snoc by auto hence "fv t' \ V \ wfrestrictedvars\<^sub>s\<^sub>t S" using vars_snd_rcv_strand_subset2(4)[of S] by blast hence "fv t' \ V \ wfrestrictedvars\<^sub>s\<^sub>t (S@S')" by force moreover have disj': "bvars\<^sub>s\<^sub>t (S@S') \ (subst_domain \ \ range_vars \) = {}" "set (bvars\<^sub>s\<^sub>t\<^sub>p x) \ (subst_domain \ \ range_vars \) = {}" "bvars\<^sub>s\<^sub>t (S@Equality ac t t'#S') \ (subst_domain \ \ range_vars \) = {}" using snoc.prems(3) by auto ultimately have "fv (t' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" by (metis inf_sup_aci(5) subst_apply_fv_subset_strand_trm2) moreover have "fv (t \ \) = fv (t' \ \)" by (metis MGU_is_Unifier[OF mgu_gives_MGU[OF \]]) ultimately have *: "fv (t \ \) \ fv (t' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" by simp from snoc.prems(1) have "wf\<^sub>s\<^sub>t V (S@Equality ac t t'#S')" using wf_prefix[of V "S@Equality ac t t'#S'"] by simp hence **: "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) ((S@S') \\<^sub>s\<^sub>t \)" by (metis snoc.IH \ disj'(3)) from snoc.prems(1) have ***: "wf\<^sub>s\<^sub>t (V \ wfvarsoccs\<^sub>s\<^sub>t (S@Equality ac t t'#S')) [x]" using wf_append_exec[of V "(S@Equality ac t t'#S')" "[x]"] by simp show ?thesis proof (cases x) case (Send t) thus ?thesis using wf_snd_append[of "fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" "(S@S') \\<^sub>s\<^sub>t \"] ** by auto next case (Receive s) hence "fv\<^sub>s\<^sub>t\<^sub>p x \ V \ wfvarsoccs\<^sub>s\<^sub>t (S@Equality ac t t'#S')" using *** by auto hence "fv\<^sub>s\<^sub>t\<^sub>p x \ V \ wfrestrictedvars\<^sub>s\<^sub>t (S@Equality ac t t'#S')" using vars_snd_rcv_strand_subset2(4)[of "S@Equality ac t t'#S'"] by blast hence "fv\<^sub>s\<^sub>t\<^sub>p x \ V \ fv t \ fv t' \ wfrestrictedvars\<^sub>s\<^sub>t (S@S')" by (cases ac) auto hence "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv (t \ \) \ fv (t' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" using subst_apply_fv_subset_strand2[of fv\<^sub>s\<^sub>t\<^sub>p] by (metis (no_types) inf_sup_aci(5) subst_apply_fv_union disj'(1,2)) hence "fv\<^sub>s\<^sub>t\<^sub>p (x \\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" when "ac = Assign" using * that by blast hence "fv (s \ \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V))" when "ac = Assign" using \x = Receive s\ that by auto hence "wf\<^sub>s\<^sub>t (fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)) (((S@S') \\<^sub>s\<^sub>t \)@[Receive (s \ \)])" when "ac = Assign" using wf_rcv_append'''[OF **, of "s \ \"] that by metis thus ?thesis using \x = Receive s\ Assign by auto next case (Equality ac' s s') show ?thesis proof (cases ac') case Assign hence "fv s' \ V \ wfvarsoccs\<^sub>s\<^sub>t (S@Equality ac t t'#S')" using *** Equality by auto hence "fv s' \ V \ wfrestrictedvars\<^sub>s\<^sub>t (S@Equality ac t t'#S')" using vars_snd_rcv_strand_subset2(4)[of "S@Equality ac t t'#S'"] by blast hence "fv s' \ V \ fv t \ fv t' \ wfrestrictedvars\<^sub>s\<^sub>t (S@S')" by (cases ac) auto moreover have "fv s' = fv_r\<^sub>e\<^sub>q ac' x" "fv (s' \ \) = fv_r\<^sub>e\<^sub>q ac' (x \\<^sub>s\<^sub>t\<^sub>p \)" using Equality by simp_all ultimately have "fv (s' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv (t \ \) \ fv (t' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" using subst_apply_fv_subset_strand2[of "fv_r\<^sub>e\<^sub>q ac'" ac' x] by (metis disj'(1) subst_apply_fv_subset_strand_trm2 subst_apply_fv_union sup_commute) hence "fv (s' \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" using * \ac = Assign\ by blast hence ****: "fv (s' \ \) \ wfrestrictedvars\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" using \x = Equality ac' s s'\ \ac = Assign\ by auto thus ?thesis using \x = Equality ac' s s'\ ** **** wf_eq_append' \ac = Assign\ by (metis (no_types, lifting) append.assoc append_Nil2 strand_step.case(3) strand_subst_hom subst_apply_strand_step_def) next case Check thus ?thesis using wf_eq_check_append''[OF **] Equality by simp qed next case (Inequality s s') thus ?thesis using wf_ineq_append''[OF **] by simp qed qed (metis snoc.prems(1) wf_eq_check_append wf_subst_apply) qed lemma wf_rcv_prefix_ground: "wf\<^sub>s\<^sub>t {} ((map Receive M)@S) \ vars\<^sub>s\<^sub>t (map Receive M) = {}" by (induct M) auto lemma simple_wfvarsoccs\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>n\<^sub>d: assumes "simple S" shows "wfvarsoccs\<^sub>s\<^sub>t S = \(set (map fv\<^sub>s\<^sub>n\<^sub>d S))" using assms unfolding simple_def proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp lemma wf\<^sub>s\<^sub>t_simple_induct[consumes 2, case_names Nil ConsSnd ConsRcv ConsIneq]: fixes S::"('a,'b) strand" assumes "wf\<^sub>s\<^sub>t V S" "simple S" "P []" "\v S. \wf\<^sub>s\<^sub>t V S; simple S; P S\ \ P (S@[Send (Var v)])" "\t S. \wf\<^sub>s\<^sub>t V S; simple S; P S; fv t \ V \ \(set (map fv\<^sub>s\<^sub>n\<^sub>d S))\ \ P (S@[Receive t])" "\X F S. \wf\<^sub>s\<^sub>t V S; simple S; P S\ \ P (S@[Inequality X F])" shows "P S" using assms proof (induction S rule: wf\<^sub>s\<^sub>t_induct') case (ConsSnd t S) hence "P S" by auto obtain v where "t = Var v" using simple_snd_is_var[OF _ \simple (S@[Send t])\] by auto thus ?case using ConsSnd.prems(3)[OF \wf\<^sub>s\<^sub>t V S\ _ \P S\] \simple (S@[Send t])\ by auto next case (ConsRcv t S) thus ?case using simple_wfvarsoccs\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>n\<^sub>d[of "S@[Receive t]"] by auto qed (auto simp add: simple_def) lemma wf_trm_stp_dom_fv_disjoint: "\wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \; t \ trms\<^sub>s\<^sub>t S\ \ subst_domain \ \ fv t = {}" unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by force lemma wf_constr_bvars_disj: "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \ \ (subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" unfolding range_vars_alt_def wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by fastforce lemma wf_constr_bvars_disj': assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t S" shows "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" (is ?A) and "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) = {}" (is ?B) proof - have "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" using assms(1) unfolding range_vars_alt_def wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by fastforce+ thus ?A and ?B using assms(2) bvars_subst_ident[of S \] by blast+ qed lemma (in intruder_model) wf_simple_strand_first_Send_var_split: assumes "wf\<^sub>s\<^sub>t {} S" "simple S" "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ = \ v" shows "\v S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f. S = S\<^sub>p\<^sub>r\<^sub>e@Send (Var v)#S\<^sub>s\<^sub>u\<^sub>f \ t \ \ = \ v \ \(\w \ wfrestrictedvars\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e. t \ \ = \ w)" (is "?P S") using assms proof (induction S rule: wf\<^sub>s\<^sub>t_simple_induct) case (ConsSnd v S) show ?case proof (cases "\w \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ = \ w") case True thus ?thesis using ConsSnd.IH by fastforce next case False thus ?thesis using ConsSnd.prems by auto qed next case (ConsRcv t' S) have "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S" using ConsRcv.hyps(3) vars_snd_rcv_strand_subset2(1) by force hence "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ = \ v" using ConsRcv.prems(1) by fastforce hence "?P S" by (metis ConsRcv.IH) thus ?case by fastforce next case (ConsIneq X F S) moreover have "wfrestrictedvars\<^sub>s\<^sub>t (S @ [Inequality X F]) = wfrestrictedvars\<^sub>s\<^sub>t S" by auto ultimately have "?P S" by blast thus ?case by fastforce qed simp lemma (in intruder_model) wf_strand_first_Send_var_split: assumes "wf\<^sub>s\<^sub>t {} S" "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ \ \ v" shows "\S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f. \(\w \ wfrestrictedvars\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e. t \ \ \ \ w) \ ((\t'. S = S\<^sub>p\<^sub>r\<^sub>e@Send t'#S\<^sub>s\<^sub>u\<^sub>f \ t \ \ \ t' \ \) \ (\t' t''. S = S\<^sub>p\<^sub>r\<^sub>e@Equality Assign t' t''#S\<^sub>s\<^sub>u\<^sub>f \ t \ \ \ t' \ \))" (is "\S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f. ?P S\<^sub>p\<^sub>r\<^sub>e \ ?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f") using assms proof (induction S rule: wf\<^sub>s\<^sub>t_induct') case (ConsSnd t' S) show ?case proof (cases "\w \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ \ \ w") case True then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsSnd.IH by moura thus ?thesis by fastforce next case False then obtain v where v: "v \ fv t'" "t \ \ \ \ v" using ConsSnd.prems by auto hence "t \ \ \ t' \ \" using subst_mono[of "Var v" t' \] vars_iff_subtermeq[of v t'] term.order_trans by auto thus ?thesis using False v by auto qed next case (ConsRcv t' S) have "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t S" using ConsRcv.hyps vars_snd_rcv_strand_subset2(4)[of S] by blast hence "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ \ \ v" using ConsRcv.prems by fastforce then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsRcv.IH by moura thus ?case by fastforce next case (ConsEq s s' S) have *: "fv s' \ wfrestrictedvars\<^sub>s\<^sub>t S" using ConsEq.hyps vars_snd_rcv_strand_subset2(4)[of S] by blast show ?case proof (cases "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ \ \ v") case True then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsEq.IH by moura thus ?thesis by fastforce next case False then obtain v where "v \ fv s" "t \ \ \ \ v" using ConsEq.prems * by auto hence "t \ \ \ s \ \" using vars_iff_subtermeq[of v s] subst_mono[of "Var v" s \] term.order_trans by auto thus ?thesis using False by fastforce qed next case (ConsEq2 s s' S) have "wfrestrictedvars\<^sub>s\<^sub>t (S@[Equality Check s s']) = wfrestrictedvars\<^sub>s\<^sub>t S" by auto hence "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ \ \ v" using ConsEq2.prems by metis then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsEq2.IH by moura thus ?case by fastforce next case (ConsIneq X F S) hence "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. t \ \ \ \ v" by fastforce then obtain S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where "?P S\<^sub>p\<^sub>r\<^sub>e" "?Q S S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f" using ConsIneq.IH by moura thus ?case by fastforce qed simp subsection \Constraint Semantics\ context intruder_model begin subsubsection \Definitions\ text \The constraint semantics in which the intruder is limited to composition only\ fun strand_sem_c::"('fun,'var) terms \ ('fun,'var) strand \ ('fun,'var) subst \ bool" ("\_; _\\<^sub>c") where "\M; []\\<^sub>c = (\\. True)" | "\M; Send t#S\\<^sub>c = (\\. M \\<^sub>c t \ \ \ \M; S\\<^sub>c \)" | "\M; Receive t#S\\<^sub>c = (\\. \insert (t \ \) M; S\\<^sub>c \)" | "\M; Equality _ t t'#S\\<^sub>c = (\\. t \ \ = t' \ \ \ \M; S\\<^sub>c \)" | "\M; Inequality X F#S\\<^sub>c = (\\. ineq_model \ X F \ \M; S\\<^sub>c \)" definition constr_sem_c ("_ \\<^sub>c \_,_\") where "\ \\<^sub>c \S,\\ \ (\ supports \ \ \{}; S\\<^sub>c \)" abbreviation constr_sem_c' ("_ \\<^sub>c \_\" 90) where "\ \\<^sub>c \S\ \ \ \\<^sub>c \S,Var\" text \The full constraint semantics\ fun strand_sem_d::"('fun,'var) terms \ ('fun,'var) strand \ ('fun,'var) subst \ bool" ("\_; _\\<^sub>d") where "\M; []\\<^sub>d = (\\. True)" | "\M; Send t#S\\<^sub>d = (\\. M \ t \ \ \ \M; S\\<^sub>d \)" | "\M; Receive t#S\\<^sub>d = (\\. \insert (t \ \) M; S\\<^sub>d \)" | "\M; Equality _ t t'#S\\<^sub>d = (\\. t \ \ = t' \ \ \ \M; S\\<^sub>d \)" | "\M; Inequality X F#S\\<^sub>d = (\\. ineq_model \ X F \ \M; S\\<^sub>d \)" definition constr_sem_d ("_ \ \_,_\") where "\ \ \S,\\ \ (\ supports \ \ \{}; S\\<^sub>d \)" abbreviation constr_sem_d' ("_ \ \_\" 90) where "\ \ \S\ \ \ \ \S,Var\" lemmas strand_sem_induct = strand_sem_c.induct[case_names Nil ConsSnd ConsRcv ConsEq ConsIneq] subsubsection \Lemmata\ lemma strand_sem_d_if_c: "\ \\<^sub>c \S,\\ \ \ \ \S,\\" proof - assume *: "\ \\<^sub>c \S,\\" { fix M have "\M; S\\<^sub>c \ \ \M; S\\<^sub>d \" proof (induction S rule: strand_sem_induct) case (ConsSnd M t S) hence "M \\<^sub>c t \ \" "\M; S\\<^sub>d \" by auto thus ?case using strand_sem_d.simps(2)[of M t S] by auto qed (auto simp add: ineq_model_def) } thus ?thesis using * by (simp add: constr_sem_c_def constr_sem_d_def) qed lemma strand_sem_mono_ik: "\M \ M'; \M; S\\<^sub>c \\ \ \M'; S\\<^sub>c \" (is "\?A'; ?A''\ \ ?A") "\M \ M'; \M; S\\<^sub>d \\ \ \M'; S\\<^sub>d \" (is "\?B'; ?B''\ \ ?B") proof - show "\?A'; ?A''\ \ ?A" proof (induction M S arbitrary: M M' rule: strand_sem_induct) case (ConsRcv M t S) thus ?case using ConsRcv.IH[of "insert (t \ \) M" "insert (t \ \) M'"] by auto next case (ConsSnd M t S) hence "M \\<^sub>c t \ \" "\M'; S\\<^sub>c \" by auto hence "M' \\<^sub>c t \ \" using ideduct_synth_mono \M \ M'\ by metis thus ?case using \\M'; S\\<^sub>c \\ by simp qed auto show "\?B'; ?B''\ \ ?B" proof (induction M S arbitrary: M M' rule: strand_sem_induct) case (ConsRcv M t S) thus ?case using ConsRcv.IH[of "insert (t \ \) M" "insert (t \ \) M'"] by auto next case (ConsSnd M t S) hence "M \ t \ \" "\M'; S\\<^sub>d \" by auto hence "M' \ t \ \" using ideduct_mono \M \ M'\ by metis thus ?case using \\M'; S\\<^sub>d \\ by simp qed auto qed context begin private lemma strand_sem_split_left: "\M; S@S'\\<^sub>c \ \ \M; S\\<^sub>c \" "\M; S@S'\\<^sub>d \ \ \M; S\\<^sub>d \" proof (induct S arbitrary: M) case (Cons x S) { case 1 thus ?case using Cons by (cases x) simp_all } { case 2 thus ?case using Cons by (cases x) simp_all } qed simp_all private lemma strand_sem_split_right: "\M; S@S'\\<^sub>c \ \ \M \ (ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \); S'\\<^sub>c \" "\M; S@S'\\<^sub>d \ \ \M \ (ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \); S'\\<^sub>d \" proof (induction S arbitrary: M rule: ik\<^sub>s\<^sub>t_induct) case (ConsRcv t S) { case 1 thus ?case using ConsRcv.IH[of "insert (t \ \) M"] by simp } { case 2 thus ?case using ConsRcv.IH[of "insert (t \ \) M"] by simp } qed simp_all lemmas strand_sem_split[dest] = strand_sem_split_left(1) strand_sem_split_right(1) strand_sem_split_left(2) strand_sem_split_right(2) end lemma strand_sem_Send_split[dest]: "\\M; map Send T\\<^sub>c \; t \ set T\ \ \M; [Send t]\\<^sub>c \" (is "\?A'; ?A''\ \ ?A") "\\M; map Send T\\<^sub>d \; t \ set T\ \ \M; [Send t]\\<^sub>d \" (is "\?B'; ?B''\ \ ?B") "\\M; map Send T@S\\<^sub>c \; t \ set T\ \ \M; Send t#S\\<^sub>c \" (is "\?C'; ?C''\ \ ?C") "\\M; map Send T@S\\<^sub>d \; t \ set T\ \ \M; Send t#S\\<^sub>d \" (is "\?D'; ?D''\ \ ?D") proof - show A: "\?A'; ?A''\ \ ?A" by (induct "map Send T" arbitrary: T rule: strand_sem_c.induct) auto show B: "\?B'; ?B''\ \ ?B" by (induct "map Send T" arbitrary: T rule: strand_sem_d.induct) auto show "\?C'; ?C''\ \ ?C" "\?D'; ?D''\ \ ?D" using list.set_map list.simps(8) set_empty ik_snd_empty sup_bot.right_neutral by (metis (no_types, lifting) A strand_sem_split(1,2) strand_sem_c.simps(2), metis (no_types, lifting) B strand_sem_split(3,4) strand_sem_d.simps(2)) qed lemma strand_sem_Send_map: "(\t. t \ set T \ \M; [Send t]\\<^sub>c \) \ \M; map Send T\\<^sub>c \" "(\t. t \ set T \ \M; [Send t]\\<^sub>d \) \ \M; map Send T\\<^sub>d \" by (induct T) auto lemma strand_sem_Receive_map: "\M; map Receive T\\<^sub>c \" "\M; map Receive T\\<^sub>d \" by (induct T arbitrary: M) auto lemma strand_sem_append[intro]: "\\M; S\\<^sub>c \; \M \ (ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \); S'\\<^sub>c \\ \ \M; S@S'\\<^sub>c \" "\\M; S\\<^sub>d \; \M \ (ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \); S'\\<^sub>d \\ \ \M; S@S'\\<^sub>d \" proof (induction S arbitrary: M) case (Cons x S) { case 1 thus ?case using Cons by (cases x) auto } { case 2 thus ?case using Cons by (cases x) auto } qed simp_all lemma ineq_model_subst: fixes F::"(('a,'b) term \ ('a,'b) term) list" assumes "(subst_domain \ \ range_vars \) \ set X = {}" and "ineq_model (\ \\<^sub>s \) X F" shows "ineq_model \ X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" proof - { fix \::"('a,'b) subst" and t t' assume \: "subst_domain \ = set X" "ground (subst_range \)" and *: "list_ex (\f. fst f \ (\ \\<^sub>s (\ \\<^sub>s \)) \ snd f \ (\ \\<^sub>s (\ \\<^sub>s \))) F" obtain f where f: "f \ set F" "fst f \ \ \\<^sub>s (\ \\<^sub>s \) \ snd f \ \ \\<^sub>s (\ \\<^sub>s \)" using * by (induct F) auto have "\ \\<^sub>s (\ \\<^sub>s \) = \ \\<^sub>s (\ \\<^sub>s \)" by (metis (no_types, lifting) \ subst_compose_assoc assms(1) inf_sup_aci(1) subst_comp_eq_if_disjoint_vars sup_inf_absorb range_vars_alt_def) hence "(fst f \ \) \ \ \\<^sub>s \ \ (snd f \ \) \ \ \\<^sub>s \" using f by auto moreover have "(fst f \ \, snd f \ \) \ set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using f(1) by (auto simp add: subst_apply_pairs_def) ultimately have "list_ex (\f. fst f \ (\ \\<^sub>s \) \ snd f \ (\ \\<^sub>s \)) (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using f(1) Bex_set by fastforce } thus ?thesis using assms unfolding ineq_model_def by simp qed lemma ineq_model_subst': fixes F::"(('a,'b) term \ ('a,'b) term) list" assumes "(subst_domain \ \ range_vars \) \ set X = {}" and "ineq_model \ X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" shows "ineq_model (\ \\<^sub>s \) X F" proof - { fix \::"('a,'b) subst" and t t' assume \: "subst_domain \ = set X" "ground (subst_range \)" and *: "list_ex (\f. fst f \ (\ \\<^sub>s \) \ snd f \ (\ \\<^sub>s \)) (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" obtain f where f: "f \ set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" "fst f \ \ \\<^sub>s \ \ snd f \ \ \\<^sub>s \" using * by (induct F) (auto simp add: subst_apply_pairs_def) then obtain g where g: "g \ set F" "f = g \\<^sub>p \" by (auto simp add: subst_apply_pairs_def) have "\ \\<^sub>s (\ \\<^sub>s \) = \ \\<^sub>s (\ \\<^sub>s \)" by (metis (no_types, lifting) \ subst_compose_assoc assms(1) inf_sup_aci(1) subst_comp_eq_if_disjoint_vars sup_inf_absorb range_vars_alt_def) hence "fst g \ \ \\<^sub>s (\ \\<^sub>s \) \ snd g \ \ \\<^sub>s (\ \\<^sub>s \)" using f(2) g by (simp add: prod.case_eq_if) hence "list_ex (\f. fst f \ (\ \\<^sub>s (\ \\<^sub>s \)) \ snd f \ (\ \\<^sub>s (\ \\<^sub>s \))) F" using g Bex_set by fastforce } thus ?thesis using assms unfolding ineq_model_def by simp qed lemma ineq_model_ground_subst: fixes F::"(('a,'b) term \ ('a,'b) term) list" assumes "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ subst_domain \" and "ground (subst_range \)" and "ineq_model \ X F" shows "ineq_model (\ \\<^sub>s \) X F" proof - { fix \::"('a,'b) subst" and t t' assume \: "subst_domain \ = set X" "ground (subst_range \)" and *: "list_ex (\f. fst f \ (\ \\<^sub>s \) \ snd f \ (\ \\<^sub>s \ )) F" obtain f where f: "f \ set F" "fst f \ \ \\<^sub>s \ \ snd f \ \ \\<^sub>s \" using * by (induct F) auto hence "fv (fst f) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" "fv (snd f) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" by auto hence "fv (fst f) - set X \ subst_domain \" "fv (snd f) - set X \ subst_domain \" using assms(1) by auto hence "fv (fst f \ \) \ subst_domain \" "fv (snd f \ \) \ subst_domain \" using \ by (simp_all add: range_vars_alt_def subst_fv_unfold_ground_img) hence "fv (fst f \ \ \\<^sub>s \) = {}" "fv (snd f \ \ \\<^sub>s \) = {}" using assms(2) by (simp_all add: subst_fv_dom_ground_if_ground_img) hence "fst f \ \ \\<^sub>s (\ \\<^sub>s \) \ snd f \ \ \\<^sub>s (\ \\<^sub>s \)" using f(2) subst_ground_ident by fastforce hence "list_ex (\f. fst f \ (\ \\<^sub>s (\ \\<^sub>s \)) \ snd f \ (\ \\<^sub>s (\ \\<^sub>s \))) F" using f(1) Bex_set by fastforce } thus ?thesis using assms unfolding ineq_model_def by simp qed context begin private lemma strand_sem_subst_c: assumes "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "\M; S\\<^sub>c (\ \\<^sub>s \) \ \M; S \\<^sub>s\<^sub>t \\\<^sub>c \" using assms proof (induction S arbitrary: \ M rule: strand_sem_induct) case (ConsSnd M t S) hence "\M; S \\<^sub>s\<^sub>t \\\<^sub>c \" "M \\<^sub>c t \ (\ \\<^sub>s \)" by auto hence "M \\<^sub>c (t \ \) \ \" using subst_comp_all[of \ \ M] subst_subst_compose[of t \ \] by simp thus ?case using \\M; S \\<^sub>s\<^sub>t \\\<^sub>c \\ unfolding subst_apply_strand_def by simp next case (ConsRcv M t S) have *: "\insert (t \ \ \\<^sub>s \) M; S\\<^sub>c (\ \\<^sub>s \)" using ConsRcv.prems(1) by simp have "bvars\<^sub>s\<^sub>t (Receive t#S) = bvars\<^sub>s\<^sub>t S" by auto hence **: "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" using ConsRcv.prems(2) by blast have "\M; Receive (t \ \)#(S \\<^sub>s\<^sub>t \)\\<^sub>c \" using ConsRcv.IH[OF * **] by (simp add: subst_all_insert) thus ?case by simp next case (ConsIneq M X F S) hence *: "\M; S \\<^sub>s\<^sub>t \\\<^sub>c \" and ***: "(subst_domain \ \ range_vars \) \ set X = {}" unfolding bvars\<^sub>s\<^sub>t_def ineq_model_def by auto have **: "ineq_model (\ \\<^sub>s \) X F" using ConsIneq by (auto simp add: subst_compose_assoc ineq_model_def) have "\\. subst_domain \ = set X \ ground (subst_range \) \ (subst_domain \ \ range_vars \) \ (subst_domain \ \ range_vars \) = {}" using * ** *** unfolding range_vars_alt_def by auto hence "\\. subst_domain \ = set X \ ground (subst_range \) \ \ \\<^sub>s \ = \ \\<^sub>s \" by (metis subst_comp_eq_if_disjoint_vars) hence "ineq_model \ X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using ineq_model_subst[OF *** **] by blast moreover have "rm_vars (set X) \ = \" using ConsIneq.prems(2) by force ultimately show ?case using * by auto qed simp_all private lemma strand_sem_subst_c': assumes "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "\M; S \\<^sub>s\<^sub>t \\\<^sub>c \ \ \M; S\\<^sub>c (\ \\<^sub>s \)" using assms proof (induction S arbitrary: \ M rule: strand_sem_induct) case (ConsSnd M t S) hence "\M; [Send t] \\<^sub>s\<^sub>t \\\<^sub>c \" "\M; S \\<^sub>s\<^sub>t \\\<^sub>c \" by auto hence "\M; S\\<^sub>c (\ \\<^sub>s \)" using ConsSnd.IH[OF _] ConsSnd.prems(2) by auto moreover have "\M; [Send t]\\<^sub>c (\ \\<^sub>s \)" proof - have "M \\<^sub>c t \ \ \ \" using \\M; [Send t] \\<^sub>s\<^sub>t \\\<^sub>c \\ by auto hence "M \\<^sub>c t \ (\ \\<^sub>s \)" using subst_subst_compose by metis thus "\M; [Send t]\\<^sub>c (\ \\<^sub>s \)" by auto qed ultimately show ?case by auto next case (ConsRcv M t S) hence "\(insert (t \ \ \ \) M); S \\<^sub>s\<^sub>t \\\<^sub>c \" by (simp add: subst_all_insert) thus ?case using ConsRcv.IH ConsRcv.prems(2) by auto next case (ConsIneq M X F S) have \: "rm_vars (set X) \ = \" using ConsIneq.prems(2) by force hence *: "\M; S\\<^sub>c (\ \\<^sub>s \)" and ***: "(subst_domain \ \ range_vars \) \ set X = {}" using ConsIneq unfolding bvars\<^sub>s\<^sub>t_def ineq_model_def by auto have **: "ineq_model \ X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using ConsIneq.prems(1) \ by (auto simp add: subst_compose_assoc ineq_model_def) have "\\. subst_domain \ = set X \ ground (subst_range \) \ (subst_domain \ \ range_vars \) \ (subst_domain \ \ range_vars \) = {}" using * ** *** unfolding range_vars_alt_def by auto hence "\\. subst_domain \ = set X \ ground (subst_range \) \ \ \\<^sub>s \ = \ \\<^sub>s \" by (metis subst_comp_eq_if_disjoint_vars) hence "ineq_model (\ \\<^sub>s \) X F" using ineq_model_subst'[OF *** **] by blast thus ?case using * by auto next case ConsEq thus ?case unfolding bvars\<^sub>s\<^sub>t_def by auto qed simp_all private lemma strand_sem_subst_d: assumes "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "\M; S\\<^sub>d (\ \\<^sub>s \) \ \M; S \\<^sub>s\<^sub>t \\\<^sub>d \" using assms proof (induction S arbitrary: \ M rule: strand_sem_induct) case (ConsSnd M t S) hence "\M; S \\<^sub>s\<^sub>t \\\<^sub>d \" "M \ t \ (\ \\<^sub>s \)" by auto hence "M \ (t \ \) \ \" using subst_comp_all[of \ \ M] subst_subst_compose[of t \ \] by simp thus ?case using \\M; S \\<^sub>s\<^sub>t \\\<^sub>d \\ by simp next case (ConsRcv M t S) have *: "\insert (t \ \ \\<^sub>s \) M; S\\<^sub>d (\ \\<^sub>s \)" using ConsRcv.prems(1) by simp have "bvars\<^sub>s\<^sub>t (Receive t#S) = bvars\<^sub>s\<^sub>t S" by auto hence **: "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" using ConsRcv.prems(2) by blast have "\M; Receive (t \ \)#(S \\<^sub>s\<^sub>t \)\\<^sub>d \" using ConsRcv.IH[OF * **] by (simp add: subst_all_insert) thus ?case by simp next case (ConsIneq M X F S) hence *: "\M; S \\<^sub>s\<^sub>t \\\<^sub>d \" and ***: "(subst_domain \ \ range_vars \) \ set X = {}" unfolding bvars\<^sub>s\<^sub>t_def ineq_model_def by auto have **: "ineq_model (\ \\<^sub>s \) X F" using ConsIneq by (auto simp add: subst_compose_assoc ineq_model_def) have "\\. subst_domain \ = set X \ ground (subst_range \) \ (subst_domain \ \ range_vars \) \ (subst_domain \ \ range_vars \) = {}" using * ** *** unfolding range_vars_alt_def by auto hence "\\. subst_domain \ = set X \ ground (subst_range \) \ \ \\<^sub>s \ = \ \\<^sub>s \" by (metis subst_comp_eq_if_disjoint_vars) hence "ineq_model \ X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using ineq_model_subst[OF *** **] by blast moreover have "rm_vars (set X) \ = \" using ConsIneq.prems(2) by force ultimately show ?case using * by auto next case ConsEq thus ?case unfolding bvars\<^sub>s\<^sub>t_def by auto qed simp_all private lemma strand_sem_subst_d': assumes "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "\M; S \\<^sub>s\<^sub>t \\\<^sub>d \ \ \M; S\\<^sub>d (\ \\<^sub>s \)" using assms proof (induction S arbitrary: \ M rule: strand_sem_induct) case (ConsSnd M t S) hence "\M; [Send t] \\<^sub>s\<^sub>t \\\<^sub>d \" "\M; S \\<^sub>s\<^sub>t \\\<^sub>d \" by auto hence "\M; S\\<^sub>d (\ \\<^sub>s \)" using ConsSnd.IH[OF _] ConsSnd.prems(2) by auto moreover have "\M; [Send t]\\<^sub>d (\ \\<^sub>s \)" proof - have "M \ t \ \ \ \" using \\M; [Send t] \\<^sub>s\<^sub>t \\\<^sub>d \\ by auto hence "M \ t \ (\ \\<^sub>s \)" using subst_subst_compose by metis thus "\M; [Send t]\\<^sub>d (\ \\<^sub>s \)" by auto qed ultimately show ?case by auto next case (ConsRcv M t S) hence "\insert (t \ \ \ \) M; S \\<^sub>s\<^sub>t \\\<^sub>d \" by (simp add: subst_all_insert) thus ?case using ConsRcv.IH ConsRcv.prems(2) by auto next case (ConsIneq M X F S) have \: "rm_vars (set X) \ = \" using ConsIneq.prems(2) by force hence *: "\M; S\\<^sub>d (\ \\<^sub>s \)" and ***: "(subst_domain \ \ range_vars \) \ set X = {}" using ConsIneq unfolding bvars\<^sub>s\<^sub>t_def ineq_model_def by auto have **: "ineq_model \ X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using ConsIneq.prems(1) \ by (auto simp add: subst_compose_assoc ineq_model_def) have "\\. subst_domain \ = set X \ ground (subst_range \) \ (subst_domain \ \ range_vars \) \ (subst_domain \ \ range_vars \) = {}" using * ** *** unfolding range_vars_alt_def by auto hence "\\. subst_domain \ = set X \ ground (subst_range \) \ \ \\<^sub>s \ = \ \\<^sub>s \" by (metis subst_comp_eq_if_disjoint_vars) hence "ineq_model (\ \\<^sub>s \) X F" using ineq_model_subst'[OF *** **] by blast thus ?case using * by auto next case ConsEq thus ?case unfolding bvars\<^sub>s\<^sub>t_def by auto qed simp_all lemmas strand_sem_subst = strand_sem_subst_c strand_sem_subst_c' strand_sem_subst_d strand_sem_subst_d' end lemma strand_sem_subst_subst_idem: assumes \: "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" shows "\\M; S \\<^sub>s\<^sub>t \\\<^sub>c (\ \\<^sub>s \); subst_idem \\ \ \M; S\\<^sub>c (\ \\<^sub>s \)" using strand_sem_subst(2)[OF assms, of M "\ \\<^sub>s \"] subst_compose_assoc[of \ \ \] unfolding subst_idem_def by argo lemma strand_sem_subst_comp: assumes "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t S = {}" and "\M; S\\<^sub>c \" "subst_domain \ \ (vars\<^sub>s\<^sub>t S \ fv\<^sub>s\<^sub>e\<^sub>t M) = {}" shows "\M; S\\<^sub>c (\ \\<^sub>s \)" proof - from assms(3) have "subst_domain \ \ vars\<^sub>s\<^sub>t S = {}" "subst_domain \ \ fv\<^sub>s\<^sub>e\<^sub>t M = {}" by auto hence "S \\<^sub>s\<^sub>t \ = S" "M \\<^sub>s\<^sub>e\<^sub>t \ = M" using strand_substI set_subst_ident[of M \] by (blast, blast) thus ?thesis using assms(2) by (auto simp add: strand_sem_subst(2)[OF assms(1)]) qed lemma strand_sem_c_imp_ineqs_neq: assumes "\M; S\\<^sub>c \" "Inequality X [(t,t')] \ set S" shows "t \ t' \ (\\. subst_domain \ = set X \ ground (subst_range \) \ t \ \ \ t' \ \ \ t \ \ \ \ \ t' \ \ \ \)" using assms proof (induction rule: strand_sem_induct) case (ConsIneq M Y F S) thus ?case proof (cases "Inequality X [(t,t')] \ set S") case False hence "X = Y" "F = [(t,t')]" using ConsIneq by auto hence *: "\\. subst_domain \ = set X \ ground (subst_range \) \ t \ \ \ \ \ t' \ \ \ \" using ConsIneq by (auto simp add: ineq_model_def) then obtain \ where \: "subst_domain \ = set X" "ground (subst_range \)" "t \ \ \ \ \ t' \ \ \ \" using interpretation_subst_exists'[of "set X"] by moura hence "t \ t'" by auto moreover have "\\ \. t \ \ \ \ \ t' \ \ \ \ \ t \ \ \ t' \ \" by auto ultimately show ?thesis using * by auto qed simp qed simp_all lemma strand_sem_c_imp_ineq_model: assumes "\M; S\\<^sub>c \" "Inequality X F \ set S" shows "ineq_model \ X F" using assms by (induct S rule: strand_sem_induct) force+ lemma strand_sem_wf_simple_fv_sat: assumes "wf\<^sub>s\<^sub>t {} S" "simple S" "\{}; S\\<^sub>c \" shows "\v. v \ wfrestrictedvars\<^sub>s\<^sub>t S \ ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c \ v" using assms proof (induction S rule: wf\<^sub>s\<^sub>t_simple_induct) case (ConsRcv t S) have "v \ wfrestrictedvars\<^sub>s\<^sub>t S" using ConsRcv.hyps(3) ConsRcv.prems(1) vars_snd_rcv_strand2 by fastforce moreover have "\{}; S\\<^sub>c \" using \\{}; S@[Receive t]\\<^sub>c \\ by blast moreover have "ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \ ik\<^sub>s\<^sub>t (S@[Receive t]) \\<^sub>s\<^sub>e\<^sub>t \" by auto ultimately show ?case using ConsRcv.IH ideduct_synth_mono by meson next case (ConsIneq X F S) hence "v \ wfrestrictedvars\<^sub>s\<^sub>t S" by fastforce moreover have "\{}; S\\<^sub>c \" using \\{}; S@[Inequality X F]\\<^sub>c \\ by blast moreover have "ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \ ik\<^sub>s\<^sub>t (S@[Inequality X F]) \\<^sub>s\<^sub>e\<^sub>t \" by auto ultimately show ?case using ConsIneq.IH ideduct_synth_mono by meson next case (ConsSnd w S) hence *: "\{}; S\\<^sub>c \" "ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c \ w" by auto have **: "ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \ ik\<^sub>s\<^sub>t (S@[Send (Var w)]) \\<^sub>s\<^sub>e\<^sub>t \" by simp show ?case proof (cases "v = w") case True thus ?thesis using *(2) ideduct_synth_mono[OF _ **] by meson next case False hence "v \ wfrestrictedvars\<^sub>s\<^sub>t S" using ConsSnd.prems(1) by auto thus ?thesis using ConsSnd.IH[OF _ *(1)] ideduct_synth_mono[OF _ **] by metis qed qed simp lemma strand_sem_wf_ik_or_assignment_rhs_fun_subterm: assumes "wf\<^sub>s\<^sub>t {} A" "\{}; A\\<^sub>c \" "Var x \ ik\<^sub>s\<^sub>t A" "\ x = Fun f T" "t\<^sub>i \ set T" "\ik\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t\<^sub>i" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" obtains S where "Fun f S \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t A) \ Fun f S \ subterms\<^sub>s\<^sub>e\<^sub>t (assignment_rhs\<^sub>s\<^sub>t A)" "Fun f T = Fun f S \ \" proof - have "x \ wfrestrictedvars\<^sub>s\<^sub>t A" by (metis (no_types) assms(3) set_rev_mp term.set_intros(3) vars_subset_if_in_strand_ik2) moreover have "Fun f T \ \ = Fun f T" by (metis subst_ground_ident interpretation_grounds_all assms(4,7)) ultimately obtain A\<^sub>p\<^sub>r\<^sub>e A\<^sub>s\<^sub>u\<^sub>f where *: "\(\w \ wfrestrictedvars\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e. Fun f T \ \ w)" "(\t. A = A\<^sub>p\<^sub>r\<^sub>e@Send t#A\<^sub>s\<^sub>u\<^sub>f \ Fun f T \ t \ \) \ (\t t'. A = A\<^sub>p\<^sub>r\<^sub>e@Equality Assign t t'#A\<^sub>s\<^sub>u\<^sub>f \ Fun f T \ t \ \)" using wf_strand_first_Send_var_split[OF assms(1)] assms(4) subtermeqI' by metis moreover { fix t assume **: "A = A\<^sub>p\<^sub>r\<^sub>e@Send t#A\<^sub>s\<^sub>u\<^sub>f" "Fun f T \ t \ \" hence "ik\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" "\ik\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t\<^sub>i" using assms(2,6) by (auto intro: ideduct_synth_mono) then obtain s where s: "s \ ik\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e" "Fun f T \ s \ \" using assms(5) **(2) by (induct rule: intruder_synth_induct) auto then obtain g S where gS: "Fun g S \ s" "Fun f T = Fun g S \ \" using subterm_subst_not_img_subterm[OF s(2)] *(1) by force hence ?thesis using that **(1) s(1) by force } moreover { fix t t' assume **: "A = A\<^sub>p\<^sub>r\<^sub>e@Equality Assign t t'#A\<^sub>s\<^sub>u\<^sub>f" "Fun f T \ t \ \" with assms(2) have "t \ \ = t' \ \" by auto hence "Fun f T \ t' \ \" using **(2) by auto from assms(1) **(1) have "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t A\<^sub>p\<^sub>r\<^sub>e" using wf_eq_fv[of "{}" A\<^sub>p\<^sub>r\<^sub>e t t' A\<^sub>s\<^sub>u\<^sub>f] vars_snd_rcv_strand_subset2(4)[of A\<^sub>p\<^sub>r\<^sub>e] by blast then obtain g S where gS: "Fun g S \ t'" "Fun f T = Fun g S \ \" using subterm_subst_not_img_subterm[OF \Fun f T \ t' \ \\] *(1) by fastforce hence ?thesis using that **(1) by auto } ultimately show ?thesis by auto qed lemma strand_sem_not_unif_is_sat_ineq: assumes "\\. Unifier \ t t'" shows "\M; [Inequality X [(t,t')]]\\<^sub>c \" "\M; [Inequality X [(t,t')]]\\<^sub>d \" using assms list_ex_simps(1)[of _ "(t,t')" "[]"] prod.sel[of t t'] strand_sem_c.simps(1,5) strand_sem_d.simps(1,5) unfolding ineq_model_def by presburger+ lemma ineq_model_singleI[intro]: assumes "\\. subst_domain \ = set X \ ground (subst_range \) \ t \ \ \ \ \ t' \ \ \ \" shows "ineq_model \ X [(t,t')]" using assms unfolding ineq_model_def by auto lemma ineq_model_singleE: assumes "ineq_model \ X [(t,t')]" shows "\\. subst_domain \ = set X \ ground (subst_range \) \ t \ \ \ \ \ t' \ \ \ \" using assms unfolding ineq_model_def by auto lemma ineq_model_single_iff: fixes F::"(('a,'b) term \ ('a,'b) term) list" shows "ineq_model \ X F \ ineq_model \ X [(Fun f (Fun c []#map fst F),Fun f (Fun c []#map snd F))]" (is "?A \ ?B") proof - let ?P = "\\ f. fst f \ (\ \\<^sub>s \) \ snd f \ (\ \\<^sub>s \)" let ?Q = "\\ t t'. t \ (\ \\<^sub>s \) \ t' \ (\ \\<^sub>s \)" let ?T = "\g. Fun c []#map g F" let ?S = "\\ g. map (\x. x \ (\ \\<^sub>s \)) (Fun c []#map g F)" let ?t = "Fun f (?T fst)" let ?t' = "Fun f (?T snd)" have len: "\g h. length (?T g) = length (?T h)" "\g h \. length (?S \ g) = length (?T h)" "\g h \. length (?S \ g) = length (?T h)" "\g h \ \. length (?S \ g) = length (?S \ h)" by simp_all { fix \::"('a,'b) subst" assume \: "subst_domain \ = set X" "ground (subst_range \)" have "list_ex (?P \) F \ ?Q \ ?t ?t'" proof assume "list_ex (?P \) F" then obtain a where a: "a \ set F" "?P \ a" by (metis (mono_tags, lifting) Bex_set) thus "?Q \ ?t ?t'" by auto qed (fastforce simp add: Bex_set) } thus ?thesis unfolding ineq_model_def by auto qed subsection \Constraint Semantics (Alternative, Equivalent Version)\ text \These are the constraint semantics used in the CSF 2017 paper\ fun strand_sem_c'::"('fun,'var) terms \ ('fun,'var) strand \ ('fun,'var) subst \ bool" ("\_; _\\<^sub>c''") where "\M; []\\<^sub>c' = (\\. True)" | "\M; Send t#S\\<^sub>c' = (\\. M \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \ \ \M; S\\<^sub>c' \)" | "\M; Receive t#S\\<^sub>c' = \insert t M; S\\<^sub>c'" | "\M; Equality _ t t'#S\\<^sub>c' = (\\. t \ \ = t' \ \ \ \M; S\\<^sub>c' \)" | "\M; Inequality X F#S\\<^sub>c' = (\\. ineq_model \ X F \ \M; S\\<^sub>c' \)" fun strand_sem_d'::"('fun,'var) terms \ ('fun,'var) strand \ ('fun,'var) subst \ bool" ("\_; _\\<^sub>d''") where "\M; []\\<^sub>d' = (\\. True)" | "\M; Send t#S\\<^sub>d' = (\\. M \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \ \ \M; S\\<^sub>d' \)" | "\M; Receive t#S\\<^sub>d' = \insert t M; S\\<^sub>d'" | "\M; Equality _ t t'#S\\<^sub>d' = (\\. t \ \ = t' \ \ \ \M; S\\<^sub>d' \)" | "\M; Inequality X F#S\\<^sub>d' = (\\. ineq_model \ X F \ \M; S\\<^sub>d' \)" lemma strand_sem_eq_defs: "\M; \\\<^sub>c' \ = \M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>c \" "\M; \\\<^sub>d' \ = \M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>d \" proof - have 1: "\M; \\\<^sub>c' \ \ \M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>c \" by (induct \ arbitrary: M rule: strand_sem_induct) force+ have 2: "\M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>c \ \ \M; \\\<^sub>c' \" by (induct \ arbitrary: M rule: strand_sem_c'.induct) auto have 3: "\M; \\\<^sub>d' \ \ \M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>d \" by (induct \ arbitrary: M rule: strand_sem_induct) force+ have 4: "\M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>d \ \ \M; \\\<^sub>d' \" by (induct \ arbitrary: M rule: strand_sem_d'.induct) auto show "\M; \\\<^sub>c' \ = \M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>c \" "\M; \\\<^sub>d' \ = \M \\<^sub>s\<^sub>e\<^sub>t \; \\\<^sub>d \" by (metis 1 2, metis 3 4) qed lemma strand_sem_split'[dest]: "\M; S@S'\\<^sub>c' \ \ \M; S\\<^sub>c' \" "\M; S@S'\\<^sub>c' \ \ \M \ ik\<^sub>s\<^sub>t S; S'\\<^sub>c' \" "\M; S@S'\\<^sub>d' \ \ \M; S\\<^sub>d' \" "\M; S@S'\\<^sub>d' \ \ \M \ ik\<^sub>s\<^sub>t S; S'\\<^sub>d' \" using strand_sem_eq_defs[of M "S@S'" \] strand_sem_eq_defs[of M S \] strand_sem_eq_defs[of "M \ ik\<^sub>s\<^sub>t S" S' \] strand_sem_split(2,4) by (auto simp add: image_Un) lemma strand_sem_append'[intro]: "\M; S\\<^sub>c' \ \ \M \ ik\<^sub>s\<^sub>t S; S'\\<^sub>c' \ \ \M; S@S'\\<^sub>c' \" "\M; S\\<^sub>d' \ \ \M \ ik\<^sub>s\<^sub>t S; S'\\<^sub>d' \ \ \M; S@S'\\<^sub>d' \" using strand_sem_eq_defs[of M "S@S'" \] strand_sem_eq_defs[of M S \] strand_sem_eq_defs[of "M \ ik\<^sub>s\<^sub>t S" S' \] by (auto simp add: image_Un) end subsection \Dual Strands\ fun dual\<^sub>s\<^sub>t::"('a,'b) strand \ ('a,'b) strand" where "dual\<^sub>s\<^sub>t [] = []" | "dual\<^sub>s\<^sub>t (Receive t#S) = Send t#(dual\<^sub>s\<^sub>t S)" | "dual\<^sub>s\<^sub>t (Send t#S) = Receive t#(dual\<^sub>s\<^sub>t S)" | "dual\<^sub>s\<^sub>t (x#S) = x#(dual\<^sub>s\<^sub>t S)" lemma dual\<^sub>s\<^sub>t_append: "dual\<^sub>s\<^sub>t (A@B) = (dual\<^sub>s\<^sub>t A)@(dual\<^sub>s\<^sub>t B)" by (induct A rule: dual\<^sub>s\<^sub>t.induct) auto lemma dual\<^sub>s\<^sub>t_self_inverse: "dual\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t S) = S" proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp lemma dual\<^sub>s\<^sub>t_trms_eq: "trms\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t S) = trms\<^sub>s\<^sub>t S" proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp lemma dual\<^sub>s\<^sub>t_fv: "fv\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t A) = fv\<^sub>s\<^sub>t A" by (induct A rule: dual\<^sub>s\<^sub>t.induct) auto lemma dual\<^sub>s\<^sub>t_bvars: "bvars\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t A) = bvars\<^sub>s\<^sub>t A" by (induct A rule: dual\<^sub>s\<^sub>t.induct) fastforce+ end