diff --git a/thys/First_Order_Terms/Unification.thy b/thys/First_Order_Terms/Unification.thy --- a/thys/First_Order_Terms/Unification.thy +++ b/thys/First_Order_Terms/Unification.thy @@ -1,579 +1,579 @@ (* Author: Christian Sternagel Author: René Thiemann License: LGPL *) subsection \A Concrete Unification Algorithm\ theory Unification imports Abstract_Unification Option_Monad begin definition "decompose s t = (case (s, t) of (Fun f ss, Fun g ts) \ if f = g then zip_option ss ts else None | _ \ None)" lemma decompose_same_Fun[simp]: \<^marker>\contributor \Martin Desharnais\\ "decompose (Fun f ss) (Fun f ss) = Some (zip ss ss)" by (simp add: decompose_def) lemma decompose_Some [dest]: "decompose (Fun f ss) (Fun g ts) = Some E \ f = g \ length ss = length ts \ E = zip ss ts" by (cases "f = g") (auto simp: decompose_def) lemma decompose_None [dest]: "decompose (Fun f ss) (Fun g ts) = None \ f \ g \ length ss \ length ts" by (cases "f = g") (auto simp: decompose_def) text \Applying a substitution to a list of equations.\ definition subst_list :: "('f, 'v) subst \ ('f, 'v) equation list \ ('f, 'v) equation list" where "subst_list \ ys = map (\p. (fst p \ \, snd p \ \)) ys" lemma mset_subst_list [simp]: "mset (subst_list (subst x t) ys) = subst_mset (subst x t) (mset ys)" by (auto simp: subst_mset_def subst_list_def) lemma subst_list_append: "subst_list \ (xs @ ys) = subst_list \ xs @ subst_list \ ys" by (auto simp: subst_list_def) function (sequential) unify :: "('f, 'v) equation list \ ('v \ ('f, 'v) term) list \ ('v \ ('f, 'v) term) list option" where "unify [] bs = Some bs" | "unify ((Fun f ss, Fun g ts) # E) bs = (case decompose (Fun f ss) (Fun g ts) of None \ None | Some us \ unify (us @ E) bs)" | "unify ((Var x, t) # E) bs = (if t = Var x then unify E bs else if x \ vars_term t then None else unify (subst_list (subst x t) E) ((x, t) # bs))" | "unify ((t, Var x) # E) bs = (if x \ vars_term t then None else unify (subst_list (subst x t) E) ((x, t) # bs))" by pat_completeness auto termination by (standard, rule wf_inv_image [of "unif\" "mset \ fst", OF wf_converse_unif]) (force intro: UNIF1.intros simp: unif_def union_commute)+ lemma unify_append_prefix_same: \<^marker>\contributor \Martin Desharnais\\ "(\e \ set es1. fst e = snd e) \ unify (es1 @ es2) bs = unify es2 bs" proof (induction "es1 @ es2" bs arbitrary: es1 es2 bs rule: unify.induct) case (1 bs) thus ?case by simp next case (2 f ss g ts E bs) show ?case proof (cases es1) case Nil thus ?thesis by simp next case (Cons e es1') hence e_def: "e = (Fun f ss, Fun g ts)" and E_def: "E = es1' @ es2" using "2" by simp_all hence "f = g" and "ss = ts" using "2.prems" local.Cons by auto hence "unify (es1 @ es2) bs = unify ((zip ts ts @ es1') @ es2) bs" by (simp add: Cons e_def) also have "\ = unify es2 bs" proof (rule "2.hyps"(1)) show "decompose (Fun f ss) (Fun g ts) = Some (zip ts ts)" by (simp add: \f = g\ \ss = ts\) next show "zip ts ts @ E = (zip ts ts @ es1') @ es2" by (simp add: E_def) next show "\e\set (zip ts ts @ es1'). fst e = snd e" using "2.prems" by (auto simp: Cons zip_same) qed finally show ?thesis . qed next case (3 x t E bs) show ?case proof (cases es1) case Nil thus ?thesis by simp next case (Cons e es1') hence e_def: "e = (Var x, t)" and E_def: "E = es1' @ es2" using 3 by simp_all show ?thesis proof (cases "t = Var x") case True show ?thesis using 3(1)[OF True E_def] using "3.hyps"(3) "3.prems" local.Cons by fastforce next case False thus ?thesis using "3.prems" e_def local.Cons by force qed qed next case (4 v va x E bs) then show ?case proof (cases es1) case Nil thus ?thesis by simp next case (Cons e es1') hence e_def: "e = (Fun v va, Var x)" and E_def: "E = es1' @ es2" using 4 by simp_all thus ?thesis using "4.prems" local.Cons by fastforce qed qed corollary unify_Cons_same: \<^marker>\contributor \Martin Desharnais\\ "fst e = snd e \ unify (e # es) bs = unify es bs" by (rule unify_append_prefix_same[of "[_]", simplified]) corollary unify_same: \<^marker>\contributor \Martin Desharnais\\ "(\e \ set es. fst e = snd e) \ unify es bs = Some bs" by (rule unify_append_prefix_same[of _ "[]", simplified]) definition subst_of :: "('v \ ('f, 'v) term) list \ ('f, 'v) subst" where "subst_of ss = List.foldr (\(x, t) \. \ \\<^sub>s subst x t) ss Var" text \Computing the mgu of two terms.\ -fun mgu :: "('f, 'v) term \ ('f, 'v) term \ ('f, 'v) subst option" where +definition mgu :: "('f, 'v) term \ ('f, 'v) term \ ('f, 'v) subst option" where "mgu s t = (case unify [(s, t)] [] of None \ None | Some res \ Some (subst_of res))" lemma subst_of_simps [simp]: "subst_of [] = Var" "subst_of ((x, Var x) # ss) = subst_of ss" "subst_of (b # ss) = subst_of ss \\<^sub>s subst (fst b) (snd b)" by (simp_all add: subst_of_def split: prod.splits) lemma subst_of_append [simp]: "subst_of (ss @ ts) = subst_of ts \\<^sub>s subst_of ss" by (induct ss) (auto simp: ac_simps) text \The concrete algorithm \unify\ can be simulated by the inference rules of \UNIF\.\ lemma unify_Some_UNIF: assumes "unify E bs = Some cs" shows "\ds ss. cs = ds @ bs \ subst_of ds = compose ss \ UNIF ss (mset E) {#}" using assms proof (induction E bs arbitrary: cs rule: unify.induct) case (2 f ss g ts E bs) then obtain us where "decompose (Fun f ss) (Fun g ts) = Some us" and [simp]: "f = g" "length ss = length ts" "us = zip ss ts" and "unify (us @ E) bs = Some cs" by (auto split: option.splits) from "2.IH" [OF this(1, 5)] obtain xs ys where "cs = xs @ bs" and [simp]: "subst_of xs = compose ys" and *: "UNIF ys (mset (us @ E)) {#}" by auto then have "UNIF (Var # ys) (mset ((Fun f ss, Fun g ts) # E)) {#}" by (force intro: UNIF1.decomp simp: ac_simps) moreover have "cs = xs @ bs" by fact moreover have "subst_of xs = compose (Var # ys)" by simp ultimately show ?case by blast next case (3 x t E bs) show ?case proof (cases "t = Var x") assume "t = Var x" then show ?case using 3 by auto (metis UNIF.step compose_simps(2) UNIF1.trivial) next assume "t \ Var x" with 3 obtain xs ys where [simp]: "cs = (ys @ [(x, t)]) @ bs" and [simp]: "subst_of ys = compose xs" and "x \ vars_term t" and "UNIF xs (mset (subst_list (subst x t) E)) {#}" by (cases "x \ vars_term t") force+ then have "UNIF (subst x t # xs) (mset ((Var x, t) # E)) {#}" by (force intro: UNIF1.Var_left simp: ac_simps) moreover have "cs = (ys @ [(x, t)]) @ bs" by simp moreover have "subst_of (ys @ [(x, t)]) = compose (subst x t # xs)" by simp ultimately show ?case by blast qed next case (4 f ss x E bs) with 4 obtain xs ys where [simp]: "cs = (ys @ [(x, Fun f ss)]) @ bs" and [simp]: "subst_of ys = compose xs" and "x \ vars_term (Fun f ss)" and "UNIF xs (mset (subst_list (subst x (Fun f ss)) E)) {#}" by (cases "x \ vars_term (Fun f ss)") force+ then have "UNIF (subst x (Fun f ss) # xs) (mset ((Fun f ss, Var x) # E)) {#}" by (force intro: UNIF1.Var_right simp: ac_simps) moreover have "cs = (ys @ [(x, Fun f ss)]) @ bs" by simp moreover have "subst_of (ys @ [(x, Fun f ss)]) = compose (subst x (Fun f ss) # xs)" by simp ultimately show ?case by blast qed force lemma unify_sound: assumes "unify E [] = Some cs" shows "is_imgu (subst_of cs) (set E)" proof - from unify_Some_UNIF [OF assms] obtain ss where "subst_of cs = compose ss" and "UNIF ss (mset E) {#}" by auto with UNIF_empty_imp_is_mgu_compose [OF this(2)] and UNIF_idemp [OF this(2)] show ?thesis by (auto simp add: is_imgu_def is_mgu_def) (metis subst_compose_assoc) qed text \If \unify\ gives up, then the given set of equations cannot be reduced to the empty set by \UNIF\.\ lemma unify_None: assumes "unify E ss = None" shows "\E'. E' \ {#} \ (mset E, E') \ unif\<^sup>!" using assms proof (induction E ss rule: unify.induct) case (1 bs) then show ?case by simp next case (2 f ss g ts E bs) moreover { assume *: "decompose (Fun f ss) (Fun g ts) = None" have ?case proof (cases "unifiable (set E)") case True then have "(mset E, {#}) \ unif\<^sup>*" by (simp add: unifiable_imp_empty) from unif_rtrancl_mono [OF this, of "{#(Fun f ss, Fun g ts)#}"] obtain \ where "(mset E + {#(Fun f ss, Fun g ts)#}, {#(Fun f ss \ \, Fun g ts \ \)#}) \ unif\<^sup>*" by (auto simp: subst_mset_def) moreover have "{#(Fun f ss \ \, Fun g ts \ \)#} \ NF unif" using decompose_None [OF *] by (auto simp: single_is_union NF_def unif_def elim!: UNIF1.cases) (metis length_map) ultimately show ?thesis by auto (metis normalizability_I add_mset_not_empty) next case False moreover have "\ unifiable {(Fun f ss, Fun g ts)}" using * by (auto simp: unifiable_def) ultimately have "\ unifiable (set ((Fun f ss, Fun g ts) # E))" by (auto simp: unifiable_def unifiers_def) then show ?thesis by (simp add: not_unifiable_imp_not_empty_NF) qed } moreover { fix us assume *: "decompose (Fun f ss) (Fun g ts) = Some us" and "unify (us @ E) bs = None" from "2.IH" [OF this] obtain E' where "E' \ {#}" and "(mset (us @ E), E') \ unif\<^sup>!" by blast moreover have "(mset ((Fun f ss, Fun g ts) # E), mset (us @ E)) \ unif" proof - have "g = f" and "length ss = length ts" and "us = zip ss ts" using * by auto then show ?thesis by (auto intro: UNIF1.decomp simp: unif_def ac_simps) qed ultimately have ?case by auto } ultimately show ?case by (auto split: option.splits) next case (3 x t E bs) { assume [simp]: "t = Var x" obtain E' where "E' \ {#}" and "(mset E, E') \ unif\<^sup>!" using 3 by auto moreover have "(mset ((Var x, t) # E), mset E) \ unif" by (auto intro: UNIF1.trivial simp: unif_def) ultimately have ?case by auto } moreover { assume *: "t \ Var x" "x \ vars_term t" then obtain E' where "E' \ {#}" and "(mset (subst_list (subst x t) E), E') \ unif\<^sup>!" using 3 by auto moreover have "(mset ((Var x, t) # E), mset (subst_list (subst x t) E)) \ unif" using * by (auto intro: UNIF1.Var_left simp: unif_def) ultimately have ?case by auto } moreover { assume *: "t \ Var x" "x \ vars_term t" then have "x \ vars_term t" "is_Fun t" by auto then have "\ unifiable {(Var x, t)}" by (rule in_vars_is_Fun_not_unifiable) then have **: "\ unifiable {(Var x \ \, t \ \)}" for \ :: "('b, 'a) subst" using subst_set_reflects_unifiable [of \ "{(Var x, t)}"] by (auto simp: subst_set_def) have ?case proof (cases "unifiable (set E)") case True then have "(mset E, {#}) \ unif\<^sup>*" by (simp add: unifiable_imp_empty) from unif_rtrancl_mono [OF this, of "{#(Var x, t)#}"] obtain \ where "(mset E + {#(Var x, t)#}, {#(Var x \ \, t \ \)#}) \ unif\<^sup>*" by (auto simp: subst_mset_def) moreover obtain E' where "E' \ {#}" and "({#(Var x \ \, t \ \)#}, E') \ unif\<^sup>!" using not_unifiable_imp_not_empty_NF and ** by (metis set_mset_single) ultimately show ?thesis by auto next case False moreover have "\ unifiable {(Var x, t)}" using * by (force simp: unifiable_def) ultimately have "\ unifiable (set ((Var x, t) # E))" by (auto simp: unifiable_def unifiers_def) then show ?thesis by (simp add: not_unifiable_imp_not_empty_NF) qed } ultimately show ?case by blast next case (4 f ss x E bs) define t where "t = Fun f ss" { assume *: "x \ vars_term t" then obtain E' where "E' \ {#}" and "(mset (subst_list (subst x t) E), E') \ unif\<^sup>!" using 4 by (auto simp: t_def) moreover have "(mset ((t, Var x) # E), mset (subst_list (subst x t) E)) \ unif" using * by (auto intro: UNIF1.Var_right simp: unif_def) ultimately have ?case by (auto simp: t_def) } moreover { assume "x \ vars_term t" then have *: "x \ vars_term t" "t \ Var x" by (auto simp: t_def) then have "x \ vars_term t" "is_Fun t" by auto then have "\ unifiable {(Var x, t)}" by (rule in_vars_is_Fun_not_unifiable) then have **: "\ unifiable {(Var x \ \, t \ \)}" for \ :: "('b, 'a) subst" using subst_set_reflects_unifiable [of \ "{(Var x, t)}"] by (auto simp: subst_set_def) have ?case proof (cases "unifiable (set E)") case True then have "(mset E, {#}) \ unif\<^sup>*" by (simp add: unifiable_imp_empty) from unif_rtrancl_mono [OF this, of "{#(t, Var x)#}"] obtain \ where "(mset E + {#(t, Var x)#}, {#(t \ \, Var x \ \)#}) \ unif\<^sup>*" by (auto simp: subst_mset_def) moreover obtain E' where "E' \ {#}" and "({#(t \ \, Var x \ \)#}, E') \ unif\<^sup>!" using not_unifiable_imp_not_empty_NF and ** by (metis unifiable_insert_swap set_mset_single) ultimately show ?thesis by (auto simp: t_def) next case False moreover have "\ unifiable {(t, Var x)}" using * by (simp add: unifiable_def) ultimately have "\ unifiable (set ((t, Var x) # E))" by (auto simp: unifiable_def unifiers_def) then show ?thesis by (simp add: not_unifiable_imp_not_empty_NF t_def) qed } ultimately show ?case by blast qed lemma unify_complete: assumes "unify E bs = None" shows "unifiers (set E) = {}" proof - from unify_None [OF assms] obtain E' where "E' \ {#}" and "(mset E, E') \ unif\<^sup>!" by blast then have "\ unifiable (set E)" using irreducible_reachable_imp_not_unifiable by force then show ?thesis by (auto simp: unifiable_def) qed lemma mgu_complete: "mgu s t = None \ unifiers {(s, t)} = {}" proof - assume "mgu s t = None" - then have "unify [(s, t)] [] = None" by (cases "unify [(s, t)] []", auto) + then have "unify [(s, t)] [] = None" by (cases "unify [(s, t)] []", auto simp: mgu_def) then have "unifiers (set [(s, t)]) = {}" by (rule unify_complete) then show ?thesis by simp qed lemma finite_subst_domain_subst_of: "finite (subst_domain (subst_of xs))" proof (induct xs) case (Cons x xs) moreover have "finite (subst_domain (subst (fst x) (snd x)))" by (metis finite_subst_domain_subst) ultimately show ?case using subst_domain_compose [of "subst_of xs" "subst (fst x) (snd x)"] by (simp del: subst_subst_domain) (metis finite_subset infinite_Un) qed simp lemma unify_subst_domain: \<^marker>\contributor \Martin Desharnais\\ assumes unif: "unify E [] = Some xs" shows "subst_domain (subst_of xs) \ (\e \ set E. vars_term (fst e) \ vars_term (snd e))" proof - from unify_Some_UNIF[OF unif] obtain xs' where "subst_of xs = compose xs'" and "UNIF xs' (mset E) {#}" by auto thus ?thesis using UNIF_subst_domain_subset by (metis (mono_tags, lifting) multiset.set_map set_mset_mset vars_mset_def) qed lemma mgu_subst_domain: assumes "mgu s t = Some \" shows "subst_domain \ \ vars_term s \ vars_term t" proof - obtain xs where "unify [(s, t)] [] = Some xs" and "\ = subst_of xs" - using assms by (simp split: option.splits) + using assms by (simp add: mgu_def split: option.splits) thus ?thesis using unify_subst_domain by fastforce qed lemma mgu_finite_subst_domain: "mgu s t = Some \ \ finite (subst_domain \)" by (drule mgu_subst_domain) (simp add: finite_subset) lemma unify_range_vars: \<^marker>\contributor \Martin Desharnais\\ assumes unif: "unify E [] = Some xs" shows "range_vars (subst_of xs) \ (\e \ set E. vars_term (fst e) \ vars_term (snd e))" proof - from unify_Some_UNIF[OF unif] obtain xs' where "subst_of xs = compose xs'" and "UNIF xs' (mset E) {#}" by auto thus ?thesis using UNIF_range_vars_subset by (metis (mono_tags, lifting) multiset.set_map set_mset_mset vars_mset_def) qed lemma mgu_range_vars: \<^marker>\contributor \Martin Desharnais\\ assumes "mgu s t = Some \" shows "range_vars \ \ vars_term s \ vars_term t" proof - obtain xs where "unify [(s, t)] [] = Some xs" and "\ = subst_of xs" - using assms by (simp split: option.splits) + using assms by (simp add: mgu_def split: option.splits) thus ?thesis using unify_range_vars by fastforce qed lemma unify_subst_domain_range_vars_disjoint: \<^marker>\contributor \Martin Desharnais\\ assumes unif: "unify E [] = Some xs" shows "subst_domain (subst_of xs) \ range_vars (subst_of xs) = {}" proof - from unify_Some_UNIF[OF unif] obtain xs' where "subst_of xs = compose xs'" and "UNIF xs' (mset E) {#}" by auto thus ?thesis using UNIF_subst_domain_range_vars_Int by metis qed lemma mgu_subst_domain_range_vars_disjoint: \<^marker>\contributor \Martin Desharnais\\ assumes "mgu s t = Some \" shows "subst_domain \ \ range_vars \ = {}" proof - obtain xs where "unify [(s, t)] [] = Some xs" and "\ = subst_of xs" - using assms by (simp split: option.splits) + using assms by (simp add: mgu_def split: option.splits) thus ?thesis using unify_subst_domain_range_vars_disjoint by metis qed lemma mgu_sound: assumes "mgu s t = Some \" shows "is_imgu \ {(s, t)}" proof - obtain ss where "unify [(s, t)] [] = Some ss" and "\ = subst_of ss" - using assms by (auto split: option.splits) + using assms by (auto simp: mgu_def split: option.splits) then have "is_imgu \ (set [(s, t)])" by (metis unify_sound) then show ?thesis by simp qed corollary subst_apply_term_eq_subst_apply_term_if_mgu: \<^marker>\contributor \Martin Desharnais\\ assumes mgu_t_u: "mgu t u = Some \" shows "t \ \ = u \ \" using mgu_sound[OF mgu_t_u] by (simp add: is_imgu_def unifiers_def) lemma mgu_same: "mgu t t = Some Var" \<^marker>\contributor \Martin Desharnais\\ - by (simp add: unify_same) + by (simp add: mgu_def unify_same) lemma ex_mgu_if_subst_apply_term_eq_subst_apply_term: \<^marker>\contributor \Martin Desharnais\\ fixes t u :: "('f, 'v) Term.term" and \ :: "('f, 'v) subst" assumes t_eq_u: "t \ \ = u \ \" shows "\\ :: ('f, 'v) subst. Unification.mgu t u = Some \" proof - from t_eq_u have "unifiers {(t, u)} \ {}" unfolding unifiers_def by auto then obtain xs where unify: "unify [(t, u)] [] = Some xs" using unify_complete by (metis list.set(1) list.set(2) not_Some_eq) show ?thesis proof (rule exI) show "Unification.mgu t u = Some (subst_of xs)" - using unify by simp + using unify by (simp add: mgu_def) qed qed lemma mgu_is_Var_if_not_in_equations: \<^marker>\contributor \Martin Desharnais\\ fixes \ :: "('f, 'v) subst" and E :: "('f, 'v) equations" and x :: 'v assumes mgu_\: "is_mgu \ E" and x_not_in: "x \ (\e\E. vars_term (fst e) \ vars_term (snd e))" shows "is_Var (\ x)" proof - from mgu_\ have unif_\: "\ \ unifiers E" and minimal_\: "\\ \ unifiers E. \\. \ = \ \\<^sub>s \" by (simp_all add: is_mgu_def) define \ :: "('f, 'v) subst" where "\ = (\x. if x \ (\e \ E. vars_term (fst e) \ vars_term (snd e)) then \ x else Var x)" have \\ \ unifiers E\ unfolding unifiers_def mem_Collect_eq proof (rule ballI) fix e assume "e \ E" with unif_\ have "fst e \ \ = snd e \ \" by blast moreover from \e \ E\ have "fst e \ \ = fst e \ \" and "snd e \ \ = snd e \ \" unfolding term_subst_eq_conv by (auto simp: \_def) ultimately show "fst e \ \ = snd e \ \" by simp qed with minimal_\ obtain \ where "\ \\<^sub>s \ = \" by auto with x_not_in have "(\ \\<^sub>s \) x = Var x" by (simp add: \_def) thus "is_Var (\ x)" by (metis subst_apply_eq_Var subst_compose term.disc(1)) qed corollary mgu_ball_is_Var: \<^marker>\contributor \Martin Desharnais\\ "is_mgu \ E \ \x \ - (\e\E. vars_term (fst e) \ vars_term (snd e)). is_Var (\ x)" by (rule ballI) (rule mgu_is_Var_if_not_in_equations[folded Compl_iff]) lemma mgu_inj_on: \<^marker>\contributor \Martin Desharnais\\ fixes \ :: "('f, 'v) subst" and E :: "('f, 'v) equations" assumes mgu_\: "is_mgu \ E" shows "inj_on \ (- (\e \ E. vars_term (fst e) \ vars_term (snd e)))" proof (rule inj_onI) fix x y assume x_in: "x \ - (\e\E. vars_term (fst e) \ vars_term (snd e))" and y_in: "y \ - (\e\E. vars_term (fst e) \ vars_term (snd e))" and "\ x = \ y" from mgu_\ have unif_\: "\ \ unifiers E" and minimal_\: "\\ \ unifiers E. \\. \ = \ \\<^sub>s \" by (simp_all add: is_mgu_def) define \ :: "('f, 'v) subst" where "\ = (\x. if x \ (\e \ E. vars_term (fst e) \ vars_term (snd e)) then \ x else Var x)" have \\ \ unifiers E\ unfolding unifiers_def mem_Collect_eq proof (rule ballI) fix e assume "e \ E" with unif_\ have "fst e \ \ = snd e \ \" by blast moreover from \e \ E\ have "fst e \ \ = fst e \ \" and "snd e \ \ = snd e \ \" unfolding term_subst_eq_conv by (auto simp: \_def) ultimately show "fst e \ \ = snd e \ \" by simp qed with minimal_\ obtain \ where "\ \\<^sub>s \ = \" by auto hence "(\ \\<^sub>s \) x = Var x" and "(\ \\<^sub>s \) y = Var y" using ComplD[OF x_in] ComplD[OF y_in] by (simp_all add: \_def) with \\ x = \ y\ show "x = y" by (simp add: subst_compose_def) qed end diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Intruder_Deduction.thy b/thys/Stateful_Protocol_Composition_and_Typing/Intruder_Deduction.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/Intruder_Deduction.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/Intruder_Deduction.thy @@ -1,1193 +1,1193 @@ (* Title: Intruder_Deduction.thy Author: Andreas Viktor Hess, DTU SPDX-License-Identifier: BSD-3-Clause *) section \Dolev-Yao Intruder Model\ theory Intruder_Deduction imports Messages More_Unification begin subsection \Syntax for the Intruder Deduction Relations\ consts INTRUDER_SYNTH::"('f,'v) terms \ ('f,'v) term \ bool" (infix "\\<^sub>c" 50) consts INTRUDER_DEDUCT::"('f,'v) terms \ ('f,'v) term \ bool" (infix "\" 50) subsection \Intruder Model Locale\ text \ The intruder model is parameterized over arbitrary function symbols (e.g, cryptographic operators) and variables. It requires three functions: - \arity\ that assigns an arity to each function symbol. - \public\ that partitions the function symbols into those that will be available to the intruder and those that will not. - \Ana\, the analysis interface, that defines how messages can be decomposed (e.g., decryption). \ locale intruder_model = fixes arity :: "'fun \ nat" and public :: "'fun \ bool" and Ana :: "('fun,'var) term \ (('fun,'var) term list \ ('fun,'var) term list)" assumes Ana_keys_fv: "\t K R. Ana t = (K,R) \ fv\<^sub>s\<^sub>e\<^sub>t (set K) \ fv t" and Ana_keys_wf: "\t k K R f T. Ana t = (K,R) \ (\g S. Fun g S \ t \ length S = arity g) \ k \ set K \ Fun f T \ k \ length T = arity f" and Ana_var[simp]: "\x. Ana (Var x) = ([],[])" and Ana_fun_subterm: "\f T K R. Ana (Fun f T) = (K,R) \ set R \ set T" and Ana_subst: "\t \ K R. \Ana t = (K,R); K \ [] \ R \ []\ \ Ana (t \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \,R \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" begin lemma Ana_subterm: assumes "Ana t = (K,T)" shows "set T \ subterms t" using assms by (cases t) (simp add: psubsetI, metis Ana_fun_subterm Fun_gt_params UN_I term.order_refl params_subterms psubsetI subset_antisym subset_trans) lemma Ana_subterm': "s \ set (snd (Ana t)) \ s \ t" using Ana_subterm by (cases "Ana t") auto lemma Ana_vars: assumes "Ana t = (K,M)" shows "fv\<^sub>s\<^sub>e\<^sub>t (set K) \ fv t" "fv\<^sub>s\<^sub>e\<^sub>t (set M) \ fv t" by (rule Ana_keys_fv[OF assms]) (use Ana_subterm[OF assms] subtermeq_vars_subset in auto) abbreviation \ where "\ \ UNIV::'var set" abbreviation \n ("\\<^sup>_") where "\\<^sup>n \ {f::'fun. arity f = n}" abbreviation \npub ("\\<^sub>p\<^sub>u\<^sub>b\<^sup>_") where "\\<^sub>p\<^sub>u\<^sub>b\<^sup>n \ {f. public f} \ \\<^sup>n" abbreviation \npriv ("\\<^sub>p\<^sub>r\<^sub>i\<^sub>v\<^sup>_") where "\\<^sub>p\<^sub>r\<^sub>i\<^sub>v\<^sup>n \ {f. \public f} \ \\<^sup>n" abbreviation \\<^sub>p\<^sub>u\<^sub>b where "\\<^sub>p\<^sub>u\<^sub>b \ (\n. \\<^sub>p\<^sub>u\<^sub>b\<^sup>n)" abbreviation \\<^sub>p\<^sub>r\<^sub>i\<^sub>v where "\\<^sub>p\<^sub>r\<^sub>i\<^sub>v \ (\n. \\<^sub>p\<^sub>r\<^sub>i\<^sub>v\<^sup>n)" abbreviation \ where "\ \ (\n. \\<^sup>n)" abbreviation \ where "\ \ \\<^sup>0" abbreviation \\<^sub>p\<^sub>u\<^sub>b where "\\<^sub>p\<^sub>u\<^sub>b \ {f. public f} \ \" abbreviation \\<^sub>p\<^sub>r\<^sub>i\<^sub>v where "\\<^sub>p\<^sub>r\<^sub>i\<^sub>v \ {f. \public f} \ \" abbreviation \\<^sub>f where "\\<^sub>f \ \ - \" abbreviation \\<^sub>f\<^sub>p\<^sub>u\<^sub>b where "\\<^sub>f\<^sub>p\<^sub>u\<^sub>b \ \\<^sub>f \ \\<^sub>p\<^sub>u\<^sub>b" abbreviation \\<^sub>f\<^sub>p\<^sub>r\<^sub>i\<^sub>v where "\\<^sub>f\<^sub>p\<^sub>r\<^sub>i\<^sub>v \ \\<^sub>f \ \\<^sub>p\<^sub>r\<^sub>i\<^sub>v" lemma disjoint_fun_syms: "\\<^sub>f \ \ = {}" by auto lemma id_union_univ: "\\<^sub>f \ \ = UNIV" "\ = UNIV" by auto lemma const_arity_eq_zero[dest]: "c \ \ \ arity c = 0" by simp lemma const_pub_arity_eq_zero[dest]: "c \ \\<^sub>p\<^sub>u\<^sub>b \ arity c = 0 \ public c" by simp lemma const_priv_arity_eq_zero[dest]: "c \ \\<^sub>p\<^sub>r\<^sub>i\<^sub>v \ arity c = 0 \ \public c" by simp lemma fun_arity_gt_zero[dest]: "f \ \\<^sub>f \ arity f > 0" by fastforce lemma pub_fun_public[dest]: "f \ \\<^sub>f\<^sub>p\<^sub>u\<^sub>b \ public f" by fastforce lemma pub_fun_arity_gt_zero[dest]: "f \ \\<^sub>f\<^sub>p\<^sub>u\<^sub>b \ arity f > 0" by fastforce lemma \\<^sub>f_unfold: "\\<^sub>f = {f::'fun. arity f > 0}" by auto lemma \_unfold: "\ = {f::'fun. arity f = 0}" by auto lemma \pub_unfold: "\\<^sub>p\<^sub>u\<^sub>b = {f::'fun. arity f = 0 \ public f}" by auto lemma \priv_unfold: "\\<^sub>p\<^sub>r\<^sub>i\<^sub>v = {f::'fun. arity f = 0 \ \public f}" by auto lemma \npub_unfold: "(\\<^sub>p\<^sub>u\<^sub>b\<^sup>n) = {f::'fun. arity f = n \ public f}" by auto lemma \npriv_unfold: "(\\<^sub>p\<^sub>r\<^sub>i\<^sub>v\<^sup>n) = {f::'fun. arity f = n \ \public f}" by auto lemma \fpub_unfold: "\\<^sub>f\<^sub>p\<^sub>u\<^sub>b = {f::'fun. arity f > 0 \ public f}" by auto lemma \fpriv_unfold: "\\<^sub>f\<^sub>p\<^sub>r\<^sub>i\<^sub>v = {f::'fun. arity f > 0 \ \public f}" by auto lemma \n_m_eq: "\(\\<^sup>n) \ {}; (\\<^sup>n) = (\\<^sup>m)\ \ n = m" by auto subsection \Term Well-formedness\ definition "wf\<^sub>t\<^sub>r\<^sub>m t \ \f T. Fun f T \ t \ length T = arity f" abbreviation "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s T \ \t \ T. wf\<^sub>t\<^sub>r\<^sub>m t" lemma Ana_keys_wf': "Ana t = (K,T) \ wf\<^sub>t\<^sub>r\<^sub>m t \ k \ set K \ wf\<^sub>t\<^sub>r\<^sub>m k" using Ana_keys_wf unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by metis lemma wf_trm_Var[simp]: "wf\<^sub>t\<^sub>r\<^sub>m (Var x)" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by simp lemma wf_trm_subst_range_Var[simp]: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range Var)" by simp lemma wf_trm_subst_range_iff: "(\x. wf\<^sub>t\<^sub>r\<^sub>m (\ x)) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by force lemma wf_trm_subst_rangeD: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ wf\<^sub>t\<^sub>r\<^sub>m (\ x)" by (metis wf_trm_subst_range_iff) lemma wf_trm_subst_rangeI[intro]: "(\x. wf\<^sub>t\<^sub>r\<^sub>m (\ x)) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" by (metis wf_trm_subst_range_iff) lemma wf_trmI[intro]: assumes "\t. t \ set T \ wf\<^sub>t\<^sub>r\<^sub>m t" "length T = arity f" shows "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" using assms unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto lemma wf_trm_subterm: "\wf\<^sub>t\<^sub>r\<^sub>m t; s \ t\ \ wf\<^sub>t\<^sub>r\<^sub>m s" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by (induct t) auto lemma wf_trm_subtermeq: assumes "wf\<^sub>t\<^sub>r\<^sub>m t" "s \ t" shows "wf\<^sub>t\<^sub>r\<^sub>m s" proof (cases "s = t") case False thus "wf\<^sub>t\<^sub>r\<^sub>m s" using assms(2) wf_trm_subterm[OF assms(1)] by simp qed (metis assms(1)) lemma wf_trm_param: assumes "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" "t \ set T" shows "wf\<^sub>t\<^sub>r\<^sub>m t" by (meson assms subtermeqI'' wf_trm_subtermeq) lemma wf_trm_param_idx: assumes "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" and "i < length T" shows "wf\<^sub>t\<^sub>r\<^sub>m (T ! i)" using wf_trm_param[OF assms(1), of "T ! i"] assms(2) by fastforce lemma wf_trm_subst: assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "wf\<^sub>t\<^sub>r\<^sub>m t = wf\<^sub>t\<^sub>r\<^sub>m (t \ \)" proof show "wf\<^sub>t\<^sub>r\<^sub>m t \ wf\<^sub>t\<^sub>r\<^sub>m (t \ \)" proof (induction t) case (Fun f T) hence "\t. t \ set T \ wf\<^sub>t\<^sub>r\<^sub>m t" by (meson wf\<^sub>t\<^sub>r\<^sub>m_def Fun_param_is_subterm term.order_trans) hence "\t. t \ set T \ wf\<^sub>t\<^sub>r\<^sub>m (t \ \)" using Fun.IH by auto moreover have "length (map (\t. t \ \) T) = arity f" using Fun.prems unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto ultimately show ?case by fastforce qed (simp add: wf_trm_subst_rangeD[OF assms]) show "wf\<^sub>t\<^sub>r\<^sub>m (t \ \) \ wf\<^sub>t\<^sub>r\<^sub>m t" proof (induction t) case (Fun f T) hence "wf\<^sub>t\<^sub>r\<^sub>m t" when "t \ set (map (\s. s \ \) T)" for t by (metis that wf\<^sub>t\<^sub>r\<^sub>m_def Fun_param_is_subterm term.order_trans subst_apply_term.simps(2)) hence "wf\<^sub>t\<^sub>r\<^sub>m t" when "t \ set T" for t using that Fun.IH by auto moreover have "length (map (\t. t \ \) T) = arity f" using Fun.prems unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto ultimately show ?case by fastforce qed (simp add: assms) qed lemma wf_trm_subst_singleton: assumes "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m t'" shows "wf\<^sub>t\<^sub>r\<^sub>m (t \ Var(v := t'))" proof - have "wf\<^sub>t\<^sub>r\<^sub>m ((Var(v := t')) w)" for w using assms(2) unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by simp thus ?thesis using assms(1) wf_trm_subst[of "Var(v := t')" t, OF wf_trm_subst_rangeI] by simp qed lemma wf_trm_subst_rm_vars: assumes "wf\<^sub>t\<^sub>r\<^sub>m (t \ \)" shows "wf\<^sub>t\<^sub>r\<^sub>m (t \ rm_vars X \)" using assms proof (induction t) case (Fun f T) have "wf\<^sub>t\<^sub>r\<^sub>m (t \ \)" when "t \ set T" for t using that wf_trm_param[of f "map (\t. t \ \) T"] Fun.prems by auto hence "wf\<^sub>t\<^sub>r\<^sub>m (t \ rm_vars X \)" when "t \ set T" for t using that Fun.IH by simp moreover have "length T = arity f" using Fun.prems unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto ultimately show ?case unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto qed simp lemma wf_trm_subst_rm_vars': "wf\<^sub>t\<^sub>r\<^sub>m (\ v) \ wf\<^sub>t\<^sub>r\<^sub>m (rm_vars X \ v)" by auto lemma wf_trms_subst: assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (M \\<^sub>s\<^sub>e\<^sub>t \)" by (metis (no_types, lifting) assms imageE wf_trm_subst) lemma wf_trms_subst_rm_vars: assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (M \\<^sub>s\<^sub>e\<^sub>t \)" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (M \\<^sub>s\<^sub>e\<^sub>t rm_vars X \)" using assms wf_trm_subst_rm_vars by blast lemma wf_trms_subst_rm_vars': assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (rm_vars X \))" using assms by force lemma wf_trms_subst_compose: assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using assms subst_img_comp_subset' wf_trm_subst by blast lemma wf_trm_subst_compose: fixes \::"('fun, 'v) subst" assumes "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" "\x. wf\<^sub>t\<^sub>r\<^sub>m (\ x)" shows "wf\<^sub>t\<^sub>r\<^sub>m ((\ \\<^sub>s \) x)" using wf_trm_subst[of \ "\ x", OF wf_trm_subst_rangeI[OF assms(2)]] assms(1) subst_subst_compose[of "Var x" \ \] subst_apply_term.simps(1)[of x \] subst_apply_term.simps(1)[of x "\ \\<^sub>s \"] by argo lemma wf_trms_Var_range: assumes "subst_range \ \ range Var" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using assms by fastforce lemma wf_trms_subst_compose_Var_range: assumes "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and "subst_range \ \ range Var" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using assms wf_trms_subst_compose wf_trms_Var_range by metis+ lemma wf_trm_subst_inv: "wf\<^sub>t\<^sub>r\<^sub>m (t \ \) \ wf\<^sub>t\<^sub>r\<^sub>m t" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by (induct t) auto lemma wf_trms_subst_inv: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (M \\<^sub>s\<^sub>e\<^sub>t \) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" using wf_trm_subst_inv by fast lemma wf_trm_subterms: "wf\<^sub>t\<^sub>r\<^sub>m t \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subterms t)" using wf_trm_subterm by blast lemma wf_trms_subterms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subterms\<^sub>s\<^sub>e\<^sub>t M)" using wf_trm_subterms by blast lemma wf_trm_arity: "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T) \ length T = arity f" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by blast lemma wf_trm_subterm_arity: "wf\<^sub>t\<^sub>r\<^sub>m t \ Fun f T \ t \ length T = arity f" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by blast lemma unify_list_wf_trm: assumes "Unification.unify E B = Some U" "\(s,t) \ set E. wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t" and "\(v,t) \ set B. wf\<^sub>t\<^sub>r\<^sub>m t" shows "\(v,t) \ set U. wf\<^sub>t\<^sub>r\<^sub>m t" using assms proof (induction E B arbitrary: U rule: Unification.unify.induct) case (1 B U) thus ?case by auto next case (2 f T g S E B U) have wf_fun: "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" "wf\<^sub>t\<^sub>r\<^sub>m (Fun g S)" using "2.prems"(2) by auto from "2.prems"(1) obtain E' where *: "decompose (Fun f T) (Fun g S) = Some E'" and [simp]: "f = g" "length T = length S" "E' = zip T S" and **: "Unification.unify (E'@E) B = Some U" by (auto split: option.splits) hence "t \ Fun f T" "t' \ Fun g S" when "(t,t') \ set E'" for t t' using that by (metis zip_arg_subterm(1), metis zip_arg_subterm(2)) hence "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m t'" when "(t,t') \ set E'" for t t' using wf_trm_subterm wf_fun \f = g\ that by blast+ thus ?case using "2.IH"[OF * ** _ "2.prems"(3)] "2.prems"(2) by fastforce next case (3 v t E B) hence *: "\(w,x) \ set ((v, t) # B). wf\<^sub>t\<^sub>r\<^sub>m x" and **: "\(s,t) \ set E. wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m t" by auto show ?case proof (cases "t = Var v") case True thus ?thesis using "3.prems" "3.IH"(1) by auto next case False hence "v \ fv t" using "3.prems"(1) by auto hence "Unification.unify (subst_list (subst v t) E) ((v, t)#B) = Some U" using \t \ Var v\ "3.prems"(1) by auto moreover have "\(s, t) \ set (subst_list (subst v t) E). wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t" using wf_trm_subst_singleton[OF _ \wf\<^sub>t\<^sub>r\<^sub>m t\] "3.prems"(2) unfolding subst_list_def subst_def by auto ultimately show ?thesis using "3.IH"(2)[OF \t \ Var v\ \v \ fv t\ _ _ *] by metis qed next case (4 f T v E B U) hence *: "\(w,x) \ set ((v, Fun f T) # B). wf\<^sub>t\<^sub>r\<^sub>m x" and **: "\(s,t) \ set E. wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" by auto have "v \ fv (Fun f T)" using "4.prems"(1) by force hence "Unification.unify (subst_list (subst v (Fun f T)) E) ((v, Fun f T)#B) = Some U" using "4.prems"(1) by auto moreover have "\(s, t) \ set (subst_list (subst v (Fun f T)) E). wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t" using wf_trm_subst_singleton[OF _ \wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)\] "4.prems"(2) unfolding subst_list_def subst_def by auto ultimately show ?case using "4.IH"[OF \v \ fv (Fun f T)\ _ _ *] by metis qed lemma mgu_wf_trm: assumes "mgu s t = Some \" "wf\<^sub>t\<^sub>r\<^sub>m s" "wf\<^sub>t\<^sub>r\<^sub>m t" shows "wf\<^sub>t\<^sub>r\<^sub>m (\ v)" proof - from assms obtain \' where "subst_of \' = \" "\(v,t) \ set \'. wf\<^sub>t\<^sub>r\<^sub>m t" - using unify_list_wf_trm[of "[(s,t)]" "[]"] by (auto split: option.splits) + using unify_list_wf_trm[of "[(s,t)]" "[]"] by (auto simp: mgu_def split: option.splits) thus ?thesis proof (induction \' arbitrary: \ v rule: List.rev_induct) case (snoc x \' \ v) define \ where "\ = subst_of \'" hence "wf\<^sub>t\<^sub>r\<^sub>m (\ v)" for v using snoc.prems(2) snoc.IH[of \] by fastforce moreover obtain w t where x: "x = (w,t)" by (metis surj_pair) hence \: "\ = Var(w := t) \\<^sub>s \" using snoc.prems(1) by (simp add: subst_def \_def) moreover have "wf\<^sub>t\<^sub>r\<^sub>m t" using snoc.prems(2) x by auto ultimately show ?case using wf_trm_subst[of _ t] unfolding subst_compose_def by auto qed (simp add: wf\<^sub>t\<^sub>r\<^sub>m_def) qed lemma mgu_wf_trms: assumes "mgu s t = Some \" "wf\<^sub>t\<^sub>r\<^sub>m s" "wf\<^sub>t\<^sub>r\<^sub>m t" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_wf_trm[OF assms] by simp subsection \Definitions: Intruder Deduction Relations\ text \ A standard Dolev-Yao intruder. \ inductive intruder_deduct::"('fun,'var) terms \ ('fun,'var) term \ bool" where Axiom[simp]: "t \ M \ intruder_deduct M t" | Compose[simp]: "\length T = arity f; public f; \t. t \ set T \ intruder_deduct M t\ \ intruder_deduct M (Fun f T)" | Decompose: "\intruder_deduct M t; Ana t = (K, T); \k. k \ set K \ intruder_deduct M k; t\<^sub>i \ set T\ \ intruder_deduct M t\<^sub>i" text \ A variant of the intruder relation which limits the intruder to composition only. \ inductive intruder_synth::"('fun,'var) terms \ ('fun,'var) term \ bool" where AxiomC[simp]: "t \ M \ intruder_synth M t" | ComposeC[simp]: "\length T = arity f; public f; \t. t \ set T \ intruder_synth M t\ \ intruder_synth M (Fun f T)" adhoc_overloading INTRUDER_DEDUCT intruder_deduct adhoc_overloading INTRUDER_SYNTH intruder_synth lemma intruder_deduct_induct[consumes 1, case_names Axiom Compose Decompose]: assumes "M \ t" "\t. t \ M \ P M t" "\T f. \length T = arity f; public f; \t. t \ set T \ M \ t; \t. t \ set T \ P M t\ \ P M (Fun f T)" "\t K T t\<^sub>i. \M \ t; P M t; Ana t = (K, T); \k. k \ set K \ M \ k; \k. k \ set K \ P M k; t\<^sub>i \ set T\ \ P M t\<^sub>i" shows "P M t" using assms by (induct rule: intruder_deduct.induct) blast+ lemma intruder_synth_induct[consumes 1, case_names AxiomC ComposeC]: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" assumes "M \\<^sub>c t" "\t. t \ M \ P M t" "\T f. \length T = arity f; public f; \t. t \ set T \ M \\<^sub>c t; \t. t \ set T \ P M t\ \ P M (Fun f T)" shows "P M t" using assms by (induct rule: intruder_synth.induct) auto subsection \Definitions: Analyzed Knowledge and Public Ground Well-formed Terms (PGWTs)\ definition analyzed::"('fun,'var) terms \ bool" where "analyzed M \ \t. M \ t \ M \\<^sub>c t" definition analyzed_in where "analyzed_in t M \ \K R. (Ana t = (K,R) \ (\k \ set K. M \\<^sub>c k)) \ (\r \ set R. M \\<^sub>c r)" definition decomp_closure::"('fun,'var) terms \ ('fun,'var) terms \ bool" where "decomp_closure M M' \ \t. M \ t \ (\t' \ M. t \ t') \ t \ M'" inductive public_ground_wf_term::"('fun,'var) term \ bool" where PGWT[simp]: "\public f; arity f = length T; \t. t \ set T \ public_ground_wf_term t\ \ public_ground_wf_term (Fun f T)" abbreviation "public_ground_wf_terms \ {t. public_ground_wf_term t}" lemma public_const_deduct: assumes "c \ \\<^sub>p\<^sub>u\<^sub>b" shows "M \ Fun c []" "M \\<^sub>c Fun c []" proof - have "arity c = 0" "public c" using const_arity_eq_zero \c \ \\<^sub>p\<^sub>u\<^sub>b\ by auto thus "M \ Fun c []" "M \\<^sub>c Fun c []" using intruder_synth.ComposeC[OF _ \public c\, of "[]"] intruder_deduct.Compose[OF _ \public c\, of "[]"] by auto qed lemma public_const_deduct'[simp]: assumes "arity c = 0" "public c" shows "M \ Fun c []" "M \\<^sub>c Fun c []" using intruder_deduct.Compose[of "[]" c] intruder_synth.ComposeC[of "[]" c] assms by simp_all lemma private_fun_deduct_in_ik: assumes t: "M \ t" "Fun f T \ subterms t" and f: "\public f" shows "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t M" using t proof (induction t rule: intruder_deduct.induct) case Decompose thus ?case by (meson Ana_subterm psubsetD term.order_trans) qed (auto simp add: f in_subterms_Union) lemma private_fun_deduct_in_ik': assumes t: "M \ Fun f T" and f: "\public f" shows "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t M" by (rule private_fun_deduct_in_ik[OF t term.order_refl f]) lemma pgwt_public: "\public_ground_wf_term t; Fun f T \ t\ \ public f" by (induct t rule: public_ground_wf_term.induct) auto lemma pgwt_ground: "public_ground_wf_term t \ fv t = {}" by (induct t rule: public_ground_wf_term.induct) auto lemma pgwt_fun: "public_ground_wf_term t \ \f T. t = Fun f T" using pgwt_ground[of t] by (cases t) auto lemma pgwt_arity: "\public_ground_wf_term t; Fun f T \ t\ \ arity f = length T" by (induct t rule: public_ground_wf_term.induct) auto lemma pgwt_wellformed: "public_ground_wf_term t \ wf\<^sub>t\<^sub>r\<^sub>m t" by (induct t rule: public_ground_wf_term.induct) auto lemma pgwt_deducible: "public_ground_wf_term t \ M \\<^sub>c t" by (induct t rule: public_ground_wf_term.induct) auto lemma pgwt_is_empty_synth: "public_ground_wf_term t \ {} \\<^sub>c t" proof - { fix M::"('fun,'var) term set" assume "M \\<^sub>c t" "M = {}" hence "public_ground_wf_term t" by (induct t rule: intruder_synth.induct) auto } thus ?thesis using pgwt_deducible by auto qed lemma ideduct_synth_subst_apply: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" assumes "{} \\<^sub>c t" "\v. M \\<^sub>c \ v" shows "M \\<^sub>c t \ \" proof - { fix M'::"('fun,'var) term set" assume "M' \\<^sub>c t" "M' = {}" hence "M \\<^sub>c t \ \" proof (induction t rule: intruder_synth.induct) case (ComposeC T f M') hence "length (map (\t. t \ \) T) = arity f" "\x. x \ set (map (\t. t \ \) T) \ M \\<^sub>c x" by auto thus ?case using intruder_synth.ComposeC[of "map (\t. t \ \) T" f M] \public f\ by fastforce qed simp } thus ?thesis using assms by metis qed subsection \Lemmata: Monotonicity, Deduction of Private Constants, etc.\ context begin lemma ideduct_mono: "\M \ t; M \ M'\ \ M' \ t" proof (induction rule: intruder_deduct.induct) case (Decompose M t K T t\<^sub>i) have "\k. k \ set K \ M' \ k" using Decompose.IH \M \ M'\ by simp moreover have "M' \ t" using Decompose.IH \M \ M'\ by simp ultimately show ?case using Decompose.hyps intruder_deduct.Decompose by blast qed auto lemma ideduct_synth_mono: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" shows "\M \\<^sub>c t; M \ M'\ \ M' \\<^sub>c t" by (induct rule: intruder_synth.induct) auto context begin \ \Used by \inductive_set\\ private lemma ideduct_mono_set[mono_set]: "M \ N \ M \ t \ N \ t" "M \ N \ M \\<^sub>c t \ N \\<^sub>c t" using ideduct_mono ideduct_synth_mono by (blast, blast) end lemma ideduct_reduce: "\M \ M' \ t; \t'. t' \ M' \ M \ t'\ \ M \ t" proof (induction rule: intruder_deduct_induct) case Decompose thus ?case using intruder_deduct.Decompose by blast qed auto lemma ideduct_synth_reduce: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" shows "\M \ M' \\<^sub>c t; \t'. t' \ M' \ M \\<^sub>c t'\ \ M \\<^sub>c t" by (induct rule: intruder_synth_induct) auto lemma ideduct_mono_eq: assumes "\t. M \ t \ M' \ t" shows "M \ N \ t \ M' \ N \ t" proof show "M \ N \ t \ M' \ N \ t" proof (induction t rule: intruder_deduct_induct) case (Axiom t) thus ?case proof (cases "t \ M") case True hence "M \ t" using intruder_deduct.Axiom by metis thus ?thesis using assms ideduct_mono[of M' t "M' \ N"] by simp qed auto next case (Compose T f) thus ?case using intruder_deduct.Compose by auto next case (Decompose t K T t\<^sub>i) thus ?case using intruder_deduct.Decompose[of "M' \ N" t K T] by auto qed show "M' \ N \ t \ M \ N \ t" proof (induction t rule: intruder_deduct_induct) case (Axiom t) thus ?case proof (cases "t \ M'") case True hence "M' \ t" using intruder_deduct.Axiom by metis thus ?thesis using assms ideduct_mono[of M t "M \ N"] by simp qed auto next case (Compose T f) thus ?case using intruder_deduct.Compose by auto next case (Decompose t K T t\<^sub>i) thus ?case using intruder_deduct.Decompose[of "M \ N" t K T] by auto qed qed lemma deduct_synth_subterm: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" assumes "M \\<^sub>c t" "s \ subterms t" "\m \ M. \s \ subterms m. M \\<^sub>c s" shows "M \\<^sub>c s" using assms by (induct t rule: intruder_synth.induct) auto lemma deduct_if_synth[intro, dest]: "M \\<^sub>c t \ M \ t" by (induct rule: intruder_synth.induct) auto private lemma ideduct_ik_eq: assumes "\t \ M. M' \ t" shows "M' \ t \ M' \ M \ t" by (meson assms ideduct_mono ideduct_reduce sup_ge1) private lemma synth_if_deduct_empty: "{} \ t \ {} \\<^sub>c t" proof (induction t rule: intruder_deduct_induct) case (Decompose t K M m) then obtain f T where "t = Fun f T" "m \ set T" using Ana_fun_subterm Ana_var by (cases t) fastforce+ with Decompose.IH(1) show ?case by (induction rule: intruder_synth_induct) auto qed auto private lemma ideduct_deduct_synth_mono_eq: assumes "\t. M \ t \ M' \\<^sub>c t" "M \ M'" and "\t. M' \ N \ t \ M' \ N \ D \\<^sub>c t" shows "M \ N \ t \ M' \ N \ D \\<^sub>c t" proof - have "\m \ M'. M \ m" using assms(1) by auto hence "\t. M \ t \ M' \ t" by (metis assms(1,2) deduct_if_synth ideduct_reduce sup.absorb2) hence "\t. M' \ N \ t \ M \ N \ t" by (meson ideduct_mono_eq) thus ?thesis by (meson assms(3)) qed lemma ideduct_subst: "M \ t \ M \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" proof (induction t rule: intruder_deduct_induct) case (Compose T f) hence "length (map (\t. t \ \) T) = arity f" "\t. t \ set T \ M \\<^sub>s\<^sub>e\<^sub>t \ \ t \ \" by auto thus ?case using intruder_deduct.Compose[OF _ Compose.hyps(2), of "map (\t. t \ \) T"] by auto next case (Decompose t K M' m') hence "Ana (t \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \, M' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" "\k. k \ set (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) \ M \\<^sub>s\<^sub>e\<^sub>t \ \ k" "m' \ \ \ set (M' \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Ana_subst[OF Decompose.hyps(2)] by fastforce+ thus ?case using intruder_deduct.Decompose[OF Decompose.IH(1)] by metis qed simp lemma ideduct_synth_subst: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" and \::"('fun,'var) subst" shows "M \\<^sub>c t \ M \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" proof (induction t rule: intruder_synth_induct) case (ComposeC T f) hence "length (map (\t. t \ \) T) = arity f" "\t. t \ set T \ M \\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>c t \ \" by auto thus ?case using intruder_synth.ComposeC[OF _ ComposeC.hyps(2), of "map (\t. t \ \) T"] by auto qed simp lemma ideduct_vars: assumes "M \ t" shows "fv t \ fv\<^sub>s\<^sub>e\<^sub>t M" using assms proof (induction t rule: intruder_deduct_induct) case (Decompose t K T t\<^sub>i) thus ?case using Ana_vars(2) fv_subset by blast qed auto lemma ideduct_synth_vars: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" assumes "M \\<^sub>c t" shows "fv t \ fv\<^sub>s\<^sub>e\<^sub>t M" using assms by (induct t rule: intruder_synth_induct) auto lemma ideduct_synth_priv_fun_in_ik: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" assumes "M \\<^sub>c t" "f \ funs_term t" "\public f" shows "f \ \(funs_term ` M)" using assms by (induct t rule: intruder_synth_induct) auto lemma ideduct_synth_priv_const_in_ik: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" assumes "M \\<^sub>c Fun c []" "\public c" shows "Fun c [] \ M" using intruder_synth.cases[OF assms(1)] assms(2) by fast lemma ideduct_synth_ik_replace: fixes M::"('fun,'var) terms" and t::"('fun,'var) term" assumes "\t \ M. N \\<^sub>c t" and "M \\<^sub>c t" shows "N \\<^sub>c t" using assms(2,1) by (induct t rule: intruder_synth.induct) auto end subsection \Lemmata: Analyzed Intruder Knowledge Closure\ lemma deducts_eq_if_analyzed: "analyzed M \ M \ t \ M \\<^sub>c t" unfolding analyzed_def by auto lemma closure_is_superset: "decomp_closure M M' \ M \ M'" unfolding decomp_closure_def by force lemma deduct_if_closure_deduct: "\M' \ t; decomp_closure M M'\ \ M \ t" proof (induction t rule: intruder_deduct.induct) case (Decompose M' t K T t\<^sub>i) thus ?case using intruder_deduct.Decompose[OF _ \Ana t = (K,T)\ _ \t\<^sub>i \ set T\] by simp qed (auto simp add: decomp_closure_def) lemma deduct_if_closure_synth: "\decomp_closure M M'; M' \\<^sub>c t\ \ M \ t" using deduct_if_closure_deduct by blast lemma decomp_closure_subterms_composable: assumes "decomp_closure M M'" and "M' \\<^sub>c t'" "M' \ t" "t \ t'" shows "M' \\<^sub>c t" using \M' \\<^sub>c t'\ assms proof (induction t' rule: intruder_synth.induct) case (AxiomC t' M') have "M \ t" using \M' \ t\ deduct_if_closure_deduct AxiomC.prems(1) by blast moreover { have "\s \ M. t' \ s" using \t' \ M'\ AxiomC.prems(1) unfolding decomp_closure_def by blast hence "\s \ M. t \ s" using \t \ t'\ term.order_trans by auto } ultimately have "t \ M'" using AxiomC.prems(1) unfolding decomp_closure_def by blast thus ?case by simp next case (ComposeC T f M') let ?t' = "Fun f T" { assume "t = ?t'" have "M' \\<^sub>c t" using \M' \\<^sub>c ?t'\ \t = ?t'\ by simp } moreover { assume "t \ ?t'" have "\x \ set T. t \ x" using \t \ ?t'\ \t \ ?t'\ by simp hence "M' \\<^sub>c t" using ComposeC.IH ComposeC.prems(1,3) ComposeC.hyps(3) by blast } ultimately show ?case using cases_simp[of "t = ?t'" "M' \\<^sub>c t"] by simp qed lemma decomp_closure_analyzed: assumes "decomp_closure M M'" shows "analyzed M'" proof - { fix t assume "M' \ t" have "M' \\<^sub>c t" using \M' \ t\ assms proof (induction t rule: intruder_deduct.induct) case (Decompose M' t K T t\<^sub>i) hence "M' \ t\<^sub>i" using Decompose.hyps intruder_deduct.Decompose by blast moreover have "t\<^sub>i \ t" using Decompose.hyps(4) Ana_subterm[OF Decompose.hyps(2)] by blast moreover have "M' \\<^sub>c t" using Decompose.IH(1) Decompose.prems by blast ultimately show "M' \\<^sub>c t\<^sub>i" using decomp_closure_subterms_composable Decompose.prems by blast qed auto } moreover have "\t. M \\<^sub>c t \ M \ t" by auto ultimately show ?thesis by (auto simp add: decomp_closure_def analyzed_def) qed lemma analyzed_if_all_analyzed_in: assumes M: "\t \ M. analyzed_in t M" shows "analyzed M" proof (unfold analyzed_def, intro allI iffI) fix t assume t: "M \ t" thus "M \\<^sub>c t" proof (induction t rule: intruder_deduct_induct) case (Decompose t K T t\<^sub>i) { assume "t \ M" hence ?case using M Decompose.IH(2) Decompose.hyps(2,4) unfolding analyzed_in_def by fastforce } moreover { fix f S assume "t = Fun f S" "\s. s \ set S \ M \\<^sub>c s" hence ?case using Ana_fun_subterm[of f S] Decompose.hyps(2,4) by blast } ultimately show ?case using intruder_synth.cases[OF Decompose.IH(1), of ?case] by blast qed simp_all qed auto lemma analyzed_is_all_analyzed_in: "(\t \ M. analyzed_in t M) \ analyzed M" proof show "analyzed M \ \t \ M. analyzed_in t M" unfolding analyzed_in_def analyzed_def by (auto intro: intruder_deduct.Decompose[OF intruder_deduct.Axiom]) qed (rule analyzed_if_all_analyzed_in) lemma ik_has_synth_ik_closure: fixes M :: "('fun,'var) terms" shows "\M'. (\t. M \ t \ M' \\<^sub>c t) \ decomp_closure M M' \ (finite M \ finite M')" proof - let ?M' = "{t. M \ t \ (\t' \ M. t \ t')}" have M'_closes: "decomp_closure M ?M'" unfolding decomp_closure_def by simp hence "M \ ?M'" using closure_is_superset by simp have "\t. ?M' \\<^sub>c t \ M \ t" using deduct_if_closure_synth[OF M'_closes] by blast moreover have "\t. M \ t \ ?M' \ t" using ideduct_mono[OF _ \M \ ?M'\] by simp moreover have "analyzed ?M'" using decomp_closure_analyzed[OF M'_closes] . ultimately have "\t. M \ t \ ?M' \\<^sub>c t" unfolding analyzed_def by blast moreover have "finite M \ finite ?M'" by auto ultimately show ?thesis using M'_closes by blast qed lemma deducts_eq_if_empty_ik: "{} \ t \ {} \\<^sub>c t" using analyzed_is_all_analyzed_in[of "{}"] deducts_eq_if_analyzed[of "{}" t] by blast subsection \Intruder Variants: Numbered and Composition-Restricted Intruder Deduction Relations\ text \ A variant of the intruder relation which restricts composition to only those terms that satisfy a given predicate Q. \ inductive intruder_deduct_restricted:: "('fun,'var) terms \ (('fun,'var) term \ bool) \ ('fun,'var) term \ bool" ("\_;_\ \\<^sub>r _" 50) where AxiomR[simp]: "t \ M \ \M; Q\ \\<^sub>r t" | ComposeR[simp]: "\length T = arity f; public f; \t. t \ set T \ \M; Q\ \\<^sub>r t; Q (Fun f T)\ \ \M; Q\ \\<^sub>r Fun f T" | DecomposeR: "\\M; Q\ \\<^sub>r t; Ana t = (K, T); \k. k \ set K \ \M; Q\ \\<^sub>r k; t\<^sub>i \ set T\ \ \M; Q\ \\<^sub>r t\<^sub>i" text \ A variant of the intruder relation equipped with a number representing the height of the derivation tree (i.e., \\M; k\ \\<^sub>n t\ iff k is the maximum number of applications of the compose and decompose rules in any path of the derivation tree for \M \ t\). \ inductive intruder_deduct_num:: "('fun,'var) terms \ nat \ ('fun,'var) term \ bool" ("\_; _\ \\<^sub>n _" 50) where AxiomN[simp]: "t \ M \ \M; 0\ \\<^sub>n t" | ComposeN[simp]: "\length T = arity f; public f; \t. t \ set T \ \M; steps t\ \\<^sub>n t\ \ \M; Suc (Max (insert 0 (steps ` set T)))\ \\<^sub>n Fun f T" | DecomposeN: "\\M; n\ \\<^sub>n t; Ana t = (K, T); \k. k \ set K \ \M; steps k\ \\<^sub>n k; t\<^sub>i \ set T\ \ \M; Suc (Max (insert n (steps ` set K)))\ \\<^sub>n t\<^sub>i" lemma intruder_deduct_restricted_induct[consumes 1, case_names AxiomR ComposeR DecomposeR]: assumes "\M; Q\ \\<^sub>r t" "\t. t \ M \ P M Q t" "\T f. \length T = arity f; public f; \t. t \ set T \ \M; Q\ \\<^sub>r t; \t. t \ set T \ P M Q t; Q (Fun f T) \ \ P M Q (Fun f T)" "\t K T t\<^sub>i. \\M; Q\ \\<^sub>r t; P M Q t; Ana t = (K, T); \k. k \ set K \ \M; Q\ \\<^sub>r k; \k. k \ set K \ P M Q k; t\<^sub>i \ set T\ \ P M Q t\<^sub>i" shows "P M Q t" using assms by (induct t rule: intruder_deduct_restricted.induct) blast+ lemma intruder_deduct_num_induct[consumes 1, case_names AxiomN ComposeN DecomposeN]: assumes "\M; n\ \\<^sub>n t" "\t. t \ M \ P M 0 t" "\T f steps. \length T = arity f; public f; \t. t \ set T \ \M; steps t\ \\<^sub>n t; \t. t \ set T \ P M (steps t) t\ \ P M (Suc (Max (insert 0 (steps ` set T)))) (Fun f T)" "\t K T t\<^sub>i steps n. \\M; n\ \\<^sub>n t; P M n t; Ana t = (K, T); \k. k \ set K \ \M; steps k\ \\<^sub>n k; t\<^sub>i \ set T; \k. k \ set K \ P M (steps k) k\ \ P M (Suc (Max (insert n (steps ` set K)))) t\<^sub>i" shows "P M n t" using assms by (induct rule: intruder_deduct_num.induct) blast+ lemma ideduct_restricted_mono: "\\M; P\ \\<^sub>r t; M \ M'\ \ \M'; P\ \\<^sub>r t" proof (induction rule: intruder_deduct_restricted_induct) case (DecomposeR t K T t\<^sub>i) have "\k. k \ set K \ \M'; P\ \\<^sub>r k" using DecomposeR.IH \M \ M'\ by simp moreover have "\M'; P\ \\<^sub>r t" using DecomposeR.IH \M \ M'\ by simp ultimately show ?case using DecomposeR intruder_deduct_restricted.DecomposeR[OF _ DecomposeR.hyps(2) _ DecomposeR.hyps(4)] by blast qed auto subsection \Lemmata: Intruder Deduction Equivalences\ lemma deduct_if_restricted_deduct: "\M;P\ \\<^sub>r m \ M \ m" proof (induction m rule: intruder_deduct_restricted_induct) case (DecomposeR t K T t\<^sub>i) thus ?case using intruder_deduct.Decompose by blast qed simp_all lemma restricted_deduct_if_restricted_ik: assumes "\M;P\ \\<^sub>r m" "\m \ M. P m" and P: "\t t'. P t \ t' \ t \ P t'" shows "P m" using assms(1) proof (induction m rule: intruder_deduct_restricted_induct) case (DecomposeR t K T t\<^sub>i) obtain f S where "t = Fun f S" using Ana_var \t\<^sub>i \ set T\ \Ana t = (K, T)\ by (cases t) auto thus ?case using DecomposeR assms(2) P Ana_subterm by blast qed (simp_all add: assms(2)) lemma deduct_restricted_if_synth: assumes P: "P m" "\t t'. P t \ t' \ t \ P t'" and m: "M \\<^sub>c m" shows "\M; P\ \\<^sub>r m" using m P(1) proof (induction m rule: intruder_synth_induct) case (ComposeC T f) hence "\M; P\ \\<^sub>r t" when t: "t \ set T" for t using t P(2) subtermeqI''[of _ T f] by fastforce thus ?case using intruder_deduct_restricted.ComposeR[OF ComposeC.hyps(1,2)] ComposeC.prems(1) by metis qed simp lemma deduct_zero_in_ik: assumes "\M; 0\ \\<^sub>n t" shows "t \ M" proof - { fix k assume "\M; k\ \\<^sub>n t" hence "k > 0 \ t \ M" by (induct t) auto } thus ?thesis using assms by auto qed lemma deduct_if_deduct_num: "\M; k\ \\<^sub>n t \ M \ t" by (induct t rule: intruder_deduct_num.induct) (metis intruder_deduct.Axiom, metis intruder_deduct.Compose, metis intruder_deduct.Decompose) lemma deduct_num_if_deduct: "M \ t \ \k. \M; k\ \\<^sub>n t" proof (induction t rule: intruder_deduct_induct) case (Compose T f) then obtain steps where *: "\t \ set T. \M; steps t\ \\<^sub>n t" by moura then obtain n where "\t \ set T. steps t \ n" using finite_nat_set_iff_bounded_le[of "steps ` set T"] by auto thus ?case using ComposeN[OF Compose.hyps(1,2), of M steps] * by force next case (Decompose t K T t\<^sub>i) hence "\u. u \ insert t (set K) \ \k. \M; k\ \\<^sub>n u" by auto then obtain steps where *: "\M; steps t\ \\<^sub>n t" "\t \ set K. \M; steps t\ \\<^sub>n t" by moura then obtain n where "steps t \ n" "\t \ set K. steps t \ n" using finite_nat_set_iff_bounded_le[of "steps ` insert t (set K)"] by auto thus ?case using DecomposeN[OF _ Decompose.hyps(2) _ Decompose.hyps(4), of M _ steps] * by force qed (metis AxiomN) lemma deduct_normalize: assumes M: "\m \ M. \f T. Fun f T \ m \ P f T" and t: "\M; k\ \\<^sub>n t" "Fun f T \ t" "\P f T" shows "\l \ k. (\M; l\ \\<^sub>n Fun f T) \ (\t \ set T. \j < l. \M; j\ \\<^sub>n t)" using t proof (induction t rule: intruder_deduct_num_induct) case (AxiomN t) thus ?case using M by auto next case (ComposeN T' f' steps) thus ?case proof (cases "Fun f' T' = Fun f T") case True hence "\M; Suc (Max (insert 0 (steps ` set T')))\ \\<^sub>n Fun f T" "T = T'" using intruder_deduct_num.ComposeN[OF ComposeN.hyps] by auto moreover have "\t. t \ set T \ \M; steps t\ \\<^sub>n t" using True ComposeN.hyps(3) by auto moreover have "\t. t \ set T \ steps t < Suc (Max (insert 0 (steps ` set T)))" using Max_less_iff[of "insert 0 (steps ` set T)" "Suc (Max (insert 0 (steps ` set T)))"] by auto ultimately show ?thesis by auto next case False then obtain t' where t': "t' \ set T'" "Fun f T \ t'" using ComposeN by auto hence "\l \ steps t'. (\M; l\ \\<^sub>n Fun f T) \ (\t \ set T. \j < l. \M; j\ \\<^sub>n t)" using ComposeN.IH[OF _ _ ComposeN.prems(2)] by auto moreover have "steps t' < Suc (Max (insert 0 (steps ` set T')))" using Max_less_iff[of "insert 0 (steps ` set T')" "Suc (Max (insert 0 (steps ` set T')))"] using t'(1) by auto ultimately show ?thesis using ComposeN.hyps(3)[OF t'(1)] by (meson Suc_le_eq le_Suc_eq le_trans) qed next case (DecomposeN t K T' t\<^sub>i steps n) hence *: "Fun f T \ t" using term.order_trans[of "Fun f T" t\<^sub>i t] Ana_subterm[of t K T'] by blast have "\l \ n. (\M; l\ \\<^sub>n Fun f T) \ (\t' \ set T. \j < l. \M; j\ \\<^sub>n t')" using DecomposeN.IH(1)[OF * DecomposeN.prems(2)] by auto moreover have "n < Suc (Max (insert n (steps ` set K)))" using Max_less_iff[of "insert n (steps ` set K)" "Suc (Max (insert n (steps ` set K)))"] by auto ultimately show ?case using DecomposeN.hyps(4) by (meson Suc_le_eq le_Suc_eq le_trans) qed lemma deduct_inv: assumes "\M; n\ \\<^sub>n t" shows "t \ M \ (\f T. t = Fun f T \ public f \ length T = arity f \ (\t \ set T. \l < n. \M; l\ \\<^sub>n t)) \ (\m \ subterms\<^sub>s\<^sub>e\<^sub>t M. (\l < n. \M; l\ \\<^sub>n m) \ (\k \ set (fst (Ana m)). \l < n. \M; l\ \\<^sub>n k) \ t \ set (snd (Ana m)))" (is "?P t n \ ?Q t n \ ?R t n") using assms proof (induction n arbitrary: t rule: nat_less_induct) case (1 n t) thus ?case proof (cases n) case 0 hence "t \ M" using deduct_zero_in_ik "1.prems"(1) by metis thus ?thesis by auto next case (Suc n') hence "\M; Suc n'\ \\<^sub>n t" "\m < Suc n'. \x. (\M; m\ \\<^sub>n x) \ ?P x m \ ?Q x m \ ?R x m" using "1.prems" "1.IH" by blast+ hence "?P t (Suc n') \ ?Q t (Suc n') \ ?R t (Suc n')" proof (induction t rule: intruder_deduct_num_induct) case (AxiomN t) thus ?case by simp next case (ComposeN T f steps) have "\t. t \ set T \ steps t < Suc (Max (insert 0 (steps ` set T)))" using Max_less_iff[of "insert 0 (steps ` set T)" "Suc (Max (insert 0 (steps ` set T)))"] by auto thus ?case using ComposeN.hyps by metis next case (DecomposeN t K T t\<^sub>i steps n) have 0: "n < Suc (Max (insert n (steps ` set K)))" "\k. k \ set K \ steps k < Suc (Max (insert n (steps ` set K)))" using Max_less_iff[of "insert n (steps ` set K)" "Suc (Max (insert n (steps ` set K)))"] by auto have IH1: "?P t j \ ?Q t j \ ?R t j" when jt: "j < n" "\M; j\ \\<^sub>n t" for j t using jt DecomposeN.prems(1) 0(1) by simp have IH2: "?P t n \ ?Q t n \ ?R t n" using DecomposeN.IH(1) IH1 by simp have 1: "\k \ set (fst (Ana t)). \l < Suc (Max (insert n (steps ` set K))). \M; l\ \\<^sub>n k" using DecomposeN.hyps(1,2,3) 0(2) by auto have 2: "t\<^sub>i \ set (snd (Ana t))" using DecomposeN.hyps(2,4) by fastforce have 3: "t \ subterms\<^sub>s\<^sub>e\<^sub>t M" when "t \ set (snd (Ana m))" "m \\<^sub>s\<^sub>e\<^sub>t M" for m using that(1) Ana_subterm[of m _ "snd (Ana m)"] in_subterms_subset_Union[OF that(2)] by (metis (no_types, lifting) prod.collapse psubsetD subsetCE subsetD) have 4: "?R t\<^sub>i (Suc (Max (insert n (steps ` set K))))" when "?R t n" using that 0(1) 1 2 3 DecomposeN.hyps(1) by (metis (no_types, lifting)) have 5: "?R t\<^sub>i (Suc (Max (insert n (steps ` set K))))" when "?P t n" using that 0(1) 1 2 DecomposeN.hyps(1) by blast have 6: ?case when *: "?Q t n" proof - obtain g S where g: "t = Fun g S" "public g" "length S = arity g" "\t \ set S. \l < n. \M; l\ \\<^sub>n t" using * by moura then obtain l where l: "l < n" "\M; l\ \\<^sub>n t\<^sub>i" using 0(1) DecomposeN.hyps(2,4) Ana_fun_subterm[of g S K T] by blast have **: "l < Suc (Max (insert n (steps ` set K)))" using l(1) 0(1) by simp show ?thesis using IH1[OF l] less_trans[OF _ **] by fastforce qed show ?case using IH2 4 5 6 by argo qed thus ?thesis using Suc by fast qed qed lemma deduct_inv': assumes "M \ Fun f ts" shows "Fun f ts \\<^sub>s\<^sub>e\<^sub>t M \ (\t \ set ts. M \ t)" proof - obtain k where k: "intruder_deduct_num M k (Fun f ts)" using deduct_num_if_deduct[OF assms] by fast have "Fun f ts \\<^sub>s\<^sub>e\<^sub>t M \ (\t \ set ts. \l. intruder_deduct_num M l t)" using deduct_inv[OF k] Ana_subterm'[of "Fun f ts"] in_subterms_subset_Union by blast thus ?thesis using deduct_if_deduct_num by blast qed lemma restricted_deduct_if_deduct: assumes M: "\m \ M. \f T. Fun f T \ m \ P (Fun f T)" and P_subterm: "\f T t. M \ Fun f T \ P (Fun f T) \ t \ set T \ P t" and P_Ana_key: "\t K T k. M \ t \ P t \ Ana t = (K, T) \ M \ k \ k \ set K \ P k" and m: "M \ m" "P m" shows "\M; P\ \\<^sub>r m" proof - { fix k assume "\M; k\ \\<^sub>n m" hence ?thesis using m(2) proof (induction k arbitrary: m rule: nat_less_induct) case (1 n m) thus ?case proof (cases n) case 0 hence "m \ M" using deduct_zero_in_ik "1.prems"(1) by metis thus ?thesis by auto next case (Suc n') hence "\M; Suc n'\ \\<^sub>n m" "\m < Suc n'. \x. (\M; m\ \\<^sub>n x) \ P x \ \M;P\ \\<^sub>r x" using "1.prems" "1.IH" by blast+ thus ?thesis using "1.prems"(2) proof (induction m rule: intruder_deduct_num_induct) case (ComposeN T f steps) have *: "steps t < Suc (Max (insert 0 (steps ` set T)))" when "t \ set T" for t using Max_less_iff[of "insert 0 (steps ` set T)"] that by blast have **: "P t" when "t \ set T" for t using P_subterm ComposeN.prems(2) that Fun_param_is_subterm[OF that] intruder_deduct.Compose[OF ComposeN.hyps(1,2)] deduct_if_deduct_num[OF ComposeN.hyps(3)] by blast have "\M; P\ \\<^sub>r t" when "t \ set T" for t using ComposeN.prems(1) ComposeN.hyps(3)[OF that] *[OF that] **[OF that] by blast thus ?case by (metis intruder_deduct_restricted.ComposeR[OF ComposeN.hyps(1,2)] ComposeN.prems(2)) next case (DecomposeN t K T t\<^sub>i steps l) show ?case proof (cases "P t") case True hence "\k. k \ set K \ P k" using P_Ana_key DecomposeN.hyps(1,2,3) deduct_if_deduct_num by blast moreover have "\k m x. k \ set K \ m < steps k \ \M; m\ \\<^sub>n x \ P x \ \M;P\ \\<^sub>r x" proof - fix k m x assume *: "k \ set K" "m < steps k" "\M; m\ \\<^sub>n x" "P x" have "steps k \ insert l (steps ` set K)" using *(1) by simp hence "m < Suc (Max (insert l (steps ` set K)))" using less_trans[OF *(2), of "Suc (Max (insert l (steps ` set K)))"] Max_less_iff[of "insert l (steps ` set K)" "Suc (Max (insert l (steps ` set K)))"] by auto thus "\M;P\ \\<^sub>r x" using DecomposeN.prems(1) *(3,4) by simp qed ultimately have "\k. k \ set K \ \M; P\ \\<^sub>r k" using DecomposeN.IH(2) by auto moreover have "\M; P\ \\<^sub>r t" using True DecomposeN.prems(1) DecomposeN.hyps(1) le_imp_less_Suc Max_less_iff[of "insert l (steps ` set K)" "Suc (Max (insert l (steps ` set K)))"] by blast ultimately show ?thesis using intruder_deduct_restricted.DecomposeR[OF _ DecomposeN.hyps(2) _ DecomposeN.hyps(4)] by metis next case False obtain g S where gS: "t = Fun g S" using DecomposeN.hyps(2,4) by (cases t) moura+ hence *: "Fun g S \ t" "\P (Fun g S)" using False by force+ have "\jM; j\ \\<^sub>n t\<^sub>i" using gS DecomposeN.hyps(2,4) Ana_fun_subterm[of g S K T] deduct_normalize[of M "\f T. P (Fun f T)", OF M DecomposeN.hyps(1) *] by force hence "\jM; j\ \\<^sub>n t\<^sub>i" using Max_less_iff[of "insert l (steps ` set K)" "Suc (Max (insert l (steps ` set K)))"] less_trans[of _ l "Suc (Max (insert l (steps ` set K)))"] by blast thus ?thesis using DecomposeN.prems(1,2) by meson qed qed auto qed qed } thus ?thesis using deduct_num_if_deduct m(1) by metis qed lemma restricted_deduct_if_deduct': assumes "\m \ M. P m" and "\t t'. P t \ t' \ t \ P t'" and "\t K T k. P t \ Ana t = (K, T) \ k \ set K \ P k" and "M \ m" "P m" shows "\M; P\ \\<^sub>r m" using restricted_deduct_if_deduct[of M P m] assms by blast lemma private_const_deduct: assumes c: "\public c" "M \ (Fun c []::('fun,'var) term)" shows "Fun c [] \ M \ (\m \ subterms\<^sub>s\<^sub>e\<^sub>t M. M \ m \ (\k \ set (fst (Ana m)). M \ m) \ Fun c [] \ set (snd (Ana m)))" proof - obtain n where "\M; n\ \\<^sub>n Fun c []" using c(2) deduct_num_if_deduct by moura hence "Fun c [] \ M \ (\m \ subterms\<^sub>s\<^sub>e\<^sub>t M. (\l < n. \M; l\ \\<^sub>n m) \ (\k \ set (fst (Ana m)). \l < n. \M; l\ \\<^sub>n k) \ Fun c [] \ set (snd (Ana m)))" using deduct_inv[of M n "Fun c []"] c(1) by fast thus ?thesis using deduct_if_deduct_num[of M] by blast qed lemma private_fun_deduct_in_ik'': assumes t: "M \ Fun f T" "Fun c [] \ set T" "\m \ subterms\<^sub>s\<^sub>e\<^sub>t M. Fun f T \ set (snd (Ana m))" and c: "\public c" "Fun c [] \ M" "\m \ subterms\<^sub>s\<^sub>e\<^sub>t M. Fun c [] \ set (snd (Ana m))" shows "Fun f T \ M" proof - have *: "\n. \M; n\ \\<^sub>n Fun c []" using private_const_deduct[OF c(1)] c(2,3) deduct_if_deduct_num by blast obtain n where n: "\M; n\ \\<^sub>n Fun f T" using t(1) deduct_num_if_deduct by blast show ?thesis using deduct_inv[OF n] t(2,3) * by blast qed end subsection \Executable Definitions for Code Generation\ fun intruder_synth' where "intruder_synth' pu ar M (Var x) = (Var x \ M)" | "intruder_synth' pu ar M (Fun f T) = ( Fun f T \ M \ (pu f \ length T = ar f \ list_all (intruder_synth' pu ar M) T))" definition "wf\<^sub>t\<^sub>r\<^sub>m' ar t \ (\s \ subterms t. is_Fun s \ ar (the_Fun s) = length (args s))" definition "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' ar M \ (\t \ M. wf\<^sub>t\<^sub>r\<^sub>m' ar t)" definition "analyzed_in' An pu ar t M \ (case An t of (K,T) \ (\k \ set K. intruder_synth' pu ar M k) \ (\s \ set T. intruder_synth' pu ar M s))" lemma (in intruder_model) intruder_synth'_induct[consumes 1, case_names Var Fun]: assumes "intruder_synth' public arity M t" "\x. intruder_synth' public arity M (Var x) \ P (Var x)" "\f T. (\z. z \ set T \ intruder_synth' public arity M z \ P z) \ intruder_synth' public arity M (Fun f T) \ P (Fun f T) " shows "P t" using assms by (induct public arity M t rule: intruder_synth'.induct) auto lemma (in intruder_model) wf\<^sub>t\<^sub>r\<^sub>m_code[code_unfold]: "wf\<^sub>t\<^sub>r\<^sub>m t = wf\<^sub>t\<^sub>r\<^sub>m' arity t" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def wf\<^sub>t\<^sub>r\<^sub>m'_def by auto lemma (in intruder_model) wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[code_unfold]: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M = wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity M" using wf\<^sub>t\<^sub>r\<^sub>m_code unfolding wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s'_def by auto lemma (in intruder_model) intruder_synth_code[code_unfold]: "intruder_synth M t = intruder_synth' public arity M t" (is "?A \ ?B") proof show "?A \ ?B" proof (induction t rule: intruder_synth_induct) case (AxiomC t) thus ?case by (cases t) auto qed (fastforce simp add: list_all_iff) show "?B \ ?A" proof (induction t rule: intruder_synth'_induct) case (Fun f T) thus ?case proof (cases "Fun f T \ M") case False hence "public f" "length T = arity f" "list_all (intruder_synth' public arity M) T" using Fun.hyps by fastforce+ thus ?thesis using Fun.IH intruder_synth.ComposeC[of T f M] Ball_set[of T] by blast qed simp qed simp qed lemma (in intruder_model) analyzed_in_code[code_unfold]: "analyzed_in t M = analyzed_in' Ana public arity t M" using intruder_synth_code[of M] unfolding analyzed_in_def analyzed_in'_def by fastforce end diff --git a/thys/Stateful_Protocol_Composition_and_Typing/More_Unification.thy b/thys/Stateful_Protocol_Composition_and_Typing/More_Unification.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/More_Unification.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/More_Unification.thy @@ -1,3335 +1,3299 @@ (* Based on src/HOL/ex/Unification.thy packaged with Isabelle/HOL 2015 having the following license: ISABELLE COPYRIGHT NOTICE, LICENCE AND DISCLAIMER. Copyright (c) 1986-2015, University of Cambridge, Technische Universitaet Muenchen, and contributors. 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 University of Cambridge or the Technische Universitaet Muenchen nor the names of their 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: More_Unification.thy Author: Andreas Viktor Hess, DTU SPDX-License-Identifier: BSD-3-Clause Originally based on src/HOL/ex/Unification.thy (Isabelle/HOL 2015) by: Author: Martin Coen, Cambridge University Computer Laboratory Author: Konrad Slind, TUM & Cambridge University Computer Laboratory Author: Alexander Krauss, TUM *) section \Definitions and Properties Related to Substitutions and Unification\ theory More_Unification imports Messages "First_Order_Terms.Unification" begin subsection \Substitutions\ abbreviation subst_apply_list (infix "\\<^sub>l\<^sub>i\<^sub>s\<^sub>t" 51) where "T \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ \ map (\t. t \ \) T" abbreviation subst_apply_pair (infixl "\\<^sub>p" 60) where "d \\<^sub>p \ \ (case d of (t,t') \ (t \ \, t' \ \))" abbreviation subst_apply_pair_set (infixl "\\<^sub>p\<^sub>s\<^sub>e\<^sub>t" 60) where "M \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \ (\d. d \\<^sub>p \) ` M" definition subst_apply_pairs (infix "\\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s" 51) where "F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \ \ map (\f. f \\<^sub>p \) F" abbreviation subst_more_general_than (infixl "\\<^sub>\" 50) where "\ \\<^sub>\ \ \ \\. \ = \ \\<^sub>s \" abbreviation subst_support (infix "supports" 50) where "\ supports \ \ (\x. \ x \ \ = \ x)" abbreviation rm_var where "rm_var v s \ s(v := Var v)" abbreviation rm_vars where "rm_vars vs \ \ (\v. if v \ vs then Var v else \ v)" definition subst_elim where "subst_elim \ v \ \t. v \ fv (t \ \)" definition subst_idem where "subst_idem s \ s \\<^sub>s s = s" lemma subst_support_def: "\ supports \ \ \ = \ \\<^sub>s \" unfolding subst_compose_def by metis lemma subst_supportD: "\ supports \ \ \ \\<^sub>\ \" using subst_support_def by auto lemma rm_vars_empty[simp]: "rm_vars {} s = s" "rm_vars (set []) s = s" by simp_all lemma rm_vars_singleton: "rm_vars {v} s = rm_var v s" by auto lemma subst_apply_terms_empty: "M \\<^sub>s\<^sub>e\<^sub>t Var = M" by simp lemma subst_agreement: "(t \ r = t \ s) \ (\v \ fv t. Var v \ r = Var v \ s)" by (induct t) auto lemma repl_invariance[dest?]: "v \ fv t \ t \ s(v := u) = t \ s" by (simp add: subst_agreement) lemma subst_idx_map: assumes "\i \ set I. i < length T" shows "(map ((!) T) I) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = map ((!) (map (\t. t \ \) T)) I" using assms by auto lemma subst_idx_map': assumes "\i \ fv\<^sub>s\<^sub>e\<^sub>t (set K). i < length T" shows "(K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t (!) T) \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \ = K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t ((!) (map (\t. t \ \) T))" (is "?A = ?B") proof - have "T ! i \ \ = (map (\t. t \ \) T) ! i" when "i < length T" for i using that by auto hence "T ! i \ \ = (map (\t. t \ \) T) ! i" when "i \ fv\<^sub>s\<^sub>e\<^sub>t (set K)" for i using that assms by auto hence "k \ (!) T \ \ = k \ (!) (map (\t. t \ \) T)" when "fv k \ fv\<^sub>s\<^sub>e\<^sub>t (set K)" for k using that by (induction k) force+ thus ?thesis by auto qed lemma subst_remove_var: "v \ fv s \ v \ fv (t \ Var(v := s))" by (induct t) simp_all lemma subst_set_map: "x \ set X \ x \ s \ set (map (\x. x \ s) X)" by simp lemma subst_set_idx_map: assumes "\i \ I. i < length T" shows "(!) T ` I \\<^sub>s\<^sub>e\<^sub>t \ = (!) (map (\t. t \ \) T) ` I" (is "?A = ?B") proof have *: "T ! i \ \ = (map (\t. t \ \) T) ! i" when "i < length T" for i using that by auto show "?A \ ?B" using * assms by blast show "?B \ ?A" using * assms by auto qed lemma subst_set_idx_map': assumes "\i \ fv\<^sub>s\<^sub>e\<^sub>t K. i < length T" shows "K \\<^sub>s\<^sub>e\<^sub>t (!) T \\<^sub>s\<^sub>e\<^sub>t \ = K \\<^sub>s\<^sub>e\<^sub>t (!) (map (\t. t \ \) T)" (is "?A = ?B") proof have "T ! i \ \ = (map (\t. t \ \) T) ! i" when "i < length T" for i using that by auto hence "T ! i \ \ = (map (\t. t \ \) T) ! i" when "i \ fv\<^sub>s\<^sub>e\<^sub>t K" for i using that assms by auto hence *: "k \ (!) T \ \ = k \ (!) (map (\t. t \ \) T)" when "fv k \ fv\<^sub>s\<^sub>e\<^sub>t K" for k using that by (induction k) force+ show "?A \ ?B" using * by auto show "?B \ ?A" using * by force qed lemma subst_term_list_obtain: assumes "\i < length T. \s. P (T ! i) s \ S ! i = s \ \" and "length T = length S" shows "\U. length T = length U \ (\i < length T. P (T ! i) (U ! i)) \ S = map (\u. u \ \) U" using assms proof (induction T arbitrary: S) case (Cons t T S') then obtain s S where S': "S' = s#S" by (cases S') auto have "\i < length T. \s. P (T ! i) s \ S ! i = s \ \" "length T = length S" using Cons.prems S' by force+ then obtain U where U: "length T = length U" "\i < length T. P (T ! i) (U ! i)" "S = map (\u. u \ \) U" using Cons.IH by moura obtain u where u: "P t u" "s = u \ \" using Cons.prems(1) S' by auto have 1: "length (t#T) = length (u#U)" using Cons.prems(2) U(1) by fastforce have 2: "\i < length (t#T). P ((t#T) ! i) ((u#U) ! i)" using u(1) U(2) by (simp add: nth_Cons') have 3: "S' = map (\u. u \ \) (u#U)" using U u S' by simp show ?case using 1 2 3 by blast qed simp lemma subst_mono: "t \ u \ t \ s \ u \ s" by (induct u) auto lemma subst_mono_fv: "x \ fv t \ s x \ t \ s" by (induct t) auto lemma subst_mono_neq: assumes "t \ u" shows "t \ s \ u \ s" proof (cases u) case (Var v) hence False using \t \ u\ by simp thus ?thesis .. next case (Fun f X) then obtain x where "x \ set X" "t \ x" using \t \ u\ by auto hence "t \ s \ x \ s" using subst_mono by metis obtain Y where "Fun f X \ s = Fun f Y" by auto hence "x \ s \ set Y" using \x \ set X\ by auto hence "x \ s \ Fun f X \ s" using \Fun f X \ s = Fun f Y\ Fun_param_is_subterm by simp hence "t \ s \ Fun f X \ s" using \t \ s \ x \ s\ by (metis term.dual_order.trans term.order.eq_iff) thus ?thesis using \u = Fun f X\ \t \ u\ by metis qed lemma subst_no_occs[dest]: "\Var v \ t \ t \ Var(v := s) = t" by (induct t) (simp_all add: map_idI) lemma var_comp[simp]: "\ \\<^sub>s Var = \" "Var \\<^sub>s \ = \" unfolding subst_compose_def by simp_all lemma subst_comp_all: "M \\<^sub>s\<^sub>e\<^sub>t (\ \\<^sub>s \) = (M \\<^sub>s\<^sub>e\<^sub>t \) \\<^sub>s\<^sub>e\<^sub>t \" using subst_subst_compose[of _ \ \] by auto lemma subst_all_mono: "M \ M' \ M \\<^sub>s\<^sub>e\<^sub>t s \ M' \\<^sub>s\<^sub>e\<^sub>t s" by auto lemma subst_comp_set_image: "(\ \\<^sub>s \) ` X = \ ` X \\<^sub>s\<^sub>e\<^sub>t \" using subst_compose by fastforce lemma subst_ground_ident[dest?]: "fv t = {} \ t \ s = t" by (induct t, simp, metis subst_agreement empty_iff subst_apply_term_empty) lemma subst_ground_ident_compose: "fv (\ x) = {} \ (\ \\<^sub>s \) x = \ x" "fv (t \ \) = {} \ t \ (\ \\<^sub>s \) = t \ \" using subst_subst_compose[of t \ \] by (simp_all add: subst_compose_def subst_ground_ident) lemma subst_all_ground_ident[dest?]: "ground M \ M \\<^sub>s\<^sub>e\<^sub>t s = M" proof - assume "ground M" hence "\t. t \ M \ fv t = {}" by auto hence "\t. t \ M \ t \ s = t" by (metis subst_ground_ident) moreover have "\t. t \ M \ t \ s \ M \\<^sub>s\<^sub>e\<^sub>t s" by (metis imageI) ultimately show "M \\<^sub>s\<^sub>e\<^sub>t s = M" by (simp add: image_cong) qed lemma subst_eqI[intro]: "(\t. t \ \ = t \ \) \ \ = \" proof - assume "\t. t \ \ = t \ \" hence "\v. Var v \ \ = Var v \ \" by auto thus "\ = \" by auto qed lemma subst_cong: "\\ = \'; \ = \'\ \ (\ \\<^sub>s \) = (\' \\<^sub>s \')" by auto lemma subst_mgt_bot[simp]: "Var \\<^sub>\ \" by simp lemma subst_mgt_refl[simp]: "\ \\<^sub>\ \" by (metis var_comp(1)) lemma subst_mgt_trans: "\\ \\<^sub>\ \; \ \\<^sub>\ \\ \ \ \\<^sub>\ \" by (metis subst_compose_assoc) lemma subst_mgt_comp: "\ \\<^sub>\ \ \\<^sub>s \" by auto lemma subst_mgt_comp': "\ \\<^sub>s \ \\<^sub>\ \ \ \ \\<^sub>\ \" by (metis subst_compose_assoc) lemma var_self: "(\w. if w = v then Var v else Var w) = Var" using subst_agreement by auto lemma var_same[simp]: "Var(v := t) = Var \ t = Var v" by (intro iffI, metis fun_upd_same, simp add: var_self) lemma subst_eq_if_eq_vars: "(\v. (Var v) \ \ = (Var v) \ \) \ \ = \" by (auto simp add: subst_agreement) lemma subst_all_empty[simp]: "{} \\<^sub>s\<^sub>e\<^sub>t \ = {}" by simp lemma subst_all_insert:"(insert t M) \\<^sub>s\<^sub>e\<^sub>t \ = insert (t \ \) (M \\<^sub>s\<^sub>e\<^sub>t \)" by auto lemma subst_apply_fv_subset: "fv t \ V \ fv (t \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` V)" by (induct t) auto lemma subst_apply_fv_empty: assumes "fv t = {}" shows "fv (t \ \) = {}" using assms subst_apply_fv_subset[of t "{}" \] by auto lemma subst_compose_fv: assumes "fv (\ x) = {}" shows "fv ((\ \\<^sub>s \) x) = {}" using assms subst_apply_fv_empty unfolding subst_compose_def by fast lemma subst_compose_fv': fixes \ \::"('a,'b) subst" assumes "y \ fv ((\ \\<^sub>s \) x)" shows "\z. z \ fv (\ x)" using assms subst_compose_fv by fast lemma subst_apply_fv_unfold: "fv (t \ \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv t)" by (induct t) auto lemma subst_apply_fv_unfold_set: "fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>s\<^sub>e\<^sub>t (set ts)) = fv\<^sub>s\<^sub>e\<^sub>t (set ts \\<^sub>s\<^sub>e\<^sub>t \)" by (simp add: subst_apply_fv_unfold) lemma subst_apply_fv_unfold': "fv (t \ \) = (\v \ fv t. fv (\ v))" using subst_apply_fv_unfold by simp lemma subst_apply_fv_union: "fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv (t \ \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ fv t))" proof - have "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (V \ fv t)) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` V) \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv t)" by auto thus ?thesis using subst_apply_fv_unfold by metis qed lemma subst_elimI[intro]: "(\t. v \ fv (t \ \)) \ subst_elim \ v" by (auto simp add: subst_elim_def) lemma subst_elimI'[intro]: "(\w. v \ fv (Var w \ \)) \ subst_elim \ v" by (simp add: subst_elim_def subst_apply_fv_unfold') lemma subst_elimD[dest]: "subst_elim \ v \ v \ fv (t \ \)" by (auto simp add: subst_elim_def) lemma subst_elimD'[dest]: "subst_elim \ v \ \ v \ Var v" by (metis subst_elim_def subst_apply_term.simps(1) term.set_intros(3)) lemma subst_elimD''[dest]: "subst_elim \ v \ v \ fv (\ w)" by (metis subst_elim_def subst_apply_term.simps(1)) lemma subst_elim_rm_vars_dest[dest]: "subst_elim (\::('a,'b) subst) v \ v \ vs \ subst_elim (rm_vars vs \) v" proof - assume assms: "subst_elim \ v" "v \ vs" obtain f::"('a, 'b) subst \ 'b \ 'b" where "\\ v. (\w. v \ fv (Var w \ \)) = (v \ fv (Var (f \ v) \ \))" by moura hence *: "\a \. a \ fv (Var (f \ a) \ \) \ subst_elim \ a" by blast have "Var (f (rm_vars vs \) v) \ \ \ Var (f (rm_vars vs \) v) \ rm_vars vs \ \ v \ fv (Var (f (rm_vars vs \) v) \ rm_vars vs \)" using assms(1) by fastforce moreover { assume "Var (f (rm_vars vs \) v) \ \ \ Var (f (rm_vars vs \) v) \ rm_vars vs \" hence "rm_vars vs \ (f (rm_vars vs \) v) \ \ (f (rm_vars vs \) v)" by auto hence "f (rm_vars vs \) v \ vs" by meson hence ?thesis using * assms(2) by force } ultimately show ?thesis using * by blast qed lemma occs_subst_elim: "\Var v \ t \ subst_elim (Var(v := t)) v \ (Var(v := t)) = Var" proof (cases "Var v = t") assume "Var v \ t" "\Var v \ t" hence "v \ fv t" by (simp add: vars_iff_subterm_or_eq) thus ?thesis by (auto simp add: subst_remove_var) qed auto lemma occs_subst_elim': "\Var v \ t \ subst_elim (Var(v := t)) v" proof - assume "\Var v \ t" hence "v \ fv t" by (auto simp add: vars_iff_subterm_or_eq) thus "subst_elim (Var(v := t)) v" by (simp add: subst_elim_def subst_remove_var) qed lemma subst_elim_comp: "subst_elim \ v \ subst_elim (\ \\<^sub>s \) v" by (auto simp add: subst_elim_def) lemma var_subst_idem: "subst_idem Var" by (simp add: subst_idem_def) lemma var_upd_subst_idem: assumes "\Var v \ t" shows "subst_idem (Var(v := t))" unfolding subst_idem_def proof let ?\ = "Var(v := t)" from assms have t_\_id: "t \ ?\ = t" by blast fix s show "s \ (?\ \\<^sub>s ?\) = s \ ?\" unfolding subst_compose_def by (induction s, metis t_\_id fun_upd_def subst_apply_term.simps(1), simp) qed lemma zip_map_subst: "zip xs (xs \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) = map (\t. (t, t \ \)) xs" by (induction xs) auto lemma map2_map_subst: "map2 f xs (xs \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \) = map (\t. f t (t \ \)) xs" by (induction xs) auto subsection \Lemmata: Domain and Range of Substitutions\ lemma range_vars_alt_def: "range_vars s \ fv\<^sub>s\<^sub>e\<^sub>t (subst_range s)" unfolding range_vars_def by simp lemma subst_dom_var_finite[simp]: "finite (subst_domain Var)" by simp lemma subst_range_Var[simp]: "subst_range Var = {}" by simp lemma range_vars_Var[simp]: "range_vars Var = {}" by fastforce lemma finite_subst_img_if_finite_dom: "finite (subst_domain \) \ finite (range_vars \)" unfolding range_vars_alt_def by auto lemma finite_subst_img_if_finite_dom': "finite (subst_domain \) \ finite (subst_range \)" by auto lemma subst_img_alt_def: "subst_range s = {t. \v. s v = t \ t \ Var v}" by (auto simp add: subst_domain_def) lemma subst_fv_img_alt_def: "range_vars s = (\t \ {t. \v. s v = t \ t \ Var v}. fv t)" unfolding range_vars_alt_def by (auto simp add: subst_domain_def) lemma subst_domI[intro]: "\ v \ Var v \ v \ subst_domain \" by (simp add: subst_domain_def) lemma subst_imgI[intro]: "\ v \ Var v \ \ v \ subst_range \" by (simp add: subst_domain_def) lemma subst_fv_imgI[intro]: "\ v \ Var v \ fv (\ v) \ range_vars \" unfolding range_vars_alt_def by auto lemma subst_eqI': assumes "t \ \ = t \ \" "subst_domain \ = subst_domain \" "subst_domain \ \ fv t" shows "\ = \" by (metis assms(2,3) term_subst_eq_rev[OF assms(1)] in_mono ext subst_domI) lemma subst_domain_subst_Fun_single[simp]: "subst_domain (Var(x := Fun f T)) = {x}" (is "?A = ?B") unfolding subst_domain_def by simp lemma subst_range_subst_Fun_single[simp]: "subst_range (Var(x := Fun f T)) = {Fun f T}" (is "?A = ?B") by simp lemma range_vars_subst_Fun_single[simp]: "range_vars (Var(x := Fun f T)) = fv (Fun f T)" unfolding range_vars_alt_def by force lemma var_renaming_is_Fun_iff: assumes "subst_range \ \ range Var" shows "is_Fun t = is_Fun (t \ \)" proof (cases t) case (Var x) hence "\y. \ x = Var y" using assms by auto thus ?thesis using Var by auto qed simp lemma subst_fv_dom_img_subset: "fv t \ subst_domain \ \ fv (t \ \) \ range_vars \" unfolding range_vars_alt_def by (induct t) auto lemma subst_fv_dom_img_subset_set: "fv\<^sub>s\<^sub>e\<^sub>t M \ subst_domain \ \ fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) \ range_vars \" proof - assume assms: "fv\<^sub>s\<^sub>e\<^sub>t M \ subst_domain \" obtain f::"'a set \ (('b, 'a) term \ 'a set) \ ('b, 'a) terms \ ('b, 'a) term" where "\x y z. (\v. v \ z \ \ y v \ x) \ (f x y z \ z \ \ y (f x y z) \ x)" by moura hence *: "\T g A. (\ \ (g ` T) \ A \ (\t. t \ T \ g t \ A)) \ (\ (g ` T) \ A \ f A g T \ T \ \ g (f A g T) \ A)" by (metis (no_types) SUP_le_iff) hence **: "\t. t \ M \ fv t \ subst_domain \" by (metis (no_types) assms fv\<^sub>s\<^sub>e\<^sub>t.simps) have "\t::('b, 'a) term. \f T. t \ f ` T \ (\t'::('b, 'a) term. t = f t' \ t' \ T)" by blast hence "f (range_vars \) fv (M \\<^sub>s\<^sub>e\<^sub>t \) \ M \\<^sub>s\<^sub>e\<^sub>t \ \ fv (f (range_vars \) fv (M \\<^sub>s\<^sub>e\<^sub>t \)) \ range_vars \" by (metis (full_types) ** subst_fv_dom_img_subset) thus ?thesis by (metis (no_types) * fv\<^sub>s\<^sub>e\<^sub>t.simps) qed lemma subst_fv_dom_ground_if_ground_img: assumes "fv t \ subst_domain s" "ground (subst_range s)" shows "fv (t \ s) = {}" using subst_fv_dom_img_subset[OF assms(1)] assms(2) by force lemma subst_fv_dom_ground_if_ground_img': assumes "fv t \ subst_domain s" "\x. x \ subst_domain s \ fv (s x) = {}" shows "fv (t \ s) = {}" using subst_fv_dom_ground_if_ground_img[OF assms(1)] assms(2) by auto lemma subst_fv_unfold: "fv (t \ s) = (fv t - subst_domain s) \ fv\<^sub>s\<^sub>e\<^sub>t (s ` (fv t \ subst_domain s))" proof (induction t) case (Var v) thus ?case proof (cases "v \ subst_domain s") case True thus ?thesis by auto next case False hence "fv (Var v \ s) = {v}" "fv (Var v) \ subst_domain s = {}" by auto thus ?thesis by auto qed next case Fun thus ?case by auto qed lemma subst_fv_unfold_ground_img: "range_vars s = {} \ fv (t \ s) = fv t - subst_domain s" using subst_fv_unfold[of t s] unfolding range_vars_alt_def by auto lemma subst_img_update: "\\ v = Var v; t \ Var v\ \ range_vars (\(v := t)) = range_vars \ \ fv t" proof - assume "\ v = Var v" "t \ Var v" hence "(\s \ {s. \w. (\(v := t)) w = s \ s \ Var w}. fv s) = fv t \ range_vars \" unfolding range_vars_alt_def by (auto simp add: subst_domain_def) thus "range_vars (\(v := t)) = range_vars \ \ fv t" by (metis Un_commute subst_fv_img_alt_def) qed lemma subst_dom_update1: "v \ subst_domain \ \ subst_domain (\(v := Var v)) = subst_domain \" by (auto simp add: subst_domain_def) lemma subst_dom_update2: "t \ Var v \ subst_domain (\(v := t)) = insert v (subst_domain \)" by (auto simp add: subst_domain_def) lemma subst_dom_update3: "t = Var v \ subst_domain (\(v := t)) = subst_domain \ - {v}" by (auto simp add: subst_domain_def) lemma var_not_in_subst_dom[elim]: "v \ subst_domain s \ s v = Var v" by (simp add: subst_domain_def) lemma subst_dom_vars_in_subst[elim]: "v \ subst_domain s \ s v \ Var v" by (simp add: subst_domain_def) lemma subst_not_dom_fixed: "\v \ fv t; v \ subst_domain s\ \ v \ fv (t \ s)" by (induct t) auto lemma subst_not_img_fixed: "\v \ fv (t \ s); v \ range_vars s\ \ v \ fv t" unfolding range_vars_alt_def by (induct t) force+ lemma ground_range_vars[intro]: "ground (subst_range s) \ range_vars s = {}" unfolding range_vars_alt_def by metis lemma ground_subst_no_var[intro]: "ground (subst_range s) \ x \ range_vars s" using ground_range_vars[of s] by blast lemma ground_img_obtain_fun: assumes "ground (subst_range s)" "x \ subst_domain s" obtains f T where "s x = Fun f T" "Fun f T \ subst_range s" "fv (Fun f T) = {}" proof - from assms(2) obtain t where t: "s x = t" "t \ subst_range s" by moura hence "fv t = {}" using assms(1) by auto thus ?thesis using t that by (cases t) simp_all qed lemma ground_term_subst_domain_fv_subset: "fv (t \ \) = {} \ fv t \ subst_domain \" by (induct t) auto lemma ground_subst_range_empty_fv: "ground (subst_range \) \ x \ subst_domain \ \ fv (\ x) = {}" by simp lemma subst_Var_notin_img: "x \ range_vars s \ t \ s = Var x \ t = Var x" using subst_not_img_fixed[of x t s] by (induct t) auto lemma fv_in_subst_img: "\s v = t; t \ Var v\ \ fv t \ range_vars s" unfolding range_vars_alt_def by auto lemma empty_dom_iff_empty_subst: "subst_domain \ = {} \ \ = Var" by auto lemma subst_dom_cong: "(\v t. \ v = t \ \ v = t) \ subst_domain \ \ subst_domain \" by (auto simp add: subst_domain_def) lemma subst_img_cong: "(\v t. \ v = t \ \ v = t) \ range_vars \ \ range_vars \" unfolding range_vars_alt_def by (auto simp add: subst_domain_def) lemma subst_dom_elim: "subst_domain s \ range_vars s = {} \ fv (t \ s) \ subst_domain s = {}" proof (induction t) case (Var v) thus ?case using fv_in_subst_img[of s] by (cases "s v = Var v") (auto simp add: subst_domain_def) next case Fun thus ?case by auto qed lemma subst_dom_insert_finite: "finite (subst_domain s) = finite (subst_domain (s(v := t)))" proof assume "finite (subst_domain s)" have "subst_domain (s(v := t)) \ insert v (subst_domain s)" by (auto simp add: subst_domain_def) thus "finite (subst_domain (s(v := t)))" by (meson \finite (subst_domain s)\ finite_insert rev_finite_subset) next assume *: "finite (subst_domain (s(v := t)))" hence "finite (insert v (subst_domain s))" proof (cases "t = Var v") case True hence "finite (subst_domain s - {v})" by (metis * subst_dom_update3) thus ?thesis by simp qed (metis * subst_dom_update2[of t v s]) thus "finite (subst_domain s)" by simp qed lemma trm_subst_disj: "t \ \ = t \ fv t \ subst_domain \ = {}" proof (induction t) case (Fun f X) hence "map (\x. x \ \) X = X" by simp hence "\x. x \ set X \ x \ \ = x" using map_eq_conv by fastforce thus ?case using Fun.IH by auto qed (simp add: subst_domain_def) lemma trm_subst_ident[intro]: "fv t \ subst_domain \ = {} \ t \ \ = t" proof - assume "fv t \ subst_domain \ = {}" hence "\v \ fv t. \w \ subst_domain \. v \ w" by auto thus ?thesis by (metis subst_agreement subst_apply_term.simps(1) subst_apply_term_empty subst_domI) qed lemma trm_subst_ident'[intro]: "v \ subst_domain \ \ (Var v) \ \ = Var v" using trm_subst_ident by (simp add: subst_domain_def) lemma trm_subst_ident''[intro]: "(\x. x \ fv t \ \ x = Var x) \ t \ \ = t" proof - assume "\x. x \ fv t \ \ x = Var x" hence "fv t \ subst_domain \ = {}" by (auto simp add: subst_domain_def) thus ?thesis using trm_subst_ident by auto qed lemma set_subst_ident: "fv\<^sub>s\<^sub>e\<^sub>t M \ subst_domain \ = {} \ M \\<^sub>s\<^sub>e\<^sub>t \ = M" proof - assume "fv\<^sub>s\<^sub>e\<^sub>t M \ subst_domain \ = {}" hence "\t \ M. t \ \ = t" by auto thus ?thesis by force qed lemma trm_subst_ident_subterms[intro]: "fv t \ subst_domain \ = {} \ subterms t \\<^sub>s\<^sub>e\<^sub>t \ = subterms t" using set_subst_ident[of "subterms t" \] fv_subterms[of t] by simp lemma trm_subst_ident_subterms'[intro]: "v \ fv t \ subterms t \\<^sub>s\<^sub>e\<^sub>t Var(v := s) = subterms t" using trm_subst_ident_subterms[of t "Var(v := s)"] by (meson subst_no_occs trm_subst_disj vars_iff_subtermeq) lemma const_mem_subst_cases: assumes "Fun c [] \ M \\<^sub>s\<^sub>e\<^sub>t \" shows "Fun c [] \ M \ Fun c [] \ \ ` fv\<^sub>s\<^sub>e\<^sub>t M" proof - obtain m where m: "m \ M" "m \ \ = Fun c []" using assms by auto thus ?thesis by (cases m) force+ qed lemma const_mem_subst_cases': assumes "Fun c [] \ M \\<^sub>s\<^sub>e\<^sub>t \" shows "Fun c [] \ M \ Fun c [] \ subst_range \" using const_mem_subst_cases[OF assms] by force lemma fv_subterms_substI[intro]: "y \ fv t \ \ y \ subterms t \\<^sub>s\<^sub>e\<^sub>t \" using image_iff vars_iff_subtermeq by fastforce lemma fv_subterms_subst_eq[simp]: "fv\<^sub>s\<^sub>e\<^sub>t (subterms (t \ \)) = fv\<^sub>s\<^sub>e\<^sub>t (subterms t \\<^sub>s\<^sub>e\<^sub>t \)" using fv_subterms by (induct t) force+ lemma fv_subterms_set_subst: "fv\<^sub>s\<^sub>e\<^sub>t (subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \) = fv\<^sub>s\<^sub>e\<^sub>t (subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \))" using fv_subterms_subst_eq[of _ \] by auto lemma fv_subterms_set_subst': "fv\<^sub>s\<^sub>e\<^sub>t (subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \) = fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" using fv_subterms_set[of "M \\<^sub>s\<^sub>e\<^sub>t \"] fv_subterms_set_subst[of \ M] by simp lemma fv_subst_subset: "x \ fv t \ fv (\ x) \ fv (t \ \)" by (metis fv_subset image_eqI subst_apply_fv_unfold) lemma fv_subst_subset': "fv s \ fv t \ fv (s \ \) \ fv (t \ \)" using fv_subst_subset by (induct s) force+ lemma fv_subst_obtain_var: fixes \::"('a,'b) subst" assumes "x \ fv (t \ \)" shows "\y \ fv t. x \ fv (\ y)" using assms by (induct t) force+ lemma set_subst_all_ident: "fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) \ subst_domain \ = {} \ M \\<^sub>s\<^sub>e\<^sub>t (\ \\<^sub>s \) = M \\<^sub>s\<^sub>e\<^sub>t \" by (metis set_subst_ident subst_comp_all) lemma subterms_subst: "subterms (t \ d) = (subterms t \\<^sub>s\<^sub>e\<^sub>t d) \ subterms\<^sub>s\<^sub>e\<^sub>t (d ` (fv t \ subst_domain d))" by (induct t) (auto simp add: subst_domain_def) lemma subterms_subst': fixes \::"('a,'b) subst" assumes "\x \ fv t. (\f. \ x = Fun f []) \ (\y. \ x = Var y)" shows "subterms (t \ \) = subterms t \\<^sub>s\<^sub>e\<^sub>t \" using assms proof (induction t) case (Var x) thus ?case proof (cases "x \ subst_domain \") case True hence "(\f. \ x = Fun f []) \ (\y. \ x = Var y)" using Var by simp hence "subterms (\ x) = {\ x}" by auto thus ?thesis by simp qed auto qed auto lemma subterms_subst'': fixes \::"('a,'b) subst" assumes "\x \ fv\<^sub>s\<^sub>e\<^sub>t M. (\f. \ x = Fun f []) \ (\y. \ x = Var y)" shows "subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) = subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" using subterms_subst'[of _ \] assms by auto lemma subterms_subst_subterm: fixes \::"('a,'b) subst" assumes "\x \ fv a. (\f. \ x = Fun f []) \ (\y. \ x = Var y)" and "b \ subterms (a \ \)" shows "\c \ subterms a. c \ \ = b" using subterms_subst'[OF assms(1)] assms(2) by auto lemma subterms_subst_subset: "subterms t \\<^sub>s\<^sub>e\<^sub>t \ \ subterms (t \ \)" by (induct t) auto lemma subterms_subst_subset': "subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \ \ subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" using subterms_subst_subset by fast lemma subterms\<^sub>s\<^sub>e\<^sub>t_subst: fixes \::"('a,'b) subst" assumes "t \ subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" shows "t \ subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \ \ (\x \ fv\<^sub>s\<^sub>e\<^sub>t M. t \ subterms (\ x))" using assms subterms_subst[of _ \] by auto lemma rm_vars_dom: "subst_domain (rm_vars V s) = subst_domain s - V" by (auto simp add: subst_domain_def) lemma rm_vars_dom_subset: "subst_domain (rm_vars V s) \ subst_domain s" by (auto simp add: subst_domain_def) lemma rm_vars_dom_eq': "subst_domain (rm_vars (UNIV - V) s) = subst_domain s \ V" using rm_vars_dom[of "UNIV - V" s] by blast lemma rm_vars_dom_eqI: assumes "t \ \ = t \ \" shows "subst_domain (rm_vars (UNIV - fv t) \) = subst_domain (rm_vars (UNIV - fv t) \)" by (meson assms Diff_iff UNIV_I term_subst_eq_rev) lemma rm_vars_img: "subst_range (rm_vars V s) = s ` subst_domain (rm_vars V s)" by (auto simp add: subst_domain_def) lemma rm_vars_img_subset: "subst_range (rm_vars V s) \ subst_range s" by (auto simp add: subst_domain_def) lemma rm_vars_img_fv_subset: "range_vars (rm_vars V s) \ range_vars s" unfolding range_vars_alt_def by (auto simp add: subst_domain_def) lemma rm_vars_fv_obtain: assumes "x \ fv (t \ rm_vars X \) - X" shows "\y \ fv t - X. x \ fv (rm_vars X \ y)" using assms by (induct t) (fastforce, force) lemma rm_vars_apply: "v \ subst_domain (rm_vars V s) \ (rm_vars V s) v = s v" by (auto simp add: subst_domain_def) lemma rm_vars_apply': "subst_domain \ \ vs = {} \ rm_vars vs \ = \" by force lemma rm_vars_ident: "fv t \ vs = {} \ t \ (rm_vars vs \) = t \ \" by (induct t) auto lemma rm_vars_fv_subset: "fv (t \ rm_vars X \) \ fv t \ fv (t \ \)" by (induct t) auto lemma rm_vars_fv_disj: assumes "fv t \ X = {}" "fv (t \ \) \ X = {}" shows "fv (t \ rm_vars X \) \ X = {}" using rm_vars_ident[OF assms(1)] assms(2) by auto lemma rm_vars_ground_supports: assumes "ground (subst_range \)" shows "rm_vars X \ supports \" proof fix x have *: "ground (subst_range (rm_vars X \))" using rm_vars_img_subset[of X \] assms by (auto simp add: subst_domain_def) show "rm_vars X \ x \ \ = \ x " proof (cases "x \ subst_domain (rm_vars X \)") case True hence "fv (rm_vars X \ x) = {}" using * by auto thus ?thesis using True by auto qed (simp add: subst_domain_def) qed lemma rm_vars_split: assumes "ground (subst_range \)" shows "\ = rm_vars X \ \\<^sub>s rm_vars (subst_domain \ - X) \" proof - let ?s1 = "rm_vars X \" let ?s2 = "rm_vars (subst_domain \ - X) \" have doms: "subst_domain ?s1 \ subst_domain \" "subst_domain ?s2 \ subst_domain \" by (auto simp add: subst_domain_def) { fix x assume "x \ subst_domain \" hence "\ x = Var x" "?s1 x = Var x" "?s2 x = Var x" using doms by auto hence "\ x = (?s1 \\<^sub>s ?s2) x" by (simp add: subst_compose_def) } moreover { fix x assume "x \ subst_domain \" "x \ X" hence "?s1 x = Var x" "?s2 x = \ x" using doms by auto hence "\ x = (?s1 \\<^sub>s ?s2) x" by (simp add: subst_compose_def) } moreover { fix x assume "x \ subst_domain \" "x \ X" hence "?s1 x = \ x" "fv (\ x) = {}" using assms doms by auto hence "\ x = (?s1 \\<^sub>s ?s2) x" by (simp add: subst_compose subst_ground_ident) } ultimately show ?thesis by blast qed lemma rm_vars_fv_img_disj: assumes "fv t \ X = {}" "X \ range_vars \ = {}" shows "fv (t \ rm_vars X \) \ X = {}" using assms proof (induction t) case (Var x) hence *: "(rm_vars X \) x = \ x" by auto show ?case proof (cases "x \ subst_domain \") case True hence "\ x \ subst_range \" by auto hence "fv (\ x) \ X = {}" using Var.prems(2) unfolding range_vars_alt_def by fastforce thus ?thesis using * by auto next case False thus ?thesis using Var.prems(1) by auto qed next case Fun thus ?case by auto qed lemma subst_apply_dom_ident: "t \ \ = t \ subst_domain \ \ subst_domain \ \ t \ \ = t" proof (induction t) case (Fun f T) thus ?case by (induct T) auto qed (auto simp add: subst_domain_def) lemma rm_vars_subst_apply_ident: assumes "t \ \ = t" shows "t \ (rm_vars vs \) = t" using rm_vars_dom[of vs \] subst_apply_dom_ident[OF assms, of "rm_vars vs \"] by auto lemma rm_vars_subst_eq: "t \ \ = t \ rm_vars (subst_domain \ - subst_domain \ \ fv t) \" by (auto intro: term_subst_eq) lemma rm_vars_subst_eq': "t \ \ = t \ rm_vars (UNIV - fv t) \" by (auto intro: term_subst_eq) lemma rm_vars_comp: assumes "range_vars \ \ vs = {}" shows "t \ rm_vars vs (\ \\<^sub>s \) = t \ (rm_vars vs \ \\<^sub>s rm_vars vs \)" using assms proof (induction t) case (Var x) thus ?case proof (cases "x \ vs") case True thus ?thesis using Var by auto next case False have "subst_domain (rm_vars vs \) \ vs = {}" by (auto simp add: subst_domain_def) moreover have "fv (\ x) \ vs = {}" using Var False unfolding range_vars_alt_def by force ultimately have "\ x \ (rm_vars vs \) = \ x \ \" using rm_vars_ident by (simp add: subst_domain_def) moreover have "(rm_vars vs (\ \\<^sub>s \)) x = (\ \\<^sub>s \) x" by (metis False) ultimately show ?thesis using subst_compose by auto qed next case Fun thus ?case by auto qed lemma rm_vars_fv\<^sub>s\<^sub>e\<^sub>t_subst: assumes "x \ fv\<^sub>s\<^sub>e\<^sub>t (rm_vars X \ ` Y)" shows "x \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` Y) \ x \ X" using assms by auto lemma disj_dom_img_var_notin: assumes "subst_domain \ \ range_vars \ = {}" "\ v = t" "t \ Var v" shows "v \ fv t" "\v \ fv (t \ \). v \ subst_domain \" proof - have "v \ subst_domain \" "fv t \ range_vars \" using fv_in_subst_img[of \ v t, OF assms(2)] assms(2,3) by (auto simp add: subst_domain_def) thus "v \ fv t" using assms(1) by auto have *: "fv t \ subst_domain \ = {}" using assms(1) \fv t \ range_vars \\ by auto hence "t \ \ = t" by blast thus "\v \ fv (t \ \). v \ subst_domain \" using * by auto qed lemma subst_sends_dom_to_img: "v \ subst_domain \ \ fv (Var v \ \) \ range_vars \" unfolding range_vars_alt_def by auto lemma subst_sends_fv_to_img: "fv (t \ s) \ fv t \ range_vars s" proof (induction t) case (Var v) thus ?case proof (cases "Var v \ s = Var v") case True thus ?thesis by simp next case False hence "v \ subst_domain s" by (meson trm_subst_ident') hence "fv (Var v \ s) \ range_vars s" using subst_sends_dom_to_img by simp thus ?thesis by auto qed next case Fun thus ?case by auto qed lemma ident_comp_subst_trm_if_disj: assumes "subst_domain \ \ range_vars \ = {}" "v \ subst_domain \" shows "(\ \\<^sub>s \) v = \ v" proof - from assms have " subst_domain \ \ fv (\ v) = {}" using fv_in_subst_img unfolding range_vars_alt_def by auto thus "(\ \\<^sub>s \) v = \ v" unfolding subst_compose_def by blast qed lemma ident_comp_subst_trm_if_disj': "fv (\ v) \ subst_domain \ = {} \ (\ \\<^sub>s \) v = \ v" unfolding subst_compose_def by blast lemma subst_idemI[intro]: "subst_domain \ \ range_vars \ = {} \ subst_idem \" using ident_comp_subst_trm_if_disj[of \ \] var_not_in_subst_dom[of _ \] subst_eq_if_eq_vars[of \] by (metis subst_idem_def subst_compose_def var_comp(2)) lemma subst_idemI'[intro]: "ground (subst_range \) \ subst_idem \" proof (intro subst_idemI) assume "ground (subst_range \)" hence "range_vars \ = {}" by (metis ground_range_vars) thus "subst_domain \ \ range_vars \ = {}" by blast qed lemma subst_idemE: "subst_idem \ \ subst_domain \ \ range_vars \ = {}" proof - assume "subst_idem \" hence "\v. fv (\ v) \ subst_domain \ = {}" unfolding subst_idem_def subst_compose_def by (metis trm_subst_disj) thus ?thesis unfolding range_vars_alt_def by auto qed lemma subst_idem_rm_vars: "subst_idem \ \ subst_idem (rm_vars X \)" proof - assume "subst_idem \" hence "subst_domain \ \ range_vars \ = {}" by (metis subst_idemE) moreover have "subst_domain (rm_vars X \) \ subst_domain \" "range_vars (rm_vars X \) \ range_vars \" unfolding range_vars_alt_def by (auto simp add: subst_domain_def) ultimately show ?thesis by blast qed lemma subst_fv_bounded_if_img_bounded: "range_vars \ \ fv t \ V \ fv (t \ \) \ fv t \ V" proof (induction t) case (Var v) thus ?case unfolding range_vars_alt_def by (cases "\ v = Var v") auto qed (metis (no_types, lifting) Un_assoc Un_commute subst_sends_fv_to_img sup.absorb_iff2) lemma subst_fv_bound_singleton: "fv (t \ Var(v := t')) \ fv t \ fv t'" using subst_fv_bounded_if_img_bounded[of "Var(v := t')" t "fv t'"] unfolding range_vars_alt_def by (auto simp add: subst_domain_def) lemma subst_fv_bounded_if_img_bounded': assumes "range_vars \ \ fv\<^sub>s\<^sub>e\<^sub>t M" shows "fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t M" proof fix v assume *: "v \ fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" obtain t where t: "t \ M" "t \ \ \ M \\<^sub>s\<^sub>e\<^sub>t \" "v \ fv (t \ \)" proof - assume **: "\t. \t \ M; t \ \ \ M \\<^sub>s\<^sub>e\<^sub>t \; v \ fv (t \ \)\ \ thesis" have "v \ \ (fv ` ((\t. t \ \) ` M))" using * by (metis fv\<^sub>s\<^sub>e\<^sub>t.simps) hence "\t. t \ M \ v \ fv (t \ \)" by blast thus ?thesis using ** imageI by blast qed from \t \ M\ obtain M' where "t \ M'" "M = insert t M'" by (meson Set.set_insert) hence "fv\<^sub>s\<^sub>e\<^sub>t M = fv t \ fv\<^sub>s\<^sub>e\<^sub>t M'" by simp hence "fv (t \ \) \ fv\<^sub>s\<^sub>e\<^sub>t M" using subst_fv_bounded_if_img_bounded assms by simp thus "v \ fv\<^sub>s\<^sub>e\<^sub>t M" using assms \v \ fv (t \ \)\ by auto qed lemma ground_img_if_ground_subst: "(\v t. s v = t \ fv t = {}) \ range_vars s = {}" unfolding range_vars_alt_def by auto lemma ground_subst_fv_subset: "ground (subst_range \) \ fv (t \ \) \ fv t" using subst_fv_bounded_if_img_bounded[of \] unfolding range_vars_alt_def by force lemma ground_subst_fv_subset': "ground (subst_range \) \ fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) \ fv\<^sub>s\<^sub>e\<^sub>t M" using subst_fv_bounded_if_img_bounded'[of \ M] unfolding range_vars_alt_def by auto lemma subst_to_var_is_var[elim]: "t \ s = Var v \ \w. t = Var w" using subst_apply_term.elims by blast lemma subst_dom_comp_inI: assumes "y \ subst_domain \" and "y \ subst_domain \" shows "y \ subst_domain (\ \\<^sub>s \)" using assms subst_domain_subst_compose[of \ \] by blast lemma subst_comp_notin_dom_eq: "x \ subst_domain \1 \ (\1 \\<^sub>s \2) x = \2 x" unfolding subst_compose_def by fastforce lemma subst_dom_comp_eq: assumes "subst_domain \ \ range_vars \ = {}" shows "subst_domain (\ \\<^sub>s \) = subst_domain \ \ subst_domain \" proof (rule ccontr) assume "subst_domain (\ \\<^sub>s \) \ subst_domain \ \ subst_domain \" hence "subst_domain (\ \\<^sub>s \) \ subst_domain \ \ subst_domain \" using subst_domain_compose[of \ \] by (simp add: subst_domain_def) then obtain v where "v \ subst_domain (\ \\<^sub>s \)" "v \ subst_domain \ \ subst_domain \" by auto hence v_in_some_subst: "\ v \ Var v \ \ v \ Var v" and "\ v \ \ = Var v" unfolding subst_compose_def by (auto simp add: subst_domain_def) then obtain w where "\ v = Var w" using subst_to_var_is_var by fastforce show False proof (cases "v = w") case True hence "\ v = Var v" using \\ v = Var w\ by simp hence "\ v \ Var v" using v_in_some_subst by simp thus False using \\ v = Var v\ \\ v \ \ = Var v\ by simp next case False hence "v \ subst_domain \" using v_in_some_subst \\ v \ \ = Var v\ by auto hence "v \ range_vars \" using assms by auto moreover have "\ w = Var v" using \\ v \ \ = Var v\ \\ v = Var w\ by simp hence "v \ range_vars \" using \v \ w\ subst_fv_imgI[of \ w] by simp ultimately show False .. qed qed lemma subst_img_comp_subset[simp]: "range_vars (\1 \\<^sub>s \2) \ range_vars \1 \ range_vars \2" proof let ?img = "range_vars" fix x assume "x \ ?img (\1 \\<^sub>s \2)" then obtain v t where vt: "x \ fv t" "t = (\1 \\<^sub>s \2) v" "t \ Var v" unfolding range_vars_alt_def subst_compose_def by (auto simp add: subst_domain_def) { assume "x \ ?img \1" hence "x \ ?img \2" by (metis (no_types, opaque_lifting) fv_in_subst_img Un_iff subst_compose_def vt subsetCE subst_apply_term.simps(1) subst_sends_fv_to_img) } thus "x \ ?img \1 \ ?img \2" by auto qed lemma subst_img_comp_subset': assumes "t \ subst_range (\1 \\<^sub>s \2)" shows "t \ subst_range \2 \ (\t' \ subst_range \1. t = t' \ \2)" proof - obtain x where x: "x \ subst_domain (\1 \\<^sub>s \2)" "(\1 \\<^sub>s \2) x = t" "t \ Var x" using assms by (auto simp add: subst_domain_def) { assume "x \ subst_domain \1" hence "(\1 \\<^sub>s \2) x = \2 x" unfolding subst_compose_def by auto hence ?thesis using x by auto } moreover { assume "x \ subst_domain \1" hence ?thesis using subst_compose x(2) by fastforce } ultimately show ?thesis by metis qed lemma subst_img_comp_subset'': "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range (\1 \\<^sub>s \2)) \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \2) \ ((subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \1)) \\<^sub>s\<^sub>e\<^sub>t \2)" proof fix t assume "t \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range (\1 \\<^sub>s \2))" then obtain x where x: "x \ subst_domain (\1 \\<^sub>s \2)" "t \ subterms ((\1 \\<^sub>s \2) x)" by auto show "t \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \2) \ (subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \1) \\<^sub>s\<^sub>e\<^sub>t \2)" proof (cases "x \ subst_domain \1") case True thus ?thesis using subst_compose[of \1 \2] x(2) subterms_subst by fastforce next case False hence "(\1 \\<^sub>s \2) x = \2 x" unfolding subst_compose_def by auto thus ?thesis using x by (auto simp add: subst_domain_def) qed qed lemma subst_img_comp_subset''': "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range (\1 \\<^sub>s \2)) - range Var \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \2) - range Var \ ((subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \1) - range Var) \\<^sub>s\<^sub>e\<^sub>t \2)" proof fix t assume t: "t \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range (\1 \\<^sub>s \2)) - range Var" then obtain f T where fT: "t = Fun f T" by (cases t) simp_all then obtain x where x: "x \ subst_domain (\1 \\<^sub>s \2)" "Fun f T \ subterms ((\1 \\<^sub>s \2) x)" using t by auto have "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \2) \ (subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \1) - range Var \\<^sub>s\<^sub>e\<^sub>t \2)" proof (cases "x \ subst_domain \1") case True hence "Fun f T \ (subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \2)) \ (subterms (\1 x) \\<^sub>s\<^sub>e\<^sub>t \2)" using x(2) subterms_subst[of "\1 x" \2] unfolding subst_compose[of \1 \2 x] by auto moreover have ?thesis when *: "Fun f T \ subterms (\1 x) \\<^sub>s\<^sub>e\<^sub>t \2" proof - obtain s where s: "s \ subterms (\1 x)" "Fun f T = s \ \2" using * by moura show ?thesis proof (cases s) case (Var y) hence "Fun f T \ subst_range \2" using s by force thus ?thesis by blast next case (Fun g S) hence "Fun f T \ (subterms (\1 x) - range Var) \\<^sub>s\<^sub>e\<^sub>t \2" using s by blast thus ?thesis using True by auto qed qed ultimately show ?thesis by blast next case False hence "(\1 \\<^sub>s \2) x = \2 x" unfolding subst_compose_def by auto thus ?thesis using x by (auto simp add: subst_domain_def) qed thus "t \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \2) - range Var \ (subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \1) - range Var \\<^sub>s\<^sub>e\<^sub>t \2)" using fT by auto qed lemma subst_img_comp_subset_const: assumes "Fun c [] \ subst_range (\1 \\<^sub>s \2)" shows "Fun c [] \ subst_range \2 \ Fun c [] \ subst_range \1 \ (\x. Var x \ subst_range \1 \ \2 x = Fun c [])" proof (cases "Fun c [] \ subst_range \2") case False then obtain t where t: "t \ subst_range \1" "Fun c [] = t \ \2" using subst_img_comp_subset'[OF assms] by auto thus ?thesis by (cases t) auto qed (simp add: subst_img_comp_subset'[OF assms]) lemma subst_img_comp_subset_const': fixes \ \::"('f,'v) subst" assumes "(\ \\<^sub>s \) x = Fun c []" shows "\ x = Fun c [] \ (\z. \ x = Var z \ \ z = Fun c [])" proof (cases "\ x = Fun c []") case False then obtain t where "\ x = t" "t \ \ = Fun c []" using assms unfolding subst_compose_def by auto thus ?thesis by (cases t) auto qed simp lemma subst_img_comp_subset_ground: assumes "ground (subst_range \1)" shows "subst_range (\1 \\<^sub>s \2) \ subst_range \1 \ subst_range \2" proof fix t assume t: "t \ subst_range (\1 \\<^sub>s \2)" then obtain x where x: "x \ subst_domain (\1 \\<^sub>s \2)" "t = (\1 \\<^sub>s \2) x" by auto show "t \ subst_range \1 \ subst_range \2" proof (cases "x \ subst_domain \1") case True hence "fv (\1 x) = {}" using assms ground_subst_range_empty_fv by fast hence "t = \1 x" using x(2) unfolding subst_compose_def by blast thus ?thesis using True by simp next case False hence "t = \2 x" "x \ subst_domain \2" using x subst_domain_compose[of \1 \2] by (metis subst_comp_notin_dom_eq, blast) thus ?thesis using x by simp qed qed lemma subst_fv_dom_img_single: assumes "v \ fv t" "\ v = t" "\w. v \ w \ \ w = Var w" shows "subst_domain \ = {v}" "range_vars \ = fv t" proof - show "subst_domain \ = {v}" using assms by (fastforce simp add: subst_domain_def) have "fv t \ range_vars \" by (metis fv_in_subst_img assms(1,2) vars_iff_subterm_or_eq) moreover have "\v. \ v \ Var v \ \ v = t" using assms by fastforce ultimately show "range_vars \ = fv t" unfolding range_vars_alt_def by (auto simp add: subst_domain_def) qed lemma subst_comp_upd1: "\(v := t) \\<^sub>s \ = (\ \\<^sub>s \)(v := t \ \)" unfolding subst_compose_def by auto lemma subst_comp_upd2: assumes "v \ subst_domain s" "v \ range_vars s" shows "s(v := t) = s \\<^sub>s (Var(v := t))" unfolding subst_compose_def proof - { fix w have "(s(v := t)) w = s w \ Var(v := t)" proof (cases "w = v") case True hence "s w = Var w" using \v \ subst_domain s\ by (simp add: subst_domain_def) thus ?thesis using \w = v\ by simp next case False hence "(s(v := t)) w = s w" by simp moreover have "s w \ Var(v := t) = s w" using \w \ v\ \v \ range_vars s\ by (metis fv_in_subst_img fun_upd_apply insert_absorb insert_subset repl_invariance subst_apply_term.simps(1) subst_apply_term_empty) ultimately show ?thesis .. qed } thus "s(v := t) = (\w. s w \ Var(v := t))" by auto qed lemma ground_subst_dom_iff_img: "ground (subst_range \) \ x \ subst_domain \ \ \ x \ subst_range \" by (auto simp add: subst_domain_def) lemma finite_dom_subst_exists: "finite S \ \\::('f,'v) subst. subst_domain \ = S" proof (induction S rule: finite.induct) case (insertI A a) then obtain \::"('f,'v) subst" where "subst_domain \ = A" by blast fix f::'f have "subst_domain (\(a := Fun f [])) = insert a A" using \subst_domain \ = A\ by (auto simp add: subst_domain_def) thus ?case by metis qed (auto simp add: subst_domain_def) lemma subst_inj_is_bij_betw_dom_img_if_ground_img: assumes "ground (subst_range \)" shows "inj \ \ bij_betw \ (subst_domain \) (subst_range \)" (is "?A \ ?B") proof show "?A \ ?B" by (metis bij_betw_def injD inj_onI subst_range.simps) next assume ?B hence "inj_on \ (subst_domain \)" unfolding bij_betw_def by auto moreover have "\x. x \ UNIV - subst_domain \ \ \ x = Var x" by auto hence "inj_on \ (UNIV - subst_domain \)" using inj_onI[of "UNIV - subst_domain \"] by (metis term.inject(1)) moreover have "\x y. x \ subst_domain \ \ y \ subst_domain \ \ \ x \ \ y" using assms by (auto simp add: subst_domain_def) ultimately show ?A by (metis injI inj_onD subst_domI term.inject(1)) qed lemma bij_finite_ground_subst_exists: assumes "finite (S::'v set)" "infinite (U::('f,'v) term set)" "ground U" shows "\\::('f,'v) subst. subst_domain \ = S \ bij_betw \ (subst_domain \) (subst_range \) \ subst_range \ \ U" proof - obtain T' where "T' \ U" "card T' = card S" "finite T'" by (meson assms(2) finite_Diff2 infinite_arbitrarily_large) then obtain f::"'v \ ('f,'v) term" where f_bij: "bij_betw f S T'" using finite_same_card_bij[OF assms(1)] by metis hence *: "\v. v \ S \ f v \ Var v" using \ground U\ \T' \ U\ bij_betwE by fastforce let ?\ = "\v. if v \ S then f v else Var v" have "subst_domain ?\ = S" proof show "subst_domain ?\ \ S" by (auto simp add: subst_domain_def) { fix v assume "v \ S" "v \ subst_domain ?\" hence "f v = Var v" by (simp add: subst_domain_def) hence False using *[OF \v \ S\] by metis } thus "S \ subst_domain ?\" by blast qed hence "\v w. \v \ subst_domain ?\; w \ subst_domain ?\\ \ ?\ w \ ?\ v" using \ground U\ bij_betwE[OF f_bij] set_rev_mp[OF _ \T' \ U\] by (metis (no_types, lifting) UN_iff empty_iff vars_iff_subterm_or_eq fv\<^sub>s\<^sub>e\<^sub>t.simps) hence "inj_on ?\ (subst_domain ?\)" using f_bij \subst_domain ?\ = S\ unfolding bij_betw_def inj_on_def by metis hence "bij_betw ?\ (subst_domain ?\) (subst_range ?\)" using inj_on_imp_bij_betw[of ?\] by simp moreover have "subst_range ?\ = T'" using \bij_betw f S T'\ \subst_domain ?\ = S\ unfolding bij_betw_def by auto hence "subst_range ?\ \ U" using \T' \ U\ by auto ultimately show ?thesis using \subst_domain ?\ = S\ by (metis (lifting)) qed lemma bij_finite_const_subst_exists: assumes "finite (S::'v set)" "finite (T::'f set)" "infinite (U::'f set)" shows "\\::('f,'v) subst. subst_domain \ = S \ bij_betw \ (subst_domain \) (subst_range \) \ subst_range \ \ (\c. Fun c []) ` (U - T)" proof - obtain T' where "T' \ U - T" "card T' = card S" "finite T'" by (meson assms(2,3) finite_Diff2 infinite_arbitrarily_large) then obtain f::"'v \ 'f" where f_bij: "bij_betw f S T'" using finite_same_card_bij[OF assms(1)] by metis let ?\ = "\v. if v \ S then Fun (f v) [] else Var v" have "subst_domain ?\ = S" by (simp add: subst_domain_def) moreover have "\v w. \v \ subst_domain ?\; w \ subst_domain ?\\ \ ?\ w \ ?\ v" by auto hence "inj_on ?\ (subst_domain ?\)" using f_bij unfolding bij_betw_def inj_on_def by (metis \subst_domain ?\ = S\ term.inject(2)) hence "bij_betw ?\ (subst_domain ?\) (subst_range ?\)" using inj_on_imp_bij_betw[of ?\] by simp moreover have "subst_range ?\ = ((\c. Fun c []) ` T')" using \bij_betw f S T'\ unfolding bij_betw_def inj_on_def by (auto simp add: subst_domain_def) hence "subst_range ?\ \ ((\c. Fun c []) ` (U - T))" using \T' \ U - T\ by auto ultimately show ?thesis by (metis (lifting)) qed lemma bij_finite_const_subst_exists': assumes "finite (S::'v set)" "finite (T::('f,'v) terms)" "infinite (U::'f set)" shows "\\::('f,'v) subst. subst_domain \ = S \ bij_betw \ (subst_domain \) (subst_range \) \ subst_range \ \ ((\c. Fun c []) ` U) - T" proof - have "finite (\(funs_term ` T))" using assms(2) by auto then obtain \ where \: "subst_domain \ = S" "bij_betw \ (subst_domain \) (subst_range \)" "subst_range \ \ (\c. Fun c []) ` (U - (\(funs_term ` T)))" using bij_finite_const_subst_exists[OF assms(1) _ assms(3)] by blast moreover have "(\c. Fun c []) ` (U - (\(funs_term ` T))) \ ((\c. Fun c []) ` U) - T" by auto ultimately show ?thesis by blast qed lemma bij_betw_iteI: assumes "bij_betw f A B" "bij_betw g C D" "A \ C = {}" "B \ D = {}" shows "bij_betw (\x. if x \ A then f x else g x) (A \ C) (B \ D)" proof - have "bij_betw (\x. if x \ A then f x else g x) A B" by (metis bij_betw_cong[of A f "\x. if x \ A then f x else g x" B] assms(1)) moreover have "bij_betw (\x. if x \ A then f x else g x) C D" using bij_betw_cong[of C g "\x. if x \ A then f x else g x" D] assms(2,3) by force ultimately show ?thesis using bij_betw_combine[OF _ _ assms(4)] by metis qed lemma subst_comp_split: assumes "subst_domain \ \ range_vars \ = {}" shows "\ = (rm_vars (subst_domain \ - V) \) \\<^sub>s (rm_vars V \)" (is ?P) and "\ = (rm_vars V \) \\<^sub>s (rm_vars (subst_domain \ - V) \)" (is ?Q) proof - let ?rm1 = "rm_vars (subst_domain \ - V) \" and ?rm2 = "rm_vars V \" have "subst_domain ?rm2 \ range_vars ?rm1 = {}" "subst_domain ?rm1 \ range_vars ?rm2 = {}" using assms unfolding range_vars_alt_def by (force simp add: subst_domain_def)+ hence *: "\v. v \ subst_domain ?rm1 \ (?rm1 \\<^sub>s ?rm2) v = \ v" "\v. v \ subst_domain ?rm2 \ (?rm2 \\<^sub>s ?rm1) v = \ v" using ident_comp_subst_trm_if_disj[of ?rm2 ?rm1] ident_comp_subst_trm_if_disj[of ?rm1 ?rm2] by (auto simp add: subst_domain_def) hence "\v. v \ subst_domain ?rm1 \ (?rm1 \\<^sub>s ?rm2) v = \ v" "\v. v \ subst_domain ?rm2 \ (?rm2 \\<^sub>s ?rm1) v = \ v" unfolding subst_compose_def by (auto simp add: subst_domain_def) hence "\v. (?rm1 \\<^sub>s ?rm2) v = \ v" "\v. (?rm2 \\<^sub>s ?rm1) v = \ v" using * by blast+ thus ?P ?Q by auto qed lemma subst_comp_eq_if_disjoint_vars: assumes "(subst_domain \ \ range_vars \) \ (subst_domain \ \ range_vars \) = {}" shows "\ \\<^sub>s \ = \ \\<^sub>s \" proof - { fix x assume "x \ subst_domain \" hence "(\ \\<^sub>s \) x = \ x" "(\ \\<^sub>s \) x = \ x" using assms unfolding range_vars_alt_def by (force simp add: subst_compose)+ hence "(\ \\<^sub>s \) x = (\ \\<^sub>s \) x" by metis } moreover { fix x assume "x \ subst_domain \" hence "(\ \\<^sub>s \) x = \ x" "(\ \\<^sub>s \) x = \ x" using assms unfolding range_vars_alt_def by (auto simp add: subst_compose subst_domain_def) hence "(\ \\<^sub>s \) x = (\ \\<^sub>s \) x" by metis } moreover { fix x assume "x \ subst_domain \" "x \ subst_domain \" hence "(\ \\<^sub>s \) x = (\ \\<^sub>s \) x" by (simp add: subst_compose subst_domain_def) } ultimately show ?thesis by auto qed lemma subst_eq_if_disjoint_vars_ground: fixes \ \::"('f,'v) subst" assumes "subst_domain \ \ subst_domain \ = {}" "ground (subst_range \)" "ground (subst_range \)" shows "t \ \ \ \ = t \ \ \ \" by (metis assms subst_comp_eq_if_disjoint_vars range_vars_alt_def subst_subst_compose sup_bot.right_neutral) lemma subst_img_bound: "subst_domain \ \ range_vars \ \ fv t \ range_vars \ \ fv (t \ \)" proof - assume "subst_domain \ \ range_vars \ \ fv t" hence "subst_domain \ \ fv t" by blast thus ?thesis by (metis (no_types) range_vars_alt_def le_iff_sup subst_apply_fv_unfold subst_apply_fv_union subst_range.simps) qed lemma subst_all_fv_subset: "fv t \ fv\<^sub>s\<^sub>e\<^sub>t M \ fv (t \ \) \ fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" proof - assume *: "fv t \ fv\<^sub>s\<^sub>e\<^sub>t M" { fix v assume "v \ fv t" hence "v \ fv\<^sub>s\<^sub>e\<^sub>t M" using * by auto then obtain t' where "t' \ M" "v \ fv t'" by auto hence "fv (\ v) \ fv (t' \ \)" by (metis subst_apply_term.simps(1) subst_apply_fv_subset subst_apply_fv_unfold subtermeq_vars_subset vars_iff_subtermeq) hence "fv (\ v) \ fv\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" using \t' \ M\ by auto } thus ?thesis using subst_apply_fv_unfold[of t \] by auto qed lemma subst_support_if_mgt_subst_idem: assumes "\ \\<^sub>\ \" "subst_idem \" shows "\ supports \" proof - from \\ \\<^sub>\ \\ obtain \ where \: "\ = \ \\<^sub>s \" by blast hence "\v. \ v \ \ = Var v \ (\ \\<^sub>s \ \\<^sub>s \)" by simp hence "\v. \ v \ \ = Var v \ (\ \\<^sub>s \)" using \subst_idem \ \ unfolding subst_idem_def by simp hence "\v. \ v \ \ = Var v \ \" using \ by simp thus "\ supports \" by simp qed lemma subst_support_iff_mgt_if_subst_idem: assumes "subst_idem \" shows "\ \\<^sub>\ \ \ \ supports \" proof show "\ \\<^sub>\ \ \ \ supports \" by (fact subst_support_if_mgt_subst_idem[OF _ \subst_idem \\]) show "\ supports \ \ \ \\<^sub>\ \" by (fact subst_supportD) qed lemma subst_support_comp: fixes \ \ \::"('a,'b) subst" assumes "\ supports \" "\ supports \" shows "(\ \\<^sub>s \) supports \" by (metis (no_types) assms subst_agreement subst_apply_term.simps(1) subst_subst_compose) lemma subst_support_comp': fixes \ \ \::"('a,'b) subst" assumes "\ supports \" shows "\ supports (\ \\<^sub>s \)" "\ supports \ \ \ supports (\ \\<^sub>s \)" using assms unfolding subst_support_def by (metis subst_compose_assoc, metis) lemma subst_support_comp_split: fixes \ \ \::"('a,'b) subst" assumes "(\ \\<^sub>s \) supports \" shows "subst_domain \ \ range_vars \ = {} \ \ supports \" and "subst_domain \ \ subst_domain \ = {} \ \ supports \" proof - assume "subst_domain \ \ range_vars \ = {}" hence "subst_idem \" by (metis subst_idemI) have "\ \\<^sub>\ \" using assms subst_compose_assoc[of \ \ \] unfolding subst_compose_def by metis show "\ supports \" using subst_support_if_mgt_subst_idem[OF \\ \\<^sub>\ \\ \subst_idem \\] by auto next assume "subst_domain \ \ subst_domain \ = {}" moreover have "\v \ subst_domain (\ \\<^sub>s \). (\ \\<^sub>s \) v \ \ = \ v" using assms by metis ultimately have "\v \ subst_domain \. \ v \ \ = \ v" using var_not_in_subst_dom unfolding subst_compose_def by (metis IntI empty_iff subst_apply_term.simps(1)) thus "\ supports \" by force qed lemma subst_idem_support: "subst_idem \ \ \ supports \ \\<^sub>s \" unfolding subst_idem_def by (metis subst_support_def subst_compose_assoc) lemma subst_idem_iff_self_support: "subst_idem \ \ \ supports \" using subst_support_def[of \ \] unfolding subst_idem_def by auto lemma subterm_subst_neq: "t \ t' \ t \ s \ t' \ s" by (metis subst_mono_neq) lemma fv_Fun_subst_neq: "x \ fv (Fun f T) \ \ x \ Fun f T \ \" using subterm_subst_neq[of "Var x" "Fun f T"] vars_iff_subterm_or_eq[of x "Fun f T"] by auto lemma subterm_subst_unfold: assumes "t \ s \ \" shows "(\s'. s' \ s \ t = s' \ \) \ (\x \ fv s. t \ \ x)" using assms proof (induction s) case (Fun f T) thus ?case proof (cases "t = Fun f T \ \") case True thus ?thesis using Fun by auto next case False then obtain s' where s': "s' \ set T" "t \ s' \ \" using Fun by auto hence "(\s''. s'' \ s' \ t = s'' \ \) \ (\x \ fv s'. t \ \ x)" by (metis Fun.IH) thus ?thesis using s'(1) by auto qed qed simp lemma subterm_subst_img_subterm: assumes "t \ s \ \" "\s'. s' \ s \ t \ s' \ \" shows "\w \ fv s. t \ \ w" using subterm_subst_unfold[OF assms(1)] assms(2) by force lemma subterm_subst_not_img_subterm: assumes "t \ s \ \" "\(\w \ fv s. t \ \ w)" shows "\f T. Fun f T \ s \ t = Fun f T \ \" proof (rule ccontr) assume "\(\f T. Fun f T \ s \ t = Fun f T \ \)" hence "\f T. Fun f T \ s \ t \ Fun f T \ \" by simp moreover have "\x. Var x \ s \ t \ Var x \ \" using assms(2) vars_iff_subtermeq by force ultimately have "\s'. s' \ s \ t \ s' \ \" by (metis "term.exhaust") thus False using assms subterm_subst_img_subterm by blast qed lemma subst_apply_img_var: assumes "v \ fv (t \ \)" "v \ fv t" obtains w where "w \ fv t" "v \ fv (\ w)" using assms by (induct t) auto lemma subst_apply_img_var': assumes "x \ fv (t \ \)" "x \ fv t" shows "\y \ fv t. x \ fv (\ y)" by (metis assms subst_apply_img_var) lemma nth_map_subst: fixes \::"('f,'v) subst" and T::"('f,'v) term list" and i::nat shows "i < length T \ (map (\t. t \ \) T) ! i = (T ! i) \ \" by (fact nth_map) lemma subst_subterm: assumes "Fun f T \ t \ \" shows "(\S. Fun f S \ t \ Fun f S \ \ = Fun f T) \ (\s \ subst_range \. Fun f T \ s)" using assms subterm_subst_not_img_subterm by (cases "\s \ subst_range \. Fun f T \ s") fastforce+ lemma subst_subterm': assumes "Fun f T \ t \ \" shows "\S. length S = length T \ (Fun f S \ t \ (\s \ subst_range \. Fun f S \ s))" using subst_subterm[OF assms] by auto lemma subst_subterm'': assumes "s \ subterms (t \ \)" shows "(\u \ subterms t. s = u \ \) \ s \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \)" proof (cases s) case (Var x) thus ?thesis using assms subterm_subst_not_img_subterm vars_iff_subtermeq by (cases "s = t \ \") fastforce+ next case (Fun f T) thus ?thesis using subst_subterm[of f T t \] assms by fastforce qed subsection \More Small Lemmata\ lemma funs_term_subst: "funs_term (t \ \) = funs_term t \ (\x \ fv t. funs_term (\ x))" by (induct t) auto lemma fv\<^sub>s\<^sub>e\<^sub>t_subst_img_eq: assumes "X \ (subst_domain \ \ range_vars \) = {}" shows "fv\<^sub>s\<^sub>e\<^sub>t (\ ` (Y - X)) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` Y) - X" using assms unfolding range_vars_alt_def by force lemma subst_Fun_index_eq: assumes "i < length T" "Fun f T \ \ = Fun g T' \ \" shows "T ! i \ \ = T' ! i \ \" proof - have "map (\x. x \ \) T = map (\x. x \ \) T'" using assms by simp thus ?thesis by (metis assms(1) length_map nth_map) qed lemma fv_exists_if_unifiable_and_neq: fixes t t'::"('a,'b) term" and \ \::"('a,'b) subst" assumes "t \ t'" "t \ \ = t' \ \" shows "fv t \ fv t' \ {}" proof assume "fv t \ fv t' = {}" hence "fv t = {}" "fv t' = {}" by auto hence "t \ \ = t" "t' \ \ = t'" by auto hence "t = t'" using assms(2) by metis thus False using assms(1) by auto qed lemma const_subterm_subst: "Fun c [] \ t \ Fun c [] \ t \ \" by (induct t) auto lemma const_subterm_subst_var_obtain: assumes "Fun c [] \ t \ \" "\Fun c [] \ t" obtains x where "x \ fv t" "Fun c [] \ \ x" using assms by (induct t) auto lemma const_subterm_subst_cases: assumes "Fun c [] \ t \ \" shows "Fun c [] \ t \ (\x \ fv t. x \ subst_domain \ \ Fun c [] \ \ x)" proof (cases "Fun c [] \ t") case False then obtain x where "x \ fv t" "Fun c [] \ \ x" using const_subterm_subst_var_obtain[OF assms] by moura thus ?thesis by (cases "x \ subst_domain \") auto qed simp lemma const_subterms_subst_cases: assumes "Fun c [] \\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" shows "Fun c [] \\<^sub>s\<^sub>e\<^sub>t M \ (\x \ fv\<^sub>s\<^sub>e\<^sub>t M. x \ subst_domain \ \ Fun c [] \ \ x)" using assms const_subterm_subst_cases[of c _ \] by auto lemma const_subterms_subst_cases': assumes "Fun c [] \\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" shows "Fun c [] \\<^sub>s\<^sub>e\<^sub>t M \ Fun c [] \\<^sub>s\<^sub>e\<^sub>t subst_range \" using const_subterms_subst_cases[OF assms] by auto lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_fv_subset: assumes "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" shows "fv (\ x) \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" using assms proof (induction F) case (Cons f F) then obtain t t' where f: "f = (t,t')" by (metis surj_pair) show ?case proof (cases "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F") case True thus ?thesis using Cons.IH unfolding subst_apply_pairs_def by auto next case False hence "x \ fv t \ fv t'" using Cons.prems f by simp hence "fv (\ x) \ fv (t \ \) \ fv (t' \ \)" using fv_subst_subset[of x] by force thus ?thesis using f unfolding subst_apply_pairs_def by auto qed qed simp lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_step_subst: "fv\<^sub>s\<^sub>e\<^sub>t (\ ` fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F) = fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" proof (induction F) case (Cons f F) obtain t t' where "f = (t,t')" by moura thus ?case using Cons by (simp add: subst_apply_pairs_def subst_apply_fv_unfold) qed (simp_all add: subst_apply_pairs_def) lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_obtain_var: fixes \::"('a,'b) subst" assumes "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" shows "\y \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F. x \ fv (\ y)" using assms proof (induction F) case (Cons f F) then obtain t s where f: "f = (t,s)" by (metis surj_pair) from Cons.IH show ?case proof (cases "x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)") case False hence "x \ fv (t \ \) \ x \ fv (s \ \)" using f Cons.prems by (simp add: subst_apply_pairs_def) hence "(\y \ fv t. x \ fv (\ y)) \ (\y \ fv s. x \ fv (\ y))" by (metis fv_subst_obtain_var) thus ?thesis using f by (auto simp add: subst_apply_pairs_def) qed (auto simp add: Cons.IH) qed (simp add: subst_apply_pairs_def) lemma pair_subst_ident[intro]: "(fv t \ fv t') \ subst_domain \ = {} \ (t,t') \\<^sub>p \ = (t,t')" by auto lemma pairs_substI[intro]: assumes "subst_domain \ \ (\(s,t) \ M. fv s \ fv t) = {}" shows "M \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ = M" proof - { fix m assume M: "m \ M" then obtain s t where m: "m = (s,t)" by (metis surj_pair) hence "(fv s \ fv t) \ subst_domain \ = {}" using assms M by auto hence "m \\<^sub>p \ = m" using m by auto } thus ?thesis by (simp add: image_cong) qed lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) = fv\<^sub>s\<^sub>e\<^sub>t (\ ` (fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F))" proof (induction F) case (Cons g G) obtain t t' where "g = (t,t')" by (metis surj_pair) thus ?case using Cons.IH by (simp add: subst_apply_pairs_def subst_apply_fv_unfold) qed (simp add: subst_apply_pairs_def) lemma fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst_subset: assumes "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \ subst_domain \" shows "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F \ subst_domain \ \ subst_domain \" using assms proof (induction F) case (Cons g G) hence IH: "fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s G \ subst_domain \ \ subst_domain \" by (simp add: subst_apply_pairs_def) 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 (simp_all add: subst_apply_pairs_def) hence "fv t \ subst_domain \ \ subst_domain \" "fv t' \ subst_domain \ \ subst_domain \" using subst_apply_fv_unfold[of _ \] by force+ thus ?case using IH g by (simp add: subst_apply_pairs_def) qed (simp add: subst_apply_pairs_def) lemma pairs_subst_comp: "F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \ \\<^sub>s \ = ((F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \)" by (induct F) (auto simp add: subst_apply_pairs_def) lemma pairs_substI'[intro]: "subst_domain \ \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F = {} \ F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \ = F" by (induct F) (force simp add: subst_apply_pairs_def)+ lemma subst_pair_compose[simp]: "d \\<^sub>p (\ \\<^sub>s \) = d \\<^sub>p \ \\<^sub>p \" proof - obtain t s where "d = (t,s)" by moura thus ?thesis by auto qed lemma subst_pairs_compose[simp]: "D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t (\ \\<^sub>s \) = D \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \ \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" by auto lemma subst_apply_pair_pair: "(t, s) \\<^sub>p \ = (t \ \, s \ \)" by (rule prod.case) lemma subst_apply_pairs_nil[simp]: "[] \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \ = []" unfolding subst_apply_pairs_def by simp lemma subst_apply_pairs_singleton[simp]: "[(t,s)] \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \ = [(t \ \,s \ \)]" unfolding subst_apply_pairs_def by simp lemma subst_apply_pairs_Var[iff]: "F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s Var = F" by (simp add: subst_apply_pairs_def) lemma subst_apply_pairs_pset_subst: "set (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s \) = set F \\<^sub>p\<^sub>s\<^sub>e\<^sub>t \" unfolding subst_apply_pairs_def by force subsection \Finite Substitutions\ inductive_set fsubst::"('a,'b) subst set" where fvar: "Var \ fsubst" | FUpdate: "\\ \ fsubst; v \ subst_domain \; t \ Var v\ \ \(v := t) \ fsubst" lemma finite_dom_iff_fsubst: "finite (subst_domain \) \ \ \ fsubst" proof assume "finite (subst_domain \)" thus "\ \ fsubst" proof (induction "subst_domain \" arbitrary: \ rule: finite.induct) case emptyI hence "\ = Var" using empty_dom_iff_empty_subst by metis thus ?case using fvar by simp next case (insertI \'\<^sub>d\<^sub>o\<^sub>m v) thus ?case proof (cases "v \ \'\<^sub>d\<^sub>o\<^sub>m") case True hence "\'\<^sub>d\<^sub>o\<^sub>m = subst_domain \" using \insert v \'\<^sub>d\<^sub>o\<^sub>m = subst_domain \\ by auto thus ?thesis using insertI.hyps(2) by metis next case False let ?\' = "\w. if w \ \'\<^sub>d\<^sub>o\<^sub>m then \ w else Var w" have "subst_domain ?\' = \'\<^sub>d\<^sub>o\<^sub>m" using \v \ \'\<^sub>d\<^sub>o\<^sub>m\ \insert v \'\<^sub>d\<^sub>o\<^sub>m = subst_domain \\ by (auto simp add: subst_domain_def) hence "?\' \ fsubst" using insertI.hyps(2) by simp moreover have "?\'(v := \ v) = (\w. if w \ insert v \'\<^sub>d\<^sub>o\<^sub>m then \ w else Var w)" by auto hence "?\'(v := \ v) = \" using \insert v \'\<^sub>d\<^sub>o\<^sub>m = subst_domain \\ by (auto simp add: subst_domain_def) ultimately show ?thesis using FUpdate[of ?\' v "\ v"] False insertI.hyps(3) by (auto simp add: subst_domain_def) qed qed next assume "\ \ fsubst" thus "finite (subst_domain \)" by (induct \, simp, metis subst_dom_insert_finite) qed lemma fsubst_induct[case_names fvar FUpdate, induct set: finite]: assumes "finite (subst_domain \)" "P Var" and "\\ v t. \finite (subst_domain \); v \ subst_domain \; t \ Var v; P \\ \ P (\(v := t))" shows "P \" using assms finite_dom_iff_fsubst fsubst.induct by metis lemma fun_upd_fsubst: "s(v := t) \ fsubst \ s \ fsubst" using subst_dom_insert_finite[of s] finite_dom_iff_fsubst by blast lemma finite_img_if_fsubst: "s \ fsubst \ finite (subst_range s)" using finite_dom_iff_fsubst finite_subst_img_if_finite_dom' by blast subsection \Unifiers and Most General Unifiers (MGUs)\ abbreviation Unifier::"('f,'v) subst \ ('f,'v) term \ ('f,'v) term \ bool" where "Unifier \ t u \ (t \ \ = u \ \)" abbreviation MGU::"('f,'v) subst \ ('f,'v) term \ ('f,'v) term \ bool" where "MGU \ t u \ Unifier \ t u \ (\\. Unifier \ t u \ \ \\<^sub>\ \)" lemma MGUI[intro]: shows "\t \ \ = u \ \; \\::('f,'v) subst. t \ \ = u \ \ \ \ \\<^sub>\ \\ \ MGU \ t u" by auto lemma UnifierD[dest]: fixes \::"('f,'v) subst" and f g::'f and X Y::"('f,'v) term list" assumes "Unifier \ (Fun f X) (Fun g Y)" shows "f = g" "length X = length Y" proof - from assms show "f = g" by auto from assms have "Fun f X \ \ = Fun g Y \ \" by auto hence "length (map (\x. x \ \) X) = length (map (\x. x \ \) Y)" by auto thus "length X = length Y" by auto qed lemma MGUD[dest]: fixes \::"('f,'v) subst" and f g::'f and X Y::"('f,'v) term list" assumes "MGU \ (Fun f X) (Fun g Y)" shows "f = g" "length X = length Y" using assms by (auto intro!: UnifierD[of f X \ g Y]) lemma MGU_sym[sym]: "MGU \ s t \ MGU \ t s" by auto lemma Unifier_sym[sym]: "Unifier \ s t \ Unifier \ t s" by auto lemma MGU_nil: "MGU Var s t \ s = t" by fastforce lemma Unifier_comp: "Unifier (\ \\<^sub>s \) t u \ Unifier \ (t \ \) (u \ \)" by simp lemma Unifier_comp': "Unifier \ (t \ \) (u \ \) \ Unifier (\ \\<^sub>s \) t u" by simp lemma Unifier_excludes_subterm: assumes \: "Unifier \ t u" shows "\t \ u" proof assume "t \ u" hence "t \ \ \ u \ \" using subst_mono_neq by metis hence "t \ \ \ u \ \" by simp moreover from \ have "t \ \ = u \ \" by auto ultimately show False .. qed lemma MGU_is_Unifier: "MGU \ t u \ Unifier \ t u" by (rule conjunct1) lemma MGU_Var1: assumes "\Var v \ t" shows "MGU (Var(v := t)) (Var v) t" proof (intro MGUI exI) show "Var v \ (Var(v := t)) = t \ (Var(v := t))" using assms subst_no_occs by fastforce next fix \::"('a,'b) subst" assume th: "Var v \ \ = t \ \" show "\ = (Var(v := t)) \\<^sub>s \" proof fix s show "s \ \ = s \ ((Var(v := t)) \\<^sub>s \)" using th by (induct s) auto qed qed lemma MGU_Var2: "v \ fv t \ MGU (Var(v := t)) (Var v) t" by (metis (no_types) MGU_Var1 vars_iff_subterm_or_eq) lemma MGU_Var3: "MGU Var (Var v) (Var w) \ v = w" by fastforce lemma MGU_Const1: "MGU Var (Fun c []) (Fun d []) \ c = d" by fastforce lemma MGU_Const2: "MGU \ (Fun c []) (Fun d []) \ c = d" by auto lemma MGU_Fun: assumes "MGU \ (Fun f X) (Fun g Y)" shows "f = g" "length X = length Y" proof - let ?F = "\\ X. map (\x. x \ \) X" from assms have "\f = g; ?F \ X = ?F \ Y; \\'. f = g \ ?F \' X = ?F \' Y \ \ \\<^sub>\ \'\ \ length X = length Y" using map_eq_imp_length_eq by auto thus "f = g" "length X = length Y" using assms by auto qed lemma Unifier_Fun: assumes "Unifier \ (Fun f (x#X)) (Fun g (y#Y))" shows "Unifier \ x y" "Unifier \ (Fun f X) (Fun g Y)" using assms by simp_all lemma Unifier_subst_idem_subst: "subst_idem r \ Unifier s (t \ r) (u \ r) \ Unifier (r \\<^sub>s s) (t \ r) (u \ r)" by (metis (no_types, lifting) subst_idem_def subst_subst_compose) lemma subst_idem_comp: "subst_idem r \ Unifier s (t \ r) (u \ r) \ (\q. Unifier q (t \ r) (u \ r) \ s \\<^sub>s q = q) \ subst_idem (r \\<^sub>s s)" by (frule Unifier_subst_idem_subst, blast, metis subst_idem_def subst_compose_assoc) lemma Unifier_mgt: "\Unifier \ t u; \ \\<^sub>\ \\ \ Unifier \ t u" by auto lemma Unifier_support: "\Unifier \ t u; \ supports \\ \ Unifier \ t u" using subst_supportD Unifier_mgt by metis lemma MGU_mgt: "\MGU \ t u; MGU \ t u\ \ \ \\<^sub>\ \" by auto lemma Unifier_trm_fv_bound: "\Unifier s t u; v \ fv t\ \ v \ subst_domain s \ range_vars s \ fv u" proof (induction t arbitrary: s u) case (Fun f X) hence "v \ fv (u \ s) \ v \ subst_domain s" by (metis subst_not_dom_fixed) thus ?case by (metis (no_types) Un_iff contra_subsetD subst_sends_fv_to_img) qed (metis (no_types) UnI1 UnI2 subsetCE no_var_subterm subst_sends_dom_to_img subst_to_var_is_var trm_subst_ident' vars_iff_subterm_or_eq) lemma Unifier_rm_var: "\Unifier \ s t; v \ fv s \ fv t\ \ Unifier (rm_var v \) s t" by (auto simp add: repl_invariance) lemma Unifier_ground_rm_vars: assumes "ground (subst_range s)" "Unifier (rm_vars X s) t t'" shows "Unifier s t t'" by (rule Unifier_support[OF assms(2) rm_vars_ground_supports[OF assms(1)]]) lemma Unifier_dom_restrict: assumes "Unifier s t t'" "fv t \ fv t' \ S" shows "Unifier (rm_vars (UNIV - S) s) t t'" proof - let ?s = "rm_vars (UNIV - S) s" show ?thesis using term_subst_eq_conv[of t s ?s] term_subst_eq_conv[of t' s ?s] assms by auto qed subsection \Well-formedness of Substitutions and Unifiers\ inductive_set wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set::"('a,'b) subst set" where Empty[simp]: "Var \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set" | Insert[simp]: "\\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set; v \ subst_domain \; v \ range_vars \; fv t \ (insert v (subst_domain \)) = {}\ \ \(v := t) \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set" definition wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t::"('a,'b) subst \ bool" where "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ subst_domain \ \ range_vars \ = {} \ finite (subst_domain \)" definition wf\<^sub>M\<^sub>G\<^sub>U::"('a,'b) subst \ ('a,'b) term \ ('a,'b) term \ bool" where "wf\<^sub>M\<^sub>G\<^sub>U \ s t \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ MGU \ s t \ subst_domain \ \ range_vars \ \ fv s \ fv t" lemma wf_subst_subst_idem: "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ subst_idem \" using subst_idemI[of \] unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by fast lemma wf_subst_properties: "\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set = wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" proof show "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ \ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set" unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def proof - assume "subst_domain \ \ range_vars \ = {} \ finite (subst_domain \)" hence "finite (subst_domain \)" "subst_domain \ \ range_vars \ = {}" by auto thus "\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set" proof (induction \ rule: fsubst_induct) case fvar thus ?case by simp next case (FUpdate \ v t) have "subst_domain \ \ subst_domain (\(v := t))" "range_vars \ \ range_vars (\(v := t))" using FUpdate.hyps(2,3) subst_img_update unfolding range_vars_alt_def by (fastforce simp add: subst_domain_def)+ hence "subst_domain \ \ range_vars \ = {}" using FUpdate.prems(1) by blast hence "\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set" using FUpdate.IH by metis have *: "range_vars (\(v := t)) = range_vars \ \ fv t" using FUpdate.hyps(2) subst_img_update[OF _ FUpdate.hyps(3)] by fastforce hence "fv t \ insert v (subst_domain \) = {}" using FUpdate.prems subst_dom_update2[OF FUpdate.hyps(3)] by blast moreover have "subst_domain (\(v := t)) = insert v (subst_domain \)" by (meson FUpdate.hyps(3) subst_dom_update2) hence "v \ range_vars \" using FUpdate.prems * by blast ultimately show ?case using Insert[OF \\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set\ \v \ subst_domain \\] by metis qed qed show "\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def proof (induction \ rule: wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set.induct) case Empty thus ?case by simp next case (Insert \ v t) hence 1: "subst_domain \ \ range_vars \ = {}" by simp hence 2: "subst_domain (\(v := t)) \ range_vars \ = {}" using Insert.hyps(3) by (auto simp add: subst_domain_def) have 3: "fv t \ subst_domain (\(v := t)) = {}" using Insert.hyps(4) by (auto simp add: subst_domain_def) have 4: "\ v = Var v" using \v \ subst_domain \\ by (simp add: subst_domain_def) from Insert.IH have "finite (subst_domain \)" by simp hence 5: "finite (subst_domain (\(v := t)))" using subst_dom_insert_finite[of \] by simp have "subst_domain (\(v := t)) \ range_vars (\(v := t)) = {}" proof (cases "t = Var v") case True hence "range_vars (\(v := t)) = range_vars \" using 4 fun_upd_triv term.inject(1) unfolding range_vars_alt_def by (auto simp add: subst_domain_def) thus "subst_domain (\(v := t)) \ range_vars (\(v := t)) = {}" using 1 2 3 by auto next case False hence "range_vars (\(v := t)) = fv t \ (range_vars \)" using 4 subst_img_update[of \ v] by auto thus "subst_domain (\(v := t)) \ range_vars (\(v := t)) = {}" using 1 2 3 by blast qed thus ?case using 5 by blast qed qed lemma wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_induct[consumes 1, case_names Empty Insert]: assumes "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "P Var" and "\\ v t. \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \; P \; v \ subst_domain \; v \ range_vars \; fv t \ insert v (subst_domain \) = {}\ \ P (\(v := t))" shows "P \" proof - from assms(1,3) wf_subst_properties have "\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set" "\\ v t. \\ \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set; P \; v \ subst_domain \; v \ range_vars \; fv t \ insert v (subst_domain \) = {}\ \ P (\(v := t))" by blast+ thus "P \" using wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set.induct assms(2) by blast qed lemma wf_subst_fsubst: "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ \ \ fsubst" unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def using finite_dom_iff_fsubst by blast lemma wf_subst_nil: "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by simp lemma wf_MGU_nil: "MGU Var s t \ wf\<^sub>M\<^sub>G\<^sub>U Var s t" using wf_subst_nil subst_domain_Var range_vars_Var unfolding wf\<^sub>M\<^sub>G\<^sub>U_def by fast lemma wf_MGU_dom_bound: "wf\<^sub>M\<^sub>G\<^sub>U \ s t \ subst_domain \ \ fv s \ fv t" unfolding wf\<^sub>M\<^sub>G\<^sub>U_def by blast lemma wf_subst_single: assumes "v \ fv t" "\ v = t" "\w. v \ w \ \ w = Var w" shows "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" proof - have *: "subst_domain \ = {v}" by (metis subst_fv_dom_img_single(1)[OF assms]) have "subst_domain \ \ range_vars \ = {}" using * assms subst_fv_dom_img_single(2) by (metis inf_bot_left insert_disjoint(1)) moreover have "finite (subst_domain \)" using * by simp ultimately show ?thesis by (metis wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) qed lemma wf_subst_reduction: "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t s \ wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (rm_var v s)" proof - assume "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t s" moreover have "subst_domain (rm_var v s) \ subst_domain s" by (auto simp add: subst_domain_def) moreover have "range_vars (rm_var v s) \ range_vars s" unfolding range_vars_alt_def by (auto simp add: subst_domain_def) ultimately have "subst_domain (rm_var v s) \ range_vars (rm_var v s) = {}" by (meson compl_le_compl_iff disjoint_eq_subset_Compl subset_trans wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) moreover have "finite (subst_domain (rm_var v s))" using \subst_domain (rm_var v s) \ subst_domain s\ \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t s\ rev_finite_subset unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by blast ultimately show "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (rm_var v s)" by (metis wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) qed lemma wf_subst_compose: assumes "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \1" "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \2" and "subst_domain \1 \ subst_domain \2 = {}" and "subst_domain \1 \ range_vars \2 = {}" shows "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\1 \\<^sub>s \2)" using assms proof (induction \1 rule: wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_induct) case Empty thus ?case unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by simp next case (Insert \1 v t) have "t \ Var v" using Insert.hyps(4) by auto hence dom1v_unfold: "subst_domain (\1(v := t)) = insert v (subst_domain \1)" using subst_dom_update2 by metis hence doms_disj: "subst_domain \1 \ subst_domain \2 = {}" using Insert.prems(2) disjoint_insert(1) by blast moreover have dom_img_disj: "subst_domain \1 \ range_vars \2 = {}" using Insert.hyps(2) Insert.prems(3) by (fastforce simp add: subst_domain_def) ultimately have "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\1 \\<^sub>s \2)" using Insert.IH[OF \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \2\] by metis have dom_comp_is_union: "subst_domain (\1 \\<^sub>s \2) = subst_domain \1 \ subst_domain \2" using subst_dom_comp_eq[OF dom_img_disj] . have "v \ subst_domain \2" using Insert.prems(2) \t \ Var v\ by (fastforce simp add: subst_domain_def) hence "\2 v = Var v" "\1 v = Var v" using Insert.hyps(2) by (simp_all add: subst_domain_def) hence "(\1 \\<^sub>s \2) v = Var v" "(\1(v := t) \\<^sub>s \2) v = t \ \2" "((\1 \\<^sub>s \2)(v := t)) v = t" unfolding subst_compose_def by simp_all have fv_t2_bound: "fv (t \ \2) \ fv t \ range_vars \2" by (meson subst_sends_fv_to_img) have 1: "v \ subst_domain (\1 \\<^sub>s \2)" using \(\1 \\<^sub>s \2) v = Var v\ by (auto simp add: subst_domain_def) have "insert v (subst_domain \1) \ range_vars \2 = {}" using Insert.prems(3) dom1v_unfold by blast hence "v \ range_vars \1 \ range_vars \2" using Insert.hyps(3) by blast hence 2: "v \ range_vars (\1 \\<^sub>s \2)" by (meson set_rev_mp subst_img_comp_subset) have "subst_domain \2 \ range_vars \2 = {}" using \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \2\ unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by simp hence "fv (t \ \2) \ subst_domain \2 = {}" using subst_dom_elim unfolding range_vars_alt_def by simp moreover have "v \ range_vars \2" using Insert.prems(3) dom1v_unfold by blast hence "v \ fv t \ range_vars \2" using Insert.hyps(4) by blast hence "v \ fv (t \ \2)" using \fv (t \ \2) \ fv t \ range_vars \2\ by blast moreover have "fv (t \ \2) \ subst_domain \1 = {}" using dom_img_disj fv_t2_bound \fv t \ insert v (subst_domain \1) = {}\ by blast ultimately have 3: "fv (t \ \2) \ insert v (subst_domain (\1 \\<^sub>s \2)) = {}" using dom_comp_is_union by blast have "\1(v := t) \\<^sub>s \2 = (\1 \\<^sub>s \2)(v := t \ \2)" using subst_comp_upd1[of \1 v t \2] . moreover have "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t ((\1 \\<^sub>s \2)(v := t \ \2))" using "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set.Insert"[OF _ 1 2 3] \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\1 \\<^sub>s \2)\ wf_subst_properties by metis ultimately show ?case by presburger qed lemma wf_subst_append: fixes \1 \2::"('f,'v) subst" assumes "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \1" "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \2" and "subst_domain \1 \ subst_domain \2 = {}" and "subst_domain \1 \ range_vars \2 = {}" and "range_vars \1 \ subst_domain \2 = {}" shows "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\v. if \1 v = Var v then \2 v else \1 v)" using assms proof (induction \1 rule: wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_induct) case Empty thus ?case unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by simp next case (Insert \1 v t) let ?if = "\w. if \1 w = Var w then \2 w else \1 w" let ?if_upd = "\w. if (\1(v := t)) w = Var w then \2 w else (\1(v := t)) w" from Insert.hyps(4) have "?if_upd = ?if(v := t)" by fastforce have dom_insert: "subst_domain (\1(v := t)) = insert v (subst_domain \1)" using Insert.hyps(4) by (auto simp add: subst_domain_def) have "\1 v = Var v" "t \ Var v" using Insert.hyps(2,4) by auto hence img_insert: "range_vars (\1(v := t)) = range_vars \1 \ fv t" using subst_img_update by metis from Insert.prems(2) dom_insert have "subst_domain \1 \ subst_domain \2 = {}" by (auto simp add: subst_domain_def) moreover have "subst_domain \1 \ range_vars \2 = {}" using Insert.prems(3) dom_insert by (simp add: subst_domain_def) moreover have "range_vars \1 \ subst_domain \2 = {}" using Insert.prems(4) img_insert by blast ultimately have "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t ?if" using Insert.IH[OF Insert.prems(1)] by metis have dom_union: "subst_domain ?if = subst_domain \1 \ subst_domain \2" by (auto simp add: subst_domain_def) hence "v \ subst_domain ?if" using Insert.hyps(2) Insert.prems(2) dom_insert by (auto simp add: subst_domain_def) moreover have "v \ range_vars ?if" using Insert.prems(3) Insert.hyps(3) dom_insert unfolding range_vars_alt_def by (auto simp add: subst_domain_def) moreover have "fv t \ insert v (subst_domain ?if) = {}" using Insert.hyps(4) Insert.prems(4) img_insert unfolding range_vars_alt_def by (fastforce simp add: subst_domain_def) ultimately show ?case using wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_set.Insert \wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t ?if\ \?if_upd = ?if(v := t)\ wf_subst_properties by (metis (no_types, lifting)) qed lemma wf_subst_elim_append: assumes "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "subst_elim \ v" "v \ fv t" shows "subst_elim (\(w := t)) v" using assms proof (induction \ rule: wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_induct) case (Insert \ v' t') hence "\q. v \ fv (Var q \ \(v' := t'))" using subst_elimD by blast hence "\q. v \ fv (Var q \ \(v' := t', w := t))" using \v \ fv t\ by simp thus ?case by (metis subst_elimI' subst_apply_term.simps(1)) qed (simp add: subst_elim_def) lemma wf_subst_elim_dom: assumes "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\v \ subst_domain \. subst_elim \ v" using assms proof (induction \ rule: wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_induct) case (Insert \ w t) have dom_insert: "subst_domain (\(w := t)) \ insert w (subst_domain \)" by (auto simp add: subst_domain_def) hence "\v \ subst_domain \. subst_elim (\(w := t)) v" using Insert.IH Insert.hyps(2,4) by (metis Insert.hyps(1) IntI disjoint_insert(2) empty_iff wf_subst_elim_append) moreover have "w \ fv t" using Insert.hyps(4) by simp hence "\q. w \ fv (Var q \ \(w := t))" by (metis fv_simps(1) fv_in_subst_img Insert.hyps(3) contra_subsetD fun_upd_def singletonD subst_apply_term.simps(1)) hence "subst_elim (\(w := t)) w" by (metis subst_elimI') ultimately show ?case using dom_insert by blast qed simp lemma wf_subst_support_iff_mgt: "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ \ supports \ \ \ \\<^sub>\ \" using subst_support_def subst_support_if_mgt_subst_idem wf_subst_subst_idem by blast subsection \Interpretations\ abbreviation interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t::"('a,'b) subst \ bool" where "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ subst_domain \ = UNIV \ ground (subst_range \)" lemma interpretation_substI: "(\v. fv (\ v) = {}) \ interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" proof - assume "\v. fv (\ v) = {}" moreover { fix v assume "fv (\ v) = {}" hence "v \ subst_domain \" by auto } ultimately show ?thesis by auto qed lemma interpretation_grounds[simp]: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ fv (t \ \) = {}" using subst_fv_dom_ground_if_ground_img[of t \] by blast lemma interpretation_grounds_all: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ (\v. fv (\ v) = {})" by (metis range_vars_alt_def UNIV_I fv_in_subst_img subset_empty subst_dom_vars_in_subst) lemma interpretation_grounds_all': "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ ground (M \\<^sub>s\<^sub>e\<^sub>t \)" using subst_fv_dom_ground_if_ground_img[of _ \] by simp lemma interpretation_comp: assumes "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" proof - have \_fv: "fv (\ v) = {}" for v using interpretation_grounds_all[OF assms] by simp hence \_fv': "fv (t \ \) = {}" for t by (metis all_not_in_conv subst_elimD subst_elimI' subst_apply_term.simps(1)) from assms have "(\ \\<^sub>s \) v \ Var v" for v unfolding subst_compose_def by (metis fv_simps(1) \_fv' insert_not_empty) hence "subst_domain (\ \\<^sub>s \) = UNIV" by (simp add: subst_domain_def) moreover have "fv ((\ \\<^sub>s \) v) = {}" for v unfolding subst_compose_def using \_fv' by simp hence "ground (subst_range (\ \\<^sub>s \))" by simp ultimately show "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" .. from assms have "(\ \\<^sub>s \) v \ Var v" for v unfolding subst_compose_def by (metis fv_simps(1) \_fv insert_not_empty subst_to_var_is_var) hence "subst_domain (\ \\<^sub>s \) = UNIV" by (simp add: subst_domain_def) moreover have "fv ((\ \\<^sub>s \) v) = {}" for v unfolding subst_compose_def by (simp add: \_fv trm_subst_ident) hence "ground (subst_range (\ \\<^sub>s \))" by simp ultimately show "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" .. qed lemma interpretation_subst_exists: "\\::('f,'v) subst. interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" proof - obtain c::"'f" where "c \ UNIV" by simp then obtain \::"('f,'v) subst" where "\v. \ v = Fun c []" by simp hence "subst_domain \ = UNIV" "ground (subst_range \)" by (simp_all add: subst_domain_def) thus ?thesis by auto qed lemma interpretation_subst_exists': "\\::('f,'v) subst. subst_domain \ = X \ ground (subst_range \)" proof - obtain \::"('f,'v) subst" where \: "subst_domain \ = UNIV" "ground (subst_range \)" using interpretation_subst_exists by moura let ?\ = "rm_vars (UNIV - X) \" have 1: "subst_domain ?\ = X" using \ by (auto simp add: subst_domain_def) hence 2: "ground (subst_range ?\)" using \ by force show ?thesis using 1 2 by blast qed lemma interpretation_subst_idem: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ subst_idem \" unfolding subst_idem_def using interpretation_grounds_all[of \] trm_subst_ident subst_eq_if_eq_vars by fastforce lemma subst_idem_comp_upd_eq: assumes "v \ subst_domain \" "subst_idem \" shows "\ \\<^sub>s \ = \(v := \ v) \\<^sub>s \" proof - from assms(1) have "(\ \\<^sub>s \) v = \ v" unfolding subst_compose_def by auto moreover have "\w. w \ v \ (\ \\<^sub>s \) w = (\(v := \ v) \\<^sub>s \) w" unfolding subst_compose_def by auto moreover have "(\(v := \ v) \\<^sub>s \) v = \ v" using assms(2) unfolding subst_idem_def subst_compose_def by (metis fun_upd_same) ultimately show ?thesis by (metis fun_upd_same fun_upd_triv subst_comp_upd1) qed lemma interpretation_dom_img_disjoint: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ subst_domain \ \ range_vars \ = {}" unfolding range_vars_alt_def by auto subsection \Basic Properties of MGUs\ lemma MGU_is_mgu_singleton: "MGU \ t u = is_mgu \ {(t,u)}" unfolding is_mgu_def unifiers_def by auto lemma Unifier_in_unifiers_singleton: "Unifier \ s t \ \ \ unifiers {(s,t)}" unfolding unifiers_def by auto lemma subst_list_singleton_fv_subset: "(\x \ set (subst_list (subst v t) E). fv (fst x) \ fv (snd x)) \ fv t \ (\x \ set E. fv (fst x) \ fv (snd x))" proof (induction E) case (Cons x E) let ?fvs = "\L. \x \ set L. fv (fst x) \ fv (snd x)" let ?fvx = "fv (fst x) \ fv (snd x)" let ?fvxsubst = "fv (fst x \ Var(v := t)) \ fv (snd x \ Var(v := t))" have "?fvs (subst_list (subst v t) (x#E)) = ?fvxsubst \ ?fvs (subst_list (subst v t) E)" unfolding subst_list_def subst_def by auto hence "?fvs (subst_list (subst v t) (x#E)) \ ?fvxsubst \ fv t \ ?fvs E" using Cons.IH by blast moreover have "?fvs (x#E) = ?fvx \ ?fvs E" by auto moreover have "?fvxsubst \ ?fvx \ fv t" using subst_fv_bound_singleton[of _ v t] by blast ultimately show ?case unfolding range_vars_alt_def by auto qed (simp add: subst_list_def) lemma subst_of_dom_subset: "subst_domain (subst_of L) \ set (map fst L)" proof (induction L rule: List.rev_induct) case (snoc x L) then obtain v t where x: "x = (v,t)" by (metis surj_pair) hence "subst_of (L@[x]) = Var(v := t) \\<^sub>s subst_of L" unfolding subst_of_def subst_def by (induct L) auto hence "subst_domain (subst_of (L@[x])) \ insert v (subst_domain (subst_of L))" using x subst_domain_compose[of "Var(v := t)" "subst_of L"] by (auto simp add: subst_domain_def) thus ?case using snoc.IH x by auto qed simp lemma wf_MGU_is_imgu_singleton: "wf\<^sub>M\<^sub>G\<^sub>U \ s t \ is_imgu \ {(s,t)}" proof - assume 1: "wf\<^sub>M\<^sub>G\<^sub>U \ s t" have 2: "subst_idem \" by (metis wf_subst_subst_idem 1 wf\<^sub>M\<^sub>G\<^sub>U_def) have 3: "\\' \ unifiers {(s,t)}. \ \\<^sub>\ \'" "\ \ unifiers {(s,t)}" by (metis 1 Unifier_in_unifiers_singleton wf\<^sub>M\<^sub>G\<^sub>U_def)+ have "\\ \ unifiers {(s,t)}. \ = \ \\<^sub>s \" by (metis 2 3 subst_idem_def subst_compose_assoc) thus "is_imgu \ {(s,t)}" by (metis is_imgu_def \\ \ unifiers {(s,t)}\) qed -lemma mgu_subst_range_vars: - assumes "mgu s t = Some \" shows "range_vars \ \ vars_term s \ vars_term t" -proof - - obtain xs where *: "Unification.unify [(s, t)] [] = Some xs" and [simp]: "subst_of xs = \" - using assms by (simp split: option.splits) - from unify_Some_UNIF [OF *] obtain ss - where "compose ss = \" and "UNIF ss {#(s, t)#} {#}" by auto - with UNIF_range_vars_subset [of ss "{#(s, t)#}" "{#}"] - show ?thesis by (metis vars_mset_singleton fst_conv snd_conv) -qed - -lemma mgu_subst_domain_range_vars_disjoint: - assumes "mgu s t = Some \" shows "subst_domain \ \ range_vars \ = {}" -proof - - have "is_imgu \ {(s, t)}" using assms mgu_sound by simp - hence "\ = \ \\<^sub>s \" unfolding is_imgu_def by blast - thus ?thesis by (metis subst_idemp_iff) -qed - -lemma mgu_same_empty: "mgu (t::('a,'b) term) t = Some Var" -proof - - { fix E::"('a,'b) equation list" and U::"('b \ ('a,'b) term) list" - assume "\(s,t) \ set E. s = t" - hence "Unification.unify E U = Some U" - proof (induction E U rule: Unification.unify.induct) - case (2 f S g T E U) - hence *: "f = g" "S = T" by auto - moreover have "\(s,t) \ set (zip T T). s = t" by (induct T) auto - hence "\(s,t) \ set (zip T T@E). s = t" using "2.prems"(1) by auto - moreover have "zip_option S T = Some (zip S T)" using \S = T\ by auto - hence **: "decompose (Fun f S) (Fun g T) = Some (zip S T)" - using \f = g\ unfolding decompose_def by auto - ultimately have "Unification.unify (zip S T@E) U = Some U" using "2.IH" * by auto - thus ?case using ** by auto - qed auto - } - hence "Unification.unify [(t,t)] [] = Some []" by auto - thus ?thesis by auto -qed +lemmas mgu_subst_range_vars = mgu_range_vars + +lemmas mgu_same_empty = mgu_same lemma mgu_var: assumes "x \ fv t" shows "mgu (Var x) t = Some (Var(x := t))" proof - have "unify [(Var x,t)] [] = Some [(x,t)]" using assms by (auto simp add: subst_list_def) moreover have "subst_of [(x,t)] = Var(x := t)" unfolding subst_of_def subst_def by simp - ultimately show ?thesis by simp + ultimately show ?thesis by (simp add: mgu_def) qed lemma mgu_gives_wellformed_subst: assumes "mgu s t = Some \" shows "wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using mgu_finite_subst_domain[OF assms] mgu_subst_domain_range_vars_disjoint[OF assms] unfolding wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by auto lemma mgu_gives_wellformed_MGU: assumes "mgu s t = Some \" shows "wf\<^sub>M\<^sub>G\<^sub>U \ s t" using mgu_subst_domain[OF assms] mgu_sound[OF assms] mgu_subst_range_vars [OF assms] MGU_is_mgu_singleton[of s \ t] is_imgu_imp_is_mgu[of \ "{(s,t)}"] mgu_gives_wellformed_subst[OF assms] unfolding wf\<^sub>M\<^sub>G\<^sub>U_def by blast lemma mgu_gives_subst_idem: "mgu s t = Some \ \ subst_idem \" using mgu_sound[of s t \] unfolding is_imgu_def subst_idem_def by auto lemma mgu_always_unifies: "Unifier \ M N \ \\. mgu M N = Some \" using mgu_complete Unifier_in_unifiers_singleton by blast lemma mgu_gives_MGU: "mgu s t = Some \ \ MGU \ s t" using mgu_sound[of s t \, THEN is_imgu_imp_is_mgu] MGU_is_mgu_singleton by metis lemma mgu_vars_bounded[dest?]: "mgu M N = Some \ \ subst_domain \ \ range_vars \ \ fv M \ fv N" using mgu_gives_wellformed_MGU unfolding wf\<^sub>M\<^sub>G\<^sub>U_def by blast lemma mgu_vars_bounded': assumes \: "mgu M N = Some \" and MN: "fv M = {} \ fv N = {}" shows "subst_domain \ = fv M \ fv N" (is ?A) and "range_vars \ = {}" (is ?B) proof - let ?C = "\t. subst_domain \ = fv t" have 0: "fv N = {} \ subst_domain \ \ fv M" "fv N = {} \ range_vars \ \ fv M" "fv M = {} \ subst_domain \ \ fv N" "fv M = {} \ range_vars \ \ fv N" using mgu_vars_bounded[OF \] by simp_all note 1 = mgu_gives_MGU[OF \] mgu_subst_domain_range_vars_disjoint[OF \] note 2 = subst_fv_imgI[of \] subst_dom_vars_in_subst[of _ \] note 3 = ground_term_subst_domain_fv_subset[of _ \] note 4 = subst_apply_fv_empty[of _ \] have "fv (\ x) = {}" when x: "x \ fv M" and N: "fv N = {}" for x using x N 0(1,2) 1 2[of x] 3[of M] 4[of N] by auto hence "?C M" ?B when N: "fv N = {}" using 0(1,2)[OF N] N by (fastforce, fastforce) moreover have "fv (\ x) = {}" when x: "x \ fv N" and M: "fv M = {}" for x using x M 0(3,4) 1 2[of x] 3[of N] 4[of M] by auto hence "?C N" ?B when M: "fv M = {}" using 0(3,4)[OF M] M by (fastforce, fastforce) ultimately show ?A ?B using MN by auto qed lemma mgu_eliminates[dest?]: assumes "mgu M N = Some \" shows "(\v \ fv M \ fv N. subst_elim \ v) \ \ = Var" (is "?P M N \") proof (cases "\ = Var") case False then obtain v where v: "v \ subst_domain \" by auto hence "v \ fv M \ fv N" using mgu_vars_bounded[OF assms] by blast thus ?thesis using wf_subst_elim_dom[OF mgu_gives_wellformed_subst[OF assms]] v by blast qed simp lemma mgu_eliminates_dom: assumes "mgu x y = Some \" "v \ subst_domain \" shows "subst_elim \ v" using mgu_gives_wellformed_subst[OF assms(1)] unfolding wf\<^sub>M\<^sub>G\<^sub>U_def wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def subst_elim_def by (metis disjoint_iff_not_equal subst_dom_elim assms(2)) lemma unify_list_distinct: assumes "Unification.unify E B = Some U" "distinct (map fst B)" and "(\x \ set E. fv (fst x) \ fv (snd x)) \ set (map fst B) = {}" shows "distinct (map fst U)" using assms proof (induction E B arbitrary: U rule: Unification.unify.induct) case 1 thus ?case by simp next case (2 f X g Y E B U) let ?fvs = "\L. \x \ set L. fv (fst x) \ fv (snd x)" from "2.prems"(1) obtain E' where *: "decompose (Fun f X) (Fun g Y) = Some E'" and [simp]: "f = g" "length X = length Y" "E' = zip X Y" and **: "Unification.unify (E'@E) B = Some U" by (auto split: option.splits) hence "\t t'. (t,t') \ set E' \ fv t \ fv (Fun f X) \ fv t' \ fv (Fun g Y)" by (metis zip_arg_subterm subtermeq_vars_subset) hence "?fvs E' \ fv (Fun f X) \ fv (Fun g Y)" by fastforce moreover have "fv (Fun f X) \ set (map fst B) = {}" "fv (Fun g Y) \ set (map fst B) = {}" using "2.prems"(3) by auto ultimately have "?fvs E' \ set (map fst B) = {}" by blast moreover have "?fvs E \ set (map fst B) = {}" using "2.prems"(3) by auto ultimately have "?fvs (E'@E) \ set (map fst B) = {}" by auto thus ?case using "2.IH"[OF * ** "2.prems"(2)] by metis next case (3 v t E B) let ?fvs = "\L. \x \ set L. fv (fst x) \ fv (snd x)" let ?E' = "subst_list (subst v t) E" from "3.prems"(3) have "v \ set (map fst B)" "fv t \ set (map fst B) = {}" by force+ hence *: "distinct (map fst ((v, t)#B))" using "3.prems"(2) by auto show ?case proof (cases "t = Var v") case True thus ?thesis using "3.prems" "3.IH"(1) by auto next case False hence "v \ fv t" using "3.prems"(1) by auto hence "Unification.unify (subst_list (subst v t) E) ((v, t)#B) = Some U" using \t \ Var v\ "3.prems"(1) by auto moreover have "?fvs ?E' \ set (map fst ((v, t)#B)) = {}" proof - have "v \ ?fvs ?E'" unfolding subst_list_def subst_def by (simp add: \v \ fv t\ subst_remove_var) moreover have "?fvs ?E' \ fv t \ ?fvs E" by (metis subst_list_singleton_fv_subset) hence "?fvs ?E' \ set (map fst B) = {}" using "3.prems"(3) by auto ultimately show ?thesis by auto qed ultimately show ?thesis using "3.IH"(2)[OF \t \ Var v\ \v \ fv t\ _ *] by metis qed next case (4 f X v E B U) let ?fvs = "\L. \x \ set L. fv (fst x) \ fv (snd x)" let ?E' = "subst_list (subst v (Fun f X)) E" have *: "?fvs E \ set (map fst B) = {}" using "4.prems"(3) by auto from "4.prems"(1) have "v \ fv (Fun f X)" by force from "4.prems"(3) have **: "v \ set (map fst B)" "fv (Fun f X) \ set (map fst B) = {}" by force+ hence ***: "distinct (map fst ((v, Fun f X)#B))" using "4.prems"(2) by auto from "4.prems"(3) have ****: "?fvs ?E' \ set (map fst ((v, Fun f X)#B)) = {}" proof - have "v \ ?fvs ?E'" unfolding subst_list_def subst_def using \v \ fv (Fun f X)\ subst_remove_var[of v "Fun f X"] by simp moreover have "?fvs ?E' \ fv (Fun f X) \ ?fvs E" by (metis subst_list_singleton_fv_subset) hence "?fvs ?E' \ set (map fst B) = {}" using * ** by blast ultimately show ?thesis by auto qed have "Unification.unify (subst_list (subst v (Fun f X)) E) ((v, Fun f X) # B) = Some U" using \v \ fv (Fun f X)\ "4.prems"(1) by auto thus ?case using "4.IH"[OF \v \ fv (Fun f X)\ _ *** ****] by metis qed lemma mgu_None_is_subst_neq: fixes s t::"('a,'b) term" and \::"('a,'b) subst" assumes "mgu s t = None" shows "s \ \ \ t \ \" using assms mgu_always_unifies by force lemma mgu_None_if_neq_ground: assumes "t \ t'" "fv t = {}" "fv t' = {}" shows "mgu t t' = None" proof (rule ccontr) assume "mgu t t' \ None" then obtain \ where \: "mgu t t' = Some \" by auto hence "t \ \ = t" "t' \ \ = t'" using assms subst_ground_ident by auto thus False using assms(1) MGU_is_Unifier[OF mgu_gives_MGU[OF \]] by auto qed lemma mgu_None_commutes: "mgu s t = None \ mgu t s = None" using mgu_complete[of s t] Unifier_in_unifiers_singleton[of s _ t] Unifier_sym[of t _ s] Unifier_in_unifiers_singleton[of t _ s] mgu_sound[of t s] unfolding is_imgu_def by fastforce lemma mgu_img_subterm_subst: fixes \::"('f,'v) subst" and s t u::"('f,'v) term" assumes "mgu s t = Some \" "u \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) - range Var" shows "u \ ((subterms s \ subterms t) - range Var) \\<^sub>s\<^sub>e\<^sub>t \" proof - define subterms_tuples::"('f,'v) equation list \ ('f,'v) terms" where subtt_def: "subterms_tuples \ \E. subterms\<^sub>s\<^sub>e\<^sub>t (fst ` set E) \ subterms\<^sub>s\<^sub>e\<^sub>t (snd ` set E)" define subterms_img::"('f,'v) subst \ ('f,'v) terms" where subti_def: "subterms_img \ \d. subterms\<^sub>s\<^sub>e\<^sub>t (subst_range d)" define d where "d \ \v t. subst v t::('f,'v) subst" define V where "V \ range Var::('f,'v) terms" define R where "R \ \d::('f,'v) subst. ((subterms s \ subterms t) - V) \\<^sub>s\<^sub>e\<^sub>t d" define M where "M \ \E d. subterms_tuples E \ subterms_img d" define Q where "Q \ (\E d. M E d - V \ R d - V)" define Q' where "Q' \ (\E d d'. (M E d - V) \\<^sub>s\<^sub>e\<^sub>t d' \ (R d - V) \\<^sub>s\<^sub>e\<^sub>t (d'::('f,'v) subst))" have Q_subst: "Q (subst_list (subst v t') E) (subst_of ((v, t')#B))" when v_fv: "v \ fv t'" and Q_assm: "Q ((Var v, t')#E) (subst_of B)" for v t' E B proof - define E' where "E' \ subst_list (subst v t') E" define B' where "B' \ subst_of ((v, t')#B)" have E': "E' = subst_list (d v t') E" and B': "B' = subst_of B \\<^sub>s d v t'" using subst_of_simps(3)[of "(v, t')"] unfolding subst_def E'_def B'_def d_def by simp_all have vt_img_subt: "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range (d v t')) = subterms t'" and vt_dom: "subst_domain (d v t') = {v}" using v_fv by (auto simp add: subst_domain_def d_def subst_def) have *: "subterms u1 \ subterms\<^sub>s\<^sub>e\<^sub>t (fst ` set E)" "subterms u2 \ subterms\<^sub>s\<^sub>e\<^sub>t (snd ` set E)" when "(u1,u2) \ set E" for u1 u2 using that by auto have **: "subterms\<^sub>s\<^sub>e\<^sub>t (d v t' ` (fv u \ subst_domain (d v t'))) \ subterms t'" for u::"('f,'v) term" using vt_dom unfolding d_def by force have 1: "subterms_tuples E' - V \ (subterms t' - V) \ (subterms_tuples E - V \\<^sub>s\<^sub>e\<^sub>t d v t')" (is "?A \ ?B") proof fix u assume "u \ ?A" then obtain u1 u2 where u12: "(u1,u2) \ set E" "u \ (subterms (u1 \ (d v t')) - V) \ (subterms (u2 \ (d v t')) - V)" unfolding subtt_def subst_list_def E'_def d_def by moura hence "u \ (subterms t' - V) \ (((subterms_tuples E) \\<^sub>s\<^sub>e\<^sub>t d v t') - V)" using subterms_subst[of u1 "d v t'"] subterms_subst[of u2 "d v t'"] *[OF u12(1)] **[of u1] **[of u2] unfolding subtt_def subst_list_def by auto moreover have "(subterms_tuples E \\<^sub>s\<^sub>e\<^sub>t d v t') - V \ (subterms_tuples E - V \\<^sub>s\<^sub>e\<^sub>t d v t') \ {t'}" unfolding subst_def subtt_def V_def d_def by force ultimately show "u \ ?B" using u12 v_fv by auto qed have 2: "subterms_img B' - V \ (subterms t' - V) \ (subterms_img (subst_of B) - V \\<^sub>s\<^sub>e\<^sub>t d v t')" using B' vt_img_subt subst_img_comp_subset'''[of "subst_of B" "d v t'"] unfolding subti_def subst_def V_def by argo have 3: "subterms_tuples ((Var v, t')#E) - V = (subterms t' - V) \ (subterms_tuples E - V)" by (auto simp add: subst_def subtt_def V_def) have "fv\<^sub>s\<^sub>e\<^sub>t (subterms t' - V) \ subst_domain (d v t') = {}" using v_fv vt_dom fv_subterms[of t'] by fastforce hence 4: "subterms t' - V \\<^sub>s\<^sub>e\<^sub>t d v t' = subterms t' - V" using set_subst_ident[of "subterms t' - range Var" "d v t'"] by (simp add: V_def) have "M E' B' - V \ M ((Var v, t')#E) (subst_of B) - V \\<^sub>s\<^sub>e\<^sub>t d v t'" using 1 2 3 4 unfolding M_def by blast moreover have "Q' ((Var v, t')#E) (subst_of B) (d v t')" using Q_assm unfolding Q_def Q'_def by auto moreover have "R (subst_of B) \\<^sub>s\<^sub>e\<^sub>t d v t' = R (subst_of ((v,t')#B))" unfolding R_def d_def by auto ultimately have "M (subst_list (d v t') E) (subst_of ((v, t')#B)) - V \ R (subst_of ((v, t')#B)) - V" unfolding Q'_def E'_def B'_def d_def by blast thus ?thesis unfolding Q_def M_def R_def d_def by blast qed have "u \ subterms s \ subterms t - V \\<^sub>s\<^sub>e\<^sub>t subst_of U" when assms': "unify E B = Some U" "u \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range (subst_of U)) - V" "Q E (subst_of B)" for E B U and T::"('f,'v) term list" using assms' proof (induction E B arbitrary: U rule: Unification.unify.induct) case (1 B) thus ?case by (auto simp add: Q_def M_def R_def subti_def) next case (2 g X h Y E B U) from "2.prems"(1) obtain E' where E': "decompose (Fun g X) (Fun h Y) = Some E'" "g = h" "length X = length Y" "E' = zip X Y" "Unification.unify (E'@E) B = Some U" by (auto split: option.splits) moreover have "subterms_tuples (E'@E) \ subterms_tuples ((Fun g X, Fun h Y)#E)" proof fix u assume "u \ subterms_tuples (E'@E)" then obtain u1 u2 where u12: "(u1,u2) \ set (E'@E)" "u \ subterms u1 \ subterms u2" unfolding subtt_def by fastforce thus "u \ subterms_tuples ((Fun g X, Fun h Y)#E)" proof (cases "(u1,u2) \ set E'") case True hence "subterms u1 \ subterms (Fun g X)" "subterms u2 \ subterms (Fun h Y)" using E'(4) subterms_subset params_subterms subsetCE by (metis set_zip_leftD, metis set_zip_rightD) thus ?thesis using u12 unfolding subtt_def by auto next case False thus ?thesis using u12 unfolding subtt_def by fastforce qed qed hence "Q (E'@E) (subst_of B)" using "2.prems"(3) unfolding Q_def M_def by blast ultimately show ?case using "2.IH"[of E' U] "2.prems" by meson next case (3 v t' E B) show ?case proof (cases "t' = Var v") case True thus ?thesis using "3.prems" "3.IH"(1) unfolding Q_def M_def V_def subtt_def by auto next case False hence 1: "v \ fv t'" using "3.prems"(1) by auto hence "unify (subst_list (subst v t') E) ((v, t')#B) = Some U" using False "3.prems"(1) by auto thus ?thesis using Q_subst[OF 1 "3.prems"(3)] "3.IH"(2)[OF False 1 _ "3.prems"(2)] by metis qed next case (4 g X v E B U) have 1: "v \ fv (Fun g X)" using "4.prems"(1) not_None_eq by fastforce hence 2: "unify (subst_list (subst v (Fun g X)) E) ((v, Fun g X)#B) = Some U" using "4.prems"(1) by auto have 3: "Q ((Var v, Fun g X)#E) (subst_of B)" using "4.prems"(3) unfolding Q_def M_def subtt_def by auto show ?case using Q_subst[OF 1 3] "4.IH"[OF 1 2 "4.prems"(2)] by metis qed moreover obtain D where "unify [(s, t)] [] = Some D" "\ = subst_of D" - using assms(1) by (auto split: option.splits) + using assms(1) by (auto simp: mgu_def split: option.splits) moreover have "Q [(s,t)] (subst_of [])" unfolding Q_def M_def R_def subtt_def subti_def by force ultimately show ?thesis using assms(2) unfolding V_def by auto qed lemma mgu_img_consts: fixes \::"('f,'v) subst" and s t::"('f,'v) term" and c::'f and z::'v assumes "mgu s t = Some \" "Fun c [] \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \)" shows "Fun c [] \ subterms s \ subterms t" proof - obtain u where "u \ (subterms s \ subterms t) - range Var" "u \ \ = Fun c []" using mgu_img_subterm_subst[OF assms(1), of "Fun c []"] assms(2) by force thus ?thesis by (cases u) auto qed lemma mgu_img_consts': fixes \::"('f,'v) subst" and s t::"('f,'v) term" and c::'f and z::'v assumes "mgu s t = Some \" "\ z = Fun c []" shows "Fun c [] \ s \ Fun c [] \ t" using mgu_img_consts[OF assms(1)] assms(2) by (metis Un_iff in_subterms_Union subst_imgI term.distinct(1)) lemma mgu_img_composed_var_term: fixes \::"('f,'v) subst" and s t::"('f,'v) term" and f::'f and Z::"'v list" assumes "mgu s t = Some \" "Fun f (map Var Z) \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \)" shows "\Z'. map \ Z' = map Var Z \ Fun f (map Var Z') \ subterms s \ subterms t" proof - obtain u where u: "u \ (subterms s \ subterms t) - range Var" "u \ \ = Fun f (map Var Z)" using mgu_img_subterm_subst[OF assms(1), of "Fun f (map Var Z)"] assms(2) by fastforce then obtain T where T: "u = Fun f T" "map (\t. t \ \) T = map Var Z" by (cases u) auto have "\t \ set T. \x. t = Var x" using T(2) by (induct T arbitrary: Z) auto then obtain Z' where Z': "map Var Z' = T" by (metis ex_map_conv) hence "map \ Z' = map Var Z" using T(2) by (induct Z' arbitrary: T Z) auto thus ?thesis using u(1) T(1) Z' by auto qed lemma mgu_ground_instance_case: assumes t: "fv (t \ \) = {}" shows "mgu t (t \ \) = Some (rm_vars (UNIV - fv t) \)" (is ?A) and mgu_ground_commutes: "mgu t (t \ \) = mgu (t \ \) t" (is ?B) proof - define \ where "\ \ rm_vars (UNIV - fv t) \" have \: "t \ \ = t \ \" using rm_vars_subst_eq'[of t \] unfolding \_def by metis have 0: "Unifier \ t (t \ \)" using subst_ground_ident[OF t, of \] term_subst_eq[of t \ \] unfolding \_def by (metis Diff_iff) obtain \ where \: "mgu t (t \ \) = Some \" "MGU \ t (t \ \)" using mgu_always_unifies[OF 0] mgu_gives_MGU[of t "t \ \"] unfolding \ by blast have 1: "subst_domain \ = fv t" using t MGU_is_Unifier[OF \(2)] subset_antisym[OF mgu_subst_domain[OF \(1)]] ground_term_subst_domain_fv_subset[of t \] subst_apply_fv_empty[OF t, of \] unfolding \ by auto have 2: "subst_domain \ = fv t" using 0 rm_vars_dom[of "UNIV - fv t" \] ground_term_subst_domain_fv_subset[of t \] subst_apply_fv_empty[OF t, of \] unfolding \_def by auto have "\ x = \ x" for x using 1 2 MGU_is_Unifier[OF \(2)] term_subst_eq_conv[of t \ \] subst_ground_ident[OF t[unfolded \], of \] subst_domI[of _ x] by metis hence "\ = \" by presburger thus A: ?A using \(1) unfolding \ \_def by blast have "Unifier \ (t \ \) t" using 0 by simp then obtain \' where \': "mgu (t \ \) t = Some \'" "MGU \' (t \ \) t" using mgu_always_unifies mgu_gives_MGU[of "t \ \" t] unfolding \ by fastforce have 3: "subst_domain \' = fv t" using t MGU_is_Unifier[OF \'(2)] subset_antisym[OF mgu_subst_domain[OF \'(1)]] ground_term_subst_domain_fv_subset[of t \'] subst_apply_fv_empty[OF t, of \'] unfolding \ by auto have "\' x = \ x" for x using 2 3 MGU_is_Unifier[OF \'(2)] term_subst_eq_conv[of t \' \] subst_ground_ident[OF t[unfolded \], of \'] subst_domI[of _ x] by metis hence "\' = \" by presburger thus ?B using A \'(1) unfolding \ \_def by argo qed subsection \Lemmata: The "Inequality Lemmata"\ text \Subterm injectivity (a stronger injectivity property)\ definition subterm_inj_on where "subterm_inj_on f A \ \x\A. \y\A. (\v. v \ f x \ v \ f y) \ x = y" lemma subterm_inj_on_imp_inj_on: "subterm_inj_on f A \ inj_on f A" unfolding subterm_inj_on_def inj_on_def by fastforce lemma subst_inj_on_is_bij_betw: "inj_on \ (subst_domain \) = bij_betw \ (subst_domain \) (subst_range \)" unfolding inj_on_def bij_betw_def by auto lemma subterm_inj_on_alt_def: "subterm_inj_on f A \ (inj_on f A \ (\s \ f`A. \u \ f`A. (\v. v \ s \ v \ u) \ s = u))" (is "?A \ ?B") unfolding subterm_inj_on_def inj_on_def by fastforce lemma subterm_inj_on_alt_def': "subterm_inj_on \ (subst_domain \) \ (inj_on \ (subst_domain \) \ (\s \ subst_range \. \u \ subst_range \. (\v. v \ s \ v \ u) \ s = u))" (is "?A \ ?B") by (metis subterm_inj_on_alt_def subst_range.simps) lemma subterm_inj_on_subset: assumes "subterm_inj_on f A" and "B \ A" shows "subterm_inj_on f B" proof - have "inj_on f A" "\s\f ` A. \u\f ` A. (\v. v \ s \ v \ u) \ s = u" using subterm_inj_on_alt_def[of f A] assms(1) by auto moreover have "f ` B \ f ` A" using assms(2) by auto ultimately have "inj_on f B" "\s\f ` B. \u\f ` B. (\v. v \ s \ v \ u) \ s = u" using inj_on_subset[of f A] assms(2) by blast+ thus ?thesis by (metis subterm_inj_on_alt_def) qed lemma inj_subst_unif_consts: fixes \ \ \::"('f,'v) subst" and s t::"('f,'v) term" assumes \: "subterm_inj_on \ (subst_domain \)" "\x \ (fv s \ fv t) - X. \c. \ x = Fun c []" "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ (subterms s \ subterms t) = {}" "ground (subst_range \)" "subst_domain \ \ X = {}" and \: "ground (subst_range \)" "subst_domain \ = subst_domain \" and unif: "Unifier \ (s \ \) (t \ \)" shows "\\. Unifier \ (s \ \) (t \ \)" proof - let ?xs = "subst_domain \" let ?ys = "(fv s \ fv t) - ?xs" have "\\::('f,'v) subst. s \ \ = t \ \" by (metis subst_subst_compose unif) then obtain \::"('f,'v) subst" where \: "mgu s t = Some \" using mgu_always_unifies by moura have 1: "\\::('f,'v) subst. s \ \ \ \ = t \ \ \ \" by (metis unif) have 2: "\\::('f,'v) subst. s \ \ \ \ = t \ \ \ \ \ \ \\<^sub>\ \ \\<^sub>s \" using mgu_gives_MGU[OF \] by simp have 3: "\(z::'v) (c::'f). \ z = Fun c [] \ Fun c [] \ s \ Fun c [] \ t" by (rule mgu_img_consts'[OF \]) have 4: "subst_domain \ \ range_vars \ = {}" by (metis mgu_gives_wellformed_subst[OF \] wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) have 5: "subst_domain \ \ range_vars \ \ fv s \ fv t" by (metis mgu_gives_wellformed_MGU[OF \] wf\<^sub>M\<^sub>G\<^sub>U_def) { fix x and \::"('f,'v) subst" assume "x \ subst_domain \" hence "(\ \\<^sub>s \) x = \ x" using \(4) ident_comp_subst_trm_if_disj[of \ \] unfolding range_vars_alt_def by fast } then obtain \::"('f,'v) subst" where \: "\x \ subst_domain \. \ x = (\ \\<^sub>s \) x" using 1 2 by moura have *: "\x. x \ subst_domain \ \ subst_domain \ \ \y \ ?ys. \ x = Var y" proof - fix x assume "x \ subst_domain \ \ ?xs" hence x: "x \ subst_domain \" "x \ subst_domain \" by auto then obtain c where c: "\ x = Fun c []" using \(2,5) 5 by moura hence *: "(\ \\<^sub>s \) x = Fun c []" using \ x by fastforce hence **: "x \ subst_domain (\ \\<^sub>s \)" "Fun c [] \ subst_range (\ \\<^sub>s \)" by (auto simp add: subst_domain_def) have "\ x = Fun c [] \ (\z. \ x = Var z \ \ z = Fun c [])" by (rule subst_img_comp_subset_const'[OF *]) moreover have "\ x \ Fun c []" proof (rule ccontr) assume "\\ x \ Fun c []" hence "Fun c [] \ s \ Fun c [] \ t" using 3 by metis moreover have "\u \ subst_range \. u \ subterms s \ subterms t" using \(3) by force hence "Fun c [] \ subterms s \ subterms t" by (metis c \ground (subst_range \)\x(2) ground_subst_dom_iff_img) ultimately show False by auto qed moreover have "\x' \ subst_domain \. \ x \ Var x'" proof (rule ccontr) assume "\(\x' \ subst_domain \. \ x \ Var x')" then obtain x' where x': "x' \ subst_domain \" "\ x = Var x'" by moura hence "\ x' = Fun c []" "(\ \\<^sub>s \) x = Fun c []" using * unfolding subst_compose_def by auto moreover have "x \ x'" using x(1) x'(2) 4 by (auto simp add: subst_domain_def) moreover have "x' \ subst_domain \" using x'(2) mgu_eliminates_dom[OF \] by (metis (no_types) subst_elim_def subst_apply_term.simps(1) vars_iff_subterm_or_eq) moreover have "(\ \\<^sub>s \) x = \ x" "(\ \\<^sub>s \) x' = \ x'" using \ x(2) x'(1) by auto ultimately show False using subterm_inj_on_imp_inj_on[OF \(1)] * by (simp add: inj_on_def subst_compose_def x'(2) subst_domain_def) qed ultimately show "\y \ ?ys. \ x = Var y" by (metis 5 x(2) subtermeqI' vars_iff_subtermeq DiffI Un_iff subst_fv_imgI sup.orderE) qed have **: "inj_on \ (subst_domain \ \ ?xs)" proof (intro inj_onI) fix x y assume *: "x \ subst_domain \ \ subst_domain \" "y \ subst_domain \ \ subst_domain \" "\ x = \ y" hence "(\ \\<^sub>s \) x = (\ \\<^sub>s \) y" unfolding subst_compose_def by auto hence "\ x = \ y" using \ * by auto thus "x = y" using inj_onD[OF subterm_inj_on_imp_inj_on[OF \(1)]] *(1,2) by simp qed define \ where "\ = (\y'. if Var y' \ \ ` (subst_domain \ \ ?xs) then Var ((inv_into (subst_domain \ \ ?xs) \) (Var y')) else Var y'::('f,'v) term)" have a1: "Unifier (\ \\<^sub>s \) s t" using mgu_gives_MGU[OF \] by auto define \' where "\' = \ \\<^sub>s \" have d1: "subst_domain \' \ ?ys" proof fix z assume z: "z \ subst_domain \'" have "z \ ?xs \ z \ subst_domain \'" proof (cases "z \ subst_domain \") case True moreover assume "z \ ?xs" ultimately have z_in: "z \ subst_domain \ \ ?xs" by simp then obtain y where y: "\ z = Var y" "y \ ?ys" using * by moura hence "\ y = Var ((inv_into (subst_domain \ \ ?xs) \) (Var y))" using \_def z_in by simp hence "\ y = Var z" by (metis y(1) z_in ** inv_into_f_eq) hence "\' z = Var z" using \'_def y(1) subst_compose_def[of \ \] by simp thus ?thesis by (simp add: subst_domain_def) next case False hence "\ z = Var z" by (simp add: subst_domain_def) moreover assume "z \ ?xs" hence "\ z = Var z" using \_def * by force ultimately show ?thesis using \'_def subst_compose_def[of \ \] by (simp add: subst_domain_def) qed moreover have "subst_domain \ \ range_vars \" unfolding \'_def \_def range_vars_alt_def by (auto simp add: subst_domain_def) hence "subst_domain \' \ subst_domain \ \ range_vars \" using subst_domain_compose[of \ \] unfolding \'_def by blast ultimately show "z \ ?ys" using 5 z by auto qed have d2: "Unifier (\' \\<^sub>s \) s t" using a1 \'_def by auto have d3: "\ \\<^sub>s \' \\<^sub>s \ = \' \\<^sub>s \" proof - { fix z::'v assume z: "z \ ?xs" then obtain u where u: "\ z = u" "fv u = {}" using \ by auto hence "(\ \\<^sub>s \' \\<^sub>s \) z = u" by (simp add: subst_compose subst_ground_ident) moreover have "z \ subst_domain \'" using d1 z by auto hence "\' z = Var z" by (simp add: subst_domain_def) hence "(\' \\<^sub>s \) z = u" using u(1) by (simp add: subst_compose) ultimately have "(\ \\<^sub>s \' \\<^sub>s \) z = (\' \\<^sub>s \) z" by metis } moreover { fix z::'v assume "z \ ?ys" hence "z \ subst_domain \" using \(2) by auto hence "(\ \\<^sub>s \' \\<^sub>s \) z = (\' \\<^sub>s \) z" by (simp add: subst_compose subst_domain_def) } moreover { fix z::'v assume "z \ ?xs" "z \ ?ys" hence "\ z = Var z" "\' z = Var z" using \(2) d1 by blast+ hence "(\ \\<^sub>s \' \\<^sub>s \) z = (\' \\<^sub>s \) z" by (simp add: subst_compose) } ultimately show ?thesis by auto qed from d2 d3 have "Unifier (\' \\<^sub>s \) (s \ \) (t \ \)" by (metis subst_subst_compose) thus ?thesis by metis qed lemma inj_subst_unif_comp_terms: fixes \ \ \::"('f,'v) subst" and s t::"('f,'v) term" assumes \: "subterm_inj_on \ (subst_domain \)" "ground (subst_range \)" "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ (subterms s \ subterms t) = {}" "(fv s \ fv t) - subst_domain \ \ X" and tfr: "\f U. Fun f U \ subterms s \ subterms t \ U = [] \ (\u \ set U. u \ Var ` X)" and \: "ground (subst_range \)" "subst_domain \ = subst_domain \" and unif: "Unifier \ (s \ \) (t \ \)" shows "\\. Unifier \ (s \ \) (t \ \)" proof - let ?xs = "subst_domain \" let ?ys = "(fv s \ fv t) - ?xs" have "ground (subst_range \)" using \(2) by auto have "\\::('f,'v) subst. s \ \ = t \ \" by (metis subst_subst_compose unif) then obtain \::"('f,'v) subst" where \: "mgu s t = Some \" using mgu_always_unifies by moura have 1: "\\::('f,'v) subst. s \ \ \ \ = t \ \ \ \" by (metis unif) have 2: "\\::('f,'v) subst. s \ \ \ \ = t \ \ \ \ \ \ \\<^sub>\ \ \\<^sub>s \" using mgu_gives_MGU[OF \] by simp have 3: "\(z::'v) (c::'f). Fun c [] \ \ z \ Fun c [] \ s \ Fun c [] \ t" using mgu_img_consts[OF \] by force have 4: "subst_domain \ \ range_vars \ = {}" using mgu_gives_wellformed_subst[OF \] by (metis wf\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) have 5: "subst_domain \ \ range_vars \ \ fv s \ fv t" using mgu_gives_wellformed_MGU[OF \] by (metis wf\<^sub>M\<^sub>G\<^sub>U_def) { fix x and \::"('f,'v) subst" assume "x \ subst_domain \" hence "(\ \\<^sub>s \) x = \ x" using \ground (subst_range \)\ ident_comp_subst_trm_if_disj[of \ \ x] unfolding range_vars_alt_def by blast } then obtain \::"('f,'v) subst" where \: "\x \ subst_domain \. \ x = (\ \\<^sub>s \) x" using 1 2 by moura have ***: "\x. x \ subst_domain \ \ subst_domain \ \ fv (\ x) \ ?ys" proof - fix x assume "x \ subst_domain \ \ ?xs" hence x: "x \ subst_domain \" "x \ subst_domain \" by auto moreover have "\(\x' \ ?xs. x' \ fv (\ x))" proof (rule ccontr) assume "\\(\x' \ ?xs. x' \ fv (\ x))" then obtain x' where x': "x' \ fv (\ x)" "x' \ ?xs" by metis have "x \ x'" "x' \ subst_domain \" "\ x' = Var x'" using 4 x(1) x'(1) unfolding range_vars_alt_def by auto hence "(\ \\<^sub>s \) x' \ (\ \\<^sub>s \) x" "\ x' = (\ \\<^sub>s \) x'" using \ x(2) x'(2) by (metis subst_compose subst_mono vars_iff_subtermeq x'(1), metis subst_apply_term.simps(1) subst_compose_def) hence "\ x' \ \ x" using \ x(2) x'(2) by auto thus False using \(1) x'(2) x(2) \x \ x'\ unfolding subterm_inj_on_def by (meson subtermeqI') qed ultimately show "fv (\ x) \ ?ys" using 5 subst_dom_vars_in_subst[of x \] subst_fv_imgI[of \ x] by blast qed have **: "inj_on \ (subst_domain \ \ ?xs)" proof (intro inj_onI) fix x y assume *: "x \ subst_domain \ \ subst_domain \" "y \ subst_domain \ \ subst_domain \" "\ x = \ y" hence "(\ \\<^sub>s \) x = (\ \\<^sub>s \) y" unfolding subst_compose_def by auto hence "\ x = \ y" using \ * by auto thus "x = y" using inj_onD[OF subterm_inj_on_imp_inj_on[OF \(1)]] *(1,2) by simp qed have *: "\x. x \ subst_domain \ \ subst_domain \ \ \y \ ?ys. \ x = Var y" proof (rule ccontr) fix xi assume xi_assms: "xi \ subst_domain \ \ subst_domain \" "\(\y \ ?ys. \ xi = Var y)" hence xi_\: "xi \ subst_domain \" and \_xi_comp: "\(\y. \ xi = Var y)" using ***[of xi] 5 by auto then obtain f T where f: "\ xi = Fun f T" by (cases "\ xi") moura have "\g Y'. Y' \ [] \ Fun g (map Var Y') \ \ xi \ set Y' \ ?ys" proof - have "\c. Fun c [] \ \ xi \ Fun c [] \ \ xi" using \ xi_\ by (metis const_subterm_subst subst_compose) hence 1: "\c. \(Fun c [] \ \ xi)" using 3[of _ xi] xi_\ \(3) by auto have "\(\x. \ xi = Var x)" using f by auto hence "\g S. Fun g S \ \ xi \ (\s \ set S. (\c. s = Fun c []) \ (\x. s = Var x))" using nonvar_term_has_composed_shallow_term[of "\ xi"] by auto then obtain g S where gS: "Fun g S \ \ xi" "\s \ set S. (\c. s = Fun c []) \ (\x. s = Var x)" by moura have "\s \ set S. \x. s = Var x" using 1 term.order_trans gS by (metis (no_types, lifting) UN_I term.order_refl subsetCE subterms.simps(2) sup_ge2) then obtain S' where 2: "map Var S' = S" by (metis ex_map_conv) have "S \ []" using 1 term.order_trans[OF _ gS(1)] by fastforce hence 3: "S' \ []" "Fun g (map Var S') \ \ xi" using gS(1) 2 by auto have "set S' \ fv (Fun g (map Var S'))" by simp hence 4: "set S' \ fv (\ xi)" using 3(2) fv_subterms by force show ?thesis using ***[OF xi_assms(1)] 2 3 4 by auto qed then obtain g Y' where g: "Y' \ []" "Fun g (map Var Y') \ \ xi" "set Y' \ ?ys" by moura then obtain X where X: "map \ X = map Var Y'" "Fun g (map Var X) \ subterms s \ subterms t" using mgu_img_composed_var_term[OF \, of g Y'] by force hence "\(u::('f,'v) term) \ set (map Var X). u \ Var ` ?ys" using \(4) tfr g(1) by fastforce then obtain j where j: "j < length X" "X ! j \ ?ys" by (metis image_iff[of _ Var "fv s \ fv t - subst_domain \"] nth_map[of _ X Var] in_set_conv_nth[of _ "map Var X"] length_map[of Var X]) define yj' where yj': "yj' \ Y' ! j" define xj where xj: "xj \ X ! j" have "xj \ fv s \ fv t" using j X(1) g(3) 5 xj yj' by (metis length_map nth_map term.simps(1) in_set_conv_nth le_supE subsetCE subst_domI) hence xj_\: "xj \ subst_domain \" using j unfolding xj by simp have len: "length X = length Y'" by (rule map_eq_imp_length_eq[OF X(1)]) have "Var yj' \ \ xi" using term.order_trans[OF _ g(2)] j(1) len unfolding yj' by auto hence "\ yj' \ \ xi" using \ xi_\ by (metis subst_apply_term.simps(1) subst_compose_def subst_mono) moreover have \_xj_var: "Var yj' = \ xj" using X(1) len j(1) nth_map unfolding xj yj' by metis hence "\ yj' = \ xj" using \ xj_\ by (metis subst_apply_term.simps(1) subst_compose_def) moreover have "xi \ xj" using \_xi_comp \_xj_var by auto ultimately show False using \(1) xi_\ xj_\ unfolding subterm_inj_on_def by blast qed define \ where "\ = (\y'. if Var y' \ \ ` (subst_domain \ \ ?xs) then Var ((inv_into (subst_domain \ \ ?xs) \) (Var y')) else Var y'::('f,'v) term)" have a1: "Unifier (\ \\<^sub>s \) s t" using mgu_gives_MGU[OF \] by auto define \' where "\' = \ \\<^sub>s \" have d1: "subst_domain \' \ ?ys" proof fix z assume z: "z \ subst_domain \'" have "z \ ?xs \ z \ subst_domain \'" proof (cases "z \ subst_domain \") case True moreover assume "z \ ?xs" ultimately have z_in: "z \ subst_domain \ \ ?xs" by simp then obtain y where y: "\ z = Var y" "y \ ?ys" using * by moura hence "\ y = Var ((inv_into (subst_domain \ \ ?xs) \) (Var y))" using \_def z_in by simp hence "\ y = Var z" by (metis y(1) z_in ** inv_into_f_eq) hence "\' z = Var z" using \'_def y(1) subst_compose_def[of \ \] by simp thus ?thesis by (simp add: subst_domain_def) next case False hence "\ z = Var z" by (simp add: subst_domain_def) moreover assume "z \ ?xs" hence "\ z = Var z" using \_def * by force ultimately show ?thesis using \'_def subst_compose_def[of \ \] by (simp add: subst_domain_def) qed moreover have "subst_domain \ \ range_vars \" unfolding \'_def \_def range_vars_alt_def subst_domain_def by auto hence "subst_domain \' \ subst_domain \ \ range_vars \" using subst_domain_compose[of \ \] unfolding \'_def by blast ultimately show "z \ ?ys" using 5 z by blast qed have d2: "Unifier (\' \\<^sub>s \) s t" using a1 \'_def by auto have d3: "\ \\<^sub>s \' \\<^sub>s \ = \' \\<^sub>s \" proof - { fix z::'v assume z: "z \ ?xs" then obtain u where u: "\ z = u" "fv u = {}" using \ by auto hence "(\ \\<^sub>s \' \\<^sub>s \) z = u" by (simp add: subst_compose subst_ground_ident) moreover have "z \ subst_domain \'" using d1 z by auto hence "\' z = Var z" by (simp add: subst_domain_def) hence "(\' \\<^sub>s \) z = u" using u(1) by (simp add: subst_compose) ultimately have "(\ \\<^sub>s \' \\<^sub>s \) z = (\' \\<^sub>s \) z" by metis } moreover { fix z::'v assume "z \ ?ys" hence "z \ subst_domain \" using \(2) by auto hence "(\ \\<^sub>s \' \\<^sub>s \) z = (\' \\<^sub>s \) z" by (simp add: subst_compose subst_domain_def) } moreover { fix z::'v assume "z \ ?xs" "z \ ?ys" hence "\ z = Var z" "\' z = Var z" using \(2) d1 by blast+ hence "(\ \\<^sub>s \' \\<^sub>s \) z = (\' \\<^sub>s \) z" by (simp add: subst_compose) } ultimately show ?thesis by auto qed from d2 d3 have "Unifier (\' \\<^sub>s \) (s \ \) (t \ \)" by (metis subst_subst_compose) thus ?thesis by metis qed context begin private lemma sat_ineq_subterm_inj_subst_aux: fixes \::"('f,'v) subst" assumes "Unifier \ (s \ \) (t \ \)" "ground (subst_range \)" "(fv s \ fv t) - X \ subst_domain \" "subst_domain \ \ X = {}" shows "\\::('f,'v) subst. subst_domain \ = X \ ground (subst_range \) \ s \ \ \ \ = t \ \ \ \" proof - have "\\. Unifier \ (s \ \) (t \ \) \ interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" proof - obtain \'::"('f,'v) subst" where *: "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \'" using interpretation_subst_exists by metis hence "Unifier (\ \\<^sub>s \') (s \ \) (t \ \)" using assms(1) by simp thus ?thesis using * interpretation_comp by blast qed then obtain \' where \': "Unifier \' (s \ \) (t \ \)" "interpretation\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \'" by moura define \'' where "\'' = rm_vars (UNIV - X) \'" have *: "fv (s \ \) \ X" "fv (t \ \) \ X" using assms(2,3) subst_fv_unfold_ground_img[of \] unfolding range_vars_alt_def by (simp_all add: Diff_subset_conv Un_commute) hence **: "subst_domain \'' = X" "ground (subst_range \'')" using rm_vars_img_subset[of "UNIV - X" \'] rm_vars_dom[of "UNIV - X" \'] \'(2) unfolding \''_def by auto hence "\t. t \ \ \ \'' = t \ \'' \ \" using subst_eq_if_disjoint_vars_ground[OF _ _ assms(2)] assms(4) by blast moreover have "Unifier \'' (s \ \) (t \ \)" using Unifier_dom_restrict[OF \'(1)] \''_def * by blast ultimately show ?thesis using ** by auto qed text \ The "inequality lemma": This lemma gives sufficient syntactic conditions for finding substitutions \\\ under which terms \s\ and \t\ are not unifiable. This is useful later when establishing the typing results since we there want to find well-typed solutions to inequality constraints / "negative checks" constraints, and this lemma gives conditions for protocols under which such constraints are well-typed satisfiable if satisfiable. \ lemma sat_ineq_subterm_inj_subst: fixes \ \ \::"('f,'v) subst" assumes \: "subterm_inj_on \ (subst_domain \)" "ground (subst_range \)" "subst_domain \ \ X = {}" "subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ (subterms s \ subterms t) = {}" "(fv s \ fv t) - subst_domain \ \ X" and tfr: "(\x \ (fv s \ fv t) - X. \c. \ x = Fun c []) \ (\f U. Fun f U \ subterms s \ subterms t \ U = [] \ (\u \ set U. u \ Var ` X))" and \: "\\::('f,'v) subst. subst_domain \ = X \ ground (subst_range \) \ s \ \ \ \ \ t \ \ \ \" "(fv s \ fv t) - X \ subst_domain \" "subst_domain \ \ X = {}" "ground (subst_range \)" "subst_domain \ = subst_domain \" and \: "subst_domain \ = X" "ground (subst_range \)" shows "s \ \ \ \ \ t \ \ \ \" proof - have "\\. \Unifier \ (s \ \) (t \ \)" by (metis \(1) sat_ineq_subterm_inj_subst_aux[OF _ \(4,2,3)]) hence "\Unifier \ (s \ \) (t \ \)" using inj_subst_unif_consts[OF \(1) _ \(4,2,3) \(4,5)] inj_subst_unif_comp_terms[OF \(1,2,4,5) _ \(4,5)] tfr by metis moreover have "subst_domain \ \ subst_domain \ = {}" using \(2,3) \(1) by auto ultimately show ?thesis using \ subst_eq_if_disjoint_vars_ground[OF _ \(2) \(2)] by metis qed end lemma ineq_subterm_inj_cond_subst: assumes "X \ range_vars \ = {}" and "\f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t S \ T = [] \ (\u \ set T. u \ Var`X)" shows "\f T. Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (S \\<^sub>s\<^sub>e\<^sub>t \) \ T = [] \ (\u \ set T. u \ Var`X)" proof (intro allI impI) let ?M = "\S. subterms\<^sub>s\<^sub>e\<^sub>t S \\<^sub>s\<^sub>e\<^sub>t \" let ?N = "\S. subterms\<^sub>s\<^sub>e\<^sub>t (\ ` (fv\<^sub>s\<^sub>e\<^sub>t S \ subst_domain \))" fix f T assume "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (S \\<^sub>s\<^sub>e\<^sub>t \)" hence 1: "Fun f T \ ?M S \ Fun f T \ ?N S" using subterms_subst[of _ \] by auto have 2: "Fun f T \ subterms\<^sub>s\<^sub>e\<^sub>t (subst_range \) \ \u \ set T. u \ Var`X" using fv_subset_subterms[of "Fun f T" "subst_range \"] assms(1) unfolding range_vars_alt_def by force have 3: "\x \ subst_domain \. \ x \ Var`X" proof fix x assume "x \ subst_domain \" hence "fv (\ x) \ range_vars \" using subst_dom_vars_in_subst subst_fv_imgI unfolding range_vars_alt_def by auto thus "\ x \ Var`X" using assms(1) by auto qed show "T = [] \ (\s \ set T. s \ Var`X)" using 1 proof assume "Fun f T \ ?M S" then obtain u where u: "u \ subterms\<^sub>s\<^sub>e\<^sub>t S" "u \ \ = Fun f T" by fastforce show ?thesis proof (cases u) case (Var x) hence "Fun f T \ subst_range \" using u(2) by (simp add: subst_domain_def) hence "\u \ set T. u \ Var`X" using 2 by force thus ?thesis by auto next case (Fun g S) hence "S = [] \ (\u \ set S. u \ Var`X)" using assms(2) u(1) by metis thus ?thesis proof assume "S = []" thus ?thesis using u(2) Fun by simp next assume "\u \ set S. u \ Var`X" then obtain u' where u': "u' \ set S" "u' \ Var`X" by moura hence "u' \ \ \ set T" using u(2) Fun by auto thus ?thesis using u'(2) 3 by (cases u') force+ qed qed next assume "Fun f T \ ?N S" thus ?thesis using 2 by force qed qed subsection \Lemmata: Sufficient Conditions for Term Matching\ definition subst_var_inv::"('a,'b) subst \ 'b set \ ('a,'b) subst" where "subst_var_inv \ X \ (\x. if Var x \ \ ` X then Var ((inv_into X \) (Var x)) else Var x)" lemma subst_var_inv_subst_domain: assumes "x \ subst_domain (subst_var_inv \ X)" shows "Var x \ \ ` X" by (meson assms subst_dom_vars_in_subst subst_var_inv_def) lemma subst_var_inv_subst_domain': assumes "X \ subst_domain \" shows "x \ subst_domain (subst_var_inv \ X) \ Var x \ \ ` X" proof show "Var x \ \ ` X \ x \ subst_domain (subst_var_inv \ X)" by (metis (no_types, lifting) assms f_inv_into_f in_mono inv_into_into subst_domI subst_dom_vars_in_subst subst_var_inv_def term.inject(1)) qed (rule subst_var_inv_subst_domain) lemma subst_var_inv_Var_range: "subst_range (subst_var_inv \ X) \ range Var" unfolding subst_var_inv_def by auto text \Injective substitutions from variables to variables are invertible\ lemma inj_var_ran_subst_is_invertible: assumes \_inj_on_X: "inj_on \ X" and \_var_on_X: "\ ` X \ range Var" and fv_t: "fv t \ X" shows "t = t \ \ \\<^sub>s subst_var_inv \ X" proof - have "\ x \ subst_var_inv \ X = Var x" when x: "x \ X" for x proof - obtain y where y: "\ x = Var y" using x \_var_on_X fv_t by auto hence "Var y \ \ ` X" using x by simp thus ?thesis using y inv_into_f_eq[OF \_inj_on_X x y] unfolding subst_var_inv_def by simp qed thus ?thesis using fv_t by (simp add: subst_compose_def trm_subst_ident'' subset_eq) qed lemma inj_var_ran_subst_is_invertible': assumes \_inj_on_t: "inj_on \ (fv t)" and \_var_on_t: "\ ` fv t \ range Var" shows "t = t \ \ \\<^sub>s subst_var_inv \ (fv t)" using assms inj_var_ran_subst_is_invertible by fast text \Sufficient conditions for matching unifiable terms\ lemma inj_var_ran_unifiable_has_subst_match: assumes "t \ \ = s \ \" "inj_on \ (fv t)" "\ ` fv t \ range Var" shows "t = s \ \ \\<^sub>s subst_var_inv \ (fv t)" using assms inj_var_ran_subst_is_invertible by fastforce end diff --git a/thys/Stateful_Protocol_Composition_and_Typing/Typed_Model.thy b/thys/Stateful_Protocol_Composition_and_Typing/Typed_Model.thy --- a/thys/Stateful_Protocol_Composition_and_Typing/Typed_Model.thy +++ b/thys/Stateful_Protocol_Composition_and_Typing/Typed_Model.thy @@ -1,2203 +1,2203 @@ (* Title: Typed_Model.thy Author: Andreas Viktor Hess, DTU SPDX-License-Identifier: BSD-3-Clause *) section \The Typed Model\ theory Typed_Model imports Lazy_Intruder begin text \Term types\ type_synonym ('f,'v) term_type = "('f,'v) term" text \Constructors for term types\ abbreviation (input) TAtom::"'v \ ('f,'v) term_type" where "TAtom a \ Var a" abbreviation (input) TComp::"['f, ('f,'v) term_type list] \ ('f,'v) term_type" where "TComp f ts \ Fun f ts" text \ The typed model extends the intruder model with a typing function \\\ that assigns types to terms. \ locale typed_model = intruder_model arity public Ana for arity::"'fun \ nat" and public::"'fun \ bool" and Ana::"('fun,'var) term \ (('fun,'var) term list \ ('fun,'var) term list)" + fixes \::"('fun,'var) term \ ('fun,'atom::finite) term_type" assumes const_type: "\c. arity c = 0 \ \a. \ts. \ (Fun c ts) = TAtom a" and fun_type: "\f ts. arity f > 0 \ \ (Fun f ts) = TComp f (map \ ts)" and \_wf: "\x f ts. TComp f ts \ \ (Var x) \ arity f > 0" "\x. wf\<^sub>t\<^sub>r\<^sub>m (\ (Var x))" begin subsection \Definitions\ text \The set of atomic types\ abbreviation "\\<^sub>a \ UNIV::('atom set)" text \Well-typed substitutions\ definition wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t where "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ (\v. \ (Var v) = \ (\ v))" text \The set of sub-message patterns (SMP)\ inductive_set SMP::"('fun,'var) terms \ ('fun,'var) terms" for M where MP[intro]: "t \ M \ t \ SMP M" | Subterm[intro]: "\t \ SMP M; t' \ t\ \ t' \ SMP M" | Substitution[intro]: "\t \ SMP M; wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \; wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)\ \ (t \ \) \ SMP M" | Ana[intro]: "\t \ SMP M; Ana t = (K,T); k \ set K\ \ k \ SMP M" text \ Type-flaw resistance for sets: Unifiable sub-message patterns must have the same type (unless they are variables) \ definition tfr\<^sub>s\<^sub>e\<^sub>t where "tfr\<^sub>s\<^sub>e\<^sub>t M \ (\s \ SMP M - (Var`\). \t \ SMP M - (Var`\). (\\. Unifier \ s t) \ \ s = \ t)" text \ Type-flaw resistance for strand steps: - The terms in a satisfiable equality step must have the same types - Inequality steps must satisfy the conditions of the "inequality lemma"\ fun tfr\<^sub>s\<^sub>t\<^sub>p where "tfr\<^sub>s\<^sub>t\<^sub>p (Equality a t t') = ((\\. Unifier \ t t') \ \ t = \ t')" | "tfr\<^sub>s\<^sub>t\<^sub>p (Inequality X F) = ( (\x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X. \a. \ (Var x) = TAtom a) \ (\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)))" | "tfr\<^sub>s\<^sub>t\<^sub>p _ = True" text \ Type-flaw resistance for strands: - The set of terms in strands must be type-flaw resistant - The steps of strands must be type-flaw resistant \ definition tfr\<^sub>s\<^sub>t where "tfr\<^sub>s\<^sub>t S \ tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S) \ list_all tfr\<^sub>s\<^sub>t\<^sub>p S" subsection \Small Lemmata\ lemma tfr\<^sub>s\<^sub>t\<^sub>p_list_all_alt_def: "list_all tfr\<^sub>s\<^sub>t\<^sub>p S \ ((\a t t'. Equality a t t' \ set S \ (\\. Unifier \ t t') \ \ t = \ t') \ (\X F. Inequality X F \ set S \ (\x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X. \a. \ (Var x) = TAtom a) \ (\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))))" (is "?P S \ ?Q S") proof show "?P S \ ?Q S" proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp show "?Q S \ ?P S" proof (induction S) case (Cons x S) thus ?case by (cases x) auto qed simp qed lemma \_wf'': "TComp f T \ \ t \ arity f > 0" proof (induction t) case (Var x) thus ?case using \_wf(1)[of f T x] by blast next case (Fun g S) thus ?case using fun_type[of g S] const_type[of g] by (cases "arity g") auto qed lemma \_wf': "wf\<^sub>t\<^sub>r\<^sub>m t \ wf\<^sub>t\<^sub>r\<^sub>m (\ t)" proof (induction t) case (Fun f T) hence *: "arity f = length T" "\t. t \ set T \ wf\<^sub>t\<^sub>r\<^sub>m (\ t)" unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto { assume "arity f = 0" hence ?case using const_type[of f] by auto } moreover { assume "arity f > 0" hence ?case using fun_type[of f] * by force } ultimately show ?case by auto qed (metis \_wf(2)) lemma fun_type_inv: assumes "\ t = TComp f T" shows "arity f > 0" using \_wf''(1)[of f T t] assms by simp_all lemma fun_type_inv_wf: assumes "\ t = TComp f T" "wf\<^sub>t\<^sub>r\<^sub>m t" shows "arity f = length T" using \_wf'[OF assms(2)] assms(1) unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto lemma const_type_inv: "\ (Fun c X) = TAtom a \ arity c = 0" by (rule ccontr, simp add: fun_type) lemma const_type_inv_wf: assumes "\ (Fun c X) = TAtom a" and "wf\<^sub>t\<^sub>r\<^sub>m (Fun c X)" shows "X = []" by (metis assms const_type_inv length_0_conv subtermeqI' wf\<^sub>t\<^sub>r\<^sub>m_def) lemma const_type': "\c \ \. \a \ \\<^sub>a. \X. \ (Fun c X) = TAtom a" using const_type by simp lemma fun_type': "\f \ \\<^sub>f. \X. \ (Fun f X) = TComp f (map \ X)" using fun_type by simp lemma fun_type_id_eq: "\ (Fun f X) = TComp g Y \ f = g" by (metis const_type fun_type neq0_conv "term.inject"(2) "term.simps"(4)) lemma fun_type_length_eq: "\ (Fun f X) = TComp g Y \ length X = length Y" by (metis fun_type fun_type_id_eq fun_type_inv(1) length_map term.inject(2)) lemma pgwt_type_map: assumes "public_ground_wf_term t" shows "\ t = TAtom a \ \f. t = Fun f []" "\ t = TComp g Y \ \X. t = Fun g X \ map \ X = Y" proof - let ?A = "\ t = TAtom a \ (\f. t = Fun f [])" let ?B = "\ t = TComp g Y \ (\X. t = Fun g X \ map \ X = Y)" have "?A \ ?B" proof (cases "\ t") case (Var a) obtain f X where "t = Fun f X" "arity f = length X" using pgwt_fun[OF assms(1)] pgwt_arity[OF assms(1)] by fastforce+ thus ?thesis using const_type_inv \\ t = TAtom a\ by auto next case (Fun g Y) obtain f X where *: "t = Fun f X" using pgwt_fun[OF assms(1)] by force hence "f = g" "map \ X = Y" using fun_type_id_eq \\ t = TComp g Y\ fun_type[OF fun_type_inv(1)[OF \\ t = TComp g Y\]] by fastforce+ thus ?thesis using *(1) \\ t = TComp g Y\ by auto qed thus "\ t = TAtom a \ \f. t = Fun f []" "\ t = TComp g Y \ \X. t = Fun g X \ map \ X = Y" by auto qed lemma wt_subst_Var[simp]: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t Var" by (metis wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) lemma wt_subst_trm: "(\v. v \ fv t \ \ (Var v) = \ (\ v)) \ \ t = \ (t \ \)" proof (induction t) case (Fun f X) hence *: "\x. x \ set X \ \ x = \ (x \ \)" by auto show ?case proof (cases "f \ \\<^sub>f") case True hence "\X. \ (Fun f X) = TComp f (map \ X)" using fun_type' by auto thus ?thesis using * by auto next case False hence "\a \ \\<^sub>a. \X. \ (Fun f X) = TAtom a" using const_type' by auto thus ?thesis by auto qed qed auto lemma wt_subst_trm': "\wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \; \ s = \ t\ \ \ (s \ \) = \ (t \ \)" by (metis wt_subst_trm wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) lemma wt_subst_trm'': "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ \ t = \ (t \ \)" by (metis wt_subst_trm wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) lemma wt_subst_compose: assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" proof - have "\v. \ (\ v) = \ (\ v \ \)" using wt_subst_trm \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by metis moreover have "\v. \ (Var v) = \ (\ v)" using \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \\ unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by metis ultimately have "\v. \ (Var v) = \ (\ v \ \)" by metis thus ?thesis unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def subst_compose_def by metis qed lemma wt_subst_TAtom_Var_cases: assumes \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and x: "\ (Var x) = TAtom a" shows "(\y. \ x = Var y) \ (\c. \ x = Fun c [])" proof (cases "(\y. \ x = Var y)") case False then obtain c T where c: "\ x = Fun c T" by (cases "\ x") simp_all hence "wf\<^sub>t\<^sub>r\<^sub>m (Fun c T)" using \(2) by fastforce hence "T = []" using const_type_inv_wf[of c T a] x c wt_subst_trm''[OF \(1), of "Var x"] by fastforce thus ?thesis using c by blast qed simp lemma wt_subst_TAtom_fv: assumes \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "\x. wf\<^sub>t\<^sub>r\<^sub>m (\ x)" and "\x \ fv t - X. \a. \ (Var x) = TAtom a" shows "\x \ fv (t \ \) - fv\<^sub>s\<^sub>e\<^sub>t (\ ` X). \a. \ (Var x) = TAtom a" using assms(3) proof (induction t) case (Var x) thus ?case proof (cases "x \ X") case False with Var obtain a where "\ (Var x) = TAtom a" by moura hence *: "\ (\ x) = TAtom a" "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" using \ unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by auto show ?thesis proof (cases "\ x") case (Var y) thus ?thesis using * by auto next case (Fun f T) hence "T = []" using * const_type_inv[of f T a] unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto thus ?thesis using Fun by auto qed qed auto qed fastforce lemma wt_subst_TAtom_subterms_subst: assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "\x \ fv t. \a. \ (Var x) = TAtom a" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\ ` fv t)" shows "subterms (t \ \) = subterms t \\<^sub>s\<^sub>e\<^sub>t \" using assms(2,3) proof (induction t) case (Var x) obtain a where a: "\ (Var x) = TAtom a" using Var.prems(1) by moura hence "\ (\ x) = TAtom a" using wt_subst_trm''[OF assms(1), of "Var x"] by simp hence "(\y. \ x = Var y) \ (\c. \ x = Fun c [])" using const_type_inv_wf Var.prems(2) by (cases "\ x") auto thus ?case by auto next case (Fun f T) have "subterms (t \ \) = subterms t \\<^sub>s\<^sub>e\<^sub>t \" when "t \ set T" for t using that Fun.prems(1,2) Fun.IH[OF that] by auto thus ?case by auto qed lemma wt_subst_TAtom_subterms_set_subst: assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "\x \ fv\<^sub>s\<^sub>e\<^sub>t M. \a. \ (Var x) = TAtom a" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\ ` fv\<^sub>s\<^sub>e\<^sub>t M)" shows "subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) = subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" proof show "subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \) \ subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" proof fix t assume "t \ subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" then obtain s where s: "s \ M" "t \ subterms (s \ \)" by auto thus "t \ subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" using assms(2,3) wt_subst_TAtom_subterms_subst[OF assms(1), of s] by auto qed show "subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \ \ subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" proof fix t assume "t \ subterms\<^sub>s\<^sub>e\<^sub>t M \\<^sub>s\<^sub>e\<^sub>t \" then obtain s where s: "s \ M" "t \ subterms s \\<^sub>s\<^sub>e\<^sub>t \" by auto thus "t \ subterms\<^sub>s\<^sub>e\<^sub>t (M \\<^sub>s\<^sub>e\<^sub>t \)" using assms(2,3) wt_subst_TAtom_subterms_subst[OF assms(1), of s] by auto qed qed lemma wt_subst_subst_upd: assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "\ (Var x) = \ t" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\(x := t))" using assms unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by (metis fun_upd_other fun_upd_same) lemma wt_subst_const_fv_type_eq: assumes "\x \ fv t. \a. \ (Var x) = TAtom a" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "\x \ fv (t \ \). \y \ fv t. \ (Var x) = \ (Var y)" using assms(1) proof (induction t) case (Var x) then obtain a where a: "\ (Var x) = TAtom a" by moura show ?case proof (cases "\ x") case (Fun f T) hence "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" "\ (Fun f T) = TAtom a" using a wt_subst_trm''[OF \(1), of "Var x"] \(2) by fastforce+ thus ?thesis using const_type_inv_wf Fun by fastforce qed (use a wt_subst_trm''[OF \(1), of "Var x"] in simp) qed fastforce lemma TComp_term_cases: assumes "wf\<^sub>t\<^sub>r\<^sub>m t" "\ t = TComp f T" shows "(\v. t = Var v) \ (\T'. t = Fun f T' \ T = map \ T' \ T' \ [])" proof (cases "\v. t = Var v") case False then obtain T' where T': "t = Fun f T'" "T = map \ T'" using assms fun_type[OF fun_type_inv(1)[OF assms(2)]] fun_type_id_eq by (cases t) force+ thus ?thesis using assms fun_type_inv(1) fun_type_inv_wf by fastforce qed metis lemma TAtom_term_cases: assumes "wf\<^sub>t\<^sub>r\<^sub>m t" "\ t = TAtom \" shows "(\v. t = Var v) \ (\f. t = Fun f [])" using assms const_type_inv unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by (cases t) auto lemma subtermeq_imp_subtermtypeeq: assumes "wf\<^sub>t\<^sub>r\<^sub>m t" "s \ t" shows "\ s \ \ t" using assms(2,1) proof (induction t) case (Fun f T) thus ?case proof (cases "s = Fun f T") case False then obtain x where x: "x \ set T" "s \ x" using Fun.prems(1) by auto hence "wf\<^sub>t\<^sub>r\<^sub>m x" using wf_trm_subtermeq[OF Fun.prems(2)] Fun_param_is_subterm[of _ T f] by auto hence "\ s \ \ x" using Fun.IH[OF x] by simp moreover have "arity f > 0" using x fun_type_inv_wf Fun.prems by (metis length_pos_if_in_set term.order_refl wf\<^sub>t\<^sub>r\<^sub>m_def) ultimately show ?thesis using x Fun.prems fun_type[of f T] by auto qed simp qed simp lemma subterm_funs_term_in_type: assumes "wf\<^sub>t\<^sub>r\<^sub>m t" "Fun f T \ t" "\ (Fun f T) = TComp f (map \ T)" shows "f \ funs_term (\ t)" using assms(2,1,3) proof (induction t) case (Fun f' T') hence [simp]: "wf\<^sub>t\<^sub>r\<^sub>m (Fun f T)" by (metis wf_trm_subtermeq) { fix a assume \: "\ (Fun f' T') = TAtom a" hence "Fun f T = Fun f' T'" using Fun TAtom_term_cases subtermeq_Var_const by metis hence False using Fun.prems(3) \ by simp } moreover { fix g S assume \: "\ (Fun f' T') = TComp g S" hence "g = f'" "S = map \ T'" using Fun.prems(2) fun_type_id_eq[OF \] fun_type[OF fun_type_inv(1)[OF \]] by auto hence \': "\ (Fun f' T') = TComp f' (map \ T')" using \ by auto hence "g \ funs_term (\ (Fun f' T'))" using \ by auto moreover { assume "Fun f T \ Fun f' T'" then obtain x where "x \ set T'" "Fun f T \ x" using Fun.prems(1) by auto hence "f \ funs_term (\ x)" using Fun.IH[OF _ _ _ Fun.prems(3), of x] wf_trm_subtermeq[OF \wf\<^sub>t\<^sub>r\<^sub>m (Fun f' T')\, of x] by force moreover have "\ x \ set (map \ T')" using \' \x \ set T'\ by auto ultimately have "f \ funs_term (\ (Fun f' T'))" using \' by auto } ultimately have ?case by (cases "Fun f T = Fun f' T'") (auto simp add: \g = f'\) } ultimately show ?case by (cases "\ (Fun f' T')") auto qed simp lemma wt_subst_fv_termtype_subterm: assumes "x \ fv (\ y)" and "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "wf\<^sub>t\<^sub>r\<^sub>m (\ y)" shows "\ (Var x) \ \ (Var y)" using subtermeq_imp_subtermtypeeq[OF assms(3) var_is_subterm[OF assms(1)]] wt_subst_trm''[OF assms(2), of "Var y"] by auto lemma wt_subst_fv\<^sub>s\<^sub>e\<^sub>t_termtype_subterm: assumes "x \ fv\<^sub>s\<^sub>e\<^sub>t (\ ` Y)" and "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "\y \ Y. \ (Var x) \ \ (Var y)" using wt_subst_fv_termtype_subterm[OF _ assms(2), of x] assms(1,3) by fastforce lemma funs_term_type_iff: assumes t: "wf\<^sub>t\<^sub>r\<^sub>m t" and f: "arity f > 0" shows "f \ funs_term (\ t) \ (f \ funs_term t \ (\x \ fv t. f \ funs_term (\ (Var x))))" (is "?P t \ ?Q t") using t proof (induction t) case (Fun g T) hence IH: "?P s \ ?Q s" when "s \ set T" for s using that wf_trm_subterm[OF _ Fun_param_is_subterm] by blast have 0: "arity g = length T" using Fun.prems unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto show ?case proof (cases "f = g") case True thus ?thesis using fun_type[OF f] by simp next case False have "?P (Fun g T) \ (\s \ set T. ?P s)" proof assume *: "?P (Fun g T)" hence "\ (Fun g T) = TComp g (map \ T)" using const_type[of g] fun_type[of g] by force thus "\s \ set T. ?P s" using False * by force next assume *: "\s \ set T. ?P s" hence "\ (Fun g T) = TComp g (map \ T)" using 0 const_type[of g] fun_type[of g] by force thus "?P (Fun g T)" using False * by force qed thus ?thesis using False f IH by auto qed qed simp lemma funs_term_type_iff': assumes M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and f: "arity f > 0" shows "f \ \(funs_term ` \ ` M) \ (f \ \(funs_term ` M) \ (\x \ fv\<^sub>s\<^sub>e\<^sub>t M. f \ funs_term (\ (Var x))))" (is "?A \ ?B") proof assume ?A then obtain t where "t \ M" "wf\<^sub>t\<^sub>r\<^sub>m t" "f \ funs_term (\ t)" using M by moura thus ?B using funs_term_type_iff[OF _ f, of t] by auto next assume ?B then obtain t where "t \ M" "wf\<^sub>t\<^sub>r\<^sub>m t" "f \ funs_term t \ (\x \ fv t. f \ funs_term (\ (Var x)))" using M by auto thus ?A using funs_term_type_iff[OF _ f, of t] by blast qed lemma Ana_subterm_type: assumes "Ana t = (K,M)" and "wf\<^sub>t\<^sub>r\<^sub>m t" and "m \ set M" shows "\ m \ \ t" proof - have "m \ t" using Ana_subterm[OF assms(1)] assms(3) by auto thus ?thesis using subtermeq_imp_subtermtypeeq[OF assms(2)] by simp qed lemma wf_trm_TAtom_subterms: assumes "wf\<^sub>t\<^sub>r\<^sub>m t" "\ t = TAtom \" shows "subterms t = {t}" using assms const_type_inv unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by (cases t) force+ lemma wf_trm_TComp_subterm: assumes "wf\<^sub>t\<^sub>r\<^sub>m s" "t \ s" obtains f T where "\ s = TComp f T" proof (cases s) case (Var x) thus ?thesis using \t \ s\ by simp next case (Fun g S) hence "length S > 0" using assms Fun_subterm_inside_params[of t g S] by auto hence "arity g > 0" by (metis \wf\<^sub>t\<^sub>r\<^sub>m s\ \s = Fun g S\ term.order_refl wf\<^sub>t\<^sub>r\<^sub>m_def) thus ?thesis using fun_type \s = Fun g S\ that by auto qed lemma SMP_empty[simp]: "SMP {} = {}" proof (rule ccontr) assume "SMP {} \ {}" then obtain t where "t \ SMP {}" by auto thus False by (induct t rule: SMP.induct) auto qed lemma SMP_I: assumes "s \ M" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "t \ s \ \" "\v. wf\<^sub>t\<^sub>r\<^sub>m (\ v)" shows "t \ SMP M" using SMP.Substitution[OF SMP.MP[OF assms(1)] assms(2)] SMP.Subterm[of "s \ \" M t] assms(3,4) by (cases "t = s \ \") simp_all lemma SMP_wf_trm: assumes "t \ SMP M" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" shows "wf\<^sub>t\<^sub>r\<^sub>m t" using assms(1) by (induct t rule: SMP.induct) (use assms(2) in blast, use wf_trm_subtermeq in blast, use wf_trm_subst in blast, use Ana_keys_wf' in blast) lemma SMP_ikI[intro]: "t \ ik\<^sub>s\<^sub>t S \ t \ SMP (trms\<^sub>s\<^sub>t S)" by force lemma MP_setI[intro]: "x \ set S \ trms\<^sub>s\<^sub>t\<^sub>p x \ trms\<^sub>s\<^sub>t S" by force lemma SMP_setI[intro]: "x \ set S \ trms\<^sub>s\<^sub>t\<^sub>p x \ SMP (trms\<^sub>s\<^sub>t S)" by force lemma SMP_subset_I: assumes M: "\t \ M. \s \. s \ N \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = s \ \" shows "SMP M \ SMP N" proof fix t show "t \ SMP M \ t \ SMP N" proof (induction t rule: SMP.induct) case (MP t) then obtain s \ where s: "s \ N" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t = s \ \" using M by moura show ?case using SMP_I[OF s(1,2), of "s \ \"] s(3,4) wf_trm_subst_range_iff by fast qed (auto intro!: SMP.Substitution[of _ N]) qed lemma SMP_union: "SMP (A \ B) = SMP A \ SMP B" proof show "SMP (A \ B) \ SMP A \ SMP B" proof fix t assume "t \ SMP (A \ B)" thus "t \ SMP A \ SMP B" by (induct rule: SMP.induct) blast+ qed { fix t assume "t \ SMP A" hence "t \ SMP (A \ B)" by (induct rule: SMP.induct) blast+ } moreover { fix t assume "t \ SMP B" hence "t \ SMP (A \ B)" by (induct rule: SMP.induct) blast+ } ultimately show "SMP A \ SMP B \ SMP (A \ B)" by blast qed lemma SMP_append[simp]: "SMP (trms\<^sub>s\<^sub>t (S@S')) = SMP (trms\<^sub>s\<^sub>t S) \ SMP (trms\<^sub>s\<^sub>t S')" (is "?A = ?B") using SMP_union by simp lemma SMP_mono: "A \ B \ SMP A \ SMP B" proof - assume "A \ B" then obtain C where "B = A \ C" by moura thus "SMP A \ SMP B" by (simp add: SMP_union) qed lemma SMP_Union: "SMP (\m \ M. f m) = (\m \ M. SMP (f m))" proof show "SMP (\m\M. f m) \ (\m\M. SMP (f m))" proof fix t assume "t \ SMP (\m\M. f m)" thus "t \ (\m\M. SMP (f m))" by (induct t rule: SMP.induct) force+ qed show "(\m\M. SMP (f m)) \ SMP (\m\M. f m)" proof fix t assume "t \ (\m\M. SMP (f m))" then obtain m where "m \ M" "t \ SMP (f m)" by moura thus "t \ SMP (\m\M. f m)" using SMP_mono[of "f m" "\m\M. f m"] by auto qed qed lemma SMP_singleton_ex: "t \ SMP M \ (\m \ M. t \ SMP {m})" "m \ M \ t \ SMP {m} \ t \ SMP M" using SMP_Union[of "\t. {t}" M] by auto lemma SMP_Cons: "SMP (trms\<^sub>s\<^sub>t (x#S)) = SMP (trms\<^sub>s\<^sub>t [x]) \ SMP (trms\<^sub>s\<^sub>t S)" using SMP_append[of "[x]" S] by auto lemma SMP_Nil[simp]: "SMP (trms\<^sub>s\<^sub>t []) = {}" proof - { fix t assume "t \ SMP (trms\<^sub>s\<^sub>t [])" hence False by induct auto } thus ?thesis by blast qed lemma SMP_subset_union_eq: assumes "M \ SMP N" shows "SMP N = SMP (M \ N)" proof - { fix t assume "t \ SMP (M \ N)" hence "t \ SMP N" using assms by (induction rule: SMP.induct) blast+ } thus ?thesis using SMP_union by auto qed lemma SMP_subterms_subset: "subterms\<^sub>s\<^sub>e\<^sub>t M \ SMP M" proof fix t assume "t \ subterms\<^sub>s\<^sub>e\<^sub>t M" then obtain m where "m \ M" "t \ m" by auto thus "t \ SMP M" using SMP_I[of _ _ Var] by auto qed lemma SMP_SMP_subset: "N \ SMP M \ SMP N \ SMP M" by (metis SMP_mono SMP_subset_union_eq Un_commute Un_upper2) lemma wt_subst_rm_vars: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (rm_vars X \)" using rm_vars_dom unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by auto lemma wt_subst_SMP_subset: assumes "trms\<^sub>s\<^sub>t S \ SMP S'" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \) \ SMP S'" proof fix t assume *: "t \ trms\<^sub>s\<^sub>t (S \\<^sub>s\<^sub>t \)" show "t \ SMP S'" using trm_strand_subst_cong(2)[OF *] proof assume "\t'. t = t' \ \ \ t' \ trms\<^sub>s\<^sub>t S" thus "t \ SMP S'" using assms SMP.Substitution by auto next assume "\X F. Inequality X F \ set S \ (\t'\trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F. t = t' \ rm_vars (set X) \)" then obtain X F t' where **: "Inequality X F \ set S" "t'\trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F" "t = t' \ rm_vars (set X) \" by force then obtain s where s: "s \ trms\<^sub>s\<^sub>t\<^sub>p (Inequality X F)" "t = s \ rm_vars (set X) \" by moura hence "s \ SMP (trms\<^sub>s\<^sub>t S)" using **(1) by force hence "t \ SMP (trms\<^sub>s\<^sub>t S)" using SMP.Substitution[OF _ wt_subst_rm_vars[OF assms(2)] wf_trms_subst_rm_vars'[OF assms(3)]] unfolding s(2) by blast thus "t \ SMP S'" by (metis SMP_union SMP_subset_union_eq UnCI assms(1)) qed qed lemma MP_subset_SMP: "\(trms\<^sub>s\<^sub>t\<^sub>p ` set S) \ SMP (trms\<^sub>s\<^sub>t S)" "trms\<^sub>s\<^sub>t S \ SMP (trms\<^sub>s\<^sub>t S)" "M \ SMP M" by auto lemma SMP_fun_map_snd_subset: "SMP (trms\<^sub>s\<^sub>t (map Send1 X)) \ SMP (trms\<^sub>s\<^sub>t [Send1 (Fun f X)])" proof fix t assume "t \ SMP (trms\<^sub>s\<^sub>t (map Send1 X))" thus "t \ SMP (trms\<^sub>s\<^sub>t [Send1 (Fun f X)])" proof (induction t rule: SMP.induct) case (MP t) hence "t \ set X" by auto hence "t \ Fun f X" by (metis subtermI') thus ?case using SMP.Subterm[of "Fun f X" "trms\<^sub>s\<^sub>t [Send1 (Fun f X)]" t] using SMP.MP by auto qed blast+ qed lemma SMP_wt_subst_subset: assumes "t \ SMP (M \\<^sub>s\<^sub>e\<^sub>t \)" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "t \ SMP M" using assms wf_trm_subst_range_iff[of \] by (induct t rule: SMP.induct) blast+ lemma SMP_wt_instances_subset: assumes "\t \ M. \s \ N. \\. t = s \ \ \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and "t \ SMP M" shows "t \ SMP N" proof - obtain m where m: "m \ M" "t \ SMP {m}" using SMP_singleton_ex(1)[OF assms(2)] by blast then obtain n \ where n: "n \ N" "m = n \ \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using assms(1) by fast have "t \ SMP (N \\<^sub>s\<^sub>e\<^sub>t \)" using n(1,2) SMP_singleton_ex(2)[of m "N \\<^sub>s\<^sub>e\<^sub>t \", OF _ m(2)] by fast thus ?thesis using SMP_wt_subst_subset[OF _ n(3,4)] by blast qed lemma SMP_consts: assumes "\t \ M. \c. t = Fun c []" and "\t \ M. Ana t = ([], [])" shows "SMP M = M" proof show "SMP M \ M" proof fix t show "t \ SMP M \ t \ M" apply (induction t rule: SMP.induct) by (use assms in auto) qed qed auto lemma SMP_subterms_eq: "SMP (subterms\<^sub>s\<^sub>e\<^sub>t M) = SMP M" proof show "SMP M \ SMP (subterms\<^sub>s\<^sub>e\<^sub>t M)" using SMP_mono[of M "subterms\<^sub>s\<^sub>e\<^sub>t M"] by blast show "SMP (subterms\<^sub>s\<^sub>e\<^sub>t M) \ SMP M" proof fix t show "t \ SMP (subterms\<^sub>s\<^sub>e\<^sub>t M) \ t \ SMP M" by (induction t rule: SMP.induct) blast+ qed qed lemma SMP_funs_term: assumes t: "t \ SMP M" "f \ funs_term t \ (\x \ fv t. f \ funs_term (\ (Var x)))" and f: "arity f > 0" and M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and Ana_f: "\s K T. Ana s = (K,T) \ f \ \(funs_term ` set K) \ f \ funs_term s" shows "f \ \(funs_term ` M) \ (\x \ fv\<^sub>s\<^sub>e\<^sub>t M. f \ funs_term (\ (Var x)))" using t proof (induction t rule: SMP.induct) case (Subterm t t') thus ?case by (metis UN_I vars_iff_subtermeq funs_term_subterms_eq(1) term.order_trans) next case (Substitution t \) show ?case using M SMP_wf_trm[OF Substitution.hyps(1)] wf_trm_subst[of \ t, OF Substitution.hyps(3)] funs_term_type_iff[OF _ f] wt_subst_trm''[OF Substitution.hyps(2), of t] Substitution.prems Substitution.IH by metis next case (Ana t K T t') thus ?case using Ana_f[OF Ana.hyps(2)] Ana_keys_fv[OF Ana.hyps(2)] by fastforce qed auto lemma id_type_eq: assumes "\ (Fun f X) = \ (Fun g Y)" shows "f \ \ \ g \ \" "f \ \\<^sub>f \ g \ \\<^sub>f" using assms const_type' fun_type' id_union_univ(1) by (metis UNIV_I UnE "term.distinct"(1))+ lemma fun_type_arg_cong: assumes "f \ \\<^sub>f" "g \ \\<^sub>f" "\ (Fun f (x#X)) = \ (Fun g (y#Y))" shows "\ x = \ y" "\ (Fun f X) = \ (Fun g Y)" using assms fun_type' by auto lemma fun_type_arg_cong': assumes "f \ \\<^sub>f" "g \ \\<^sub>f" "\ (Fun f (X@x#X')) = \ (Fun g (Y@y#Y'))" "length X = length Y" shows "\ x = \ y" using assms proof (induction X arbitrary: Y) case Nil thus ?case using fun_type_arg_cong(1)[of f g x X' y Y'] by auto next case (Cons x' X Y'') then obtain y' Y where "Y'' = y'#Y" by (metis length_Suc_conv) hence "\ (Fun f (X@x#X')) = \ (Fun g (Y@y#Y'))" "length X = length Y" using Cons.prems(3,4) fun_type_arg_cong(2)[OF Cons.prems(1,2), of x' "X@x#X'"] by auto thus ?thesis using Cons.IH[OF Cons.prems(1,2)] by auto qed lemma fun_type_param_idx: "\ (Fun f T) = Fun g S \ i < length T \ \ (T ! i) = S ! i" by (metis fun_type fun_type_id_eq fun_type_inv(1) nth_map term.inject(2)) lemma fun_type_param_ex: assumes "\ (Fun f T) = Fun g (map \ S)" "t \ set S" shows "\s \ set T. \ s = \ t" using fun_type_length_eq[OF assms(1)] length_map[of \ S] assms(2) fun_type_param_idx[OF assms(1)] nth_map in_set_conv_nth by metis lemma tfr_stp_all_split: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (x#S) \ list_all tfr\<^sub>s\<^sub>t\<^sub>p [x]" "list_all tfr\<^sub>s\<^sub>t\<^sub>p (x#S) \ list_all tfr\<^sub>s\<^sub>t\<^sub>p S" "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S') \ list_all tfr\<^sub>s\<^sub>t\<^sub>p S" "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S') \ list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@x#S') \ list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S')" by fastforce+ lemma tfr_stp_all_append: assumes "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" "list_all tfr\<^sub>s\<^sub>t\<^sub>p S'" shows "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@S')" using assms by fastforce lemma tfr_stp_all_wt_subst_apply: assumes "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "bvars\<^sub>s\<^sub>t S \ range_vars \ = {}" shows "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S \\<^sub>s\<^sub>t \)" using assms(1,4) proof (induction S) case (Cons x S) hence IH: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S \\<^sub>s\<^sub>t \)" using tfr_stp_all_split(2)[of x S] unfolding range_vars_alt_def by fastforce thus ?case proof (cases x) case (Equality a t t') hence "(\\. Unifier \ t t') \ \ t = \ t'" using Cons.prems by auto hence "(\\. Unifier \ (t \ \) (t' \ \)) \ \ (t \ \) = \ (t' \ \)" by (metis Unifier_comp' wt_subst_trm'[OF assms(2)]) moreover have "(x#S) \\<^sub>s\<^sub>t \ = Equality a (t \ \) (t' \ \)#(S \\<^sub>s\<^sub>t \)" using \x = Equality a t t'\ by auto ultimately show ?thesis using IH by auto next case (Inequality X F) let ?\ = "rm_vars (set X) \" let ?G = "F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\" let ?P = "\F X. \x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X. \a. \ (Var x) = TAtom a" let ?Q = "\F 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 0: "set X \ range_vars ?\ = {}" using Cons.prems(2) Inequality rm_vars_img_subset[of "set X"] by (auto simp add: subst_domain_def range_vars_alt_def) have 1: "?P F X \ ?Q F X" using Inequality Cons.prems by simp have 2: "fv\<^sub>s\<^sub>e\<^sub>t (?\ ` set X) = set X" by auto have "?P ?G X" when "?P F X" using that proof (induction F) case (Cons g G) obtain t t' where g: "g = (t,t')" by (metis surj_pair) have "\x \ (fv (t \ ?\) \ fv (t' \ ?\)) - set X. \a. \ (Var x) = Var a" proof - have *: "\x \ fv t - set X. \a. \ (Var x) = Var a" "\x \ fv t' - set X. \a. \ (Var x) = Var a" using g Cons.prems by simp_all have **: "\x. wf\<^sub>t\<^sub>r\<^sub>m (?\ x)" using \(2) wf_trm_subst_range_iff[of \] wf_trm_subst_rm_vars'[of \ _ "set X"] by simp show ?thesis using wt_subst_TAtom_fv[OF wt_subst_rm_vars[OF \(1)] ** *(1)] wt_subst_TAtom_fv[OF wt_subst_rm_vars[OF \(1)] ** *(2)] wt_subst_trm'[OF wt_subst_rm_vars[OF \(1), of "set X"]] 2 by blast qed moreover have "\x\fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s (G \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\) - set X. \a. \ (Var x) = Var a" using Cons by auto ultimately show ?case using g by (auto simp add: subst_apply_pairs_def) qed (simp add: subst_apply_pairs_def) hence "?P ?G X \ ?Q ?G X" using 1 ineq_subterm_inj_cond_subst[OF 0, of "trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F"] trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s_subst[of F ?\] by presburger moreover have "(x#S) \\<^sub>s\<^sub>t \ = Inequality X (F \\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s ?\)#(S \\<^sub>s\<^sub>t \)" using \x = Inequality X F\ by auto ultimately show ?thesis using IH by simp qed auto qed simp lemma tfr_stp_all_same_type: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (S@Equality a t t'#S') \ Unifier \ t t' \ \ t = \ t'" by force+ lemma tfr_subset: "\A B. tfr\<^sub>s\<^sub>e\<^sub>t (A \ B) \ tfr\<^sub>s\<^sub>e\<^sub>t A" "\A B. tfr\<^sub>s\<^sub>e\<^sub>t B \ A \ B \ tfr\<^sub>s\<^sub>e\<^sub>t A" "\A B. tfr\<^sub>s\<^sub>e\<^sub>t B \ SMP A \ SMP B \ tfr\<^sub>s\<^sub>e\<^sub>t A" proof - show 1: "tfr\<^sub>s\<^sub>e\<^sub>t (A \ B) \ tfr\<^sub>s\<^sub>e\<^sub>t A" for A B using SMP_union[of A B] unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by simp fix A B assume B: "tfr\<^sub>s\<^sub>e\<^sub>t B" show "A \ B \ tfr\<^sub>s\<^sub>e\<^sub>t A" proof - assume "A \ B" then obtain C where "B = A \ C" by moura thus ?thesis using B 1 by blast qed show "SMP A \ SMP B \ tfr\<^sub>s\<^sub>e\<^sub>t A" proof - assume "SMP A \ SMP B" then obtain C where "SMP B = SMP A \ C" by moura thus ?thesis using B unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast qed qed lemma tfr_empty[simp]: "tfr\<^sub>s\<^sub>e\<^sub>t {}" unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by simp lemma tfr_consts_mono: assumes "\t \ M. \c. t = Fun c []" and "\t \ M. Ana t = ([], [])" and "tfr\<^sub>s\<^sub>e\<^sub>t N" shows "tfr\<^sub>s\<^sub>e\<^sub>t (N \ M)" proof - { fix s t assume *: "s \ SMP (N \ M) - range Var" "t \ SMP (N \ M) - range Var" "\\. Unifier \ s t" hence **: "is_Fun s" "is_Fun t" "s \ SMP N \ s \ M" "t \ SMP N \ t \ M" using assms(3) SMP_consts[OF assms(1,2)] SMP_union[of N M] by auto moreover have "\ s = \ t" when "s \ SMP N" "t \ SMP N" using that assms(3) *(3) **(1,2) unfolding tfr\<^sub>s\<^sub>e\<^sub>t_def by blast moreover have "\ s = \ t" when st: "s \ M" "t \ M" proof - obtain c d where "s = Fun c []" "t = Fun d []" using st assms(1) by moura hence "s = t" using *(3) by fast thus ?thesis by metis qed moreover have "\ s = \ t" when st: "s \ SMP N" "t \ M" proof - obtain c where "t = Fun c []" using st assms(1) by moura hence "s = t" using *(3) **(1,2) by auto thus ?thesis by metis qed moreover have "\ s = \ t" when st: "s \ M" "t \ SMP N" proof - obtain c where "s = Fun c []" using st assms(1) by moura hence "s = t" using *(3) **(1,2) by auto thus ?thesis by metis qed ultimately have "\ s = \ t" by metis } thus ?thesis by (metis tfr\<^sub>s\<^sub>e\<^sub>t_def) qed lemma dual\<^sub>s\<^sub>t_tfr\<^sub>s\<^sub>t\<^sub>p: "list_all tfr\<^sub>s\<^sub>t\<^sub>p S \ list_all tfr\<^sub>s\<^sub>t\<^sub>p (dual\<^sub>s\<^sub>t S)" proof (induction S) case (Cons x S) have "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" using Cons.prems by simp hence IH: "list_all tfr\<^sub>s\<^sub>t\<^sub>p (dual\<^sub>s\<^sub>t S)" using Cons.IH by metis from Cons show ?case proof (cases x) case (Equality a t t') hence "(\\. Unifier \ t t') \ \ t = \ t'" using Cons by auto thus ?thesis using Equality IH by fastforce next case (Inequality X F) have "set (dual\<^sub>s\<^sub>t (x#S)) = insert x (set (dual\<^sub>s\<^sub>t S))" using Inequality by auto moreover have "(\x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X. \a. \ (Var x) = Var a) \ (\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))" using Cons.prems Inequality by auto ultimately show ?thesis using Inequality IH by auto qed auto qed simp lemma subst_var_inv_wt: assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (subst_var_inv \ X)" using assms f_inv_into_f[of _ \ X] unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def subst_var_inv_def by presburger lemma subst_var_inv_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (subst_var_inv \ X))" using f_inv_into_f[of _ \ X] unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def subst_var_inv_def by auto lemma unify_list_wt_if_same_type: assumes "Unification.unify E B = Some U" "\(s,t) \ set E. wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t \ \ s = \ t" and "\(v,t) \ set B. \ (Var v) = \ t" shows "\(v,t) \ set U. \ (Var v) = \ t" using assms proof (induction E B arbitrary: U rule: Unification.unify.induct) case (2 f X g Y E B U) hence "wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)" "wf\<^sub>t\<^sub>r\<^sub>m (Fun g Y)" "\ (Fun f X) = \ (Fun g Y)" by auto from "2.prems"(1) obtain E' where *: "decompose (Fun f X) (Fun g Y) = Some E'" and [simp]: "f = g" "length X = length Y" "E' = zip X Y" and **: "Unification.unify (E'@E) B = Some U" by (auto split: option.splits) have "\(s,t) \ set E'. wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t \ \ s = \ t" proof - { fix s t assume "(s,t) \ set E'" then obtain X' X'' Y' Y'' where "X = X'@s#X''" "Y = Y'@t#Y''" "length X' = length Y'" using zip_arg_subterm_split[of s t X Y] \E' = zip X Y\ by metis hence "\ (Fun f (X'@s#X'')) = \ (Fun g (Y'@t#Y''))" by (metis \\ (Fun f X) = \ (Fun g Y)\) from \E' = zip X Y\ have "\(s,t) \ set E'. s \ Fun f X \ t \ Fun g Y" using zip_arg_subterm[of _ _ X Y] by blast with \(s,t) \ set E'\ have "wf\<^sub>t\<^sub>r\<^sub>m s" "wf\<^sub>t\<^sub>r\<^sub>m t" using wf_trm_subterm \wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)\ \wf\<^sub>t\<^sub>r\<^sub>m (Fun g Y)\ by (blast,blast) moreover have "f \ \\<^sub>f" proof (rule ccontr) assume "f \ \\<^sub>f" hence "f \ \" "arity f = 0" using const_arity_eq_zero[of f] by simp_all thus False using \wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)\ * \(s,t) \ set E'\ unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by auto qed hence "\ s = \ t" using fun_type_arg_cong' \f \ \\<^sub>f\ \\ (Fun f (X'@s#X'')) = \ (Fun g (Y'@t#Y''))\ \length X' = length Y'\ \f = g\ by metis ultimately have "wf\<^sub>t\<^sub>r\<^sub>m s" "wf\<^sub>t\<^sub>r\<^sub>m t" "\ s = \ t" by metis+ } thus ?thesis by blast qed moreover have "\(s,t) \ set E. wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t \ \ s = \ t" using "2.prems"(2) by auto ultimately show ?case using "2.IH"[OF * ** _ "2.prems"(3)] by fastforce next case (3 v t E B U) hence "\ (Var v) = \ t" "wf\<^sub>t\<^sub>r\<^sub>m t" by auto hence "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (subst v t)" and *: "\(v, t) \ set ((v,t)#B). \ (Var v) = \ t" "\t t'. (t,t') \ set E \ \ t = \ t'" using "3.prems"(2,3) unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def subst_def by auto show ?case proof (cases "t = Var v") assume "t = Var v" thus ?case using 3 by auto next assume "t \ Var v" hence "v \ fv t" using "3.prems"(1) by auto hence **: "Unification.unify (subst_list (subst v t) E) ((v, t)#B) = Some U" using Unification.unify.simps(3)[of v t E B] "3.prems"(1) \t \ Var v\ by auto have "\(s, t) \ set (subst_list (subst v t) E). wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t" using wf_trm_subst_singleton[OF _ \wf\<^sub>t\<^sub>r\<^sub>m t\] "3.prems"(2) unfolding subst_list_def subst_def by auto moreover have "\(s, t) \ set (subst_list (subst v t) E). \ s = \ t" using *(2)[THEN wt_subst_trm'[OF \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (subst v t)\]] by (simp add: subst_list_def) ultimately show ?thesis using "3.IH"(2)[OF \t \ Var v\ \v \ fv t\ ** _ *(1)] by auto qed next case (4 f X v E B U) hence "\ (Var v) = \ (Fun f X)" "wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)" by auto hence "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (subst v (Fun f X))" and *: "\(v, t) \ set ((v,(Fun f X))#B). \ (Var v) = \ t" "\t t'. (t,t') \ set E \ \ t = \ t'" using "4.prems"(2,3) unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def subst_def by auto have "v \ fv (Fun f X)" using "4.prems"(1) by force hence **: "Unification.unify (subst_list (subst v (Fun f X)) E) ((v, (Fun f X))#B) = Some U" using Unification.unify.simps(3)[of v "Fun f X" E B] "4.prems"(1) by auto have "\(s, t) \ set (subst_list (subst v (Fun f X)) E). wf\<^sub>t\<^sub>r\<^sub>m s \ wf\<^sub>t\<^sub>r\<^sub>m t" using wf_trm_subst_singleton[OF _ \wf\<^sub>t\<^sub>r\<^sub>m (Fun f X)\] "4.prems"(2) unfolding subst_list_def subst_def by auto moreover have "\(s, t) \ set (subst_list (subst v (Fun f X)) E). \ s = \ t" using *(2)[THEN wt_subst_trm'[OF \wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (subst v (Fun f X))\]] by (simp add: subst_list_def) ultimately show ?case using "4.IH"[OF \v \ fv (Fun f X)\ ** _ *(1)] by auto qed auto lemma mgu_wt_if_same_type: assumes "mgu s t = Some \" "wf\<^sub>t\<^sub>r\<^sub>m s" "wf\<^sub>t\<^sub>r\<^sub>m t" "\ s = \ t" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" proof - let ?fv_disj = "\v t S. \(\(v',t') \ S - {(v,t)}. (insert v (fv t)) \ (insert v' (fv t')) \ {})" from assms(1) obtain \' where "Unification.unify [(s,t)] [] = Some \'" "subst_of \' = \" - by (auto split: option.splits) + by (auto simp: mgu_def split: option.splits) hence "\(v,t) \ set \'. \ (Var v) = \ t" "distinct (map fst \')" using assms(2,3,4) unify_list_wt_if_same_type unify_list_distinct[of "[(s,t)]"] by auto thus "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using \subst_of \' = \\ unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def proof (induction \' arbitrary: \ rule: List.rev_induct) case (snoc tt \' \) then obtain v t where tt: "tt = (v,t)" by (metis surj_pair) hence \: "\ = subst v t \\<^sub>s subst_of \'" using snoc.prems(3) by simp have "\(v,t) \ set \'. \ (Var v) = \ t" "distinct (map fst \')" using snoc.prems(1,2) by auto then obtain \'' where \'': "subst_of \' = \''" "\v. \ (Var v) = \ (\'' v)" by (metis snoc.IH) hence "\ t = \ (t \ \'')" for t using wt_subst_trm by blast hence "\ (Var v) = \ (\'' v)" "\ t = \ (t \ \'')" using \''(2) by auto moreover have "\ (Var v) = \ t" using snoc.prems(1) tt by simp moreover have \2: "\ = Var(v := t) \\<^sub>s \'' " using \ \''(1) unfolding subst_def by simp ultimately have "\ (Var v) = \ (\ v)" unfolding subst_compose_def by simp have "subst_domain (subst v t) \ {v}" unfolding subst_def by (auto simp add: subst_domain_def) hence *: "subst_domain \ \ insert v (subst_domain \'')" using tt \ \''(1) snoc.prems(2) subst_domain_compose[of _ \''] by (auto simp add: subst_domain_def) have "v \ set (map fst \')" using tt snoc.prems(2) by auto hence "v \ subst_domain \''" using \''(1) subst_of_dom_subset[of \'] by auto { fix w assume "w \ subst_domain \''" hence "\ w = \'' w" using \2 \''(1) \v \ subst_domain \''\ unfolding subst_compose_def by auto hence "\ (Var w) = \ (\ w)" using \''(2) by simp } thus ?case using \\ (Var v) = \ (\ v)\ * by force qed simp qed lemma wt_Unifier_if_Unifier: assumes s_t: "wf\<^sub>t\<^sub>r\<^sub>m s" "wf\<^sub>t\<^sub>r\<^sub>m t" "\ s = \ t" and \: "Unifier \ s t" shows "\\. Unifier \ s t \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using mgu_always_unifies[OF \] mgu_gives_MGU[THEN MGU_is_Unifier[of s _ t]] mgu_wt_if_same_type[OF _ s_t] mgu_wf_trm[OF _ s_t(1,2)] wf_trm_subst_range_iff by fast end subsection \Automatically Proving Type-Flaw Resistance\ subsubsection \Definitions: Variable Renaming\ abbreviation "max_var t \ Max (insert 0 (snd ` fv t))" abbreviation "max_var_set X \ Max (insert 0 (snd ` X))" definition "var_rename n v \ Var (fst v, snd v + Suc n)" definition "var_rename_inv n v \ Var (fst v, snd v - Suc n)" subsubsection \Definitions: Computing a Finite Representation of the Sub-Message Patterns\ text \A sufficient requirement for a term to be a well-typed instance of another term\ definition is_wt_instance_of_cond where "is_wt_instance_of_cond \ t s \ ( \ t = \ s \ (case mgu t s of None \ False | Some \ \ inj_on \ (fv t) \ (\x \ fv t. is_Var (\ x))))" definition has_all_wt_instances_of where "has_all_wt_instances_of \ N M \ \t \ N. \s \ M. is_wt_instance_of_cond \ t s" text \This function computes a finite representation of the set of sub-message patterns\ definition SMP0 where "SMP0 Ana \ M \ let f = \t. Fun (the_Fun (\ t)) (map Var (zip (args (\ t)) [0.. t))])); g = \M'. map f (filter (\t. is_Var t \ is_Fun (\ t)) M')@ concat (map (fst \ Ana) M')@concat (map subterms_list M'); h = remdups \ g in while (\A. set (h A) \ set A) h M" text \These definitions are useful to refine an SMP representation set\ fun generalize_term where "generalize_term _ _ n (Var x) = (Var x, n)" | "generalize_term \ p n (Fun f T) = (let \ = \ (Fun f T) in if p \ then (Var (\, n), Suc n) else let (T',n') = foldr (\t (S,m). let (t',m') = generalize_term \ p m t in (t'#S,m')) T ([],n) in (Fun f T', n'))" definition generalize_terms where "generalize_terms \ p \ map (fst \ generalize_term \ p 0)" definition remove_superfluous_terms where "remove_superfluous_terms \ T \ let f = \S t R. \s \ set S - R. s \ t \ is_wt_instance_of_cond \ t s; g = \S t (U,R). if f S t R then (U, insert t R) else (t#U, R); h = \S. remdups (fst (foldr (g S) S ([],{}))) in while (\S. h S \ S) h T" subsubsection \Definitions: Checking Type-Flaw Resistance\ definition is_TComp_var_instance_closed where "is_TComp_var_instance_closed \ M \ \x \ fv\<^sub>s\<^sub>e\<^sub>t M. is_Fun (\ (Var x)) \ (\t \ M. is_Fun t \ \ t = \ (Var x) \ list_all is_Var (args t) \ distinct (args t))" definition finite_SMP_representation where "finite_SMP_representation arity Ana \ M \ (M = {} \ card M > 0) \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s' arity M \ has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M \ has_all_wt_instances_of \ (\((set \ fst \ Ana) ` M)) M \ is_TComp_var_instance_closed \ M" definition comp_tfr\<^sub>s\<^sub>e\<^sub>t where "comp_tfr\<^sub>s\<^sub>e\<^sub>t arity Ana \ M \ finite_SMP_representation arity Ana \ M \ (let \ = var_rename (max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M)) in \s \ M. \t \ M. is_Fun s \ is_Fun t \ \ s \ \ t \ mgu s (t \ \) = None)" fun comp_tfr\<^sub>s\<^sub>t\<^sub>p where "comp_tfr\<^sub>s\<^sub>t\<^sub>p \ (\_: t \ t'\\<^sub>s\<^sub>t) = (mgu t t' \ None \ \ t = \ t')" | "comp_tfr\<^sub>s\<^sub>t\<^sub>p \ (\X\\\: F\\<^sub>s\<^sub>t) = ( (\x \ fv\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F - set X. is_Var (\ (Var x))) \ (\u \ subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F). is_Fun u \ (args u = [] \ (\s \ set (args u). s \ Var ` set X))))" | "comp_tfr\<^sub>s\<^sub>t\<^sub>p _ _ = True" definition comp_tfr\<^sub>s\<^sub>t where "comp_tfr\<^sub>s\<^sub>t arity Ana \ M S \ list_all (comp_tfr\<^sub>s\<^sub>t\<^sub>p \) S \ list_all (wf\<^sub>t\<^sub>r\<^sub>m' arity) (trms_list\<^sub>s\<^sub>t S) \ has_all_wt_instances_of \ (trms\<^sub>s\<^sub>t S) M \ comp_tfr\<^sub>s\<^sub>e\<^sub>t arity Ana \ M" subsubsection \Small Lemmata\ lemma max_var_set_mono: assumes "finite N" and "M \ N" shows "max_var_set M \ max_var_set N" by (meson assms Max.subset_imp finite.insertI finite_imageI image_mono insert_mono insert_not_empty) lemma less_Suc_max_var_set: assumes z: "z \ X" and X: "finite X" shows "snd z < Suc (max_var_set X)" proof - have "snd z \ snd ` X" using z by simp hence "snd z \ Max (insert 0 (snd ` X))" using X by simp thus ?thesis using X by simp qed lemma (in typed_model) finite_SMP_representationD: assumes "finite_SMP_representation arity Ana \ M" shows "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" and "has_all_wt_instances_of \ (\((set \ fst \ Ana) ` M)) M" and "is_TComp_var_instance_closed \ M" and "finite M" using assms wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code[of M] unfolding finite_SMP_representation_def list_all_iff card_gt_0_iff by blast+ lemma (in typed_model) is_wt_instance_of_condD: assumes t_instance_s: "is_wt_instance_of_cond \ t s" obtains \ where "\ t = \ s" "mgu t s = Some \" "inj_on \ (fv t)" "\ ` (fv t) \ range Var" using t_instance_s unfolding is_wt_instance_of_cond_def Let_def by (cases "mgu t s") fastforce+ lemma (in typed_model) is_wt_instance_of_condD': assumes t_wf_trm: "wf\<^sub>t\<^sub>r\<^sub>m t" and s_wf_trm: "wf\<^sub>t\<^sub>r\<^sub>m s" and t_instance_s: "is_wt_instance_of_cond \ t s" shows "\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = s \ \" proof - obtain \ where s: "\ t = \ s" "mgu t s = Some \" "inj_on \ (fv t)" "\ ` (fv t) \ range Var" by (metis is_wt_instance_of_condD[OF t_instance_s]) have 0: "wf\<^sub>t\<^sub>r\<^sub>m t" "wf\<^sub>t\<^sub>r\<^sub>m s" using s(1) t_wf_trm s_wf_trm by auto note 1 = mgu_wt_if_same_type[OF s(2) 0 s(1)] note 2 = conjunct1[OF mgu_gives_MGU[OF s(2)]] show ?thesis using s(1) inj_var_ran_unifiable_has_subst_match[OF 2 s(3,4)] wt_subst_compose[OF 1 subst_var_inv_wt[OF 1, of "fv t"]] wf_trms_subst_compose[OF mgu_wf_trms[OF s(2) 0] subst_var_inv_wf_trms[of \ "fv t"]] by auto qed lemma (in typed_model) is_wt_instance_of_condD'': assumes s_wf_trm: "wf\<^sub>t\<^sub>r\<^sub>m s" and t_instance_s: "is_wt_instance_of_cond \ t s" and t_var: "t = Var x" shows "\y. s = Var y \ \ (Var y) = \ (Var x)" proof - obtain \ where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and s: "Var x = s \ \" using is_wt_instance_of_condD'[OF _ s_wf_trm t_instance_s] t_var by auto obtain y where y: "s = Var y" using s by (cases s) auto show ?thesis using wt_subst_trm''[OF \] s y by metis qed lemma (in typed_model) has_all_wt_instances_ofD: assumes N_instance_M: "has_all_wt_instances_of \ N M" and t_in_N: "t \ N" obtains s \ where "s \ M" "\ t = \ s" "mgu t s = Some \" "inj_on \ (fv t)" "\ ` (fv t) \ range Var" by (metis t_in_N N_instance_M is_wt_instance_of_condD has_all_wt_instances_of_def) lemma (in typed_model) has_all_wt_instances_ofD': assumes N_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and M_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and N_instance_M: "has_all_wt_instances_of \ N M" and t_in_N: "t \ N" shows "\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t \ M \\<^sub>s\<^sub>e\<^sub>t \" using assms is_wt_instance_of_condD' unfolding has_all_wt_instances_of_def by fast lemma (in typed_model) has_all_wt_instances_ofD'': assumes N_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and M_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and N_instance_M: "has_all_wt_instances_of \ N M" and t_in_N: "Var x \ N" shows "\y. Var y \ M \ \ (Var y) = \ (Var x)" using assms is_wt_instance_of_condD'' unfolding has_all_wt_instances_of_def by fast lemma (in typed_model) has_all_instances_of_if_subset: assumes "N \ M" shows "has_all_wt_instances_of \ N M" using assms inj_onI mgu_same_empty unfolding has_all_wt_instances_of_def is_wt_instance_of_cond_def by (smt option.case_eq_if option.discI option.sel subsetD term.discI(1) term.inject(1)) lemma (in typed_model) SMP_I': assumes N_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s N" and M_wf_trms: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and N_instance_M: "has_all_wt_instances_of \ N M" and t_in_N: "t \ N" shows "t \ SMP M" using has_all_wt_instances_ofD'[OF N_wf_trms M_wf_trms N_instance_M t_in_N] SMP.Substitution[OF SMP.MP[of _ M]] by blast subsubsection \Lemma: Proving Type-Flaw Resistance\ locale typed_model' = typed_model arity public Ana \ 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" + assumes \_Var_fst: "\\ n m. \ (Var (\,n)) = \ (Var (\,m))" and Ana_const: "\c T. arity c = 0 \ Ana (Fun c T) = ([],[])" and Ana_subst'_or_Ana_keys_subterm: "(\f T \ K R. Ana (Fun f T) = (K,R) \ Ana (Fun f T \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \,R \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)) \ (\t K R k. Ana t = (K,R) \ k \ set K \ k \ t)" begin lemma var_rename_inv_comp: "t \ (var_rename n \\<^sub>s var_rename_inv n) = t" proof (induction t) case (Fun f T) hence "map (\t. t \ var_rename n \\<^sub>s var_rename_inv n) T = T" by (simp add: map_idI) thus ?case by (metis subst_apply_term.simps(2)) qed (simp add: var_rename_def var_rename_inv_def) lemma var_rename_fv_disjoint: "fv s \ fv (t \ var_rename (max_var s)) = {}" proof - have 1: "\v \ fv s. snd v \ max_var s" by simp have 2: "\v \ fv (t \ var_rename n). snd v > n" for n unfolding var_rename_def by (induct t) auto show ?thesis using 1 2 by force qed lemma var_rename_fv_set_disjoint: assumes "finite M" "s \ M" shows "fv s \ fv (t \ var_rename (max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M))) = {}" proof - have 1: "\v \ fv s. snd v \ max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M)" using assms proof (induction M rule: finite_induct) case (insert t M) thus ?case proof (cases "t = s") case False hence "\v \ fv s. snd v \ max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M)" using insert by simp moreover have "max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M) \ max_var_set (fv\<^sub>s\<^sub>e\<^sub>t (insert t M))" using insert.hyps(1) insert.prems by force ultimately show ?thesis by auto qed simp qed simp have 2: "\v \ fv (t \ var_rename n). snd v > n" for n unfolding var_rename_def by (induct t) auto show ?thesis using 1 2 by force qed lemma var_rename_fv_set_disjoint': assumes "finite M" shows "fv\<^sub>s\<^sub>e\<^sub>t M \ fv\<^sub>s\<^sub>e\<^sub>t (N \\<^sub>s\<^sub>e\<^sub>t var_rename (max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M))) = {}" using var_rename_fv_set_disjoint[OF assms] by auto lemma var_rename_is_renaming[simp]: "subst_range (var_rename n) \ range Var" "subst_range (var_rename_inv n) \ range Var" unfolding var_rename_def var_rename_inv_def by auto lemma var_rename_wt[simp]: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (var_rename n)" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (var_rename_inv n)" by (auto simp add: var_rename_def var_rename_inv_def wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def \_Var_fst) lemma var_rename_wt': assumes "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "s = m \ \" shows "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (var_rename_inv n \\<^sub>s \)" "s = m \ var_rename n \ var_rename_inv n \\<^sub>s \" using assms(2) wt_subst_compose[OF var_rename_wt(2)[of n] assms(1)] var_rename_inv_comp[of m n] by force+ lemma var_rename_wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_range[simp]: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (var_rename n))" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (var_rename_inv n))" using var_rename_is_renaming by fastforce+ lemma Fun_range_case: "(\f T. Fun f T \ M \ P f T) \ (\u \ M. case u of Fun f T \ P f T | _ \ True)" "(\f T. Fun f T \ M \ P f T) \ (\u \ M. is_Fun u \ P (the_Fun u) (args u))" by (auto split: "term.splits") lemma is_TComp_var_instance_closedD: assumes x: "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var x) = \ (Var y)" "\ (Var x) = TComp f T" and closed: "is_TComp_var_instance_closed \ M" shows "\g U. Fun g U \ M \ \ (Fun g U) = \ (Var x) \ (\u \ set U. is_Var u) \ distinct U" using assms unfolding is_TComp_var_instance_closed_def list_all_iff list_ex_iff by fastforce lemma is_TComp_var_instance_closedD': assumes "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var x) = \ (Var y)" "TComp f T \ \ (Var x)" and closed: "is_TComp_var_instance_closed \ M" and wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" shows "\g U. Fun g U \ M \ \ (Fun g U) = TComp f T \ (\u \ set U. is_Var u) \ distinct U" using assms(1,2) proof (induction "\ (Var x)" arbitrary: x) case (Fun g U) note IH = Fun.hyps(1) have g: "arity g > 0" using Fun.hyps(2) fun_type_inv[of "Var x"] \_Var_fst by simp_all then obtain V where V: "Fun g V \ M" "\ (Fun g V) = \ (Var x)" "\v \ set V. \x. v = Var x" "distinct V" "length U = length V" using is_TComp_var_instance_closedD[OF Fun.prems(1) Fun.hyps(2)[symmetric] closed(1)] by (metis Fun.hyps(2) fun_type_id_eq fun_type_length_eq is_VarE) hence U: "U = map \ V" using fun_type[OF g(1), of V] Fun.hyps(2) by simp hence 1: "\ v \ set U" when v: "v \ set V" for v using v by simp have 2: "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var z) = \ (Var y)" when z: "Var z \ set V" for z using V(1) fv_subset_subterms Fun_param_in_subterms[OF z] by fastforce show ?case proof (cases "TComp f T = \ (Var x)") case False then obtain u where u: "u \ set U" "TComp f T \ u" using Fun.prems(2) Fun.hyps(2) by moura then obtain y where y: "Var y \ set V" "\ (Var y) = u" using U V(3) \_Var_fst by auto show ?thesis using IH[OF _ 2[OF y(1)]] u y(2) by metis qed (use V in fastforce) qed simp lemma TComp_var_instance_wt_subst_exists: assumes gT: "\ (Fun g T) = TComp g (map \ U)" "wf\<^sub>t\<^sub>r\<^sub>m (Fun g T)" and U: "\u \ set U. \y. u = Var y" "distinct U" shows "\\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ Fun g T = Fun g U \ \" proof - define the_i where "the_i \ \y. THE x. x < length U \ U ! x = Var y" define \ where \: "\ \ \y. if Var y \ set U then T ! the_i y else Var y" have g: "arity g > 0" using gT(1,2) fun_type_inv(1) by blast have UT: "length U = length T" using fun_type_length_eq gT(1) by fastforce have 1: "the_i y < length U \ U ! the_i y = Var y" when y: "Var y \ set U" for y using theI'[OF distinct_Ex1[OF U(2) y]] unfolding the_i_def by simp have 2: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using \ 1 gT(1) fun_type[OF g] UT unfolding wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by (metis (no_types, lifting) nth_map term.inject(2)) have "\i \ = T ! i" using \ 1 U(1) UT distinct_Ex1[OF U(2)] in_set_conv_nth by (metis (no_types, lifting) subst_apply_term.simps(1)) hence "T = map (\t. t \ \) U" by (simp add: UT nth_equalityI) hence 3: "Fun g T = Fun g U \ \" by simp have "subst_range \ \ set T" using \ 1 U(1) UT by (auto simp add: subst_domain_def) hence 4: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using gT(2) wf_trm_param by auto show ?thesis by (metis 2 3 4) qed lemma TComp_var_instance_closed_has_Var: assumes closed: "is_TComp_var_instance_closed \ M" and wf_M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and wf_\x: "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" and y_ex: "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var x) = \ (Var y)" and t: "t \ \ x" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" shows "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var y) = \ t" proof (cases "\ (Var x)") case (Var a) hence "t = \ x" using t wf_\x \_wt by (metis (full_types) const_type_inv_wf fun_if_subterm subtermeq_Var_const(2) wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) thus ?thesis using y_ex wt_subst_trm''[OF \_wt, of "Var x"] by fastforce next case (Fun f T) hence \_\x: "\ (\ x) = TComp f T" using wt_subst_trm''[OF \_wt, of "Var x"] by auto show ?thesis proof (cases "t = \ x") case False hence t_subt_\x: "t \ \ x" using t(1) \_\x by fastforce obtain T' where T': "\ x = Fun f T'" using \_\x t_subt_\x fun_type_id_eq by (cases "\ x") auto obtain g S where gS: "Fun g S \ \ x" "t \ set S" using Fun_ex_if_subterm[OF t_subt_\x] by blast have gS_wf: "wf\<^sub>t\<^sub>r\<^sub>m (Fun g S)" by (rule wf_trm_subtermeq[OF wf_\x gS(1)]) hence "arity g > 0" using gS(2) by (metis length_pos_if_in_set wf_trm_arity) hence gS_\: "\ (Fun g S) = TComp g (map \ S)" using fun_type by blast obtain h U where hU: "Fun h U \ M" "\ (Fun h U) = Fun g (map \ S)" "\u \ set U. is_Var u" using is_TComp_var_instance_closedD'[OF y_ex _ closed wf_M] subtermeq_imp_subtermtypeeq[OF wf_\x] gS \_\x Fun gS_\ by metis obtain y where y: "Var y \ set U" "\ (Var y) = \ t" using hU(3) fun_type_param_ex[OF hU(2) gS(2)] by fast have "y \ fv\<^sub>s\<^sub>e\<^sub>t M" using hU(1) y(1) by force thus ?thesis using y(2) closed by metis qed (metis y_ex Fun \_\x) qed lemma TComp_var_instance_closed_has_Fun: assumes closed: "is_TComp_var_instance_closed \ M" and wf_M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and wf_\x: "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" and y_ex: "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var x) = \ (Var y)" and t: "t \ \ x" and \_wt: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" and t_\: "\ t = TComp g T" and t_fun: "is_Fun t" shows "\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = m \ \ \ is_Fun m" proof - obtain T'' where T'': "t = Fun g T''" using t_\ t_fun fun_type_id_eq by blast have g: "arity g > 0" using t_\ fun_type_inv[of t] by simp_all have "TComp g T \ \ (Var x)" using \_wt t t_\ by (metis wf_\x subtermeq_imp_subtermtypeeq wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) then obtain U where U: "Fun g U \ M" "\ (Fun g U) = TComp g T" "\u \ set U. \y. u = Var y" "distinct U" "length T'' = length U" using is_TComp_var_instance_closedD'[OF y_ex _ closed wf_M] by (metis t_\ T'' fun_type_id_eq fun_type_length_eq is_VarE) hence UT': "T = map \ U" using fun_type[OF g, of U] by simp show ?thesis using TComp_var_instance_wt_subst_exists UT' T'' U(1,3,4) t t_\ wf_\x wf_trm_subtermeq by (metis term.disc(2)) qed lemma TComp_var_and_subterm_instance_closed_has_subterms_instances: assumes M_var_inst_cl: "is_TComp_var_instance_closed \ M" and M_subterms_cl: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" and M_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and t: "t \\<^sub>s\<^sub>e\<^sub>t M" and s: "s \ t \ \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" shows "\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ s = m \ \" using subterm_subst_unfold[OF s] proof assume "\s'. s' \ t \ s = s' \ \" then obtain s' where s': "s' \ t" "s = s' \ \" by blast then obtain \ where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "s' \ M \\<^sub>s\<^sub>e\<^sub>t \" using t has_all_wt_instances_ofD'[OF wf_trms_subterms[OF M_wf] M_wf M_subterms_cl] term.order_trans[of s' t] by blast then obtain m where m: "m \ M" "s' = m \ \" by blast have "s = m \ (\ \\<^sub>s \)" using s'(2) m(2) by simp thus ?thesis using m(1) wt_subst_compose[OF \(1) \(1)] wf_trms_subst_compose[OF \(2) \(2)] by blast next assume "\x \ fv t. s \ \ x" then obtain x where x: "x \ fv t" "s \ \ x" "s \ \ x" by blast note 0 = TComp_var_instance_closed_has_Var[OF M_var_inst_cl M_wf] note 1 = has_all_wt_instances_ofD''[OF wf_trms_subterms[OF M_wf] M_wf M_subterms_cl] have \x_wf: "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" and s_wf_trm: "wf\<^sub>t\<^sub>r\<^sub>m s" using \(2) wf_trm_subterm[OF _ x(2)] by fastforce+ have x_fv_ex: "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var x) = \ (Var y)" using x(1) s fv_subset_subterms[OF t] by auto obtain y where y: "y \ fv\<^sub>s\<^sub>e\<^sub>t M" "\ (Var y) = \ s" using 0[of \ x s, OF \x_wf x_fv_ex x(3) \(1)] by metis then obtain z where z: "Var z \ M" "\ (Var z) = \ s" using 1[of y] vars_iff_subtermeq_set[of y M] by metis define \ where "\ \ Var(z := s)::('fun, ('fun, 'atom) term \ nat) subst" have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "s = Var z \ \" using z(2) s_wf_trm unfolding \_def wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by force+ thus ?thesis using z(1) by blast qed context begin private lemma SMP_D_aux1: assumes "t \ SMP M" and closed: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" "is_TComp_var_instance_closed \ M" and wf_M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" shows "\x \ fv t. \y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var y) = \ (Var x)" using assms(1) proof (induction t rule: SMP.induct) case (MP t) show ?case proof fix x assume x: "x \ fv t" hence "Var x \ subterms\<^sub>s\<^sub>e\<^sub>t M" using MP.hyps vars_iff_subtermeq by fastforce then obtain \ s where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and s: "s \ M" "Var x = s \ \" using has_all_wt_instances_ofD'[OF wf_trms_subterms[OF wf_M] wf_M closed(1)] by blast then obtain y where y: "s = Var y" by (cases s) auto thus "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var y) = \ (Var x)" using s wt_subst_trm''[OF \(1), of "Var y"] by force qed next case (Subterm t t') hence "fv t' \ fv t" using subtermeq_vars_subset by auto thus ?case using Subterm.IH by auto next case (Substitution t \) note IH = Substitution.IH show ?case proof fix x assume x: "x \ fv (t \ \)" then obtain y where y: "y \ fv t" "\ (Var x) \ \ (Var y)" using Substitution.hyps(2,3) by (metis subst_apply_img_var subtermeqI' subtermeq_imp_subtermtypeeq vars_iff_subtermeq wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def wf_trm_subst_rangeD) let ?P = "\x. \y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var y) = \ (Var x)" show "?P x" using y IH proof (induction "\ (Var y)" arbitrary: y t) case (Var a) hence "\ (Var x) = \ (Var y)" by auto thus ?case using Var(2,4) by auto next case (Fun f T) obtain z where z: "\w \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var z) = \ (Var w)" "\ (Var z) = \ (Var y)" using Fun.prems(1,3) by blast show ?case proof (cases "\ (Var x) = \ (Var y)") case True thus ?thesis using Fun.prems by auto next case False then obtain \ where \: "\ \ set T" "\ (Var x) \ \" using Fun.prems(2) Fun.hyps(2) by auto then obtain U where U: "Fun f U \ M" "\ (Fun f U) = \ (Var z)" "\u \ set U. \v. u = Var v" "distinct U" using is_TComp_var_instance_closedD'[OF z(1) _ closed(2) wf_M] Fun.hyps(2) z(2) by (metis fun_type_id_eq subtermeqI' is_VarE) hence 1: "\x \ fv (Fun f U). \y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var y) = \ (Var x)" by force have "arity f > 0" using U(2) z(2) Fun.hyps(2) fun_type_inv(1) by metis hence "\ (Fun f U) = TComp f (map \ U)" using fun_type by auto then obtain u where u: "Var u \ set U" "\ (Var u) = \" using \(1) U(2,3) z(2) Fun.hyps(2) by auto show ?thesis using Fun.hyps(1)[of u "Fun f U"] u \ 1 by force qed qed qed next case (Ana t K T k) have "fv k \ fv t" using Ana_keys_fv[OF Ana.hyps(2)] Ana.hyps(3) by auto thus ?case using Ana.IH by auto qed private lemma SMP_D_aux2: fixes t::"('fun, ('fun, 'atom) term \ nat) term" assumes t_SMP: "t \ SMP M" and t_Var: "\x. t = Var x" and M_SMP_repr: "finite_SMP_representation arity Ana \ M" shows "\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = m \ \" proof - have M_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and M_var_inst_cl: "is_TComp_var_instance_closed \ M" and M_subterms_cl: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" and M_Ana_cl: "has_all_wt_instances_of \ (\((set \ fst \ Ana) ` M)) M" using finite_SMP_representationD[OF M_SMP_repr] by blast+ have M_Ana_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\ ((set \ fst \ Ana) ` M))" proof fix k assume "k \ \((set \ fst \ Ana) ` M)" then obtain m where m: "m \ M" "k \ set (fst (Ana m))" by force thus "wf\<^sub>t\<^sub>r\<^sub>m k" using M_wf Ana_keys_wf'[of m "fst (Ana m)" _ k] surjective_pairing by blast qed note 0 = has_all_wt_instances_ofD'[OF wf_trms_subterms[OF M_wf] M_wf M_subterms_cl] note 1 = has_all_wt_instances_ofD'[OF M_Ana_wf M_wf M_Ana_cl] obtain x y where x: "t = Var x" and y: "y \ fv\<^sub>s\<^sub>e\<^sub>t M" "\ (Var y) = \ (Var x)" using t_Var SMP_D_aux1[OF t_SMP M_subterms_cl M_var_inst_cl M_wf] by fastforce then obtain m \ where m: "m \ M" "m \ \ = Var y" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using 0[of "Var y"] vars_iff_subtermeq_set[of y M] by fastforce obtain z where z: "m = Var z" using m(2) by (cases m) auto define \ where "\ \ Var(z := Var x)::('fun, ('fun, 'atom) term \ nat) subst" have "\ (Var z) = \ (Var x)" using y(2) m(2) z wt_subst_trm''[OF \, of m] by argo hence "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" unfolding \_def wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by force+ moreover have "t = m \ \" using x z unfolding \_def by simp ultimately show ?thesis using m(1) by blast qed private lemma SMP_D_aux3: assumes hyps: "t' \ t" and wf_t: "wf\<^sub>t\<^sub>r\<^sub>m t" and prems: "is_Fun t'" and IH: "((\f. t = Fun f []) \ (\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = m \ \)) \ (\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = m \ \ \ is_Fun m)" and M_SMP_repr: "finite_SMP_representation arity Ana \ M" shows "((\f. t' = Fun f []) \ (\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t' = m \ \)) \ (\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t' = m \ \ \ is_Fun m)" proof (cases "\f. t = Fun f [] \ t' = Fun f []") case True have M_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and M_var_inst_cl: "is_TComp_var_instance_closed \ M" and M_subterms_cl: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" and M_Ana_cl: "has_all_wt_instances_of \ (\((set \ fst \ Ana) ` M)) M" using finite_SMP_representationD[OF M_SMP_repr] by blast+ note 0 = has_all_wt_instances_ofD'[OF wf_trms_subterms[OF M_wf] M_wf M_subterms_cl] note 1 = TComp_var_instance_closed_has_Fun[OF M_var_inst_cl M_wf] note 2 = TComp_var_and_subterm_instance_closed_has_subterms_instances[ OF M_var_inst_cl M_subterms_cl M_wf] have wf_t': "wf\<^sub>t\<^sub>r\<^sub>m t'" using hyps wf_t wf_trm_subterm by blast obtain c where "t = Fun c [] \ t' = Fun c []" using True by moura thus ?thesis proof assume c: "t' = Fun c []" show ?thesis proof (cases "\f. t = Fun f []") case True hence "t = t'" using c hyps by force thus ?thesis using IH by fast next case False note F = this then obtain m \ where m: "m \ M" "t = 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 IH by blast show ?thesis using subterm_subst_unfold[OF hyps[unfolded m(2)]] proof assume "\m'. m' \ m \ t' = m' \ \" then obtain m' where m': "m' \ m" "t' = m' \ \" by moura obtain n \ where n: "n \ M" "m' = n \ \" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using 0[of m'] m(1) m'(1) by blast have "t' = n \ (\ \\<^sub>s \)" using m'(2) n(2) by auto thus ?thesis using c n(1) wt_subst_compose[OF \(1) \(1)] wf_trms_subst_compose[OF \(2) \(2)] by blast next assume "\x \ fv m. t' \ \ x" then obtain x where x: "x \ fv m" "t' \ \ x" "t' \ \ x" by moura have \x_wf: "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" using \(2) by fastforce have x_fv_ex: "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var x) = \ (Var y)" using x m by auto show ?thesis proof (cases "\ t'") case (Var a) show ?thesis using c m 2[OF _ hyps[unfolded m(2)] \] by fast next case (Fun g S) show ?thesis using c 1[of \ x t', OF \x_wf x_fv_ex x(3) \(1) Fun] by blast qed qed qed qed (use IH hyps in simp) next case False note F = False then obtain m \ where m: "m \ M" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "t = m \ \" "is_Fun m" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using IH by moura obtain f T where fT: "t' = Fun f T" "arity f > 0" "\ t' = TComp f (map \ T)" using F prems fun_type wf_trm_subtermeq[OF wf_t hyps] by (metis is_FunE length_greater_0_conv subtermeqI' wf\<^sub>t\<^sub>r\<^sub>m_def) have closed: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" "is_TComp_var_instance_closed \ M" using M_SMP_repr unfolding finite_SMP_representation_def by metis+ have M_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" using finite_SMP_representationD[OF M_SMP_repr] by blast show ?thesis proof (cases "\x \ fv m. t' \ \ x") case True then obtain x where x: "x \ fv m" "t' \ \ x" by moura have 1: "x \ fv\<^sub>s\<^sub>e\<^sub>t M" using m(1) x(1) by auto have 2: "is_Fun (\ x)" using prems x(2) by auto have 3: "wf\<^sub>t\<^sub>r\<^sub>m (\ x)" using m(5) by (simp add: wf_trm_subst_rangeD) have "\(\f. \ x = Fun f [])" using F x(2) by auto hence "\f T. \ (Var x) = TComp f T" using 2 3 m(2) by (metis (no_types) fun_type is_FunE length_greater_0_conv subtermeqI' wf\<^sub>t\<^sub>r\<^sub>m_def wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def) moreover have "\f T. \ t' = Fun f T" using False prems wf_trm_subtermeq[OF wf_t hyps] by (metis (no_types) fun_type is_FunE length_greater_0_conv subtermeqI' wf\<^sub>t\<^sub>r\<^sub>m_def) ultimately show ?thesis using TComp_var_instance_closed_has_Fun 1 x(2) m(2) prems closed 3 M_wf by metis next case False then obtain m' where m': "m' \ m" "t' = m' \ \" "is_Fun m'" using hyps m(3) subterm_subst_not_img_subterm by blast then obtain \ m'' where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "m'' \ M" "m' = m'' \ \" using m(1) has_all_wt_instances_ofD'[OF wf_trms_subterms[OF M_wf] M_wf closed(1)] by blast hence t'_m'': "t' = m'' \ \ \\<^sub>s \" using m'(2) by fastforce note \\ = wt_subst_compose[OF \(1) m(2)] wf_trms_subst_compose[OF \(2) m(5)] show ?thesis proof (cases "is_Fun m''") case True thus ?thesis using \(3,4) m'(2,3) m(4) fT t'_m'' \\ by blast next case False then obtain x where x: "m'' = Var x" by moura hence "\y \ fv\<^sub>s\<^sub>e\<^sub>t M. \ (Var x) = \ (Var y)" "t' \ (\ \\<^sub>s \) x" "\ (Var x) = Fun f (map \ T)" "wf\<^sub>t\<^sub>r\<^sub>m ((\ \\<^sub>s \) x)" using \\ t'_m'' \(3) fv_subset[OF \(3)] fT(3) subst_apply_term.simps(1)[of x "\ \\<^sub>s \"] wt_subst_trm''[OF \\(1), of "Var x"] by (fastforce, blast, argo, fastforce) thus ?thesis using x TComp_var_instance_closed_has_Fun[ of M "\ \\<^sub>s \" x t' f "map \ T", OF closed(2) M_wf _ _ _ \\(1) fT(3) prems] by blast qed qed qed lemma SMP_D: assumes "t \ SMP M" "is_Fun t" and M_SMP_repr: "finite_SMP_representation arity Ana \ M" shows "((\f. t = Fun f []) \ (\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = m \ \)) \ (\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = m \ \ \ is_Fun m)" proof - have wf_M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and closed: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" "has_all_wt_instances_of \ (\((set \ fst \ Ana) ` M)) M" "is_TComp_var_instance_closed \ M" using finite_SMP_representationD[OF M_SMP_repr] by blast+ show ?thesis using assms(1,2) proof (induction t rule: SMP.induct) case (MP t) 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)" "t = t \ Var" by simp_all ultimately show ?case by blast next case (Subterm t t') hence t_fun: "is_Fun t" by auto note * = Subterm.hyps(2) SMP_wf_trm[OF Subterm.hyps(1) wf_M(1)] Subterm.prems Subterm.IH[OF t_fun] M_SMP_repr show ?case by (rule SMP_D_aux3[OF *]) next case (Substitution t \) have wf: "wf\<^sub>t\<^sub>r\<^sub>m t" by (metis Substitution.hyps(1) wf_M(1) SMP_wf_trm) hence wf': "wf\<^sub>t\<^sub>r\<^sub>m (t \ \)" using Substitution.hyps(3) wf_trm_subst by blast show ?case proof (cases "\ t") case (Var a) hence 1: "\ (t \ \) = TAtom a" using Substitution.hyps(2) by (metis wt_subst_trm'') then obtain c where c: "t \ \ = Fun c []" using TAtom_term_cases[OF wf' 1] Substitution.prems by fastforce hence "(\x. t = Var x) \ t = t \ \" by (cases t) auto thus ?thesis proof assume t_Var: "\x. t = Var x" then obtain x where x: "t = Var x" "\ x = Fun c []" "\ (Var x) = TAtom a" using c 1 wt_subst_trm''[OF Substitution.hyps(2), of t] by force obtain m \ where m: "m \ M" "t = 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 SMP_D_aux2[OF Substitution.hyps(1) t_Var M_SMP_repr] by moura have "m \ (\ \\<^sub>s \) = Fun c []" using c m(2) by auto thus ?thesis using c m(1) wt_subst_compose[OF \(1) Substitution.hyps(2)] wf_trms_subst_compose[OF \(2) Substitution.hyps(3)] by metis qed (use c Substitution.IH in auto) next case (Fun f T) hence 1: "\ (t \ \) = TComp f T" using Substitution.hyps(2) by (metis wt_subst_trm'') have 2: "\(\f. t = Fun f [])" using Fun TComp_term_cases[OF wf] by auto obtain T'' where T'': "t \ \ = Fun f T''" using 1 2 fun_type_id_eq Fun Substitution.prems by fastforce have f: "arity f > 0" using fun_type_inv[OF 1] by metis show ?thesis proof (cases t) case (Fun g U) then obtain m \ where m: "m \ M" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "t = m \ \" "is_Fun m" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using Substitution.IH Fun 2 by moura have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" "t \ \ = m \ (\ \\<^sub>s \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using wt_subst_compose[OF m(2) Substitution.hyps(2)] m(3) wf_trms_subst_compose[OF m(5) Substitution.hyps(3)] by auto thus ?thesis using m(1,4) by metis next case (Var x) then obtain y where y: "y \ fv\<^sub>s\<^sub>e\<^sub>t M" "\ (Var y) = \ (Var x)" using SMP_D_aux1[OF Substitution.hyps(1) closed(1,3) wf_M] Fun by moura hence 3: "\ (Var y) = TComp f T" using Var Fun \_Var_fst by simp obtain h V where V: "Fun h V \ M" "\ (Fun h V) = \ (Var y)" "\u \ set V. \z. u = Var z" "distinct V" by (metis is_VarE is_TComp_var_instance_closedD[OF _ 3 closed(3)] y(1)) moreover have "length T'' = length V" using 3 V(2) fun_type_length_eq 1 T'' by metis ultimately have TV: "T = map \ V" by (metis fun_type[OF f(1)] 3 fun_type_id_eq term.inject(2)) obtain \ where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t \ \ = Fun h V \ \" using TComp_var_instance_wt_subst_exists 1 3 T'' TV V(2,3,4) wf' by (metis fun_type_id_eq) have 9: "\ (Fun h V) = \ (\ x)" using y(2) Substitution.hyps(2) V(2) 1 3 Var by auto show ?thesis using Var \ 9 V(1) by force qed qed next case (Ana t K T k) have 1: "is_Fun t" using Ana.hyps(2,3) by auto then obtain f U where U: "t = Fun f U" by moura have 2: "fv k \ fv t" using Ana_keys_fv[OF Ana.hyps(2)] Ana.hyps(3) by auto have wf_t: "wf\<^sub>t\<^sub>r\<^sub>m t" using SMP_wf_trm[OF Ana.hyps(1)] wf\<^sub>t\<^sub>r\<^sub>m_code wf_M by auto hence wf_k: "wf\<^sub>t\<^sub>r\<^sub>m k" using Ana_keys_wf'[OF Ana.hyps(2)] wf\<^sub>t\<^sub>r\<^sub>m_code Ana.hyps(3) by auto have wf_M_keys: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\((set \ fst \ Ana) ` M))" proof fix t assume "t \ (\((set \ fst \ Ana) ` M))" then obtain s where s: "s \ M" "t \ (set \ fst \ Ana) s" by blast obtain K R where KR: "Ana s = (K,R)" by (metis surj_pair) hence "t \ set K" using s(2) by simp thus "wf\<^sub>t\<^sub>r\<^sub>m t" using Ana_keys_wf'[OF KR] wf_M s(1) by blast qed show ?case using Ana_subst'_or_Ana_keys_subterm proof assume "\t K T k. Ana t = (K, T) \ k \ set K \ k \ t" hence *: "k \ t" using Ana.hyps(2,3) by auto show ?thesis by (rule SMP_D_aux3[OF * wf_t Ana.prems Ana.IH[OF 1] M_SMP_repr]) next assume Ana_subst': "\f T \ K M. 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 \)" have "arity f > 0" using Ana_const[of f U] U Ana.hyps(2,3) by fastforce hence "U \ []" using wf_t U unfolding wf\<^sub>t\<^sub>r\<^sub>m_def by force then obtain m \ where m: "m \ M" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t = m \ \" "is_Fun m" using Ana.IH[OF 1] U by auto hence "Ana (t \ \) = (K \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \,T \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Ana_subst' U Ana.hyps(2) by auto obtain Km Tm where Ana_m: "Ana m = (Km,Tm)" by moura hence "Ana (m \ \) = (Km \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \,Tm \\<^sub>l\<^sub>i\<^sub>s\<^sub>t \)" using Ana_subst' U m(4) is_FunE[OF m(5)] Ana.hyps(2) by metis then obtain km where km: "km \ set Km" "k = km \ \" using Ana.hyps(2,3) m(4) by auto then obtain \ km' where \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" and km': "km' \ M" "km = km' \ \" using Ana_m m(1) has_all_wt_instances_ofD'[OF wf_M_keys wf_M closed(2), of km] by force have k\\: "k = km' \ \ \\<^sub>s \" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t (\ \\<^sub>s \)" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range (\ \\<^sub>s \))" using km(2) km' wt_subst_compose[OF \(1) m(2)] wf_trms_subst_compose[OF \(2) m(3)] by auto show ?case proof (cases "is_Fun km'") case True thus ?thesis using k\\ km'(1) by blast next case False note F = False then obtain x where x: "km' = Var x" by auto hence 3: "x \ fv\<^sub>s\<^sub>e\<^sub>t M" using fv_subset[OF km'(1)] by auto obtain kf kT where kf: "k = Fun kf kT" using Ana.prems by auto show ?thesis proof (cases "kT = []") case True thus ?thesis using k\\(1) k\\(2) k\\(3) kf km'(1) by blast next case False hence 4: "arity kf > 0" using wf_k kf TAtom_term_cases const_type by fastforce then obtain kT' where kT': "\ k = TComp kf kT'" by (simp add: fun_type kf) then obtain V where V: "Fun kf V \ M" "\ (Fun kf V) = \ (Var x)" "\u \ set V. \v. u = Var v" "distinct V" "is_Fun (Fun kf V)" using is_TComp_var_instance_closedD[OF _ _ closed(3), of x] x m(2) k\\(1) 3 wt_subst_trm''[OF k\\(2)] by (metis fun_type_id_eq term.disc(2) is_VarE) have 5: "kT' = map \ V" using fun_type[OF 4] x kT' k\\ m(2) V(2) by (metis term.inject(2) wt_subst_trm'') thus ?thesis using TComp_var_instance_wt_subst_exists wf_k kf 4 V(3,4) kT' V(1,5) by metis qed qed qed qed qed lemma SMP_D': fixes M defines "\ \ var_rename (max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M))" assumes M_SMP_repr: "finite_SMP_representation arity Ana \ M" and s: "s \ SMP M" "is_Fun s" "\f. s = Fun f []" and t: "t \ SMP M" "is_Fun t" "\f. t = Fun f []" obtains \ s0 \ t0 where "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "s0 \ M" "is_Fun s0" "s = s0 \ \" "\ s = \ s0" and "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t0 \ M" "is_Fun t0" "t = t0 \ \ \ \" "\ t = \ t0" proof - obtain \ s0 where s0: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "s0 \ M" "s = s0 \ \" "is_Fun s0" using s(3) SMP_D[OF s(1,2) M_SMP_repr] unfolding \_def by metis obtain \ t0 where t0: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t0 \ M" "t = t0 \ \ \ \" "is_Fun t0" using t(3) SMP_D[OF t(1,2) M_SMP_repr] var_rename_wt'[of _ t] wf_trms_subst_compose_Var_range(1)[OF _ var_rename_is_renaming(2)] unfolding \_def by metis have "\ s = \ s0" "\ t = \ (t0 \ \)" "\ (t0 \ \) = \ t0" using s0 t0 wt_subst_trm'' by (metis, metis, metis \_def var_rename_wt(1)) thus ?thesis using s0 t0 that by simp qed lemma SMP_D'': fixes t::"('fun, ('fun, 'atom) term \ nat) term" assumes t_SMP: "t \ SMP M" and M_SMP_repr: "finite_SMP_representation arity Ana \ M" shows "\m \ M. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ t = m \ \" proof (cases "(\x. t = Var x) \ (\c. t = Fun c [])") case True have M_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and M_var_inst_cl: "is_TComp_var_instance_closed \ M" and M_subterms_cl: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" and M_Ana_cl: "has_all_wt_instances_of \ (\((set \ fst \ Ana) ` M)) M" using finite_SMP_representationD[OF M_SMP_repr] by blast+ have M_Ana_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (\ ((set \ fst \ Ana) ` M))" proof fix k assume "k \ \((set \ fst \ Ana) ` M)" then obtain m where m: "m \ M" "k \ set (fst (Ana m))" by force thus "wf\<^sub>t\<^sub>r\<^sub>m k" using M_wf Ana_keys_wf'[of m "fst (Ana m)" _ k] surjective_pairing by blast qed show ?thesis using True proof assume "\x. t = Var x" then obtain x y where x: "t = Var x" and y: "y \ fv\<^sub>s\<^sub>e\<^sub>t M" "\ (Var y) = \ (Var x)" using SMP_D_aux1[OF t_SMP M_subterms_cl M_var_inst_cl M_wf] by fastforce then obtain m \ where m: "m \ M" "m \ \ = Var y" and \: "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" using has_all_wt_instances_ofD'[OF wf_trms_subterms[OF M_wf] M_wf M_subterms_cl, of "Var y"] vars_iff_subtermeq_set[of y M] by fastforce obtain z where z: "m = Var z" using m(2) by (cases m) auto define \ where "\ \ Var(z := Var x)::('fun, ('fun, 'atom) term \ nat) subst" have "\ (Var z) = \ (Var x)" using y(2) m(2) z wt_subst_trm''[OF \, of m] by argo hence "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" unfolding \_def wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t_def by force+ moreover have "t = m \ \" using x z unfolding \_def by simp ultimately show ?thesis using m(1) by blast qed (use SMP_D[OF t_SMP _ M_SMP_repr] in blast) qed (use SMP_D[OF t_SMP _ M_SMP_repr] in blast) end lemma tfr\<^sub>s\<^sub>e\<^sub>t_if_comp_tfr\<^sub>s\<^sub>e\<^sub>t: assumes "comp_tfr\<^sub>s\<^sub>e\<^sub>t arity Ana \ M" shows "tfr\<^sub>s\<^sub>e\<^sub>t M" proof - let ?\ = "var_rename (max_var_set (fv\<^sub>s\<^sub>e\<^sub>t M))" have M_SMP_repr: "finite_SMP_representation arity Ana \ M" by (metis comp_tfr\<^sub>s\<^sub>e\<^sub>t_def assms) have M_finite: "finite M" using assms card_gt_0_iff unfolding comp_tfr\<^sub>s\<^sub>e\<^sub>t_def finite_SMP_representation_def by blast show ?thesis proof (unfold tfr\<^sub>s\<^sub>e\<^sub>t_def; intro ballI impI) fix s t assume "s \ SMP M - Var`\" "t \ SMP M - Var`\" hence st: "s \ SMP M" "is_Fun s" "t \ SMP M" "is_Fun t" by auto have "\(\\. Unifier \ s t)" when st_type_neq: "\ s \ \ t" proof (cases "\f. s = Fun f [] \ t = Fun f []") case False then obtain \ s0 \ t0 where s0: "s0 \ M" "is_Fun s0" "s = s0 \ \" "\ s = \ s0" and t0: "t0 \ M" "is_Fun t0" "t = t0 \ ?\ \ \" "\ t = \ t0" using SMP_D'[OF M_SMP_repr st(1,2) _ st(3,4)] by metis hence "\(\\. Unifier \ s0 (t0 \ ?\))" using assms mgu_None_is_subst_neq st_type_neq wt_subst_trm''[OF var_rename_wt(1)] unfolding comp_tfr\<^sub>s\<^sub>e\<^sub>t_def Let_def by metis thus ?thesis using vars_term_disjoint_imp_unifier[OF var_rename_fv_set_disjoint[OF M_finite]] s0(1) t0(1) unfolding s0(3) t0(3) by (metis (no_types, opaque_lifting) subst_subst_compose) qed (use st_type_neq st(2,4) in auto) thus "\ s = \ t" when "\\. Unifier \ s t" by (metis that) qed qed lemma tfr\<^sub>s\<^sub>e\<^sub>t_if_comp_tfr\<^sub>s\<^sub>e\<^sub>t': assumes "let N = SMP0 Ana \ M in set M \ set N \ comp_tfr\<^sub>s\<^sub>e\<^sub>t arity Ana \ (set N)" shows "tfr\<^sub>s\<^sub>e\<^sub>t (set M)" by (rule tfr_subset(2)[ OF tfr\<^sub>s\<^sub>e\<^sub>t_if_comp_tfr\<^sub>s\<^sub>e\<^sub>t[OF conjunct2[OF assms[unfolded Let_def]]] conjunct1[OF assms[unfolded Let_def]]]) lemma tfr\<^sub>s\<^sub>t\<^sub>p_is_comp_tfr\<^sub>s\<^sub>t\<^sub>p: "tfr\<^sub>s\<^sub>t\<^sub>p a = comp_tfr\<^sub>s\<^sub>t\<^sub>p \ a" proof (cases a) case (Equality ac t t') thus ?thesis using mgu_always_unifies[of t _ t'] mgu_gives_MGU[of t t'] by auto next case (Inequality X F) thus ?thesis using tfr\<^sub>s\<^sub>t\<^sub>p.simps(2)[of X F] comp_tfr\<^sub>s\<^sub>t\<^sub>p.simps(2)[of \ X F] Fun_range_case(2)[of "subterms\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>p\<^sub>a\<^sub>i\<^sub>r\<^sub>s F)"] unfolding is_Var_def by auto qed auto lemma tfr\<^sub>s\<^sub>t_if_comp_tfr\<^sub>s\<^sub>t: assumes "comp_tfr\<^sub>s\<^sub>t arity Ana \ M S" shows "tfr\<^sub>s\<^sub>t S" unfolding tfr\<^sub>s\<^sub>t_def proof have comp_tfr\<^sub>s\<^sub>e\<^sub>t_M: "comp_tfr\<^sub>s\<^sub>e\<^sub>t arity Ana \ M" using assms unfolding comp_tfr\<^sub>s\<^sub>t_def by blast have wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_M: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_S: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (trms\<^sub>s\<^sub>t S)" and S_trms_instance_M: "has_all_wt_instances_of \ (trms\<^sub>s\<^sub>t S) M" using assms wf\<^sub>t\<^sub>r\<^sub>m_code wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_code trms_list\<^sub>s\<^sub>t_is_trms\<^sub>s\<^sub>t unfolding comp_tfr\<^sub>s\<^sub>t_def comp_tfr\<^sub>s\<^sub>e\<^sub>t_def finite_SMP_representation_def list_all_iff by blast+ show "tfr\<^sub>s\<^sub>e\<^sub>t (trms\<^sub>s\<^sub>t S)" using tfr_subset(3)[OF tfr\<^sub>s\<^sub>e\<^sub>t_if_comp_tfr\<^sub>s\<^sub>e\<^sub>t[OF comp_tfr\<^sub>s\<^sub>e\<^sub>t_M] SMP_SMP_subset] SMP_I'[OF wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_S wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s_M S_trms_instance_M] by blast have "list_all (comp_tfr\<^sub>s\<^sub>t\<^sub>p \) S" by (metis assms comp_tfr\<^sub>s\<^sub>t_def) thus "list_all tfr\<^sub>s\<^sub>t\<^sub>p S" by (induct S) (simp_all add: tfr\<^sub>s\<^sub>t\<^sub>p_is_comp_tfr\<^sub>s\<^sub>t\<^sub>p) qed lemma tfr\<^sub>s\<^sub>t_if_comp_tfr\<^sub>s\<^sub>t': assumes "comp_tfr\<^sub>s\<^sub>t arity Ana \ (set (SMP0 Ana \ (trms_list\<^sub>s\<^sub>t S))) S" shows "tfr\<^sub>s\<^sub>t S" by (rule tfr\<^sub>s\<^sub>t_if_comp_tfr\<^sub>s\<^sub>t[OF assms]) subsubsection \Lemmata for Checking Ground SMP (GSMP) Disjointness\ context begin private lemma ground_SMP_disjointI_aux1: fixes M::"('fun, ('fun, 'atom) term \ nat) term set" assumes f_def: "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and g_def: "g \ \M. {t \ M. fv t = {}}" shows "f (SMP M) = g (SMP M)" proof have "t \ f (SMP M)" when t: "t \ SMP M" "fv t = {}" for t proof - define \ where "\ \ Var::('fun, ('fun, 'atom) term \ nat) subst" have "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" "t = t \ \" using subst_apply_term_empty[of t] that(2) wt_subst_Var wf_trm_subst_range_Var unfolding \_def by auto thus ?thesis using SMP.Substitution[OF t(1), of \] t(2) unfolding f_def by fastforce qed thus "g (SMP M) \ f (SMP M)" unfolding g_def by blast qed (use f_def g_def in blast) private lemma ground_SMP_disjointI_aux2: fixes M::"('fun, ('fun, 'atom) term \ nat) term set" assumes f_def: "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and M_SMP_repr: "finite_SMP_representation arity Ana \ M" shows "f M = f (SMP M)" proof have M_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s M" and M_var_inst_cl: "is_TComp_var_instance_closed \ M" and M_subterms_cl: "has_all_wt_instances_of \ (subterms\<^sub>s\<^sub>e\<^sub>t M) M" and M_Ana_cl: "has_all_wt_instances_of \ (\((set \ fst \ Ana) ` M)) M" using finite_SMP_representationD[OF M_SMP_repr] by blast+ show "f (SMP M) \ f M" proof fix t assume "t \ f (SMP M)" then obtain s \ where s: "t = s \ \" "s \ SMP M" "fv (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 have t_wf: "wf\<^sub>t\<^sub>r\<^sub>m t" using SMP_wf_trm[OF s(2) M_wf] s(1) wf_trm_subst[OF \(2)] by blast obtain m \ where m: "m \ M" "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 SMP_D''[OF s(2) M_SMP_repr] by blast have "t = m \ (\ \\<^sub>s \)" "fv (m \ (\ \\<^sub>s \)) = {}" using s(1,3) m(2) by simp_all thus "t \ f M" using m(1) wt_subst_compose[OF \(1) \(1)] wf_trms_subst_compose[OF \(2) \(2)] unfolding f_def by blast qed qed (auto simp add: f_def) private lemma ground_SMP_disjointI_aux3: fixes A B C::"('fun, ('fun, 'atom) term \ nat) term set" defines "P \ \t s. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ Unifier \ t s" assumes f_def: "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and Q_def: "Q \ \t. intruder_synth' public arity {} t" and R_def: "R \ \t. \u \ C. is_wt_instance_of_cond \ t u" and AB: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s A" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s B" "fv\<^sub>s\<^sub>e\<^sub>t A \ fv\<^sub>s\<^sub>e\<^sub>t B = {}" and C: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s C" and ABC: "\t \ A. \s \ B. P t s \ Q t \ R t" shows "f A \ f B \ f C \ {m. {} \\<^sub>c m}" proof fix t assume "t \ f A \ f B" then obtain ta tb \a \b where ta: "t = ta \ \a" "ta \ A" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \a" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \a)" "fv (ta \ \a) = {}" and tb: "t = tb \ \b" "tb \ B" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \b" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \b)" "fv (tb \ \b) = {}" unfolding f_def by blast have ta_tb_wf: "wf\<^sub>t\<^sub>r\<^sub>m ta" "wf\<^sub>t\<^sub>r\<^sub>m tb" "fv ta \ fv tb = {}" "\ ta = \ tb" using ta(1,2) tb(1,2) AB fv_subset_subterms wt_subst_trm''[OF ta(3), of ta] wt_subst_trm''[OF tb(3), of tb] by (fast, fast, blast, simp) obtain \ where \: "Unifier \ ta tb" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \)" using vars_term_disjoint_imp_unifier[OF ta_tb_wf(3), of \a \b] ta(1) tb(1) wt_Unifier_if_Unifier[OF ta_tb_wf(1,2,4)] by blast hence "Q ta \ R ta" using ABC ta(2) tb(2) unfolding P_def by blast+ thus "t \ f C \ {m. {} \\<^sub>c m}" proof show "Q ta \ ?thesis" using ta(1) pgwt_ground[of ta] pgwt_is_empty_synth[of ta] subst_ground_ident[of ta \a] unfolding Q_def f_def intruder_synth_code[symmetric] by simp next assume "R ta" then obtain ua \a where ua: "ta = ua \ \a" "ua \ C" "wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \a" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \a)" using \ ABC ta_tb_wf(1,2) ta(2) tb(2) C is_wt_instance_of_condD' unfolding P_def R_def by metis have "t = ua \ (\a \\<^sub>s \a)" "fv t = {}" using ua(1) ta(1,5) tb(1,5) by auto thus ?thesis using ua(2) wt_subst_compose[OF ua(3) ta(3)] wf_trms_subst_compose[OF ua(4) ta(4)] unfolding f_def by blast qed qed lemma ground_SMP_disjointI: fixes A B::"('fun, ('fun, 'atom) term \ nat) term set" and C defines "f \ \M. {t \ \ | t \. t \ M \ wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ fv (t \ \) = {}}" and "g \ \M. {t \ M. fv t = {}}" and "Q \ \t. intruder_synth' public arity {} t" and "R \ \t. \u \ C. is_wt_instance_of_cond \ t u" assumes AB_fv_disj: "fv\<^sub>s\<^sub>e\<^sub>t A \ fv\<^sub>s\<^sub>e\<^sub>t B = {}" and A_SMP_repr: "finite_SMP_representation arity Ana \ A" and B_SMP_repr: "finite_SMP_representation arity Ana \ B" and C_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s C" and ABC: "\t \ A. \s \ B. \ t = \ s \ mgu t s \ None \ Q t \ R t" shows "g (SMP A) \ g (SMP B) \ f C \ {m. {} \\<^sub>c m}" proof - have AB_wf: "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s A" "wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s B" using A_SMP_repr B_SMP_repr finite_SMP_representationD(1) by (blast, blast) let ?P = "\t s. \\. wt\<^sub>s\<^sub>u\<^sub>b\<^sub>s\<^sub>t \ \ wf\<^sub>t\<^sub>r\<^sub>m\<^sub>s (subst_range \) \ Unifier \ t s" have ABC': "\t \ A. \s \ B. ?P t s \ Q t \ R t" by (metis (no_types) ABC mgu_None_is_subst_neq wt_subst_trm'') show ?thesis using ground_SMP_disjointI_aux1[OF f_def g_def, of A] ground_SMP_disjointI_aux1[OF f_def g_def, of B] ground_SMP_disjointI_aux2[OF f_def A_SMP_repr] ground_SMP_disjointI_aux2[OF f_def B_SMP_repr] ground_SMP_disjointI_aux3[OF f_def Q_def R_def AB_wf AB_fv_disj C_wf ABC'] by argo qed end end end