diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Compositionality.thy b/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Compositionality.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Compositionality.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/Stateful_Compositionality.thy @@ -1,3086 +1,3086 @@ (* (C) Copyright Andreas Viktor Hess, DTU, 2018-2020 All Rights Reserved. Redistribution and use in source and binary forms, with or without modification, are permitted provided that the following conditions are met: - Redistributions of source code must retain the above copyright notice, this list of conditions and the following disclaimer. - Redistributions in binary form must reproduce the above copyright notice, this list of conditions and the following disclaimer in the documentation and/or other materials provided with the distribution. - Neither the name of the copyright holder nor the names of its contributors may be used to endorse or promote products derived from this software without specific prior written permission. THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. *) (* Title: Stateful_Compositionality.thy Author: Andreas Viktor Hess, DTU *) section \Stateful Protocol Compositionality\ theory Stateful_Compositionality imports Stateful_Typing Parallel_Compositionality Labeled_Stateful_Strands begin subsection \Small Lemmata\ lemma (in typed_model) wt_subst_sstp_vars_type_subset: fixes a::"('fun,'var) stateful_strand_step" assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "\t \ subst_range \. fv t = {} \ (\x. t = Var x)" shows "\ ` Var ` fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ \ ` Var ` fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" (is ?A) and "\ ` Var ` set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)) = \ ` Var ` set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a)" (is ?B) and "\ ` Var ` vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ \ ` Var ` vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" (is ?C) proof - show ?A proof fix \ assume \: "\ \ \ ` Var ` fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" then obtain x where x: "x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "\ (Var x) = \" by moura - + show "\ \ \ ` Var ` fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" proof (cases "x \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a") case False hence "\y \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a. \ y = Var x" proof (cases a) case (NegChecks X F G) hence *: "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" "x \ set X" using fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck(1)[of X "F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \" "G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \"] fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck(1)[of X F G] False x(1) by fastforce+ - + obtain y where y: "y \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" "x \ fv (rm_vars (set X) \ y)" using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_obtain_var[of _ _ "rm_vars (set X) \"] fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_obtain_var[of _ _ "rm_vars (set X) \"] *(1) by blast - + have "fv (rm_vars (set X) \ z) = {} \ (\u. rm_vars (set X) \ z = Var u)" for z using assms(2) rm_vars_img_subset[of "set X" \] by blast hence "rm_vars (set X) \ y = Var x" using y(2) by fastforce hence "\y \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a. rm_vars (set X) \ y = Var x" using y fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck(1)[of X F G] NegChecks *(2) by fastforce thus ?thesis by (metis (full_types) *(2) term.inject(1)) qed (use assms(2) x(1) subst_apply_img_var'[of x _ \] in fastforce)+ then obtain y where y: "y \ fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" "\ y = Var x" by moura hence "\ (Var y) = \" using x(2) assms(1) by (simp add: wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) thus ?thesis using y(1) by auto qed (use x in auto) qed show ?B by (metis bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_subst) show ?C proof fix \ assume \: "\ \ \ ` Var ` vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" then obtain x where x: "x \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" "\ (Var x) = \" by moura - + show "\ \ \ ` Var ` vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" proof (cases "x \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a") case False hence "\y \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a. \ y = Var x" proof (cases a) case (NegChecks X F G) hence *: "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" "x \ set X" using vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck[of X "F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \" "G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \"] vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck[of X F G] False x(1) by (fastforce, blast) - + obtain y where y: "y \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" "x \ fv (rm_vars (set X) \ y)" using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_obtain_var[of _ _ "rm_vars (set X) \"] fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_obtain_var[of _ _ "rm_vars (set X) \"] *(1) by blast - + have "fv (rm_vars (set X) \ z) = {} \ (\u. rm_vars (set X) \ z = Var u)" for z using assms(2) rm_vars_img_subset[of "set X" \] by blast hence "rm_vars (set X) \ y = Var x" using y(2) by fastforce hence "\y \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a. rm_vars (set X) \ y = Var x" using y vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p_NegCheck[of X F G] NegChecks by blast thus ?thesis by (metis (full_types) *(2) term.inject(1)) qed (use assms(2) x(1) subst_apply_img_var'[of x _ \] in fastforce)+ then obtain y where y: "y \ vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" "\ y = Var x" by moura hence "\ (Var y) = \" using x(2) assms(1) by (simp add: wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) thus ?thesis using y(1) by auto qed (use x in auto) qed qed lemma (in typed_model) wt_subst_lsst_vars_type_subset: fixes A::"('fun,'var,'a) labeled_stateful_strand" assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "\t \ subst_range \. fv t = {} \ (\x. t = Var x)" shows "\ ` Var ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ \ ` Var ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" (is ?A) and "\ ` Var ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = \ ` Var ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" (is ?B) and "\ ` Var ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) \ \ ` Var ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" (is ?C) proof - have "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) = vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \ vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) = fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \ fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) = set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" when "a = (l,b)" for a l b and A::"('fun,'var,'a) labeled_stateful_strand" using that unlabel_Cons(1)[of l b A] unlabel_subst[of "a#A" \] subst_lsst_cons[of a A \] subst_sst_cons[of b "unlabel A" \] subst_apply_labeled_stateful_strand_step.simps(1)[of l b \] vars\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons[of l b A] vars\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons[of l "b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" "A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] fv\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons[of l b A] fv\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons[of l "b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" "A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons[of l b A] bvars\<^sub>s\<^sub>s\<^sub>t_unlabel_Cons[of l "b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" "A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \"] by simp_all hence *: "\ ` Var ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = \ ` Var ` vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ \ ` Var ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "\ ` Var ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) = \ ` Var ` vars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \ \ ` Var ` vars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "\ ` Var ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = \ ` Var ` fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \) \ \ ` Var ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "\ ` Var ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) = \ ` Var ` fv\<^sub>s\<^sub>s\<^sub>t\<^sub>p b \ \ ` Var ` fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "\ ` Var ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = \ ` Var ` set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (b \\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)) \ \ ` Var ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" "\ ` Var ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) = \ ` Var ` set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p b) \ \ ` Var ` bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" when "a = (l,b)" for a l b and A::"('fun,'var,'a) labeled_stateful_strand" using that by fast+ have "?A \ ?B \ ?C" proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) - + show ?case using Cons.IH wt_subst_sstp_vars_type_subset[OF assms, of b] *[OF a, of A] by (metis Un_mono) qed simp thus ?A ?B ?C by metis+ qed lemma (in stateful_typed_model) fv_pair_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subset: assumes "d \ set D" shows "fv (pair (snd d)) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" using assms unfolding pair_def by (induct D) (auto simp add: unlabel_def) lemma (in stateful_typed_model) labeled_sat_ineq_lift: assumes "\M; map (\d. \X\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\dbproj i D. d \ set Di]\\<^sub>d \" (is "?R1 D") and "\(j,p) \ {(i,t,s)} \ set D \ set Di. \(k,q) \ {(i,t,s)} \ set D \ set Di. (\\. Unifier \ (pair p) (pair q)) \ j = k" (is "?R2 D") shows "\M; map (\d. \X\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\D. d \ set Di]\\<^sub>d \" using assms proof (induction D) case (Cons dl D) obtain d l where dl: "dl = (l,d)" by (metis surj_pair) have 1: "?R1 D" proof (cases "i = l") case True thus ?thesis using Cons.prems(1) dl by (cases "dl \ set Di") auto next case False thus ?thesis using Cons.prems(1) dl by auto qed have "set D \ set (dl#D)" by auto hence 2: "?R2 D" using Cons.prems(2) by blast have "i \ l \ dl \ set Di \ \M; [\X\\\: [(pair (t,s), pair (snd dl))]\\<^sub>s\<^sub>t]\\<^sub>d \" using Cons.prems(1) dl by (auto simp add: ineq_model_def) moreover have "\\. Unifier \ (pair (t,s)) (pair d) \ i = l" using Cons.prems(2) dl by force ultimately have 3: "dl \ set Di \ \M; [\X\\\: [(pair (t,s), pair (snd dl))]\\<^sub>s\<^sub>t]\\<^sub>d \" using strand_sem_not_unif_is_sat_ineq[of "pair (t,s)" "pair d"] dl by fastforce show ?case using Cons.IH[OF 1 2] 3 dl by auto qed simp lemma (in stateful_typed_model) labeled_sat_ineq_dbproj: assumes "\M; map (\d. \X\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\D. d \ set Di]\\<^sub>d \" (is "?P D") shows "\M; map (\d. \X\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\dbproj i D. d \ set Di]\\<^sub>d \" (is "?Q D") using assms proof (induction D) case (Cons di D) obtain d j where di: "di = (j,d)" by (metis surj_pair) have "?P D" using Cons.prems by (cases "di \ set Di") auto hence IH: "?Q D" by (metis Cons.IH) show ?case using di IH proof (cases "i = j \ di \ set Di") case True have 1: "\M; [\X\\\: [(pair (t,s), pair (snd di))]\\<^sub>s\<^sub>t]\\<^sub>d \" using Cons.prems True by auto have 2: "dbproj i (di#D) = di#dbproj i D" using True dbproj_Cons(1) di by auto show ?thesis using 1 2 IH by auto qed auto qed simp lemma (in stateful_typed_model) labeled_sat_ineq_dbproj_sem_equiv: assumes "\(j,p) \ ((\(t, s). (i, t, s)) ` set F') \ set D. \(k,q) \ ((\(t, s). (i, t, s)) ` set F') \ set D. (\\. Unifier \ (pair p) (pair q)) \ j = k" and "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (map snd D) \ set X = {}" shows "\M; map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd D))\\<^sub>d \ \ \M; map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))\\<^sub>d \" proof - let ?A = "set (map snd D) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" let ?B = "set (map snd (dbproj i D)) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" let ?C = "set (map snd D) - set (map snd (dbproj i D))" let ?F = "(\(t, s). (i, t, s)) ` set F'" let ?P = "\\. subst_domain \ = set X \ ground (subst_range \)" have 1: "\(t, t') \ set (map snd D). (fv t \ fv t') \ set X = {}" "\(t, t') \ set (map snd (dbproj i D)). (fv t \ fv t') \ set X = {}" using assms(2) dbproj_subset[of i D] unfolding unlabel_def by force+ have 2: "?B \ ?A" by auto have 3: "\Unifier \ (pair f) (pair d)" when f: "f \ set F'" and d: "d \ set (map snd D) - set (map snd (dbproj i D))" for f d and \::"('fun,'var) subst" proof - obtain k where k: "(k,d) \ set D - set (dbproj i D)" using d by force - + have "(i,f) \ ((\(t, s). (i, t, s)) ` set F') \ set D" "(k,d) \ ((\(t, s). (i, t, s)) ` set F') \ set D" using f k by auto hence "i = k" when "Unifier \ (pair f) (pair d)" for \ using assms(1) that by blast moreover have "k \ i" using k d by simp ultimately show ?thesis by metis qed have "f \\<^sub>p \ \ d \\<^sub>p \" when "f \ set F'" "d \ ?C" for f d and \::"('fun,'var) subst" by (metis fun_pair_eq_subst 3[OF that]) hence "f \\<^sub>p (\ \\<^sub>s \) \ ?C \\<^sub>p\<^sub>s\<^sub>e\<^sub>t (\ \\<^sub>s \)" when "f \ set F'" for f and \::"('fun,'var) subst" using that by blast moreover have "?C \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ = ?C \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" when "?P \" for \ using assms(2) that pairs_substI[of \ "(set (map snd D) - set (map snd (dbproj i D)))"] by blast ultimately have 4: "f \\<^sub>p (\ \\<^sub>s \) \ ?C \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" when "f \ set F'" "?P \" for f and \::"('fun,'var) subst" by (metis that subst_pairs_compose) { fix f and \::"('fun,'var) subst" assume "f \ set F'" "?P \" hence "f \\<^sub>p (\ \\<^sub>s \) \ ?C \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" by (metis 4) hence "f \\<^sub>p (\ \\<^sub>s \) \ ?A - ?B" by force } hence 5: "\f\set F'. \\. ?P \ \ f \\<^sub>p (\ \\<^sub>s \) \ ?A - ?B" by metis show ?thesis using negchecks_model_db_subset[OF 2] negchecks_model_db_supset[OF 2 5] tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_sem_equiv[OF 1(1)] tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_sem_equiv[OF 1(2)] tr_NegChecks_constr_iff(1) strand_sem_eq_defs(2) by (metis (no_types, lifting)) qed lemma (in stateful_typed_model) labeled_sat_eqs_list_all: assumes "\(j, p) \ {(i,t,s)} \ set D. \(k,q) \ {(i,t,s)} \ set D. (\\. Unifier \ (pair p) (pair q)) \ j = k" (is "?P D") and "\M; map (\d. \ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t) D\\<^sub>d \" (is "?Q D") shows "list_all (\d. fst d = i) D" using assms proof (induction D rule: List.rev_induct) case (snoc di D) obtain d j where di: "di = (j,d)" by (metis surj_pair) have "pair (t,s) \ \ = pair d \ \" using di snoc.prems(2) by auto hence "\\. Unifier \ (pair (t,s)) (pair d)" by auto hence 1: "i = j" using snoc.prems(1) di by fastforce - + have "set D \ set (D@[di])" by auto hence 2: "?P D" using snoc.prems(1) by blast have 3: "?Q D" using snoc.prems(2) by auto - + show ?case using di 1 snoc.IH[OF 2 3] by simp qed simp lemma (in stateful_typed_model) labeled_sat_eqs_subseqs: assumes "Di \ set (subseqs D)" and "\(j, p) \ {(i,t,s)} \ set D. \(k, q) \ {(i,t,s)} \ set D. (\\. Unifier \ (pair p) (pair q)) \ j = k" (is "?P D") and "\M; map (\d. \ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t) Di\\<^sub>d \" shows "Di \ set (subseqs (dbproj i D))" proof - have "set Di \ set D" by (rule subseqs_subset[OF assms(1)]) hence "?P Di" using assms(2) by blast thus ?thesis using labeled_sat_eqs_list_all[OF _ assms(3)] subseqs_mem_dbproj[OF assms(1)] by simp qed lemma (in stateful_typed_model) dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p: assumes "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel S)" shows "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S))" using assms proof (induction S) case (Cons a S) have prems: "tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd a)" "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel S)" using Cons.prems unlabel_Cons(2)[of a S] by simp_all hence IH: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S))" by (metis Cons.IH) obtain l b where a: "a = (l,b)" by (metis surj_pair) with Cons show ?case proof (cases b) case (Equality c t t') hence "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#S) = a#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S" by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_Cons(3) a) thus ?thesis using a IH prems by fastforce next case (NegChecks X F G) hence "dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#S) = a#dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t S" by (metis dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_Cons(7) a) thus ?thesis using a IH prems by fastforce qed auto qed simp lemma (in stateful_typed_model) setops\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq: "setops\<^sub>s\<^sub>s\<^sub>t (unlabel (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)) = setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) thus ?case using Cons.IH by (cases b) (simp_all add: setops\<^sub>s\<^sub>s\<^sub>t_def) qed simp subsection \Locale Setup and Definitions\ locale labeled_stateful_typed_model = stateful_typed_model arity public Ana \ Pair + labeled_typed_model arity public Ana \ label_witness1 label_witness2 for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,'var) term \ (('fun,'var) term list \ ('fun,'var) term list)" and \::"('fun,'var) term \ ('fun,'atom::finite) term_type" and Pair::"'fun" and label_witness1::"'lbl" and label_witness2::"'lbl" begin definition lpair where "lpair lp \ case lp of (i,p) \ (i,pair p)" lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_pair_image[simp]: "lpair ` (setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,send\t\)) = {}" "lpair ` (setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,receive\t\)) = {}" "lpair ` (setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,\ac: t \ t'\)) = {}" "lpair ` (setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,insert\t,s\)) = {(i, pair (t,s))}" "lpair ` (setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,delete\t,s\)) = {(i, pair (t,s))}" "lpair ` (setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,\ac: t \ s\)) = {(i, pair (t,s))}" "lpair ` (setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (i,\X\\\: F \\: F'\)) = ((\(t,s). (i, pair (t,s))) ` set F')" unfolding lpair_def by force+ definition par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t where - "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\::('fun,'var,'lbl) labeled_stateful_strand) (Secrets::('fun,'var) terms) \ + "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (\::('fun,'var,'lbl) labeled_stateful_strand) (Secrets::('fun,'var) terms) \ (\l1 l2. l1 \ l2 \ GSMP_disjoint (trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l1 \)) (trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l2 \)) Secrets) \ ground Secrets \ (\s \ Secrets. \s' \ subterms s. {} \\<^sub>c s' \ s' \ Secrets) \ (\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t \. (\\. Unifier \ (pair p) (pair q)) \ i = j)" definition declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ \ {t. \\, receive\t\\ \ set \} \\<^sub>s\<^sub>e\<^sub>t \" definition strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t ("_ leaks _ under _") where "(\::('fun,'var,'lbl) labeled_stateful_strand) leaks Secrets under \ \ (\t \ Secrets - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \. \n. \ \\<^sub>s (proj_unl n \@[send\t\]))" definition typing_cond\<^sub>s\<^sub>s\<^sub>t where "typing_cond\<^sub>s\<^sub>s\<^sub>t \ \ wf\<^sub>s\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t \) \ tfr\<^sub>s\<^sub>s\<^sub>t \" type_synonym ('a,'b,'c) labeleddbstate = "('c strand_label \ (('a,'b) term \ ('a,'b) term)) set" type_synonym ('a,'b,'c) labeleddbstatelist = "('c strand_label \ (('a,'b) term \ ('a,'b) term)) list" text \ For proving the compositionality theorem for stateful constraints the idea is to first define a variant of the reduction technique that was used to establish the stateful typing result. This variant performs database-state projections, and it allows us to reduce the compositionality problem for stateful constraints to ordinary constraints. \ fun tr\<^sub>p\<^sub>c:: "('fun,'var,'lbl) labeled_stateful_strand \ ('fun,'var,'lbl) labeleddbstatelist \ ('fun,'var,'lbl) labeled_strand list" where "tr\<^sub>p\<^sub>c [] D = [[]]" | "tr\<^sub>p\<^sub>c ((i,send\t\)#A) D = map ((#) (i,send\t\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>c A D)" | "tr\<^sub>p\<^sub>c ((i,receive\t\)#A) D = map ((#) (i,receive\t\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>c A D)" | "tr\<^sub>p\<^sub>c ((i,\ac: t \ t'\)#A) D = map ((#) (i,\ac: t \ t'\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>c A D)" | "tr\<^sub>p\<^sub>c ((i,insert\t,s\)#A) D = tr\<^sub>p\<^sub>c A (List.insert (i,(t,s)) D)" | "tr\<^sub>p\<^sub>c ((i,delete\t,s\)#A) D = ( concat (map (\Di. map (\B. (map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\dbproj i D. d \ set Di])@B) (tr\<^sub>p\<^sub>c A [d\D. d \ set Di])) (subseqs (dbproj i D))))" | "tr\<^sub>p\<^sub>c ((i,\ac: t \ s\)#A) D = concat (map (\B. map (\d. (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#B) (dbproj i D)) (tr\<^sub>p\<^sub>c A D))" | "tr\<^sub>p\<^sub>c ((i,\X\\\: F \\: F' \)#A) D = map ((@) (map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D))))) (tr\<^sub>p\<^sub>c A D)" subsection \Small Lemmata\ lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_nil: assumes "ground Sec" "\s \ Sec. \s'\subterms s. {} \\<^sub>c s' \ s' \ Sec" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t [] Sec" using assms unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by simp lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subset: assumes A: "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A Sec" and BA: "set B \ set A" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t B Sec" proof - let ?L = "\n A. trms\<^sub>s\<^sub>s\<^sub>t (proj_unl n A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl n A)" have "?L n B \ ?L n A" for n using trms\<^sub>s\<^sub>s\<^sub>t_mono[OF proj_set_mono(2)[OF BA]] setops\<^sub>s\<^sub>s\<^sub>t_mono[OF proj_set_mono(2)[OF BA]] by blast hence "GSMP_disjoint (?L m B) (?L n B) Sec" when nm: "m \ n" for n m::'lbl using GSMP_disjoint_subset[of "?L m A" "?L n A" Sec "?L m B" "?L n B"] A nm unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by simp thus "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t B Sec" using A setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_mono[OF BA] unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by blast qed lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_split: assumes "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A@B) Sec" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A Sec" "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t B Sec" using par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subset[OF assms] by simp_all lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj: assumes "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A Sec" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A) Sec" using par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subset[OF assms] by simp lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t: assumes A: "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A S" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A) S" proof (unfold par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def case_prod_unfold; intro conjI) show "ground S" "\s \ S. \s' \ subterms s. {} \\<^sub>c s' \ s' \ S" using A unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by fast+ let ?M = "\l B. (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l B) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l B))" let ?P = "\B. \l1 l2. l1 \ l2 \ GSMP_disjoint (?M l1 B) (?M l2 B) S" let ?Q = "\B. \p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B. \q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B. (\\. Unifier \ (pair (snd p)) (pair (snd q))) \ fst p = fst q" have "?P A" "?Q A" using A unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def case_prod_unfold by blast+ thus "?P (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" "?Q (dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" by (metis setops\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq trms\<^sub>s\<^sub>s\<^sub>t_unlabel_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq proj_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t, metis setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_dual\<^sub>l\<^sub>s\<^sub>s\<^sub>t_eq) qed lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_subst: assumes A: "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A S" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "subst_domain \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = {}" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) S" proof (unfold par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def case_prod_unfold; intro conjI) show "ground S" "\s \ S. \s' \ subterms s. {} \\<^sub>c s' \ s' \ S" using A unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by fast+ let ?N = "\l B. trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l B) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l B)" define M where "M \ \l (B::('fun,'var,'lbl) labeled_stateful_strand). ?N l B" let ?P = "\p q. \\. Unifier \ (pair (snd p)) (pair (snd q))" let ?Q = "\B. \p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B. \q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B. ?P p q \ fst p = fst q" let ?R = "\B. \l1 l2. l1 \ l2 \ GSMP_disjoint (?N l1 B) (?N l2 B) S" have d: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l A) \ subst_domain \ = {}" for l using \(3) unfolding proj_def bvars\<^sub>s\<^sub>s\<^sub>t_def unlabel_def by auto have "GSMP_disjoint (M l1 A) (M l2 A) S" when l: "l1 \ l2" for l1 l2 using l A unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def M_def by presburger moreover have "M l (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) = (M l A) \\<^sub>s\<^sub>e\<^sub>t \" for l using fun_pair_subst_set[of \ "setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l A)", symmetric] trms\<^sub>s\<^sub>s\<^sub>t_subst[OF d[of l]] setops\<^sub>s\<^sub>s\<^sub>t_subst[OF d[of l]] proj_subst[of l A \] unfolding M_def unlabel_subst by auto ultimately have "GSMP_disjoint (M l1 (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) (M l2 (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)) S" when l: "l1 \ l2" for l1 l2 using l GSMP_wt_subst_subset[OF _ \(1,2), of _ "M l1 A"] GSMP_wt_subst_subset[OF _ \(1,2), of _ "M l2 A"] unfolding GSMP_disjoint_def by fastforce thus "?R (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" unfolding M_def by blast - have "?Q A" using A unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by force + have "?Q A" using A unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by force thus "?Q (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using \(3) proof (induction A) case (Cons a A) obtain l b where a: "a = (l,b)" by (metis surj_pair) have 0: "bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A) = set (bvars\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd a)) \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" unfolding bvars\<^sub>s\<^sub>s\<^sub>t_def unlabel_def by simp have "?Q A" "subst_domain \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t A = {}" using Cons.prems 0 unfolding setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto hence IH: "?Q (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using Cons.IH unfolding setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by blast - + have 1: "fst p = fst q" when p: "p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" and q: "q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" and pq: "?P p q" for p q using a p q pq by (cases b) auto have 2: "fst p = fst q" when p: "p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" and q: "q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p (a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \)" and pq: "?P p q" for p q proof - obtain p' X where p': - "p' \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "fst p = fst p'" + "p' \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "fst p = fst p'" "X \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A)" "snd p = snd p' \\<^sub>p rm_vars X \" using setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_subst[OF p] 0 by blast obtain q' Y where q': - "q' \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" "fst q = fst q'" + "q' \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p a" "fst q = fst q'" "Y \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A)" "snd q = snd q' \\<^sub>p rm_vars Y \" using setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p_in_subst[OF q] 0 by blast have "pair (snd p) = pair (snd p') \ \" "pair (snd q) = pair (snd q') \ \" using fun_pair_subst[of "snd p'" "rm_vars X \"] fun_pair_subst[of "snd q'" "rm_vars Y \"] p'(3,4) q'(3,4) Cons.prems(2) rm_vars_apply'[of \ X] rm_vars_apply'[of \ Y] by fastforce+ hence "\\. Unifier \ (pair (snd p')) (pair (snd q'))" using pq Unifier_comp' by metis thus ?thesis using Cons.prems p'(1,2) q'(1,2) by simp qed show ?case by (metis 1 2 IH Un_iff setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_cons subst_lsst_cons) qed simp qed lemma wf_pair_negchecks_map': assumes "wf\<^sub>s\<^sub>t X (unlabel A)" shows "wf\<^sub>s\<^sub>t X (unlabel (map (\G. (i,\Y\\\: (F@G)\\<^sub>s\<^sub>t)) M@A))" using assms by (induct M) auto lemma wf_pair_eqs_ineqs_map': fixes A::"('fun,'var,'lbl) labeled_strand" assumes "wf\<^sub>s\<^sub>t X (unlabel A)" "Di \ set (subseqs (dbproj i D))" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X" shows "wf\<^sub>s\<^sub>t X (unlabel ( (map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\dbproj i D. d \ set Di])@A))" proof - let ?f = "[d\dbproj i D. d \ set Di]" define c1 where c1: "c1 = map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di" define c2 where c2: "c2 = map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) ?f" define c3 where c3: "c3 = map (\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t) (unlabel Di)" define c4 where c4: "c4 = map (\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) (unlabel ?f)" have ci_eqs: "c3 = unlabel c1" "c4 = unlabel c2" unfolding c1 c2 c3 c4 unlabel_def by auto have 1: "wf\<^sub>s\<^sub>t X (unlabel (c2@A))" using wf_fun_pair_ineqs_map[OF assms(1)] ci_eqs(2) unlabel_append[of c2 A] c4 by metis - have 2: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel Di) \ X" + have 2: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel Di) \ X" using assms(3) subseqs_set_subset(1)[OF assms(2)] - unfolding unlabel_def + unfolding unlabel_def by fastforce { fix B::"('fun,'var) strand" assume "wf\<^sub>s\<^sub>t X B" hence "wf\<^sub>s\<^sub>t X (unlabel c1@B)" using 2 unfolding c1 unlabel_def by (induct Di) auto } thus ?thesis using 1 unfolding c1 c2 unlabel_def by simp qed lemma trms\<^sub>s\<^sub>s\<^sub>t_setops\<^sub>s\<^sub>s\<^sub>t_wt_instance_ex: defines "M \ \A. trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" assumes B: "\b \ set B. \a \ set A. \\. b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "\t \ M B. \s \ M A. \\. t = s \ \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" proof let ?P = "\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" fix t assume "t \ M B" then obtain b where b: "b \ set B" "t \ trms\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd b) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t\<^sub>p (snd b)" unfolding M_def unfolding unlabel_def trms\<^sub>s\<^sub>s\<^sub>t_def setops\<^sub>s\<^sub>s\<^sub>t_def by auto then obtain a \ where a: "a \ set A" "b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using B by meson note \' = wt_subst_rm_vars[OF \(1)] wf_trms_subst_rm_vars'[OF \(2)] have "t \ M (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using b(2) a unfolding M_def subst_apply_labeled_stateful_strand_def unlabel_def trms\<^sub>s\<^sub>s\<^sub>t_def setops\<^sub>s\<^sub>s\<^sub>t_def by auto moreover have "\s \ M A. \\. t = s \ \ \ ?P \" when "t \ trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subst'[OF that] \' unfolding M_def by blast moreover have "\s \ M A. \\. t = s \ \ \ ?P \" when t: "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A \\<^sub>s\<^sub>s\<^sub>t \)" proof - obtain p where p: "p \ setops\<^sub>s\<^sub>s\<^sub>t (unlabel A \\<^sub>s\<^sub>s\<^sub>t \)" "t = pair p" using t by blast then obtain q X where q: "q \ setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "p = q \\<^sub>p rm_vars (set X) \" using setops\<^sub>s\<^sub>s\<^sub>t_subst'[OF p(1)] by blast hence "t = pair q \ rm_vars (set X) \" using fun_pair_subst[of q "rm_vars (set X) \"] p(2) by presburger thus ?thesis using \'[of "set X"] q(1) unfolding M_def by blast qed ultimately show "\s \ M A. \\. t = s \ \ \ ?P \" unfolding M_def unlabel_subst by fast qed lemma setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_wt_instance_ex: assumes B: "\b \ set B. \a \ set A. \\. b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "\p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B. \q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \\. fst p = fst q \ snd p = snd q \\<^sub>p \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" proof let ?P = "\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" fix p assume "p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B" then obtain b where b: "b \ set B" "p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p b" unfolding setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by blast then obtain a \ where a: "a \ set A" "b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using B by meson hence p: "p \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (A \\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" using b(2) unfolding setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def subst_apply_labeled_stateful_strand_def by auto - + obtain X q where q: "q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "fst p = fst q" "snd p = snd q \\<^sub>p rm_vars X \" using setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_in_subst[OF p] by blast show "\q \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \\. fst p = fst q \ snd p = snd q \\<^sub>p \ \ ?P \" using q wt_subst_rm_vars[OF \(1)] wf_trms_subst_rm_vars'[OF \(2)] by blast qed subsection \Lemmata: Properties of the Constraint Translation Function\ lemma tr_par_labeled_rcv_iff: "B \ set (tr\<^sub>p\<^sub>c A D) \ (i, receive\t\\<^sub>s\<^sub>t) \ set B \ (i, receive\t\) \ set A" by (induct A D arbitrary: B rule: tr\<^sub>p\<^sub>c.induct) auto lemma tr_par_declassified_eq: "B \ set (tr\<^sub>p\<^sub>c A D) \ declassified\<^sub>l\<^sub>s\<^sub>t B I = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I" using tr_par_labeled_rcv_iff unfolding declassified\<^sub>l\<^sub>s\<^sub>t_def declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by simp lemma tr_par_ik_eq: assumes "B \ set (tr\<^sub>p\<^sub>c A D)" shows "ik\<^sub>s\<^sub>t (unlabel B) = ik\<^sub>s\<^sub>s\<^sub>t (unlabel A)" proof - have "{t. \i. (i, receive\t\\<^sub>s\<^sub>t) \ set B} = {t. \i. (i, receive\t\) \ set A}" using tr_par_labeled_rcv_iff[OF assms] by simp moreover have "\C. {t. \i. (i, receive\t\\<^sub>s\<^sub>t) \ set C} = {t. receive\t\\<^sub>s\<^sub>t \ set (unlabel C)}" "\C. {t. \i. (i, receive\t\) \ set C} = {t. receive\t\ \ set (unlabel C)}" unfolding unlabel_def by force+ ultimately show ?thesis by (metis ik\<^sub>s\<^sub>s\<^sub>t_def ik\<^sub>s\<^sub>t_is_rcv_set) qed lemma tr_par_deduct_iff: assumes "B \ set (tr\<^sub>p\<^sub>c A D)" shows "ik\<^sub>s\<^sub>t (unlabel B) \\<^sub>s\<^sub>e\<^sub>t I \ t \ ik\<^sub>s\<^sub>s\<^sub>t (unlabel A) \\<^sub>s\<^sub>e\<^sub>t I \ t" using tr_par_ik_eq[OF assms] by metis lemma tr_par_vars_subset: assumes "A' \ set (tr\<^sub>p\<^sub>c A D)" shows "fv\<^sub>l\<^sub>s\<^sub>t A' \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" (is ?P) and "bvars\<^sub>l\<^sub>s\<^sub>t A' \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A)" (is ?Q) proof - show ?P using assms proof (induction "unlabel A" arbitrary: A A' D rule: strand_sem_stateful_induct) case (ConsIn A' D ac t s AA A A') then obtain i B where iB: "A = (i,\ac: t \ s\)#B" "AA = unlabel B" unfolding unlabel_def by moura then obtain A'' d where *: "d \ set (dbproj i D)" "A' = (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c B D)" using ConsIn.prems(1) by moura hence "fv\<^sub>l\<^sub>s\<^sub>t A'' \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel B) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" "fv (pair (snd d)) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" apply (metis ConsIn.hyps(1)[OF iB(2)]) using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_mono[OF dbproj_subset[of i D]] fv_pair_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subset[OF *(1)] by blast thus ?case using * iB unfolding pair_def by auto next case (ConsDel A' D t s AA A A') then obtain i B where iB: "A = (i,delete\t,s\)#B" "AA = unlabel B" unfolding unlabel_def by moura define fltD1 where "fltD1 = (\Di. filter (\d. d \ set Di) D)" define fltD2 where "fltD2 = (\Di. filter (\d. d \ set Di) (dbproj i D))" define constr where "constr = (\Di. (map (\d. (i, \check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i, \[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) (fltD2 Di)))" from iB obtain A'' Di where *: "Di \ set (subseqs (dbproj i D))" "A' = (constr Di)@A''" "A'' \ set (tr\<^sub>p\<^sub>c B (fltD1 Di))" using ConsDel.prems(1) unfolding constr_def fltD1_def fltD2_def by moura hence "fv\<^sub>l\<^sub>s\<^sub>t A'' \ fv\<^sub>s\<^sub>s\<^sub>t AA \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel (fltD1 Di))" unfolding constr_def fltD1_def by (metis ConsDel.hyps(1) iB(2)) hence 1: "fv\<^sub>l\<^sub>s\<^sub>t A'' \ fv\<^sub>s\<^sub>s\<^sub>t AA \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_mono[of "unlabel (fltD1 Di)" "unlabel D"] unfolding unlabel_def fltD1_def by force - + have 2: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel Di) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel (fltD1 Di)) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" using subseqs_set_subset(1)[OF *(1)] unfolding fltD1_def unlabel_def by auto have 5: "fv\<^sub>l\<^sub>s\<^sub>t A' = fv\<^sub>l\<^sub>s\<^sub>t (constr Di) \ fv\<^sub>l\<^sub>s\<^sub>t A''" using * unfolding unlabel_def by force have "fv\<^sub>l\<^sub>s\<^sub>t (constr Di) \ fv t \ fv s \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel Di) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel (fltD1 Di))" unfolding unlabel_def constr_def fltD1_def fltD2_def pair_def by auto hence 3: "fv\<^sub>l\<^sub>s\<^sub>t (constr Di) \ fv t \ fv s \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" using 2 by blast have 4: "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) = fv t \ fv s \ fv\<^sub>s\<^sub>s\<^sub>t AA" using iB by auto - + have "fv\<^sub>s\<^sub>t (unlabel A') \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" using 1 3 4 5 by blast thus ?case by metis next case (ConsNegChecks A' D X F F' AA A A') then obtain i B where iB: "A = (i,NegChecks X F F')#B" "AA = unlabel B" unfolding unlabel_def by moura define D' where "D' \ \(fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ` set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (unlabel (dbproj i D))))" define constr where "constr = map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" from iB obtain A'' where *: "A'' \ set (tr\<^sub>p\<^sub>c B D)" "A' = constr@A''" using ConsNegChecks.prems(1) unfolding constr_def by moura hence "fv\<^sub>l\<^sub>s\<^sub>t A'' \ fv\<^sub>s\<^sub>s\<^sub>t AA \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" by (metis ConsNegChecks.hyps(1) iB(2)) hence **: "fv\<^sub>l\<^sub>s\<^sub>t A'' \ fv\<^sub>s\<^sub>s\<^sub>t AA \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" by auto have 1: "fv\<^sub>l\<^sub>s\<^sub>t constr \ (D' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) - set X" unfolding D'_def constr_def unlabel_def by auto have "set (unlabel (dbproj i D)) \ set (unlabel D)" unfolding unlabel_def by auto hence 2: "D' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" using tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_vars_subset'[of F' "unlabel (dbproj i D)"] fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_mono unfolding D'_def by blast have 3: "fv\<^sub>l\<^sub>s\<^sub>t A' \ ((fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) - set X) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ fv\<^sub>l\<^sub>s\<^sub>t A''" using 1 2 *(2) unfolding unlabel_def by fastforce have 4: "fv\<^sub>s\<^sub>s\<^sub>t AA \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel A)" by (metis ConsNegChecks.hyps(2) fv\<^sub>s\<^sub>s\<^sub>t_cons_subset) have 5: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel A)" using ConsNegChecks.hyps(2) unfolding unlabel_def by force show ?case using ** 3 4 5 by blast qed (fastforce simp add: unlabel_def)+ show ?Q using assms apply (induct "unlabel A" arbitrary: A A' D rule: strand_sem_stateful_induct) by (fastforce simp add: unlabel_def)+ qed lemma tr_par_vars_disj: assumes "A' \ set (tr\<^sub>p\<^sub>c A D)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" and "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" shows "fv\<^sub>l\<^sub>s\<^sub>t A' \ bvars\<^sub>l\<^sub>s\<^sub>t A' = {}" using assms tr_par_vars_subset by fast lemma tr_par_trms_subset: assumes "A' \ set (tr\<^sub>p\<^sub>c A D)" shows "trms\<^sub>l\<^sub>s\<^sub>t A' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D" using assms proof (induction A D arbitrary: A' rule: tr\<^sub>p\<^sub>c.induct) case 1 thus ?case by simp next case (2 i t A D) then obtain A'' where A'': "A' = (i,send\t\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" by moura hence "trms\<^sub>l\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D" by (metis "2.IH") thus ?case using A'' by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) next case (3 i t A D) then obtain A'' where A'': "A' = (i,receive\t\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" by moura hence "trms\<^sub>l\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D" by (metis "3.IH") thus ?case using A'' by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) next case (4 i ac t t' A D) then obtain A'' where A'': "A' = (i,\ac: t \ t'\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" by moura hence "trms\<^sub>l\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D" by (metis "4.IH") thus ?case using A'' by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) next case (5 i t s A D) hence "A' \ set (tr\<^sub>p\<^sub>c A (List.insert (i,t,s) D))" by simp hence "trms\<^sub>l\<^sub>s\<^sub>t A' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set (List.insert (i,t,s) D)" by (metis "5.IH") thus ?case by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) next case (6 i t s A D) from 6 obtain Di A'' B C where A'': "Di \ set (subseqs (dbproj i D))" "A'' \ set (tr\<^sub>p\<^sub>c A [d\D. d \ set Di])" "A' = (B@C)@A''" "B = map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di" "C = map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\dbproj i D. d \ set Di]" by moura hence "trms\<^sub>l\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set [d\D. d \ set Di]" by (metis "6.IH") moreover have "set [d\D. d \ set Di] \ set D" using set_filter by auto ultimately have "trms\<^sub>l\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D" by blast hence "trms\<^sub>l\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel ((i,delete\t,s\)#A)) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel ((i,delete\t,s\)#A)) \ pair ` snd ` set D" using setops\<^sub>s\<^sub>s\<^sub>t_cons_subset trms\<^sub>s\<^sub>s\<^sub>t_cons by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) moreover have "set Di \ set D" "set [d\dbproj i D . d \ set Di] \ set D" using subseqs_set_subset[OF A''(1)] by auto hence "trms\<^sub>s\<^sub>t (unlabel B) \ insert (pair (t, s)) (pair ` snd ` set D)" "trms\<^sub>s\<^sub>t (unlabel C) \ insert (pair (t, s)) (pair ` snd ` set D)" using A''(4,5) unfolding unlabel_def by auto hence "trms\<^sub>s\<^sub>t (unlabel (B@C)) \ insert (pair (t,s)) (pair ` snd ` set D)" using unlabel_append[of B C] by auto moreover have "pair (t,s) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (delete\t,s\#unlabel A)" by (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) ultimately show ?case using A''(3) trms\<^sub>s\<^sub>t_append[of "unlabel (B@C)" "unlabel A'"] unlabel_append[of "B@C" A''] by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) next case (7 i ac t s A D) from 7 obtain d A'' where A'': "d \ set (dbproj i D)" "A'' \ set (tr\<^sub>p\<^sub>c A D)" "A' = (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#A''" by moura hence "trms\<^sub>l\<^sub>s\<^sub>t A'' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D" by (metis "7.IH") moreover have "trms\<^sub>s\<^sub>t (unlabel A') = {pair (t,s), pair (snd d)} \ trms\<^sub>s\<^sub>t (unlabel A'')" using A''(1,3) by auto ultimately show ?case using A''(1) by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) next case (8 i X F F' A D) define constr where "constr = map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" define B where "B \ \(trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ` set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D))))" from 8 obtain A'' where A'': "A'' \ set (tr\<^sub>p\<^sub>c A D)" "A' = constr@A''" unfolding constr_def by moura have "trms\<^sub>s\<^sub>t (unlabel A'') \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair`snd`set D" by (metis A''(1) "8.IH") moreover have "trms\<^sub>s\<^sub>t (unlabel constr) \ B \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair ` snd ` set D" unfolding unlabel_def constr_def B_def by auto ultimately have "trms\<^sub>s\<^sub>t (unlabel A') \ B \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D" using A'' unlabel_append[of constr A''] by auto moreover have "set (dbproj i D) \ set D" by auto hence "B \ pair ` set F' \ pair ` snd ` set D" using tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_trms_subset'[of F' "map snd (dbproj i D)"] unfolding B_def by force moreover have "pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel ((i, \X\\\: F \\: F'\)#A)) = pair ` set F' \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" by auto ultimately show ?case by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) qed lemma tr_par_wf_trms: assumes "A' \ set (tr\<^sub>p\<^sub>c A [])" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t (unlabel A))" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>t A')" using tr_par_trms_subset[OF assms(1)] setops\<^sub>s\<^sub>s\<^sub>t_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s(2)[OF assms(2)] by auto lemma tr_par_wf': assumes "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" and "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X" and "wf'\<^sub>s\<^sub>s\<^sub>t X (unlabel A)" "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" and "A' \ set (tr\<^sub>p\<^sub>c A D)" shows "wf\<^sub>l\<^sub>s\<^sub>t X A'" proof - define P where "P = (\(D::('fun,'var,'lbl) labeleddbstatelist) (A::('fun,'var,'lbl) labeled_stateful_strand). (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}) \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {})" have "P D A" using assms(1,4) by (simp add: P_def) with assms(5,3,2) show ?thesis proof (induction A arbitrary: X A' D) case Nil thus ?case by simp next case (Cons a A) obtain i s where i: "a = (i,s)" by (metis surj_pair) note prems = Cons.prems note IH = Cons.IH show ?case proof (cases s) case (Receive t) note si = Receive i then obtain A'' where A'': "A' = (i,receive\t\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" "fv t \ X" using prems unlabel_Cons(1)[of i s A] by moura have *: "wf'\<^sub>s\<^sub>s\<^sub>t X (unlabel A)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X" "P D A" using prems si apply (force, force) using prems(4) si unfolding P_def by fastforce show ?thesis using IH[OF A''(2) *] A''(1,3) by simp next case (Send t) note si = Send i then obtain A'' where A'': "A' = (i,send\t\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" using prems by moura have *: "wf'\<^sub>s\<^sub>s\<^sub>t (X \ fv t) (unlabel A)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X \ fv t" "P D A" using prems si apply (force, force) using prems(4) si unfolding P_def by fastforce show ?thesis using IH[OF A''(2) *] A''(1) by simp next case (Equality ac t t') note si = Equality i then obtain A'' where A'': "A' = (i,\ac: t \ t'\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" "ac = Assign \ fv t' \ X" using prems unlabel_Cons(1)[of i s] by moura have *: "ac = Assign \ wf'\<^sub>s\<^sub>s\<^sub>t (X \ fv t) (unlabel A)" "ac = Check \ wf'\<^sub>s\<^sub>s\<^sub>t X (unlabel A)" "ac = Assign \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X \ fv t" "ac = Check \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X" "P D A" using prems si apply (force, force, force, force) using prems(4) si unfolding P_def by fastforce show ?thesis using IH[OF A''(2) *(1,3,5)] IH[OF A''(2) *(2,4,5)] A''(1,3) by (cases ac) simp_all next case (Insert t t') note si = Insert i hence A': "A' \ set (tr\<^sub>p\<^sub>c A (List.insert (i,t,t') D))" "fv t \ X" "fv t' \ X" using prems by auto have *: "wf'\<^sub>s\<^sub>s\<^sub>t X (unlabel A)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel (List.insert (i,t,t') D)) \ X" using prems si by (auto simp add: unlabel_def) have **: "P (List.insert (i,t,t') D) A" using prems(4) si unfolding P_def unlabel_def by fastforce show ?thesis using IH[OF A'(1) * **] A'(2,3) by simp next case (Delete t t') note si = Delete i define constr where "constr = (\Di. (map (\d. (i,\check: (pair (t,t')) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i,\[]\\\: [(pair (t,t'), pair (snd d))]\\<^sub>s\<^sub>t)) [d\dbproj i D. d \ set Di]))" from prems si obtain Di A'' where A'': "A' = constr Di@A''" "A'' \ set (tr\<^sub>p\<^sub>c A [d\D. d \ set Di])" "Di \ set (subseqs (dbproj i D))" unfolding constr_def by auto have *: "wf'\<^sub>s\<^sub>s\<^sub>t X (unlabel A)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel (filter (\d. d \ set Di) D)) \ X" using prems si apply simp using prems si by (fastforce simp add: unlabel_def) have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel (filter (\d. d \ set Di) D)) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D)" by (auto simp add: unlabel_def) hence **: "P [d\D. d \ set Di] A" using prems si unfolding P_def by fastforce have ***: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X" using prems si by auto show ?thesis using IH[OF A''(2) * **] A''(1) wf_pair_eqs_ineqs_map'[OF _ A''(3) ***] unfolding constr_def by simp next case (InSet ac t t') note si = InSet i then obtain d A'' where A'': "A' = (i,\ac: (pair (t,t')) \ (pair (snd d))\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" "d \ set D" using prems by moura have *: "ac = Assign \ wf'\<^sub>s\<^sub>s\<^sub>t (X \ fv t \ fv t') (unlabel A)" "ac = Check \ wf'\<^sub>s\<^sub>s\<^sub>t X (unlabel A)" "ac = Assign \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X \ fv t \ fv t'" "ac = Check \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X" "P D A" using prems si apply (force, force, force, force) using prems(4) si unfolding P_def by fastforce have **: "fv (pair (snd d)) \ X" using A''(3) prems(3) fv_pair_fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subset by fast have ***: "fv (pair (t,t')) = fv t \ fv t'" unfolding pair_def by auto show ?thesis using IH[OF A''(2) *(1,3,5)] IH[OF A''(2) *(2,4,5)] A''(1) ** *** by (cases ac) (simp_all add: Un_assoc) next case (NegChecks Y F F') note si = NegChecks i then obtain A'' where A'': "A' = (map (\G. (i,\Y\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D))))@A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" using prems by moura have *: "wf'\<^sub>s\<^sub>s\<^sub>t X (unlabel A)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ X" using prems si by auto - + have "bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel ((i,\Y\\\: F \\: F'\)#A))" "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel ((i,\Y\\\: F \\: F'\)#A))" by auto hence **: "P D A" using prems si unfolding P_def by blast - + show ?thesis using IH[OF A''(2) * **] A''(1) wf_pair_negchecks_map' by simp qed qed qed lemma tr_par_wf: assumes "A' \ set (tr\<^sub>p\<^sub>c A [])" and "wf\<^sub>s\<^sub>s\<^sub>t (unlabel A)" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A)" shows "wf\<^sub>l\<^sub>s\<^sub>t {} A'" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>t A')" and "fv\<^sub>l\<^sub>s\<^sub>t A' \ bvars\<^sub>l\<^sub>s\<^sub>t A' = {}" using tr_par_wf'[OF _ _ _ _ assms(1)] tr_par_wf_trms[OF assms(1,3)] tr_par_vars_disj[OF assms(1)] assms(2) by fastforce+ lemma tr_par_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p: assumes "A' \ set (tr\<^sub>p\<^sub>c A D)" "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel A)" and "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" (is "?P0 A D") and "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel D) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" (is "?P1 A D") and "\t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D. \t' \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` snd ` set D. (\\. Unifier \ t t') \ \ t = \ t'" (is "?P3 A D") shows "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A')" proof - have sublmm: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel A)" "?P0 A D" "?P1 A D" "?P3 A D" when p: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel (a#A))" "?P0 (a#A) D" "?P1 (a#A) D" "?P3 (a#A) D" for a A D proof - show "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel A)" using p(1) by (simp add: unlabel_def tfr\<^sub>s\<^sub>s\<^sub>t_def) show "?P0 A D" using p(2) fv\<^sub>s\<^sub>s\<^sub>t_cons_subset unfolding unlabel_def by fastforce show "?P1 A D" using p(3) bvars\<^sub>s\<^sub>s\<^sub>t_cons_subset unfolding unlabel_def by fastforce have "setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ setops\<^sub>s\<^sub>s\<^sub>t (unlabel (a#A))" using setops\<^sub>s\<^sub>s\<^sub>t_cons_subset unfolding unlabel_def by auto thus "?P3 A D" using p(4) by blast qed show ?thesis using assms proof (induction A D arbitrary: A' rule: tr\<^sub>p\<^sub>c.induct) case 1 thus ?case by simp next case (2 i t A D) note prems = "2.prems" note IH = "2.IH" from prems(1) obtain A'' where A'': "A' = (i,send\t\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" by moura have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A'')" using IH[OF A''(2)] prems(5) sublmm[OF prems(2,3,4,5)] by meson thus ?case using A''(1) by simp next case (3 i t A D) note prems = "3.prems" note IH = "3.IH" from prems(1) obtain A'' where A'': "A' = (i,receive\t\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" by moura have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A'')" using IH[OF A''(2)] prems(5) sublmm[OF prems(2,3,4,5)] by meson thus ?case using A''(1) by simp next case (4 i ac t t' A D) note prems = "4.prems" note IH = "4.IH" from prems(1) obtain A'' where A'': "A' = (i,\ac: t \ t'\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" by moura have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A'')" using IH[OF A''(2)] prems(5) sublmm[OF prems(2,3,4,5)] by meson thus ?case using A''(1) prems(2) by simp next case (5 i t s A D) note prems = "5.prems" note IH = "5.IH" from prems(1) have A': "A' \ set (tr\<^sub>p\<^sub>c A (List.insert (i,t,s) D))" by simp have 1: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel A)" using sublmm[OF prems(2,3,4,5)] by simp - + have "pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel ((i,insert\t,s\)#A)) \ pair`snd`set D = pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair`snd`set (List.insert (i,t,s) D)" by (auto simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) hence 3: "?P3 A (List.insert (i,t,s) D)" using prems(5) by metis moreover have "?P1 A (List.insert (i,t,s) D)" using prems(3,4) bvars\<^sub>s\<^sub>s\<^sub>t_cons_subset[of "unlabel A" "insert\t,s\"] unfolding unlabel_def by fastforce ultimately have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A')" using IH[OF A' sublmm(1,2)[OF prems(2,3,4,5)] _ 3] by metis thus ?case using A'(1) by auto next case (6 i t s A D) note prems = "6.prems" note IH = "6.IH" define constr where constr: "constr \ (\Di. (map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) (filter (\d. d \ set Di) (dbproj i D))))" from prems(1) obtain Di A'' where A'': "A' = constr Di@A''" "A'' \ set (tr\<^sub>p\<^sub>c A (filter (\d. d \ set Di) D))" "Di \ set (subseqs (dbproj i D))" unfolding constr by fastforce define Q1 where "Q1 \ (\(F::(('fun,'var) term \ ('fun,'var) term) list) X. \x \ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) - set X. \a. \ (Var x) = TAtom a)" define Q2 where "Q2 \ (\(F::(('fun,'var) term \ ('fun,'var) term) list) X. \f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ T = [] \ (\s \ set T. s \ Var ` set X))" have "pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair`snd`set [d\D. d \ set Di] \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel ((i,delete\t,s\)#A)) \ pair`snd`set D" using subseqs_set_subset[OF A''(3)] by (force simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) moreover have "\a\M. \b\M. P a b" when "M \ N" "\a\N. \b\N. P a b" for M N::"('fun, 'var) terms" and P using that by blast ultimately have *: "?P3 A (filter (\d. d \ set Di) D)" using prems(5) by presburger have **: "?P1 A (filter (\d. d \ set Di) D)" using prems(4) bvars\<^sub>s\<^sub>s\<^sub>t_cons_subset[of "unlabel A" "delete\t,s\"] unfolding unlabel_def by fastforce - + have 1: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A'')" using IH[OF A''(3,2) sublmm(1,2)[OF prems(2,3,4,5)] ** *] by metis have 2: "\ac: u \ u'\\<^sub>s\<^sub>t \ set (unlabel A'') \ (\d \ set Di. u = pair (t,s) \ u' = pair (snd d))" when "\ac: u \ u'\\<^sub>s\<^sub>t \ set (unlabel A')" for ac u u' using that A''(1) unfolding constr unlabel_def by force have 3: "\X\\\: u\\<^sub>s\<^sub>t \ set (unlabel A'') \ (\d \ set (filter (\d. d \ set Di) D). u = [(pair (t,s), pair (snd d))] \ Q2 u X)" when "\X\\\: u\\<^sub>s\<^sub>t \ set (unlabel A')" for X u using that A''(1) unfolding Q2_def constr unlabel_def by force have 4: "\d\set D. (\\. Unifier \ (pair (t,s)) (pair (snd d))) \ \ (pair (t,s)) = \ (pair (snd d))" using prems(5) by (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) { fix ac u u' assume a: "\ac: u \ u'\\<^sub>s\<^sub>t \ set (unlabel A')" "\\. Unifier \ u u'" hence "\ac: u \ u'\\<^sub>s\<^sub>t \ set (unlabel A'') \ (\d \ set Di. u = pair (t,s) \ u' = pair (snd d))" using 2 by metis moreover { assume "\ac: u \ u'\\<^sub>s\<^sub>t \ set (unlabel A'')" hence "tfr\<^sub>s\<^sub>t\<^sub>p (\ac: u \ u'\\<^sub>s\<^sub>t)" using 1 Ball_set_list_all[of "unlabel A''" tfr\<^sub>s\<^sub>t\<^sub>p] by fast } moreover { fix d assume "d \ set Di" "u = pair (t,s)" "u' = pair (snd d)" hence "\\. Unifier \ u u' \ \ u = \ u'" using 4 dbproj_subseq_subset A''(3) by fast hence "tfr\<^sub>s\<^sub>t\<^sub>p (\ac: u \ u'\\<^sub>s\<^sub>t)" using Ball_set_list_all[of "unlabel A''" tfr\<^sub>s\<^sub>t\<^sub>p] by simp hence "\ u = \ u'" using tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def[of "unlabel A''"] using a(2) unfolding unlabel_def by auto } ultimately have "\ u = \ u'" using tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def[of "unlabel A''"] a(2) unfolding unlabel_def by auto } moreover { fix u U assume "\U\\\: u\\<^sub>s\<^sub>t \ set (unlabel A')" hence "\U\\\: u\\<^sub>s\<^sub>t \ set (unlabel A'') \ (\d \ set (filter (\d. d \ set Di) D). u = [(pair (t,s), pair (snd d))] \ Q2 u U)" using 3 by metis hence "Q1 u U \ Q2 u U" using 1 4 subseqs_set_subset[OF A''(3)] tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def[of "unlabel A''"] unfolding Q1_def Q2_def by blast } ultimately show ?case using tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def[of "unlabel A'"] unfolding Q1_def Q2_def unlabel_def by blast next case (7 i ac t s A D) note prems = "7.prems" note IH = "7.IH" from prems(1) obtain d A'' where A'': "A' = (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" "d \ set (dbproj i D)" by moura have 1: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A'')" - using IH[OF A''(2) sublmm(1,2,3)[OF prems(2,3,4,5)] sublmm(4)[OF prems(2,3,4,5)]] + using IH[OF A''(2) sublmm(1,2,3)[OF prems(2,3,4,5)] sublmm(4)[OF prems(2,3,4,5)]] by metis - + have 2: "\ (pair (t,s)) = \ (pair (snd d))" when "\\. Unifier \ (pair (t,s)) (pair (snd d))" using that prems(2,5) A''(3) unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by (simp add: setops\<^sub>s\<^sub>s\<^sub>t_def) - + show ?case using A''(1) 1 2 by fastforce next case (8 i X F F' A D) note prems = "8.prems" note IH = "8.IH" define constr where "constr = map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" define Q1 where "Q1 \ (\(F::(('fun,'var) term \ ('fun,'var) term) list) X. \x \ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) - set X. \a. \ (Var x) = TAtom a)" define Q2 where "Q2 \ (\(M::('fun,'var) terms) X. \f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t M \ T = [] \ (\s \ set T. s \ Var ` set X))" have Q2_subset: "Q2 M' X" when "M' \ M" "Q2 M X" for X M M' using that unfolding Q2_def by auto have Q2_supset: "Q2 (M \ M') X" when "Q2 M X" "Q2 M' X" for X M M' using that unfolding Q2_def by auto from prems obtain A'' where A'': "A' = constr@A''" "A'' \ set (tr\<^sub>p\<^sub>c A D)" using constr_def by moura have 0: "constr = [(i,\X\\\: F\\<^sub>s\<^sub>t)]" when "F' = []" using that unfolding constr_def by simp have 1: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A'')" using IH[OF A''(2) sublmm(1,2,3)[OF prems(2,3,4,5)] sublmm(4)[OF prems(2,3,4,5)]] by metis have 2: "(F' = [] \ Q1 F X) \ Q2 (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair ` set F') X" using prems(2) unfolding Q1_def Q2_def by simp - + have 3: "F' = [] \ Q1 F X \ list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel constr)" using 0 2 tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def[of "unlabel constr"] unfolding Q1_def by auto { fix c assume "c \ set (unlabel constr)" hence "\G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D))). c = \X\\\: (F@G)\\<^sub>s\<^sub>t" unfolding constr_def unlabel_def by force } moreover { fix G assume G: "G \ set (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" and c: "\X\\\: (F@G)\\<^sub>s\<^sub>t \ set (unlabel constr)" and e: "Q2 (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair ` set F') X" have d_Q2: "Q2 (pair ` set (map snd D)) X" unfolding Q2_def proof (intro allI impI) fix f T assume "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (pair ` set (map snd D))" then obtain d where d: "d \ set (map snd D)" "Fun f T \ subterms (pair d)" by force hence "fv (pair d) \ set X = {}" using prems(4) unfolding pair_def by (force simp add: unlabel_def) thus "T = [] \ (\s \ set T. s \ Var ` set X)" by (metis fv_disj_Fun_subterm_param_cases d(2)) qed have "trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F@G) \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair ` set F' \ pair ` set (map snd D)" using tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_trms_subset[OF G] by force hence "Q2 (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F@G)) X" using Q2_subset[OF _ Q2_supset[OF e d_Q2]] by metis hence "tfr\<^sub>s\<^sub>t\<^sub>p (\X\\\: (F@G)\\<^sub>s\<^sub>t)" by (metis Q2_def tfr\<^sub>s\<^sub>t\<^sub>p.simps(2)) } ultimately have 4: "Q2 (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ pair ` set F') X \ list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel constr)" using Ball_set by blast have 5: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel constr)" using 2 3 4 by metis show ?case using 1 5 A''(1) by (simp add: unlabel_def) qed qed lemma tr_par_tfr: assumes "A' \ set (tr\<^sub>p\<^sub>c A [])" and "tfr\<^sub>s\<^sub>s\<^sub>t (unlabel A)" and "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" shows "tfr\<^sub>s\<^sub>t (unlabel A')" proof - have *: "trms\<^sub>l\<^sub>s\<^sub>t A' \ trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" using tr_par_trms_subset[OF assms(1)] by simp hence "SMP (trms\<^sub>l\<^sub>s\<^sub>t A') \ SMP (trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A))" using SMP_mono by simp moreover have "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A))" using assms(2) unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by fast ultimately have 1: "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>t A')" by (metis tfr_subset(2)[OF _ *]) have **: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel A)" using assms(2) unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def by fast have "pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ SMP (trms\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)) - Var`\" using setops\<^sub>s\<^sub>s\<^sub>t_are_pairs unfolding pair_def by auto hence "\ t = \ t'" when "\\. Unifier \ t t'" "t \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "t' \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A)" for t t' using that assms(2) unfolding tfr\<^sub>s\<^sub>s\<^sub>t_def tfr\<^sub>s\<^sub>e\<^sub>t_def by blast moreover have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (unlabel []) = {}" "pair ` snd ` set [] = {}" by auto ultimately have 2: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (unlabel A')" using tr_par_tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p[OF assms(1) ** assms(3)] by simp show ?thesis by (metis 1 2 tfr\<^sub>s\<^sub>t_def) qed lemma tr_par_proj: assumes "B \ set (tr\<^sub>p\<^sub>c A D)" shows "proj n B \ set (tr\<^sub>p\<^sub>c (proj n A) (proj n D))" using assms proof (induction A D arbitrary: B rule: tr\<^sub>p\<^sub>c.induct) case (5 i t s S D) note prems = "5.prems" note IH = "5.IH" have IH': "proj n B \ set (tr\<^sub>p\<^sub>c (proj n S) (proj n (List.insert (i,t,s) D)))" using prems IH by auto show ?case proof (cases "(i = ln n) \ (i = \)") case True thus ?thesis using IH' proj_list_insert(1,2)[of n "(t,s)" D] proj_list_Cons(1,2)[of n _ S] by auto next case False then obtain m where "i = ln m" "n \ m" by (cases i) simp_all thus ?thesis using IH' proj_list_insert(3)[of n _ "(t,s)" D] proj_list_Cons(3)[of n _ "insert\t,s\" S] by auto qed next case (6 i t s S D) note prems = "6.prems" note IH = "6.IH" define constr where "constr = (\Di D. (map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\dbproj i D. d \ set Di]))" obtain Di B' where B': "B = constr Di D@B'" "Di \ set (subseqs (dbproj i D))" "B' \ set (tr\<^sub>p\<^sub>c S [d\D. d \ set Di])" using prems constr_def by fastforce hence "proj n B' \ set (tr\<^sub>p\<^sub>c (proj n S) (proj n [d\D. d \ set Di]))" using IH by simp hence IH': "proj n B' \ set (tr\<^sub>p\<^sub>c (proj n S) [d\proj n D. d \ set Di])" by (metis proj_filter) show ?case proof (cases "(i = ln n) \ (i = \)") case True hence "proj n B = constr Di D@proj n B'" "Di \ set (subseqs (dbproj i (proj n D)))" using B'(1,2) proj_dbproj(1,2)[of n D] unfolding proj_def constr_def by auto moreover have "constr Di (proj n D) = constr Di D" using True proj_dbproj(1,2)[of n D] unfolding constr_def by presburger ultimately have "proj n B \ set (tr\<^sub>p\<^sub>c ((i, delete\t,s\)#proj n S) (proj n D))" using IH' unfolding constr_def by force thus ?thesis by (metis proj_list_Cons(1,2) True) next case False then obtain m where m: "i = ln m" "n \ m" by (cases i) simp_all hence *: "(ln n) \ i" by simp have "proj n B = proj n B'" using B'(1) False unfolding constr_def proj_def by auto moreover have "[d\proj n D. d \ set Di] = proj n D" using proj_subseq[OF _ m(2)[symmetric]] m(1) B'(2) by simp ultimately show ?thesis using m(1) IH' proj_list_Cons(3)[OF m(2), of _ S] by auto qed next case (7 i ac t s S D) note prems = "7.prems" note IH = "7.IH" define constr where "constr = ( \d::'lbl strand_label \ ('fun,'var) term \ ('fun,'var) term. (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t))" obtain d B' where B': "B = constr d#B'" "d \ set (dbproj i D)" "B' \ set (tr\<^sub>p\<^sub>c S D)" using prems constr_def by fastforce hence IH': "proj n B' \ set (tr\<^sub>p\<^sub>c (proj n S) (proj n D))" using IH by auto show ?case proof (cases "(i = ln n) \ (i = \)") case True hence "proj n B = constr d#proj n B'" "d \ set (dbproj i (proj n D))" using B' proj_list_Cons(1,2)[of n _ B'] unfolding constr_def by (force, metis proj_dbproj(1,2)) hence "proj n B \ set (tr\<^sub>p\<^sub>c ((i, InSet ac t s)#proj n S) (proj n D))" using IH' unfolding constr_def by auto thus ?thesis using proj_list_Cons(1,2)[of n _ S] True by metis next case False then obtain m where m: "i = ln m" "n \ m" by (cases i) simp_all hence "proj n B = proj n B'" using B'(1) proj_list_Cons(3) unfolding constr_def by auto thus ?thesis using IH' m proj_list_Cons(3)[OF m(2), of "InSet ac t s" S] unfolding constr_def by auto qed next case (8 i X F F' S D) note prems = "8.prems" note IH = "8.IH" define constr where "constr = (\D. map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D))))" - + obtain B' where B': "B = constr D@B'" "B' \ set (tr\<^sub>p\<^sub>c S D)" using prems constr_def by fastforce hence IH': "proj n B' \ set (tr\<^sub>p\<^sub>c (proj n S) (proj n D))" using IH by auto show ?case proof (cases "(i = ln n) \ (i = \)") case True hence "proj n B = constr (proj n D)@proj n B'" using B'(1,2) proj_dbproj(1,2)[of n D] unfolding proj_def constr_def by auto hence "proj n B \ set (tr\<^sub>p\<^sub>c ((i, NegChecks X F F')#proj n S) (proj n D))" using IH' unfolding constr_def by auto thus ?thesis using proj_list_Cons(1,2)[of n _ S] True by metis next case False then obtain m where m: "i = ln m" "n \ m" by (cases i) simp_all hence "proj n B = proj n B'" using B'(1) unfolding constr_def proj_def by auto thus ?thesis using IH' m proj_list_Cons(3)[OF m(2), of "NegChecks X F F'" S] unfolding constr_def by auto qed qed (force simp add: proj_def)+ lemma tr_par_preserves_typing_cond: assumes "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A Sec" "typing_cond\<^sub>s\<^sub>s\<^sub>t (unlabel A)" "A' \ set (tr\<^sub>p\<^sub>c A [])" shows "typing_cond (unlabel A')" proof - have "wf'\<^sub>s\<^sub>s\<^sub>t {} (unlabel A)" "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>s\<^sub>t (unlabel A))" using assms(2) unfolding typing_cond\<^sub>s\<^sub>s\<^sub>t_def by simp_all hence 1: "wf\<^sub>s\<^sub>t {} (unlabel A')" "fv\<^sub>s\<^sub>t (unlabel A') \ bvars\<^sub>s\<^sub>t (unlabel A') = {}" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t (unlabel A'))" "Ana_invar_subst (ik\<^sub>s\<^sub>t (unlabel A') \ assignment_rhs\<^sub>s\<^sub>t (unlabel A'))" using tr_par_wf[OF assms(3)] Ana_invar_subst' by metis+ have 2: "tfr\<^sub>s\<^sub>t (unlabel A')" by (metis tr_par_tfr assms(2,3) typing_cond\<^sub>s\<^sub>s\<^sub>t_def) show ?thesis by (metis 1 2 typing_cond_def) qed lemma tr_par_preserves_par_comp: assumes "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A Sec" "A' \ set (tr\<^sub>p\<^sub>c A [])" shows "par_comp A' Sec" proof - let ?M = "\l. trms\<^sub>s\<^sub>s\<^sub>t (proj_unl l A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l A)" let ?N = "\l. trms_proj\<^sub>l\<^sub>s\<^sub>t l A'" have 0: "\l1 l2. l1 \ l2 \ GSMP_disjoint (?M l1) (?M l2) Sec" using assms(1) unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by simp_all { fix l1 l2::'lbl assume *: "l1 \ l2" hence "GSMP_disjoint (?M l1) (?M l2) Sec" using 0(1) by metis moreover have "pair ` snd ` set (proj n []) = {}" for n::'lbl unfolding proj_def by simp hence "?N l1 \ ?M l1" "?N l2 \ ?M l2" using tr_par_trms_subset[OF tr_par_proj[OF assms(2)]] by (metis Un_empty_right)+ ultimately have "GSMP_disjoint (?N l1) (?N l2) Sec" using GSMP_disjoint_subset by presburger } hence 1: "\l1 l2. l1 \ l2 \ GSMP_disjoint (trms_proj\<^sub>l\<^sub>s\<^sub>t l1 A') (trms_proj\<^sub>l\<^sub>s\<^sub>t l2 A') Sec" using 0(1) by metis have 2: "ground Sec" "\s \ Sec. \s' \ subterms s. {} \\<^sub>c s' \ s' \ Sec" using assms(1) unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by metis+ show ?thesis using 1 2 unfolding par_comp_def by metis qed lemma tr_leaking_prefix_exists: assumes "A' \ set (tr\<^sub>p\<^sub>c A [])" "prefix B A'" "ik\<^sub>s\<^sub>t (proj_unl n B) \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" shows "\C D. prefix C B \ prefix D A \ C \ set (tr\<^sub>p\<^sub>c D []) \ (ik\<^sub>s\<^sub>t (proj_unl n C) \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \)" proof - let ?P = "\B C C'. B = C@C' \ (\n t. (n, receive\t\\<^sub>s\<^sub>t) \ set C') \ (C = [] \ (\n t. suffix [(n,receive\t\\<^sub>s\<^sub>t)] C))" have "\C C'. ?P B C C'" proof (induction B) case (Cons b B) then obtain C C' n s where *: "?P B C C'" "b = (n,s)" by moura show ?case proof (cases "C = []") case True note T = True show ?thesis proof (cases "\t. s = receive\t\\<^sub>s\<^sub>t") case True hence "?P (b#B) [b] C'" using * T by auto thus ?thesis by metis next case False hence "?P (b#B) [] (b#C')" using * T by auto thus ?thesis by metis qed next case False hence "?P (b#B) (b#C) C'" using * unfolding suffix_def by auto thus ?thesis by metis qed qed simp then obtain C C' where C: "B = C@C'" "\n t. (n, receive\t\\<^sub>s\<^sub>t) \ set C'" "C = [] \ (\n t. suffix [(n,receive\t\\<^sub>s\<^sub>t)] C)" by moura hence 1: "prefix C B" by simp hence 2: "prefix C A'" using assms(2) by simp have "\m t. (m,receive\t\\<^sub>s\<^sub>t) \ set B \ (m,receive\t\\<^sub>s\<^sub>t) \ set C" using C by auto hence "\t. receive\t\\<^sub>s\<^sub>t \ set (proj_unl n B) \ receive\t\\<^sub>s\<^sub>t \ set (proj_unl n C)" unfolding unlabel_def proj_def by force hence "ik\<^sub>s\<^sub>t (proj_unl n B) \ ik\<^sub>s\<^sub>t (proj_unl n C)" using ik\<^sub>s\<^sub>t_is_rcv_set by auto hence 3: "ik\<^sub>s\<^sub>t (proj_unl n C) \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" by (metis ideduct_mono[OF assms(3)] subst_all_mono) { fix D E m t assume "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E" "prefix E A'" "A' \ set (tr\<^sub>p\<^sub>c A D)" hence "\F. prefix F A \ E \ set (tr\<^sub>p\<^sub>c F D)" proof (induction A D arbitrary: A' E rule: tr\<^sub>p\<^sub>c.induct) case (1 D) thus ?case by simp next case (2 i t' S D) note prems = "2.prems" note IH = "2.IH" obtain A'' where *: "A' = (i,send\t'\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c S D)" using prems(3) by auto have "E \ []" using prems(1) by auto then obtain E' where **: "E = (i,send\t'\\<^sub>s\<^sub>t)#E'" using *(1) prems(2) by (cases E) auto hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'" "prefix E' A''" using *(1) prems(1,2) suffix_Cons[of _ _ E'] by auto then obtain F where "prefix F S" "E' \ set (tr\<^sub>p\<^sub>c F D)" using *(2) ** IH by metis hence "prefix ((i,Send t')#F) ((i,Send t')#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,Send t')#F) D)" using ** by auto thus ?case by metis next case (3 i t' S D) note prems = "3.prems" note IH = "3.IH" obtain A'' where *: "A' = (i,receive\t'\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c S D)" using prems(3) by auto have "E \ []" using prems(1) by auto then obtain E' where **: "E = (i,receive\t'\\<^sub>s\<^sub>t)#E'" using *(1) prems(2) by (cases E) auto show ?case proof (cases "(m, receive\t\\<^sub>s\<^sub>t) = (i, receive\t'\\<^sub>s\<^sub>t)") case True note T = True show ?thesis proof (cases "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'") case True hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'" "prefix E' A''" using ** *(1) prems(1,2) by auto then obtain F where "prefix F S" "E' \ set (tr\<^sub>p\<^sub>c F D)" using *(2) ** IH by metis hence "prefix ((i,receive\t'\)#F) ((i,receive\t'\)#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,receive\t'\)#F) D)" using ** by auto thus ?thesis by metis next case False hence "E' = []" using **(1) T prems(1) suffix_Cons[of "[(m, receive\t\\<^sub>s\<^sub>t)]" "(m, receive\t\\<^sub>s\<^sub>t)" E'] by auto hence "prefix [(i,receive\t'\)] ((i,receive\t'\) # S) \ E \ set (tr\<^sub>p\<^sub>c [(i,receive\t'\)] D)" using * ** prems by auto thus ?thesis by metis qed next case False hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'" "prefix E' A''" using ** *(1) prems(1,2) suffix_Cons[of _ _ E'] by auto then obtain F where "prefix F S" "E' \ set (tr\<^sub>p\<^sub>c F D)" using *(2) ** IH by metis hence "prefix ((i,receive\t'\)#F) ((i,receive\t'\)#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,receive\t'\)#F) D)" using ** by auto thus ?thesis by metis qed next case (4 i ac t' t'' S D) note prems = "4.prems" note IH = "4.IH" obtain A'' where *: "A' = (i,\ac: t' \ t''\\<^sub>s\<^sub>t)#A''" "A'' \ set (tr\<^sub>p\<^sub>c S D)" using prems(3) by auto have "E \ []" using prems(1) by auto then obtain E' where **: "E = (i,\ac: t' \ t''\\<^sub>s\<^sub>t)#E'" using *(1) prems(2) by (cases E) auto hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'" "prefix E' A''" using *(1) prems(1,2) suffix_Cons[of _ _ E'] by auto then obtain F where "prefix F S" "E' \ set (tr\<^sub>p\<^sub>c F D)" using *(2) ** IH by metis hence "prefix ((i,Equality ac t' t'')#F) ((i,Equality ac t' t'')#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,Equality ac t' t'')#F) D)" using ** by auto thus ?case by metis next case (5 i t' s S D) note prems = "5.prems" note IH = "5.IH" have *: "A' \ set (tr\<^sub>p\<^sub>c S (List.insert (i,t',s) D))" using prems(3) by auto have "E \ []" using prems(1) by auto hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E" "prefix E A'" using *(1) prems(1,2) suffix_Cons[of _ _ E] by auto then obtain F where "prefix F S" "E \ set (tr\<^sub>p\<^sub>c F (List.insert (i,t',s) D))" using * IH by metis hence "prefix ((i,insert\t',s\)#F) ((i,insert\t',s\)#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,insert\t',s\)#F) D)" by auto thus ?case by metis next case (6 i t' s S D) note prems = "6.prems" note IH = "6.IH" define constr where "constr = (\Di. (map (\d. (i,\check: (pair (t',s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i,\[]\\\: [(pair (t',s), pair (snd d))]\\<^sub>s\<^sub>t)) (filter (\d. d \ set Di) (dbproj i D))))" obtain A'' Di where *: "A' = constr Di@A''" "A'' \ set (tr\<^sub>p\<^sub>c S (filter (\d. d \ set Di) D))" "Di \ set (subseqs (dbproj i D))" using prems(3) constr_def by auto have ***: "(m, receive\t\\<^sub>s\<^sub>t) \ set (constr Di)" using constr_def by auto have "E \ []" using prems(1) by auto then obtain E' where **: "E = constr Di@E'" using *(1) prems(1,2) *** by (metis (mono_tags, lifting) Un_iff list.set_intros(1) prefixI prefix_def prefix_same_cases set_append suffix_def) hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'" "prefix E' A''" using *(1) prems(1,2) suffix_append[of "[(m,receive\t\\<^sub>s\<^sub>t)]" "constr Di" E'] *** by (metis (no_types, hide_lams) Nil_suffix append_Nil2 in_set_conv_decomp rev_exhaust snoc_suffix_snoc suffix_appendD, auto) then obtain F where "prefix F S" "E' \ set (tr\<^sub>p\<^sub>c F (filter (\d. d \ set Di) D))" using *(2,3) ** IH by metis hence "prefix ((i,delete\t',s\)#F) ((i,delete\t',s\)#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,delete\t',s\)#F) D)" using *(3) ** constr_def by auto thus ?case by metis next case (7 i ac t' s S D) note prems = "7.prems" note IH = "7.IH" define constr where "constr = ( \d::(('lbl strand_label \ ('fun,'var) term \ ('fun,'var) term)). (i,\ac: (pair (t',s)) \ (pair (snd d))\\<^sub>s\<^sub>t))" - + obtain A'' d where *: "A' = constr d#A''" "A'' \ set (tr\<^sub>p\<^sub>c S D)" "d \ set (dbproj i D)" using prems(3) constr_def by auto have "E \ []" using prems(1) by auto then obtain E' where **: "E = constr d#E'" using *(1) prems(2) by (cases E) auto hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'" "prefix E' A''" using *(1) prems(1,2) suffix_Cons[of _ _ E'] using constr_def by auto then obtain F where "prefix F S" "E' \ set (tr\<^sub>p\<^sub>c F D)" using *(2) ** IH by metis hence "prefix ((i,InSet ac t' s)#F) ((i,InSet ac t' s)#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,InSet ac t' s)#F) D)" using *(3) ** unfolding constr_def by auto thus ?case by metis next case (8 i X G G' S D) note prems = "8.prems" note IH = "8.IH" define constr where "constr = map (\H. (i,\X\\\: (G@H)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G' (map snd (dbproj i D)))" - + obtain A'' where *: "A' = constr@A''" "A'' \ set (tr\<^sub>p\<^sub>c S D)" using prems(3) constr_def by auto have ***: "(m, receive\t\\<^sub>s\<^sub>t) \ set constr" using constr_def by auto have "E \ []" using prems(1) by auto then obtain E' where **: "E = constr@E'" using *(1) prems(1,2) *** by (metis (mono_tags, lifting) Un_iff list.set_intros(1) prefixI prefix_def prefix_same_cases set_append suffix_def) hence "suffix [(m, receive\t\\<^sub>s\<^sub>t)] E'" "prefix E' A''" using *(1) prems(1,2) suffix_append[of "[(m,receive\t\\<^sub>s\<^sub>t)]" constr E'] *** by (metis (no_types, hide_lams) Nil_suffix append_Nil2 in_set_conv_decomp rev_exhaust snoc_suffix_snoc suffix_appendD, auto) then obtain F where "prefix F S" "E' \ set (tr\<^sub>p\<^sub>c F D)" using *(2) ** IH by metis hence "prefix ((i,NegChecks X G G')#F) ((i,NegChecks X G G')#S)" "E \ set (tr\<^sub>p\<^sub>c ((i,NegChecks X G G')#F) D)" using ** constr_def by auto thus ?case by metis qed } moreover have "prefix [] A" "[] \ set (tr\<^sub>p\<^sub>c [] [])" by auto ultimately have 4: "\D. prefix D A \ C \ set (tr\<^sub>p\<^sub>c D [])" using C(3) assms(1) 2 by blast show ?thesis by (metis 1 3 4) qed subsection \Theorem: Semantic Equivalence of Translation\ context begin text \ An alternative version of the translation that does not perform database-state projections. It is used as an intermediate step in the proof of semantic equivalence. \ private fun tr'\<^sub>p\<^sub>c:: "('fun,'var,'lbl) labeled_stateful_strand \ ('fun,'var,'lbl) labeleddbstatelist \ ('fun,'var,'lbl) labeled_strand list" where "tr'\<^sub>p\<^sub>c [] D = [[]]" | "tr'\<^sub>p\<^sub>c ((i,send\t\)#A) D = map ((#) (i,send\t\\<^sub>s\<^sub>t)) (tr'\<^sub>p\<^sub>c A D)" | "tr'\<^sub>p\<^sub>c ((i,receive\t\)#A) D = map ((#) (i,receive\t\\<^sub>s\<^sub>t)) (tr'\<^sub>p\<^sub>c A D)" | "tr'\<^sub>p\<^sub>c ((i,\ac: t \ t'\)#A) D = map ((#) (i,\ac: t \ t'\\<^sub>s\<^sub>t)) (tr'\<^sub>p\<^sub>c A D)" | "tr'\<^sub>p\<^sub>c ((i,insert\t,s\)#A) D = tr'\<^sub>p\<^sub>c A (List.insert (i,(t,s)) D)" | "tr'\<^sub>p\<^sub>c ((i,delete\t,s\)#A) D = ( concat (map (\Di. map (\B. (map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di)@ (map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\D. d \ set Di])@B) (tr'\<^sub>p\<^sub>c A [d\D. d \ set Di])) (subseqs D)))" | "tr'\<^sub>p\<^sub>c ((i,\ac: t \ s\)#A) D = concat (map (\B. map (\d. (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#B) D) (tr'\<^sub>p\<^sub>c A D))" | "tr'\<^sub>p\<^sub>c ((i,\X\\\: F \\: F'\)#A) D = map ((@) (map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd D)))) (tr'\<^sub>p\<^sub>c A D)" subsubsection \Part 1\ private lemma tr'_par_iff_unlabel_tr: assumes "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. p = q \ i = j" shows "(\C \ set (tr'\<^sub>p\<^sub>c A D). B = unlabel C) \ B \ set (tr (unlabel A) (unlabel D))" (is "?A \ ?B") proof { fix C have "C \ set (tr'\<^sub>p\<^sub>c A D) \ unlabel C \ set (tr (unlabel A) (unlabel D))" using assms proof (induction A D arbitrary: C rule: tr'\<^sub>p\<^sub>c.induct) case (5 i t s S D) hence "unlabel C \ set (tr (unlabel S) (unlabel (List.insert (i, t, s) D)))" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) moreover have "insert (i,t,s) (set D) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i,insert\t,s\)#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "\(j,p) \ insert (i,t,s) (set D). \(k,q) \ insert (i,t,s) (set D). p = q \ j = k" using "5.prems"(2) by blast hence "unlabel (List.insert (i, t, s) D) = (List.insert (t, s) (unlabel D))" using map_snd_list_insert_distrib[of "(i,t,s)" D] unfolding unlabel_def by simp ultimately show ?case by auto next case (6 i t s S D) let ?f1 = "\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t" let ?g1 = "\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t" let ?f2 = "\d. (i, ?f1 (snd d))" let ?g2 = "\d. (i, ?g1 (snd d))" - + define constr1 where "constr1 = (\Di. (map ?f1 Di)@(map ?g1 [d\unlabel D. d \ set Di]))" define constr2 where "constr2 = (\Di. (map ?f2 Di)@(map ?g2 [d\D. d \ set Di]))" - + obtain C' Di where C': "Di \ set (subseqs D)" "C = constr2 Di@C'" "C' \ set (tr'\<^sub>p\<^sub>c S [d\D. d \ set Di])" using "6.prems"(1) unfolding constr2_def by moura - + have 0: "set [d\D. d \ set Di] \ set D" "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S)" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence 1: "\(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set [d\D. d \ set Di]. \(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set [d\D. d \ set Di]. p = q \ j = k" using "6.prems"(2) by blast - + have "\(i,p) \ set D \ set Di. \(j,q) \ set D \ set Di. p = q \ i = j" using "6.prems"(2) subseqs_set_subset(1)[OF C'(1)] by blast hence 2: "unlabel [d\D. d \ set Di] = [d\unlabel D. d \ set (unlabel Di)]" using unlabel_filter_eq[of D "set Di"] unfolding unlabel_def by simp - + have 3: "\f g::('a \ 'a \ 'c). \A B::(('b \ 'a \ 'a) list). map snd ((map (\d. (i, f (snd d))) A)@(map (\d. (i, g (snd d))) B)) = map f (map snd A)@map g (map snd B)" by simp have "unlabel (constr2 Di) = constr1 (unlabel Di)" using 2 3[of ?f1 Di ?g1 "[d\D. d \ set Di]"] by (simp add: constr1_def constr2_def unlabel_def) hence 4: "unlabel C = constr1 (unlabel Di)@unlabel C'" using C'(2) unlabel_append by metis - + have "unlabel Di \ set (map unlabel (subseqs D))" using C'(1) unfolding unlabel_def by simp hence 5: "unlabel Di \ set (subseqs (unlabel D))" using map_subseqs[of snd D] unfolding unlabel_def by simp - + show ?case using "6.IH"[OF C'(1,3) 1] 2 4 5 unfolding constr1_def by auto next case (7 i ac t s S D) obtain C' d where C': "C = (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#C'" "C' \ set (tr'\<^sub>p\<^sub>c S D)" "d \ set D" using "7.prems"(1) by moura - + have "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i,InSet ac t s)#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "\(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D. \(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D. p = q \ j = k" using "7.prems"(2) by blast hence "unlabel C' \ set (tr (unlabel S) (unlabel D))" using "7.IH"[OF C'(2)] by auto thus ?case using C' unfolding unlabel_def by force next case (8 i X F F' S D) obtain C' where C': "C = map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd D))@C'" "C' \ set (tr'\<^sub>p\<^sub>c S D)" using "8.prems"(1) by moura - + have "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i,NegChecks X F F')#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "\(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D. \(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D. p = q \ j = k" using "8.prems"(2) by blast hence "unlabel C' \ set (tr (unlabel S) (unlabel D))" using "8.IH"[OF C'(2)] by auto thus ?case using C' unfolding unlabel_def by auto qed (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) } thus "?A \ ?B" by blast show "?B \ ?A" using assms proof (induction A arbitrary: B D) case (Cons a A) obtain ia sa where a: "a = (ia,sa)" by moura have "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (a#A)" using a by (cases sa) (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence 1: "\(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. \(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. p = q \ j = k" using Cons.prems(2) by blast show ?case proof (cases sa) case (Send t) then obtain B' where B': "B = send\t\\<^sub>s\<^sub>t#B'" "B' \ set (tr (unlabel A) (unlabel D))" using Cons.prems(1) a by auto thus ?thesis using Cons.IH[OF B'(2) 1] a B'(1) Send by auto next case (Receive t) then obtain B' where B': "B = receive\t\\<^sub>s\<^sub>t#B'" "B' \ set (tr (unlabel A) (unlabel D))" using Cons.prems(1) a by auto thus ?thesis using Cons.IH[OF B'(2) 1] a B'(1) Receive by auto next case (Equality ac t t') then obtain B' where B': "B = \ac: t \ t'\\<^sub>s\<^sub>t#B'" "B' \ set (tr (unlabel A) (unlabel D))" using Cons.prems(1) a by auto thus ?thesis using Cons.IH[OF B'(2) 1] a B'(1) Equality by auto next case (Insert t s) hence B: "B \ set (tr (unlabel A) (List.insert (t,s) (unlabel D)))" using Cons.prems(1) a by auto let ?P = "\i. List.insert (t,s) (unlabel D) = unlabel (List.insert (i,t,s) D)" { obtain j where j: "?P j" "j = ia \ (j,t,s) \ set D" using labeled_list_insert_eq_ex_cases[of "(t,s)" D ia] by moura hence "j = ia" using Cons.prems(2) a Insert by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "?P ia" using j(1) by metis } hence j: "?P ia" by metis - + have 2: "\(k1, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set (List.insert (ia,t,s) D). \(k2, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set (List.insert (ia,t,s) D). p = q \ k1 = k2" using Cons.prems(2) a Insert by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) show ?thesis using Cons.IH[OF _ 2] j(1) B Insert a by auto next case (Delete t s) define c where "c \ (\(i::'lbl strand_label) Di. map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di@ map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\D. d \ set Di])" - + define d where "d \ (\Di. map (\d. \check: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t) Di@ map (\d. \[]\\\: [(pair (t,s), pair d)]\\<^sub>s\<^sub>t) [d\unlabel D. d \ set Di])" obtain B' Di where B': "B = d Di@B'" "Di \ set (subseqs (unlabel D))" "B' \ set (tr (unlabel A) [d\unlabel D. d \ set Di])" using Cons.prems(1) a Delete unfolding d_def by auto obtain Di' where Di': "Di' \ set (subseqs D)" "unlabel Di' = Di" using unlabel_subseqsD[OF B'(2)] by moura have 2: "\(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set [d\D. d \ set Di']. \(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set [d\D. d \ set Di']. p = q \ j = k" using 1 subseqs_subset[OF Di'(1)] filter_is_subset[of "\d. d \ set Di'"] by blast have "set Di' \ set D" by (rule subseqs_subset[OF Di'(1)]) hence "\(j, p)\set D \ set Di'. \(k, q)\set D \ set Di'. p = q \ j = k" using Cons.prems(2) by blast hence 3: "[d\unlabel D. d \ set Di] = unlabel [d\D. d \ set Di']" using Di'(2) unlabel_filter_eq[of D "set Di'"] unfolding unlabel_def by auto - + obtain C where C: "C \ set (tr'\<^sub>p\<^sub>c A [d\D. d \ set Di'])" "B' = unlabel C" using 3 Cons.IH[OF _ 2] B'(3) by auto hence 4: "c ia Di'@C \ set (tr'\<^sub>p\<^sub>c (a#A) D)" using Di'(1) a Delete unfolding c_def by auto have "unlabel (c ia Di') = d Di" using Di' 3 unfolding c_def d_def unlabel_def by auto hence 5: "B = unlabel (c ia Di'@C)" using B'(1) C(2) unlabel_append[of "c ia Di'" C] by simp - + show ?thesis using 4 5 by blast next case (InSet ac t s) then obtain B' d where B': "B = \ac: (pair (t,s)) \ (pair d)\\<^sub>s\<^sub>t#B'" "B' \ set (tr (unlabel A) (unlabel D))" "d \ set (unlabel D)" using Cons.prems(1) a by auto thus ?thesis using Cons.IH[OF _ 1] a InSet unfolding unlabel_def by auto next case (NegChecks X F F') - then obtain B' where B': + then obtain B' where B': "B = map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (unlabel D))@B'" "B' \ set (tr (unlabel A) (unlabel D))" using Cons.prems(1) a by auto thus ?thesis using Cons.IH[OF _ 1] a NegChecks unfolding unlabel_def by auto qed qed simp qed subsubsection \Part 2\ private lemma tr_par_iff_tr'_par: assumes "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. (\\. Unifier \ (pair p) (pair q)) \ i = j" (is "?R3 A D") and "\(l,t,s) \ set D. (fv t \ fv s) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" (is "?R4 A D") and "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" (is "?R5 A D") shows "(\B \ set (tr\<^sub>p\<^sub>c A D). \M; unlabel B\\<^sub>d \) \ (\C \ set (tr'\<^sub>p\<^sub>c A D). \M; unlabel C\\<^sub>d \)" (is "?P \ ?Q") proof { fix B assume "B \ set (tr\<^sub>p\<^sub>c A D)" "\M; unlabel B\\<^sub>d \" hence ?Q using assms proof (induction A D arbitrary: B M rule: tr\<^sub>p\<^sub>c.induct) case (1 D) thus ?case by simp next case (2 i t S D) note prems = "2.prems" note IH = "2.IH" obtain B' where B': "B = (i,send\t\\<^sub>s\<^sub>t)#B'" "B' \ set (tr\<^sub>p\<^sub>c S D)" using prems(1) by moura - + have 1: "\M; unlabel B'\\<^sub>d \" using prems(2) B'(1) by simp have 4: "?R3 S D" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force have 7: "M \ t \ \" using prems(2) B'(1) by simp obtain C where C: "C \ set (tr'\<^sub>p\<^sub>c S D)" "\M; unlabel C\\<^sub>d \" using IH[OF B'(2) 1 4 5 6] by moura hence "((i,send\t\\<^sub>s\<^sub>t)#C) \ set (tr'\<^sub>p\<^sub>c ((i,Send t)#S) D)" "\M; unlabel ((i,send\t\\<^sub>s\<^sub>t)#C)\\<^sub>d \" using 7 by auto thus ?case by metis next case (3 i t S D) note prems = "3.prems" note IH = "3.IH" obtain B' where B': "B = (i,receive\t\\<^sub>s\<^sub>t)#B'" "B' \ set (tr\<^sub>p\<^sub>c S D)" using prems(1) by moura - + have 1: "\insert (t \ \) M; unlabel B'\\<^sub>d \ " using prems(2) B'(1) by simp have 4: "?R3 S D" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force obtain C where C: "C \ set (tr'\<^sub>p\<^sub>c S D)" "\insert (t \ \) M; unlabel C\\<^sub>d \" using IH[OF B'(2) 1 4 5 6] by moura hence "((i,receive\t\\<^sub>s\<^sub>t)#C) \ set (tr'\<^sub>p\<^sub>c ((i,receive\t\)#S) D)" "\insert (t \ \) M; unlabel ((i,receive\t\\<^sub>s\<^sub>t)#C)\\<^sub>d \" by auto thus ?case by auto next case (4 i ac t t' S D) note prems = "4.prems" note IH = "4.IH" obtain B' where B': "B = (i,\ac: t \ t'\\<^sub>s\<^sub>t)#B'" "B' \ set (tr\<^sub>p\<^sub>c S D)" using prems(1) by moura - + have 1: "\M; unlabel B'\\<^sub>d \ " using prems(2) B'(1) by simp have 4: "?R3 S D" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force have 7: "t \ \ = t' \ \" using prems(2) B'(1) by simp obtain C where C: "C \ set (tr'\<^sub>p\<^sub>c S D)" "\M; unlabel C\\<^sub>d \" using IH[OF B'(2) 1 4 5 6] by moura hence "((i,\ac: t \ t'\\<^sub>s\<^sub>t)#C) \ set (tr'\<^sub>p\<^sub>c ((i,Equality ac t t')#S) D)" "\M; unlabel ((i,\ac: t \ t'\\<^sub>s\<^sub>t)#C)\\<^sub>d \" using 7 by auto thus ?case by metis next case (5 i t s S D) note prems = "5.prems" note IH = "5.IH" have B: "B \ set (tr\<^sub>p\<^sub>c S (List.insert (i,t,s) D))" using prems(1) by simp - + have 1: "\M; unlabel B\\<^sub>d \ " using prems(2) B(1) by simp have 4: "?R3 S (List.insert (i,t,s) D)" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S (List.insert (i,t,s) D)" using prems(4,5) by force have 6: "?R5 S D" using prems(5) by force show ?case using IH[OF B(1) 1 4 5 6] by simp next case (6 i t s S D) note prems = "6.prems" note IH = "6.IH" let ?cl1 = "\Di. map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di" let ?cu1 = "\Di. map (\d. \check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t) Di" let ?cl2 = "\Di. map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\dbproj i D. d\set Di]" let ?cu2 = "\Di. map (\d. \[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\dbproj i D. d\set Di]" let ?dl1 = "\Di. map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di" let ?du1 = "\Di. map (\d. \check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t) Di" let ?dl2 = "\Di. map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\D. d\set Di]" let ?du2 = "\Di. map (\d. \[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\D. d\set Di]" define c where c: "c = (\Di. ?cl1 Di@?cl2 Di)" define d where d: "d = (\Di. ?dl1 Di@?dl2 Di)" obtain B' Di where B': "Di \ set (subseqs (dbproj i D))" "B = c Di@B'" "B' \ set (tr\<^sub>p\<^sub>c S [d\D. d \ set Di])" using prems(1) c by moura have 0: "ik\<^sub>s\<^sub>t (unlabel (c Di)) = {}" "ik\<^sub>s\<^sub>t (unlabel (d Di)) = {}" "unlabel (?cl1 Di) = ?cu1 Di" "unlabel (?cl2 Di) = ?cu2 Di" "unlabel (?dl1 Di) = ?du1 Di" "unlabel (?dl2 Di) = ?du2 Di" unfolding c d unlabel_def by force+ have 1: "\M; unlabel B'\\<^sub>d \ " using prems(2) B'(2) 0(1) unfolding unlabel_def by auto { fix j p k q assume "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set [d\D. d \ set Di]" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set [d\D. d \ set Di]" hence "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" using dbproj_subseq_subset[OF B'(1)] by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence 4: "?R3 S [d\D. d \ set Di]" by blast - + have 5: "?R4 S (filter (\d. d \ set Di) D)" using prems(4) by force have 6: "?R5 S D" using prems(5) by force obtain C where C: "C \ set (tr'\<^sub>p\<^sub>c S [d\D . d \ set Di])" "\M; unlabel C\\<^sub>d \" using IH[OF B'(1,3) 1 4 5 6] by moura have 7: "\M; unlabel (c Di)\\<^sub>d \" "\M; unlabel B'\\<^sub>d \" using prems(2) B'(2) 0(1) strand_sem_split(3,4)[of M "unlabel (c Di)" "unlabel B'"] unfolding c unlabel_def by auto - have "\M; unlabel (?cl2 Di)\\<^sub>d \" using 7(1) 0(1) unfolding c unlabel_def by auto + have "\M; unlabel (?cl2 Di)\\<^sub>d \" using 7(1) 0(1) unfolding c unlabel_def by auto hence "\M; ?cu2 Di\\<^sub>d \" by (metis 0(4)) moreover { fix j p k q assume "(j, p) \ {(i, t, s)} \ set D \ set Di" "(k, q) \ {(i, t, s)} \ set D \ set Di" hence "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" using dbproj_subseq_subset[OF B'(1)] by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence "\(j, p) \ {(i, t, s)} \ set D \ set Di. \(k, q) \ {(i, t, s)} \ set D \ set Di. (\\. Unifier \ (pair p) (pair q)) \ j = k" by blast ultimately have "\M; ?du2 Di\\<^sub>d \" using labeled_sat_ineq_lift by simp hence "\M; unlabel (?dl2 Di)\\<^sub>d \" by (metis 0(6)) moreover have "\M; unlabel (?cl1 Di)\\<^sub>d \" using 7(1) unfolding c unlabel_def by auto hence "\M; unlabel (?dl1 Di)\\<^sub>d \" by (metis 0(3,5)) ultimately have "\M; unlabel (d Di)\\<^sub>d \" using 0(2) unfolding c d unlabel_def by force hence 8: "\M; unlabel (d Di@C)\\<^sub>d \" using 0(2) C(2) unfolding unlabel_def by auto have 9: "d Di@C \ set (tr'\<^sub>p\<^sub>c ((i,delete\t,s\)#S) D)" using C(1) dbproj_subseq_in_subseqs[OF B'(1)] unfolding d unlabel_def by auto show ?case by (metis 8 9) next case (7 i ac t s S D) note prems = "7.prems" note IH = "7.IH" obtain B' d where B': "B = (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#B'" "B' \ set (tr\<^sub>p\<^sub>c S D)" "d \ set (dbproj i D)" using prems(1) by moura have 1: "\M; unlabel B'\\<^sub>d \ " using prems(2) B'(1) by simp { fix j p k q assume "(j,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D" "(k,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D" hence "(j,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, InSet ac t s)#S) \ set D" "(k,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, InSet ac t s)#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence 4: "?R3 S D" by blast have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force have 7: "pair (t,s) \ \ = pair (snd d) \ \" using prems(2) B'(1) by simp obtain C where C: "C \ set (tr'\<^sub>p\<^sub>c S D)" "\M; unlabel C\\<^sub>d \" using IH[OF B'(2) 1 4 5 6] by moura hence "((i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#C) \ set (tr'\<^sub>p\<^sub>c ((i,InSet ac t s)#S) D)" "\M; unlabel ((i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#C)\\<^sub>d \" using 7 B'(3) by auto thus ?case by metis next case (8 i X F F' S D) note prems = "8.prems" note IH = "8.IH" let ?cl = "map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" let ?cu = "map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" let ?dl = "map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd D))" let ?du = "map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd D))" define c where c: "c = ?cl" define d where d: "d = ?dl" obtain B' where B': "B = c@B'" "B' \ set (tr\<^sub>p\<^sub>c S D)" using prems(1) c by moura have 0: "ik\<^sub>s\<^sub>t (unlabel c) = {}" "ik\<^sub>s\<^sub>t (unlabel d) = {}" - "unlabel ?cl = ?cu" "unlabel ?dl = ?du" + "unlabel ?cl = ?cu" "unlabel ?dl = ?du" unfolding c d unlabel_def by force+ have "ik\<^sub>s\<^sub>t (unlabel c) = {}" unfolding c unlabel_def by force hence 1: "\M; unlabel B'\\<^sub>d \ " using prems(2) B'(1) unfolding unlabel_def by auto have "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, NegChecks X F F')#S)" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence 4: "?R3 S D" using prems(3) by blast have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force obtain C where C: "C \ set (tr'\<^sub>p\<^sub>c S D)" "\M; unlabel C\\<^sub>d \" using IH[OF B'(2) 1 4 5 6] by moura have 7: "\M; unlabel c\\<^sub>d \" "\M; unlabel B'\\<^sub>d \" using prems(2) B'(1) 0(1) strand_sem_split(3,4)[of M "unlabel c" "unlabel B'"] unfolding c unlabel_def by auto have 8: "d@C \ set (tr'\<^sub>p\<^sub>c ((i,NegChecks X F F')#S) D)" using C(1) unfolding d unlabel_def by auto - have "\M; unlabel ?cl\\<^sub>d \" using 7(1) unfolding c unlabel_def by auto + have "\M; unlabel ?cl\\<^sub>d \" using 7(1) unfolding c unlabel_def by auto hence "\M; ?cu\\<^sub>d \" by (metis 0(3)) moreover { fix j p k q assume "(j, p) \ ((\(t,s). (i,t,s)) ` set F') \ set D" "(k, q) \ ((\(t,s). (i,t,s)) ` set F') \ set D" hence "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, NegChecks X F F')#S) \ set D" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, NegChecks X F F')#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence "\(j, p) \ ((\(t,s). (i,t,s)) ` set F') \ set D. \(k, q) \ ((\(t,s). (i,t,s)) ` set F') \ set D. (\\. Unifier \ (pair p) (pair q)) \ j = k" by blast moreover have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (map snd D) \ set X = {}" using prems(4) by fastforce ultimately have "\M; ?du\\<^sub>d \" using labeled_sat_ineq_dbproj_sem_equiv[of i] by simp hence "\M; unlabel ?dl\\<^sub>d \" by (metis 0(4)) hence "\M; unlabel d\\<^sub>d \" using 0(2) unfolding c d unlabel_def by force hence 9: "\M; unlabel (d@C)\\<^sub>d \" using 0(2) C(2) unfolding unlabel_def by auto show ?case by (metis 8 9) qed } thus "?P \ ?Q" by metis { fix C assume "C \ set (tr'\<^sub>p\<^sub>c A D)" "\M; unlabel C\\<^sub>d \" hence ?P using assms proof (induction A D arbitrary: C M rule: tr'\<^sub>p\<^sub>c.induct) case (1 D) thus ?case by simp next case (2 i t S D) note prems = "2.prems" note IH = "2.IH" obtain C' where C': "C = (i,send\t\\<^sub>s\<^sub>t)#C'" "C' \ set (tr'\<^sub>p\<^sub>c S D)" using prems(1) by moura - + have 1: "\M; unlabel C'\\<^sub>d \ " using prems(2) C'(1) by simp have 4: "?R3 S D" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force have 7: "M \ t \ \" using prems(2) C'(1) by simp obtain B where B: "B \ set (tr\<^sub>p\<^sub>c S D)" "\M; unlabel B\\<^sub>d \" using IH[OF C'(2) 1 4 5 6] by moura hence "((i,send\t\\<^sub>s\<^sub>t)#B) \ set (tr\<^sub>p\<^sub>c ((i,Send t)#S) D)" "\M; unlabel ((i,send\t\\<^sub>s\<^sub>t)#B)\\<^sub>d \" using 7 by auto thus ?case by metis next case (3 i t S D) note prems = "3.prems" note IH = "3.IH" obtain C' where C': "C = (i,receive\t\\<^sub>s\<^sub>t)#C'" "C' \ set (tr'\<^sub>p\<^sub>c S D)" using prems(1) by moura - + have 1: "\insert (t \ \) M; unlabel C'\\<^sub>d \ " using prems(2) C'(1) by simp have 4: "?R3 S D" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force obtain B where B: "B \ set (tr\<^sub>p\<^sub>c S D)" "\insert (t \ \) M; unlabel B\\<^sub>d \" using IH[OF C'(2) 1 4 5 6] by moura hence "((i,receive\t\\<^sub>s\<^sub>t)#B) \ set (tr\<^sub>p\<^sub>c ((i,receive\t\)#S) D)" "\insert (t \ \) M; unlabel ((i,receive\t\\<^sub>s\<^sub>t)#B)\\<^sub>d \" by auto thus ?case by auto next case (4 i ac t t' S D) note prems = "4.prems" note IH = "4.IH" obtain C' where C': "C = (i,\ac: t \ t'\\<^sub>s\<^sub>t)#C'" "C' \ set (tr'\<^sub>p\<^sub>c S D)" using prems(1) by moura - + have 1: "\M; unlabel C'\\<^sub>d \ " using prems(2) C'(1) by simp have 4: "?R3 S D" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force have 7: "t \ \ = t' \ \" using prems(2) C'(1) by simp obtain B where B: "B \ set (tr\<^sub>p\<^sub>c S D)" "\M; unlabel B\\<^sub>d \" using IH[OF C'(2) 1 4 5 6] by moura hence "((i,\ac: t \ t'\\<^sub>s\<^sub>t)#B) \ set (tr\<^sub>p\<^sub>c ((i,Equality ac t t')#S) D)" "\M; unlabel ((i,\ac: t \ t'\\<^sub>s\<^sub>t)#B)\\<^sub>d \" using 7 by auto thus ?case by metis next case (5 i t s S D) note prems = "5.prems" note IH = "5.IH" have C: "C \ set (tr'\<^sub>p\<^sub>c S (List.insert (i,t,s) D))" using prems(1) by simp - + have 1: "\M; unlabel C\\<^sub>d \ " using prems(2) C(1) by simp have 4: "?R3 S (List.insert (i,t,s) D)" using prems(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) have 5: "?R4 S (List.insert (i,t,s) D)" using prems(4,5) by force have 6: "?R5 S (List.insert (i,t,s) D)" using prems(5) by force show ?case using IH[OF C(1) 1 4 5 6] by simp next case (6 i t s S D) note prems = "6.prems" note IH = "6.IH" let ?dl1 = "\Di. map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di" let ?du1 = "\Di. map (\d. \check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t) Di" let ?dl2 = "\Di. map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\dbproj i D. d\set Di]" let ?du2 = "\Di. map (\d. \[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\dbproj i D. d\set Di]" let ?cl1 = "\Di. map (\d. (i,\check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)) Di" let ?cu1 = "\Di. map (\d. \check: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t) Di" let ?cl2 = "\Di. map (\d. (i,\[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t)) [d\D. d\set Di]" let ?cu2 = "\Di. map (\d. \[]\\\: [(pair (t,s), pair (snd d))]\\<^sub>s\<^sub>t) [d\D. d\set Di]" define c where c: "c = (\Di. ?cl1 Di@?cl2 Di)" define d where d: "d = (\Di. ?dl1 Di@?dl2 Di)" obtain C' Di where C': "Di \ set (subseqs D)" "C = c Di@C'" "C' \ set (tr'\<^sub>p\<^sub>c S [d\D. d \ set Di])" using prems(1) c by moura have 0: "ik\<^sub>s\<^sub>t (unlabel (c Di)) = {}" "ik\<^sub>s\<^sub>t (unlabel (d Di)) = {}" "unlabel (?cl1 Di) = ?cu1 Di" "unlabel (?cl2 Di) = ?cu2 Di" "unlabel (?dl1 Di) = ?du1 Di" "unlabel (?dl2 Di) = ?du2 Di" unfolding c d unlabel_def by force+ have 1: "\M; unlabel C'\\<^sub>d \ " using prems(2) C'(2) 0(1) unfolding unlabel_def by auto { fix j p k q assume "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set [d\D. d \ set Di]" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set [d\D. d \ set Di]" hence "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence 4: "?R3 S [d\D. d \ set Di]" by blast have 5: "?R4 S (filter (\d. d \ set Di) D)" using prems(4) by force have 6: "?R5 S D" using prems(5) by force obtain B where B: "B \ set (tr\<^sub>p\<^sub>c S [d\D. d \ set Di])" "\M; unlabel B\\<^sub>d \" using IH[OF C'(1,3) 1 4 5 6] by moura have 7: "\M; unlabel (c Di)\\<^sub>d \" "\M; unlabel C'\\<^sub>d \" using prems(2) C'(2) 0(1) strand_sem_split(3,4)[of M "unlabel (c Di)" "unlabel C'"] unfolding c unlabel_def by auto { fix j p k q assume "(j, p) \ {(i, t, s)} \ set D" "(k, q) \ {(i, t, s)} \ set D" hence "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, delete\t,s\)#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence "\(j, p) \ {(i, t, s)} \ set D. \(k, q) \ {(i, t, s)} \ set D. (\\. Unifier \ (pair p) (pair q)) \ j = k" by blast - moreover have "\M; unlabel (?cl1 Di)\\<^sub>d \" using 7(1) unfolding c unlabel_append by auto + moreover have "\M; unlabel (?cl1 Di)\\<^sub>d \" using 7(1) unfolding c unlabel_append by auto hence "\M; ?cu1 Di\\<^sub>d \" by (metis 0(3)) ultimately have *: "Di \ set (subseqs (dbproj i D))" using labeled_sat_eqs_subseqs[OF C'(1)] by simp hence 8: "d Di@B \ set (tr\<^sub>p\<^sub>c ((i,delete\t,s\)#S) D)" using B(1) unfolding d unlabel_def by auto - have "\M; unlabel (?cl2 Di)\\<^sub>d \" using 7(1) 0(1) unfolding c unlabel_def by auto + have "\M; unlabel (?cl2 Di)\\<^sub>d \" using 7(1) 0(1) unfolding c unlabel_def by auto hence "\M; ?cu2 Di\\<^sub>d \" by (metis 0(4)) hence "\M; ?du2 Di\\<^sub>d \" by (metis labeled_sat_ineq_dbproj) hence "\M; unlabel (?dl2 Di)\\<^sub>d \" by (metis 0(6)) - moreover have "\M; unlabel (?cl1 Di)\\<^sub>d \" using 7(1) unfolding c unlabel_def by auto + moreover have "\M; unlabel (?cl1 Di)\\<^sub>d \" using 7(1) unfolding c unlabel_def by auto hence "\M; unlabel (?dl1 Di)\\<^sub>d \" by (metis 0(3,5)) ultimately have "\M; unlabel (d Di)\\<^sub>d \" using 0(2) unfolding c d unlabel_def by force hence 9: "\M; unlabel (d Di@B)\\<^sub>d \" using 0(2) B(2) unfolding unlabel_def by auto show ?case by (metis 8 9) next case (7 i ac t s S D) note prems = "7.prems" note IH = "7.IH" obtain C' d where C': "C = (i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#C'" "C' \ set (tr'\<^sub>p\<^sub>c S D)" "d \ set D" using prems(1) by moura - + have 1: "\M; unlabel C'\\<^sub>d \ " using prems(2) C'(1) by simp { fix j p k q assume "(j,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D" "(k,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ set D" hence "(j,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, InSet ac t s)#S) \ set D" "(k,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, InSet ac t s)#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence 4: "?R3 S D" by blast have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force obtain B where B: "B \ set (tr\<^sub>p\<^sub>c S D)" "\M; unlabel B\\<^sub>d \" using IH[OF C'(2) 1 4 5 6] by moura have 7: "pair (t,s) \ \ = pair (snd d) \ \" using prems(2) C'(1) by simp have "(i,t,s) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, InSet ac t s)#S) \ set D" "(fst d, snd d) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, InSet ac t s)#S) \ set D" using C'(3) by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "\\. Unifier \ (pair (t,s)) (pair (snd d)) \ i = fst d" using prems(3) by blast hence "fst d = i" using 7 by auto hence 8: "d \ set (dbproj i D)" using C'(3) by auto have 9: "((i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#B) \ set (tr\<^sub>p\<^sub>c ((i,InSet ac t s)#S) D)" using B 8 by auto have 10: "\M; unlabel ((i,\ac: (pair (t,s)) \ (pair (snd d))\\<^sub>s\<^sub>t)#B)\\<^sub>d \" using B 7 8 by auto show ?case by (metis 9 10) next case (8 i X F F' S D) note prems = "8.prems" note IH = "8.IH" let ?dl = "map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" let ?du = "map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd (dbproj i D)))" let ?cl = "map (\G. (i,\X\\\: (F@G)\\<^sub>s\<^sub>t)) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd D))" let ?cu = "map (\G. \X\\\: (F@G)\\<^sub>s\<^sub>t) (tr\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F' (map snd D))" define c where c: "c = ?cl" define d where d: "d = ?dl" obtain C' where C': "C = c@C'" "C' \ set (tr'\<^sub>p\<^sub>c S D)" using prems(1) c by moura have 0: "ik\<^sub>s\<^sub>t (unlabel c) = {}" "ik\<^sub>s\<^sub>t (unlabel d) = {}" - "unlabel ?cl = ?cu" "unlabel ?dl = ?du" + "unlabel ?cl = ?cu" "unlabel ?dl = ?du" unfolding c d unlabel_def by force+ have "ik\<^sub>s\<^sub>t (unlabel c) = {}" unfolding c unlabel_def by force hence 1: "\M; unlabel C'\\<^sub>d \ " using prems(2) C'(1) unfolding unlabel_def by auto have "setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t S \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, NegChecks X F F')#S)" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence 4: "?R3 S D" using prems(3) by blast have 5: "?R4 S D" using prems(4) by force have 6: "?R5 S D" using prems(5) by force obtain B where B: "B \ set (tr\<^sub>p\<^sub>c S D)" "\M; unlabel B\\<^sub>d \" using IH[OF C'(2) 1 4 5 6] by moura have 7: "\M; unlabel c\\<^sub>d \" "\M; unlabel C'\\<^sub>d \" using prems(2) C'(1) 0(1) strand_sem_split(3,4)[of M "unlabel c" "unlabel C'"] unfolding c unlabel_def by auto have 8: "d@B \ set (tr\<^sub>p\<^sub>c ((i,NegChecks X F F')#S) D)" using B(1) unfolding d unlabel_def by auto - have "\M; unlabel ?cl\\<^sub>d \" using 7(1) unfolding c unlabel_def by auto + have "\M; unlabel ?cl\\<^sub>d \" using 7(1) unfolding c unlabel_def by auto hence "\M; ?cu\\<^sub>d \" by (metis 0(3)) moreover { fix j p k q assume "(j, p) \ ((\(t,s). (i,t,s)) ` set F') \ set D" "(k, q) \ ((\(t,s). (i,t,s)) ` set F') \ set D" hence "(j, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, NegChecks X F F')#S) \ set D" "(k, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t ((i, NegChecks X F F')#S) \ set D" by (auto simp add: setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def) hence "(\\. Unifier \ (pair p) (pair q)) \ j = k" using prems(3) by blast } hence "\(j, p) \ ((\(t,s). (i,t,s)) ` set F') \ set D. \(k, q) \ ((\(t,s). (i,t,s)) ` set F') \ set D. (\\. Unifier \ (pair p) (pair q)) \ j = k" by blast moreover have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (map snd D) \ set X = {}" using prems(4) by fastforce ultimately have "\M; ?du\\<^sub>d \" using labeled_sat_ineq_dbproj_sem_equiv[of i] by simp hence "\M; unlabel ?dl\\<^sub>d \" by (metis 0(4)) hence "\M; unlabel d\\<^sub>d \" using 0(2) unfolding c d unlabel_def by force hence 9: "\M; unlabel (d@B)\\<^sub>d \" using 0(2) B(2) unfolding unlabel_def by auto show ?case by (metis 8 9) qed } thus "?Q \ ?P" by metis qed subsubsection \Part 3\ private lemma tr'_par_sem_equiv: assumes "\(l,t,s) \ set D. (fv t \ fv s) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" and "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" "ground M" and "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. (\\. Unifier \ (pair p) (pair q)) \ i = j" (is "?R A D") and \: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\M; set (unlabel D) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \; unlabel A\\<^sub>s \ \ (\B \ set (tr'\<^sub>p\<^sub>c A D). \M; unlabel B\\<^sub>d \)" (is "?P \ ?Q") proof - have 1: "\(t,s) \ set (unlabel D). (fv t \ fv s) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" using assms(1) unfolding unlabel_def by force have 2: "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. p = q \ i = j" using assms(4) subst_apply_term_empty by blast show ?thesis by (metis tr_sem_equiv'[OF 1 assms(2,3) \] tr'_par_iff_unlabel_tr[OF 2]) qed subsubsection \Part 4\ lemma tr_par_sem_equiv: assumes "\(l,t,s) \ set D. (fv t \ fv s) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" and "fv\<^sub>s\<^sub>s\<^sub>t (unlabel A) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel A) = {}" "ground M" and "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ set D. (\\. Unifier \ (pair p) (pair q)) \ i = j" and \: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\M; set (unlabel D) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \; unlabel A\\<^sub>s \ \ (\B \ set (tr\<^sub>p\<^sub>c A D). \M; unlabel B\\<^sub>d \)" (is "?P \ ?Q") using tr_par_iff_tr'_par[OF assms(4,1,2), of M \] tr'_par_sem_equiv[OF assms] by metis end subsection \Theorem: The Stateful Compositionality Result, on the Constraint Level\ theorem par_comp_constr_stateful: assumes \: "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ Sec" "typing_cond\<^sub>s\<^sub>s\<^sub>t (unlabel \)" and \: "\ \\<^sub>s unlabel \" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\\\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\) \ (\\<^sub>\ \\<^sub>s unlabel \) \ ((\n. \\<^sub>\ \\<^sub>s proj_unl n \) \ (\\'. prefix \' \ \ (\' leaks Sec under \\<^sub>\)))" proof - let ?P = "\n A D. \(i, p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A) \ set D. \(j, q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A) \ set D. (\\. Unifier \ (pair p) (pair q)) \ i = j" have 1: "\(l, t, t')\set []. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel \) = {}" "fv\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel \) = {}" "ground {}" using \(2) unfolding typing_cond\<^sub>s\<^sub>s\<^sub>t_def by simp_all have 2: "\n. \(l, t, t')\set []. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n \) = {}" "\n. fv\<^sub>s\<^sub>s\<^sub>t (proj_unl n \) \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n \) = {}" using 1(1,2) sst_vars_proj_subset[of _ \] by fast+ have 3: "\n. par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n \) Sec" using par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj[OF \(1)] by metis have 4: "\{}; set (unlabel []) \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \'; unlabel \\\<^sub>s \' \ (\B\set (tr\<^sub>p\<^sub>c \ []). \{}; unlabel B\\<^sub>d \')" when \': "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \'" for \' using tr_par_sem_equiv[OF 1 _ \'] \(1) unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def constr_sem_d_def by auto obtain \' where \': "\' \ set (tr\<^sub>p\<^sub>c \ [])" "\ \ \unlabel \'\" using 4[OF \(2)] \(1) unfolding constr_sem_d_def by moura obtain \\<^sub>\ where \\<^sub>\: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" "\\<^sub>\ \ \unlabel \'\" "(\n. (\\<^sub>\ \ \proj_unl n \'\)) \ (\\''. prefix \'' \' \ (strand_leaks\<^sub>l\<^sub>s\<^sub>t \'' Sec \\<^sub>\))" using par_comp_constr[OF tr_par_preserves_par_comp[OF \(1) \'(1)] tr_par_preserves_typing_cond[OF \ \'(1)] \'(2) \(2)] by moura have \\<^sub>\': "\\<^sub>\ \\<^sub>s unlabel \" using 4[OF \\<^sub>\(1)] \'(1) \\<^sub>\(4) unfolding constr_sem_d_def by auto show ?thesis proof (cases "\n. (\\<^sub>\ \ \proj_unl n \'\)") case True { fix n assume "\\<^sub>\ \ \proj_unl n \'\" hence "\{}; {}; unlabel (proj n \)\\<^sub>s \\<^sub>\" using tr_par_proj[OF \'(1), of n] tr_par_sem_equiv[OF 2(1,2) 1(3) _ \\<^sub>\(1), of n] 3(1) unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def proj_def constr_sem_d_def by force } thus ?thesis using True \\<^sub>\(1,2,3) \\<^sub>\' by metis next case False then obtain \''::"('fun,'var,'lbl) labeled_strand" where \'': "prefix \'' \'" "strand_leaks\<^sub>l\<^sub>s\<^sub>t \'' Sec \\<^sub>\" using \\<^sub>\ by blast moreover { fix t l assume *: "\{}; unlabel (proj l \'')@[send\t\\<^sub>s\<^sub>t]\\<^sub>d \\<^sub>\" have "\\<^sub>\ \ \unlabel (proj l \'')\" "ik\<^sub>s\<^sub>t (unlabel (proj l \'')) \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\ \ t \ \\<^sub>\" using strand_sem_split(3,4)[OF *] unfolding constr_sem_d_def by auto } ultimately have "\t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t \'' \\<^sub>\. \l. (\\<^sub>\ \ \unlabel (proj l \'')\) \ ik\<^sub>s\<^sub>t (unlabel (proj l \'')) \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\ \ t \ \\<^sub>\" unfolding strand_leaks\<^sub>l\<^sub>s\<^sub>t_def constr_sem_d_def by metis then obtain s m where sm: "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t \'' \\<^sub>\" "\\<^sub>\ \ \unlabel (proj m \'')\" "ik\<^sub>s\<^sub>t (unlabel (proj m \'')) \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\ \ s \ \\<^sub>\" by moura \ \ We now need to show that there is some prefix \B\ of \\''\ that also leaks and where \B \ set (tr C D)\ for some prefix \C\ of \\\ \ obtain B::"('fun,'var,'lbl) labeled_strand" and C::"('fun,'var,'lbl) labeled_stateful_strand" where BC: "prefix B \'" "prefix C \" "B \ set (tr\<^sub>p\<^sub>c C [])" "ik\<^sub>s\<^sub>t (unlabel (proj m B)) \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\ \ s \ \\<^sub>\" "prefix B \''" using tr_leaking_prefix_exists[OF \'(1) \''(1) sm(3)] prefix_order.order_trans[OF _ \''(1)] by auto have "\{}; unlabel (proj m B)\\<^sub>d \\<^sub>\" using sm(2) BC(5) unfolding prefix_def unlabel_def proj_def constr_sem_d_def by auto hence BC': "\\<^sub>\ \ \proj_unl m B@[send\s\\<^sub>s\<^sub>t]\" using BC(4) unfolding constr_sem_d_def by auto have BC'': "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>t B \\<^sub>\" using BC(5) sm(1) unfolding prefix_def declassified\<^sub>l\<^sub>s\<^sub>t_def by auto have 5: "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n C) Sec" for n using \(1) BC(2) par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_split(1)[THEN par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_proj] unfolding prefix_def by auto have "fv\<^sub>s\<^sub>s\<^sub>t (unlabel \) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel \) = {}" "fv\<^sub>s\<^sub>s\<^sub>t (unlabel C) \ fv\<^sub>s\<^sub>s\<^sub>t (unlabel \)" "bvars\<^sub>s\<^sub>s\<^sub>t (unlabel C) \ bvars\<^sub>s\<^sub>s\<^sub>t (unlabel \)" using \(2) BC(2) sst_vars_append_subset(1,2)[of "unlabel C"] unfolding typing_cond\<^sub>s\<^sub>s\<^sub>t_def prefix_def unlabel_def by auto hence "fv\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) = {}" for n using sst_vars_proj_subset[of _ C] sst_vars_proj_subset[of _ \] by blast hence 6: "\(l, t, t')\set []. (fv t \ fv t') \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) = {}" "fv\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) \ bvars\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) = {}" "ground {}" for n using 2 by auto have 7: "?P n C []" for n using 5 unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by simp have "s \ \\<^sub>\ = s" using \\<^sub>\(1) BC'' \(1) unfolding par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto hence "\n. (\\<^sub>\ \\<^sub>s proj_unl n C) \ ik\<^sub>s\<^sub>s\<^sub>t (proj_unl n C) \\<^sub>s\<^sub>e\<^sub>t \\<^sub>\ \ s \ \\<^sub>\" using tr_par_proj[OF BC(3), of m] BC'(1) tr_par_sem_equiv[OF 6 7 \\<^sub>\(1), of m] tr_par_deduct_iff[OF tr_par_proj(1)[OF BC(3)], of \\<^sub>\ m s] unfolding proj_def constr_sem_d_def by auto hence "\n. \\<^sub>\ \\<^sub>s (proj_unl n C@[Send s])" using strand_sem_append_stateful by simp moreover have "s \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t C \\<^sub>\" by (metis tr_par_declassified_eq BC(3) BC'') ultimately show ?thesis using \\<^sub>\(1,2,3) \\<^sub>\' BC(2) unfolding strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by metis qed qed subsection \Theorem: The Stateful Compositionality Result, on the Protocol Level\ abbreviation wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t V \ \ wf'\<^sub>s\<^sub>s\<^sub>t V (unlabel \)" text \ We state our result on the level of protocol traces (i.e., the constraints reachable in a symbolic execution of the actual protocol). Hence, we do not need to convert protocol strands to intruder constraints in the following well-formedness definitions. \ definition wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s::"('fun,'var,'lbl) labeled_stateful_strand set \ bool" where "wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s \ \ (\\ \ \. wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t {} \) \ (\\ \ \. \\' \ \. fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' = {})" definition wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s':: "('fun,'var,'lbl) labeled_stateful_strand set \ ('fun,'var,'lbl) labeled_stateful_strand \ bool" where "wf\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>s' \ \ \ (\\' \ \. wf'\<^sub>s\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \) (unlabel \')) \ (\\' \ \. \\'' \ \. fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \'' = {}) \ (\\' \ \. fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ = {}) \ (\\' \ \. fv\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ bvars\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' = {})" definition typing_cond_prot_stateful where "typing_cond_prot_stateful \

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

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

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

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

)) \ (\S \ \

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

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

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

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

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

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

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

Sec" "\ \ \

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

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

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

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

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

" "\ \ \

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

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

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

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

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

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

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

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

)"] unfolding typing_cond_prot_stateful_def by (metis, metis, auto) hence 2: "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel \))" and 3: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t \)" unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by (meson subsetD)+ have 4: "list_all tfr\<^sub>s\<^sub>s\<^sub>t\<^sub>p (unlabel \)" using assms unfolding typing_cond_prot_stateful_def by auto show ?thesis using 1 2 3 4 unfolding typing_cond\<^sub>s\<^sub>s\<^sub>t_def tfr\<^sub>s\<^sub>s\<^sub>t_def by blast qed theorem par_comp_constr_prot_stateful: assumes P: "P = composed_prot Pi" "par_comp_prot_stateful P Sec" "\n. component_prot n (Pi n)" and left_secure: "component_secure_prot_stateful n (Pi n) Sec attack" shows "\\ \ P. suffix [(ln n, Send (Fun attack []))] \ \ unsat_stateful \ \ (\m. n \ m \ component_leaks_stateful m \ Sec)" proof - { fix \ \' assume \: "\ = \'@[(ln n, Send (Fun attack []))]" "\ \ P" let ?P = "\\' \\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\) \ prefix \' \ \ (\t \ Sec-declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\. \m. n \ m \ (\\<^sub>\ \\<^sub>s (proj_unl m \'@[Send t])))" have tcp: "typing_cond_prot_stateful P" using P(2) unfolding par_comp_prot_stateful_def by simp have par_comp: "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t \ Sec" "typing_cond\<^sub>s\<^sub>s\<^sub>t (unlabel \)" using par_comp_prot_impl_par_comp_stateful[OF P(2) \(2)] typing_cond_prot_impl_typing_cond_stateful[OF tcp \(2)] by metis+ - + have "unlabel (proj n \) = proj_unl n \" "proj_unl n \ = proj_unl n (proj n \)" - "\A. A \ Pi n \ proj n A = A" + "\A. A \ Pi n \ proj n A = A" "proj n \ = (proj n \')@[(ln n, Send (Fun attack []))]" using P(1,3) \ by (auto simp add: proj_def unlabel_def component_prot_def composed_prot_def) moreover have "proj n \ \ Pi n" using P(1) \ unfolding composed_prot_def by blast moreover { fix A assume "prefix A \" hence *: "prefix (proj n A) (proj n \)" unfolding proj_def prefix_def by force hence "proj_unl n A = proj_unl n (proj n A)" "\I. declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj n A) I" unfolding proj_def declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by auto hence "\B. prefix B (proj n \) \ proj_unl n A = proj_unl n B \ (\I. declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t A I = declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t B I)" using * by metis } - ultimately have *: + ultimately have *: "\\\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\) \ \(\\<^sub>\ \\<^sub>s (proj_unl n \)) \ (\\'. prefix \' \ \ (\t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\. \(\\<^sub>\ \\<^sub>s (proj_unl n \'@[Send t]))))" using left_secure unfolding component_secure_prot_stateful_def composed_prot_def suffix_def by metis { fix \ assume \: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "\ \\<^sub>s unlabel \" obtain \\<^sub>\ where \\<^sub>\: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" "\\'. prefix \' \ \ (\' leaks Sec under \\<^sub>\)" using par_comp_constr_stateful[OF par_comp \(2,1)] * by moura hence "\\'. prefix \' \ \ (\t \ Sec - declassified\<^sub>l\<^sub>s\<^sub>s\<^sub>t \' \\<^sub>\. \m. n \ m \ (\\<^sub>\ \\<^sub>s (proj_unl m \'@[Send t])))" using \\<^sub>\(4) * unfolding strand_leaks\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def by metis hence ?P using \\<^sub>\(1,2,3) by auto } hence "unsat_stateful \ \ (\m. n \ m \ component_leaks_stateful m \ Sec)" by (metis unsat_stateful_def component_leaks_stateful_def) } thus ?thesis unfolding suffix_def by metis qed end subsection \Automated Compositionality Conditions\ definition comp_GSMP_disjoint where "comp_GSMP_disjoint public arity Ana \ A' B' A B C \ let B\ = B \\<^sub>l\<^sub>i\<^sub>s\<^sub>t var_rename (max_var_set (fv\<^sub>s\<^sub>e\<^sub>t (set A))) in has_all_wt_instances_of \ (set A') (set A) \ has_all_wt_instances_of \ (set B') (set B\) \ finite_SMP_representation arity Ana \ A \ finite_SMP_representation arity Ana \ B\ \ (\t \ set A. \s \ set B\. \ t = \ s \ mgu t s \ None \ (intruder_synth' public arity {} t \ intruder_synth' public arity {} s) \ (\u \ set C. is_wt_instance_of_cond \ t u) \ (\u \ set C. is_wt_instance_of_cond \ s u))" definition comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t where "comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ pair_fun A M C \ let L = remdups (map (the_LabelN \ fst) (filter (Not \ is_LabelS) A)); MP0 = \B. remdups (trms_list\<^sub>s\<^sub>s\<^sub>t B@map (pair' pair_fun) (setops_list\<^sub>s\<^sub>s\<^sub>t B)); pr = \l. MP0 (proj_unl l A) in length L > 1 \ list_all (wf\<^sub>t\<^sub>r\<^sub>m' arity) (MP0 (unlabel A)) \ list_all (wf\<^sub>t\<^sub>r\<^sub>m' arity) C \ has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t (set C)) (set C) \ is_TComp_var_instance_closed \ C \ (\i \ set L. \j \ set L. i \ j \ comp_GSMP_disjoint public arity Ana \ (pr i) (pr j) (M i) (M j) C) \ (\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. i \ j \ (let s = pair' pair_fun p; t = pair' pair_fun q in mgu s (t \ var_rename (max_var s)) = None))" locale labeled_stateful_typed_model' = stateful_typed_model' arity public Ana \ Pair + labeled_typed_model' arity public Ana \ label_witness1 label_witness2 for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,(('fun,'atom::finite) term_type \ nat)) term \ (('fun,(('fun,'atom) term_type \ nat)) term list \ ('fun,(('fun,'atom) term_type \ nat)) term list)" and \::"('fun,(('fun,'atom) term_type \ nat)) term \ ('fun,'atom) term_type" and Pair::"'fun" and label_witness1::"'lbl" and label_witness2::"'lbl" begin sublocale labeled_stateful_typed_model by unfold_locales lemma GSMP_disjoint_if_comp_GSMP_disjoint: defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" assumes AB'_wf: "list_all (wf\<^sub>t\<^sub>r\<^sub>m' arity) A'" "list_all (wf\<^sub>t\<^sub>r\<^sub>m' arity) B'" and C_wf: "list_all (wf\<^sub>t\<^sub>r\<^sub>m' arity) C" and AB'_disj: "comp_GSMP_disjoint public arity Ana \ A' B' A B C" shows "GSMP_disjoint (set A') (set B') ((f (set C)) - {m. {} \\<^sub>c m})" using GSMP_disjointI[of A' B' A B] AB'_wf AB'_disj C_wf unfolding comp_GSMP_disjoint_def f_def wf\<^sub>t\<^sub>r\<^sub>m_code list_all_iff Let_def by fast lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_if_comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t: defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" assumes A: "comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ Pair A M C" shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t A ((f (set C)) - {m. {} \\<^sub>c m})" proof (unfold par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def; intro conjI) let ?Sec = "(f (set C)) - {m. {} \\<^sub>c m}" let ?L = "remdups (map (the_LabelN \ fst) (filter (Not \ is_LabelS) A))" let ?N1 = "\B. remdups (trms_list\<^sub>s\<^sub>s\<^sub>t B@map (pair' Pair) (setops_list\<^sub>s\<^sub>s\<^sub>t B))" let ?N2 = "\B. trms\<^sub>s\<^sub>s\<^sub>t B \ pair ` setops\<^sub>s\<^sub>s\<^sub>t B" let ?pr = "\l. ?N1 (proj_unl l A)" let ?\ = "\p. var_rename (max_var (pair p))" have 0: "length ?L > 1" "list_all (wf\<^sub>t\<^sub>r\<^sub>m' arity) (?N1 (unlabel A))" "list_all (wf\<^sub>t\<^sub>r\<^sub>m' arity) C" "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t (set C)) (set C)" "is_TComp_var_instance_closed \ C" "\i \ set ?L. \j \ set ?L. i \ j \ comp_GSMP_disjoint public arity Ana \ (?pr i) (?pr j) (M i) (M j) C" "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. i \ j \ mgu (pair p) (pair q \ ?\ p) = None" using A unfolding comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def pair_code by meson+ have L_in_iff: "l \ set ?L \ (\a \ set A. is_LabelN l a)" for l by force have A_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t A \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (unlabel A))" using 0(2) unfolding pair_code wf\<^sub>t\<^sub>r\<^sub>m_code list_all_iff trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t setops_list\<^sub>s\<^sub>s\<^sub>t_is_setops\<^sub>s\<^sub>s\<^sub>t by auto hence A_proj_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>l\<^sub>s\<^sub>s\<^sub>t (proj l A) \ pair ` setops\<^sub>s\<^sub>s\<^sub>t (proj_unl l A))" for l using trms\<^sub>s\<^sub>s\<^sub>t_proj_subset(1)[of l A] setops\<^sub>s\<^sub>s\<^sub>t_proj_subset(1)[of l A] by blast hence A_proj_wf_trms': "list_all (wf\<^sub>t\<^sub>r\<^sub>m' arity) (?N1 (proj_unl l A))" for l unfolding pair_code wf\<^sub>t\<^sub>r\<^sub>m_code list_all_iff trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t setops_list\<^sub>s\<^sub>s\<^sub>t_is_setops\<^sub>s\<^sub>s\<^sub>t by auto note C_wf_trms = 0(3)[unfolded list_all_iff wf\<^sub>t\<^sub>r\<^sub>m_code[symmetric]] note 1 = has_all_wt_instances_ofD'[OF wf_trms_subterms[OF C_wf_trms] C_wf_trms 0(4)] have 2: "GSMP (?N2 (proj_unl l A)) \ GSMP (?N2 (proj_unl l' A))" when "l \ set ?L" for l l' using that L_in_iff GSMP_mono[of "?N2 (proj_unl l A)" "?N2 (proj_unl l' A)"] trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subset_if_no_label[of l A] setops\<^sub>s\<^sub>s\<^sub>t_unlabel_subset_if_no_label[of l A] unfolding list_ex_iff by fast have 3: "GSMP_disjoint (?N2 (proj_unl l1 A)) (?N2 (proj_unl l2 A)) ?Sec" when "l1 \ set ?L" "l2 \ set ?L" "l1 \ l2" for l1 l2 proof - have "GSMP_disjoint (set (?N1 (proj_unl l1 A))) (set (?N1 (proj_unl l2 A))) ?Sec" using 0(6) that GSMP_disjoint_if_comp_GSMP_disjoint[ OF A_proj_wf_trms'[of l1] A_proj_wf_trms'[of l2] 0(3), of "M l1" "M l2"] unfolding f_def by blast thus ?thesis unfolding pair_code trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t setops_list\<^sub>s\<^sub>s\<^sub>t_is_setops\<^sub>s\<^sub>s\<^sub>t by simp qed obtain a1 a2 where a: "a1 \ set ?L" "a2 \ set ?L" "a1 \ a2" using remdups_ex2[OF 0(1)] by moura show "ground ?Sec" unfolding f_def by fastforce { fix i p j q assume p: "(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" and q: "(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" and pq: "\\. Unifier \ (pair p) (pair q)" have "\\. Unifier \ (pair p) (pair q \ ?\ p)" using pq vars_term_disjoint_imp_unifier[OF var_rename_fv_disjoint[of "pair p"], of _ "pair q"] by (metis (no_types, lifting) subst_subst_compose var_rename_inv_comp) hence "i = j" using 0(7) mgu_None_is_subst_neq[of "pair p" "pair q \ ?\ p"] p q by fast } thus "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. (\\. Unifier \ (pair p) (pair q)) \ i = j" by blast show "\l1 l2. l1 \ l2 \ GSMP_disjoint (?N2 (proj_unl l1 A)) (?N2 (proj_unl l2 A)) ?Sec" using 2 3 3[OF a] unfolding GSMP_disjoint_def by blast show "\s \ ?Sec. \s' \ subterms s. {} \\<^sub>c s' \ s' \ ?Sec" proof (intro ballI) fix s s' assume s: "s \ ?Sec" and s': "s' \ s" then obtain t \ where t: "t \ set C" "s = t \ \" "fv s = {}" "\{} \\<^sub>c s" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" unfolding f_def by blast obtain m \ where m: "m \ set C" "s' = m \ \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using TComp_var_and_subterm_instance_closed_has_subterms_instances[ OF 0(5,4) C_wf_trms in_subterms_Union[OF t(1)] s'[unfolded t(2)] \] by blast thus "{} \\<^sub>c s' \ s' \ ?Sec" using ground_subterm[OF t(3) s'] unfolding f_def by blast qed qed lemma par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_if_comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t': defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" assumes a: "comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t public arity Ana \ Pair A M C" and B: "\b \ set B. \a \ set A. \\. b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" (is "\b \ set B. \a \ set A. \\. b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \ ?D \") shows "par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t B ((f (set C)) - {m. {} \\<^sub>c m})" proof (unfold par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def; intro conjI) define N1 where "N1 \ \B::('fun, ('fun,'atom) term_type \ nat) stateful_strand. remdups (trms_list\<^sub>s\<^sub>s\<^sub>t B@map (pair' Pair) (setops_list\<^sub>s\<^sub>s\<^sub>t B))" define N2 where "N2 \ \B::('fun, ('fun,'atom) term_type \ nat) stateful_strand. trms\<^sub>s\<^sub>s\<^sub>t B \ pair ` setops\<^sub>s\<^sub>s\<^sub>t B" define L where "L \ \A::('fun, ('fun,'atom) term_type \ nat, 'lbl) labeled_stateful_strand. remdups (map (the_LabelN \ fst) (filter (Not \ is_LabelS) A))" define \ where "\ \ \p. var_rename (max_var (pair p::('fun, ('fun,'atom) term_type \ nat) term)) ::('fun, ('fun,'atom) term_type \ nat) subst" let ?Sec = "(f (set C)) - {m. {} \\<^sub>c m}" have 0: "length (L A) > 1" "list_all (wf\<^sub>t\<^sub>r\<^sub>m' arity) (N1 (unlabel A))" "list_all (wf\<^sub>t\<^sub>r\<^sub>m' arity) C" "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t (set C)) (set C)" "is_TComp_var_instance_closed \ C" "\i \ set (L A). \j \ set (L A). i \ j \ comp_GSMP_disjoint public arity Ana \ (N1 (proj_unl i A)) (N1 (proj_unl j A)) (M i) (M j) C" "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A. i \ j \ mgu (pair p) (pair q \ \ p) = None" using a unfolding comp_par_comp\<^sub>l\<^sub>s\<^sub>s\<^sub>t_def pair_code L_def N1_def \_def by meson+ note 1 = trms\<^sub>s\<^sub>s\<^sub>t_proj_subset(1) setops\<^sub>s\<^sub>s\<^sub>t_proj_subset(1) have N1_iff_N2: "set (N1 A) = N2 A" for A unfolding pair_code trms_list\<^sub>s\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>s\<^sub>t setops_list\<^sub>s\<^sub>s\<^sub>t_is_setops\<^sub>s\<^sub>s\<^sub>t N1_def N2_def by simp have N2_proj_subset: "N2 (proj_unl l A) \ N2 (unlabel A)" for l::'lbl and A::"('fun, ('fun,'atom) term_type \ nat, 'lbl) labeled_stateful_strand" using 1(1)[of l A] image_mono[OF 1(2)[of l A], of pair] unfolding N2_def by blast have L_in_iff: "l \ set (L A) \ (\a \ set A. is_LabelN l a)" for l A unfolding L_def by force have L_B_subset_A: "l \ set (L A)" when l: "l \ set (L B)" for l using L_in_iff[of l B] L_in_iff[of l A] B l by fastforce note B_setops = setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t_wt_instance_ex[OF B] have B_proj: "\b \ set (proj l B). \a \ set (proj l A). \\. b = a \\<^sub>l\<^sub>s\<^sub>s\<^sub>t\<^sub>p \ \ ?D \" for l using proj_instance_ex[OF B] by fast have B': "\t \ N2 (unlabel B). \s \ N2 (unlabel A). \\. t = s \ \ \ ?D \" using trms\<^sub>s\<^sub>s\<^sub>t_setops\<^sub>s\<^sub>s\<^sub>t_wt_instance_ex[OF B] unfolding N2_def by blast have B'_proj: "\t \ N2 (proj_unl l B). \s \ N2 (proj_unl l A). \\. t = s \ \ \ ?D \" for l using trms\<^sub>s\<^sub>s\<^sub>t_setops\<^sub>s\<^sub>s\<^sub>t_wt_instance_ex[OF B_proj] unfolding N2_def by presburger have A_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (N2 (unlabel A))" using N1_iff_N2[of "unlabel A"] 0(2) unfolding wf\<^sub>t\<^sub>r\<^sub>m_code list_all_iff by auto hence A_proj_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (N2 (proj_unl l A))" for l using 1[of l] unfolding N2_def by blast hence A_proj_wf_trms': "list_all (wf\<^sub>t\<^sub>r\<^sub>m' arity) (N1 (proj_unl l A))" for l using N1_iff_N2[of "proj_unl l A"] unfolding wf\<^sub>t\<^sub>r\<^sub>m_code list_all_iff by presburger note C_wf_trms = 0(3)[unfolded list_all_iff wf\<^sub>t\<^sub>r\<^sub>m_code[symmetric]] have 2: "GSMP (N2 (proj_unl l A)) \ GSMP (N2 (proj_unl l' A))" when "l \ set (L A)" for l l' and A::"('fun, ('fun,'atom) term_type \ nat, 'lbl) labeled_stateful_strand" using that L_in_iff[of _ A] GSMP_mono[of "N2 (proj_unl l A)" "N2 (proj_unl l' A)"] trms\<^sub>s\<^sub>s\<^sub>t_unlabel_subset_if_no_label[of l A] setops\<^sub>s\<^sub>s\<^sub>t_unlabel_subset_if_no_label[of l A] unfolding list_ex_iff N2_def by fast have 3: "GSMP (N2 (proj_unl l B)) \ GSMP (N2 (proj_unl l A))" (is "?X \ ?Y") for l proof fix t assume "t \ ?X" hence t: "t \ SMP (N2 (proj_unl l B))" "fv t = {}" unfolding GSMP_def by simp_all have "t \ SMP (N2 (proj_unl l A))" using t(1) B'_proj[of l] SMP_wt_instances_subset[of "N2 (proj_unl l B)" "N2 (proj_unl l A)"] by metis thus "t \ ?Y" using t(2) unfolding GSMP_def by fast qed have "GSMP_disjoint (N2 (proj_unl l1 A)) (N2 (proj_unl l2 A)) ?Sec" when "l1 \ set (L A)" "l2 \ set (L A)" "l1 \ l2" for l1 l2 proof - have "GSMP_disjoint (set (N1 (proj_unl l1 A))) (set (N1 (proj_unl l2 A))) ?Sec" using 0(6) that GSMP_disjoint_if_comp_GSMP_disjoint[ OF A_proj_wf_trms'[of l1] A_proj_wf_trms'[of l2] 0(3), of "M l1" "M l2"] unfolding f_def by blast thus ?thesis using N1_iff_N2 by simp qed hence 4: "GSMP_disjoint (N2 (proj_unl l1 B)) (N2 (proj_unl l2 B)) ?Sec" when "l1 \ set (L A)" "l2 \ set (L A)" "l1 \ l2" for l1 l2 using that 3 unfolding GSMP_disjoint_def by blast { fix i p j q assume p: "(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B" and q: "(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B" and pq: "\\. Unifier \ (pair p) (pair q)" obtain p' \p where p': "(i,p') \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "p = p' \\<^sub>p \p" "pair p = pair p' \ \p" using p B_setops unfolding pair_def by auto obtain q' \q where q': "(j,q') \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t A" "q = q' \\<^sub>p \q" "pair q = pair q' \ \q" using q B_setops unfolding pair_def by auto obtain \ where "Unifier \ (pair p) (pair q)" using pq by blast hence "\\. Unifier \ (pair p') (pair q' \ \ p')" using p'(3) q'(3) var_rename_inv_comp[of "pair q'"] subst_subst_compose vars_term_disjoint_imp_unifier[ OF var_rename_fv_disjoint[of "pair p'"], of "\p \\<^sub>s \" "pair q'" "var_rename_inv (max_var_set (fv (pair p'))) \\<^sub>s \q \\<^sub>s \"] unfolding \_def by fastforce hence "i = j" using mgu_None_is_subst_neq[of "pair p'" "pair q' \ \ p'"] p'(1) q'(1) 0(7) unfolding \_def by fast } thus "\(i,p) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B. \(j,q) \ setops\<^sub>l\<^sub>s\<^sub>s\<^sub>t B. (\\. Unifier \ (pair p) (pair q)) \ i = j" by blast obtain a1 a2 where a: "a1 \ set (L A)" "a2 \ set (L A)" "a1 \ a2" using remdups_ex2[OF 0(1)[unfolded L_def]] unfolding L_def by moura show "\l1 l2. l1 \ l2 \ GSMP_disjoint (N2 (proj_unl l1 B)) (N2 (proj_unl l2 B)) ?Sec" using 2[of _ B] 4 4[OF a] L_B_subset_A unfolding GSMP_disjoint_def by blast show "ground ?Sec" unfolding f_def by fastforce show "\s \ ?Sec. \s' \ subterms s. {} \\<^sub>c s' \ s' \ ?Sec" proof (intro ballI) fix s s' assume s: "s \ ?Sec" and s': "s' \ s" then obtain t \ where t: "t \ set C" "s = t \ \" "fv s = {}" "\{} \\<^sub>c s" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" unfolding f_def by blast obtain m \ where m: "m \ set C" "s' = m \ \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using TComp_var_and_subterm_instance_closed_has_subterms_instances[ OF 0(5,4) C_wf_trms in_subterms_Union[OF t(1)] s'[unfolded t(2)] \] by blast thus "{} \\<^sub>c s' \ s' \ ?Sec" using ground_subterm[OF t(3) s'] unfolding f_def by blast qed qed end end diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Typing_Result.thy b/thys/Stateful_Protocol_Composition_and_Typing/Typing_Result.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/Typing_Result.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/Typing_Result.thy @@ -1,3463 +1,3463 @@ (* (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: Typing_Result.thy Author: Andreas Viktor Hess, DTU *) section \The Typing Result\ theory Typing_Result imports Typed_Model begin subsection \The Typing Result for the Composition-Only Intruder\ context typed_model begin subsubsection \Well-typedness and Type-Flaw Resistance Preservation\ context begin private lemma LI_preserves_tfr_stp_all_single: assumes "(S,\) \ (S',\')" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" shows "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" using assms proof (induction rule: LI_rel.induct) case (Compose S X f S' \) hence "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" by simp_all moreover have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (map Send X)" by (induct X) auto ultimately show ?case by simp next case (Unify S f Y \ X S' \) hence "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S')" by simp have "fv\<^sub>s\<^sub>t (S@Send (Fun f X)#S') \ bvars\<^sub>s\<^sub>t (S@S') = {}" using Unify.prems(1) by (auto simp add: wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) moreover have "fv (Fun f X) \ fv\<^sub>s\<^sub>t (S@Send (Fun f X)#S')" by auto moreover have "fv (Fun f Y) \ fv\<^sub>s\<^sub>t (S@Send (Fun f X)#S')" using Unify.hyps(2) fv_subset_if_in_strand_ik'[of "Fun f Y" S] by force ultimately have bvars_disj: "bvars\<^sub>s\<^sub>t (S@S') \ fv (Fun f X) = {}" "bvars\<^sub>s\<^sub>t (S@S') \ fv (Fun f Y) = {}" by blast+ have "wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)" using Unify.prems(5) by simp moreover have "wf\<^sub>t\<^sub>r\<^sub>m (Fun f Y)" proof - obtain x where "x \ set S" "Fun f Y \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t\<^sub>p x)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p x)" using Unify.hyps(2) Unify.prems(5) by force+ thus ?thesis using wf_trm_subterm by auto qed moreover have "Fun f X \ SMP (trms\<^sub>s\<^sub>t (S@Send (Fun f X)#S'))" "Fun f Y \ SMP (trms\<^sub>s\<^sub>t (S@Send (Fun f X)#S'))" using SMP_append[of S "Send (Fun f X)#S'"] SMP_Cons[of "Send (Fun f X)" S'] SMP_ikI[OF Unify.hyps(2)] by auto hence "\ (Fun f X) = \ (Fun f Y)" using Unify.prems(4) mgu_gives_MGU[OF Unify.hyps(3)[symmetric]] unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast ultimately have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using mgu_wt_if_same_type[OF Unify.hyps(3)[symmetric]] by metis moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_wf_trm[OF Unify.hyps(3)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)\ \wf\<^sub>t\<^sub>r\<^sub>m (Fun f Y)\] by (metis wf_trm_subst_range_iff) moreover have "bvars\<^sub>s\<^sub>t (S@S') \ range_vars \ = {}" using mgu_vars_bounded[OF Unify.hyps(3)[symmetric]] bvars_disj by fast ultimately show ?case using tfr_stp_all_wt_subst_apply[OF \list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S')\] by metis next case (Equality S \ t t' a S' \) have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S')" "\ t = \ t'" using tfr_stp_all_same_type[of S a t t' S'] tfr_stp_all_split(5)[of S _ S'] MGU_is_Unifier[OF mgu_gives_MGU[OF Equality.hyps(2)[symmetric]]] Equality.prems(3) by blast+ moreover have "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m t'" using Equality.prems(5) by auto ultimately have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using mgu_wt_if_same_type[OF Equality.hyps(2)[symmetric]] by metis moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_wf_trm[OF Equality.hyps(2)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m t\ \wf\<^sub>t\<^sub>r\<^sub>m t'\] by (metis wf_trm_subst_range_iff) moreover have "fv\<^sub>s\<^sub>t (S@Equality a t t'#S') \ bvars\<^sub>s\<^sub>t (S@Equality a t t'#S') = {}" using Equality.prems(1) by (auto simp add: wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def) hence "bvars\<^sub>s\<^sub>t (S@S') \ fv t = {}" "bvars\<^sub>s\<^sub>t (S@S') \ fv t' = {}" by auto hence "bvars\<^sub>s\<^sub>t (S@S') \ range_vars \ = {}" using mgu_vars_bounded[OF Equality.hyps(2)[symmetric]] by fast ultimately show ?case using tfr_stp_all_wt_subst_apply[OF \list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S')\] by metis qed private lemma LI_in_SMP_subset_single: assumes "(S,\) \ (S',\')" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" and "trms\<^sub>s\<^sub>t S \ SMP M" shows "trms\<^sub>s\<^sub>t S' \ SMP M" using assms proof (induction rule: LI_rel.induct) case (Compose S X f S' \) hence "SMP (trms\<^sub>s\<^sub>t [Send (Fun f X)]) \ SMP M" proof - have "SMP (trms\<^sub>s\<^sub>t [Send (Fun f X)]) \ SMP (trms\<^sub>s\<^sub>t (S@Send (Fun f X)#S'))" using trms\<^sub>s\<^sub>t_append SMP_mono by auto thus ?thesis using SMP_union[of "trms\<^sub>s\<^sub>t (S@Send (Fun f X)#S')" M] SMP_subset_union_eq[OF Compose.prems(6)] by auto qed thus ?case using Compose.prems(6) by auto next case (Unify S f Y \ X S' \) have "Fun f X \ SMP (trms\<^sub>s\<^sub>t (S@Send (Fun f X)#S'))" by auto moreover have "MGU \ (Fun f X) (Fun f Y)" by (metis mgu_gives_MGU[OF Unify.hyps(3)[symmetric]]) moreover have "\x. x \ set S \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p x)" "wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)" using Unify.prems(4) by force+ moreover have "Fun f Y \ SMP (trms\<^sub>s\<^sub>t (S@Send (Fun f X)#S'))" by (meson SMP_ikI Unify.hyps(2) contra_subsetD ik_append_subset(1)) ultimately have "wf\<^sub>t\<^sub>r\<^sub>m (Fun f Y)" "\ (Fun f X) = \ (Fun f Y)" using ik\<^sub>s\<^sub>t_subterm_exD[OF \Fun f Y \ ik\<^sub>s\<^sub>t S\] \tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t (S@Send (Fun f X)#S'))\ unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by (metis (full_types) SMP_wf_trm Unify.prems(4), blast) hence "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" by (metis mgu_wt_if_same_type[OF Unify.hyps(3)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)\]) moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_wf_trm[OF Unify.hyps(3)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)\ \wf\<^sub>t\<^sub>r\<^sub>m (Fun f Y)\] by simp ultimately have "trms\<^sub>s\<^sub>t ((S@Send (Fun f X)#S') \\<^sub>s\<^sub>t \) \ SMP M" using SMP.Substitution Unify.prems(6) wt_subst_SMP_subset by metis thus ?case by auto next case (Equality S \ t t' a S' \) hence "\ t = \ t'" using tfr_stp_all_same_type MGU_is_Unifier[OF mgu_gives_MGU[OF Equality.hyps(2)[symmetric]]] by metis moreover have "t \ SMP (trms\<^sub>s\<^sub>t (S@Equality a t t'#S'))" "t' \ SMP (trms\<^sub>s\<^sub>t (S@Equality a t t'#S'))" using Equality.prems(1) by auto moreover have "MGU \ t t'" using mgu_gives_MGU[OF Equality.hyps(2)[symmetric]] by metis moreover have "\x. x \ set S \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p x)" "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m t'" using Equality.prems(4) by force+ ultimately have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" by (metis mgu_wt_if_same_type[OF Equality.hyps(2)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m t\]) moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_wf_trm[OF Equality.hyps(2)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m t\ \wf\<^sub>t\<^sub>r\<^sub>m t'\] by simp ultimately have "trms\<^sub>s\<^sub>t ((S@Equality a t t'#S') \\<^sub>s\<^sub>t \) \ SMP M" using SMP.Substitution Equality.prems wt_subst_SMP_subset by metis thus ?case by auto qed private lemma LI_preserves_tfr_single: assumes "(S,\) \ (S',\')" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" shows "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S') \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S')" using assms proof (induction rule: LI_rel.induct) case (Compose S X f S' \) let ?SMPmap = "SMP (trms\<^sub>s\<^sub>t (S@map Send X@S')) - (Var`\)" have "?SMPmap \ SMP (trms\<^sub>s\<^sub>t (S@Send (Fun f X)#S')) - (Var`\)" using SMP_fun_map_snd_subset[of X f] SMP_append[of "map Send X" S'] SMP_Cons[of "Send (Fun f X)" S'] SMP_append[of S "Send (Fun f X)#S'"] SMP_append[of S "map Send X@S'"] by auto hence "\s \ ?SMPmap. \t \ ?SMPmap. (\\. Unifier \ s t) \ \ s = \ t" using Compose unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by (meson subsetCE) thus ?case using LI_preserves_trm_wf[OF r_into_rtrancl[OF LI_rel.Compose[OF Compose.hyps]], of S'] Compose.prems(5) unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast next case (Unify S f Y \ X S' \) let ?SMP\ = "SMP (trms\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \)) - (Var`\)" have "SMP (trms\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \)) \ SMP (trms\<^sub>s\<^sub>t (S@Send (Fun f X)#S'))" proof fix s assume "s \ SMP (trms\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \))" thus "s \ SMP (trms\<^sub>s\<^sub>t (S@Send (Fun f X)#S'))" using LI_in_SMP_subset_single[ OF LI_rel.Unify[OF Unify.hyps] Unify.prems(1,2,4,5,6) MP_subset_SMP(2)[of "S@Send (Fun f X)#S'"]] by (metis SMP_union SMP_subset_union_eq Un_iff) qed hence "\s \ ?SMP\. \t \ ?SMP\. (\\. Unifier \ s t) \ \ s = \ t" using Unify.prems(4) unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by (meson Diff_iff subsetCE) thus ?case using LI_preserves_trm_wf[OF r_into_rtrancl[OF LI_rel.Unify[OF Unify.hyps]], of S'] Unify.prems(5) unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast next case (Equality S \ t t' a S' \) let ?SMP\ = "SMP (trms\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \)) - (Var`\)" have "SMP (trms\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \)) \ SMP (trms\<^sub>s\<^sub>t (S@Equality a t t'#S'))" proof fix s assume "s \ SMP (trms\<^sub>s\<^sub>t (S@S' \\<^sub>s\<^sub>t \))" thus "s \ SMP (trms\<^sub>s\<^sub>t (S@Equality a t t'#S'))" using LI_in_SMP_subset_single[ OF LI_rel.Equality[OF Equality.hyps] Equality.prems(1,2,4,5,6) MP_subset_SMP(2)[of "S@Equality a t t'#S'"]] by (metis SMP_union SMP_subset_union_eq Un_iff) qed hence "\s \ ?SMP\. \t \ ?SMP\. (\\. Unifier \ s t) \ \ s = \ t" using Equality.prems unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by (meson Diff_iff subsetCE) thus ?case using LI_preserves_trm_wf[OF r_into_rtrancl[OF LI_rel.Equality[OF Equality.hyps]], of _ S'] Equality.prems unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast qed private lemma LI_preserves_welltypedness_single: assumes "(S,\) \ (S',\')" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \' \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \')" using assms proof (induction rule: LI_rel.induct) case (Unify S f Y \ X S' \) have "wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)" using Unify.prems(5) unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by simp moreover have "wf\<^sub>t\<^sub>r\<^sub>m (Fun f Y)" proof - obtain x where "x \ set S" "Fun f Y \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t\<^sub>p x)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t\<^sub>p x)" using Unify.hyps(2) Unify.prems(5) unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by force thus ?thesis using wf_trm_subterm by auto qed moreover have "Fun f X \ SMP (trms\<^sub>s\<^sub>t (S@Send (Fun f X)#S'))" "Fun f Y \ SMP (trms\<^sub>s\<^sub>t (S@Send (Fun f X)#S'))" using SMP_append[of S "Send (Fun f X)#S'"] SMP_Cons[of "Send (Fun f X)" S'] SMP_ikI[OF Unify.hyps(2)] by auto hence "\ (Fun f X) = \ (Fun f Y)" using Unify.prems(4) mgu_gives_MGU[OF Unify.hyps(3)[symmetric]] unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast ultimately have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using mgu_wt_if_same_type[OF Unify.hyps(3)[symmetric]] by metis have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (meson mgu_wf_trm[OF Unify.hyps(3)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)\ \wf\<^sub>t\<^sub>r\<^sub>m (Fun f Y)\] wf_trm_subst_range_iff) hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using wf_trm_subst_range_iff wf_trm_subst \wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)\ unfolding subst_compose_def by (metis (no_types, lifting)) thus ?case by (metis wt_subst_compose[OF \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\]) next case (Equality S \ t t' a S' \) have "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m t'" using Equality.prems(5) by simp_all moreover have "\ t = \ t'" using \list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@Equality a t t'#S')\ MGU_is_Unifier[OF mgu_gives_MGU[OF Equality.hyps(2)[symmetric]]] by auto ultimately have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using mgu_wt_if_same_type[OF Equality.hyps(2)[symmetric]] by metis have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (meson mgu_wf_trm[OF Equality.hyps(2)[symmetric] \wf\<^sub>t\<^sub>r\<^sub>m t\ \wf\<^sub>t\<^sub>r\<^sub>m t'\] wf_trm_subst_range_iff) hence "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using wf_trm_subst_range_iff wf_trm_subst \wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)\ unfolding subst_compose_def by (metis (no_types, lifting)) thus ?case by (metis wt_subst_compose[OF \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\]) qed metis lemma LI_preserves_welltypedness: assumes "(S,\) \\<^sup>* (S',\')" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \'" (is "?A \'") and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \')" (is "?B \'") proof - have "?A \' \ ?B \'" using assms proof (induction S \ rule: converse_rtrancl_induct2) case (step S1 \1 S2 \2) hence "?A \2 \ ?B \2" using LI_preserves_welltypedness_single by presburger moreover have "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S2 \2" by (fact LI_preserves_wellformedness[OF r_into_rtrancl[OF step.hyps(1)] step.prems(1)]) moreover have "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S2)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S2)" using LI_preserves_tfr_single[OF step.hyps(1)] step.prems by presburger+ moreover have "list_all tfr\<^sub>s\<^sub>t\<^sub>p S2" using LI_preserves_tfr_stp_all_single[OF step.hyps(1)] step.prems by fastforce ultimately show ?case using step.IH by presburger qed simp thus "?A \'" "?B \'" by simp_all qed lemma LI_preserves_tfr: assumes "(S,\) \\<^sup>* (S',\')" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" shows "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S')" (is "?A S'") and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S')" (is "?B S'") and "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" (is "?C S'") proof - have "?A S' \ ?B S' \ ?C S'" using assms proof (induction S \ rule: converse_rtrancl_induct2) case (step S1 \1 S2 \2) have "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S2 \2" "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S2)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S2)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S2" using LI_preserves_wellformedness[OF r_into_rtrancl[OF step.hyps(1)] step.prems(1)] LI_preserves_tfr_single[OF step.hyps(1) step.prems(1,2)] LI_preserves_tfr_stp_all_single[OF step.hyps(1) step.prems(1,2)] step.prems(3,4,5,6) by metis+ moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \2" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \2)" using LI_preserves_welltypedness[OF r_into_rtrancl[OF step.hyps(1)] step.prems] by simp_all ultimately show ?case using step.IH by presburger qed blast thus "?A S'" "?B S'" "?C S'" by simp_all qed end subsubsection \Simple Constraints are Well-typed Satisfiable\ text \Proving the existence of a well-typed interpretation\ context begin -lemma wt_interpretation_exists: +lemma wt_interpretation_exists: obtains \::"('fun,'var) subst" where "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "subst_range \ \ public_ground_wf_terms" proof define \ where "\ = (\x. (SOME t. \ (Var x) = \ t \ public_ground_wf_term t))" { fix x t assume "\ x = t" hence "\ (Var x) = \ t \ public_ground_wf_term t" using someI_ex[of "\t. \ (Var x) = \ t \ public_ground_wf_term t", OF type_pgwt_inhabited[of "Var x"]] unfolding \_def wf\<^sub>t\<^sub>r\<^sub>m_def by simp } hence props: "\ v = t \ \ (Var v) = \ t \ public_ground_wf_term t" for v t by metis - have "\ v \ Var v" for v using props pgwt_ground by force + have "\ v \ Var v" for v using props pgwt_ground by (simp add: empty_fv_not_var) hence "subst_domain \ = UNIV" by auto moreover have "ground (subst_range \)" by (simp add: props pgwt_ground) ultimately show "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" by metis show "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def using props by simp show "subst_range \ \ public_ground_wf_terms" by (auto simp add: props) qed lemma wt_grounding_subst_exists: "\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}" proof - obtain \ where \: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "subst_range \ \ public_ground_wf_terms" using wt_interpretation_exists by blast show ?thesis using pgwt_wellformed interpretation_grounds[OF \(1)] \(2,3) by blast qed private fun fresh_pgwt::"'fun set \ ('fun,'atom) term_type \ ('fun,'var) term" where "fresh_pgwt S (TAtom a) = Fun (SOME c. c \ S \ \ (Fun c []) = TAtom a \ public c) []" | "fresh_pgwt S (TComp f T) = Fun f (map (fresh_pgwt S) T)" private lemma fresh_pgwt_same_type: assumes "finite S" "wf\<^sub>t\<^sub>r\<^sub>m t" shows "\ (fresh_pgwt S (\ t)) = \ t" proof - let ?P = "\\::('fun,'atom) term_type. wf\<^sub>t\<^sub>r\<^sub>m \ \ (\f T. TComp f T \ \ \ 0 < arity f)" { fix \ assume "?P \" hence "\ (fresh_pgwt S \) = \" proof (induction \) case (Var a) let ?P = "\c. c \ S \ \ (Fun c []) = Var a \ public c" let ?Q = "\c. \ (Fun c []) = Var a \ public c" have " {c. ?Q c} - S = {c. ?P c}" by auto hence "infinite {c. ?P c}" - using Diff_infinite_finite[OF assms(1) infinite_typed_consts[of a]] + using Diff_infinite_finite[OF assms(1) infinite_typed_consts[of a]] by metis hence "\c. ?P c" using not_finite_existsD by blast thus ?case using someI_ex[of ?P] by auto next case (Fun f T) have f: "0 < arity f" using Fun.prems fun_type_inv by auto have "\t. t \ set T \ ?P t" using Fun.prems wf_trm_subtermeq term.le_less_trans Fun_param_is_subterm by metis hence "\t. t \ set T \ \ (fresh_pgwt S t) = t" using Fun.prems Fun.IH by auto hence "map \ (map (fresh_pgwt S) T) = T" by (induct T) auto thus ?case using fun_type[OF f] by simp qed } thus ?thesis using assms(1) \_wf'[OF assms(2)] \_wf(1) by auto qed private lemma fresh_pgwt_empty_synth: assumes "finite S" "wf\<^sub>t\<^sub>r\<^sub>m t" shows "{} \\<^sub>c fresh_pgwt S (\ t)" proof - let ?P = "\\::('fun,'atom) term_type. wf\<^sub>t\<^sub>r\<^sub>m \ \ (\f T. TComp f T \ \ \ 0 < arity f)" { fix \ assume "?P \" hence "{} \\<^sub>c fresh_pgwt S \" proof (induction \) case (Var a) let ?P = "\c. c \ S \ \ (Fun c []) = Var a \ public c" let ?Q = "\c. \ (Fun c []) = Var a \ public c" have " {c. ?Q c} - S = {c. ?P c}" by auto hence "infinite {c. ?P c}" - using Diff_infinite_finite[OF assms(1) infinite_typed_consts[of a]] + using Diff_infinite_finite[OF assms(1) infinite_typed_consts[of a]] by metis hence "\c. ?P c" using not_finite_existsD by blast thus ?case using someI_ex[of ?P] intruder_synth.ComposeC[of "[]" _ "{}"] const_type_inv by auto next case (Fun f T) - have f: "0 < arity f" "length T = arity f" "public f" + have f: "0 < arity f" "length T = arity f" "public f" using Fun.prems fun_type_inv unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto have "\t. t \ set T \ ?P t" using Fun.prems wf_trm_subtermeq term.le_less_trans Fun_param_is_subterm by metis hence "\t. t \ set T \ {} \\<^sub>c fresh_pgwt S t" using Fun.prems Fun.IH by auto moreover have "length (map (fresh_pgwt S) T) = arity f" using f(2) by auto ultimately show ?case using intruder_synth.ComposeC[of "map (fresh_pgwt S) T" f] f by auto qed } thus ?thesis using assms(1) \_wf'[OF assms(2)] \_wf(1) by auto qed private lemma fresh_pgwt_has_fresh_const: assumes "finite S" "wf\<^sub>t\<^sub>r\<^sub>m t" obtains c where "Fun c [] \ fresh_pgwt S (\ t)" "c \ S" proof - let ?P = "\\::('fun,'atom) term_type. wf\<^sub>t\<^sub>r\<^sub>m \ \ (\f T. TComp f T \ \ \ 0 < arity f)" { fix \ assume "?P \" hence "\c. Fun c [] \ fresh_pgwt S \ \ c \ S" proof (induction \) case (Var a) let ?P = "\c. c \ S \ \ (Fun c []) = Var a \ public c" let ?Q = "\c. \ (Fun c []) = Var a \ public c" have " {c. ?Q c} - S = {c. ?P c}" by auto hence "infinite {c. ?P c}" - using Diff_infinite_finite[OF assms(1) infinite_typed_consts[of a]] + using Diff_infinite_finite[OF assms(1) infinite_typed_consts[of a]] by metis hence "\c. ?P c" using not_finite_existsD by blast thus ?case using someI_ex[of ?P] by auto next case (Fun f T) have f: "0 < arity f" "length T = arity f" "public f" "T \ []" using Fun.prems fun_type_inv unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto - obtain t' where t': "t' \ set T" by (meson all_not_in_conv f(4) set_empty) + obtain t' where t': "t' \ set T" by (meson all_not_in_conv f(4) set_empty) have "\t. t \ set T \ ?P t" using Fun.prems wf_trm_subtermeq term.le_less_trans Fun_param_is_subterm by metis hence "\t. t \ set T \ \c. Fun c [] \ fresh_pgwt S t \ c \ S" using Fun.prems Fun.IH by auto then obtain c where c: "Fun c [] \ fresh_pgwt S t'" "c \ S" using t' by metis thus ?case using t' by auto qed - } thus ?thesis using that assms \_wf'[OF assms(2)] \_wf(1) by blast + } thus ?thesis using that assms \_wf'[OF assms(2)] \_wf(1) by blast qed private lemma fresh_pgwt_subterm_fresh: assumes "finite S" "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m s" "funs_term s \ S" shows "s \ subterms (fresh_pgwt S (\ t))" proof - let ?P = "\\::('fun,'atom) term_type. wf\<^sub>t\<^sub>r\<^sub>m \ \ (\f T. TComp f T \ \ \ 0 < arity f)" { fix \ assume "?P \" hence "s \ subterms (fresh_pgwt S \)" proof (induction \) case (Var a) let ?P = "\c. c \ S \ \ (Fun c []) = Var a \ public c" let ?Q = "\c. \ (Fun c []) = Var a \ public c" have " {c. ?Q c} - S = {c. ?P c}" by auto hence "infinite {c. ?P c}" - using Diff_infinite_finite[OF assms(1) infinite_typed_consts[of a]] + using Diff_infinite_finite[OF assms(1) infinite_typed_consts[of a]] by metis hence "\c. ?P c" using not_finite_existsD by blast thus ?case using someI_ex[of ?P] assms(4) by auto next case (Fun f T) - have f: "0 < arity f" "length T = arity f" "public f" + have f: "0 < arity f" "length T = arity f" "public f" using Fun.prems fun_type_inv unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto have "\t. t \ set T \ ?P t" using Fun.prems wf_trm_subtermeq term.le_less_trans Fun_param_is_subterm by metis hence "\t. t \ set T \ s \ subterms (fresh_pgwt S t)" using Fun.prems Fun.IH by auto moreover have "s \ fresh_pgwt S (Fun f T)" proof - obtain c where c: "Fun c [] \ fresh_pgwt S (Fun f T)" "c \ S" using fresh_pgwt_has_fresh_const[OF assms(1)] type_wfttype_inhabited Fun.prems by metis hence "\Fun c [] \ s" using assms(4) subtermeq_imp_funs_term_subset by force thus ?thesis using c(1) by auto qed ultimately show ?case by auto qed } thus ?thesis using assms(1) \_wf'[OF assms(2)] \_wf(1) by auto qed private lemma wt_fresh_pgwt_term_exists: assumes "finite T" "wf\<^sub>t\<^sub>r\<^sub>m s" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s T" obtains t where "\ t = \ s" "{} \\<^sub>c t" "\s \ T. \u \ subterms s. u \ subterms t" proof - have finite_S: "finite (\(funs_term ` T))" using assms(1) by auto have 1: "\ (fresh_pgwt (\(funs_term ` T)) (\ s)) = \ s" using fresh_pgwt_same_type[OF finite_S assms(2)] by auto have 2: "{} \\<^sub>c fresh_pgwt (\(funs_term ` T)) (\ s)" using fresh_pgwt_empty_synth[OF finite_S assms(2)] by auto have 3: "\v \ T. \u \ subterms v. u \ subterms (fresh_pgwt (\(funs_term ` T)) (\ s))" using fresh_pgwt_subterm_fresh[OF finite_S assms(2)] assms(3) wf_trm_subtermeq subtermeq_imp_funs_term_subset by force show ?thesis by (rule that[OF 1 2 3]) qed lemma wt_bij_finite_subst_exists: assumes "finite (S::'var set)" "finite (T::('fun,'var) terms)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s T" shows "\\::('fun,'var) subst. subst_domain \ = S \ bij_betw \ (subst_domain \) (subst_range \) \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ {t. {} \\<^sub>c t} - T \ (\s \ subst_range \. \u \ subst_range \. (\v. v \ s \ v \ u) \ s = u) \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using assms proof (induction rule: finite_induct) case empty have "subst_domain Var = {}" "bij_betw Var (subst_domain Var) (subst_range Var)" "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range Var) \ {t. {} \\<^sub>c t} - T" "\s \ subst_range Var. \u \ subst_range Var. (\v. v \ s \ v \ u) \ s = u" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range Var)" unfolding bij_betw_def by auto thus ?case by (force simp add: subst_domain_def) next case (insert x S) then obtain \ where \: "subst_domain \ = S" "bij_betw \ (subst_domain \) (subst_range \)" "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ {t. {} \\<^sub>c t} - T" "\s \ subst_range \. \u \ subst_range \. (\v. v \ s \ v \ u) \ s = u" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (auto simp del: subst_range.simps) have *: "finite (T \ subst_range \)" using insert.prems(1) insert.hyps(1) \(1) by simp have **: "wf\<^sub>t\<^sub>r\<^sub>m (Var x)" by simp have ***: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (T \ subst_range \)" using assms(3) \(6) by blast obtain t where t: "\ t = \ (Var x)" "{} \\<^sub>c t" "\s \ T \ subst_range \. \u \ subterms s. u \ subterms t" using wt_fresh_pgwt_term_exists[OF * ** ***] by auto obtain \ where \: "\ \ \y. if x = y then t else \ y" by simp have t_ground: "fv t = {}" using t(2) pgwt_ground[of t] pgwt_is_empty_synth[of t] by auto hence x_dom: "x \ subst_domain \" "x \ subst_domain \" using insert.hyps(2) \(1) \ by auto moreover have "subst_range \ \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \)" by auto hence ground_imgs: "ground (subst_range \)" using \(3) pgwt_ground pgwt_is_empty_synth by force ultimately have x_img: "\ x \ subst_range \" using ground_subst_dom_iff_img by (auto simp add: subst_domain_def) have "ground (insert t (subst_range \))" using ground_imgs x_dom t_ground by auto have \_dom: "subst_domain \ = insert x (subst_domain \)" using \ t_ground by (auto simp add: subst_domain_def) have \_img: "subst_range \ = insert t (subst_range \)" proof show "subst_range \ \ insert t (subst_range \)" proof fix t' assume "t' \ subst_range \" then obtain y where "y \ subst_domain \" "t' = \ y" by auto thus "t' \ insert t (subst_range \)" using \ by (auto simp add: subst_domain_def) qed show "insert t (subst_range \) \ subst_range \" proof fix t' assume t': "t' \ insert t (subst_range \)" hence "fv t' = {}" using ground_imgs x_img t_ground by auto hence "t' \ Var x" by auto show "t' \ subst_range \" proof (cases "t' = t") case False hence "t' \ subst_range \" using t' by auto then obtain y where "\ y \ subst_range \" "t' = \ y" by auto hence "y \ subst_domain \" "t' \ Var y" using ground_subst_dom_iff_img[OF ground_imgs(1)] by (auto simp add: subst_domain_def simp del: subst_range.simps) hence "x \ y" using x_dom by auto hence "\ y = \ y" unfolding \ by auto thus ?thesis using \t' \ Var y\ \t' = \ y\ subst_imgI[of \ y] by auto qed (metis subst_imgI \ \t' \ Var x\) qed qed hence \_ground_img: "ground (subst_range \)" using ground_imgs t_ground by auto have "subst_domain \ = insert x S" using \_dom \(1) by auto moreover have "bij_betw \ (subst_domain \) (subst_range \)" proof (intro bij_betwI') fix y z assume *: "y \ subst_domain \" "z \ subst_domain \" hence "fv (\ y) = {}" "fv (\ z) = {}" using \_ground_img by auto { assume "\ y = \ z" hence "y = z" proof (cases "\ y \ subst_range \ \ \ z \ subst_range \") case True hence **: "y \ subst_domain \" "z \ subst_domain \" - using \ \_dom True * t(3) by (metis Un_iff term.order_refl insertE)+ + using \ \_dom True * t(3) by (metis Un_iff term.order_refl insertE)+ hence "y \ x" "z \ x" using x_dom by auto hence "\ y = \ y" "\ z = \ z" using \ by auto thus ?thesis using \\ y = \ z\ \(2) ** unfolding bij_betw_def inj_on_def by auto qed (metis \ * \\ y = \ z\ \_dom ground_imgs(1) ground_subst_dom_iff_img insertE) } thus "(\ y = \ z) = (y = z)" by auto next fix y assume "y \ subst_domain \" thus "\ y \ subst_range \" by auto next fix t assume "t \ subst_range \" thus "\z \ subst_domain \. t = \ z" by auto qed moreover have "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ {t. {} \\<^sub>c t} - T" proof - { fix s assume "s \ t" hence "s \ {t. {} \\<^sub>c t} - T" - using t(2,3) + using t(2,3) by (metis Diff_eq_empty_iff Diff_iff Un_upper1 term.order_refl - deduct_synth_subterm mem_Collect_eq) + deduct_synth_subterm mem_Collect_eq) } thus ?thesis using \(3) \ \_img by auto qed moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using \ t(1) \(5) unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by auto moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using \ \(6) t(2) pgwt_is_empty_synth pgwt_wellformed wf_trm_subst_range_iff[of \] wf_trm_subst_range_iff[of \] by metis moreover have "\s\subst_range \. \u\subst_range \. (\v. v \ s \ v \ u) \ s = u" using \(4) \_img t(3) by (auto simp del: subst_range.simps) ultimately show ?case by blast qed private lemma wt_bij_finite_tatom_subst_exists_single: assumes "finite (S::'var set)" "finite (T::('fun,'var) terms)" and "\x. x \ S \ \ (Var x) = TAtom a" shows "\\::('fun,'var) subst. subst_domain \ = S \ bij_betw \ (subst_domain \) (subst_range \) \ subst_range \ \ ((\c. Fun c []) ` {c. \ (Fun c []) = TAtom a \ public c \ arity c = 0}) - T \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" proof - let ?U = "{c. \ (Fun c []) = TAtom a \ public c \ arity c = 0}" obtain \ where \: "subst_domain \ = S" "bij_betw \ (subst_domain \) (subst_range \)" "subst_range \ \ ((\c. Fun c []) ` ?U) - T" using bij_finite_const_subst_exists'[OF assms(1,2) infinite_typed_consts'[of a]] by auto { fix x assume "x \ subst_domain \" hence "\ (Var x) = \ (\ x)" by auto } moreover { fix x assume "x \ subst_domain \" hence "\c \ ?U. \ x = Fun c [] \ arity c = 0" using \ by auto hence "\ (\ x) = TAtom a" "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" using assms(3) const_type wf_trmI[of "[]"] by auto hence "\ (Var x) = \ (\ x)" "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" using assms(3) \(1) by force+ } ultimately have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using wf_trm_subst_range_iff[of \] unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by force+ thus ?thesis using \ by auto qed lemma wt_bij_finite_tatom_subst_exists: assumes "finite (S::'var set)" "finite (T::('fun,'var) terms)" and "\x. x \ S \ \a. \ (Var x) = TAtom a" shows "\\::('fun,'var) subst. subst_domain \ = S \ bij_betw \ (subst_domain \) (subst_range \) \ subst_range \ \ ((\c. Fun c []) ` \\<^sub>p\<^sub>u\<^sub>b) - T \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using assms proof (induction rule: finite_induct) case empty have "subst_domain Var = {}" "bij_betw Var (subst_domain Var) (subst_range Var)" "subst_range Var \ ((\c. Fun c []) ` \\<^sub>p\<^sub>u\<^sub>b) - T" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range Var)" unfolding bij_betw_def by auto thus ?case by (auto simp add: subst_domain_def) next case (insert x S) then obtain a where a: "\ (Var x) = TAtom a" by fastforce from insert obtain \ where \: "subst_domain \ = S" "bij_betw \ (subst_domain \) (subst_range \)" "subst_range \ \ ((\c. Fun c []) ` \\<^sub>p\<^sub>u\<^sub>b) - T" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by auto let ?S' = "{y \ S. \ (Var y) = TAtom a}" let ?T' = "T \ subst_range \" have *: "finite (insert x ?S')" using insert by simp have **: "finite ?T'" using insert.prems(1) insert.hyps(1) \(1) by simp have ***: "\y. y \ insert x ?S' \ \ (Var y) = TAtom a" using a by auto obtain \ where \: "subst_domain \ = insert x ?S'" "bij_betw \ (subst_domain \) (subst_range \)" "subst_range \ \ ((\c. Fun c []) ` \\<^sub>p\<^sub>u\<^sub>b) - ?T'" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using wt_bij_finite_tatom_subst_exists_single[OF * ** ***] const_type_inv[of _ "[]" a] by blast obtain \ where \: "\ \ \y. if x = y then \ y else \ y" by simp have x_dom: "x \ subst_domain \" "x \ subst_domain \" "x \ subst_domain \" using insert.hyps(2) \(1) \(1) \ by (auto simp add: subst_domain_def) moreover have ground_imgs: "ground (subst_range \)" "ground (subst_range \)" using pgwt_ground \(3) \(3) by auto ultimately have x_img: "\ x \ subst_range \" "\ x \ subst_range \" using ground_subst_dom_iff_img by (auto simp add: subst_domain_def) have "ground (insert (\ x) (subst_range \))" using ground_imgs x_dom by auto have \_dom: "subst_domain \ = insert x (subst_domain \)" using \(1) \ by (auto simp add: subst_domain_def) have \_img: "subst_range \ = insert (\ x) (subst_range \)" proof show "subst_range \ \ insert (\ x) (subst_range \)" proof fix t assume "t \ subst_range \" then obtain y where "y \ subst_domain \" "t = \ y" by auto thus "t \ insert (\ x) (subst_range \)" using \ by (auto simp add: subst_domain_def) qed show "insert (\ x) (subst_range \) \ subst_range \" proof fix t assume t: "t \ insert (\ x) (subst_range \)" hence "fv t = {}" using ground_imgs x_img(2) by auto hence "t \ Var x" by auto show "t \ subst_range \" proof (cases "t = \ x") case True thus ?thesis using subst_imgI \ \t \ Var x\ by metis next case False hence "t \ subst_range \" using t by auto then obtain y where "\ y \ subst_range \" "t = \ y" by auto hence "y \ subst_domain \" "t \ Var y" using ground_subst_dom_iff_img[OF ground_imgs(1)] by (auto simp add: subst_domain_def simp del: subst_range.simps) hence "x \ y" using x_dom by auto hence "\ y = \ y" unfolding \ by auto thus ?thesis using \t \ Var y\ \t = \ y\ subst_imgI[of \ y] by auto qed qed qed hence \_ground_img: "ground (subst_range \)" using ground_imgs x_img by auto have "subst_domain \ = insert x S" using \_dom \(1) by auto moreover have "bij_betw \ (subst_domain \) (subst_range \)" proof (intro bij_betwI') fix y z assume *: "y \ subst_domain \" "z \ subst_domain \" hence "fv (\ y) = {}" "fv (\ z) = {}" using \_ground_img by auto { assume "\ y = \ z" hence "y = z" proof (cases "\ y \ subst_range \ \ \ z \ subst_range \") case True hence **: "y \ subst_domain \" "z \ subst_domain \" using \ \_dom x_img(2) \(3) True by (metis (no_types) *(1) DiffE Un_upper2 insertE subsetCE, metis (no_types) *(2) DiffE Un_upper2 insertE subsetCE) hence "y \ x" "z \ x" using x_dom by auto hence "\ y = \ y" "\ z = \ z" using \ by auto thus ?thesis using \\ y = \ z\ \(2) ** unfolding bij_betw_def inj_on_def by auto qed (metis \ * \\ y = \ z\ \_dom ground_imgs(1) ground_subst_dom_iff_img insertE) } thus "(\ y = \ z) = (y = z)" by auto next fix y assume "y \ subst_domain \" thus "\ y \ subst_range \" by auto next fix t assume "t \ subst_range \" thus "\z \ subst_domain \. t = \ z" by auto qed moreover have "subst_range \ \ (\c. Fun c []) ` \\<^sub>p\<^sub>u\<^sub>b - T" using \(3) \(3) \ by (auto simp add: subst_domain_def) moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using \(4) \(4) \ unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by auto moreover have "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using \ \(5) \(5) wf_trm_subst_range_iff[of \] wf_trm_subst_range_iff[of \] wf_trm_subst_range_iff[of \] by presburger ultimately show ?case by blast qed theorem wt_sat_if_simple: assumes "simple S" "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" and \': "\X F. Inequality X F \ set S \ ineq_model \' X F" "ground (subst_range \')" "subst_domain \' = {x \ vars\<^sub>s\<^sub>t S. \X F. Inequality X F \ set S \ x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X}" and tfr_stp_all: "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" shows "\\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ (\ \\<^sub>c \S, \\) \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" proof - from \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \\ have "wf\<^sub>s\<^sub>t {} S" "subst_idem \" and S_\_disj: "\v \ vars\<^sub>s\<^sub>t S. \ v = Var v" using subst_idemI[of \] unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by force+ - + obtain \::"('fun,'var) subst" where \: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "subst_range \ \ public_ground_wf_terms" using wt_interpretation_exists by blast hence \_deduct: "\x M. M \\<^sub>c \ x" and \_wf_trm: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using pgwt_deducible pgwt_wellformed by fastforce+ let ?P = "\\ X. subst_domain \ = set X \ ground (subst_range \)" let ?Sineqsvars = "{x \ vars\<^sub>s\<^sub>t S. \X F. Inequality X F \ set S \ x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ x \ set X}" let ?Strms = "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" have finite_vars: "finite ?Sineqsvars" "finite ?Strms" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s ?Strms" using wf_trm_subtermeq assms(5) by fastforce+ define Q1 where "Q1 = (\(F::(('fun,'var) term \ ('fun,'var) term) list) X. \x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X. \a. \ (Var x) = TAtom a)" define Q2 where "Q2 = (\(F::(('fun,'var) term \ ('fun,'var) term) list) X. \f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ T = [] \ (\s \ set T. s \ Var ` set X))" define Q1' where "Q1' = (\(t::('fun,'var) term) (t'::('fun,'var) term) X. \x \ (fv t \ fv t') - set X. \a. \ (Var x) = TAtom a)" define Q2' where "Q2' = (\(t::('fun,'var) term) (t'::('fun,'var) term) X. \f T. Fun f T \ subterms t \ subterms t' \ T = [] \ (\s \ set T. s \ Var ` set X))" have ex_P: "\X. \\. ?P \ X" using interpretation_subst_exists' by blast have tfr_ineq: "\X F. Inequality X F \ set S \ Q1 F X \ Q2 F X" using tfr_stp_all Q1_def Q2_def tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def[of S] by blast have S_fv_bvars_disj: "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" using \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \\ unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def by metis hence ineqs_vars_not_bound: "\X F x. Inequality X F \ set S \ x \ ?Sineqsvars \ x \ set X" using strand_fv_bvars_disjoint_unfold by blast have \_vars_S_bvars_disj: "(subst_domain \ \ range_vars \) \ set X = {}" when "Inequality X F \ set S" for F X using wf_constr_bvars_disj[OF \wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \\] strand_fv_bvars_disjointD(1)[OF S_fv_bvars_disj that] by blast obtain \::"('fun,'var) subst" where \_fv_dom: "subst_domain \ = ?Sineqsvars" and \_subterm_inj: "subterm_inj_on \ (subst_domain \)" and \_fresh_pub_img: "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ {t. {} \\<^sub>c t} - ?Strms" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and \_wf_trm: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using wt_bij_finite_subst_exists[OF finite_vars] subst_inj_on_is_bij_betw subterm_inj_on_alt_def' by moura have \_bij_dom_img: "bij_betw \ (subst_domain \) (subst_range \)" by (metis \_subterm_inj subst_inj_on_is_bij_betw subterm_inj_on_alt_def) have "finite (subst_domain \)" by(metis \_fv_dom finite_vars(1)) - hence \_finite_img: "finite (subst_range \)" using \_bij_dom_img bij_betw_finite by blast - + hence \_finite_img: "finite (subst_range \)" using \_bij_dom_img bij_betw_finite by blast + have \_img_subterms: "\s \ subst_range \. \u \ subst_range \. (\v. v \ s \ v \ u) \ s = u" by (metis \_subterm_inj subterm_inj_on_alt_def') have "subst_range \ \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \)" by auto hence "subst_range \ \ public_ground_wf_terms - ?Strms" and \_pgwt_img: "subst_range \ \ public_ground_wf_terms" "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ public_ground_wf_terms" using \_fresh_pub_img pgwt_is_empty_synth by blast+ have \_img_ground: "ground (subst_range \)" using \_pgwt_img pgwt_ground by auto hence \_inj: "inj \" using \_bij_dom_img subst_inj_is_bij_betw_dom_img_if_ground_img by auto have \_ineqs_fv_dom: "\X F. Inequality X F \ set S \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ subst_domain \" using \_fv_dom by fastforce have \_dom_bvars_disj: "\X F. Inequality X F \ set S \ subst_domain \ \ set X = {}" using ineqs_vars_not_bound \_fv_dom by fastforce - + have \'1: "\X F \. Inequality X F \ set S \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ subst_domain \'" using \'(3) ineqs_vars_not_bound by fastforce - + have \'2: "\X F. Inequality X F \ set S \ subst_domain \' \ set X = {}" using \'(3) ineqs_vars_not_bound by blast - + have doms_eq: "subst_domain \' = subst_domain \" using \'(3) \_fv_dom by simp have \_ineqs_neq: "ineq_model \ X F" when "Inequality X F \ set S" for X F proof - obtain a::"'fun" where a: "a \ \(funs_term ` subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \))" using exists_fun_notin_funs_terms[OF subterms_union_finite[OF \_finite_img]] by moura hence a': "\T. Fun a T \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \)" "\S. Fun a [] \ set (Fun a []#S)" "Fun a [] \ Var ` set X" by (meson a UN_I term.set_intros(1), auto) define t where "t \ Fun a (Fun a []#map fst F)" define t' where "t' \ Fun a (Fun a []#map snd F)" note F_in = that have t_fv: "fv t \ fv t' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" unfolding t_def t'_def by force have t_subterms: "subterms t \ subterms t' \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ {t, t', Fun a []}" unfolding t_def t'_def by force have "t \ \ \ \ \ t' \ \ \ \" when "?P \ X" for \ proof - have tfr_assms: "Q1 F X \ Q2 F X" using tfr_ineq F_in by metis - + have "Q1 F X \ \x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X. \c. \ x = Fun c []" proof fix x assume "Q1 F X" and x: "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X" then obtain a where "\ (Var x) = TAtom a" unfolding Q1_def by moura hence a: "\ (\ x) = TAtom a" using \_wt unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by simp - + have "x \ subst_domain \" using \_ineqs_fv_dom x F_in by auto then obtain f T where fT: "\ x = Fun f T" by (meson \_img_ground ground_img_obtain_fun) hence "T = []" using \_wf_trm a TAtom_term_cases by fastforce thus "\c. \ x = Fun c []" using fT by metis qed hence 1: "Q1 F X \ \x \ (fv t \ fv t') - set X. \c. \ x = Fun c []" using t_fv by auto - + have 2: "\Q1 F X \ Q2 F X" by (metis tfr_assms) - + have 3: "subst_domain \ \ set X = {}" using \_dom_bvars_disj F_in by auto have 4: "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ (subterms t \ subterms t') = {}" proof - define M1 where "M1 \ {t, t', Fun a []}" define M2 where "M2 \ ?Strms" have "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ M2" using F_in unfolding M2_def by force moreover have "subterms t \ subterms t' \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) \ M1" using t_subterms unfolding M1_def by blast ultimately have *: "subterms t \ subterms t' \ M2 \ M1" by auto have "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ M1 = {}" "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ M2 = {}" using a' \_fresh_pub_img unfolding t_def t'_def M1_def M2_def by blast+ thus ?thesis using * by blast qed - + have 5: "(fv t \ fv t') - subst_domain \ \ set X" using \_ineqs_fv_dom[OF F_in] t_fv by auto - + have 6: "\\. ?P \ X \ t \ \ \ \' \ t' \ \ \ \'" by (metis t_def t'_def \'(1) F_in ineq_model_singleE ineq_model_single_iff) - + have 7: "fv t \ fv t' - set X \ subst_domain \'" using \'1 F_in t_fv by force - + have 8: "subst_domain \' \ set X = {}" using \'2 F_in by auto have 9: "Q1' t t' X" when "Q1 F X" using that t_fv unfolding Q1_def Q1'_def t_def t'_def by blast have 10: "Q2' t t' X" when "Q2 F X" unfolding Q2'_def proof (intro allI impI) fix f T assume "Fun f T \ subterms t \ subterms t'" moreover { assume "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F)" hence "T = [] \ (\s\set T. s \ Var ` set X)" by (metis Q2_def that) } moreover { assume "Fun f T = t" hence "T = [] \ (\s\set T. s \ Var ` set X)" unfolding t_def using a'(2,3) by simp } moreover { assume "Fun f T = t'" hence "T = [] \ (\s\set T. s \ Var ` set X)" unfolding t'_def using a'(2,3) by simp } moreover { assume "Fun f T = Fun a []" hence "T = [] \ (\s\set T. s \ Var ` set X)" by simp } ultimately show "T = [] \ (\s\set T. s \ Var ` set X)" using t_subterms by blast qed note 11 = \_subterm_inj \_img_ground 3 4 5 - + note 12 = 6 7 8 \'(2) doms_eq - + show "t \ \ \ \ \ t' \ \ \ \" - using 1 2 9 10 that sat_ineq_subterm_inj_subst[OF 11 _ 12] + using 1 2 9 10 that sat_ineq_subterm_inj_subst[OF 11 _ 12] unfolding Q1'_def Q2'_def by metis qed thus ?thesis by (metis t_def t'_def ineq_model_singleI ineq_model_single_iff) qed have \_ineqs_fv_dom': "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \ subst_domain \" when "Inequality X F \ set S" and "?P \ X" for F \ X using \_ineqs_fv_dom[OF that(1)] proof (induction F) case (Cons g G) obtain t t' where g: "g = (t,t')" by (metis surj_pair) hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (g#G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) = fv (t \ \) \ fv (t' \ \) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (g#G) = fv t \ fv t' \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G" by (simp_all add: subst_apply_pairs_def) moreover have "fv (t \ \) = fv t - subst_domain \" "fv (t' \ \) = fv t' - subst_domain \" using g that(2) by (simp_all add: subst_fv_unfold_ground_img range_vars_alt_def) moreover have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \ subst_domain \" using Cons by auto ultimately show ?case using Cons.prems that(2) by auto qed (simp add: subst_apply_pairs_def) have \_ineqs_ground: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ((F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) = {}" when "Inequality X F \ set S" and "?P \ X" for F \ X using \_ineqs_fv_dom'[OF that] proof (induction F) case (Cons g G) obtain t t' where g: "g = (t,t')" by (metis surj_pair) hence "fv (t \ \) \ subst_domain \" "fv (t' \ \) \ subst_domain \" using Cons.prems by (auto simp add: subst_apply_pairs_def) hence "fv (t \ \ \ \) = {}" "fv (t' \ \ \ \) = {}" using subst_fv_dom_ground_if_ground_img[OF _ \_img_ground] by metis+ thus ?case using g Cons by (auto simp add: subst_apply_pairs_def) qed (simp add: subst_apply_pairs_def) - + from \_pgwt_img \_ineqs_neq have \_deduct: "M \\<^sub>c \ x" when "x \ subst_domain \" for x M using that pgwt_deducible by fastforce { fix M::"('fun,'var) terms" have "\M; S\\<^sub>c (\ \\<^sub>s \ \\<^sub>s \)" using \wf\<^sub>s\<^sub>t {} S\ \simple S\ S_\_disj \_ineqs_neq \_ineqs_fv_dom' \_vars_S_bvars_disj proof (induction S arbitrary: M rule: wf\<^sub>s\<^sub>t_simple_induct) case (ConsSnd v S) hence S_sat: "\M; S\\<^sub>c (\ \\<^sub>s \ \\<^sub>s \)" and "\ v = Var v" by auto hence "\M. M \\<^sub>c Var v \ (\ \\<^sub>s \ \\<^sub>s \)" using \_deduct \_deduct by (metis ideduct_synth_subst_apply subst_apply_term.simps(1) - subst_subst_compose trm_subst_ident') + subst_subst_compose trm_subst_ident') thus ?case using strand_sem_append(1)[OF S_sat] by (metis strand_sem_c.simps(1,2)) next case (ConsIneq X F S) have dom_disj: "subst_domain \ \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F = {}" using ConsIneq.prems(1) subst_dom_vars_in_subst by force hence *: "F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \ = F" by blast have **: "ineq_model \ X F" by (meson ConsIneq.prems(2) in_set_conv_decomp) have "\x. x \ vars\<^sub>s\<^sub>t S \ x \ vars\<^sub>s\<^sub>t (S@[Inequality X F])" "\x. x \ set S \ x \ set (S@[Inequality X F])" by auto hence IH: "\M; S\\<^sub>c (\ \\<^sub>s \ \\<^sub>s \)" by (metis ConsIneq.IH ConsIneq.prems(1,2,3,4)) have "ineq_model (\ \\<^sub>s \) X F" proof - have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \ subst_domain \" when "?P \ X" for \ using ConsIneq.prems(3)[OF _ that] by simp hence "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ subst_domain \" using fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_subset ex_P by (metis Diff_subset_conv Un_commute) thus ?thesis by (metis ineq_model_ground_subst[OF _ \_img_ground **]) qed hence "ineq_model (\ \\<^sub>s \ \\<^sub>s \) X F" using * ineq_model_subst' subst_compose_assoc ConsIneq.prems(4) by (metis UnCI list.set_intros(1) set_append) thus ?case using IH by (auto simp add: ineq_model_def) qed auto } moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \ \\<^sub>s \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \ \\<^sub>s \))" by (metis wt_subst_compose \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\, metis assms(4) \_wf_trm \_wf_trm wf_trm_subst subst_img_comp_subset') ultimately show ?thesis using interpretation_comp(1)[OF \interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\, of "\ \\<^sub>s \"] subst_idem_support[OF \subst_idem \\, of "\ \\<^sub>s \"] subst_compose_assoc unfolding constr_sem_c_def by metis qed end subsubsection \Theorem: Type-flaw resistant constraints are well-typed satisfiable (composition-only)\ text \ There exists well-typed models of satisfiable type-flaw resistant constraints in the semantics where the intruder is limited to composition only (i.e., he cannot perform decomposition/analysis of deducible messages). \ theorem wt_attack_if_tfr_attack: assumes "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "\ \\<^sub>c \S, \\" and "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" and "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "tfr\<^sub>s\<^sub>t S" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" obtains \\<^sub>\ where "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" and "\\<^sub>\ \\<^sub>c \S, \\" and "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" proof - have tfr: "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using assms(5,6) unfolding tfr\<^sub>s\<^sub>t_def by metis+ obtain S' \' where *: "simple S'" "(S,\) \\<^sup>* (S',\')" "\{}; S'\\<^sub>c \" using LI_completeness[OF assms(3,2)] unfolding constr_sem_c_def by (meson term.order_refl) - have **: "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S' \'" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \'" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S')" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \')" + have **: "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S' \'" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \'" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S')" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \')" using LI_preserves_welltypedness[OF *(2) assms(3,4,7) tfr] LI_preserves_wellformedness[OF *(2) assms(3)] LI_preserves_tfr[OF *(2) assms(3,4,7) tfr] by metis+ define A where "A \ {x \ vars\<^sub>s\<^sub>t S'. \X F. Inequality X F \ set S' \ x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ x \ set X}" define B where "B \ UNIV - A" let ?\ = "rm_vars B \" have gr\: "ground (subst_range \)" "ground (subst_range ?\)" using assms(1) rm_vars_img_subset[of B \] by (auto simp add: subst_domain_def) { fix X F assume "Inequality X F \ set S'" hence *: "ineq_model \ X F" using strand_sem_c_imp_ineq_model[OF *(3)] by (auto simp del: subst_range.simps) hence "ineq_model ?\ X F" proof - { fix \ assume 1: "subst_domain \ = set X" "ground (subst_range \)" and 2: "list_ex (\f. fst f \ \ \\<^sub>s \ \ snd f \ \ \\<^sub>s \) F" have "list_ex (\f. fst f \ \ \\<^sub>s rm_vars B \ \ snd f \ \ \\<^sub>s rm_vars B \) F" using 2 proof (induction F) case (Cons g G) obtain t t' where g: "g = (t,t')" by (metis surj_pair) thus ?case using Cons Unifier_ground_rm_vars[OF gr\(1), of "t \ \" B "t' \ \"] by auto qed simp } thus ?thesis using * unfolding ineq_model_def by simp qed } moreover have "subst_domain \ = UNIV" using assms(1) by metis hence "subst_domain ?\ = A" using rm_vars_dom[of B \] B_def by blast ultimately obtain \\<^sub>\ where "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "\\<^sub>\ \\<^sub>c \S', \'\" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" using wt_sat_if_simple[OF *(1) **(1,2,5,4) _ gr\(2) _ **(3)] A_def by (auto simp del: subst_range.simps) thus ?thesis using that LI_soundness[OF assms(3) *(2)] by metis qed text \ Contra-positive version: if a type-flaw resistant constraint does not have a well-typed model then it is unsatisfiable \ corollary secure_if_wt_secure: assumes "\(\\\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ (\\<^sub>\ \\<^sub>c \S, \\) \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\)" and "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r S \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "tfr\<^sub>s\<^sub>t S" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "\(\\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ (\ \\<^sub>c \S, \\))" using wt_attack_if_tfr_attack[OF _ _ assms(2,3,4,5,6)] assms(1) by metis end subsection \Lifting the Composition-Only Typing Result to the Full Intruder Model\ context typed_model begin subsubsection \Analysis Invariance\ definition (in typed_model) Ana_invar_subst where "Ana_invar_subst \ \ (\f T K M \. Fun f T \ (subterms\<^sub>s\<^sub>e\<^sub>t \) \ Ana (Fun f T) = (K, M) \ Ana (Fun f T \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \))" lemma (in typed_model) Ana_invar_subst_subset: assumes "Ana_invar_subst M" "N \ M" shows "Ana_invar_subst N" using assms unfolding Ana_invar_subst_def by blast lemma (in typed_model) Ana_invar_substD: assumes "Ana_invar_subst \" and "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t \" "Ana (Fun f T) = (K, M)" shows "Ana (Fun f T \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using assms Ana_invar_subst_def by blast end subsubsection \Preliminary Definitions\ text \Strands extended with "decomposition steps"\ datatype (funs\<^sub>e\<^sub>s\<^sub>t\<^sub>p: 'a, vars\<^sub>e\<^sub>s\<^sub>t\<^sub>p: 'b) extstrand_step = Step "('a,'b) strand_step" | Decomp "('a,'b) term" context typed_model begin context begin private fun trms\<^sub>e\<^sub>s\<^sub>t\<^sub>p where "trms\<^sub>e\<^sub>s\<^sub>t\<^sub>p (Step x) = trms\<^sub>s\<^sub>t\<^sub>p x" | "trms\<^sub>e\<^sub>s\<^sub>t\<^sub>p (Decomp t) = {t}" private abbreviation trms\<^sub>e\<^sub>s\<^sub>t where "trms\<^sub>e\<^sub>s\<^sub>t S \ \(trms\<^sub>e\<^sub>s\<^sub>t\<^sub>p ` set S)" private type_synonym ('a,'b) extstrand = "('a,'b) extstrand_step list" private type_synonym ('a,'b) extstrands = "('a,'b) extstrand set" private definition decomp::"('fun,'var) term \ ('fun,'var) strand" where "decomp t \ (case (Ana t) of (K,T) \ send\t\\<^sub>s\<^sub>t#map Send K@map Receive T)" private fun to_st where "to_st [] = []" | "to_st (Step x#S) = x#(to_st S)" | "to_st (Decomp t#S) = (decomp t)@(to_st S)" private fun to_est where "to_est [] = []" | "to_est (x#S) = Step x#to_est S" private abbreviation "ik\<^sub>e\<^sub>s\<^sub>t A \ ik\<^sub>s\<^sub>t (to_st A)" private abbreviation "wf\<^sub>e\<^sub>s\<^sub>t V A \ wf\<^sub>s\<^sub>t V (to_st A)" private abbreviation "assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>s\<^sub>t (to_st A)" private abbreviation "vars\<^sub>e\<^sub>s\<^sub>t A \ vars\<^sub>s\<^sub>t (to_st A)" private abbreviation "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t A \ wfrestrictedvars\<^sub>s\<^sub>t (to_st A)" private abbreviation "bvars\<^sub>e\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>t (to_st A)" private abbreviation "fv\<^sub>e\<^sub>s\<^sub>t A \ fv\<^sub>s\<^sub>t (to_st A)" private abbreviation "funs\<^sub>e\<^sub>s\<^sub>t A \ funs\<^sub>s\<^sub>t (to_st A)" private definition wf\<^sub>s\<^sub>t\<^sub>s'::"('fun,'var) strands \ ('fun,'var) extstrand \ bool" where "wf\<^sub>s\<^sub>t\<^sub>s' \ \ \ (\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \) (dual\<^sub>s\<^sub>t S)) \ (\S \ \. \S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}) \ (\S \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t \ = {}) \ (\S \ \. fv\<^sub>s\<^sub>t (to_st \) \ bvars\<^sub>s\<^sub>t S = {})" private definition wf\<^sub>s\<^sub>t\<^sub>s::"('fun,'var) strands \ bool" where "wf\<^sub>s\<^sub>t\<^sub>s \ \ (\S \ \. wf\<^sub>s\<^sub>t {} (dual\<^sub>s\<^sub>t S)) \ (\S \ \. \S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {})" private inductive well_analyzed::"('fun,'var) extstrand \ bool" where Nil[simp]: "well_analyzed []" | Step: "well_analyzed A \ well_analyzed (A@[Step x])" | Decomp: "\well_analyzed A; t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) - (Var ` \)\ \ well_analyzed (A@[Decomp t])" private fun subst_apply_extstrandstep (infix "\\<^sub>e\<^sub>s\<^sub>t\<^sub>p" 51) where "subst_apply_extstrandstep (Step x) \ = Step (x \\<^sub>s\<^sub>t\<^sub>p \)" | "subst_apply_extstrandstep (Decomp t) \ = Decomp (t \ \)" private lemma subst_apply_extstrandstep'_simps[simp]: "(Step (send\t\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \ = Step (send\t \ \\\<^sub>s\<^sub>t)" "(Step (receive\t\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \ = Step (receive\t \ \\\<^sub>s\<^sub>t)" "(Step (\a: t \ t'\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \ = Step (\a: (t \ \) \ (t' \ \)\\<^sub>s\<^sub>t)" "(Step (\X\\\: F\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \ = Step (\X\\\: (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)\\<^sub>s\<^sub>t)" by simp_all private lemma vars\<^sub>e\<^sub>s\<^sub>t\<^sub>p_subst_apply_simps[simp]: "vars\<^sub>e\<^sub>s\<^sub>t\<^sub>p ((Step (send\t\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \)" "vars\<^sub>e\<^sub>s\<^sub>t\<^sub>p ((Step (receive\t\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \)" "vars\<^sub>e\<^sub>s\<^sub>t\<^sub>p ((Step (\a: t \ t'\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \) = fv (t \ \) \ fv (t' \ \)" "vars\<^sub>e\<^sub>s\<^sub>t\<^sub>p ((Step (\X\\\: F\\<^sub>s\<^sub>t)) \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \) = set X \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s rm_vars (set X) \)" by auto private definition subst_apply_extstrand (infix "\\<^sub>e\<^sub>s\<^sub>t" 51) where "S \\<^sub>e\<^sub>s\<^sub>t \ \ map (\x. x \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \) S" private abbreviation update\<^sub>s\<^sub>t::"('fun,'var) strands \ ('fun,'var) strand \ ('fun,'var) strands" where "update\<^sub>s\<^sub>t \ S \ (case S of Nil \ \ - {S} | Cons _ S' \ insert S' (\ - {S}))" private inductive_set decomps\<^sub>e\<^sub>s\<^sub>t:: "('fun,'var) terms \ ('fun,'var) terms \ ('fun,'var) subst \ ('fun,'var) extstrands" (* \: intruder knowledge \: additional messages *) for \ and \ and \ where Nil: "[] \ decomps\<^sub>e\<^sub>s\<^sub>t \ \ \" | Decomp: "\\ \ decomps\<^sub>e\<^sub>s\<^sub>t \ \ \; Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (\ \ \); Ana (Fun f T) = (K,M); M \ []; (\ \ ik\<^sub>e\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c Fun f T \ \; \k. k \ set K \ (\ \ ik\<^sub>e\<^sub>s\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \\ \ \@[Decomp (Fun f T)] \ decomps\<^sub>e\<^sub>s\<^sub>t \ \ \" private fun decomp_rm\<^sub>e\<^sub>s\<^sub>t::"('fun,'var) extstrand \ ('fun,'var) extstrand" where "decomp_rm\<^sub>e\<^sub>s\<^sub>t [] = []" | "decomp_rm\<^sub>e\<^sub>s\<^sub>t (Decomp t#S) = decomp_rm\<^sub>e\<^sub>s\<^sub>t S" | "decomp_rm\<^sub>e\<^sub>s\<^sub>t (Step x#S) = Step x#(decomp_rm\<^sub>e\<^sub>s\<^sub>t S)" private inductive sem\<^sub>e\<^sub>s\<^sub>t_d::"('fun,'var) terms \ ('fun,'var) subst \ ('fun,'var) extstrand \ bool" where Nil[simp]: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ []" | Send: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ S \ (ik\<^sub>e\<^sub>s\<^sub>t S \ M\<^sub>0) \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \ \ sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ (S@[Step (send\t\\<^sub>s\<^sub>t)])" | Receive: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ S \ sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ (S@[Step (receive\t\\<^sub>s\<^sub>t)])" | Equality: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ S \ t \ \ = t' \ \ \ sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ (S@[Step (\a: t \ t'\\<^sub>s\<^sub>t)])" | Inequality: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ S \ ineq_model \ X F \ sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ (S@[Step (\X\\\: F\\<^sub>s\<^sub>t)])" | Decompose: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ S \ (ik\<^sub>e\<^sub>s\<^sub>t S \ M\<^sub>0) \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \ \ Ana t = (K, M) \ (\k. k \ set K \ (ik\<^sub>e\<^sub>s\<^sub>t S \ M\<^sub>0) \\<^sub>s\<^sub>e\<^sub>t \ \ k \ \) \ sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ (S@[Decomp t])" private inductive sem\<^sub>e\<^sub>s\<^sub>t_c::"('fun,'var) terms \ ('fun,'var) subst \ ('fun,'var) extstrand \ bool" where Nil[simp]: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ []" | Send: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ S \ (ik\<^sub>e\<^sub>s\<^sub>t S \ M\<^sub>0) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \ \ sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ (S@[Step (send\t\\<^sub>s\<^sub>t)])" | Receive: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ S \ sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ (S@[Step (receive\t\\<^sub>s\<^sub>t)])" | Equality: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ S \ t \ \ = t' \ \ \ sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ (S@[Step (\a: t \ t'\\<^sub>s\<^sub>t)])" | Inequality: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ S \ ineq_model \ X F \ sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ (S@[Step (\X\\\: F\\<^sub>s\<^sub>t)])" | Decompose: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ S \ (ik\<^sub>e\<^sub>s\<^sub>t S \ M\<^sub>0) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \ \ Ana t = (K, M) \ (\k. k \ set K \ (ik\<^sub>e\<^sub>s\<^sub>t S \ M\<^sub>0) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \) \ sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ (S@[Decomp t])" subsubsection \Preliminary Lemmata\ private lemma wf\<^sub>s\<^sub>t\<^sub>s_wf\<^sub>s\<^sub>t\<^sub>s': "wf\<^sub>s\<^sub>t\<^sub>s \ = wf\<^sub>s\<^sub>t\<^sub>s' \ []" by (simp add: wf\<^sub>s\<^sub>t\<^sub>s_def wf\<^sub>s\<^sub>t\<^sub>s'_def) private lemma decomp_ik: assumes "Ana t = (K,M)" shows "ik\<^sub>s\<^sub>t (decomp t) = set M" using ik_rcv_map[of _ M] ik_rcv_map'[of _ M] by (auto simp add: decomp_def inv_def assms) private lemma decomp_assignment_rhs_empty: assumes "Ana t = (K,M)" shows "assignment_rhs\<^sub>s\<^sub>t (decomp t) = {}" by (auto simp add: decomp_def inv_def assms) private lemma decomp_tfr\<^sub>s\<^sub>t\<^sub>p: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (decomp t)" by (auto simp add: decomp_def list_all_def) private lemma trms\<^sub>e\<^sub>s\<^sub>t_ikI: "t \ ik\<^sub>e\<^sub>s\<^sub>t A \ t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>e\<^sub>s\<^sub>t A)" proof (induction A rule: to_st.induct) case (2 x S) thus ?case by (cases x) auto next case (3 t' A) obtain K M where Ana: "Ana t' = (K,M)" by (metis surj_pair) show ?case using 3 decomp_ik[OF Ana] Ana_subterm[OF Ana] by auto qed simp private lemma trms\<^sub>e\<^sub>s\<^sub>t_ik_assignment_rhsI: "t \ ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \ t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>e\<^sub>s\<^sub>t A)" proof (induction A rule: to_st.induct) case (2 x S) thus ?case proof (cases x) case (Equality ac t t') thus ?thesis using 2 by (cases ac) auto qed auto next case (3 t' A) obtain K M where Ana: "Ana t' = (K,M)" by (metis surj_pair) show ?case using 3 decomp_ik[OF Ana] decomp_assignment_rhs_empty[OF Ana] Ana_subterm[OF Ana] by auto qed simp private lemma trms\<^sub>e\<^sub>s\<^sub>t_ik_subtermsI: assumes "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A)" shows "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>e\<^sub>s\<^sub>t A)" proof - obtain t' where "t' \ ik\<^sub>e\<^sub>s\<^sub>t A" "t \ t'" using trms\<^sub>e\<^sub>s\<^sub>t_ikI assms by auto thus ?thesis by (meson contra_subsetD in_subterms_subset_Union trms\<^sub>e\<^sub>s\<^sub>t_ikI) qed private lemma trms\<^sub>e\<^sub>s\<^sub>tD: assumes "t \ trms\<^sub>e\<^sub>s\<^sub>t A" shows "t \ trms\<^sub>s\<^sub>t (to_st A)" using assms proof (induction A) case (Cons a A) obtain K M where Ana: "Ana t = (K,M)" by (metis surj_pair) hence "t \ trms\<^sub>s\<^sub>t (decomp t)" unfolding decomp_def by force thus ?case using Cons.IH Cons.prems by (cases a) auto qed simp private lemma subst_apply_extstrand_nil[simp]: "[] \\<^sub>e\<^sub>s\<^sub>t \ = []" by (simp add: subst_apply_extstrand_def) private lemma subst_apply_extstrand_singleton[simp]: "[Step (receive\t\\<^sub>s\<^sub>t)] \\<^sub>e\<^sub>s\<^sub>t \ = [Step (Receive (t \ \))]" "[Step (send\t\\<^sub>s\<^sub>t)] \\<^sub>e\<^sub>s\<^sub>t \ = [Step (Send (t \ \))]" "[Step (\a: t \ t'\\<^sub>s\<^sub>t)] \\<^sub>e\<^sub>s\<^sub>t \ = [Step (Equality a (t \ \) (t' \ \))]" "[Decomp t] \\<^sub>e\<^sub>s\<^sub>t \ = [Decomp (t \ \)]" unfolding subst_apply_extstrand_def by auto private lemma extstrand_subst_hom: "(S@S') \\<^sub>e\<^sub>s\<^sub>t \ = (S \\<^sub>e\<^sub>s\<^sub>t \)@(S' \\<^sub>e\<^sub>s\<^sub>t \)" "(x#S) \\<^sub>e\<^sub>s\<^sub>t \ = (x \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \)#(S \\<^sub>e\<^sub>s\<^sub>t \)" unfolding subst_apply_extstrand_def by auto private lemma decomp_vars: "wfrestrictedvars\<^sub>s\<^sub>t (decomp t) = fv t" "vars\<^sub>s\<^sub>t (decomp t) = fv t" "bvars\<^sub>s\<^sub>t (decomp t) = {}" "fv\<^sub>s\<^sub>t (decomp t) = fv t" proof - obtain K M where Ana: "Ana t = (K,M)" by (metis surj_pair) hence "decomp t = send\t\\<^sub>s\<^sub>t#map Send K@map Receive M" unfolding decomp_def by simp moreover have "\(set (map fv K)) = fv\<^sub>s\<^sub>e\<^sub>t (set K)" "\(set (map fv M)) = fv\<^sub>s\<^sub>e\<^sub>t (set M)" by auto moreover have "fv\<^sub>s\<^sub>e\<^sub>t (set K) \ fv t" "fv\<^sub>s\<^sub>e\<^sub>t (set M) \ fv t" using Ana_subterm[OF Ana(1)] Ana_keys_fv[OF Ana(1)] by (simp_all add: UN_least psubsetD subtermeq_vars_subset) ultimately show "wfrestrictedvars\<^sub>s\<^sub>t (decomp t) = fv t" "vars\<^sub>s\<^sub>t (decomp t) = fv t" "bvars\<^sub>s\<^sub>t (decomp t) = {}" "fv\<^sub>s\<^sub>t (decomp t) = fv t" by auto qed private lemma bvars\<^sub>e\<^sub>s\<^sub>t_cons: "bvars\<^sub>e\<^sub>s\<^sub>t (x#X) = bvars\<^sub>e\<^sub>s\<^sub>t [x] \ bvars\<^sub>e\<^sub>s\<^sub>t X" by (cases x) auto private lemma bvars\<^sub>e\<^sub>s\<^sub>t_append: "bvars\<^sub>e\<^sub>s\<^sub>t (A@B) = bvars\<^sub>e\<^sub>s\<^sub>t A \ bvars\<^sub>e\<^sub>s\<^sub>t B" proof (induction A) case (Cons x A) thus ?case using bvars\<^sub>e\<^sub>s\<^sub>t_cons[of x "A@B"] bvars\<^sub>e\<^sub>s\<^sub>t_cons[of x A] by force qed simp private lemma fv\<^sub>e\<^sub>s\<^sub>t_cons: "fv\<^sub>e\<^sub>s\<^sub>t (x#X) = fv\<^sub>e\<^sub>s\<^sub>t [x] \ fv\<^sub>e\<^sub>s\<^sub>t X" by (cases x) auto private lemma fv\<^sub>e\<^sub>s\<^sub>t_append: "fv\<^sub>e\<^sub>s\<^sub>t (A@B) = fv\<^sub>e\<^sub>s\<^sub>t A \ fv\<^sub>e\<^sub>s\<^sub>t B" proof (induction A) case (Cons x A) thus ?case using fv\<^sub>e\<^sub>s\<^sub>t_cons[of x "A@B"] fv\<^sub>e\<^sub>s\<^sub>t_cons[of x A] by auto qed simp private lemma bvars_decomp: "bvars\<^sub>e\<^sub>s\<^sub>t (A@[Decomp t]) = bvars\<^sub>e\<^sub>s\<^sub>t A" "bvars\<^sub>e\<^sub>s\<^sub>t (Decomp t#A) = bvars\<^sub>e\<^sub>s\<^sub>t A" using bvars\<^sub>e\<^sub>s\<^sub>t_append decomp_vars(3) by fastforce+ private lemma bvars_decomp_rm: "bvars\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) = bvars\<^sub>e\<^sub>s\<^sub>t A" using bvars_decomp by (induct A rule: decomp_rm\<^sub>e\<^sub>s\<^sub>t.induct) simp_all+ private lemma fv_decomp_rm: "fv\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ fv\<^sub>e\<^sub>s\<^sub>t A" by (induct A rule: decomp_rm\<^sub>e\<^sub>s\<^sub>t.induct) auto private lemma ik_assignment_rhs_decomp_fv: assumes "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" shows "fv\<^sub>e\<^sub>s\<^sub>t (A@[Decomp t]) = fv\<^sub>e\<^sub>s\<^sub>t A" proof - have "fv\<^sub>e\<^sub>s\<^sub>t (A@[Decomp t]) = fv\<^sub>e\<^sub>s\<^sub>t A \ fv t" using fv\<^sub>e\<^sub>s\<^sub>t_append decomp_vars by simp moreover have "fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \ fv\<^sub>e\<^sub>s\<^sub>t A" by force moreover have "fv t \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" using fv_subset_subterms[OF assms(1)] by simp ultimately show ?thesis by blast qed private lemma wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_decomp_rm\<^sub>e\<^sub>s\<^sub>t_subset: "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t A" by (induct A rule: decomp_rm\<^sub>e\<^sub>s\<^sub>t.induct) auto+ private lemma wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_eq_wfrestrictedvars\<^sub>s\<^sub>t: "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t A = wfrestrictedvars\<^sub>s\<^sub>t (to_st A)" by simp private lemma decomp_set_unfold: assumes "Ana t = (K, M)" shows "set (decomp t) = {send\t\\<^sub>s\<^sub>t} \ (Send ` set K) \ (Receive ` set M)" using assms unfolding decomp_def by auto private lemma ik\<^sub>e\<^sub>s\<^sub>t_finite: "finite (ik\<^sub>e\<^sub>s\<^sub>t A)" by (rule finite_ik\<^sub>s\<^sub>t) private lemma assignment_rhs\<^sub>e\<^sub>s\<^sub>t_finite: "finite (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" by (rule finite_assignment_rhs\<^sub>s\<^sub>t) private lemma to_est_append: "to_est (A@B) = to_est A@to_est B" by (induct A rule: to_est.induct) auto private lemma to_st_to_est_inv: "to_st (to_est A) = A" by (induct A rule: to_est.induct) auto private lemma to_st_append: "to_st (A@B) = (to_st A)@(to_st B)" by (induct A rule: to_st.induct) auto private lemma to_st_cons: "to_st (a#B) = (to_st [a])@(to_st B)" using to_st_append[of "[a]" B] by simp private lemma wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_split: "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t (x#S) = wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t [x] \ wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t S" "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t (S@S') = wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t S \ wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t S'" using to_st_cons[of x S] to_st_append[of S S'] by auto private lemma ik\<^sub>e\<^sub>s\<^sub>t_append: "ik\<^sub>e\<^sub>s\<^sub>t (A@B) = ik\<^sub>e\<^sub>s\<^sub>t A \ ik\<^sub>e\<^sub>s\<^sub>t B" by (metis ik_append to_st_append) private lemma assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append: "assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@B) = assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t B" by (metis assignment_rhs_append to_st_append) private lemma ik\<^sub>e\<^sub>s\<^sub>t_cons: "ik\<^sub>e\<^sub>s\<^sub>t (a#A) = ik\<^sub>e\<^sub>s\<^sub>t [a] \ ik\<^sub>e\<^sub>s\<^sub>t A" -by (metis ik_append to_st_cons) +by (metis ik_append to_st_cons) private lemma ik\<^sub>e\<^sub>s\<^sub>t_append_subst: "ik\<^sub>e\<^sub>s\<^sub>t (A@B \\<^sub>e\<^sub>s\<^sub>t \) = ik\<^sub>e\<^sub>s\<^sub>t (A \\<^sub>e\<^sub>s\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t (B \\<^sub>e\<^sub>s\<^sub>t \)" "ik\<^sub>e\<^sub>s\<^sub>t (A@B) \\<^sub>s\<^sub>e\<^sub>t \ = (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t B \\<^sub>s\<^sub>e\<^sub>t \)" by (metis ik\<^sub>e\<^sub>s\<^sub>t_append extstrand_subst_hom(1), simp add: image_Un to_st_append) private lemma assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append_subst: "assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@B \\<^sub>e\<^sub>s\<^sub>t \) = assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A \\<^sub>e\<^sub>s\<^sub>t \) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t (B \\<^sub>e\<^sub>s\<^sub>t \)" "assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@B) \\<^sub>s\<^sub>e\<^sub>t \ = (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t B \\<^sub>s\<^sub>e\<^sub>t \)" by (metis assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append extstrand_subst_hom(1), use assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append in blast) private lemma ik\<^sub>e\<^sub>s\<^sub>t_cons_subst: "ik\<^sub>e\<^sub>s\<^sub>t (a#A \\<^sub>e\<^sub>s\<^sub>t \) = ik\<^sub>e\<^sub>s\<^sub>t ([a \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \]) \ ik\<^sub>e\<^sub>s\<^sub>t (A \\<^sub>e\<^sub>s\<^sub>t \)" "ik\<^sub>e\<^sub>s\<^sub>t (a#A) \\<^sub>s\<^sub>e\<^sub>t \ = (ik\<^sub>e\<^sub>s\<^sub>t [a] \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \)" by (metis ik\<^sub>e\<^sub>s\<^sub>t_cons extstrand_subst_hom(2), metis image_Un ik\<^sub>e\<^sub>s\<^sub>t_cons) private lemma decomp_rm\<^sub>e\<^sub>s\<^sub>t_append: "decomp_rm\<^sub>e\<^sub>s\<^sub>t (S@S') = (decomp_rm\<^sub>e\<^sub>s\<^sub>t S)@(decomp_rm\<^sub>e\<^sub>s\<^sub>t S')" by (induct S rule: decomp_rm\<^sub>e\<^sub>s\<^sub>t.induct) auto private lemma decomp_rm\<^sub>e\<^sub>s\<^sub>t_single[simp]: "decomp_rm\<^sub>e\<^sub>s\<^sub>t [Step (send\t\\<^sub>s\<^sub>t)] = [Step (send\t\\<^sub>s\<^sub>t)]" "decomp_rm\<^sub>e\<^sub>s\<^sub>t [Step (receive\t\\<^sub>s\<^sub>t)] = [Step (receive\t\\<^sub>s\<^sub>t)]" "decomp_rm\<^sub>e\<^sub>s\<^sub>t [Decomp t] = []" by auto private lemma decomp_rm\<^sub>e\<^sub>s\<^sub>t_ik_subset: "ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t S) \ ik\<^sub>e\<^sub>s\<^sub>t S" proof (induction S rule: decomp_rm\<^sub>e\<^sub>s\<^sub>t.induct) case (3 x S) thus ?case by (cases x) auto qed auto private lemma decomps\<^sub>e\<^sub>s\<^sub>t_ik_subset: "D \ decomps\<^sub>e\<^sub>s\<^sub>t M N \ \ ik\<^sub>e\<^sub>s\<^sub>t D \ subterms\<^sub>s\<^sub>e\<^sub>t (M \ N)" proof (induction D rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case (Decomp D f T K M') have "ik\<^sub>s\<^sub>t (decomp (Fun f T)) \ subterms (Fun f T)" "ik\<^sub>s\<^sub>t (decomp (Fun f T)) = ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun f T)]" using decomp_ik[OF Decomp.hyps(3)] Ana_subterm[OF Decomp.hyps(3)] by auto hence "ik\<^sub>s\<^sub>t (to_st [Decomp (Fun f T)]) \ subterms\<^sub>s\<^sub>e\<^sub>t (M \ N)" using in_subterms_subset_Union[OF Decomp.hyps(2)] by blast thus ?case using ik\<^sub>e\<^sub>s\<^sub>t_append[of D "[Decomp (Fun f T)]"] using Decomp.IH by auto qed simp private lemma decomps\<^sub>e\<^sub>s\<^sub>t_decomp_rm\<^sub>e\<^sub>s\<^sub>t_empty: "D \ decomps\<^sub>e\<^sub>s\<^sub>t M N \ \ decomp_rm\<^sub>e\<^sub>s\<^sub>t D = []" by (induct D rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) (auto simp add: decomp_rm\<^sub>e\<^sub>s\<^sub>t_append) private lemma decomps\<^sub>e\<^sub>s\<^sub>t_append: assumes "A \ decomps\<^sub>e\<^sub>s\<^sub>t S N \" "B \ decomps\<^sub>e\<^sub>s\<^sub>t S N \" shows "A@B \ decomps\<^sub>e\<^sub>s\<^sub>t S N \" using assms(2) proof (induction B rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case Nil show ?case using assms(1) by simp next case (Decomp B f X K T) hence "S \ ik\<^sub>e\<^sub>s\<^sub>t B \\<^sub>s\<^sub>e\<^sub>t \ \ S \ ik\<^sub>e\<^sub>s\<^sub>t (A@B) \\<^sub>s\<^sub>e\<^sub>t \" using ik\<^sub>e\<^sub>s\<^sub>t_append by auto thus ?case using decomps\<^sub>e\<^sub>s\<^sub>t.Decomp[OF Decomp.IH(1) Decomp.hyps(2,3,4)] ideduct_synth_mono[OF Decomp.hyps(5)] ideduct_synth_mono[OF Decomp.hyps(6)] by auto qed private lemma decomps\<^sub>e\<^sub>s\<^sub>t_subterms: assumes "A' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" shows "subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A') \ subterms\<^sub>s\<^sub>e\<^sub>t (M \ N)" using assms proof (induction A' rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case (Decomp D f X K T) hence "Fun f X \ subterms\<^sub>s\<^sub>e\<^sub>t (M \ N)" by auto hence "subterms\<^sub>s\<^sub>e\<^sub>t (set X) \ subterms\<^sub>s\<^sub>e\<^sub>t (M \ N)" using in_subterms_subset_Union[of "Fun f X" "M \ N"] params_subterms_Union[of X f] by blast moreover have "ik\<^sub>s\<^sub>t (to_st [Decomp (Fun f X)]) = set T" using Decomp.hyps(3) decomp_ik by simp hence "subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>s\<^sub>t (to_st [Decomp (Fun f X)])) \ subterms\<^sub>s\<^sub>e\<^sub>t (set X)" using Ana_fun_subterm[OF Decomp.hyps(3)] by auto ultimately show ?case using ik\<^sub>e\<^sub>s\<^sub>t_append[of D "[Decomp (Fun f X)]"] Decomp.IH by auto qed simp private lemma decomps\<^sub>e\<^sub>s\<^sub>t_assignment_rhs_empty: assumes "A' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" shows "assignment_rhs\<^sub>e\<^sub>s\<^sub>t A' = {}" using assms by (induction A' rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) - (simp_all add: decomp_assignment_rhs_empty assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append) + (simp_all add: decomp_assignment_rhs_empty assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append) private lemma decomps\<^sub>e\<^sub>s\<^sub>t_finite_ik_append: assumes "finite M" "M \ decomps\<^sub>e\<^sub>s\<^sub>t A N \" shows "\D \ decomps\<^sub>e\<^sub>s\<^sub>t A N \. ik\<^sub>e\<^sub>s\<^sub>t D = (\m \ M. ik\<^sub>e\<^sub>s\<^sub>t m)" using assms proof (induction M rule: finite_induct) case empty moreover have "[] \ decomps\<^sub>e\<^sub>s\<^sub>t A N \" "ik\<^sub>s\<^sub>t (to_st []) = {}" using decomps\<^sub>e\<^sub>s\<^sub>t.Nil by auto ultimately show ?case by blast next case (insert m M) then obtain D where "D \ decomps\<^sub>e\<^sub>s\<^sub>t A N \" "ik\<^sub>e\<^sub>s\<^sub>t D = (\m\M. ik\<^sub>s\<^sub>t (to_st m))" by moura moreover have "m \ decomps\<^sub>e\<^sub>s\<^sub>t A N \" using insert.prems(1) by blast ultimately show ?case using decomps\<^sub>e\<^sub>s\<^sub>t_append[of D A N \ m] ik\<^sub>e\<^sub>s\<^sub>t_append[of D m] by blast qed private lemma decomp_snd_exists[simp]: "\D. decomp t = send\t\\<^sub>s\<^sub>t#D" by (metis (mono_tags, lifting) decomp_def prod.case surj_pair) private lemma decomp_nonnil[simp]: "decomp t \ []" using decomp_snd_exists[of t] by fastforce private lemma to_st_nil_inv[dest]: "to_st A = [] \ A = []" by (induct A rule: to_st.induct) auto private lemma well_analyzedD: assumes "well_analyzed A" "Decomp t \ set A" shows "\f T. t = Fun f T" using assms proof (induction A rule: well_analyzed.induct) case (Decomp A t') hence "\f T. t' = Fun f T" by (cases t') auto moreover have "Decomp t \ set A \ t = t'" using Decomp by auto ultimately show ?case using Decomp.IH by auto qed auto private lemma well_analyzed_inv: assumes "well_analyzed (A@[Decomp t])" shows "t \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) - (Var ` \)" using assms well_analyzed.cases[of "A@[Decomp t]"] by fastforce private lemma well_analyzed_split_left_single: "well_analyzed (A@[a]) \ well_analyzed A" by (induction "A@[a]" rule: well_analyzed.induct) auto private lemma well_analyzed_split_left: "well_analyzed (A@B) \ well_analyzed A" proof (induction B rule: List.rev_induct) case (snoc b B) thus ?case using well_analyzed_split_left_single[of "A@B" b] by simp qed simp private lemma well_analyzed_append: assumes "well_analyzed A" "well_analyzed B" shows "well_analyzed (A@B)" using assms(2,1) proof (induction B rule: well_analyzed.induct) case (Step B x) show ?case using well_analyzed.Step[OF Step.IH[OF Step.prems]] by simp next case (Decomp B t) thus ?case using well_analyzed.Decomp[OF Decomp.IH[OF Decomp.prems]] ik\<^sub>e\<^sub>s\<^sub>t_append assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append by auto qed simp_all private lemma well_analyzed_singleton: "well_analyzed [Step (send\t\\<^sub>s\<^sub>t)]" "well_analyzed [Step (receive\t\\<^sub>s\<^sub>t)]" "well_analyzed [Step (\a: t \ t'\\<^sub>s\<^sub>t)]" "well_analyzed [Step (\X\\\: F\\<^sub>s\<^sub>t)]" "\well_analyzed [Decomp t]" proof - show "well_analyzed [Step (send\t\\<^sub>s\<^sub>t)]" "well_analyzed [Step (receive\t\\<^sub>s\<^sub>t)]" "well_analyzed [Step (\a: t \ t'\\<^sub>s\<^sub>t)]" "well_analyzed [Step (\X\\\: F\\<^sub>s\<^sub>t)]" using well_analyzed.Step[OF well_analyzed.Nil] by simp_all show "\well_analyzed [Decomp t]" using well_analyzed.cases[of "[Decomp t]"] by auto qed private lemma well_analyzed_decomp_rm\<^sub>e\<^sub>s\<^sub>t_fv: "well_analyzed A \ fv\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) = fv\<^sub>e\<^sub>s\<^sub>t A" proof assume "well_analyzed A" thus "fv\<^sub>e\<^sub>s\<^sub>t A \ fv\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A)" proof (induction A rule: well_analyzed.induct) case Decomp thus ?case using ik_assignment_rhs_decomp_fv decomp_rm\<^sub>e\<^sub>s\<^sub>t_append by auto next case (Step A x) have "fv\<^sub>e\<^sub>s\<^sub>t (A@[Step x]) = fv\<^sub>e\<^sub>s\<^sub>t A \ fv\<^sub>s\<^sub>t\<^sub>p x" "fv\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t (A@[Step x])) = fv\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ fv\<^sub>s\<^sub>t\<^sub>p x" using fv\<^sub>e\<^sub>s\<^sub>t_append decomp_rm\<^sub>e\<^sub>s\<^sub>t_append by auto thus ?case using Step by auto qed simp qed (rule fv_decomp_rm) private lemma sem\<^sub>e\<^sub>s\<^sub>t_d_split_left: assumes "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ (\@\')" shows "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ \" using assms sem\<^sub>e\<^sub>s\<^sub>t_d.cases by (induction \' rule: List.rev_induct) fastforce+ private lemma sem\<^sub>e\<^sub>s\<^sub>t_d_eq_sem_st: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ \ = \M\<^sub>0; to_st \\\<^sub>d' \" proof show "\M\<^sub>0; to_st \\\<^sub>d' \ \ sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ \" proof (induction \ arbitrary: M\<^sub>0 rule: List.rev_induct) case Nil show ?case using to_st_nil_inv by simp next case (snoc a \) hence IH: "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ \" and *: "\ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0; to_st [a]\\<^sub>d' \" using to_st_append by (auto simp add: sup.commute) thus ?case using snoc proof (cases a) case (Step b) thus ?thesis proof (cases b) case (Send t) thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_d.Send[OF IH] * Step by auto next case (Receive t) thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_d.Receive[OF IH] Step by auto next case (Equality a t t') thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_d.Equality[OF IH] * Step by auto next case (Inequality X F) thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_d.Inequality[OF IH] * Step by auto qed next case (Decomp t) obtain K M where Ana: "Ana t = (K,M)" by moura have "to_st [a] = decomp t" using Decomp by auto hence "to_st [a] = send\t\\<^sub>s\<^sub>t#map Send K@map Receive M" using Ana unfolding decomp_def by auto hence **: "ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" and "\ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0; map Send K\\<^sub>d' \" using * by auto hence "\k. k \ set K \ ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ k \ \" using * by (metis (full_types) strand_sem_d.simps(2) strand_sem_eq_defs(2) strand_sem_Send_split(2)) thus ?thesis using Decomp sem\<^sub>e\<^sub>s\<^sub>t_d.Decompose[OF IH ** Ana] by metis qed qed show "sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ \ \ \M\<^sub>0; to_st \\\<^sub>d' \" proof (induction rule: sem\<^sub>e\<^sub>s\<^sub>t_d.induct) case Nil thus ?case by simp next case (Send M\<^sub>0 \ \ t) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[send\t\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (send\t\\<^sub>s\<^sub>t)]"] by (simp add: sup.commute) next case (Receive M\<^sub>0 \ \ t) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[receive\t\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (receive\t\\<^sub>s\<^sub>t)]"] by (simp add: sup.commute) next case (Equality M\<^sub>0 \ \ t t' a) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[\a: t \ t'\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (\a: t \ t'\\<^sub>s\<^sub>t)]"] by (simp add: sup.commute) next case (Inequality M\<^sub>0 \ \ X F) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[\X\\\: F\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (\X\\\: F\\<^sub>s\<^sub>t)]"] by (simp add: sup.commute) next case (Decompose M\<^sub>0 \ \ t K M) have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); decomp t\\<^sub>d' \" proof - have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [send\t\\<^sub>s\<^sub>t]\\<^sub>d' \" using Decompose.hyps(2) by (auto simp add: sup.commute) moreover have "\k. k \ set K \ M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \) \\<^sub>s\<^sub>e\<^sub>t \ \ k \ \" using Decompose by (metis sup.commute) hence "\k. k \ set K \ \M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [Send k]\\<^sub>d' \" by auto hence "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); map Send K\\<^sub>d' \" using strand_sem_Send_map(2)[of K, of "M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \) \\<^sub>s\<^sub>e\<^sub>t \" \] strand_sem_eq_defs(2) by auto moreover have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); map Receive M\\<^sub>d' \" by (metis strand_sem_Receive_map(2) strand_sem_eq_defs(2)) ultimately have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); send\t\\<^sub>s\<^sub>t#map Send K@map Receive M\\<^sub>d' \" by auto thus ?thesis using Decompose.hyps(3) unfolding decomp_def by auto qed hence "\M\<^sub>0; to_st \@decomp t\\<^sub>d' \" using strand_sem_append'[of M\<^sub>0 "to_st \" \ "decomp t"] Decompose.IH by simp thus ?case using to_st_append[of \ "[Decomp t]"] by simp qed qed private lemma sem\<^sub>e\<^sub>s\<^sub>t_c_eq_sem_st: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ \ = \M\<^sub>0; to_st \\\<^sub>c' \" proof show "\M\<^sub>0; to_st \\\<^sub>c' \ \ sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ \" proof (induction \ arbitrary: M\<^sub>0 rule: List.rev_induct) case Nil show ?case using to_st_nil_inv by simp next case (snoc a \) hence IH: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ \" and *: "\ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0; to_st [a]\\<^sub>c' \" using to_st_append by (auto simp add: sup.commute) thus ?case using snoc proof (cases a) case (Step b) thus ?thesis proof (cases b) case (Send t) thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_c.Send[OF IH] * Step by auto next case (Receive t) thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_c.Receive[OF IH] Step by auto next case (Equality t) thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_c.Equality[OF IH] * Step by auto next case (Inequality t) thus ?thesis using sem\<^sub>e\<^sub>s\<^sub>t_c.Inequality[OF IH] * Step by auto qed next case (Decomp t) obtain K M where Ana: "Ana t = (K,M)" by moura have "to_st [a] = decomp t" using Decomp by auto hence "to_st [a] = send\t\\<^sub>s\<^sub>t#map Send K@map Receive M" using Ana unfolding decomp_def by auto hence **: "ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" and "\ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0; map Send K\\<^sub>c' \" using * by auto hence "\k. k \ set K \ ik\<^sub>e\<^sub>s\<^sub>t \ \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \" using * strand_sem_Send_split(1) strand_sem_eq_defs(1) by auto thus ?thesis using Decomp sem\<^sub>e\<^sub>s\<^sub>t_c.Decompose[OF IH ** Ana] by metis qed qed show "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ \ \ \M\<^sub>0; to_st \\\<^sub>c' \" proof (induction rule: sem\<^sub>e\<^sub>s\<^sub>t_c.induct) case Nil thus ?case by simp next case (Send M\<^sub>0 \ \ t) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[send\t\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (send\t\\<^sub>s\<^sub>t)]"] by (simp add: sup.commute) next case (Receive M\<^sub>0 \ \ t) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[receive\t\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (receive\t\\<^sub>s\<^sub>t)]"] by (simp add: sup.commute) next case (Equality M\<^sub>0 \ \ t t' a) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[\a: t \ t'\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (\a: t \ t'\\<^sub>s\<^sub>t)]"] by (simp add: sup.commute) next case (Inequality M\<^sub>0 \ \ X F) thus ?case using strand_sem_append'[of M\<^sub>0 "to_st \" \ "[\X\\\: F\\<^sub>s\<^sub>t]"] to_st_append[of \ "[Step (\X\\\: F\\<^sub>s\<^sub>t)]"] by (auto simp add: sup.commute) next case (Decompose M\<^sub>0 \ \ t K M) have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); decomp t\\<^sub>c' \" proof - have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [send\t\\<^sub>s\<^sub>t]\\<^sub>c' \" using Decompose.hyps(2) by (auto simp add: sup.commute) moreover have "\k. k \ set K \ M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \" using Decompose by (metis sup.commute) hence "\k. k \ set K \ \M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); [Send k]\\<^sub>c' \" by auto hence "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); map Send K\\<^sub>c' \" using strand_sem_Send_map(1)[of K, of "M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \) \\<^sub>s\<^sub>e\<^sub>t \" \] strand_sem_eq_defs(1) by auto moreover have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); map Receive M\\<^sub>c' \" by (metis strand_sem_Receive_map(1) strand_sem_eq_defs(1)) ultimately have "\M\<^sub>0 \ ik\<^sub>s\<^sub>t (to_st \); send\t\\<^sub>s\<^sub>t#map Send K@map Receive M\\<^sub>c' \" by auto thus ?thesis using Decompose.hyps(3) unfolding decomp_def by auto qed hence "\M\<^sub>0; to_st \@decomp t\\<^sub>c' \" using strand_sem_append'[of M\<^sub>0 "to_st \" \ "decomp t"] Decompose.IH by simp thus ?case using to_st_append[of \ "[Decomp t]"] by simp qed qed private lemma sem\<^sub>e\<^sub>s\<^sub>t_c_decomp_rm\<^sub>e\<^sub>s\<^sub>t_deduct_aux: assumes "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ A" "t \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" "t \ ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t \" shows "ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ t" using assms proof (induction M\<^sub>0 \ A arbitrary: t rule: sem\<^sub>e\<^sub>s\<^sub>t_c.induct) case (Send M\<^sub>0 \ A t') thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append ik\<^sub>e\<^sub>s\<^sub>t_append by auto next case (Receive M\<^sub>0 \ A t') hence "t \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" "t \ ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t \" using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append ik\<^sub>e\<^sub>s\<^sub>t_append by auto hence IH: "ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ t" using Receive.IH by auto show ?case using ideduct_mono[OF IH] decomp_rm\<^sub>e\<^sub>s\<^sub>t_append ik\<^sub>e\<^sub>s\<^sub>t_append by auto next case (Equality M\<^sub>0 \ A t') thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append ik\<^sub>e\<^sub>s\<^sub>t_append by auto next case (Inequality M\<^sub>0 \ A t') thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append ik\<^sub>e\<^sub>s\<^sub>t_append by auto next case (Decompose M\<^sub>0 \ A t' K M t) have *: "ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ t' \ \" using Decompose.hyps(2) proof (induction rule: intruder_synth_induct) case (AxiomC t'') moreover { assume "t'' \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" "t'' \ ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t \" hence ?case using Decompose.IH by auto } ultimately show ?case by force qed simp - + { fix k assume "k \ set K" hence "ik\<^sub>e\<^sub>s\<^sub>t A \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \" using Decompose.hyps by auto hence "ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ k \ \" proof (induction rule: intruder_synth_induct) case (AxiomC t'') moreover { assume "t'' \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" "t'' \ ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t \" hence ?case using Decompose.IH by auto } ultimately show ?case by force qed simp } hence **: "\k. k \ set (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ k" by auto - + show ?case proof (cases "t \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \") case True thus ?thesis using Decompose.IH Decompose.prems(2) decomp_rm\<^sub>e\<^sub>s\<^sub>t_append by auto next case False hence "t \ ik\<^sub>s\<^sub>t (decomp t') \\<^sub>s\<^sub>e\<^sub>t \" using Decompose.prems(1) ik\<^sub>e\<^sub>s\<^sub>t_append by auto hence ***: "t \ set (M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Decompose.hyps(3) decomp_ik by auto hence "M \ []" by auto hence ****: "Ana (t' \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Ana_subst[OF Decompose.hyps(3)] by auto have "ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ t" by (rule intruder_deduct.Decompose[OF * **** ** ***]) thus ?thesis using ideduct_mono decomp_rm\<^sub>e\<^sub>s\<^sub>t_append by auto qed qed simp private lemma sem\<^sub>e\<^sub>s\<^sub>t_c_decomp_rm\<^sub>e\<^sub>s\<^sub>t_deduct: assumes "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ A" "ik\<^sub>e\<^sub>s\<^sub>t A \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t" shows "ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \ \ t" using assms(2) proof (induction t rule: intruder_synth_induct) case (AxiomC t) hence "t \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ \ t \ M\<^sub>0 \\<^sub>s\<^sub>e\<^sub>t \" by auto moreover { assume "t \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" "t \ ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t \" hence ?case using ideduct_mono[OF intruder_deduct.Axiom] by auto } moreover { assume "t \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" "t \ ik\<^sub>e\<^sub>s\<^sub>t (decomp_rm\<^sub>e\<^sub>s\<^sub>t A) \\<^sub>s\<^sub>e\<^sub>t \" hence ?case using sem\<^sub>e\<^sub>s\<^sub>t_c_decomp_rm\<^sub>e\<^sub>s\<^sub>t_deduct_aux[OF assms(1)] by auto } ultimately show ?case by auto qed simp private lemma sem\<^sub>e\<^sub>s\<^sub>t_d_decomp_rm\<^sub>e\<^sub>s\<^sub>t_if_sem\<^sub>e\<^sub>s\<^sub>t_c: "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ A \ sem\<^sub>e\<^sub>s\<^sub>t_d M\<^sub>0 \ (decomp_rm\<^sub>e\<^sub>s\<^sub>t A)" proof (induction M\<^sub>0 \ A rule: sem\<^sub>e\<^sub>s\<^sub>t_c.induct) case (Send M\<^sub>0 \ A t) thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append sem\<^sub>e\<^sub>s\<^sub>t_d.Send[OF Send.IH] sem\<^sub>e\<^sub>s\<^sub>t_c_decomp_rm\<^sub>e\<^sub>s\<^sub>t_deduct by auto next case (Receive t) thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append sem\<^sub>e\<^sub>s\<^sub>t_d.Receive by auto next case (Equality M\<^sub>0 \ A t) thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append sem\<^sub>e\<^sub>s\<^sub>t_d.Equality[OF Equality.IH] sem\<^sub>e\<^sub>s\<^sub>t_c_decomp_rm\<^sub>e\<^sub>s\<^sub>t_deduct by auto next case (Inequality M\<^sub>0 \ A t) thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append sem\<^sub>e\<^sub>s\<^sub>t_d.Inequality[OF Inequality.IH] sem\<^sub>e\<^sub>s\<^sub>t_c_decomp_rm\<^sub>e\<^sub>s\<^sub>t_deduct by auto next case Decompose thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append by auto qed auto private lemma sem\<^sub>e\<^sub>s\<^sub>t_c_decomps\<^sub>e\<^sub>s\<^sub>t_append: assumes "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ A" "D \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \) \" shows "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ (A@D)" using assms(2,1) proof (induction D rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case (Decomp D f T K M) hence *: "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ (A @ D)" "ik\<^sub>e\<^sub>s\<^sub>t (A@D) \ {} \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c Fun f T \ \" "\k. k \ set K \ ik\<^sub>e\<^sub>s\<^sub>t (A @ D) \ {} \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \" using ik\<^sub>e\<^sub>s\<^sub>t_append by auto show ?case using sem\<^sub>e\<^sub>s\<^sub>t_c.Decompose[OF *(1,2) Decomp.hyps(3) *(3)] by simp qed auto private lemma decomps\<^sub>e\<^sub>s\<^sub>t_preserves_wf: assumes "D \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \" "wf\<^sub>e\<^sub>s\<^sub>t V A" shows "wf\<^sub>e\<^sub>s\<^sub>t V (A@D)" using assms proof (induction D rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case (Decomp D f T K M) have "wfrestrictedvars\<^sub>s\<^sub>t (decomp (Fun f T)) \ fv\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" using decomp_vars fv_subset_subterms[OF Decomp.hyps(2)] by fast hence "wfrestrictedvars\<^sub>s\<^sub>t (decomp (Fun f T)) \ wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t A" using ik\<^sub>s\<^sub>t_assignment_rhs\<^sub>s\<^sub>t_wfrestrictedvars_subset[of "to_st A"] by blast hence "wfrestrictedvars\<^sub>s\<^sub>t (decomp (Fun f T)) \ wfrestrictedvars\<^sub>s\<^sub>t (to_st (A@D)) \ V" using to_st_append[of A D] strand_vars_split(2)[of "to_st A" "to_st D"] by (metis le_supI1) thus ?case using wf_append_suffix[OF Decomp.IH[OF Decomp.prems], of "decomp (Fun f T)"] to_st_append[of "A@D" "[Decomp (Fun f T)]"] by auto qed auto private lemma decomps\<^sub>e\<^sub>s\<^sub>t_preserves_model_c: assumes "D \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \" "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ A" shows "sem\<^sub>e\<^sub>s\<^sub>t_c M\<^sub>0 \ (A@D)" using assms proof (induction D rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case (Decomp D f T K M) show ?case using sem\<^sub>e\<^sub>s\<^sub>t_c.Decompose[OF Decomp.IH[OF Decomp.prems] _ Decomp.hyps(3)] Decomp.hyps(5,6) ideduct_synth_mono ik\<^sub>e\<^sub>s\<^sub>t_append - by (metis (mono_tags, lifting) List.append_assoc image_Un sup_ge1) + by (metis (mono_tags, lifting) List.append_assoc image_Un sup_ge1) qed auto private lemma decomps\<^sub>e\<^sub>s\<^sub>t_exist_aux: assumes "D \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "M \ ik\<^sub>e\<^sub>s\<^sub>t D \ t" "\(M \ (ik\<^sub>e\<^sub>s\<^sub>t D) \\<^sub>c t)" obtains D' where "D@D' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "M \ ik\<^sub>e\<^sub>s\<^sub>t (D@D') \\<^sub>c t" "M \ ik\<^sub>e\<^sub>s\<^sub>t D \ M \ ik\<^sub>e\<^sub>s\<^sub>t (D@D')" proof - have "\D' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \. M \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>c t" using assms(2) proof (induction t rule: intruder_deduct_induct) case (Compose X f) from Compose.IH have "\D \ decomps\<^sub>e\<^sub>s\<^sub>t M N \. \x \ set X. M \ ik\<^sub>e\<^sub>s\<^sub>t D \\<^sub>c x" proof (induction X) case (Cons t X) then obtain D' D'' where D': "D' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "M \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>c t" and D'': "D'' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "\x \ set X. M \ ik\<^sub>e\<^sub>s\<^sub>t D'' \\<^sub>c x" by moura hence "M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>c t" "\x \ set X. M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>c x" by (auto intro: ideduct_synth_mono simp add: ik\<^sub>e\<^sub>s\<^sub>t_append) thus ?case using decomps\<^sub>e\<^sub>s\<^sub>t_append[OF D'(1) D''(1)] by (metis set_ConsD) qed (auto intro: decomps\<^sub>e\<^sub>s\<^sub>t.Nil) thus ?case using intruder_synth.ComposeC[OF Compose.hyps(1,2)] by metis next case (Decompose t K T t\<^sub>i) have "\D \ decomps\<^sub>e\<^sub>s\<^sub>t M N \. \k \ set K. M \ ik\<^sub>e\<^sub>s\<^sub>t D \\<^sub>c k" using Decompose.IH proof (induction K) case (Cons t X) then obtain D' D'' where D': "D' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "M \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>c t" and D'': "D'' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "\x \ set X. M \ ik\<^sub>e\<^sub>s\<^sub>t D'' \\<^sub>c x" using assms(1) by moura hence "M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>c t" "\x \ set X. M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>c x" by (auto intro: ideduct_synth_mono simp add: ik\<^sub>e\<^sub>s\<^sub>t_append) thus ?case using decomps\<^sub>e\<^sub>s\<^sub>t_append[OF D'(1) D''(1)] by auto qed auto then obtain D' where D': "D' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "\k. k \ set K \ M \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>c k" by metis obtain D'' where D'': "D'' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "M \ ik\<^sub>e\<^sub>s\<^sub>t D'' \\<^sub>c t" by (metis Decompose.IH(1)) obtain f X where fX: "t = Fun f X" "t\<^sub>i \ set X" using Decompose.hyps(2,4) by (cases t) (auto dest: Ana_fun_subterm) - + from decomps\<^sub>e\<^sub>s\<^sub>t_append[OF D'(1) D''(1)] D'(2) D''(2) have *: "D'@D'' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" "\k. k \ set K \ M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>c k" "M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>c t" by (auto intro: ideduct_synth_mono simp add: ik\<^sub>e\<^sub>s\<^sub>t_append) hence **: "\k. k \ set K \ M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \" using ideduct_synth_subst by auto have "t\<^sub>i \ ik\<^sub>s\<^sub>t (decomp t)" using Decompose.hyps(2,4) ik_rcv_map unfolding decomp_def by auto with *(3) fX(1) Decompose.hyps(2) show ?case proof (induction t rule: intruder_synth_induct) case (AxiomC t) hence t_in_subterms: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (M \ N)" using decomps\<^sub>e\<^sub>s\<^sub>t_ik_subset[OF *(1)] subset_subterms_Union by auto have "M \ ik\<^sub>e\<^sub>s\<^sub>t (D'@D'') \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" using ideduct_synth_subst[OF intruder_synth.AxiomC[OF AxiomC.hyps(1)]] by metis moreover have "T \ []" using decomp_ik[OF \Ana t = (K,T)\] \t\<^sub>i \ ik\<^sub>s\<^sub>t (decomp t)\ by auto ultimately have "D'@D''@[Decomp (Fun f X)] \ decomps\<^sub>e\<^sub>s\<^sub>t M N \" using AxiomC decomps\<^sub>e\<^sub>s\<^sub>t.Decomp[OF *(1) _ _ _ _ **] subset_subterms_Union t_in_subterms by (simp add: subset_eq) moreover have "decomp t = to_st [Decomp (Fun f X)]" using AxiomC.prems(1,2) by auto ultimately show ?case by (metis AxiomC.prems(3) UnCI intruder_synth.AxiomC ik\<^sub>e\<^sub>s\<^sub>t_append to_st_append) qed (auto intro!: fX(2) *(1)) qed (fastforce intro: intruder_synth.AxiomC assms(1)) hence "\D' \ decomps\<^sub>e\<^sub>s\<^sub>t M N \. M \ ik\<^sub>e\<^sub>s\<^sub>t (D@D') \\<^sub>c t" by (auto intro: ideduct_synth_mono simp add: ik\<^sub>e\<^sub>s\<^sub>t_append) thus thesis using that[OF decomps\<^sub>e\<^sub>s\<^sub>t_append[OF assms(1)]] assms ik\<^sub>e\<^sub>s\<^sub>t_append by moura qed private lemma decomps\<^sub>e\<^sub>s\<^sub>t_ik_max_exist: assumes "finite A" "finite N" shows "\D \ decomps\<^sub>e\<^sub>s\<^sub>t A N \. \D' \ decomps\<^sub>e\<^sub>s\<^sub>t A N \. ik\<^sub>e\<^sub>s\<^sub>t D' \ ik\<^sub>e\<^sub>s\<^sub>t D" proof - let ?IK = "\M. \D \ M. ik\<^sub>e\<^sub>s\<^sub>t D" have "?IK (decomps\<^sub>e\<^sub>s\<^sub>t A N \) \ (\t \ A \ N. subterms t)" by (auto dest!: decomps\<^sub>e\<^sub>s\<^sub>t_ik_subset) hence "finite (?IK (decomps\<^sub>e\<^sub>s\<^sub>t A N \))" using subterms_union_finite[OF assms(1)] subterms_union_finite[OF assms(2)] infinite_super by auto then obtain M where M: "finite M" "M \ decomps\<^sub>e\<^sub>s\<^sub>t A N \" "?IK M = ?IK (decomps\<^sub>e\<^sub>s\<^sub>t A N \)" using finite_subset_Union by moura show ?thesis using decomps\<^sub>e\<^sub>s\<^sub>t_finite_ik_append[OF M(1,2)] M(3) by auto qed private lemma decomps\<^sub>e\<^sub>s\<^sub>t_exist: assumes "finite A" "finite N" shows "\D \ decomps\<^sub>e\<^sub>s\<^sub>t A N \. \t. A \ t \ A \ ik\<^sub>e\<^sub>s\<^sub>t D \\<^sub>c t" proof (rule ccontr) assume neg: "\(\D \ decomps\<^sub>e\<^sub>s\<^sub>t A N \. \t. A \ t \ A \ ik\<^sub>e\<^sub>s\<^sub>t D \\<^sub>c t)" obtain D where D: "D \ decomps\<^sub>e\<^sub>s\<^sub>t A N \" "\D' \ decomps\<^sub>e\<^sub>s\<^sub>t A N \. ik\<^sub>e\<^sub>s\<^sub>t D' \ ik\<^sub>e\<^sub>s\<^sub>t D" using decomps\<^sub>e\<^sub>s\<^sub>t_ik_max_exist[OF assms] by moura then obtain t where t: "A \ ik\<^sub>e\<^sub>s\<^sub>t D \ t" "\(A \ ik\<^sub>e\<^sub>s\<^sub>t D \\<^sub>c t)" using neg by (fastforce intro: ideduct_mono) obtain D' where D': "D@D' \ decomps\<^sub>e\<^sub>s\<^sub>t A N \" "A \ ik\<^sub>e\<^sub>s\<^sub>t (D@D') \\<^sub>c t" "A \ ik\<^sub>e\<^sub>s\<^sub>t D \ A \ ik\<^sub>e\<^sub>s\<^sub>t (D@D')" by (metis decomps\<^sub>e\<^sub>s\<^sub>t_exist_aux t D(1)) hence "ik\<^sub>e\<^sub>s\<^sub>t D \ ik\<^sub>e\<^sub>s\<^sub>t (D@D')" using ik\<^sub>e\<^sub>s\<^sub>t_append by auto moreover have "ik\<^sub>e\<^sub>s\<^sub>t (D@D') \ ik\<^sub>e\<^sub>s\<^sub>t D" using D(2) D'(1) by auto ultimately show False by simp qed private lemma decomps\<^sub>e\<^sub>s\<^sub>t_exist_subst: assumes "ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" and "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ A" "wf\<^sub>e\<^sub>s\<^sub>t {} A" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "Ana_invar_subst (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" and "well_analyzed A" shows "\D \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \. ik\<^sub>e\<^sub>s\<^sub>t (A@D) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" proof - have ik_eq: "ik\<^sub>e\<^sub>s\<^sub>t (A \\<^sub>e\<^sub>s\<^sub>t \) = ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" using assms(5,6) proof (induction A rule: List.rev_induct) case (snoc a A) hence "Ana_invar_subst (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" using Ana_invar_subst_subset[OF snoc.prems(1)] ik\<^sub>e\<^sub>s\<^sub>t_append assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append unfolding Ana_invar_subst_def by simp with snoc have IH: "ik\<^sub>e\<^sub>s\<^sub>t (A@[a] \\<^sub>e\<^sub>s\<^sub>t \) = (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t ([a] \\<^sub>e\<^sub>s\<^sub>t \)" "ik\<^sub>e\<^sub>s\<^sub>t (A@[a]) \\<^sub>s\<^sub>e\<^sub>t \ = (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t [a] \\<^sub>s\<^sub>e\<^sub>t \)" using well_analyzed_split_left[OF snoc.prems(2)] by (auto simp add: to_st_append ik\<^sub>e\<^sub>s\<^sub>t_append_subst) - + have "ik\<^sub>e\<^sub>s\<^sub>t [a \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \] = ik\<^sub>e\<^sub>s\<^sub>t [a] \\<^sub>s\<^sub>e\<^sub>t \" proof (cases a) case (Step b) thus ?thesis by (cases b) auto next case (Decomp t) then obtain f T where t: "t = Fun f T" using well_analyzedD[OF snoc.prems(2)] by force obtain K M where Ana_t: "Ana (Fun f T) = (K,M)" by (metis surj_pair) moreover have "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t ((ik\<^sub>e\<^sub>s\<^sub>t (A@[a]) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@[a])))" using t Decomp snoc.prems(2) by (auto dest: well_analyzed_inv simp add: ik\<^sub>e\<^sub>s\<^sub>t_append assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append) hence "Ana (Fun f T \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Ana_t snoc.prems(1) - unfolding Ana_invar_subst_def by force + unfolding Ana_invar_subst_def by blast ultimately show ?thesis using Decomp t by (auto simp add: decomp_ik) qed thus ?case using IH unfolding subst_apply_extstrand_def by simp qed simp moreover have assignment_rhs_eq: "assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A \\<^sub>e\<^sub>s\<^sub>t \) = assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" using assms(5,6) proof (induction A rule: List.rev_induct) case (snoc a A) hence "Ana_invar_subst (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" using Ana_invar_subst_subset[OF snoc.prems(1)] ik\<^sub>e\<^sub>s\<^sub>t_append assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append unfolding Ana_invar_subst_def by simp hence "assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A \\<^sub>e\<^sub>s\<^sub>t \) = assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \" using snoc.IH well_analyzed_split_left[OF snoc.prems(2)] by simp hence IH: "assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@[a] \\<^sub>e\<^sub>s\<^sub>t \) = (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t ([a] \\<^sub>e\<^sub>s\<^sub>t \)" "assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@[a]) \\<^sub>s\<^sub>e\<^sub>t \ = (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t [a] \\<^sub>s\<^sub>e\<^sub>t \)" by (metis assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append_subst(1), metis assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append_subst(2)) have "assignment_rhs\<^sub>e\<^sub>s\<^sub>t [a \\<^sub>e\<^sub>s\<^sub>t\<^sub>p \] = assignment_rhs\<^sub>e\<^sub>s\<^sub>t [a] \\<^sub>s\<^sub>e\<^sub>t \" proof (cases a) case (Step b) thus ?thesis by (cases b) auto next case (Decomp t) then obtain f T where t: "t = Fun f T" using well_analyzedD[OF snoc.prems(2)] by force obtain K M where Ana_t: "Ana (Fun f T) = (K,M)" by (metis surj_pair) moreover have "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t ((ik\<^sub>e\<^sub>s\<^sub>t (A@[a]) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@[a])))" using t Decomp snoc.prems(2) by (auto dest: well_analyzed_inv simp add: ik\<^sub>e\<^sub>s\<^sub>t_append assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append) hence "Ana (Fun f T \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" - using Ana_t snoc.prems(1) unfolding Ana_invar_subst_def by force + using Ana_t snoc.prems(1) unfolding Ana_invar_subst_def by blast ultimately show ?thesis using Decomp t by (auto simp add: decomp_assignment_rhs_empty) qed thus ?case using IH unfolding subst_apply_extstrand_def by simp qed simp ultimately obtain D where D: "D \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) Var" "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D) \\<^sub>c t \ \" using decomps\<^sub>e\<^sub>s\<^sub>t_exist[OF ik\<^sub>e\<^sub>s\<^sub>t_finite assignment_rhs\<^sub>e\<^sub>s\<^sub>t_finite, of "A \\<^sub>e\<^sub>s\<^sub>t \" "A \\<^sub>e\<^sub>s\<^sub>t \"] ik\<^sub>e\<^sub>s\<^sub>t_append assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append assms(1) by force let ?P = "\D D'. \t. (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D) \\<^sub>c t \ (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t" have "\D' \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \. ?P D D'" using D(1) proof (induction D rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case Nil have "ik\<^sub>e\<^sub>s\<^sub>t [] = ik\<^sub>e\<^sub>s\<^sub>t [] \\<^sub>s\<^sub>e\<^sub>t \" by auto thus ?case by (metis decomps\<^sub>e\<^sub>s\<^sub>t.Nil) next case (Decomp D f T K M) obtain D' where D': "D' \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \" "?P D D'" using Decomp.IH by auto hence IH: "\k. k \ set K \ (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c k" "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c Fun f T" using Decomp.hyps(5,6) by auto have D'_ik: "ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \ \ subterms\<^sub>s\<^sub>e\<^sub>t ((ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)) \\<^sub>s\<^sub>e\<^sub>t \" "ik\<^sub>e\<^sub>s\<^sub>t D' \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" using decomps\<^sub>e\<^sub>s\<^sub>t_ik_subset[OF D'(1)] by (metis subst_all_mono, metis) show ?case using IH(2,1) Decomp.hyps(2,3,4) proof (induction "Fun f T" arbitrary: f T K M rule: intruder_synth_induct) case (AxiomC f T) then obtain s where s: "s \ ik\<^sub>e\<^sub>s\<^sub>t A \ ik\<^sub>e\<^sub>s\<^sub>t D'" "Fun f T = s \ \" using AxiomC.prems by blast hence fT_s_in: "Fun f T \ (subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)) \\<^sub>s\<^sub>e\<^sub>t \" "s \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" using AxiomC D'_ik subset_subterms_Union[of "ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A"] subst_all_mono[OF subset_subterms_Union, of \] by (metis (no_types) Un_iff image_eqI subset_Un_eq, metis (no_types) Un_iff subset_Un_eq) obtain Ks Ms where Ana_s: "Ana s = (Ks,Ms)" by moura have AD'_props: "wf\<^sub>e\<^sub>s\<^sub>t {} (A@D')" "\{}; to_st (A@D')\\<^sub>c \" using decomps\<^sub>e\<^sub>s\<^sub>t_preserves_model_c[OF D'(1) assms(2)] decomps\<^sub>e\<^sub>s\<^sub>t_preserves_wf[OF D'(1) assms(3)] sem\<^sub>e\<^sub>s\<^sub>t_c_eq_sem_st strand_sem_eq_defs(1) by auto show ?case proof (cases s) case (Var x) \ \In this case \\ x\ (is a subterm of something that) was derived from an "earlier intruder knowledge" because \A\ is well-formed and has \\\ as a model. So either the intruder composed \Fun f T\ himself (making \Decomp (Fun f T)\ unnecessary) or \Fun f T\ is an instance of something else in the intruder knowledge (in which case the "something" can be used in place of \Fun f T\)\ hence "Var x \ ik\<^sub>e\<^sub>s\<^sub>t (A@D')" "\ x = Fun f T" using s ik\<^sub>e\<^sub>s\<^sub>t_append by auto show ?thesis proof (cases "\m \ set M. ik\<^sub>e\<^sub>s\<^sub>t A \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c m") case True \ \All terms acquired by decomposing \Fun f T\ are already derivable. Hence there is no need to consider decomposition of \Fun f T\ at all.\ have *: "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t (D@[Decomp (Fun f T)]) = (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t D \ set M" using decomp_ik[OF \Ana (Fun f T) = (K,M)\] ik\<^sub>e\<^sub>s\<^sub>t_append[of D "[Decomp (Fun f T)]"] by auto - + { fix t' assume "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t D \ set M \\<^sub>c t'" hence "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t'" proof (induction t' rule: intruder_synth_induct) case (AxiomC t') thus ?case proof assume "t' \ set M" moreover have "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) = ik\<^sub>e\<^sub>s\<^sub>t A \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \" by auto ultimately show ?case using True by auto qed (metis D'(2) intruder_synth.AxiomC) qed auto } thus ?thesis using D'(1) * by metis next case False \ \Some term acquired by decomposition of \Fun f T\ cannot be derived in \\\<^sub>c\. \Fun f T\ must therefore be an instance of something else in the intruder knowledge, because of well-formedness.\ then obtain t\<^sub>i where t\<^sub>i: "t\<^sub>i \ set T" "\ik\<^sub>e\<^sub>s\<^sub>t (A@D') \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t\<^sub>i" using Ana_fun_subterm[OF \Ana (Fun f T) = (K,M)\] by (auto simp add: ik\<^sub>e\<^sub>s\<^sub>t_append) obtain S where fS: "Fun f S \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t (A@D')) \ Fun f S \ subterms\<^sub>s\<^sub>e\<^sub>t (assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@D'))" "\ x = Fun f S \ \" using strand_sem_wf_ik_or_assignment_rhs_fun_subterm[ OF AD'_props \Var x \ ik\<^sub>e\<^sub>s\<^sub>t (A@D')\ _ t\<^sub>i \interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\] \\ x = Fun f T\ by moura hence fS_in: "Fun f S \ \ \ ik\<^sub>e\<^sub>s\<^sub>t A \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \" "Fun f S \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A)" using imageI[OF s(1), of "\x. x \ \"] Var ik\<^sub>e\<^sub>s\<^sub>t_append[of A D'] assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append[of A D'] decomps\<^sub>e\<^sub>s\<^sub>t_subterms[OF D'(1)] decomps\<^sub>e\<^sub>s\<^sub>t_assignment_rhs_empty[OF D'(1)] by auto obtain KS MS where Ana_fS: "Ana (Fun f S) = (KS, MS)" by moura hence "K = KS \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" "M = MS \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" using Ana_invar_substD[OF assms(5) fS_in(2)] s(2) fS(2) \s = Var x\ \Ana (Fun f T) = (K,M)\ by simp_all hence "MS \ []" using \M \ []\ by simp have "\k. k \ set KS \ ik\<^sub>e\<^sub>s\<^sub>t A \ ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c k \ \" using AxiomC.prems(1) \K = KS \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\ by (simp add: image_Un) hence D'': "D'@[Decomp (Fun f S)] \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \" using decomps\<^sub>e\<^sub>s\<^sub>t.Decomp[OF D'(1) fS_in(2) Ana_fS \MS \ []\] AxiomC.prems(1) intruder_synth.AxiomC[OF fS_in(1)] by simp moreover { fix t' assume "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t (D@[Decomp (Fun f T)]) \\<^sub>c t'" hence "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t (D'@[Decomp (Fun f S)]) \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t'" proof (induction t' rule: intruder_synth_induct) case (AxiomC t') hence "t' \ (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t D \ t' \ ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun f T)]" by (simp add: ik\<^sub>e\<^sub>s\<^sub>t_append) thus ?case proof assume "t' \ ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun f T)]" hence "t' \ ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun f S)] \\<^sub>s\<^sub>e\<^sub>t \" using decomp_ik \Ana (Fun f T) = (K,M)\ \Ana (Fun f S) = (KS,MS)\ \M = MS \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \\ by simp thus ?case using ideduct_synth_mono[ OF intruder_synth.AxiomC[of t' "ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun f S)] \\<^sub>s\<^sub>e\<^sub>t \"], of "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t (D'@[Decomp (Fun f S)]) \\<^sub>s\<^sub>e\<^sub>t \)"] by (auto simp add: ik\<^sub>e\<^sub>s\<^sub>t_append) next assume "t' \ (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t D" hence "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t'" by (metis D'(2) intruder_synth.AxiomC) hence "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun f S)] \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t'" by (simp add: ideduct_synth_mono) thus ?case using ik\<^sub>e\<^sub>s\<^sub>t_append[of D' "[Decomp (Fun f S)]"] image_Un[of "\x. x \ \" "ik\<^sub>e\<^sub>s\<^sub>t D'" "ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun f S)]"] by (simp add: sup_aci(2)) qed qed auto } ultimately show ?thesis using D'' by auto qed next case (Fun g S) \ \Hence \Decomp (Fun f T)\ can be substituted for \Decomp (Fun g S)\\ hence KM: "K = Ks \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" "M = Ms \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \" "set K = set Ks \\<^sub>s\<^sub>e\<^sub>t \" "set M = set Ms \\<^sub>s\<^sub>e\<^sub>t \" using fT_s_in(2) \Ana (Fun f T) = (K,M)\ Ana_s s(2) Ana_invar_substD[OF assms(5), of g S] by auto hence Ms_nonempty: "Ms \ []" using \M \ []\ by auto { fix t' assume "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t (D@[Decomp (Fun f T)]) \\<^sub>c t'" hence "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t (D'@[Decomp (Fun g S)]) \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t'" using AxiomC proof (induction t' rule: intruder_synth_induct) case (AxiomC t') hence "t' \ ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \ \ t' \ ik\<^sub>e\<^sub>s\<^sub>t D \ t' \ set M" by (simp add: decomp_ik ik\<^sub>e\<^sub>s\<^sub>t_append) thus ?case proof (elim disjE) assume "t' \ ik\<^sub>e\<^sub>s\<^sub>t D" hence *: "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t'" using D'(2) by simp show ?case by (auto intro: ideduct_synth_mono[OF *] simp add: ik\<^sub>e\<^sub>s\<^sub>t_append_subst(2)) next assume "t' \ set M" hence "t' \ ik\<^sub>e\<^sub>s\<^sub>t [Decomp (Fun g S)] \\<^sub>s\<^sub>e\<^sub>t \" using KM(2) Fun decomp_ik[OF Ana_s] by auto thus ?case by (simp add: image_Un ik\<^sub>e\<^sub>s\<^sub>t_append) qed (simp add: ideduct_synth_mono[OF intruder_synth.AxiomC]) qed auto } thus ?thesis using s Fun Ana_s AxiomC.prems(1) KM(3) fT_s_in decomps\<^sub>e\<^sub>s\<^sub>t.Decomp[OF D'(1) _ _ Ms_nonempty, of g S Ks] by (metis AxiomC.hyps image_Un image_eqI intruder_synth.AxiomC) qed next case (ComposeC T f) have *: "\m. m \ set M \ (ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c m" using Ana_fun_subterm[OF \Ana (Fun f T) = (K, M)\] ComposeC.hyps(3) by auto - + have **: "ik\<^sub>e\<^sub>s\<^sub>t (D@[Decomp (Fun f T)]) = ik\<^sub>e\<^sub>s\<^sub>t D \ set M" using decomp_ik[OF \Ana (Fun f T) = (K, M)\] ik\<^sub>e\<^sub>s\<^sub>t_append by auto { fix t' assume "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ ik\<^sub>e\<^sub>s\<^sub>t (D@[Decomp (Fun f T)]) \\<^sub>c t'" hence "(ik\<^sub>e\<^sub>s\<^sub>t A \\<^sub>s\<^sub>e\<^sub>t \) \ (ik\<^sub>e\<^sub>s\<^sub>t D' \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>c t'" by (induct rule: intruder_synth_induct) (auto simp add: D'(2) * **) } thus ?case using D'(1) by auto qed qed thus ?thesis using D(2) assms(1) by (auto simp add: ik\<^sub>e\<^sub>s\<^sub>t_append_subst(2)) qed private lemma wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_nil: assumes "wf\<^sub>s\<^sub>t\<^sub>s' \ \" shows "wf\<^sub>s\<^sub>t\<^sub>s' (update\<^sub>s\<^sub>t \ []) \" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by auto private lemma wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_snd: assumes "wf\<^sub>s\<^sub>t\<^sub>s' \ \" "send\t\\<^sub>s\<^sub>t#S \ \" shows "wf\<^sub>s\<^sub>t\<^sub>s' (update\<^sub>s\<^sub>t \ (send\t\\<^sub>s\<^sub>t#S)) (\@[Step (receive\t\\<^sub>s\<^sub>t)])" unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def proof (intro conjI) let ?S = "send\t\\<^sub>s\<^sub>t#S" let ?A = "\@[Step (receive\t\\<^sub>s\<^sub>t)]" have \: "\S'. S' \ update\<^sub>s\<^sub>t \ ?S \ S' = S \ S' \ \" by auto have 1: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \) (dual\<^sub>s\<^sub>t S)" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by auto moreover have 2: "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A = wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \ \ fv t" using wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_split(2) by (auto simp add: Un_assoc) ultimately have 3: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by (metis wf_vars_mono) have 4: "\S \ \. \S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by simp have "wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" using 1 2 3 assms(2) by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by (metis 3 \) - + have "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" "\S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t S = {}" using 4 assms(2) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by force+ thus "\S \ update\<^sub>s\<^sub>t \ ?S. \S' \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" by (metis 4 \) have "\S' \ \. fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t ?S = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by metis+ hence 5: "fv\<^sub>e\<^sub>s\<^sub>t ?A = fv\<^sub>e\<^sub>s\<^sub>t \ \ fv t" "bvars\<^sub>e\<^sub>s\<^sub>t ?A = bvars\<^sub>e\<^sub>s\<^sub>t \" "\S' \ \. fv t \ bvars\<^sub>s\<^sub>t S' = {}" using to_st_append by fastforce+ have *: "\S \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using 5 assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fast hence "fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using assms(2) by metis hence "fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by (metis * \) have **: "\S \ \. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" using 5 assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fast hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t ?S = {}" using assms(2) by metis hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by fastforce thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by (metis ** \) qed private lemma wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_rcv: assumes "wf\<^sub>s\<^sub>t\<^sub>s' \ \" "receive\t\\<^sub>s\<^sub>t#S \ \" shows "wf\<^sub>s\<^sub>t\<^sub>s' (update\<^sub>s\<^sub>t \ (receive\t\\<^sub>s\<^sub>t#S)) (\@[Step (send\t\\<^sub>s\<^sub>t)])" unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def proof (intro conjI) let ?S = "receive\t\\<^sub>s\<^sub>t#S" let ?A = "\@[Step (send\t\\<^sub>s\<^sub>t)]" have \: "\S'. S' \ update\<^sub>s\<^sub>t \ ?S \ S' = S \ S' \ \" by auto have 1: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \) (dual\<^sub>s\<^sub>t S)" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by auto moreover have 2: "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A = wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \ \ fv t" using wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_split(2) by (auto simp add: Un_assoc) ultimately have 3: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by (metis wf_vars_mono) have 4: "\S \ \. \S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by simp have "wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" using 1 2 3 assms(2) by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by (metis 3 \) have "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" "\S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t S = {}" using 4 assms(2) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by force+ thus "\S \ update\<^sub>s\<^sub>t \ ?S. \S' \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" by (metis 4 \) have "\S' \ \. fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t ?S = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by metis+ hence 5: "fv\<^sub>e\<^sub>s\<^sub>t ?A = fv\<^sub>e\<^sub>s\<^sub>t \ \ fv t" "bvars\<^sub>e\<^sub>s\<^sub>t ?A = bvars\<^sub>e\<^sub>s\<^sub>t \" "\S' \ \. fv t \ bvars\<^sub>s\<^sub>t S' = {}" using to_st_append by fastforce+ have *: "\S \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using 5 assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fast hence "fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using assms(2) by metis hence "fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by (metis * \) have **: "\S \ \. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" using 5 assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fast hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t ?S = {}" using assms(2) by metis hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by fastforce thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by (metis ** \) qed private lemma wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_eq: assumes "wf\<^sub>s\<^sub>t\<^sub>s' \ \" "\a: t \ t'\\<^sub>s\<^sub>t#S \ \" shows "wf\<^sub>s\<^sub>t\<^sub>s' (update\<^sub>s\<^sub>t \ (\a: t \ t'\\<^sub>s\<^sub>t#S)) (\@[Step (\a: t \ t'\\<^sub>s\<^sub>t)])" unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def proof (intro conjI) let ?S = "\a: t \ t'\\<^sub>s\<^sub>t#S" let ?A = "\@[Step (\a: t \ t'\\<^sub>s\<^sub>t)]" have \: "\S'. S' \ update\<^sub>s\<^sub>t \ ?S \ S' = S \ S' \ \" by auto have 1: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \) (dual\<^sub>s\<^sub>t S)" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by auto moreover have 2: "a = Assign \ wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A = wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \ \ fv t \ fv t'" "a = Check \ wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A = wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \" using wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_split(2) by (auto simp add: Un_assoc) ultimately have 3: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by (cases a) (metis wf_vars_mono, metis) have 4: "\S \ \. \S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by simp have "wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" using 1 2 3 assms(2) by (cases a) auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by (metis 3 \) have "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" "\S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t S = {}" using 4 assms(2) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by force+ thus "\S \ update\<^sub>s\<^sub>t \ ?S. \S' \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" by (metis 4 \) have "\S' \ \. fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t ?S = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by metis+ hence 5: "fv\<^sub>e\<^sub>s\<^sub>t ?A = fv\<^sub>e\<^sub>s\<^sub>t \ \ fv t \ fv t'" "bvars\<^sub>e\<^sub>s\<^sub>t ?A = bvars\<^sub>e\<^sub>s\<^sub>t \" "\S' \ \. fv t \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv t' \ bvars\<^sub>s\<^sub>t S' = {}" using to_st_append by fastforce+ have *: "\S \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using 5 assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fast hence "fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using assms(2) by metis hence "fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by (metis * \) have **: "\S \ \. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" using 5 assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fast hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t ?S = {}" using assms(2) by metis hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by fastforce thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by (metis ** \) qed private lemma wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_ineq: assumes "wf\<^sub>s\<^sub>t\<^sub>s' \ \" "\X\\\: F\\<^sub>s\<^sub>t#S \ \" shows "wf\<^sub>s\<^sub>t\<^sub>s' (update\<^sub>s\<^sub>t \ (\X\\\: F\\<^sub>s\<^sub>t#S)) (\@[Step (\X\\\: F\\<^sub>s\<^sub>t)])" unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def proof (intro conjI) let ?S = "\X\\\: F\\<^sub>s\<^sub>t#S" let ?A = "\@[Step (\X\\\: F\\<^sub>s\<^sub>t)]" have \: "\S'. S' \ update\<^sub>s\<^sub>t \ ?S \ S' = S \ S' \ \" by auto have 1: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \) (dual\<^sub>s\<^sub>t S)" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by auto moreover have 2: "wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A = wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \" using wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_split(2) by (auto simp add: Un_assoc) ultimately have 3: "\S \ \. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by metis have 4: "\S \ \. \S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by simp have "wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" using 1 2 3 assms(2) by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t ?A) (dual\<^sub>s\<^sub>t S)" by (metis 3 \) have "fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S = {}" "\S' \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t S = {}" using 4 assms(2) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by force+ thus "\S \ update\<^sub>s\<^sub>t \ ?S. \S' \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>s\<^sub>t S' = {}" by (metis 4 \) have "\S' \ \. fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>s\<^sub>t S' = {}" "\S' \ \. fv\<^sub>s\<^sub>t S' \ bvars\<^sub>s\<^sub>t ?S = {}" using assms unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by metis+ moreover have "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X \ fv\<^sub>s\<^sub>t (\X\\\: F\\<^sub>s\<^sub>t # S)" by auto ultimately have 5: "\S' \ \. (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X) \ bvars\<^sub>s\<^sub>t S' = {}" - "fv\<^sub>e\<^sub>s\<^sub>t ?A = fv\<^sub>e\<^sub>s\<^sub>t \ \ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X)" "bvars\<^sub>e\<^sub>s\<^sub>t ?A = set X \ bvars\<^sub>e\<^sub>s\<^sub>t \" + "fv\<^sub>e\<^sub>s\<^sub>t ?A = fv\<^sub>e\<^sub>s\<^sub>t \ \ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X)" "bvars\<^sub>e\<^sub>s\<^sub>t ?A = set X \ bvars\<^sub>e\<^sub>s\<^sub>t \" "\S \ \. fv\<^sub>s\<^sub>t S \ set X = {}" using to_st_append by (blast, force, force, force) have *: "\S \ \. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using 5(3,4) assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by blast hence "fv\<^sub>s\<^sub>t ?S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" using assms(2) by metis hence "fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t ?A = {}" by (metis * \) have **: "\S \ \. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" using 5(1,2) assms(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fast hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t ?S = {}" using assms(2) by metis hence "fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by auto thus "\S \ update\<^sub>s\<^sub>t \ ?S. fv\<^sub>e\<^sub>s\<^sub>t ?A \ bvars\<^sub>s\<^sub>t S = {}" by (metis ** \) qed private lemma trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq: assumes "x#S \ \" shows "\(trms\<^sub>s\<^sub>t ` update\<^sub>s\<^sub>t \ (x#S)) \ trms\<^sub>s\<^sub>t\<^sub>p x = \(trms\<^sub>s\<^sub>t ` \)" (is "?A = ?B") proof show "?B \ ?A" proof have "trms\<^sub>s\<^sub>t\<^sub>p x \ trms\<^sub>s\<^sub>t (x#S)" by auto hence "\t'. t' \ ?B \ t' \ trms\<^sub>s\<^sub>t\<^sub>p x \ t' \ ?A" by simp moreover { fix t' assume t': "t' \ ?B" "t' \ trms\<^sub>s\<^sub>t\<^sub>p x" then obtain S' where S': "t' \ trms\<^sub>s\<^sub>t S'" "S' \ \" by auto hence "S' = x#S \ S' \ update\<^sub>s\<^sub>t \ (x#S)" by auto moreover { assume "S' = x#S" hence "t' \ trms\<^sub>s\<^sub>t S" using S' t' by simp hence "t' \ ?A" by auto } ultimately have "t' \ ?A" using t' S' by auto } ultimately show "\t'. t' \ ?B \ t' \ ?A" by metis qed show "?A \ ?B" proof have "\t'. t' \ ?A \ t' \ trms\<^sub>s\<^sub>t\<^sub>p x \ trms\<^sub>s\<^sub>t\<^sub>p x \ ?B" using assms by force+ moreover { fix t' assume t': "t' \ ?A" "t' \ trms\<^sub>s\<^sub>t\<^sub>p x" then obtain S' where "t' \ trms\<^sub>s\<^sub>t S'" "S' \ update\<^sub>s\<^sub>t \ (x#S)" by auto hence "S' = S \ S' \ \" by auto moreover have "trms\<^sub>s\<^sub>t S \ ?B" using assms trms\<^sub>s\<^sub>t_cons[of x S] by blast ultimately have "t' \ ?B" using t' by fastforce } ultimately show "\t'. t' \ ?A \ t' \ ?B" by blast qed qed private lemma trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_snd: assumes "send\t\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (send\t\\<^sub>s\<^sub>t#S)" "\' = \@[Step (receive\t\\<^sub>s\<^sub>t)]" shows "(\(trms\<^sub>s\<^sub>t ` \)) \ (trms\<^sub>e\<^sub>s\<^sub>t \) = (\(trms\<^sub>s\<^sub>t ` \')) \ (trms\<^sub>e\<^sub>s\<^sub>t \')" proof - have "(trms\<^sub>e\<^sub>s\<^sub>t \') = (trms\<^sub>e\<^sub>s\<^sub>t \) \ {t}" "\(trms\<^sub>s\<^sub>t ` \') \ {t} = \(trms\<^sub>s\<^sub>t ` \)" using to_st_append trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq[OF assms(1)] assms(2,3) by auto thus ?thesis by (metis (no_types, lifting) Un_insert_left Un_insert_right sup_bot.right_neutral) qed private lemma trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_rcv: assumes "receive\t\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (receive\t\\<^sub>s\<^sub>t#S)" "\' = \@[Step (send\t\\<^sub>s\<^sub>t)]" shows "(\(trms\<^sub>s\<^sub>t ` \)) \ (trms\<^sub>e\<^sub>s\<^sub>t \) = (\(trms\<^sub>s\<^sub>t ` \')) \ (trms\<^sub>e\<^sub>s\<^sub>t \')" proof - have "(trms\<^sub>e\<^sub>s\<^sub>t \') = (trms\<^sub>e\<^sub>s\<^sub>t \) \ {t}" "\(trms\<^sub>s\<^sub>t ` \') \ {t} = \(trms\<^sub>s\<^sub>t ` \)" using to_st_append trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq[OF assms(1)] assms(2,3) by auto thus ?thesis - by (metis (no_types, lifting) Un_insert_left Un_insert_right sup_bot.right_neutral) + by (metis (no_types, lifting) Un_insert_left Un_insert_right sup_bot.right_neutral) qed private lemma trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_eq: assumes "\a: t \ t'\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (\a: t \ t'\\<^sub>s\<^sub>t#S)" "\' = \@[Step (\a: t \ t'\\<^sub>s\<^sub>t)]" shows "(\(trms\<^sub>s\<^sub>t ` \)) \ (trms\<^sub>e\<^sub>s\<^sub>t \) = (\(trms\<^sub>s\<^sub>t ` \')) \ (trms\<^sub>e\<^sub>s\<^sub>t \')" proof - have "(trms\<^sub>e\<^sub>s\<^sub>t \') = (trms\<^sub>e\<^sub>s\<^sub>t \) \ {t,t'}" "\(trms\<^sub>s\<^sub>t ` \') \ {t,t'} = \(trms\<^sub>s\<^sub>t ` \)" using to_st_append trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq[OF assms(1)] assms(2,3) by auto thus ?thesis by (metis (no_types, lifting) Un_insert_left Un_insert_right sup_bot.right_neutral) qed private lemma trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_ineq: assumes "\X\\\: F\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (\X\\\: F\\<^sub>s\<^sub>t#S)" "\' = \@[Step (\X\\\: F\\<^sub>s\<^sub>t)]" shows "(\(trms\<^sub>s\<^sub>t ` \)) \ (trms\<^sub>e\<^sub>s\<^sub>t \) = (\(trms\<^sub>s\<^sub>t ` \')) \ (trms\<^sub>e\<^sub>s\<^sub>t \')" proof - have "(trms\<^sub>e\<^sub>s\<^sub>t \') = (trms\<^sub>e\<^sub>s\<^sub>t \) \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" "\(trms\<^sub>s\<^sub>t ` \') \ trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F = \(trms\<^sub>s\<^sub>t ` \)" using to_st_append trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq[OF assms(1)] assms(2,3) by auto thus ?thesis by (simp add: Un_commute sup_left_commute) qed private lemma ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset: assumes "x#S \ \" shows "\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` (update\<^sub>s\<^sub>t \ (x#S))) \ \(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)" (is ?A) "\(assignment_rhs\<^sub>s\<^sub>t ` (update\<^sub>s\<^sub>t \ (x#S))) \ \(assignment_rhs\<^sub>s\<^sub>t ` \)" (is ?B) proof - { fix t assume "t \ \(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` (update\<^sub>s\<^sub>t \ (x#S)))" then obtain S' where S': "S' \ update\<^sub>s\<^sub>t \ (x#S)" "t \ ik\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t S')" by auto - + have *: "ik\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t S) \ ik\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t (x#S))" using ik_append[of "dual\<^sub>s\<^sub>t [x]" "dual\<^sub>s\<^sub>t S"] dual\<^sub>s\<^sub>t_append[of "[x]" S] by auto - + hence "t \ \(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)" proof (cases "S' = S") case True thus ?thesis using * assms S' by auto next case False thus ?thesis using S' by auto qed } moreover { fix t assume "t \ \(assignment_rhs\<^sub>s\<^sub>t ` (update\<^sub>s\<^sub>t \ (x#S)))" then obtain S' where S': "S' \ update\<^sub>s\<^sub>t \ (x#S)" "t \ assignment_rhs\<^sub>s\<^sub>t S'" by auto - + have "assignment_rhs\<^sub>s\<^sub>t S \ assignment_rhs\<^sub>s\<^sub>t (x#S)" using assignment_rhs_append[of "[x]" S] by simp hence "t \ \(assignment_rhs\<^sub>s\<^sub>t ` \)" using assms S' by (cases "S' = S") auto } ultimately show ?A ?B by (metis subsetI)+ qed private lemma ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_snd: assumes "send\t\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (send\t\\<^sub>s\<^sub>t#S)" "\' = \@[Step (receive\t\\<^sub>s\<^sub>t)]" shows "(\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \') \ (\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" (is ?A) "(\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \') \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" (is ?B) proof - { fix t' assume t'_in: "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \')" hence "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \) \ {t}" using assms ik\<^sub>e\<^sub>s\<^sub>t_append by auto moreover have "t \ \(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)" using assms(1) by force ultimately have "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) by auto } moreover { fix t' assume t'_in: "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \')" hence "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" using assms assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append by auto hence "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) by auto } ultimately show ?A ?B by (metis subsetI)+ qed private lemma ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_rcv: assumes "receive\t\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (receive\t\\<^sub>s\<^sub>t#S)" "\' = \@[Step (send\t\\<^sub>s\<^sub>t)]" shows "(\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \') \ (\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" (is ?A) "(\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \') \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" (is ?B) proof - { fix t' assume t'_in: "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \')" hence "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" using assms ik\<^sub>e\<^sub>s\<^sub>t_append by auto hence "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) by auto } moreover { fix t' assume t'_in: "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \')" hence "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" using assms assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append by auto hence "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) by auto } ultimately show ?A ?B by (metis subsetI)+ qed private lemma ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_eq: assumes "\a: t \ t'\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (\a: t \ t'\\<^sub>s\<^sub>t#S)" "\' = \@[Step (\a: t \ t'\\<^sub>s\<^sub>t)]" shows "(\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \') \ (\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" (is ?A) "(\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \') \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" (is ?B) proof - have 1: "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" when "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \')" for t' proof - have "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" using that assms ik\<^sub>e\<^sub>s\<^sub>t_append by auto thus ?thesis using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) by auto qed have 2: "t'' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" when "t'' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \')" "a = Assign" for t'' proof - have "t'' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \) \ {t'}" using that assms assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append by auto moreover have "t' \ \(assignment_rhs\<^sub>s\<^sub>t ` \)" using assms(1) that by force ultimately show ?thesis using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) that by auto qed have 3: "assignment_rhs\<^sub>e\<^sub>s\<^sub>t \' = assignment_rhs\<^sub>e\<^sub>s\<^sub>t \" (is ?C) "(\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (\(assignment_rhs\<^sub>s\<^sub>t ` \))" (is ?D) when "a = Check" proof - show ?C using that assms(2,3) by (simp add: assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append) - show ?D using assms(1,2,3) ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset(2) by auto + show ?D using assms(1,2,3) ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset(2) by auto qed show ?A using 1 2 by (metis subsetI) show ?B using 1 2 3 by (cases a) blast+ qed private lemma ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_ineq: assumes "\X\\\: F\\<^sub>s\<^sub>t#S \ \" "\' = update\<^sub>s\<^sub>t \ (\X\\\: F\\<^sub>s\<^sub>t#S)" "\' = \@[Step (\X\\\: F\\<^sub>s\<^sub>t)]" shows "(\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \') \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" (is ?A) "(\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \') \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" (is ?B) proof - { fix t' assume t'_in: "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \')" hence "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \')) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" using assms ik\<^sub>e\<^sub>s\<^sub>t_append by auto hence "t' \ (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \)) \ (ik\<^sub>e\<^sub>s\<^sub>t \)" using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) by auto } moreover { fix t' assume t'_in: "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \')" hence "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \')) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" using assms assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append by auto hence "t' \ (\(assignment_rhs\<^sub>s\<^sub>t ` \)) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)" using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset[OF assms(1)] assms(2) by auto } ultimately show ?A ?B by (metis subsetI)+ qed subsubsection \Transition Systems Definitions\ inductive pts_symbolic:: "(('fun,'var) strands \ ('fun,'var) strand) \ (('fun,'var) strands \ ('fun,'var) strand) \ bool" (infix "\\<^sup>\" 50) where Nil[simp]: "[] \ \ \ (\,\) \\<^sup>\ (update\<^sub>s\<^sub>t \ [],\)" | Send[simp]: "send\t\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\ (update\<^sub>s\<^sub>t \ (send\t\\<^sub>s\<^sub>t#S),\@[receive\t\\<^sub>s\<^sub>t])" | Receive[simp]: "receive\t\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\ (update\<^sub>s\<^sub>t \ (receive\t\\<^sub>s\<^sub>t#S),\@[send\t\\<^sub>s\<^sub>t])" | Equality[simp]: "\a: t \ t'\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\ (update\<^sub>s\<^sub>t \ (\a: t \ t'\\<^sub>s\<^sub>t#S),\@[\a: t \ t'\\<^sub>s\<^sub>t])" | Inequality[simp]: "\X\\\: F\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\ (update\<^sub>s\<^sub>t \ (\X\\\: F\\<^sub>s\<^sub>t#S),\@[\X\\\: F\\<^sub>s\<^sub>t])" private inductive pts_symbolic_c:: "(('fun,'var) strands \ ('fun,'var) extstrand) \ (('fun,'var) strands \ ('fun,'var) extstrand) \ bool" (infix "\\<^sup>\\<^sub>c" 50) where Nil[simp]: "[] \ \ \ (\,\) \\<^sup>\\<^sub>c (update\<^sub>s\<^sub>t \ [],\)" | Send[simp]: "send\t\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\\<^sub>c (update\<^sub>s\<^sub>t \ (send\t\\<^sub>s\<^sub>t#S),\@[Step (receive\t\\<^sub>s\<^sub>t)])" | Receive[simp]: "receive\t\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\\<^sub>c (update\<^sub>s\<^sub>t \ (receive\t\\<^sub>s\<^sub>t#S),\@[Step (send\t\\<^sub>s\<^sub>t)])" | Equality[simp]: "\a: t \ t'\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\\<^sub>c (update\<^sub>s\<^sub>t \ (\a: t \ t'\\<^sub>s\<^sub>t#S),\@[Step (\a: t \ t'\\<^sub>s\<^sub>t)])" | Inequality[simp]: "\X\\\: F\\<^sub>s\<^sub>t#S \ \ \ (\,\) \\<^sup>\\<^sub>c (update\<^sub>s\<^sub>t \ (\X\\\: F\\<^sub>s\<^sub>t#S),\@[Step (\X\\\: F\\<^sub>s\<^sub>t)])" | Decompose[simp]: "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \ \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \) \ (\,\) \\<^sup>\\<^sub>c (\,\@[Decomp (Fun f T)])" abbreviation pts_symbolic_rtrancl (infix "\\<^sup>\\<^sup>*" 50) where "a \\<^sup>\\<^sup>* b \ pts_symbolic\<^sup>*\<^sup>* a b" private abbreviation pts_symbolic_c_rtrancl (infix "\\<^sup>\\<^sub>c\<^sup>*" 50) where "a \\<^sup>\\<^sub>c\<^sup>* b \ pts_symbolic_c\<^sup>*\<^sup>* a b" lemma pts_symbolic_induct[consumes 1, case_names Nil Send Receive Equality Inequality]: assumes "(\,\) \\<^sup>\ (\',\')" and "\[] \ \; \' = update\<^sub>s\<^sub>t \ []; \' = \\ \ P" and "\t S. \send\t\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (send\t\\<^sub>s\<^sub>t#S); \' = \@[receive\t\\<^sub>s\<^sub>t]\ \ P" and "\t S. \receive\t\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (receive\t\\<^sub>s\<^sub>t#S); \' = \@[send\t\\<^sub>s\<^sub>t]\ \ P" and "\a t t' S. \\a: t \ t'\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (\a: t \ t'\\<^sub>s\<^sub>t#S); \' = \@[\a: t \ t'\\<^sub>s\<^sub>t]\ \ P" and "\X F S. \\X\\\: F\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (\X\\\: F\\<^sub>s\<^sub>t#S); \' = \@[\X\\\: F\\<^sub>s\<^sub>t]\ \ P" shows "P" apply (rule pts_symbolic.cases[OF assms(1)]) using assms(2,3,4,5,6) by simp_all private lemma pts_symbolic_c_induct[consumes 1, case_names Nil Send Receive Equality Inequality Decompose]: assumes "(\,\) \\<^sup>\\<^sub>c (\',\')" and "\[] \ \; \' = update\<^sub>s\<^sub>t \ []; \' = \\ \ P" and "\t S. \send\t\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (send\t\\<^sub>s\<^sub>t#S); \' = \@[Step (receive\t\\<^sub>s\<^sub>t)]\ \ P" and "\t S. \receive\t\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (receive\t\\<^sub>s\<^sub>t#S); \' = \@[Step (send\t\\<^sub>s\<^sub>t)]\ \ P" and "\a t t' S. \\a: t \ t'\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (\a: t \ t'\\<^sub>s\<^sub>t#S); \' = \@[Step (\a: t \ t'\\<^sub>s\<^sub>t)]\ \ P" and "\X F S. \\X\\\: F\\<^sub>s\<^sub>t#S \ \; \' = update\<^sub>s\<^sub>t \ (\X\\\: F\\<^sub>s\<^sub>t#S); \' = \@[Step (\X\\\: F\\<^sub>s\<^sub>t)]\ \ P" and "\f T. \Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \ \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \); \' = \; \' = \@[Decomp (Fun f T)]\ \ P" shows "P" apply (rule pts_symbolic_c.cases[OF assms(1)]) using assms(2,3,4,5,6,7) by simp_all private lemma pts_symbolic_c_preserves_wf_prot: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" "wf\<^sub>s\<^sub>t\<^sub>s' \ \" shows "wf\<^sub>s\<^sub>t\<^sub>s' \' \'" using assms proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) from step.hyps(2) step.IH[OF step.prems] show ?case proof (induction rule: pts_symbolic_c_induct) case Decompose hence "fv\<^sub>e\<^sub>s\<^sub>t \2 = fv\<^sub>e\<^sub>s\<^sub>t \1" "bvars\<^sub>e\<^sub>s\<^sub>t \2 = bvars\<^sub>e\<^sub>s\<^sub>t \1" using bvars_decomp ik_assignment_rhs_decomp_fv by metis+ thus ?case using Decompose unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by (metis wf_vars_mono wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_split(2)) qed (metis wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_nil, metis wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_snd, metis wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_rcv, metis wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_eq, metis wf\<^sub>s\<^sub>t\<^sub>s'_update\<^sub>s\<^sub>t_ineq) qed metis private lemma pts_symbolic_c_preserves_wf_is: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" "wf\<^sub>s\<^sub>t\<^sub>s' \ \" "wf\<^sub>s\<^sub>t V (to_st \)" shows "wf\<^sub>s\<^sub>t V (to_st \')" using assms proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) hence "(\, \) \\<^sup>\\<^sub>c\<^sup>* (\2, \2)" by auto hence *: "wf\<^sub>s\<^sub>t\<^sub>s' \1 \1" "wf\<^sub>s\<^sub>t\<^sub>s' \2 \2" using pts_symbolic_c_preserves_wf_prot[OF _ step.prems(1)] step.hyps(1) by auto from step.hyps(2) step.IH[OF step.prems] show ?case proof (induction rule: pts_symbolic_c_induct) case Nil thus ?case by auto next case (Send t S) hence "wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \1) (receive\t\\<^sub>s\<^sub>t#(dual\<^sub>s\<^sub>t S))" using *(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fastforce hence "fv t \ wfrestrictedvars\<^sub>s\<^sub>t (to_st \1) \ V" using wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_eq_wfrestrictedvars\<^sub>s\<^sub>t by auto thus ?case using Send wf_rcv_append''' to_st_append by simp next case (Receive t) thus ?case using wf_snd_append to_st_append by simp next case (Equality a t t' S) hence "wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \1) (\a: t \ t'\\<^sub>s\<^sub>t#(dual\<^sub>s\<^sub>t S))" using *(1) unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by fastforce hence "fv t' \ wfrestrictedvars\<^sub>s\<^sub>t (to_st \1) \ V" when "a = Assign" using wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_eq_wfrestrictedvars\<^sub>s\<^sub>t that by auto thus ?case using Equality wf_eq_append''' to_st_append by (cases a) auto next case (Inequality t t' S) thus ?case using wf_ineq_append'' to_st_append by simp next case (Decompose f T) hence "fv (Fun f T) \ wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t \1" by (metis fv_subterms_set fv_subset subset_trans ik\<^sub>s\<^sub>t_assignment_rhs\<^sub>s\<^sub>t_wfrestrictedvars_subset) hence "vars\<^sub>s\<^sub>t (decomp (Fun f T)) \ wfrestrictedvars\<^sub>s\<^sub>t (to_st \1) \ V" using decomp_vars[of "Fun f T"] wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_eq_wfrestrictedvars\<^sub>s\<^sub>t[of \1] by auto thus ?case using to_st_append[of \1 "[Decomp (Fun f T)]"] wf_append_suffix[OF Decompose.prems] Decompose.hyps(3) by (metis append_Nil2 decomp_vars(1,2) to_st.simps(1,3)) qed qed metis private lemma pts_symbolic_c_preserves_tfr\<^sub>s\<^sub>e\<^sub>t: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" and "tfr\<^sub>s\<^sub>e\<^sub>t ((\(trms\<^sub>s\<^sub>t ` \)) \ (trms\<^sub>e\<^sub>s\<^sub>t \))" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s ((\(trms\<^sub>s\<^sub>t ` \)) \ (trms\<^sub>e\<^sub>s\<^sub>t \))" shows "tfr\<^sub>s\<^sub>e\<^sub>t ((\(trms\<^sub>s\<^sub>t ` \')) \ (trms\<^sub>e\<^sub>s\<^sub>t \')) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s ((\(trms\<^sub>s\<^sub>t ` \')) \ (trms\<^sub>e\<^sub>s\<^sub>t \'))" using assms proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) from step.hyps(2) step.IH[OF step.prems] show ?case proof (induction rule: pts_symbolic_c_induct) case Nil hence "\(trms\<^sub>s\<^sub>t ` \1) = \(trms\<^sub>s\<^sub>t ` \2)" by force thus ?case using Nil by metis next case (Decompose f T) obtain t where t: "t \ ik\<^sub>e\<^sub>s\<^sub>t \1 \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1" "Fun f T \ t" using Decompose.hyps(1) by auto have t_wf: "wf\<^sub>t\<^sub>r\<^sub>m t" using Decompose.prems wf_trm_subterm[of _ t] trms\<^sub>e\<^sub>s\<^sub>t_ik_assignment_rhsI[OF t(1)] unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by (metis UN_E Un_iff) have "t \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>e\<^sub>s\<^sub>t \1)" using trms\<^sub>e\<^sub>s\<^sub>t_ik_assignment_rhsI t by auto hence "Fun f T \ SMP (trms\<^sub>e\<^sub>s\<^sub>t \1)" - by (metis (no_types) SMP.MP SMP.Subterm UN_E t(2)) + by (metis (no_types) SMP.MP SMP.Subterm UN_E t(2)) hence "{Fun f T} \ SMP (trms\<^sub>e\<^sub>s\<^sub>t \1)" using SMP.Subterm[of "Fun f T"] by auto moreover have "trms\<^sub>e\<^sub>s\<^sub>t \2 = insert (Fun f T) (trms\<^sub>e\<^sub>s\<^sub>t \1)" using Decompose.hyps(3) by auto ultimately have *: "SMP (trms\<^sub>e\<^sub>s\<^sub>t \1) = SMP (trms\<^sub>e\<^sub>s\<^sub>t \2)" using SMP_subset_union_eq[of "{Fun f T}"] by (simp add: Un_commute) hence "SMP ((\(trms\<^sub>s\<^sub>t ` \1)) \ (trms\<^sub>e\<^sub>s\<^sub>t \1)) = SMP ((\(trms\<^sub>s\<^sub>t ` \2)) \ (trms\<^sub>e\<^sub>s\<^sub>t \2))" using Decompose.hyps(2) SMP_union by auto moreover have "\t \ trms\<^sub>e\<^sub>s\<^sub>t \1. wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" using Decompose.prems wf_trm_subterm t(2) t_wf unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by auto - hence "\t \ trms\<^sub>e\<^sub>s\<^sub>t \2. wf\<^sub>t\<^sub>r\<^sub>m t" by (metis * SMP.MP SMP_wf_trm) + hence "\t \ trms\<^sub>e\<^sub>s\<^sub>t \2. wf\<^sub>t\<^sub>r\<^sub>m t" by (metis * SMP.MP SMP_wf_trm) hence "\t \ (\(trms\<^sub>s\<^sub>t ` \2)) \ (trms\<^sub>e\<^sub>s\<^sub>t \2). wf\<^sub>t\<^sub>r\<^sub>m t" using Decompose.prems Decompose.hyps(2) unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by force - ultimately show ?thesis using Decompose.prems unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by presburger + ultimately show ?thesis using Decompose.prems unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by presburger qed (metis trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_snd, metis trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_rcv, metis trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_eq, metis trms\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_eq_ineq) qed metis private lemma pts_symbolic_c_preserves_tfr\<^sub>s\<^sub>t\<^sub>p: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" "\S \ \ \ {to_st \}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" shows "\S \ \' \ {to_st \'}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using assms proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) from step.hyps(2) step.IH[OF step.prems] show ?case proof (induction rule: pts_symbolic_c_induct) case Nil have 1: "\S \ {to_st \2}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Nil by simp have 2: "\2 = \1 - {[]}" "\S \ \1. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Nil by simp_all have "\S \ \2. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" proof fix S assume "S \ \2" hence "S \ \1" using 2(1) by simp thus "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using 2(2) by simp qed thus ?case using 1 by auto next case (Send t S) have 1: "\S \ {to_st \2}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Send by (simp add: to_st_append) have 2: "\2 = insert S (\1 - {send\t\\<^sub>s\<^sub>t#S})" "\S \ \1. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Send by simp_all have 3: "\S \ \2. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" proof fix S' assume "S' \ \2" hence "S' \ \1 \ S' = S" using 2(1) by auto moreover have "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Send.hyps 2(2) by auto ultimately show "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" using 2(2) by blast qed thus ?case using 1 by auto next case (Receive t S) have 1: "\S \ {to_st \2}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Receive by (simp add: to_st_append) have 2: "\2 = insert S (\1 - {receive\t\\<^sub>s\<^sub>t#S})" "\S \ \1. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Receive by simp_all have 3: "\S \ \2. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" proof fix S' assume "S' \ \2" hence "S' \ \1 \ S' = S" using 2(1) by auto moreover have "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Receive.hyps 2(2) by auto ultimately show "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" using 2(2) by blast qed show ?case using 1 3 by auto next case (Equality a t t' S) have 1: "to_st \2 = to_st \1@[\a: t \ t'\\<^sub>s\<^sub>t]" "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \1)" using Equality by (simp_all add: to_st_append) have 2: "list_all tfr\<^sub>s\<^sub>t\<^sub>p [\a: t \ t'\\<^sub>s\<^sub>t]" using Equality by fastforce have 3: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \2)" using tfr_stp_all_append[of "to_st \1" "[\a: t \ t'\\<^sub>s\<^sub>t]"] 1 2 by metis hence 4: "\S \ {to_st \2}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Equality by simp have 5: "\2 = insert S (\1 - {\a: t \ t'\\<^sub>s\<^sub>t#S})" "\S \ \1. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Equality by simp_all - have 6: "\S \ \2. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" + have 6: "\S \ \2. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" proof fix S' assume "S' \ \2" hence "S' \ \1 \ S' = S" using 5(1) by auto moreover have "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Equality.hyps 5(2) by auto ultimately show "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" using 5(2) by blast qed thus ?case using 4 by auto next case (Inequality X F S) have 1: "to_st \2 = to_st \1@[\X\\\: F\\<^sub>s\<^sub>t]" "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \1)" using Inequality by (simp_all add: to_st_append) have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (\X\\\: F\\<^sub>s\<^sub>t#S)" using Inequality(1,4) by blast hence 2: "list_all tfr\<^sub>s\<^sub>t\<^sub>p [\X\\\: F\\<^sub>s\<^sub>t]" by simp have 3: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \2)" using tfr_stp_all_append[of "to_st \1" "[\X\\\: F\\<^sub>s\<^sub>t]"] 1 2 by metis hence 4: "\S \ {to_st \2}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Inequality by simp have 5: "\2 = insert S (\1 - {\X\\\: F\\<^sub>s\<^sub>t#S})" "\S \ \1. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Inequality by simp_all have 6: "\S \ \2. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" proof fix S' assume "S' \ \2" hence "S' \ \1 \ S' = S" using 5(1) by auto moreover have "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Inequality.hyps 5(2) by auto ultimately show "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" using 5(2) by blast qed thus ?case using 4 by auto next case (Decompose f T) hence 1: "\S \ \2. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" by blast have 2: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \1)" "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st [Decomp (Fun f T)])" using Decompose.prems decomp_tfr\<^sub>s\<^sub>t\<^sub>p by auto hence "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \1@to_st [Decomp (Fun f T)])" by auto hence "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \2)" using Decompose.hyps(3) to_st_append[of \1 "[Decomp (Fun f T)]"] by auto thus ?case using 1 by blast qed qed private lemma pts_symbolic_c_preserves_well_analyzed: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" "well_analyzed \" shows "well_analyzed \'" using assms proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) from step.hyps(2) step.IH[OF step.prems] show ?case proof (induction rule: pts_symbolic_c_induct) case Receive thus ?case by (metis well_analyzed_singleton(1) well_analyzed_append) next case Send thus ?case by (metis well_analyzed_singleton(2) well_analyzed_append) next case Equality thus ?case by (metis well_analyzed_singleton(3) well_analyzed_append) next case Inequality thus ?case by (metis well_analyzed_singleton(4) well_analyzed_append) next case (Decompose f T) hence "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \1 \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1) - (Var`\)" by auto thus ?case by (metis well_analyzed.Decomp Decompose.prems Decompose.hyps(3)) qed simp qed metis private lemma pts_symbolic_c_preserves_Ana_invar_subst: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" and "Ana_invar_subst ( (\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \) \ (ik\<^sub>e\<^sub>s\<^sub>t \)) \ (\(assignment_rhs\<^sub>s\<^sub>t ` \) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \)))" shows "Ana_invar_subst ( (\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \') \ (ik\<^sub>e\<^sub>s\<^sub>t \')) \ (\(assignment_rhs\<^sub>s\<^sub>t ` \') \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \')))" using assms proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) from step.hyps(2) step.IH[OF step.prems] show ?case proof (induction rule: pts_symbolic_c_induct) case Nil hence "\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \1) = \(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \2)" "\(assignment_rhs\<^sub>s\<^sub>t ` \1) = \(assignment_rhs\<^sub>s\<^sub>t ` \2)" by force+ thus ?case using Nil by metis - next + next case Send show ?case using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_snd[OF Send.hyps] Ana_invar_subst_subset[OF Send.prems] by (metis Un_mono) next case Receive show ?case using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_rcv[OF Receive.hyps] Ana_invar_subst_subset[OF Receive.prems] by (metis Un_mono) next case Equality show ?case using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_eq[OF Equality.hyps] Ana_invar_subst_subset[OF Equality.prems] by (metis Un_mono) next case Inequality show ?case using ik\<^sub>s\<^sub>t_update\<^sub>s\<^sub>t_subset_ineq[OF Inequality.hyps] Ana_invar_subst_subset[OF Inequality.prems] by (metis Un_mono) next case (Decompose f T) let ?X = "\(assignment_rhs\<^sub>s\<^sub>t`\2) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \2" let ?Y = "\(assignment_rhs\<^sub>s\<^sub>t`\1) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1" obtain K M where Ana: "Ana (Fun f T) = (K,M)" by moura hence *: "ik\<^sub>e\<^sub>s\<^sub>t \2 = ik\<^sub>e\<^sub>s\<^sub>t \1 \ set M" "assignment_rhs\<^sub>e\<^sub>s\<^sub>t \2 = assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1" using ik\<^sub>e\<^sub>s\<^sub>t_append assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append decomp_ik decomp_assignment_rhs_empty Decompose.hyps(3) by auto { fix g S assume "Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t`\2) \ ik\<^sub>e\<^sub>s\<^sub>t \2 \ ?X)" hence "Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \1) \ ik\<^sub>e\<^sub>s\<^sub>t \1 \ set M \ ?X)" using * Decompose.hyps(2) by auto hence "Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \1)) \ Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \1) \ Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (set M) \ Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (\(assignment_rhs\<^sub>s\<^sub>t`\1)) \ Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1)" using Decompose * Ana_fun_subterm[OF Ana] by auto moreover have "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \1 \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1)" - using trms\<^sub>e\<^sub>s\<^sub>t_ik_subtermsI Decompose.hyps(1) by auto + using trms\<^sub>e\<^sub>s\<^sub>t_ik_subtermsI Decompose.hyps(1) by auto hence "subterms (Fun f T) \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \1 \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1)" by (metis in_subterms_subset_Union) hence "subterms\<^sub>s\<^sub>e\<^sub>t (set M) \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \1 \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1)" by (meson Un_upper2 Ana_subterm[OF Ana] subterms_subset_set psubsetE subset_trans) ultimately have "Fun g S \ subterms\<^sub>s\<^sub>e\<^sub>t (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t ` \1) \ ik\<^sub>e\<^sub>s\<^sub>t \1 \ ?Y)" by auto } thus ?case using Decompose unfolding Ana_invar_subst_def by metis qed qed private lemma pts_symbolic_c_preserves_constr_disj_vars: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" "wf\<^sub>s\<^sub>t\<^sub>s' \ \" "fv\<^sub>e\<^sub>s\<^sub>t \ \ bvars\<^sub>e\<^sub>s\<^sub>t \ = {}" shows "fv\<^sub>e\<^sub>s\<^sub>t \' \ bvars\<^sub>e\<^sub>s\<^sub>t \' = {}" using assms proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) have *: "\S. S \ \1 \ fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t \1 = {}" "\S. S \ \1 \ fv\<^sub>e\<^sub>s\<^sub>t \1 \ bvars\<^sub>s\<^sub>t S = {}" using pts_symbolic_c_preserves_wf_prot[OF step.hyps(1) step.prems(1)] unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def by auto from step.hyps(2) step.IH[OF step.prems] show ?case proof (induction rule: pts_symbolic_c_induct) case Nil thus ?case by auto - next + next case (Send t S) hence "fv\<^sub>e\<^sub>s\<^sub>t \2 = fv\<^sub>e\<^sub>s\<^sub>t \1 \ fv t" "bvars\<^sub>e\<^sub>s\<^sub>t \2 = bvars\<^sub>e\<^sub>s\<^sub>t \1" "fv\<^sub>s\<^sub>t (send\t\\<^sub>s\<^sub>t#S) = fv t \ fv\<^sub>s\<^sub>t S" using fv\<^sub>e\<^sub>s\<^sub>t_append bvars\<^sub>e\<^sub>s\<^sub>t_append by simp+ thus ?case using *(1)[OF Send(1)] Send(4) by auto next case (Receive t S) hence "fv\<^sub>e\<^sub>s\<^sub>t \2 = fv\<^sub>e\<^sub>s\<^sub>t \1 \ fv t" "bvars\<^sub>e\<^sub>s\<^sub>t \2 = bvars\<^sub>e\<^sub>s\<^sub>t \1" "fv\<^sub>s\<^sub>t (receive\t\\<^sub>s\<^sub>t#S) = fv t \ fv\<^sub>s\<^sub>t S" using fv\<^sub>e\<^sub>s\<^sub>t_append bvars\<^sub>e\<^sub>s\<^sub>t_append by simp+ thus ?case using *(1)[OF Receive(1)] Receive(4) by auto next case (Equality a t t' S) hence "fv\<^sub>e\<^sub>s\<^sub>t \2 = fv\<^sub>e\<^sub>s\<^sub>t \1 \ fv t \ fv t'" "bvars\<^sub>e\<^sub>s\<^sub>t \2 = bvars\<^sub>e\<^sub>s\<^sub>t \1" "fv\<^sub>s\<^sub>t (\a: t \ t'\\<^sub>s\<^sub>t#S) = fv t \ fv t' \ fv\<^sub>s\<^sub>t S" using fv\<^sub>e\<^sub>s\<^sub>t_append bvars\<^sub>e\<^sub>s\<^sub>t_append by fastforce+ thus ?case using *(1)[OF Equality(1)] Equality(4) by auto next case (Inequality X F S) hence "fv\<^sub>e\<^sub>s\<^sub>t \2 = fv\<^sub>e\<^sub>s\<^sub>t \1 \ (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X)" "bvars\<^sub>e\<^sub>s\<^sub>t \2 = bvars\<^sub>e\<^sub>s\<^sub>t \1 \ set X" "fv\<^sub>s\<^sub>t (\X\\\: F\\<^sub>s\<^sub>t#S) = (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X) \ fv\<^sub>s\<^sub>t S" using fv\<^sub>e\<^sub>s\<^sub>t_append bvars\<^sub>e\<^sub>s\<^sub>t_append strand_vars_split(3)[of "[\X\\\: F\\<^sub>s\<^sub>t]" S] by auto+ moreover have "fv\<^sub>e\<^sub>s\<^sub>t \1 \ set X = {}" using *(2)[OF Inequality(1)] by auto ultimately show ?case using *(1)[OF Inequality(1)] Inequality(4) by auto next case (Decompose f T) thus ?case using Decompose(3,4) bvars_decomp ik_assignment_rhs_decomp_fv[OF Decompose(1)] by auto qed qed subsubsection \Theorem: The Typing Result Lifted to the Transition System Level\ private lemma wf\<^sub>s\<^sub>t\<^sub>s'_decomp_rm: assumes "well_analyzed A" "wf\<^sub>s\<^sub>t\<^sub>s' S (decomp_rm\<^sub>e\<^sub>s\<^sub>t A)" shows "wf\<^sub>s\<^sub>t\<^sub>s' S A" unfolding wf\<^sub>s\<^sub>t\<^sub>s'_def proof (intro conjI) show "\S\S. wf\<^sub>s\<^sub>t (wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t A) (dual\<^sub>s\<^sub>t S)" by (metis (no_types) assms(2) wf\<^sub>s\<^sub>t\<^sub>s'_def wfrestrictedvars\<^sub>e\<^sub>s\<^sub>t_decomp_rm\<^sub>e\<^sub>s\<^sub>t_subset wf_vars_mono le_iff_sup) show "\Sa\S. \S'\S. fv\<^sub>s\<^sub>t Sa \ bvars\<^sub>s\<^sub>t S' = {}" by (metis assms(2) wf\<^sub>s\<^sub>t\<^sub>s'_def) show "\S\S. fv\<^sub>s\<^sub>t S \ bvars\<^sub>e\<^sub>s\<^sub>t A = {}" by (metis assms(2) wf\<^sub>s\<^sub>t\<^sub>s'_def bvars_decomp_rm) show "\S\S. fv\<^sub>e\<^sub>s\<^sub>t A \ bvars\<^sub>s\<^sub>t S = {}" by (metis assms wf\<^sub>s\<^sub>t\<^sub>s'_def well_analyzed_decomp_rm\<^sub>e\<^sub>s\<^sub>t_fv) qed private lemma decomps\<^sub>e\<^sub>s\<^sub>t_pts_symbolic_c: assumes "D \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \" shows "(S,A) \\<^sup>\\<^sub>c\<^sup>* (S,A@D)" using assms(1) proof (induction D rule: decomps\<^sub>e\<^sub>s\<^sub>t.induct) case (Decomp B f X K T) have "subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t A \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t A) \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t (A@B) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@B))" using ik\<^sub>e\<^sub>s\<^sub>t_append[of A B] assignment_rhs\<^sub>e\<^sub>s\<^sub>t_append[of A B] by auto hence "Fun f X \ subterms\<^sub>s\<^sub>e\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t (A@B) \ assignment_rhs\<^sub>e\<^sub>s\<^sub>t (A@B))" using Decomp.hyps by auto hence "(S,A@B) \\<^sup>\\<^sub>c (S,A@B@[Decomp (Fun f X)])" using pts_symbolic_c.Decompose[of f X "A@B"] by simp thus ?case using Decomp.IH rtrancl_into_rtrancl rtranclp_rtrancl_eq[of pts_symbolic_c "(S,A)" "(S,A@B)"] by auto qed simp private lemma pts_symbolic_to_pts_symbolic_c: assumes "(\,to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \\<^sub>d)) \\<^sup>\\<^sup>* (\',\')" "sem\<^sub>e\<^sub>s\<^sub>t_d {} \ (to_est \')" "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ \\<^sub>d" and wf: "wf\<^sub>s\<^sub>t\<^sub>s' \ (decomp_rm\<^sub>e\<^sub>s\<^sub>t \\<^sub>d)" "wf\<^sub>e\<^sub>s\<^sub>t {} \\<^sub>d" and tar: "Ana_invar_subst ((\(ik\<^sub>s\<^sub>t` dual\<^sub>s\<^sub>t` \) \ (ik\<^sub>e\<^sub>s\<^sub>t \\<^sub>d)) \ (\(assignment_rhs\<^sub>s\<^sub>t` \) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \\<^sub>d)))" and wa: "well_analyzed \\<^sub>d" and \: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\\\<^sub>d'. \' = to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \\<^sub>d') \ (\,\\<^sub>d) \\<^sup>\\<^sub>c\<^sup>* (\',\\<^sub>d') \ sem\<^sub>e\<^sub>s\<^sub>t_c {} \ \\<^sub>d'" using assms(1,2) proof (induction rule: rtranclp_induct2) case refl thus ?case using assms by auto next case (step \1 \1 \2 \2) have "sem\<^sub>e\<^sub>s\<^sub>t_d {} \ (to_est \1)" using step.hyps(2) step.prems by (induct rule: pts_symbolic_induct, metis, (metis sem\<^sub>e\<^sub>s\<^sub>t_d_split_left to_est_append)+) then obtain \1d where \1d: "\1 = to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \1d)" "(\, \\<^sub>d) \\<^sup>\\<^sub>c\<^sup>* (\1, \1d)" "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ \1d" using step.IH by moura show ?case using step.hyps(2) proof (induction rule: pts_symbolic_induct) case Nil hence "(\, \\<^sub>d) \\<^sup>\\<^sub>c\<^sup>* (\2, \1d)" using \1d pts_symbolic_c.Nil[OF Nil.hyps(1), of \1d] by simp thus ?case using \1d Nil by auto next case (Send t S) hence "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ (\1d@[Step (receive\t\\<^sub>s\<^sub>t)])" using sem\<^sub>e\<^sub>s\<^sub>t_c.Receive[OF \1d(3)] by simp moreover have "(\1, \1d) \\<^sup>\\<^sub>c (\2, \1d@[Step (receive\t\\<^sub>s\<^sub>t)])" using Send.hyps(2) pts_symbolic_c.Send[OF Send.hyps(1), of \1d] by simp moreover have "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t (\1d@[Step (receive\t\\<^sub>s\<^sub>t)])) = \2" - using Send.hyps(3) decomp_rm\<^sub>e\<^sub>s\<^sub>t_append \1d(1) by (simp add: to_st_append) - ultimately show ?case using \1d(2) by auto + using Send.hyps(3) decomp_rm\<^sub>e\<^sub>s\<^sub>t_append \1d(1) by (simp add: to_st_append) + ultimately show ?case using \1d(2) by auto next case (Equality a t t' S) hence "t \ \ = t' \ \" using step.prems sem\<^sub>e\<^sub>s\<^sub>t_d_eq_sem_st[of "{}" \ "to_est \2"] to_st_append to_est_append to_st_to_est_inv by auto hence "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ (\1d@[Step (\a: t \ t'\\<^sub>s\<^sub>t)])" using sem\<^sub>e\<^sub>s\<^sub>t_c.Equality[OF \1d(3)] by simp moreover have "(\1, \1d) \\<^sup>\\<^sub>c (\2, \1d@[Step (\a: t \ t'\\<^sub>s\<^sub>t)])" using Equality.hyps(2) pts_symbolic_c.Equality[OF Equality.hyps(1), of \1d] by simp moreover have "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t (\1d@[Step (\a: t \ t'\\<^sub>s\<^sub>t)])) = \2" - using Equality.hyps(3) decomp_rm\<^sub>e\<^sub>s\<^sub>t_append \1d(1) by (simp add: to_st_append) + using Equality.hyps(3) decomp_rm\<^sub>e\<^sub>s\<^sub>t_append \1d(1) by (simp add: to_st_append) ultimately show ?case using \1d(2) by auto next case (Inequality X F S) hence "ineq_model \ X F" using step.prems sem\<^sub>e\<^sub>s\<^sub>t_d_eq_sem_st[of "{}" \ "to_est \2"] to_st_append to_est_append to_st_to_est_inv by auto hence "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ (\1d@[Step (\X\\\: F\\<^sub>s\<^sub>t)])" using sem\<^sub>e\<^sub>s\<^sub>t_c.Inequality[OF \1d(3)] by simp moreover have "(\1, \1d) \\<^sup>\\<^sub>c (\2, \1d@[Step (\X\\\: F\\<^sub>s\<^sub>t)])" using Inequality.hyps(2) pts_symbolic_c.Inequality[OF Inequality.hyps(1), of \1d] by simp moreover have "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t (\1d@[Step (\X\\\: F\\<^sub>s\<^sub>t)])) = \2" - using Inequality.hyps(3) decomp_rm\<^sub>e\<^sub>s\<^sub>t_append \1d(1) by (simp add: to_st_append) + using Inequality.hyps(3) decomp_rm\<^sub>e\<^sub>s\<^sub>t_append \1d(1) by (simp add: to_st_append) ultimately show ?case using \1d(2) by auto next case (Receive t S) hence "ik\<^sub>s\<^sub>t \1 \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" using step.prems sem\<^sub>e\<^sub>s\<^sub>t_d_eq_sem_st[of "{}" \ "to_est \2"] strand_sem_split(4)[of "{}" \1 "[send\t\\<^sub>s\<^sub>t]" \] to_st_append to_est_append to_st_to_est_inv by auto moreover have "ik\<^sub>s\<^sub>t \1 \\<^sub>s\<^sub>e\<^sub>t \ \ ik\<^sub>e\<^sub>s\<^sub>t \1d \\<^sub>s\<^sub>e\<^sub>t \" using \1d(1) decomp_rm\<^sub>e\<^sub>s\<^sub>t_ik_subset by auto ultimately have *: "ik\<^sub>e\<^sub>s\<^sub>t \1d \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" using ideduct_mono by auto have "wf\<^sub>s\<^sub>t\<^sub>s' \ \\<^sub>d" by (rule wf\<^sub>s\<^sub>t\<^sub>s'_decomp_rm[OF wa assms(4)]) hence **: "wf\<^sub>e\<^sub>s\<^sub>t {} \1d" by (rule pts_symbolic_c_preserves_wf_is[OF \1d(2) _ assms(5)]) have "Ana_invar_subst (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t`\1) \ (ik\<^sub>e\<^sub>s\<^sub>t \1d) \ (\(assignment_rhs\<^sub>s\<^sub>t`\1) \ (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1d)))" using tar \1d(2) pts_symbolic_c_preserves_Ana_invar_subst by metis hence "Ana_invar_subst (ik\<^sub>e\<^sub>s\<^sub>t \1d)" "Ana_invar_subst (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1d)" using Ana_invar_subst_subset by blast+ moreover have "well_analyzed \1d" using pts_symbolic_c_preserves_well_analyzed[OF \1d(2) wa] by metis ultimately obtain D where D: "D \ decomps\<^sub>e\<^sub>s\<^sub>t (ik\<^sub>e\<^sub>s\<^sub>t \1d) (assignment_rhs\<^sub>e\<^sub>s\<^sub>t \1d) \" "ik\<^sub>e\<^sub>s\<^sub>t (\1d@D) \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" using decomps\<^sub>e\<^sub>s\<^sub>t_exist_subst[OF * \1d(3) ** assms(8)] unfolding Ana_invar_subst_def by auto - + have "(\, \\<^sub>d) \\<^sup>\\<^sub>c\<^sup>* (\1, \1d@D)" using \1d(2) decomps\<^sub>e\<^sub>s\<^sub>t_pts_symbolic_c[OF D(1), of \1] by auto hence "(\, \\<^sub>d) \\<^sup>\\<^sub>c\<^sup>* (\2, \1d@D@[Step (send\t\\<^sub>s\<^sub>t)])" using Receive(2) pts_symbolic_c.Receive[OF Receive.hyps(1), of "\1d@D"] by auto moreover have "\2 = to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t (\1d@D@[Step (send\t\\<^sub>s\<^sub>t)]))" using Receive.hyps(3) \1d(1) decomps\<^sub>e\<^sub>s\<^sub>t_decomp_rm\<^sub>e\<^sub>s\<^sub>t_empty[OF D(1)] decomp_rm\<^sub>e\<^sub>s\<^sub>t_append to_st_append by auto moreover have "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ (\1d@D@[Step (send\t\\<^sub>s\<^sub>t)])" using D(2) sem\<^sub>e\<^sub>s\<^sub>t_c.Send[OF sem\<^sub>e\<^sub>s\<^sub>t_c_decomps\<^sub>e\<^sub>s\<^sub>t_append[OF \1d(3) D(1)]] by simp ultimately show ?case by auto qed qed private lemma pts_symbolic_c_to_pts_symbolic: assumes "(\,\) \\<^sup>\\<^sub>c\<^sup>* (\',\')" "sem\<^sub>e\<^sub>s\<^sub>t_c {} \ \'" shows "(\,to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \)) \\<^sup>\\<^sup>* (\',to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \'))" "sem\<^sub>e\<^sub>s\<^sub>t_d {} \ (decomp_rm\<^sub>e\<^sub>s\<^sub>t \')" proof - show "(\,to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \)) \\<^sup>\\<^sup>* (\',to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \'))" using assms(1) proof (induction rule: rtranclp_induct2) case (step \1 \1 \2 \2) show ?case using step.hyps(2,1) step.IH proof (induction rule: pts_symbolic_c_induct) case Nil thus ?case using pts_symbolic.Nil[OF Nil.hyps(1), of "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \1)"] by simp next case (Send t S) thus ?case using pts_symbolic.Send[OF Send.hyps(1), of "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \1)"] by (simp add: decomp_rm\<^sub>e\<^sub>s\<^sub>t_append to_st_append) next case (Receive t S) thus ?case - using pts_symbolic.Receive[OF Receive.hyps(1), of "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \1)"] + using pts_symbolic.Receive[OF Receive.hyps(1), of "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \1)"] by (simp add: decomp_rm\<^sub>e\<^sub>s\<^sub>t_append to_st_append) next case (Equality a t t' S) thus ?case - using pts_symbolic.Equality[OF Equality.hyps(1), of "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \1)"] + using pts_symbolic.Equality[OF Equality.hyps(1), of "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \1)"] by (simp add: decomp_rm\<^sub>e\<^sub>s\<^sub>t_append to_st_append) next case (Inequality t t' S) thus ?case - using pts_symbolic.Inequality[OF Inequality.hyps(1), of "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \1)"] + using pts_symbolic.Inequality[OF Inequality.hyps(1), of "to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \1)"] by (simp add: decomp_rm\<^sub>e\<^sub>s\<^sub>t_append to_st_append) next case (Decompose t) thus ?case using decomp_rm\<^sub>e\<^sub>s\<^sub>t_append by simp qed qed simp qed (rule sem\<^sub>e\<^sub>s\<^sub>t_d_decomp_rm\<^sub>e\<^sub>s\<^sub>t_if_sem\<^sub>e\<^sub>s\<^sub>t_c[OF assms(2)]) private lemma pts_symbolic_to_pts_symbolic_c_from_initial: assumes "(\\<^sub>0,[]) \\<^sup>\\<^sup>* (\,\)" "\ \ \\\" "wf\<^sub>s\<^sub>t\<^sub>s' \\<^sub>0 []" and "Ana_invar_subst (\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \\<^sub>0) \ \(assignment_rhs\<^sub>s\<^sub>t ` \\<^sub>0))" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\\\<^sub>d. \ = to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \\<^sub>d) \ (\\<^sub>0,[]) \\<^sup>\\<^sub>c\<^sup>* (\,\\<^sub>d) \ (\ \\<^sub>c \to_st \\<^sub>d\)" using assms pts_symbolic_to_pts_symbolic_c[of \\<^sub>0 "[]" \ \ \] sem\<^sub>e\<^sub>s\<^sub>t_c_eq_sem_st[of "{}" \] sem\<^sub>e\<^sub>s\<^sub>t_d_eq_sem_st[of "{}" \] to_st_to_est_inv[of \] strand_sem_eq_defs by (auto simp add: constr_sem_c_def constr_sem_d_def simp del: subst_range.simps) private lemma pts_symbolic_c_to_pts_symbolic_from_initial: assumes "(\\<^sub>0,[]) \\<^sup>\\<^sub>c\<^sup>* (\,\)" "\ \\<^sub>c \to_st \\" shows "(\\<^sub>0,[]) \\<^sup>\\<^sup>* (\,to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \))" "\ \ \to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \)\" using assms pts_symbolic_c_to_pts_symbolic[of \\<^sub>0 "[]" \ \ \] sem\<^sub>e\<^sub>s\<^sub>t_c_eq_sem_st[of "{}" \] sem\<^sub>e\<^sub>s\<^sub>t_d_eq_sem_st[of "{}" \] strand_sem_eq_defs by (auto simp add: constr_sem_c_def constr_sem_d_def) private lemma to_st_trms_wf: assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>e\<^sub>s\<^sub>t A)" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t (to_st A))" using assms proof (induction A) case (Cons x A) hence IH: "\t \ trms\<^sub>s\<^sub>t (to_st A). wf\<^sub>t\<^sub>r\<^sub>m t" by auto with Cons show ?case proof (cases x) case (Decomp t) hence "wf\<^sub>t\<^sub>r\<^sub>m t" using Cons.prems by auto obtain K T where Ana_t: "Ana t = (K,T)" by moura hence "trms\<^sub>s\<^sub>t (decomp t) \ {t} \ set K \ set T" using decomp_set_unfold[OF Ana_t] by force moreover have "\t \ set T. wf\<^sub>t\<^sub>r\<^sub>m t" using Ana_subterm[OF Ana_t] \wf\<^sub>t\<^sub>r\<^sub>m t\ wf_trm_subterm by auto ultimately have "\t \ trms\<^sub>s\<^sub>t (decomp t). wf\<^sub>t\<^sub>r\<^sub>m t" using Ana_keys_wf'[OF Ana_t] \wf\<^sub>t\<^sub>r\<^sub>m t\ by auto thus ?thesis using IH Decomp by auto qed auto qed simp private lemma to_st_trms_SMP_subset: "trms\<^sub>s\<^sub>t (to_st A) \ SMP (trms\<^sub>e\<^sub>s\<^sub>t A)" proof fix t assume "t \ trms\<^sub>s\<^sub>t (to_st A)" thus "t \ SMP (trms\<^sub>e\<^sub>s\<^sub>t A)" proof (induction A) case (Cons x A) hence *: "t \ trms\<^sub>s\<^sub>t (to_st [x]) \ trms\<^sub>s\<^sub>t (to_st A)" using to_st_append[of "[x]" A] by auto have **: "trms\<^sub>s\<^sub>t (to_st A) \ trms\<^sub>s\<^sub>t (to_st (x#A))" "trms\<^sub>e\<^sub>s\<^sub>t A \ trms\<^sub>e\<^sub>s\<^sub>t (x#A)" using to_st_append[of "[x]" A] by auto show ?case proof (cases "t \ trms\<^sub>s\<^sub>t (to_st A)") case True thus ?thesis using Cons.IH SMP_mono[OF **(2)] by auto next case False hence ***: "t \ trms\<^sub>s\<^sub>t (to_st [x])" using * by auto thus ?thesis proof (cases x) case (Decomp t') hence ****: "t \ trms\<^sub>s\<^sub>t (decomp t')" "t' \ trms\<^sub>e\<^sub>s\<^sub>t (x#A)" using *** by auto obtain K T where Ana_t': "Ana t' = (K,T)" by moura hence "t \ {t'} \ set K \ set T" using decomp_set_unfold[OF Ana_t'] ****(1) by force moreover { assume "t = t'" hence ?thesis using SMP.MP[OF ****(2)] by simp } moreover { assume "t \ set K" hence ?thesis using SMP.Ana[OF SMP.MP[OF ****(2)] Ana_t'] by auto } moreover { assume "t \ set T" "t \ t'" hence "t \ t'" using Ana_subterm[OF Ana_t'] by blast hence ?thesis using SMP.Subterm[OF SMP.MP[OF ****(2)]] by auto - } + } ultimately show ?thesis using Decomp by auto qed auto qed qed simp qed private lemma to_st_trms_tfr\<^sub>s\<^sub>e\<^sub>t: assumes "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>e\<^sub>s\<^sub>t A)" shows "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t (to_st A))" proof - have *: "trms\<^sub>s\<^sub>t (to_st A) \ SMP (trms\<^sub>e\<^sub>s\<^sub>t A)" using to_st_trms_wf to_st_trms_SMP_subset assms unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by auto have "trms\<^sub>s\<^sub>t (to_st A) = trms\<^sub>s\<^sub>t (to_st A) \ trms\<^sub>e\<^sub>s\<^sub>t A" by (blast dest!: trms\<^sub>e\<^sub>s\<^sub>tD) hence "SMP (trms\<^sub>e\<^sub>s\<^sub>t A) = SMP (trms\<^sub>s\<^sub>t (to_st A))" using SMP_subset_union_eq[OF *] by auto thus ?thesis using * assms unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by presburger qed theorem wt_attack_if_tfr_attack_pts: assumes "wf\<^sub>s\<^sub>t\<^sub>s \\<^sub>0" "tfr\<^sub>s\<^sub>e\<^sub>t (\(trms\<^sub>s\<^sub>t ` \\<^sub>0))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\(trms\<^sub>s\<^sub>t ` \\<^sub>0))" "\S \ \\<^sub>0. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" and "Ana_invar_subst (\(ik\<^sub>s\<^sub>t ` dual\<^sub>s\<^sub>t ` \\<^sub>0) \ \(assignment_rhs\<^sub>s\<^sub>t ` \\<^sub>0))" and "(\\<^sub>0,[]) \\<^sup>\\<^sup>* (\,\)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "\ \ \\, Var\" shows "\\\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ (\\<^sub>\ \ \\, Var\) \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" proof - have "(\(trms\<^sub>s\<^sub>t ` \\<^sub>0)) \ (trms\<^sub>e\<^sub>s\<^sub>t []) = \(trms\<^sub>s\<^sub>t ` \\<^sub>0)" "to_st [] = []" "list_all tfr\<^sub>s\<^sub>t\<^sub>p []" using assms by simp_all hence *: "tfr\<^sub>s\<^sub>e\<^sub>t ((\(trms\<^sub>s\<^sub>t ` \\<^sub>0)) \ (trms\<^sub>e\<^sub>s\<^sub>t []))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s ((\(trms\<^sub>s\<^sub>t ` \\<^sub>0)) \ (trms\<^sub>e\<^sub>s\<^sub>t []))" "wf\<^sub>s\<^sub>t\<^sub>s' \\<^sub>0 []" "\S \ \\<^sub>0 \ {to_st []}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using assms wf\<^sub>s\<^sub>t\<^sub>s_wf\<^sub>s\<^sub>t\<^sub>s' by (metis, metis, metis, simp) obtain \\<^sub>d where \\<^sub>d: "\ = to_st (decomp_rm\<^sub>e\<^sub>s\<^sub>t \\<^sub>d)" "(\\<^sub>0,[]) \\<^sup>\\<^sub>c\<^sup>* (\,\\<^sub>d)" "\ \\<^sub>c \to_st \\<^sub>d\" using pts_symbolic_to_pts_symbolic_c_from_initial assms *(3) by metis hence "tfr\<^sub>s\<^sub>e\<^sub>t (\(trms\<^sub>s\<^sub>t ` \) \ (trms\<^sub>e\<^sub>s\<^sub>t \\<^sub>d))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\(trms\<^sub>s\<^sub>t ` \) \ (trms\<^sub>e\<^sub>s\<^sub>t \\<^sub>d))" using pts_symbolic_c_preserves_tfr\<^sub>s\<^sub>e\<^sub>t[OF _ *(1,2)] by blast+ hence "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>e\<^sub>s\<^sub>t \\<^sub>d)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>e\<^sub>s\<^sub>t \\<^sub>d)" unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by (metis DiffE DiffI SMP_union UnCI, metis UnCI) hence "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t (to_st \\<^sub>d))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t (to_st \\<^sub>d))" by (metis to_st_trms_tfr\<^sub>s\<^sub>e\<^sub>t, metis to_st_trms_wf) moreover have "wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r (to_st \\<^sub>d) Var" proof - have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range Var)" "subst_domain Var \ vars\<^sub>e\<^sub>s\<^sub>t \\<^sub>d = {}" "range_vars Var \ bvars\<^sub>e\<^sub>s\<^sub>t \\<^sub>d = {}" by (simp_all add: range_vars_alt_def) moreover have "wf\<^sub>e\<^sub>s\<^sub>t {} \\<^sub>d" using pts_symbolic_c_preserves_wf_is[OF \\<^sub>d(2) *(3), of "{}"] by auto moreover have "fv\<^sub>s\<^sub>t (to_st \\<^sub>d) \ bvars\<^sub>e\<^sub>s\<^sub>t \\<^sub>d = {}" using pts_symbolic_c_preserves_constr_disj_vars[OF \\<^sub>d(2)] assms(1) wf\<^sub>s\<^sub>t\<^sub>s_wf\<^sub>s\<^sub>t\<^sub>s' by fastforce ultimately show ?thesis unfolding wf\<^sub>c\<^sub>o\<^sub>n\<^sub>s\<^sub>t\<^sub>r_def wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by simp qed moreover have "list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \\<^sub>d)" using pts_symbolic_c_preserves_tfr\<^sub>s\<^sub>t\<^sub>p[OF \\<^sub>d(2) *(4)] by blast moreover have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range Var)" by simp_all ultimately obtain \\<^sub>\ where \\<^sub>\: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "\\<^sub>\ \\<^sub>c \to_st \\<^sub>d, Var\" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" using wt_attack_if_tfr_attack[OF assms(7) \\<^sub>d(3)] \tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t (to_st \\<^sub>d))\ \list_all tfr\<^sub>s\<^sub>t\<^sub>p (to_st \\<^sub>d)\ unfolding tfr\<^sub>s\<^sub>t_def by metis hence "\\<^sub>\ \ \\, Var\" using pts_symbolic_c_to_pts_symbolic_from_initial \\<^sub>d by metis thus ?thesis using \\<^sub>\(1,3,4) by metis qed subsubsection \Corollary: The Typing Result on the Level of Constraints\ text \There exists well-typed models of satisfiable type-flaw resistant constraints\ corollary wt_attack_if_tfr_attack_d: assumes "wf\<^sub>s\<^sub>t {} \" "fv\<^sub>s\<^sub>t \ \ bvars\<^sub>s\<^sub>t \ = {}" "tfr\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t \)" and "Ana_invar_subst (ik\<^sub>s\<^sub>t \ \ assignment_rhs\<^sub>s\<^sub>t \)" and "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "\ \ \\\" shows "\\\<^sub>\. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ (\\<^sub>\ \ \\\) \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\<^sub>\ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \\<^sub>\)" proof - { fix S A have "({S},A) \\<^sup>\\<^sup>* ({},A@dual\<^sub>s\<^sub>t S)" proof (induction S arbitrary: A) case Nil thus ?case using pts_symbolic.Nil[of "{[]}"] by auto next case (Cons x S) hence "({S}, A@dual\<^sub>s\<^sub>t [x]) \\<^sup>\\<^sup>* ({}, A@dual\<^sub>s\<^sub>t (x#S))" by (metis dual\<^sub>s\<^sub>t_append List.append_assoc List.append_Nil List.append_Cons) moreover have "({x#S}, A) \\<^sup>\ ({S}, A@dual\<^sub>s\<^sub>t [x])" using pts_symbolic.Send[of _ S "{x#S}"] pts_symbolic.Receive[of _ S "{x#S}"] pts_symbolic.Equality[of _ _ _ S "{x#S}"] pts_symbolic.Inequality[of _ _ S "{x#S}"] by (cases x) auto ultimately show ?case by simp qed } hence 0: "({dual\<^sub>s\<^sub>t \},[]) \\<^sup>\\<^sup>* ({},\)" using dual\<^sub>s\<^sub>t_self_inverse by (metis List.append_Nil) have "fv\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t \) \ bvars\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t \) = {}" using assms(2) dual\<^sub>s\<^sub>t_fv dual\<^sub>s\<^sub>t_bvars by metis+ hence 1: "wf\<^sub>s\<^sub>t\<^sub>s {dual\<^sub>s\<^sub>t \}" using assms(1,2) dual\<^sub>s\<^sub>t_self_inverse[of \] unfolding wf\<^sub>s\<^sub>t\<^sub>s_def by auto - + have "\(trms\<^sub>s\<^sub>t ` {\}) = trms\<^sub>s\<^sub>t \" "\(trms\<^sub>s\<^sub>t ` {dual\<^sub>s\<^sub>t \}) = trms\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t \)" by auto hence "tfr\<^sub>s\<^sub>e\<^sub>t (\(trms\<^sub>s\<^sub>t ` {\}))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\(trms\<^sub>s\<^sub>t ` {\}))" "(\(trms\<^sub>s\<^sub>t ` {\})) = \(trms\<^sub>s\<^sub>t ` {dual\<^sub>s\<^sub>t \})" using assms(3,4) unfolding tfr\<^sub>s\<^sub>t_def by (metis, metis, metis dual\<^sub>s\<^sub>t_trms_eq) hence 2: "tfr\<^sub>s\<^sub>e\<^sub>t (\(trms\<^sub>s\<^sub>t ` {dual\<^sub>s\<^sub>t \}))" and 3: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\(trms\<^sub>s\<^sub>t ` {dual\<^sub>s\<^sub>t \}))" by metis+ have 4: "\S \ {dual\<^sub>s\<^sub>t \}. list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using dual\<^sub>s\<^sub>t_tfr\<^sub>s\<^sub>t\<^sub>p assms(3) unfolding tfr\<^sub>s\<^sub>t_def by blast have "assignment_rhs\<^sub>s\<^sub>t \ = assignment_rhs\<^sub>s\<^sub>t (dual\<^sub>s\<^sub>t \)" by (induct \ rule: assignment_rhs\<^sub>s\<^sub>t.induct) auto hence 5: "Ana_invar_subst (\(ik\<^sub>s\<^sub>t`dual\<^sub>s\<^sub>t`{dual\<^sub>s\<^sub>t \}) \ \(assignment_rhs\<^sub>s\<^sub>t`{dual\<^sub>s\<^sub>t \}))" using assms(5) dual\<^sub>s\<^sub>t_self_inverse[of \] by auto show ?thesis by (rule wt_attack_if_tfr_attack_pts[OF 1 2 3 4 5 0 assms(6,7)]) qed end end end