(* (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: Lazy_Intruder.thy Author: Andreas Viktor Hess, DTU *) section \The Lazy Intruder\ theory Lazy_Intruder imports Strands_and_Constraints Intruder_Deduction begin context intruder_model begin subsection \Definition of the Lazy Intruder\ text \The lazy intruder constraint reduction system, defined as a relation on constraint states\ inductive_set LI_rel:: "((('fun,'var) strand \ (('fun,'var) subst)) \ ('fun,'var) strand \ (('fun,'var) subst)) set" and LI_rel' (infix "\" 50) and LI_rel_trancl (infix "\\<^sup>+" 50) and LI_rel_rtrancl (infix "\\<^sup>*" 50) where "A \ B \ (A,B) \ LI_rel" | "A \\<^sup>+ B \ (A,B) \ LI_rel\<^sup>+" | "A \\<^sup>* B \ (A,B) \ LI_rel\<^sup>*" | Compose: "\simple S; length T = arity f; public f\ \ (S@Send (Fun f T)#S',\) \ (S@(map Send T)@S',\)" | Unify: "\simple S; Fun f T' \ ik\<^sub>s\<^sub>t S; Some \ = mgu (Fun f T) (Fun f T')\ \ (S@Send (Fun f T)#S',\) \ ((S@S') \\<^sub>s\<^sub>t \,\ \\<^sub>s \)" | Equality: "\simple S; Some \ = mgu t t'\ \ (S@Equality _ t t'#S',\) \ ((S@S') \\<^sub>s\<^sub>t \,\ \\<^sub>s \)" subsection \Lemma: The Lazy Intruder is Well-founded\ context begin private lemma LI_compose_measure_lt: "((S@(map Send T)@S',\\<^sub>1), (S@Send (Fun f T)#S',\\<^sub>2)) \ measure\<^sub>s\<^sub>t" using strand_fv_card_map_fun_eq[of S f T S'] strand_size_map_fun_lt(2)[of T f] by (simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) private lemma LI_unify_measure_lt: assumes "Some \ = mgu (Fun f T) t" "fv t \ fv\<^sub>s\<^sub>t S" shows "(((S@S') \\<^sub>s\<^sub>t \,\\<^sub>1), (S@Send (Fun f T)#S',\\<^sub>2)) \ measure\<^sub>s\<^sub>t" proof (cases "\ = Var") assume "\ = Var" hence "(S@S') \\<^sub>s\<^sub>t \ = S@S'" by blast thus ?thesis using strand_fv_card_rm_fun_le[of S S' f T] by (auto simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) next assume "\ \ Var" then obtain v where "v \ fv (Fun f T) \ fv t" "subst_elim \ v" using mgu_eliminates[OF assms(1)[symmetric]] by metis hence v_in: "v \ fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" using assms(2) by (auto simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) have "range_vars \ \ fv (Fun f T) \ fv\<^sub>s\<^sub>t S" using assms(2) mgu_vars_bounded[OF assms(1)[symmetric]] by auto hence img_bound: "range_vars \ \ fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" by auto have finite_fv: "finite (fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S'))" by auto have "v \ fv\<^sub>s\<^sub>t ((S@Send (Fun f T)#S') \\<^sub>s\<^sub>t \)" using strand_fv_subst_subset_if_subst_elim[OF \subst_elim \ v\] v_in by metis hence v_not_in: "v \ fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" by auto have "fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" using strand_subst_fv_bounded_if_img_bounded[OF img_bound] by simp hence "fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" using v_in v_not_in by blast hence "card (fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)) < card (fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S'))" using psubset_card_mono[OF finite_fv] by simp thus ?thesis by (auto simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) qed private lemma LI_equality_measure_lt: assumes "Some \ = mgu t t'" shows "(((S@S') \\<^sub>s\<^sub>t \,\\<^sub>1), (S@Equality a t t'#S',\\<^sub>2)) \ measure\<^sub>s\<^sub>t" proof (cases "\ = Var") assume "\ = Var" hence "(S@S') \\<^sub>s\<^sub>t \ = S@S'" by blast thus ?thesis using strand_fv_card_rm_eq_le[of S S' a t t'] by (auto simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) next assume "\ \ Var" then obtain v where "v \ fv t \ fv t'" "subst_elim \ v" using mgu_eliminates[OF assms(1)[symmetric]] by metis hence v_in: "v \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using assms by auto have "range_vars \ \ fv t \ fv t' \ fv\<^sub>s\<^sub>t S" using assms mgu_vars_bounded[OF assms(1)[symmetric]] by auto hence img_bound: "range_vars \ \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" by auto have finite_fv: "finite (fv\<^sub>s\<^sub>t (S@Equality a t t'#S'))" by auto have "v \ fv\<^sub>s\<^sub>t ((S@Equality a t t'#S') \\<^sub>s\<^sub>t \)" using strand_fv_subst_subset_if_subst_elim[OF \subst_elim \ v\] v_in by metis hence v_not_in: "v \ fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)" by auto have "fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using strand_subst_fv_bounded_if_img_bounded[OF img_bound] by simp hence "fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using v_in v_not_in by blast hence "card (fv\<^sub>s\<^sub>t ((S@S') \\<^sub>s\<^sub>t \)) < card (fv\<^sub>s\<^sub>t (S@Equality a t t'#S'))" using psubset_card_mono[OF finite_fv] by simp thus ?thesis by (auto simp add: measure\<^sub>s\<^sub>t_def size\<^sub>s\<^sub>t_def) qed private lemma LI_in_measure: "(S\<^sub>1,\\<^sub>1) \ (S\<^sub>2,\\<^sub>2) \ ((S\<^sub>2,\\<^sub>2),(S\<^sub>1,\\<^sub>1)) \ measure\<^sub>s\<^sub>t" proof (induction rule: LI_rel.induct) case (Compose S T f S' \) thus ?case using LI_compose_measure_lt[of S T S'] by metis next case (Unify S f U \ T S' \) hence "fv (Fun f U) \ fv\<^sub>s\<^sub>t S" using fv_snd_rcv_strand_subset(2)[of S] by force thus ?case using LI_unify_measure_lt[OF Unify.hyps(3), of S S'] by metis qed (metis LI_equality_measure_lt) private lemma LI_in_measure_trans: "(S\<^sub>1,\\<^sub>1) \\<^sup>+ (S\<^sub>2,\\<^sub>2) \ ((S\<^sub>2,\\<^sub>2),(S\<^sub>1,\\<^sub>1)) \ measure\<^sub>s\<^sub>t" by (induction rule: trancl.induct, metis surjective_pairing LI_in_measure) (metis (no_types, lifting) surjective_pairing LI_in_measure measure\<^sub>s\<^sub>t_trans trans_def) private lemma LI_converse_wellfounded_trans: "wf ((LI_rel\<^sup>+)\)" proof - have "(LI_rel\<^sup>+)\ \ measure\<^sub>s\<^sub>t" using LI_in_measure_trans by auto thus ?thesis using measure\<^sub>s\<^sub>t_wellfounded wf_subset by metis qed private lemma LI_acyclic_trans: "acyclic (LI_rel\<^sup>+)" using wf_acyclic[OF LI_converse_wellfounded_trans] acyclic_converse by metis private lemma LI_acyclic: "acyclic LI_rel" using LI_acyclic_trans acyclic_subset by (simp add: acyclic_def) lemma LI_no_infinite_chain: "\(\f. \i. f i \\<^sup>+ f (Suc i))" proof - have "\(\f. \i. (f (Suc i), f i) \ (LI_rel\<^sup>+)\)" using wf_iff_no_infinite_down_chain LI_converse_wellfounded_trans by metis thus ?thesis by simp qed private lemma LI_unify_finite: assumes "finite M" shows "finite {((S@Send (Fun f T)#S',\), ((S@S') \\<^sub>s\<^sub>t \,\ \\<^sub>s \)) | \ T'. simple S \ Fun f T' \ M \ Some \ = mgu (Fun f T) (Fun f T')}" using assms proof (induction M rule: finite_induct) case (insert m M) thus ?case proof (cases m) case (Fun g U) let ?a = "\\. ((S@Send (Fun f T)#S',\), ((S@S') \\<^sub>s\<^sub>t \,\ \\<^sub>s \))" let ?A = "\B. {?a \ | \ T'. simple S \ Fun f T' \ B \ Some \ = mgu (Fun f T) (Fun f T')}" have "?A (insert m M) = (?A M) \ (?A {m})" by auto moreover have "finite (?A {m})" proof (cases "\\. Some \ = mgu (Fun f T) (Fun g U)") case True then obtain \ where \: "Some \ = mgu (Fun f T) (Fun g U)" by blast have A_m_eq: "\\'. ?a \' \ ?A {m} \ ?a \ = ?a \'" proof - fix \' assume "?a \' \ ?A {m}" hence "\\. Some \ = mgu (Fun f T) (Fun g U) \ ?a \ = ?a \'" using \m = Fun g U\ by auto thus "?a \ = ?a \'" by (metis \ option.inject) qed have "?A {m} = {} \ ?A {m} = {?a \}" proof (cases "simple S \ ?A {m} \ {}") case True hence "simple S" "?A {m} \ {}" by meson+ hence "?A {m} = {?a \ | \. Some \ = mgu (Fun f T) (Fun g U)}" using \m = Fun g U\ by auto hence "?a \ \ ?A {m}" using \ by auto show ?thesis proof (rule ccontr) assume "\(?A {m} = {} \ ?A {m} = {?a \})" then obtain B where B: "?A {m} = insert (?a \) B" "?a \ \ B" "B \ {}" using \?A {m} \ {}\ \?a \ \ ?A {m}\ by (metis (no_types, lifting) Set.set_insert) then obtain b where b: "?a \ \ b" "b \ B" by (metis (no_types, lifting) ex_in_conv) then obtain \' where \': "b = ?a \'" using B(1) by blast moreover have "?a \' \ ?A {m}" using B(1) b(2) \' by auto hence "?a \ = ?a \'" by (blast dest!: A_m_eq) ultimately show False using b(1) by simp qed qed auto thus ?thesis by (metis (no_types, lifting) finite.emptyI finite_insert) next case False hence "?A {m} = {}" using \m = Fun g U\ by blast thus ?thesis by (metis finite.emptyI) qed ultimately show ?thesis using insert.IH by auto qed simp qed fastforce end subsection \Lemma: The Lazy Intruder Preserves Well-formedness\ context begin private lemma LI_preserves_subst_wf_single: assumes "(S\<^sub>1,\\<^sub>1) \ (S\<^sub>2,\\<^sub>2)" "fv\<^sub>s\<^sub>t S\<^sub>1 \ bvars\<^sub>s\<^sub>t S\<^sub>1 = {}" "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>1" and "subst_domain \\<^sub>1 \ vars\<^sub>s\<^sub>t S\<^sub>1 = {}" "range_vars \\<^sub>1 \ bvars\<^sub>s\<^sub>t S\<^sub>1 = {}" shows "fv\<^sub>s\<^sub>t S\<^sub>2 \ bvars\<^sub>s\<^sub>t S\<^sub>2 = {}" "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>2" and "subst_domain \\<^sub>2 \ vars\<^sub>s\<^sub>t S\<^sub>2 = {}" "range_vars \\<^sub>2 \ bvars\<^sub>s\<^sub>t S\<^sub>2 = {}" using assms proof (induction rule: LI_rel.induct) case (Compose S X f S' \) { case 1 thus ?case using vars_st_snd_map by auto } { case 2 thus ?case using vars_st_snd_map by auto } { case 3 thus ?case using vars_st_snd_map by force } { case 4 thus ?case using vars_st_snd_map by auto } next case (Unify S f U \ T S' \) hence "fv (Fun f U) \ fv\<^sub>s\<^sub>t S" using fv_subset_if_in_strand_ik' by blast hence *: "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" using mgu_vars_bounded[OF Unify.hyps(3)[symmetric]] unfolding range_vars_alt_def by (fastforce simp del: subst_range.simps) have "fv\<^sub>s\<^sub>t (S@S') \ fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" "vars\<^sub>s\<^sub>t (S@S') \ vars\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" by auto hence **: "fv\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" "vars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) \ vars\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" using subst_sends_strand_fv_to_img[of "S@S'" \] strand_subst_vars_union_bound[of "S@S'" \] * by blast+ have "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" by (fact mgu_gives_wellformed_subst[OF Unify.hyps(3)[symmetric]]) { case 1 have "bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = bvars\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" using bvars_subst_ident[of "S@S'" \] by auto thus ?case using 1 ** by blast } { case 2 hence "subst_domain \ \ subst_domain \ = {}" "subst_domain \ \ range_vars \ = {}" using * by blast+ thus ?case by (metis wf_subst_compose[OF \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\]) } { case 3 hence "subst_domain \ \ vars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using ** by blast moreover have "v \ fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" when "v \ subst_domain \" for v using * that by blast hence "subst_domain \ \ fv\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using mgu_eliminates_dom[OF Unify.hyps(3)[symmetric], THEN strand_fv_subst_subset_if_subst_elim, of _ "S@Send (Fun f T)#S'"] unfolding subst_elim_def by auto moreover have "bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = bvars\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" using bvars_subst_ident[of "S@S'" \] by auto hence "subst_domain \ \ bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using 3(1) * by blast ultimately show ?case using ** * subst_domain_compose[of \ \] vars\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>t[of "S@S' \\<^sub>s\<^sub>t \"] by blast } { case 4 have ***: "bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = bvars\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" using bvars_subst_ident[of "S@S'" \] by auto hence "range_vars \ \ bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using 4(1) * by blast thus ?case using subst_img_comp_subset[of \ \] 4(4) *** by blast } next case (Equality S \ t t' a S' \) hence *: "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using mgu_vars_bounded[OF Equality.hyps(2)[symmetric]] unfolding range_vars_alt_def by fastforce have "fv\<^sub>s\<^sub>t (S@S') \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" "vars\<^sub>s\<^sub>t (S@S') \ vars\<^sub>s\<^sub>t (S@Equality a t t'#S')" by auto hence **: "fv\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" "vars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) \ vars\<^sub>s\<^sub>t (S@Equality a t t'#S')" using subst_sends_strand_fv_to_img[of "S@S'" \] strand_subst_vars_union_bound[of "S@S'" \] * by blast+ have "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" by (fact mgu_gives_wellformed_subst[OF Equality.hyps(2)[symmetric]]) { case 1 have "bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = bvars\<^sub>s\<^sub>t (S@Equality a t t'#S')" using bvars_subst_ident[of "S@S'" \] by auto thus ?case using 1 ** by blast } { case 2 hence "subst_domain \ \ subst_domain \ = {}" "subst_domain \ \ range_vars \ = {}" using * by blast+ thus ?case by (metis wf_subst_compose[OF \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\]) } { case 3 hence "subst_domain \ \ vars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using ** by blast moreover have "v \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" when "v \ subst_domain \" for v using * that by blast hence "subst_domain \ \ fv\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using mgu_eliminates_dom[OF Equality.hyps(2)[symmetric], THEN strand_fv_subst_subset_if_subst_elim, of _ "S@Equality a t t'#S'"] unfolding subst_elim_def by auto moreover have "bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = bvars\<^sub>s\<^sub>t (S@Equality a t t'#S')" using bvars_subst_ident[of "S@S'" \] by auto hence "subst_domain \ \ bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using 3(1) * by blast ultimately show ?case using ** * subst_domain_compose[of \ \] vars\<^sub>s\<^sub>t_is_fv\<^sub>s\<^sub>t_bvars\<^sub>s\<^sub>t[of "S@S' \\<^sub>s\<^sub>t \"] by blast } { case 4 have ***: "bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = bvars\<^sub>s\<^sub>t (S@Equality a t t'#S')" using bvars_subst_ident[of "S@S'" \] by auto hence "range_vars \ \ bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using 4(1) * by blast thus ?case using subst_img_comp_subset[of \ \] 4(4) *** by blast } qed private lemma LI_preserves_subst_wf: assumes "(S\<^sub>1,\\<^sub>1) \\<^sup>* (S\<^sub>2,\\<^sub>2)" "fv\<^sub>s\<^sub>t S\<^sub>1 \ bvars\<^sub>s\<^sub>t S\<^sub>1 = {}" "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>1" and "subst_domain \\<^sub>1 \ vars\<^sub>s\<^sub>t S\<^sub>1 = {}" "range_vars \\<^sub>1 \ bvars\<^sub>s\<^sub>t S\<^sub>1 = {}" shows "fv\<^sub>s\<^sub>t S\<^sub>2 \ bvars\<^sub>s\<^sub>t S\<^sub>2 = {}" "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>2" and "subst_domain \\<^sub>2 \ vars\<^sub>s\<^sub>t S\<^sub>2 = {}" "range_vars \\<^sub>2 \ bvars\<^sub>s\<^sub>t S\<^sub>2 = {}" using assms proof (induction S\<^sub>2 \\<^sub>2 rule: rtrancl_induct2) case (step S\<^sub>i \\<^sub>i S\<^sub>j \\<^sub>j) { case 1 thus ?case using LI_preserves_subst_wf_single[OF \(S\<^sub>i,\\<^sub>i) \ (S\<^sub>j,\\<^sub>j)\] step.IH by metis } { case 2 thus ?case using LI_preserves_subst_wf_single[OF \(S\<^sub>i,\\<^sub>i) \ (S\<^sub>j,\\<^sub>j)\] step.IH by metis } { case 3 thus ?case using LI_preserves_subst_wf_single[OF \(S\<^sub>i,\\<^sub>i) \ (S\<^sub>j,\\<^sub>j)\] step.IH by metis } { case 4 thus ?case using LI_preserves_subst_wf_single[OF \(S\<^sub>i,\\<^sub>i) \ (S\<^sub>j,\\<^sub>j)\] step.IH by metis } qed metis lemma LI_preserves_wellformedness: assumes "(S\<^sub>1,\\<^sub>1) \\<^sup>* (S\<^sub>2,\\<^sub>2)" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1" shows "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>2 \\<^sub>2" proof - have *: "wf\<^sub>s\<^sub>t {} S\<^sub>j" when "(S\<^sub>i, \\<^sub>i) \ (S\<^sub>j, \\<^sub>j)" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>i \\<^sub>i" for S\<^sub>i \\<^sub>i S\<^sub>j \\<^sub>j using that proof (induction rule: LI_rel.induct) case (Unify S f U \ T S' \) have "fv (Fun f T) \ fv (Fun f U) \ fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" using Unify.hyps(2) by force hence "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" using mgu_vars_bounded[OF Unify.hyps(3)[symmetric]] by (metis subset_trans) hence "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@Send (Fun f T)#S') = {}" using Unify.prems unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by blast thus ?case using wf_unify[OF _ Unify.hyps(2) MGU_is_Unifier[OF mgu_gives_MGU], of "{}", OF _ Unify.hyps(3)[symmetric], of S'] Unify.prems(1) by (auto simp add: wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) next case (Equality S \ t t' a S' \) have "fv t \ fv t' \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using Equality.hyps(2) by force hence "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using mgu_vars_bounded[OF Equality.hyps(2)[symmetric]] by (metis subset_trans) hence "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@Equality a t t'#S') = {}" using Equality.prems unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by blast thus ?case using wf_equality[OF _ Equality.hyps(2)[symmetric], of "{}" S a S'] Equality.prems(1) by (auto simp add: wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) qed (metis wf_send_compose wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) show ?thesis using assms proof (induction rule: rtrancl_induct2) case (step S\<^sub>i \\<^sub>i S\<^sub>j \\<^sub>j) thus ?case using LI_preserves_subst_wf_single[OF \(S\<^sub>i,\\<^sub>i) \ (S\<^sub>j,\\<^sub>j)\] *[OF \(S\<^sub>i,\\<^sub>i) \ (S\<^sub>j,\\<^sub>j)\] by (metis wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) qed simp qed lemma LI_preserves_trm_wf: assumes "(S,\) \\<^sup>* (S',\')" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S')" proof - { fix S \ S' \' assume "(S,\) \ (S',\')" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S')" proof (induction rule: LI_rel.induct) case (Compose S T f S' \) hence "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" and *: "t \ set S \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p t)" "t \ set S' \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p t)" for t by auto hence "wf\<^sub>t\<^sub>r\<^sub>m t" when "t \ set T" for t using that unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p t)" when "t \ set (map Send T)" for t using that unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto thus ?case using * by force next case (Unify S f U \ T S' \) have "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" "wf\<^sub>t\<^sub>r\<^sub>m (Fun f U)" using Unify.prems(1) Unify.hyps(2) wf_trm_subterm[of _ "Fun f U"] by (simp, force) hence range_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_wf_trm[OF Unify.hyps(3)[symmetric]] by simp { fix s assume "s \ set (S@S' \\<^sub>s\<^sub>t \)" hence "\s' \ set (S@S'). s = s' \\<^sub>s\<^sub>t\<^sub>p \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')" using Unify.prems(1) by (auto simp add: subst_apply_strand_def) moreover { fix s' assume s': "s = s' \\<^sub>s\<^sub>t\<^sub>p \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')" "s' \ set (S@S')" from s'(2) have "trms\<^sub>s\<^sub>t\<^sub>p (s' \\<^sub>s\<^sub>t\<^sub>p \) = trms\<^sub>s\<^sub>t\<^sub>p s' \\<^sub>s\<^sub>e\<^sub>t (rm_vars (set (bvars\<^sub>s\<^sub>t\<^sub>p s')) \)" proof (induction s') case (Inequality X F) thus ?case by (induct F) (auto simp add: subst_apply_pairs_def) qed auto hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s)" using wf_trm_subst[OF wf_trms_subst_rm_vars'[OF range_wf]] \wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')\ s'(1) by simp } ultimately have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s)" by auto } thus ?case by auto next case (Equality S \ t t' a S' \) hence "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m t'" by simp_all hence range_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_wf_trm[OF Equality.hyps(2)[symmetric]] by simp { fix s assume "s \ set (S@S' \\<^sub>s\<^sub>t \)" hence "\s' \ set (S@S'). s = s' \\<^sub>s\<^sub>t\<^sub>p \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')" using Equality.prems(1) by (auto simp add: subst_apply_strand_def) moreover { fix s' assume s': "s = s' \\<^sub>s\<^sub>t\<^sub>p \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')" "s' \ set (S@S')" from s'(2) have "trms\<^sub>s\<^sub>t\<^sub>p (s' \\<^sub>s\<^sub>t\<^sub>p \) = trms\<^sub>s\<^sub>t\<^sub>p s' \\<^sub>s\<^sub>e\<^sub>t (rm_vars (set (bvars\<^sub>s\<^sub>t\<^sub>p s')) \)" proof (induction s') case (Inequality X F) thus ?case by (induct F) (auto simp add: subst_apply_pairs_def) qed auto hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s)" using wf_trm_subst[OF wf_trms_subst_rm_vars'[OF range_wf]] \wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s')\ s'(1) by simp } ultimately have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p s)" by auto } thus ?case by auto qed } with assms show ?thesis by (induction rule: rtrancl_induct2) metis+ qed end subsection \Theorem: Soundness of the Lazy Intruder\ context begin private lemma LI_soundness_single: assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1" "(S\<^sub>1,\\<^sub>1) \ (S\<^sub>2,\\<^sub>2)" "\ \\<^sub>c \S\<^sub>2,\\<^sub>2\" shows "\ \\<^sub>c \S\<^sub>1,\\<^sub>1\" using assms(2,1,3) proof (induction rule: LI_rel.induct) case (Compose S T f S' \) hence *: "\{}; S\\<^sub>c \" "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; map Send T\\<^sub>c \" "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \" unfolding constr_sem_c_def by force+ have "ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c Fun f T \ \" using *(2) Compose.hyps(2) ComposeC[OF _ Compose.hyps(3), of "map (\x. x \ \) T"] unfolding subst_compose_def by force thus "\ \\<^sub>c \S@Send (Fun f T)#S',\\" using *(1,3) \\ \\<^sub>c \S@map Send T@S',\\\ by (auto simp add: constr_sem_c_def) next case (Unify S f U \ T S' \) have "(\ \\<^sub>s \) supports \" "\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \" using Unify.prems(2) unfolding constr_sem_c_def by metis+ then obtain \ where \: "\ \\<^sub>s \ \\<^sub>s \ = \" unfolding subst_compose_def by auto have \fun_id: "Fun f U \ \ = Fun f U" "Fun f T \ \ = Fun f T" using Unify.prems(1) trm_subst_ident[of "Fun f U" \] fv_subset_if_in_strand_ik[of "Fun f U" S] Unify.hyps(2) fv_snd_rcv_strand_subset(2)[of S] strand_vars_split(1)[of S "Send (Fun f T)#S'"] unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def apply blast using Unify.prems(1) trm_subst_ident[of "Fun f T" \] unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by fastforce hence \\_disj: "subst_domain \ \ subst_domain \ = {}" "subst_domain \ \ range_vars \ = {}" "subst_domain \ \ range_vars \ = {}" using trm_subst_disj mgu_vars_bounded[OF Unify.hyps(3)[symmetric]] apply (blast,blast) using Unify.prems(1) unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by blast hence \\_support: "\ supports \" "\ supports \" by (simp_all add: subst_support_comp_split[OF \(\ \\<^sub>s \) supports \\]) have "fv (Fun f T) \ fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" "fv (Fun f U) \ fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" using Unify.hyps(2) by force+ hence \_vars_bound: "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t (S@Send (Fun f T)#S')" using mgu_vars_bounded[OF Unify.hyps(3)[symmetric]] by blast have "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; [Send (Fun f T)]\\<^sub>c \" proof - from Unify.hyps(2) have "Fun f U \ \ \ ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \" by blast hence "Fun f U \ \ \ ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \" by blast moreover have "Unifier \ (Fun f T) (Fun f U)" by (fact MGU_is_Unifier[OF mgu_gives_MGU[OF Unify.hyps(3)[symmetric]]]) ultimately have "Fun f T \ \ \ ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \" using \ by (metis \fun_id subst_subst_compose) thus ?thesis by simp qed have "\{}; S\\<^sub>c \" "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \" proof - have "(S@S' \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \ = S@S' \\<^sub>s\<^sub>t \" "(S@S') \\<^sub>s\<^sub>t \ = S@S'" proof - have "subst_domain \ \ vars\<^sub>s\<^sub>t (S@S') = {}" using Unify.prems(1) by (auto simp add: wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) hence "subst_domain \ \ vars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using \\_disj(2) strand_subst_vars_union_bound[of "S@S'" \] by blast thus "(S@S' \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \ = S@S' \\<^sub>s\<^sub>t \" "(S@S') \\<^sub>s\<^sub>t \ = S@S'" using strand_subst_comp \subst_domain \ \ vars\<^sub>s\<^sub>t (S@S') = {}\ by (blast,blast) qed moreover have "subst_idem \" by (fact mgu_gives_subst_idem[OF Unify.hyps(3)[symmetric]]) moreover have "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@S') = {}" "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@S') = {}" using wf_constr_bvars_disj[OF Unify.prems(1)] wf_constr_bvars_disj'[OF Unify.prems(1) \_vars_bound] by auto ultimately have "\{}; S@S'\\<^sub>c \" using \\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \\ \ strand_sem_subst(1)[of \ "S@S' \\<^sub>s\<^sub>t \" "{}" "\ \\<^sub>s \"] strand_sem_subst(2)[of \ "S@S'" "{}" "\ \\<^sub>s \"] strand_sem_subst_subst_idem[of \ "S@S'" "{}" \] unfolding constr_sem_c_def by (metis subst_compose_assoc) thus "\{}; S\\<^sub>c \" "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \" by auto qed show "\ \\<^sub>c \S@Send (Fun f T)#S',\\" using \\_support(1) \\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; [Send (Fun f T)]\\<^sub>c \\ \\{}; S\\<^sub>c \\ \\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \\ by (auto simp add: constr_sem_c_def) next case (Equality S \ t t' a S' \) have "(\ \\<^sub>s \) supports \" "\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \" using Equality.prems(2) unfolding constr_sem_c_def by metis+ then obtain \ where \: "\ \\<^sub>s \ \\<^sub>s \ = \" unfolding subst_compose_def by auto have "fv t \ vars\<^sub>s\<^sub>t (S@Equality a t t'#S')" "fv t' \ vars\<^sub>s\<^sub>t (S@Equality a t t'#S')" by auto moreover have "subst_domain \ \ vars\<^sub>s\<^sub>t (S@Equality a t t'#S') = {}" using Equality.prems(1) unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by auto ultimately have \fun_id: "t \ \ = t" "t' \ \ = t'" using trm_subst_ident[of t \] trm_subst_ident[of t' \] by auto hence \\_disj: "subst_domain \ \ subst_domain \ = {}" "subst_domain \ \ range_vars \ = {}" "subst_domain \ \ range_vars \ = {}" using trm_subst_disj mgu_vars_bounded[OF Equality.hyps(2)[symmetric]] apply (blast,blast) using Equality.prems(1) unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by blast hence \\_support: "\ supports \" "\ supports \" by (simp_all add: subst_support_comp_split[OF \(\ \\<^sub>s \) supports \\]) have "fv t \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" "fv t' \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" by auto hence \_vars_bound: "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t (S@Equality a t t'#S')" using mgu_vars_bounded[OF Equality.hyps(2)[symmetric]] by blast have "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; [Equality a t t']\\<^sub>c \" proof - have "t \ \ = t' \ \" using MGU_is_Unifier[OF mgu_gives_MGU[OF Equality.hyps(2)[symmetric]]] by metis hence "t \ (\ \\<^sub>s \) = t' \ (\ \\<^sub>s \)" by (metis \fun_id subst_subst_compose) hence "t \ \ = t' \ \" by (metis \ subst_subst_compose) thus ?thesis by simp qed have "\{}; S\\<^sub>c \" "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \" proof - have "(S@S' \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \ = S@S' \\<^sub>s\<^sub>t \" "(S@S') \\<^sub>s\<^sub>t \ = S@S'" proof - have "subst_domain \ \ vars\<^sub>s\<^sub>t (S@S') = {}" using Equality.prems(1) by (fastforce simp add: wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def simp del: subst_range.simps) hence "subst_domain \ \ fv\<^sub>s\<^sub>t (S@S') = {}" by blast hence "subst_domain \ \ fv\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" using \\_disj(2) subst_sends_strand_fv_to_img[of "S@S'" \] by blast thus "(S@S' \\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>t \ = S@S' \\<^sub>s\<^sub>t \" "(S@S') \\<^sub>s\<^sub>t \ = S@S'" using strand_subst_comp \subst_domain \ \ vars\<^sub>s\<^sub>t (S@S') = {}\ by (blast,blast) qed moreover have "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@S') = {}" "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \) = {}" "(subst_domain \ \ range_vars \) \ bvars\<^sub>s\<^sub>t (S@S') = {}" using wf_constr_bvars_disj[OF Equality.prems(1)] wf_constr_bvars_disj'[OF Equality.prems(1) \_vars_bound] by auto ultimately have "\{}; S@S'\\<^sub>c \" using \\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \\ \ strand_sem_subst(1)[of \ "S@S' \\<^sub>s\<^sub>t \" "{}" "\ \\<^sub>s \"] strand_sem_subst(2)[of \ "S@S'" "{}" "\ \\<^sub>s \"] strand_sem_subst_subst_idem[of \ "S@S'" "{}" \] mgu_gives_subst_idem[OF Equality.hyps(2)[symmetric]] unfolding constr_sem_c_def by (metis subst_compose_assoc) thus "\{}; S\\<^sub>c \" "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \" by auto qed show "\ \\<^sub>c \S@Equality a t t'#S',\\" using \\_support(1) \\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; [Equality a t t']\\<^sub>c \\ \\{}; S\\<^sub>c \\ \\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \\ by (auto simp add: constr_sem_c_def) qed theorem LI_soundness: assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1" "(S\<^sub>1,\\<^sub>1) \\<^sup>* (S\<^sub>2,\\<^sub>2)" "\ \\<^sub>c \S\<^sub>2, \\<^sub>2\" shows "\ \\<^sub>c \S\<^sub>1, \\<^sub>1\" using assms(2,1,3) proof (induction S\<^sub>2 \\<^sub>2 rule: rtrancl_induct2) case (step S\<^sub>i \\<^sub>i S\<^sub>j \\<^sub>j) thus ?case using LI_preserves_wellformedness[OF \(S\<^sub>1, \\<^sub>1) \\<^sup>* (S\<^sub>i, \\<^sub>i)\ \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1\] LI_soundness_single[OF _ \(S\<^sub>i, \\<^sub>i) \ (S\<^sub>j, \\<^sub>j)\ \\ \\<^sub>c \S\<^sub>j, \\<^sub>j\\] step.IH[OF \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1\] by metis qed metis end subsection \Theorem: Completeness of the Lazy Intruder\ context begin private lemma LI_completeness_single: assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1" "\ \\<^sub>c \S\<^sub>1, \\<^sub>1\" "\simple S\<^sub>1" shows "\S\<^sub>2 \\<^sub>2. (S\<^sub>1,\\<^sub>1) \ (S\<^sub>2,\\<^sub>2) \ (\ \\<^sub>c \S\<^sub>2, \\<^sub>2\)" using not_simple_elim[OF \\simple S\<^sub>1\] proof - { \ \In this case \S\<^sub>1\ isn't simple because it contains an equality constraint, so we can simply proceed with the reduction by computing the MGU for the equation\ assume "\S' S'' a t t'. S\<^sub>1 = S'@Equality a t t'#S'' \ simple S'" then obtain S a t t' S' where S\<^sub>1: "S\<^sub>1 = S@Equality a t t'#S'" "simple S" by moura hence *: "wf\<^sub>s\<^sub>t {} S" "\ \\<^sub>c \S, \\<^sub>1\" "\\<^sub>1 supports \" "t \ \ = t' \ \" using \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1\ wf_eq_fv[of "{}" S t t' S'] fv_snd_rcv_strand_subset(5)[of S] by (auto simp add: constr_sem_c_def wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) from * have "Unifier \ t t'" by simp then obtain \ where \: "Some \ = mgu t t'" "subst_idem \" "subst_domain \ \ range_vars \ \ fv t \ fv t'" using mgu_always_unifies mgu_gives_subst_idem mgu_vars_bounded by metis+ have "\ \\<^sub>\ \" using mgu_gives_MGU[OF \(1)[symmetric]] by (metis \Unifier \ t t'\) hence "\ supports \" using subst_support_if_mgt_subst_idem[OF _ \(2)] by metis hence "(\\<^sub>1 \\<^sub>s \) supports \" using subst_support_comp \\\<^sub>1 supports \\ by metis have "\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \" proof - have "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t S\<^sub>1" using \(3) S\<^sub>1(1) by auto hence "\{}; S\<^sub>1 \\<^sub>s\<^sub>t \\\<^sub>c \" using \subst_idem \\ \\ \\<^sub>\ \\ \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ strand_sem_subst wf_constr_bvars_disj'(1)[OF assms(1)] unfolding subst_idem_def constr_sem_c_def by (metis (no_types) subst_compose_assoc) thus "\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \" using S\<^sub>1(1) by force qed moreover have "(S@Equality a t t'#S', \\<^sub>1) \ (S@S' \\<^sub>s\<^sub>t \, \\<^sub>1 \\<^sub>s \)" using LI_rel.Equality[OF \simple S\ \(1)] S\<^sub>1 by metis ultimately have ?thesis using S\<^sub>1(1) \(\\<^sub>1 \\<^sub>s \) supports \\ by (auto simp add: constr_sem_c_def) } moreover { \ \In this case \S\<^sub>1\ isn't simple because it contains a deduction constraint for a composed term, so we must look at how this composed term is derived under the interpretation \\\\ assume "\S' S'' f T. S\<^sub>1 = S'@Send (Fun f T)#S'' \ simple S'" with assms obtain S f T S' where S\<^sub>1: "S\<^sub>1 = S@Send (Fun f T)#S'" "simple S" by moura hence "wf\<^sub>s\<^sub>t {} S" "\ \\<^sub>c \S, \\<^sub>1\" "\\<^sub>1 supports \" using \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1\ by (auto simp add: constr_sem_c_def wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) \ \Lemma for a common subcase\ have fun_sat: "\ \\<^sub>c \S@(map Send T)@S', \\<^sub>1\" when T: "\t. t \ set T \ ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" proof - have "\t. t \ set T \ \ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; [Send t]\\<^sub>c \" using T by simp hence "\ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \; map Send T\\<^sub>c \" using \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ strand_sem_Send_map by metis moreover have "\ik\<^sub>s\<^sub>t (S@(map Send T)) \\<^sub>s\<^sub>e\<^sub>t \; S'\\<^sub>c \" using \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ S\<^sub>1 by (auto simp add: constr_sem_c_def) ultimately show ?thesis using \\ \\<^sub>c \S, \\<^sub>1\\ \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ by (force simp add: constr_sem_c_def) qed from S\<^sub>1 \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ have "ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c Fun f T \ \" by (auto simp add: constr_sem_c_def) hence ?thesis proof cases \ \Case 1: \\(f(T))\ has been derived using the \AxiomC\ rule.\ case AxiomC hence ex_t: "\t. t \ ik\<^sub>s\<^sub>t S \ Fun f T \ \ = t \ \" by auto show ?thesis proof (cases "\T'. Fun f T' \ ik\<^sub>s\<^sub>t S \ Fun f T \ \ \ Fun f T' \ \") \ \Case 1.1: \f(T)\ is equal to a variable in the intruder knowledge under \\\. Hence there must exists a deduction constraint in the simple prefix of the constraint in which this variable occurs/"is sent" for the first time. Since this variable itself cannot have been derived from the \AxiomC\ rule (because it must be equal under the interpretation to \f(T)\, which is by assumption not in the intruder knowledge under \\\) it must be the case that we can derive it using the \ComposeC\ rule. Hence we can apply the \Compose\ rule of the lazy intruder to \f(T)\.\ case True have "\v. Var v \ ik\<^sub>s\<^sub>t S \ Fun f T \ \ = \ v" proof - obtain t where "t \ ik\<^sub>s\<^sub>t S" "Fun f T \ \ = t \ \" using ex_t by moura thus ?thesis using \\T'. Fun f T' \ ik\<^sub>s\<^sub>t S \ Fun f T \ \ \ Fun f T' \ \\ by (cases t) auto qed hence "\v \ wfrestrictedvars\<^sub>s\<^sub>t S. Fun f T \ \ = \ v" using vars_subset_if_in_strand_ik2[of _ S] by fastforce then obtain v S\<^sub>p\<^sub>r\<^sub>e S\<^sub>s\<^sub>u\<^sub>f where S: "S = S\<^sub>p\<^sub>r\<^sub>e@Send (Var v)#S\<^sub>s\<^sub>u\<^sub>f" "Fun f T \ \ = \ v" "\(\w \ wfrestrictedvars\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e. Fun f T \ \ = \ w)" using \wf\<^sub>s\<^sub>t {} S\ wf_simple_strand_first_Send_var_split[OF _ \simple S\, of "Fun f T" \] by auto hence "\w. Var w \ ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \ \ v \ Var w \ \" by auto moreover have "\T'. Fun f T' \ ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \ Fun f T \ \ \ Fun f T' \ \" using \\T'. Fun f T' \ ik\<^sub>s\<^sub>t S \ Fun f T \ \ \ Fun f T' \ \\ S(1) by (meson contra_subsetD ik_append_subset(1)) hence "\g T'. Fun g T' \ ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \ \ v \ Fun g T' \ \" using S(2) by simp ultimately have "\t \ ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e. \ v \ t \ \" by (metis term.exhaust) hence "\ v \ (ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e) \\<^sub>s\<^sub>e\<^sub>t \" by auto have "ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c \ v" using S\<^sub>1(1) S(1) \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ by (auto simp add: constr_sem_c_def) hence "ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c Fun f T \ \" using \Fun f T \ \ = \ v\ by metis hence "length T = arity f" "public f" "\t. t \ set T \ ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" using \Fun f T \ \ = \ v\ \\ v \ ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t \\ intruder_synth.simps[of "ik\<^sub>s\<^sub>t S\<^sub>p\<^sub>r\<^sub>e \\<^sub>s\<^sub>e\<^sub>t \" "\ v"] by auto hence *: "\t. t \ set T \ ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" using S(1) by (auto intro: ideduct_synth_mono) hence "\ \\<^sub>c \S@(map Send T)@S', \\<^sub>1\" by (metis fun_sat) moreover have "(S@Send (Fun f T)#S', \\<^sub>1) \ (S@map Send T@S', \\<^sub>1)" by (metis LI_rel.Compose[OF \simple S\ \length T = arity f\ \public f\]) ultimately show ?thesis using S\<^sub>1 by auto next \ \Case 1.2: \\(f(T))\ can be derived from an interpreted composed term in the intruder knowledge. Use the \Unify\ rule on this composed term to further reduce the constraint.\ case False then obtain T' where t: "Fun f T' \ ik\<^sub>s\<^sub>t S" "Fun f T \ \ = Fun f T' \ \" by auto hence "fv (Fun f T') \ fv\<^sub>s\<^sub>t S\<^sub>1" using S\<^sub>1(1) fv_subset_if_in_strand_ik'[OF t(1)] fv_snd_rcv_strand_subset(2)[of S] by auto from t have "Unifier \ (Fun f T) (Fun f T')" by simp then obtain \ where \: "Some \ = mgu (Fun f T) (Fun f T')" "subst_idem \" "subst_domain \ \ range_vars \ \ fv (Fun f T) \ fv (Fun f T')" using mgu_always_unifies mgu_gives_subst_idem mgu_vars_bounded by metis+ have "\ \\<^sub>\ \" using mgu_gives_MGU[OF \(1)[symmetric]] by (metis \Unifier \ (Fun f T) (Fun f T')\) hence "\ supports \" using subst_support_if_mgt_subst_idem[OF _ \(2)] by metis hence "(\\<^sub>1 \\<^sub>s \) supports \" using subst_support_comp \\\<^sub>1 supports \\ by metis have "\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \" proof - have "subst_domain \ \ range_vars \ \ fv\<^sub>s\<^sub>t S\<^sub>1" using \(3) S\<^sub>1(1) \fv (Fun f T') \ fv\<^sub>s\<^sub>t S\<^sub>1\ unfolding range_vars_alt_def by (fastforce simp del: subst_range.simps) hence "\{}; S\<^sub>1 \\<^sub>s\<^sub>t \\\<^sub>c \" using \subst_idem \\ \\ \\<^sub>\ \\ \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\ strand_sem_subst wf_constr_bvars_disj'(1)[OF assms(1)] unfolding subst_idem_def constr_sem_c_def by (metis (no_types) subst_compose_assoc) thus "\{}; S@S' \\<^sub>s\<^sub>t \\\<^sub>c \" using S\<^sub>1(1) by force qed moreover have "(S@Send (Fun f T)#S', \\<^sub>1) \ (S@S' \\<^sub>s\<^sub>t \, \\<^sub>1 \\<^sub>s \)" using LI_rel.Unify[OF \simple S\ t(1) \(1)] S\<^sub>1 by metis ultimately show ?thesis using S\<^sub>1(1) \(\\<^sub>1 \\<^sub>s \) supports \\ by (auto simp add: constr_sem_c_def) qed next \ \Case 2: \\(f(T))\ has been derived using the \ComposeC\ rule. Simply use the \Compose\ rule of the lazy intruder to proceed with the reduction.\ case (ComposeC T' g) hence "f = g" "length T = arity f" "public f" and "\x. x \ set T \ ik\<^sub>s\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c x \ \" by auto hence "\ \\<^sub>c \S@(map Send T)@S', \\<^sub>1\" using fun_sat by metis moreover have "(S\<^sub>1, \\<^sub>1) \ (S@(map Send T)@S', \\<^sub>1)" using S\<^sub>1 LI_rel.Compose[OF \simple S\ \length T = arity f\ \public f\] by metis ultimately show ?thesis by metis qed } moreover have "\A B X F. S\<^sub>1 = A@Inequality X F#B \ ineq_model \ X F" using assms(2) by (auto simp add: constr_sem_c_def) ultimately show ?thesis using not_simple_elim[OF \\simple S\<^sub>1\] by metis qed theorem LI_completeness: assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1" "\ \\<^sub>c \S\<^sub>1, \\<^sub>1\" shows "\S\<^sub>2 \\<^sub>2. (S\<^sub>1,\\<^sub>1) \\<^sup>* (S\<^sub>2,\\<^sub>2) \ simple S\<^sub>2 \ (\ \\<^sub>c \S\<^sub>2, \\<^sub>2\)" proof (cases "simple S\<^sub>1") case False let ?Stuck = "\S\<^sub>2 \\<^sub>2. \(\S\<^sub>3 \\<^sub>3. (S\<^sub>2,\\<^sub>2) \ (S\<^sub>3,\\<^sub>3) \ (\ \\<^sub>c \S\<^sub>3, \\<^sub>3\))" let ?Sats = "{((S,\),(S',\')). (S,\) \ (S',\') \ (\ \\<^sub>c \S, \\) \ (\ \\<^sub>c \S', \'\)}" have simple_if_stuck: "\S\<^sub>2 \\<^sub>2. \(S\<^sub>1,\\<^sub>1) \\<^sup>+ (S\<^sub>2,\\<^sub>2); \ \\<^sub>c \S\<^sub>2, \\<^sub>2\; ?Stuck S\<^sub>2 \\<^sub>2\ \ simple S\<^sub>2" using LI_completeness_single LI_preserves_wellformedness \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1\ trancl_into_rtrancl by metis have base: "\b. ((S\<^sub>1,\\<^sub>1),b) \ ?Sats" using LI_completeness_single[OF assms False] assms(2) by auto have *: "\S \ S' \'. ((S,\),(S',\')) \ ?Sats\<^sup>+ \ (S,\) \\<^sup>+ (S',\') \ (\ \\<^sub>c \S', \'\)" proof - fix S \ S' \' assume "((S,\),(S',\')) \ ?Sats\<^sup>+" thus "(S,\) \\<^sup>+ (S',\') \ (\ \\<^sub>c \S', \'\)" by (induct rule: trancl_induct2) auto qed have "\S\<^sub>2 \\<^sub>2. ((S\<^sub>1,\\<^sub>1),(S\<^sub>2,\\<^sub>2)) \ ?Sats\<^sup>+ \ ?Stuck S\<^sub>2 \\<^sub>2" proof (rule ccontr) assume "\(\S\<^sub>2 \\<^sub>2. ((S\<^sub>1,\\<^sub>1),(S\<^sub>2,\\<^sub>2)) \ ?Sats\<^sup>+ \ ?Stuck S\<^sub>2 \\<^sub>2)" hence sat_not_stuck: "\S\<^sub>2 \\<^sub>2. ((S\<^sub>1,\\<^sub>1),(S\<^sub>2,\\<^sub>2)) \ ?Sats\<^sup>+ \ \?Stuck S\<^sub>2 \\<^sub>2" by blast have "\S \. ((S\<^sub>1,\\<^sub>1),(S,\)) \ ?Sats\<^sup>+ \ (\b. ((S,\),b) \ ?Sats)" proof (intro allI impI) fix S \ assume a: "((S\<^sub>1,\\<^sub>1),(S,\)) \ ?Sats\<^sup>+" have "\b. ((S\<^sub>1,\\<^sub>1),b) \ ?Sats\<^sup>+ \ \c. b \ c \ ((S\<^sub>1,\\<^sub>1),c) \ ?Sats\<^sup>+" proof - fix b assume in_sat: "((S\<^sub>1,\\<^sub>1),b) \ ?Sats\<^sup>+" hence "\c. (b,c) \ ?Sats" using * sat_not_stuck by (cases b) blast thus "\c. b \ c \ ((S\<^sub>1,\\<^sub>1),c) \ ?Sats\<^sup>+" using trancl_into_trancl[OF in_sat] by blast qed hence "\S' \'. (S,\) \ (S',\') \ ((S\<^sub>1,\\<^sub>1),(S',\')) \ ?Sats\<^sup>+" using a by auto then obtain S' \' where S'\': "(S,\) \ (S',\')" "((S\<^sub>1,\\<^sub>1),(S',\')) \ ?Sats\<^sup>+" by auto hence "\ \\<^sub>c \S', \'\" using * by blast moreover have "(S\<^sub>1, \\<^sub>1) \\<^sup>+ (S,\)" using a trancl_mono by blast ultimately have "((S,\),(S',\')) \ ?Sats" using S'\'(1) * a by blast thus "\b. ((S,\),b) \ ?Sats" using S'\'(2) by blast qed hence "\f. \i::nat. (f i, f (Suc i)) \ ?Sats" using infinite_chain_intro'[OF base] by blast moreover have "?Sats \ LI_rel\<^sup>+" by auto hence "\(\f. \i::nat. (f i, f (Suc i)) \ ?Sats)" using LI_no_infinite_chain infinite_chain_mono by blast ultimately show False by auto qed hence "\S\<^sub>2 \\<^sub>2. (S\<^sub>1, \\<^sub>1) \\<^sup>+ (S\<^sub>2, \\<^sub>2) \ simple S\<^sub>2 \ (\ \\<^sub>c \S\<^sub>2, \\<^sub>2\)" using simple_if_stuck * by blast thus ?thesis by (meson trancl_into_rtrancl) qed (blast intro: \\ \\<^sub>c \S\<^sub>1, \\<^sub>1\\) end subsection \Corollary: Soundness and Completeness as a Single Theorem\ corollary LI_soundness_and_completeness: assumes "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S\<^sub>1 \\<^sub>1" shows "\ \\<^sub>c \S\<^sub>1, \\<^sub>1\ \ (\S\<^sub>2 \\<^sub>2. (S\<^sub>1,\\<^sub>1) \\<^sup>* (S\<^sub>2,\\<^sub>2) \ simple S\<^sub>2 \ (\ \\<^sub>c \S\<^sub>2, \\<^sub>2\))" by (metis LI_soundness[OF assms] LI_completeness[OF assms]) end end