diff --git a/thys/Q0_Metatheory/Consistency.thy b/thys/Q0_Metatheory/Consistency.thy --- a/thys/Q0_Metatheory/Consistency.thy +++ b/thys/Q0_Metatheory/Consistency.thy @@ -1,244 +1,244 @@ section \Consistency\ theory Consistency imports Soundness begin definition is_inconsistent_set :: "form set \ bool" where [iff]: "is_inconsistent_set \ \ \ \ F\<^bsub>o\<^esub>" definition \\<^sub>0_is_inconsistent :: bool where [iff]: "\\<^sub>0_is_inconsistent \ \ F\<^bsub>o\<^esub>" definition is_wffo_consistent_with :: "form \ form set \ bool" where [iff]: "is_wffo_consistent_with B \ \ \ is_inconsistent_set (\ \ {B})" subsection \Existence of a standard model\ text \We construct a standard model in which \\ i\ is the set \{0}\:\ primrec singleton_standard_domain_family ("\\<^sup>S") where "\\<^sup>S i = 1" \ \i.e., \<^term>\\\<^sup>S i = set {0}\\ | "\\<^sup>S o = \" | "\\<^sup>S (\\\) = \\<^sup>S \ \ \\<^sup>S \" interpretation singleton_standard_frame: frame "\\<^sup>S" proof unfold_locales { fix \ have "\\<^sup>S \ \ 0" proof (induction \) case (TFun \ \) from \\\<^sup>S \ \ 0\ obtain y where "y \ elts (\\<^sup>S \)" by fastforce then have "(\<^bold>\z \<^bold>: \\<^sup>S \\<^bold>. y) \ elts (\\<^sup>S \ \ \\<^sup>S \)" by (intro VPi_I) then show ?case by force qed simp_all } then show "\\. \\<^sup>S \ \ 0" by (intro allI) qed simp_all definition singleton_standard_constant_denotation_function ("\\<^sup>S") where [simp]: "\\<^sup>S k = ( if \\. is_Q_constant_of_type k \ then let \ = type_of_Q_constant k in q\<^bsub>\\<^esub>\<^bsup>\\<^sup>S\<^esup> else if is_iota_constant k then \<^bold>\z \<^bold>: \\<^sup>S (i\o)\<^bold>. 0 else case k of (c, \) \ SOME z. z \ elts (\\<^sup>S \) )" interpretation singleton_standard_premodel: premodel "\\<^sup>S" "\\<^sup>S" proof (unfold_locales) show "\\. \\<^sup>S (Q_constant_of_type \) = q\<^bsub>\\<^esub>\<^bsup>\\<^sup>S\<^esup>" by simp next show "singleton_standard_frame.is_unique_member_selector (\\<^sup>S iota_constant)" unfolding singleton_standard_frame.is_unique_member_selector_def proof fix x assume "x \ elts (\\<^sup>S i)" then have "x = 0" by simp moreover have "(\<^bold>\z \<^bold>: \\<^sup>S (i\o)\<^bold>. 0) \ {0}\<^bsub>i\<^esub>\<^bsup>\\<^sup>S\<^esup> = 0" using beta[OF singleton_standard_frame.one_element_function_is_domain_respecting] unfolding singleton_standard_domain_family.simps(3) by blast ultimately show "(\\<^sup>S iota_constant) \ {x}\<^bsub>i\<^esub>\<^bsup>\\<^sup>S\<^esup> = x" by fastforce qed next show "\c \. \ is_logical_constant (c, \) \ \\<^sup>S (c, \) \ elts (\\<^sup>S \)" proof (intro allI impI) fix c and \ assume "\ is_logical_constant (c, \)" then have "\\<^sup>S (c, \) = (SOME z. z \ elts (\\<^sup>S \))" by auto moreover have "\z. z \ elts (\\<^sup>S \)" using eq0_iff and singleton_standard_frame.domain_nonemptiness by presburger then have "(SOME z. z \ elts (\\<^sup>S \)) \ elts (\\<^sup>S \)" using some_in_eq by auto ultimately show "\\<^sup>S (c, \) \ elts (\\<^sup>S \)" by auto qed qed fun singleton_standard_wff_denotation_function ("\\<^sup>S") where "\\<^sup>S \ (x\<^bsub>\\<^esub>) = \ (x, \)" | "\\<^sup>S \ (\c\\<^bsub>\\<^esub>) = \\<^sup>S (c, \)" | "\\<^sup>S \ (A \ B) = (\\<^sup>S \ A) \ (\\<^sup>S \ B)" | "\\<^sup>S \ (\x\<^bsub>\\<^esub>. A) = (\<^bold>\z \<^bold>: \\<^sup>S \\<^bold>. \\<^sup>S (\((x, \) := z)) A)" lemma singleton_standard_wff_denotation_function_closure: assumes "frame.is_assignment \\<^sup>S \" and "A \ wffs\<^bsub>\\<^esub>" shows "\\<^sup>S \ A \ elts (\\<^sup>S \)" using assms(2,1) proof (induction A arbitrary: \) case (var_is_wff \ x) then show ?case by simp next case (con_is_wff \ c) then show ?case proof (cases "(c, \)" rule: constant_cases) case non_logical then show ?thesis using singleton_standard_premodel.non_logical_constant_denotation and singleton_standard_wff_denotation_function.simps(2) by presburger next case (Q_constant \) then have "\\<^sup>S \ (\c\\<^bsub>\\<^esub>) = q\<^bsub>\\<^esub>\<^bsup>\\<^sup>S\<^esup>" by simp moreover have "q\<^bsub>\\<^esub>\<^bsup>\\<^sup>S\<^esup> \ elts (\\<^sup>S (\\\\o))" using singleton_standard_domain_family.simps(3) and singleton_standard_frame.identity_relation_is_domain_respecting by presburger ultimately show ?thesis using Q_constant by simp next case \_constant then have "\\<^sup>S \ (\c\\<^bsub>\\<^esub>) = (\<^bold>\z \<^bold>: \\<^sup>S (i\o)\<^bold>. 0)" by simp moreover have "(\<^bold>\z \<^bold>: \\<^sup>S (i\o)\<^bold>. 0) \ elts (\\<^sup>S ((i\o)\i))" by (simp add: VPi_I) ultimately show ?thesis using \_constant by simp qed next case (app_is_wff \ \ A B) have "\\<^sup>S \ (A \ B) = (\\<^sup>S \ A) \ (\\<^sup>S \ B)" using singleton_standard_wff_denotation_function.simps(3) . moreover have "\\<^sup>S \ A \ elts (\\<^sup>S (\\\))" and "\\<^sup>S \ B \ elts (\\<^sup>S \)" using app_is_wff.IH and app_is_wff.prems by simp_all ultimately show ?case by (simp only: singleton_standard_frame.app_is_domain_respecting) next case (abs_is_wff \ A \ x) have "\\<^sup>S \ (\x\<^bsub>\\<^esub>. A) = (\<^bold>\z \<^bold>: \\<^sup>S \\<^bold>. \\<^sup>S (\((x, \) := z)) A)" using singleton_standard_wff_denotation_function.simps(4) . moreover have "\\<^sup>S (\((x, \) := z)) A \ elts (\\<^sup>S \)" if "z \ elts (\\<^sup>S \)" for z using that and abs_is_wff.IH and abs_is_wff.prems by simp ultimately show ?case by (simp add: VPi_I) qed interpretation singleton_standard_model: standard_model "\\<^sup>S" "\\<^sup>S" "\\<^sup>S" proof (unfold_locales) show "singleton_standard_premodel.is_wff_denotation_function \\<^sup>S" by (simp add: singleton_standard_wff_denotation_function_closure) next show "\\ \. \\<^sup>S (\\\) = \\<^sup>S \ \ \\<^sup>S \" using singleton_standard_domain_family.simps(3) by (intro allI) qed proposition standard_model_existence: shows "\\. is_standard_model \" using singleton_standard_model.standard_model_axioms by auto subsection \Theorem 5403 (Consistency Theorem)\ proposition model_existence_implies_set_consistency: assumes "is_hyps \" and "\\. is_general_model \ \ is_model_for \ \" shows "\ is_inconsistent_set \" proof (rule ccontr) from assms(2) obtain \ and \ and \ and \ where "\ = (\, \, \)" and "is_model_for \ \" and "is_general_model \" by fastforce assume "\ \ is_inconsistent_set \" then have "\ \ F\<^bsub>o\<^esub>" by simp with \is_general_model \\ have "\ \ F\<^bsub>o\<^esub>" using thm_5402(2)[OF assms(1) \is_model_for \ \\] by simp then have "\ \ F\<^bsub>o\<^esub> = \<^bold>T" if "\ \ \" for \ using that and \\ = (\, \, \)\ by force moreover have "\ \ F\<^bsub>o\<^esub> = \<^bold>F" if "\ \ \" for \ using \\ = (\, \, \)\ and \is_general_model \\ and that and general_model.prop_5401_d by simp ultimately have "\\. \ \ \" by (auto simp add: inj_eq) moreover have "\\. \ \ \" proof - \ \ Since by definition domains are not empty then, by using the Axiom of Choice, we can specify an assignment \\\ that simply chooses some element in the respective domain for each variable. - Nonetheless, as pointed out in Footnote 11, page 19 in @{cite "andrews:1965"}, it is not + Nonetheless, as pointed out in Footnote 11, page 19 in \<^cite>\"andrews:1965"\, it is not necessary to use the Axiom of Choice to show that assignments exist since some assignments can be described explicitly. \ let ?\ = "\v. case v of (_, \) \ SOME z. z \ elts (\ \)" from \\ = (\, \, \)\ and \is_general_model \\ have "\\. elts (\ \) \ {}" using frame.domain_nonemptiness and premodel_def and general_model.axioms(1) by auto with \\ = (\, \, \)\ and \is_general_model \\ have "?\ \ \" using frame.is_assignment_def and premodel_def and general_model.axioms(1) by (metis (mono_tags) case_prod_conv some_in_eq) then show ?thesis by (intro exI) qed ultimately show False .. qed proposition \\<^sub>0_is_consistent: shows "\ \\<^sub>0_is_inconsistent" proof - have "\\. is_general_model \ \ is_model_for \ {}" using standard_model_existence and standard_model.axioms(1) by blast then show ?thesis using model_existence_implies_set_consistency by simp qed lemmas thm_5403 = \\<^sub>0_is_consistent model_existence_implies_set_consistency proposition principle_of_explosion: assumes "is_hyps \" shows "is_inconsistent_set \ \ (\A \ (wffs\<^bsub>o\<^esub>). \ \ A)" proof assume "is_inconsistent_set \" show "\A \ (wffs\<^bsub>o\<^esub>). \ \ A" proof fix A assume "A \ wffs\<^bsub>o\<^esub>" from \is_inconsistent_set \\ have "\ \ F\<^bsub>o\<^esub>" unfolding is_inconsistent_set_def . then have "\ \ \\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>" unfolding false_is_forall . with \A \ wffs\<^bsub>o\<^esub>\ have "\ \ \<^bold>S {(\, o) \ A} (\\<^bsub>o\<^esub>)" using "\I" by fastforce then show "\ \ A" by simp qed next assume "\A \ (wffs\<^bsub>o\<^esub>). \ \ A" then have "\ \ F\<^bsub>o\<^esub>" using false_wff by (elim bspec) then show "is_inconsistent_set \" unfolding is_inconsistent_set_def . qed end diff --git a/thys/Q0_Metatheory/Elementary_Logic.thy b/thys/Q0_Metatheory/Elementary_Logic.thy --- a/thys/Q0_Metatheory/Elementary_Logic.thy +++ b/thys/Q0_Metatheory/Elementary_Logic.thy @@ -1,4821 +1,4821 @@ section \Elementary Logic\ theory Elementary_Logic imports Proof_System Propositional_Wff begin no_notation funcset (infixr "\" 60) notation funcset (infixr "\" 60) subsection \Proposition 5200\ proposition prop_5200: assumes "A \ wffs\<^bsub>\\<^esub>" shows "\ A =\<^bsub>\\<^esub> A" using assms and equality_reflexivity and dv_thm by simp corollary hyp_prop_5200: assumes "is_hyps \" and "A \ wffs\<^bsub>\\<^esub>" shows "\ \ A =\<^bsub>\\<^esub> A" using derivability_implies_hyp_derivability[OF prop_5200[OF assms(2)] assms(1)] . subsection \Proposition 5201 (Equality Rules)\ proposition prop_5201_1: assumes "\ \ A" and "\ \ A \\<^sup>\ B" shows "\ \ B" proof - from assms(2) have "\ \ A =\<^bsub>o\<^esub> B" unfolding equivalence_def . with assms(1) show ?thesis by (rule rule_R'[where p = "[]"]) auto qed proposition prop_5201_2: assumes "\ \ A =\<^bsub>\\<^esub> B" shows "\ \ B =\<^bsub>\\<^esub> A" proof - have "\ \ A =\<^bsub>\\<^esub> A" proof (rule hyp_prop_5200) from assms show "is_hyps \" by (blast elim: is_derivable_from_hyps.cases) show "A \ wffs\<^bsub>\\<^esub>" by (fact hyp_derivable_form_is_wffso[OF assms, THEN wffs_from_equality(1)]) qed from this and assms show ?thesis by (rule rule_R'[where p = "[\,\]"]) (force+, fastforce dest: subforms_from_app) qed proposition prop_5201_3: assumes "\ \ A =\<^bsub>\\<^esub> B" and "\ \ B =\<^bsub>\\<^esub> C" shows "\ \ A =\<^bsub>\\<^esub> C" using assms by (rule rule_R'[where p = "[\]"]) (force+, fastforce dest: subforms_from_app) proposition prop_5201_4: assumes "\ \ A =\<^bsub>\\\\<^esub> B" and "\ \ C =\<^bsub>\\<^esub> D" shows "\ \ A \ C =\<^bsub>\\<^esub> B \ D" proof - have "\ \ A \ C =\<^bsub>\\<^esub> A \ C" proof (rule hyp_prop_5200) from assms show "is_hyps \" by (blast elim: is_derivable_from_hyps.cases) from assms have "A \ wffs\<^bsub>\\\\<^esub>" and "C \ wffs\<^bsub>\\<^esub>" using hyp_derivable_form_is_wffso and wffs_from_equality by blast+ then show "A \ C \ wffs\<^bsub>\\<^esub>" by auto qed from this and assms(1) have "\ \ A \ C =\<^bsub>\\<^esub> B \ C" by (rule rule_R'[where p = "[\,\]"]) (force+, fastforce dest: subforms_from_app) from this and assms(2) show ?thesis by (rule rule_R'[where p = "[\,\]"]) (force+, fastforce dest: subforms_from_app) qed proposition prop_5201_5: assumes "\ \ A =\<^bsub>\\\\<^esub> B" and "C \ wffs\<^bsub>\\<^esub>" shows "\ \ A \ C =\<^bsub>\\<^esub> B \ C" proof - have "\ \ A \ C =\<^bsub>\\<^esub> A \ C" proof (rule hyp_prop_5200) from assms(1) show "is_hyps \" by (blast elim: is_derivable_from_hyps.cases) have "A \ wffs\<^bsub>\\\\<^esub>" by (fact hyp_derivable_form_is_wffso[OF assms(1), THEN wffs_from_equality(1)]) with assms(2) show "A \ C \ wffs\<^bsub>\\<^esub>" by auto qed from this and assms(1) show ?thesis by (rule rule_R'[where p = "[\,\]"]) (force+, fastforce dest: subforms_from_app) qed proposition prop_5201_6: assumes "\ \ C =\<^bsub>\\<^esub> D" and "A \ wffs\<^bsub>\\\\<^esub>" shows "\ \ A \ C =\<^bsub>\\<^esub> A \ D" proof - have "\ \ A \ C =\<^bsub>\\<^esub> A \ C" proof (rule hyp_prop_5200) from assms(1) show "is_hyps \" by (blast elim: is_derivable_from_hyps.cases) have "C \ wffs\<^bsub>\\<^esub>" by (fact hyp_derivable_form_is_wffso[OF assms(1), THEN wffs_from_equality(1)]) with assms(2) show "A \ C \ wffs\<^bsub>\\<^esub>" by auto qed from this and assms(1) show ?thesis by (rule rule_R'[where p = "[\,\]"]) (force+, fastforce dest: subforms_from_app) qed lemmas Equality_Rules = prop_5201_1 prop_5201_2 prop_5201_3 prop_5201_4 prop_5201_5 prop_5201_6 subsection \Proposition 5202 (Rule RR)\ proposition prop_5202: assumes "\ A =\<^bsub>\\<^esub> B \ \ B =\<^bsub>\\<^esub> A" and "p \ positions C" and "A \\<^bsub>p\<^esub> C" and "C\p \ B\ \ D" and "\ \ C" shows "\ \ D" proof - from assms(5) have "\ C =\<^bsub>o\<^esub> C" using prop_5200 and hyp_derivable_form_is_wffso by blast moreover from assms(1) consider (a) "\ A =\<^bsub>\\<^esub> B" | (b) "\ B =\<^bsub>\\<^esub> A" by blast then have "\ A =\<^bsub>\\<^esub> B" by cases (assumption, fact Equality_Rules(2)) ultimately have "\ C =\<^bsub>o\<^esub> D" by (rule rule_R[where p = "\ # p"]) (use assms(2-4) in auto) then have "\ \ C =\<^bsub>o\<^esub> D" proof - from assms(5) have "is_hyps \" by (blast elim: is_derivable_from_hyps.cases) with \\ C =\<^bsub>o\<^esub> D\ show ?thesis by (fact derivability_implies_hyp_derivability) qed with assms(5) show ?thesis by (rule Equality_Rules(1)[unfolded equivalence_def]) qed lemmas rule_RR = prop_5202 (* synonym *) subsection \Proposition 5203\ proposition prop_5203: assumes "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" and "\v \ vars A. \ is_bound v B" shows "\ (\x\<^bsub>\\<^esub>. B) \ A =\<^bsub>\\<^esub> \<^bold>S {(x, \) \ A} B" using assms(2,1,3) proof induction case (var_is_wff \ y) then show ?case proof (cases "y\<^bsub>\\<^esub> = x\<^bsub>\\<^esub>") case True then have "\ = \" by simp moreover from assms(1) have "\ (\x\<^bsub>\\<^esub>. x\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> A" using axiom_4_2 by (intro axiom_is_derivable_from_no_hyps) moreover have "\<^bold>S {(x, \) \ A} (x\<^bsub>\\<^esub>) = A" by force ultimately show ?thesis using True by (simp only:) next case False with assms(1) have "\ (\x\<^bsub>\\<^esub>. y\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> y\<^bsub>\\<^esub>" using axiom_4_1_var by (intro axiom_is_derivable_from_no_hyps) moreover from False have "\<^bold>S {(x, \) \ A} (y\<^bsub>\\<^esub>) = y\<^bsub>\\<^esub>" by auto ultimately show ?thesis by (simp only:) qed next case (con_is_wff \ c) from assms(1) have "\ (\x\<^bsub>\\<^esub>. \c\\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> \c\\<^bsub>\\<^esub>" using axiom_4_1_con by (intro axiom_is_derivable_from_no_hyps) moreover have "\<^bold>S {(x, \) \ A} (\c\\<^bsub>\\<^esub>) = \c\\<^bsub>\\<^esub>" by auto ultimately show ?case by (simp only:) next case (app_is_wff \ \ D C) from app_is_wff.prems(2) have not_bound_subforms: "\v \ vars A. \ is_bound v D \ \ is_bound v C" using is_bound_in_app_homomorphism by fast from \D \ wffs\<^bsub>\\\\<^esub>\ have "\ (\x\<^bsub>\\<^esub>. D) \ A =\<^bsub>\\\\<^esub> \<^bold>S {(x, \) \ A} D" using app_is_wff.IH(1)[OF assms(1)] and not_bound_subforms by simp moreover from \C \ wffs\<^bsub>\\<^esub>\ have "\ (\x\<^bsub>\\<^esub>. C) \ A =\<^bsub>\\<^esub> \<^bold>S {(x, \) \ A} C" using app_is_wff.IH(2)[OF assms(1)] and not_bound_subforms by simp moreover have "\ (\x\<^bsub>\\<^esub>. D \ C) \ A =\<^bsub>\\<^esub> ((\x\<^bsub>\\<^esub>. D) \ A) \ ((\x\<^bsub>\\<^esub>. C) \ A)" using axiom_is_derivable_from_no_hyps[OF axiom_4_3[OF assms(1) \D \ wffs\<^bsub>\\\\<^esub>\ \C \ wffs\<^bsub>\\<^esub>\]] . ultimately show ?case using Equality_Rules(3,4) and substitute.simps(3) by presburger next case (abs_is_wff \ D \ y) then show ?case proof (cases "y\<^bsub>\\<^esub> = x\<^bsub>\\<^esub>") case True then have "\ (\x\<^bsub>\\<^esub>. \y\<^bsub>\\<^esub>. D) \ A =\<^bsub>\\\\<^esub> \y\<^bsub>\\<^esub>. D" using axiom_is_derivable_from_no_hyps[OF axiom_4_5[OF assms(1) abs_is_wff.hyps(1)]] by fast moreover from True have "\<^bold>S {(x, \) \ A} (\y\<^bsub>\\<^esub>. D) = \y\<^bsub>\\<^esub>. D" using empty_substitution_neutrality by (simp add: singleton_substitution_simps(4) fmdrop_fmupd_same) ultimately show ?thesis by (simp only:) next case False have "binders_at (\y\<^bsub>\\<^esub>. D) [\] = {(y, \)}" by simp then have "is_bound (y, \) (\y\<^bsub>\\<^esub>. D)" by fastforce with abs_is_wff.prems(2) have "(y, \) \ vars A" by blast with \y\<^bsub>\\<^esub> \ x\<^bsub>\\<^esub>\ have "\ (\x\<^bsub>\\<^esub>. \y\<^bsub>\\<^esub>. D) \ A =\<^bsub>\\\\<^esub> \y\<^bsub>\\<^esub>. (\x\<^bsub>\\<^esub>. D) \ A" using axiom_4_4[OF assms(1) abs_is_wff.hyps(1)] and axiom_is_derivable_from_no_hyps by blast moreover have "\ (\x\<^bsub>\\<^esub>. D) \ A =\<^bsub>\\<^esub> \<^bold>S {(x, \) \ A} D" proof - have "\p. y\<^bsub>\\<^esub> \\<^bsub>\ # p\<^esub> \y\<^bsub>\\<^esub>. D \ y\<^bsub>\\<^esub> \\<^bsub>p\<^esub> D" using subforms_from_abs by fastforce from abs_is_wff.prems(2) have "\v \ vars A. \ is_bound v D" using is_bound_in_abs_body by fast then show ?thesis by (fact abs_is_wff.IH[OF assms(1)]) qed ultimately have "\ (\x\<^bsub>\\<^esub>. \y\<^bsub>\\<^esub>. D) \ A =\<^bsub>\\\\<^esub> \y\<^bsub>\\<^esub>. \<^bold>S {(x, \) \ A} D" by (rule rule_R[where p = "[\,\]"]) force+ with False show ?thesis by simp qed qed subsection \Proposition 5204\ proposition prop_5204: assumes "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" and "C \ wffs\<^bsub>\\<^esub>" and "\ B =\<^bsub>\\<^esub> C" and "\v \ vars A. \ is_bound v B \ \ is_bound v C" shows "\ \<^bold>S {(x, \) \ A} (B =\<^bsub>\\<^esub> C)" proof - have "\ (\x\<^bsub>\\<^esub>. B) \ A =\<^bsub>\\<^esub> (\x\<^bsub>\\<^esub>. B) \ A" proof - have "(\x\<^bsub>\\<^esub>. B) \ A \ wffs\<^bsub>\\<^esub>" using assms(1,2) by auto then show ?thesis by (fact prop_5200) qed from this and assms(4) have "\ (\x\<^bsub>\\<^esub>. B) \ A =\<^bsub>\\<^esub> (\x\<^bsub>\\<^esub>. C) \ A" by (rule rule_R[where p = "[\,\,\]"]) force+ moreover from assms(1,2,5) have "\ (\x\<^bsub>\\<^esub>. B) \ A =\<^bsub>\\<^esub> \<^bold>S {(x, \) \ A} B" using prop_5203 by auto moreover from assms(1,3,5) have "\ (\x\<^bsub>\\<^esub>. C) \ A =\<^bsub>\\<^esub> \<^bold>S {(x, \) \ A} C" using prop_5203 by auto ultimately have "\ (\<^bold>S {(x, \) \ A} B) =\<^bsub>\\<^esub> (\<^bold>S {(x, \) \ A} C)" using Equality_Rules(2,3) by blast then show ?thesis by simp qed subsection \Proposition 5205 ($\eta$-conversion)\ proposition prop_5205: shows "\ \\<^bsub>\\\\<^esub> =\<^bsub>\\\\<^esub> (\y\<^bsub>\\<^esub>. \\<^bsub>\\\\<^esub> \ y\<^bsub>\\<^esub>)" proof - { fix y assume "y\<^bsub>\\<^esub> \ \\<^bsub>\\<^esub>" let ?A = "\y\<^bsub>\\<^esub>. \\<^bsub>\\\\<^esub> \ y\<^bsub>\\<^esub>" have "\ (\\<^bsub>\\\\<^esub> =\<^bsub>\\\\<^esub> ?A) =\<^bsub>o\<^esub> \\\<^bsub>\\<^esub>. (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> ?A \ \\<^bsub>\\<^esub>)" proof - have "\ (\\<^bsub>\\\\<^esub> =\<^bsub>\\\\<^esub> \\<^bsub>\\\\<^esub>) =\<^bsub>o\<^esub> \\\<^bsub>\\<^esub>. (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>)" (is "\ ?B =\<^bsub>o\<^esub> ?C") using axiom_3[unfolded equivalence_def] by (rule axiom_is_derivable_from_no_hyps) have "\ \<^bold>S {(\, \\\) \ ?A} (?B =\<^bsub>o\<^esub> ?C)" proof - have "?A \ wffs\<^bsub>\\\\<^esub>" and "?B \ wffs\<^bsub>o\<^esub>" and "?C \ wffs\<^bsub>o\<^esub>" by auto moreover have "\v \ vars ?A. \ is_bound v ?B \ \ is_bound v ?C" proof fix v assume "v \ vars ?A" have "vars ?B = {(\, \\\), (\, \\\)}" and "vars ?C = {(\, \\\), (\, \), (\, \\\)}" by force+ with \y\<^bsub>\\<^esub> \ \\<^bsub>\\<^esub>\ have "(y, \) \ vars ?B" and "(y, \) \ vars ?C" by force+ then have "\ is_bound (y, \) ?B" and "\ is_bound (y, \) ?C" using absent_var_is_not_bound by blast+ moreover have "\ is_bound (\, \\\) ?B" and "\ is_bound (\, \\\) ?C" by code_simp+ moreover from \v \ vars ?A\ have "v \ {(y, \), (\, \\\)}" by auto ultimately show "\ is_bound v ?B \ \ is_bound v ?C" by fast qed ultimately show ?thesis using \\ ?B =\<^bsub>o\<^esub> ?C\ and prop_5204 by presburger qed then show ?thesis by simp qed moreover have "\ ?A \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>" proof - have "\\<^bsub>\\<^esub> \ wffs\<^bsub>\\<^esub>" and "\\<^bsub>\\\\<^esub> \ y\<^bsub>\\<^esub> \ wffs\<^bsub>\\<^esub>" by auto moreover have "\v \ vars (\\<^bsub>\\<^esub>). \ is_bound v (\\<^bsub>\\\\<^esub> \ y\<^bsub>\\<^esub>)" using \y\<^bsub>\\<^esub> \ \\<^bsub>\\<^esub>\ by auto moreover have "\<^bold>S {(y, \) \ \\<^bsub>\\<^esub>} (\\<^bsub>\\\\<^esub> \ y\<^bsub>\\<^esub>) = \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>" by simp ultimately show ?thesis using prop_5203 by metis qed ultimately have "\ (\\<^bsub>\\\\<^esub> =\<^bsub>\\\\<^esub> ?A) =\<^bsub>o\<^esub> \\\<^bsub>\\<^esub>. (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>)" by (rule rule_R[where p = "[\,\,\,\]"]) force+ moreover have "\ (\\<^bsub>\\\\<^esub> =\<^bsub>\\\\<^esub> \\<^bsub>\\\\<^esub>) =\<^bsub>o\<^esub> \\\<^bsub>\\<^esub>. (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>)" proof - let ?A = "\\<^bsub>\\\\<^esub>" have "\ (\\<^bsub>\\\\<^esub> =\<^bsub>\\\\<^esub> \\<^bsub>\\\\<^esub>) =\<^bsub>o\<^esub> \\\<^bsub>\\<^esub>. (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>)" (is "\ ?B =\<^bsub>o\<^esub> ?C") using axiom_3[unfolded equivalence_def] by (rule axiom_is_derivable_from_no_hyps) have "\ \<^bold>S {(\, \\\) \ ?A} (?B =\<^bsub>o\<^esub> ?C)" proof - have "?A \ wffs\<^bsub>\\\\<^esub>" and "?B \ wffs\<^bsub>o\<^esub>" and "?C \ wffs\<^bsub>o\<^esub>" by auto moreover have "\v \ vars ?A. \ is_bound v ?B \ \ is_bound v ?C" proof fix v assume "v \ vars ?A" have "vars ?B = {(\, \\\), (\, \\\)}" and "vars ?C = {(\, \\\), (\, \), (\, \\\)}" by force+ with \y\<^bsub>\\<^esub> \ \\<^bsub>\\<^esub>\ have "(y, \) \ vars ?B" and "(y, \) \ vars ?C" by force+ then have "\ is_bound (y, \) ?B" and "\ is_bound (y, \) ?C" using absent_var_is_not_bound by blast+ moreover have "\ is_bound (\, \\\) ?B" and "\ is_bound (\, \\\) ?C" by code_simp+ moreover from \v \ vars ?A \have "v \ {(y, \), (\, \\\)}" by auto ultimately show "\ is_bound v ?B \ \ is_bound v ?C" by fast qed ultimately show ?thesis using \\ ?B =\<^bsub>o\<^esub> ?C\ and prop_5204 by presburger qed then show ?thesis by simp qed ultimately have "\ \\<^bsub>\\\\<^esub> =\<^bsub>\\\\<^esub> (\y\<^bsub>\\<^esub>. \\<^bsub>\\\\<^esub> \ y\<^bsub>\\<^esub>)" using Equality_Rules(1)[unfolded equivalence_def] and Equality_Rules(2) and prop_5200 by (metis wffs_of_type_intros(1)) } note x_neq_y = this then have "\
6": "\ \\<^bsub>\\\\<^esub> =\<^bsub>\\\\<^esub> \\\<^bsub>\\<^esub>. \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>" (is "\ ?B =\<^bsub>_\<^esub> ?C") by simp then have "\
7": "\ (\\\<^bsub>\\<^esub>. \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>) =\<^bsub>\\\\<^esub> (\\\<^bsub>\\<^esub>. (\\\<^bsub>\\<^esub>. \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>) \ \\<^bsub>\\<^esub>)" proof - let ?A = "\\\<^bsub>\\<^esub>. \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>" have "?A \ wffs\<^bsub>\\\\<^esub>" and "?B \ wffs\<^bsub>\\\\<^esub>" and "?C \ wffs\<^bsub>\\\\<^esub>" by auto moreover have "\v \ vars ?A. \ is_bound v ?B \ \ is_bound v ?C" proof fix v assume "v \ vars ?A" have "\ is_bound (\, \) ?B" and "\ is_bound (\, \) ?C" by code_simp+ moreover have "\ is_bound (\, \\\) ?B" and "\ is_bound (\, \\\) ?C" by code_simp+ moreover from \v \ vars ?A \have "v \ {(\, \), (\, \\\)}" by auto ultimately show "\ is_bound v ?B \ \ is_bound v ?C" by fast qed ultimately have "\ \<^bold>S {(\, \\\) \ ?A} (?B =\<^bsub>\\\\<^esub> ?C)" using "\
6" and prop_5204 by presburger then show ?thesis by simp qed have "\ (\\\<^bsub>\\<^esub>. \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>) =\<^bsub>\\\\<^esub> (\\\<^bsub>\\<^esub>. \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>)" proof - have "\ (\\\<^bsub>\\<^esub>. \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>) \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>" proof - have "\\<^bsub>\\<^esub> \ wffs\<^bsub>\\<^esub>" and "\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub> \ wffs\<^bsub>\\<^esub>" by auto moreover have "\v \ vars (\\<^bsub>\\<^esub>). \ is_bound v (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>)" by simp moreover have "\<^bold>S {(\, \) \ \\<^bsub>\\<^esub>} (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>) = \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>" by simp ultimately show ?thesis using prop_5203 by metis qed from "\
7" and this show ?thesis by (rule rule_R [where p = "[\,\]"]) force+ qed with "\
6" and x_neq_y[of y] show ?thesis using Equality_Rules(2,3) by blast qed subsection \Proposition 5206 ($\alpha$-conversion)\ proposition prop_5206: assumes "A \ wffs\<^bsub>\\<^esub>" and "(z, \) \ free_vars A" and "is_free_for (z\<^bsub>\\<^esub>) (x, \) A" shows "\ (\x\<^bsub>\\<^esub>. A) =\<^bsub>\\\\<^esub> (\z\<^bsub>\\<^esub>. \<^bold>S {(x, \) \ z\<^bsub>\\<^esub>} A)" proof - have "is_substitution {(x, \) \ z\<^bsub>\\<^esub>}" by auto from this and assms(1) have "\<^bold>S {(x, \) \ z\<^bsub>\\<^esub>} A \ wffs\<^bsub>\\<^esub>" by (fact substitution_preserves_typing) obtain y where "(y, \) \ {(x, \), (z, \)} \ vars A" proof - have "finite ({(x, \), (z, \)} \ vars A)" using vars_form_finiteness by blast with that show ?thesis using fresh_var_existence by metis qed then have "(y, \) \ (x, \)" and "(y, \) \ (z, \)" and "(y, \) \ vars A" and "(y, \) \ free_vars A" using free_vars_in_all_vars by auto have "\
1": "\ (\x\<^bsub>\\<^esub>. A) =\<^bsub>\\\\<^esub> (\y\<^bsub>\\<^esub>. (\x\<^bsub>\\<^esub>. A) \ y\<^bsub>\\<^esub>)" proof - let ?A = "\x\<^bsub>\\<^esub>. A" have *: "\ \\<^bsub>\\\\<^esub> =\<^bsub>\\\\<^esub> (\y\<^bsub>\\<^esub>. \\<^bsub>\\\\<^esub> \ y\<^bsub>\\<^esub>)" (is "\ ?B =\<^bsub>_\<^esub> ?C") by (fact prop_5205) moreover have "\ \<^bold>S {(\, \\\) \ ?A} (?B =\<^bsub>\\\\<^esub> ?C)" proof - from assms(1) have "?A \ wffs\<^bsub>\\\\<^esub>" and "?B \ wffs\<^bsub>\\\\<^esub>" and "?C \ wffs\<^bsub>\\\\<^esub>" by auto moreover have "\v \ vars ?A. \ is_bound v ?B \ \ is_bound v ?C" proof fix v assume "v \ vars ?A" then consider (a) "v = (x, \)" | (b) "v \ vars A" by fastforce then show "\ is_bound v ?B \ \ is_bound v ?C" proof cases case a then show ?thesis using \(y, \) \ (x, \)\ by force next case b then have "\ is_bound v ?B" by simp moreover have "\ is_bound v ?C" using b and \(y, \) \ vars A\ by code_simp force ultimately show ?thesis by blast qed qed ultimately show ?thesis using prop_5204 and * by presburger qed ultimately show ?thesis by simp qed then have "\
2": "\ (\x\<^bsub>\\<^esub>. A) =\<^bsub>\\\\<^esub> (\y\<^bsub>\\<^esub>. \<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} A)" proof - have "\ (\x\<^bsub>\\<^esub>. A) \ y\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} A" (is "\ (\x\<^bsub>\\<^esub>. ?B) \ ?A =\<^bsub>_\<^esub> _") proof - have "?A \ wffs\<^bsub>\\<^esub>" and "?B \ wffs\<^bsub>\\<^esub>" by blast fact moreover have "\v \ vars ?A. \ is_bound v ?B" using \(y, \) \ vars A\ and absent_var_is_not_bound by auto ultimately show ?thesis by (fact prop_5203) qed with "\
1" show ?thesis by (rule rule_R [where p = "[\,\]"]) force+ qed moreover have "\
3": "\ (\z\<^bsub>\\<^esub>. \<^bold>S {(x, \) \ z\<^bsub>\\<^esub>} A) =\<^bsub>\\\\<^esub> (\y\<^bsub>\\<^esub>. (\z\<^bsub>\\<^esub>. \<^bold>S {(x, \) \ z\<^bsub>\\<^esub>} A) \ y\<^bsub>\\<^esub>)" proof - let ?A = "\z\<^bsub>\\<^esub>. \<^bold>S {(x, \) \ z\<^bsub>\\<^esub>} A" have *: "\ \\<^bsub>\\\\<^esub> =\<^bsub>\\\\<^esub> (\y\<^bsub>\\<^esub>. \\<^bsub>\\\\<^esub> \ y\<^bsub>\\<^esub>)" (is "\ ?B =\<^bsub>_\<^esub> ?C") by (fact prop_5205) moreover have "\ \<^bold>S {(\, \\\) \ ?A} (?B =\<^bsub>\\\\<^esub> ?C)" proof - have "?A \ wffs\<^bsub>\\\\<^esub>" and "?B \ wffs\<^bsub>\\\\<^esub>" and "?C \ wffs\<^bsub>\\\\<^esub>" using \\<^bold>S {(x, \) \ z\<^bsub>\\<^esub>} A \ wffs\<^bsub>\\<^esub>\ by auto moreover have "\v \ vars ?A. \ is_bound v ?B \ \ is_bound v ?C" proof fix v assume "v \ vars ?A" then consider (a) "v = (z, \)" | (b) "v \ vars (\<^bold>S {(x, \) \ z\<^bsub>\\<^esub>} A)" by fastforce then show "\ is_bound v ?B \ \ is_bound v ?C" proof cases case a then show ?thesis using \(y, \) \ (z, \)\ by auto next case b then have "\ is_bound v ?B" by simp moreover from b and \(y, \) \ vars A\ and \(y, \) \ (z, \)\ have "v \ (y, \)" using renaming_substitution_minimal_change by blast then have "\ is_bound v ?C" by code_simp simp ultimately show ?thesis by blast qed qed ultimately show ?thesis using prop_5204 and * by presburger qed ultimately show ?thesis by simp qed then have "\
4": "\ (\z\<^bsub>\\<^esub>. \<^bold>S {(x, \) \ z\<^bsub>\\<^esub>} A) =\<^bsub>\\\\<^esub> (\y\<^bsub>\\<^esub>. \<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} A)" proof - have "\ (\z\<^bsub>\\<^esub>. \<^bold>S {(x, \) \ z\<^bsub>\\<^esub>} A) \ y\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} A" (is "\ (\z\<^bsub>\\<^esub>. ?B) \ ?A =\<^bsub>_\<^esub> _") proof - have "?A \ wffs\<^bsub>\\<^esub>" and "?B \ wffs\<^bsub>\\<^esub>" by blast fact moreover from \(y, \) \ vars A\ and \(y, \) \ (z, \)\ have "\v \ vars ?A. \ is_bound v ?B" using absent_var_is_not_bound and renaming_substitution_minimal_change by auto ultimately have "\ (\z\<^bsub>\\<^esub>. \<^bold>S {(x, \) \ z\<^bsub>\\<^esub>} A) \ y\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \<^bold>S {(z, \) \ y\<^bsub>\\<^esub>} \<^bold>S {(x, \) \ z\<^bsub>\\<^esub>} A" using prop_5203 by fast moreover have "\<^bold>S {(z, \) \ y\<^bsub>\\<^esub>} \<^bold>S {(x, \) \ z\<^bsub>\\<^esub>} A = \<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} A" by (fact renaming_substitution_composability[OF assms(2,3)]) ultimately show ?thesis by (simp only:) qed with "\
3" show ?thesis by (rule rule_R [where p = "[\,\]"]) auto qed ultimately show ?thesis using Equality_Rules(2,3) by blast qed lemmas "\" = prop_5206 (* synonym *) subsection \Proposition 5207 ($\beta$-conversion)\ context begin private lemma bound_var_renaming_equality: assumes "A \ wffs\<^bsub>\\<^esub>" and "z\<^bsub>\\<^esub> \ y\<^bsub>\\<^esub>" and "(z, \) \ vars A" shows "\ A =\<^bsub>\\<^esub> rename_bound_var (y, \) z A" using assms proof induction case (var_is_wff \ x) then show ?case using prop_5200 by force next case (con_is_wff \ c) then show ?case using prop_5200 by force next case (app_is_wff \ \ A B) then show ?case using Equality_Rules(4) by auto next case (abs_is_wff \ A \ x) then show ?case proof (cases "(y, \) = (x, \)") case True have "\ \y\<^bsub>\\<^esub>. A =\<^bsub>\\\\<^esub> \y\<^bsub>\\<^esub>. A" by (fact abs_is_wff.hyps[THEN prop_5200[OF wffs_of_type_intros(4)]]) moreover have "\ A =\<^bsub>\\<^esub> rename_bound_var (y, \) z A" using abs_is_wff.IH[OF assms(2)] and abs_is_wff.prems(2) by fastforce ultimately have "\ \y\<^bsub>\\<^esub>. A =\<^bsub>\\\\<^esub> \y\<^bsub>\\<^esub>. rename_bound_var (y, \) z A" by (rule rule_R[where p = "[\,\]"]) force+ moreover have " \ \y\<^bsub>\\<^esub>. rename_bound_var (y, \) z A =\<^bsub>\\\\<^esub> \z\<^bsub>\\<^esub>. \<^bold>S {(y, \) \ z\<^bsub>\\<^esub>} (rename_bound_var (y, \) z A)" proof - have "rename_bound_var (y, \) z A \ wffs\<^bsub>\\<^esub>" using hyp_derivable_form_is_wffso[OF \\ A =\<^bsub>\\<^esub> rename_bound_var (y, \) z A\] by (blast dest: wffs_from_equality) moreover from abs_is_wff.prems(2) have "(z, \) \ free_vars (rename_bound_var (y, \) z A)" using rename_bound_var_free_vars[OF abs_is_wff.hyps assms(2)] by simp moreover from abs_is_wff.prems(2) have "is_free_for (z\<^bsub>\\<^esub>) (y, \) (rename_bound_var (y, \) z A)" using is_free_for_in_rename_bound_var[OF abs_is_wff.hyps assms(2)] by simp ultimately show ?thesis using "\" by fast qed ultimately have "\ \y\<^bsub>\\<^esub>. A =\<^bsub>\\\\<^esub> \z\<^bsub>\\<^esub>. \<^bold>S {(y, \) \ z\<^bsub>\\<^esub>} (rename_bound_var (y, \) z A)" by (rule Equality_Rules(3)) then show ?thesis using True by auto next case False have "\ \x\<^bsub>\\<^esub>. A =\<^bsub>\\\\<^esub> \x\<^bsub>\\<^esub>. A" by (fact abs_is_wff.hyps[THEN prop_5200[OF wffs_of_type_intros(4)]]) moreover have "\ A =\<^bsub>\\<^esub> rename_bound_var (y, \) z A" using abs_is_wff.IH[OF assms(2)] and abs_is_wff.prems(2) by fastforce ultimately have "\ \x\<^bsub>\\<^esub>. A =\<^bsub>\\\\<^esub> \x\<^bsub>\\<^esub>. rename_bound_var (y, \) z A" by (rule rule_R[where p = "[\,\]"]) force+ then show ?thesis using False by auto qed qed proposition prop_5207: assumes "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" and "is_free_for A (x, \) B" shows "\ (\x\<^bsub>\\<^esub>. B) \ A =\<^bsub>\\<^esub> \<^bold>S {(x, \) \ A} B" using assms proof (induction "form_size B" arbitrary: B \ rule: less_induct) case less from less(3,1,2,4) show ?case proof (cases B rule: wffs_of_type_cases) case (var_is_wff y) then show ?thesis proof (cases "y\<^bsub>\\<^esub> = x\<^bsub>\\<^esub>") case True then have "\ = \" by simp moreover from assms(1) have "\ (\x\<^bsub>\\<^esub>. x\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> A" using axiom_4_2 by (intro axiom_is_derivable_from_no_hyps) moreover have "\<^bold>S {(x, \) \ A} (x\<^bsub>\\<^esub>) = A" by force ultimately show ?thesis unfolding True and var_is_wff by simp next case False with assms(1) have "\ (\x\<^bsub>\\<^esub>. y\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> y\<^bsub>\\<^esub>" using axiom_4_1_var by (intro axiom_is_derivable_from_no_hyps) moreover from False have "\<^bold>S {(x, \) \ A} (y\<^bsub>\\<^esub>) = y\<^bsub>\\<^esub>" by auto ultimately show ?thesis unfolding False and var_is_wff by simp qed next case (con_is_wff c) from assms(1) have "\ (\x\<^bsub>\\<^esub>. \c\\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> \c\\<^bsub>\\<^esub>" using axiom_4_1_con by (intro axiom_is_derivable_from_no_hyps) moreover have "\<^bold>S {(x, \) \ A} (\c\\<^bsub>\\<^esub>) = \c\\<^bsub>\\<^esub>" by auto ultimately show ?thesis by (simp only: con_is_wff) next case (app_is_wff \ D C) have "form_size D < form_size B" and "form_size C < form_size B" unfolding app_is_wff(1) by simp_all from less(4)[unfolded app_is_wff(1)] have "is_free_for A (x, \) D" and "is_free_for A (x, \) C" using is_free_for_from_app by iprover+ from \is_free_for A (x, \) D\ have "\ (\x\<^bsub>\\<^esub>. D) \ A =\<^bsub>\\\\<^esub> \<^bold>S {(x, \) \ A} D" by (fact less(1)[OF \form_size D < form_size B\ assms(1) app_is_wff(2)]) moreover from \is_free_for A (x, \) C\ have "\ (\x\<^bsub>\\<^esub>. C) \ A =\<^bsub>\\<^esub> \<^bold>S {(x, \) \ A} C" by (fact less(1)[OF \form_size C < form_size B\ assms(1) app_is_wff(3)]) moreover have "\ (\x\<^bsub>\\<^esub>. D \ C) \ A =\<^bsub>\\<^esub> ((\x\<^bsub>\\<^esub>. D) \ A) \ ((\x\<^bsub>\\<^esub>. C) \ A)" by (fact axiom_4_3[OF assms(1) app_is_wff(2,3), THEN axiom_is_derivable_from_no_hyps]) ultimately show ?thesis unfolding app_is_wff(1) using Equality_Rules(3,4) and substitute.simps(3) by presburger next case (abs_is_wff \ D \ y) then show ?thesis proof (cases "y\<^bsub>\\<^esub> = x\<^bsub>\\<^esub>") case True with abs_is_wff(1) have "\ (\x\<^bsub>\\<^esub>. \y\<^bsub>\\<^esub>. D) \ A =\<^bsub>\\<^esub> \y\<^bsub>\\<^esub>. D" using axiom_4_5[OF assms(1) abs_is_wff(3)] by (simp add: axiom_is_derivable_from_no_hyps) moreover have "\<^bold>S {(x, \) \ A} (\y\<^bsub>\\<^esub>. D) = \y\<^bsub>\\<^esub>. D" using True by (simp add: empty_substitution_neutrality fmdrop_fmupd_same) ultimately show ?thesis unfolding abs_is_wff(2) by (simp only:) next case False have "form_size D < form_size B" unfolding abs_is_wff(2) by simp have "is_free_for A (x, \) D" using is_free_for_from_abs[OF less(4)[unfolded abs_is_wff(2)]] and \y\<^bsub>\\<^esub> \ x\<^bsub>\\<^esub>\ by blast have "\ (\x\<^bsub>\\<^esub>. (\y\<^bsub>\\<^esub>. D)) \ A =\<^bsub>\\<^esub> \y\<^bsub>\\<^esub>. \<^bold>S {(x, \) \ A} D" proof (cases "(y, \) \ vars A") case True with \y\<^bsub>\\<^esub> \ x\<^bsub>\\<^esub>\ have "\ (\x\<^bsub>\\<^esub>. \y\<^bsub>\\<^esub>. D) \ A =\<^bsub>\\\\<^esub> \y\<^bsub>\\<^esub>. (\x\<^bsub>\\<^esub>. D) \ A" using axiom_4_4[OF assms(1) abs_is_wff(3)] and axiom_is_derivable_from_no_hyps by auto moreover have "\ (\x\<^bsub>\\<^esub>. D) \ A =\<^bsub>\\<^esub> \<^bold>S {(x, \) \ A} D" by ( fact less(1) [OF \form_size D < form_size B\ assms(1) \D \ wffs\<^bsub>\\<^esub>\ \is_free_for A (x, \) D\] ) ultimately show ?thesis unfolding abs_is_wff(1) by (rule rule_R[where p = "[\,\]"]) force+ next case False have "finite (vars {A, D})" using vars_form_finiteness and vars_form_set_finiteness by simp then obtain z where "(z, \) \ ({(x, \), (y, \)} \ vars {A, D})" using fresh_var_existence by (metis Un_insert_left finite.simps insert_is_Un) then have "z\<^bsub>\\<^esub> \ x\<^bsub>\\<^esub>" and "z\<^bsub>\\<^esub> \ y\<^bsub>\\<^esub>" and "(z, \) \ vars {A, D}" by simp_all then show ?thesis proof (cases "(x, \) \ free_vars D") case True define D' where "D' = \<^bold>S {(y, \) \ z\<^bsub>\\<^esub>} D" have "is_substitution {(y, \) \ z\<^bsub>\\<^esub>}" by auto with \D \ wffs\<^bsub>\\<^esub>\ and D'_def have "D' \ wffs\<^bsub>\\<^esub>" using substitution_preserves_typing by blast then have "\ (\x\<^bsub>\\<^esub>. \z\<^bsub>\\<^esub>. D') \ A =\<^bsub>\\\\<^esub> \z\<^bsub>\\<^esub>. (\x\<^bsub>\\<^esub>. D') \ A" using \z\<^bsub>\\<^esub> \ x\<^bsub>\\<^esub>\ and \(z, \) \ vars {A, D}\ and axiom_4_4[OF assms(1)] and axiom_is_derivable_from_no_hyps by auto moreover have "\
2": "\ (\x\<^bsub>\\<^esub>. D') \ A =\<^bsub>\\<^esub> D'" proof - have "form_size D' = form_size D" unfolding D'_def by (fact renaming_substitution_preserves_form_size) then have "form_size D' < form_size B" using \form_size D < form_size B\ by simp moreover from \z\<^bsub>\\<^esub> \ x\<^bsub>\\<^esub>\ have "is_free_for A (x, \) D'" unfolding D'_def and is_free_for_def using substitution_preserves_freeness[OF True] and is_free_at_in_free_vars by fast ultimately have "\ (\x\<^bsub>\\<^esub>. D') \ A =\<^bsub>\\<^esub> \<^bold>S {(x, \) \ A} D'" using less(1) and assms(1) and \D' \ wffs\<^bsub>\\<^esub>\ by simp moreover from \z\<^bsub>\\<^esub> \ x\<^bsub>\\<^esub>\ have "(x, \) \ free_vars D'" unfolding D'_def using substitution_preserves_freeness[OF True] by fast then have "\<^bold>S {(x, \) \ A} D' = D'" by (fact free_var_singleton_substitution_neutrality) ultimately show ?thesis by (simp only:) qed ultimately have "\
3": "\ (\x\<^bsub>\\<^esub>. \z\<^bsub>\\<^esub>. D') \ A =\<^bsub>\\\\<^esub> \z\<^bsub>\\<^esub>. D'" (is \\ ?A3\) by (rule rule_R[where p = "[\,\]"]) force+ moreover have "\
4": "\ (\y\<^bsub>\\<^esub>. D) =\<^bsub>\\\\<^esub> \z\<^bsub>\\<^esub>. D'" proof - have "(z, \) \ free_vars D" using \(z, \) \ vars {A, D}\ and free_vars_in_all_vars_set by auto moreover have "is_free_for (z\<^bsub>\\<^esub>) (y, \) D" using \(z, \) \ vars {A, D}\ and absent_var_is_free_for by force ultimately have "\ \y\<^bsub>\\<^esub>. D =\<^bsub>\\\\<^esub> \z\<^bsub>\\<^esub>. \<^bold>S {(y, \) \ z\<^bsub>\\<^esub>} D" using "\"[OF \D \ wffs\<^bsub>\\<^esub>\] by fast then show ?thesis using D'_def by blast qed ultimately have "\
5": "\ (\x\<^bsub>\\<^esub>. \y\<^bsub>\\<^esub>. D) \ A =\<^bsub>\\\\<^esub> \y\<^bsub>\\<^esub>. D" proof - note rule_RR' = rule_RR[OF disjI2] have "\
5\<^sub>1": "\ (\x\<^bsub>\\<^esub>. \y\<^bsub>\\<^esub>. D) \ A =\<^bsub>\\\\<^esub> \z\<^bsub>\\<^esub>. D'" (is \\ ?A5\<^sub>1\) by (rule rule_RR'[OF "\
4", where p = "[\,\,\,\]" and C = "?A3"]) (use "\
3" in \force+\) show ?thesis by (rule rule_RR'[OF "\
4", where p = "[\]" and C = "?A5\<^sub>1"]) (use "\
5\<^sub>1" in \force+\) qed then show ?thesis using free_var_singleton_substitution_neutrality[OF \(x, \) \ free_vars D\] by (simp only: \\ = \\\\) next case False have "(y, \) \ free_vars A" proof (rule ccontr) assume "\ (y, \) \ free_vars A" moreover from \\ (x, \) \ free_vars D\ obtain p where "p \ positions D" and "is_free_at (x, \) p D" using free_vars_in_is_free_at by blast then have "\ # p \ positions (\y\<^bsub>\\<^esub>. D)" and "is_free_at (x, \) (\ # p) (\y\<^bsub>\\<^esub>. D)" using is_free_at_to_abs[OF \is_free_at (x, \) p D\] and \y\<^bsub>\\<^esub> \ x\<^bsub>\\<^esub>\ by (simp, fast) moreover have "in_scope_of_abs (y, \) (\ # p) (\y\<^bsub>\\<^esub>. D)" by force ultimately have "\ is_free_for A (x, \) (\y\<^bsub>\\<^esub>. D)" by blast with \is_free_for A (x, \) B\[unfolded abs_is_wff(2)] show False by contradiction qed define A' where "A' = rename_bound_var (y, \) z A" have "A' \ wffs\<^bsub>\\<^esub>" unfolding A'_def by (fact rename_bound_var_preserves_typing[OF assms(1)]) from \(z, \) \ vars {A, D}\ have "(y, \) \ vars A'" using old_var_not_free_not_occurring_after_rename [ OF assms(1) \z\<^bsub>\\<^esub> \ y\<^bsub>\\<^esub>\ \(y, \) \ free_vars A\ ] unfolding A'_def by simp from A'_def have "\
6": "\ A =\<^bsub>\\<^esub> A'" using bound_var_renaming_equality[OF assms(1) \z\<^bsub>\\<^esub> \ y\<^bsub>\\<^esub>\] and \(z, \) \ vars {A, D}\ by simp moreover have "\
7": "\ (\x\<^bsub>\\<^esub>. \y\<^bsub>\\<^esub>. D) \ A' =\<^bsub>\\\\<^esub> \y\<^bsub>\\<^esub>. (\x\<^bsub>\\<^esub>. D) \ A'" (is \\ ?A7\) using axiom_4_4[OF \A' \ wffs\<^bsub>\\<^esub>\ \D \ wffs\<^bsub>\\<^esub>\] and \(y, \) \ vars A'\ and \y\<^bsub>\\<^esub> \ x\<^bsub>\\<^esub>\ and axiom_is_derivable_from_no_hyps by auto ultimately have "\
8": "\ (\x\<^bsub>\\<^esub>. \y\<^bsub>\\<^esub>. D) \ A =\<^bsub>\\\\<^esub> \y\<^bsub>\\<^esub>. (\x\<^bsub>\\<^esub>. D) \ A" proof - note rule_RR' = rule_RR[OF disjI2] have "\
8\<^sub>1": "\ (\x\<^bsub>\\<^esub>. \y\<^bsub>\\<^esub>. D) \ A =\<^bsub>\\\\<^esub> \y\<^bsub>\\<^esub>. (\x\<^bsub>\\<^esub>. D) \ A'" (is \\ ?A8\<^sub>1\) by (rule rule_RR'[OF "\
6", where p = "[\,\,\]" and C = "?A7"]) (use "\
7" in \force+\) show ?thesis by (rule rule_RR'[OF "\
6", where p = "[\,\,\]" and C = "?A8\<^sub>1"]) (use "\
8\<^sub>1" in \force+\) qed moreover have "form_size D < form_size B" unfolding abs_is_wff(2) by (simp only: form_size.simps(4) lessI) with assms(1) have "\
9": "\ (\x\<^bsub>\\<^esub>. D) \ A =\<^bsub>\\<^esub> \<^bold>S {(x, \) \ A} D" using less(1) and \D \ wffs\<^bsub>\\<^esub>\ and \is_free_for A (x, \) D\ by (simp only:) ultimately show ?thesis unfolding \\ = \\\\ by (rule rule_R[where p = "[\,\]"]) force+ qed qed then show ?thesis unfolding abs_is_wff(2) using False and singleton_substitution_simps(4) by simp qed qed qed end subsection \Proposition 5208\ proposition prop_5208: assumes "vs \ []" and "B \ wffs\<^bsub>\\<^esub>" shows "\ \\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs B) (map FVar vs) =\<^bsub>\\<^esub> B" using assms(1) proof (induction vs rule: list_nonempty_induct) case (single v) obtain x and \ where "v = (x, \)" by fastforce then have "\\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ [v] B) (map FVar [v]) = (\x\<^bsub>\\<^esub>. B) \ x\<^bsub>\\<^esub>" by simp moreover have "\ (\x\<^bsub>\\<^esub>. B) \ x\<^bsub>\\<^esub> =\<^bsub>\\<^esub> B" proof - have "is_free_for (x\<^bsub>\\<^esub>) (x, \) B" by fastforce then have "\ (\x\<^bsub>\\<^esub>. B) \ x\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \<^bold>S {(x, \) \ x\<^bsub>\\<^esub>} B" by (rule prop_5207 [OF wffs_of_type_intros(1) assms(2)]) then show ?thesis using identity_singleton_substitution_neutrality by (simp only:) qed ultimately show ?case by (simp only:) next case (cons v vs) obtain x and \ where "v = (x, \)" by fastforce have "\ \\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ (v # vs) B) (map FVar (v # vs)) =\<^bsub>\\<^esub> \\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs B) (map FVar vs)" proof - have "\\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ (v # vs) B) (map FVar (v # vs)) \ wffs\<^bsub>\\<^esub>" proof - have "\\<^sup>\\<^sub>\ (v # vs) B \ wffs\<^bsub>foldr (\) (map snd (v # vs)) \\<^esub>" using generalized_abs_wff [OF assms(2)] by blast moreover have "\k < length (map FVar (v # vs)). map FVar (v # vs) ! k \ wffs\<^bsub>map snd (v # vs) ! k\<^esub>" proof safe fix k assume *: "k < length (map FVar (v # vs))" moreover obtain x and \ where "(v # vs) ! k = (x, \)" by fastforce with * have "map FVar (v # vs) ! k = x\<^bsub>\\<^esub>" and "map snd (v # vs) ! k = \" by (metis length_map nth_map snd_conv)+ ultimately show "map FVar (v # vs) ! k \ wffs\<^bsub>map snd (v # vs) ! k\<^esub>" by fastforce qed ultimately show ?thesis using generalized_app_wff[where As = "map FVar (v # vs)" and ts = "map snd (v # vs)"] by simp qed then have " \ \\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ (v # vs) B) (map FVar (v # vs)) =\<^bsub>\\<^esub> \\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ (v # vs) B) (map FVar (v # vs))" by (fact prop_5200) then have " \ \\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ (v # vs) B) (map FVar (v # vs)) =\<^bsub>\\<^esub> \\<^sup>\\<^sub>\ ((\\<^sup>\\<^sub>\ (v # vs) B) \ FVar v) (map FVar vs)" by simp moreover have "\ (\\<^sup>\\<^sub>\ (v # vs) B) \ FVar v =\<^bsub>foldr (\) (map snd vs) \\<^esub> (\\<^sup>\\<^sub>\ vs B)" proof - have "\ (\\<^sup>\\<^sub>\ (v # vs) B) \ FVar v =\<^bsub>foldr (\) (map snd vs) \\<^esub> \<^bold>S {v \ FVar v} (\\<^sup>\\<^sub>\ vs B)" proof - from \v = (x, \)\ have "\\<^sup>\\<^sub>\ (v # vs) B = \x\<^bsub>\\<^esub>. \\<^sup>\\<^sub>\ vs B" by simp have "\\<^sup>\\<^sub>\ vs B \ wffs\<^bsub>foldr (\) (map snd vs) \\<^esub>" using generalized_abs_wff[OF assms(2)] by blast moreover have "is_free_for (x\<^bsub>\\<^esub>) (x, \) (\\<^sup>\\<^sub>\ vs B)" by fastforce ultimately have "\ (\x\<^bsub>\\<^esub>. \\<^sup>\\<^sub>\ vs B) \ x\<^bsub>\\<^esub> =\<^bsub>foldr (\) (map snd vs) \\<^esub> \<^bold>S {(x, \) \ x\<^bsub>\\<^esub>} \\<^sup>\\<^sub>\ vs B" by (rule prop_5207 [OF wffs_of_type_intros(1)]) with \v = (x, \)\ show ?thesis by simp qed then show ?thesis using identity_singleton_substitution_neutrality by (simp only:) qed ultimately show ?thesis proof (induction rule: rule_R [where p = "[\] @ replicate (length vs) \"]) case occ_subform then show ?case unfolding equality_of_type_def using leftmost_subform_in_generalized_app by (metis append_Cons append_Nil is_subform_at.simps(3) length_map) next case replacement then show ?case unfolding equality_of_type_def using leftmost_subform_in_generalized_app_replacement and is_subform_implies_in_positions and leftmost_subform_in_generalized_app by (metis append_Cons append_Nil length_map replace_right_app) qed qed moreover have "\ \\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs B) (map FVar vs) =\<^bsub>\\<^esub> B" by (fact cons.IH) ultimately show ?case by (rule rule_R [where p = "[\]"]) auto qed subsection \Proposition 5209\ proposition prop_5209: assumes "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" and "C \ wffs\<^bsub>\\<^esub>" and "\ B =\<^bsub>\\<^esub> C" and "is_free_for A (x, \) (B =\<^bsub>\\<^esub> C)" shows "\ \<^bold>S {(x, \) \ A} (B =\<^bsub>\\<^esub> C)" proof - have "\ (\x\<^bsub>\\<^esub>. B) \ A =\<^bsub>\\<^esub> (\x\<^bsub>\\<^esub>. B) \ A" proof - have "(\x\<^bsub>\\<^esub>. B) \ A \ wffs\<^bsub>\\<^esub>" using assms(1,2) by blast then show ?thesis by (fact prop_5200) qed from this and assms(4) have "\ (\x\<^bsub>\\<^esub>. B) \ A =\<^bsub>\\<^esub> (\x\<^bsub>\\<^esub>. C) \ A" by (rule rule_R [where p = "[\,\,\]"]) force+ moreover have "\ (\x\<^bsub>\\<^esub>. B) \ A =\<^bsub>\\<^esub> \<^bold>S {(x, \) \ A} B" proof - from assms(5)[unfolded equality_of_type_def] have "is_free_for A (x, \) (Q\<^bsub>\\<^esub> \ B)" by (rule is_free_for_from_app) then have "is_free_for A (x, \) B" by (rule is_free_for_from_app) with assms(1,2) show ?thesis by (rule prop_5207) qed moreover have "\ (\x\<^bsub>\\<^esub>. C) \ A =\<^bsub>\\<^esub> \<^bold>S {(x, \) \ A} C" proof - from assms(5)[unfolded equality_of_type_def] have "is_free_for A (x, \) C" by (rule is_free_for_from_app) with assms(1,3) show ?thesis by (rule prop_5207) qed ultimately have "\ (\<^bold>S {(x, \) \ A} B) =\<^bsub>\\<^esub> (\<^bold>S {(x, \) \ A} C)" using Equality_Rules(2,3) by blast then show ?thesis by simp qed subsection \Proposition 5210\ proposition prop_5210: assumes "B \ wffs\<^bsub>\\<^esub>" shows "\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> (B =\<^bsub>\\<^esub> B)" proof - have "\
1": " \ ((\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>) =\<^bsub>\\\\<^esub> (\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>)) =\<^bsub>o\<^esub> \\\<^bsub>\\<^esub>. ((\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>) \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> (\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>) \ \\<^bsub>\\<^esub>)" proof - have "\ (\\<^bsub>\\\\<^esub> =\<^bsub>\\\\<^esub> \\<^bsub>\\\\<^esub>) =\<^bsub>o\<^esub> \\\<^bsub>\\<^esub>. (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>)" (is "\ ?B =\<^bsub>o\<^esub> ?C") using axiom_3[unfolded equivalence_def] by (rule axiom_is_derivable_from_no_hyps) moreover have "(\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>) \ wffs\<^bsub>\\\\<^esub>" and "?B \ wffs\<^bsub>o\<^esub>" and "?C \ wffs\<^bsub>o\<^esub>" by auto moreover have "is_free_for (\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>) (\, \\\) (?B =\<^bsub>o\<^esub> ?C)" by simp ultimately have "\ \<^bold>S {(\, \\\) \ (\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>)} (?B =\<^bsub>o\<^esub> ?C)" (is "\ ?S") using prop_5209 by presburger moreover have "?S = ( (\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>) =\<^bsub>\\\\<^esub> \\<^bsub>\\\\<^esub>) =\<^bsub>o\<^esub> \\\<^bsub>\\<^esub>. ((\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>) \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub> )" (is "_ = ?B' =\<^bsub>o\<^esub> ?C'") by simp ultimately have "\ ?B' =\<^bsub>o\<^esub> ?C'" by (simp only:) moreover from \(\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>) \ wffs\<^bsub>\\\\<^esub>\ have "?B' \ wffs\<^bsub>o\<^esub>" and "?C' \ wffs\<^bsub>o\<^esub>" by auto moreover have "is_free_for (\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>) (\, \\\) (?B' =\<^bsub>o\<^esub> ?C')" by simp ultimately have "\ \<^bold>S {(\, \\\) \ (\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>)} (?B' =\<^bsub>o\<^esub> ?C')" (is "\ ?S'") using prop_5209[OF \(\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>) \ wffs\<^bsub>\\\\<^esub>\] by blast then show ?thesis by simp qed then have "\ (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) =\<^bsub>\\o\<^esub> (\\\<^bsub>\\<^esub>. (\\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\<^esub>))" proof - have "\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub> \ wffs\<^bsub>\\\\<^esub>" by blast then have "\ \\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub> =\<^bsub>\\\\<^esub> \\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>" by (fact prop_5200) with "\
1" have "\ \\\<^bsub>\\<^esub>. ((\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>) \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> (\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>) \ \\<^bsub>\\<^esub>)" using rule_R and is_subform_at.simps(1) by blast moreover have "\ (\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>) \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\<^esub>" using axiom_4_2[OF wffs_of_type_intros(1)] by (rule axiom_is_derivable_from_no_hyps) ultimately have "\ \\\<^bsub>\\<^esub>. (\\<^bsub>\\<^esub> =\<^bsub>\\<^esub> (\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>) \ \\<^bsub>\\<^esub>)" by (rule rule_R[where p = "[\,\,\,\]"]) auto from this and \\ (\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>) \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\<^esub>\ have "\ \\\<^bsub>\\<^esub>. (\\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\<^esub>)" by (rule rule_R[where p = "[\,\,\]"]) auto then show ?thesis unfolding forall_def and PI_def by (fold equality_of_type_def) qed from this and assms have 3: "\ (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) \ B =\<^bsub>o\<^esub> (\\\<^bsub>\\<^esub>. (\\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\<^esub>)) \ B" by (rule Equality_Rules(5)) then show ?thesis proof - have "\ (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) \ B =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using prop_5207[OF assms true_wff] by fastforce from 3 and this have "\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> (\\\<^bsub>\\<^esub>. (\\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\<^esub>)) \ B" by (rule rule_R[where p = "[\,\]"]) auto moreover have "\ (\\\<^bsub>\\<^esub>. (\\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\<^esub>)) \ B =\<^bsub>o\<^esub> (B =\<^bsub>\\<^esub> B)" proof - have "\\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\<^esub> \ wffs\<^bsub>o\<^esub>" and "is_free_for B (\, \) (\\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\<^esub>)" by (blast, intro is_free_for_in_equality is_free_for_in_var) moreover have "\<^bold>S {(\, \) \ B} (\\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\<^esub>) = (B =\<^bsub>\\<^esub> B)" by simp ultimately show ?thesis using prop_5207[OF assms] by metis qed ultimately show ?thesis by (rule rule_R [where p = "[\]"]) auto qed qed subsection \Proposition 5211\ proposition prop_5211: shows "\ (T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" proof - have const_T_wff: "(\x\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>) \ wffs\<^bsub>o\o\<^esub>" for x by blast have "\
1": "\ (\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>) \ T\<^bsub>o\<^esub> \\<^sup>\ (\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>) \ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. (\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>) \ \\<^bsub>o\<^esub>" proof - have "\ \\<^bsub>o\o\<^esub> \ T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\o\<^esub> \ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. \\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub>" (is "\ ?B =\<^bsub>o\<^esub> ?C") using axiom_1[unfolded equivalence_def] by (rule axiom_is_derivable_from_no_hyps) moreover have "?B \ wffs\<^bsub>o\<^esub>" and "?C \ wffs\<^bsub>o\<^esub>" by auto moreover have "is_free_for (\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>) (\, o\o) (?B =\<^bsub>o\<^esub> ?C)" by simp ultimately have "\ \<^bold>S {(\, o\o) \ (\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>)} (?B =\<^bsub>o\<^esub> ?C)" using const_T_wff and prop_5209 by presburger then show ?thesis by simp qed then have "\ T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>" proof - have T_\_redex: "\ (\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>) \ A =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" if "A \ wffs\<^bsub>o\<^esub>" for A using that and prop_5207[OF that true_wff] by fastforce from "\
1" and T_\_redex[OF true_wff] have "\ T\<^bsub>o\<^esub> \\<^sup>\ (\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>) \ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. (\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>) \ \\<^bsub>o\<^esub>" by (rule rule_R[where p = "[\,\,\,\]"]) force+ from this and T_\_redex[OF false_wff] have "\ T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. (\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>) \ \\<^bsub>o\<^esub>" by (rule rule_R[where p = "[\,\,\]"]) force+ from this and T_\_redex[OF wffs_of_type_intros(1)] show ?thesis by (rule rule_R[where p = "[\,\,\]"]) force+ qed moreover have "\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>" using prop_5210[OF const_T_wff] by simp ultimately show ?thesis using Equality_Rules(2,3) by blast qed lemma true_is_derivable: shows "\ T\<^bsub>o\<^esub>" unfolding true_def using Q_wff by (rule prop_5200) subsection \Proposition 5212\ proposition prop_5212: shows "\ T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>" proof - have "\ T\<^bsub>o\<^esub>" by (fact true_is_derivable) moreover have "\ (T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by (fact prop_5211) then have "\ T\<^bsub>o\<^esub> \\<^sup>\ (T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>)" unfolding equivalence_def by (fact Equality_Rules(2)) ultimately show ?thesis by (rule Equality_Rules(1)) qed subsection \Proposition 5213\ proposition prop_5213: assumes "\ A =\<^bsub>\\<^esub> B" and "\ C =\<^bsub>\\<^esub> D" shows "\ (A =\<^bsub>\\<^esub> B) \\<^sup>\ (C =\<^bsub>\\<^esub> D)" proof - from assms have "A \ wffs\<^bsub>\\<^esub>" and "C \ wffs\<^bsub>\\<^esub>" using hyp_derivable_form_is_wffso and wffs_from_equality by blast+ have "\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> (A =\<^bsub>\\<^esub> A)" by (fact prop_5210[OF \A \ wffs\<^bsub>\\<^esub>\]) moreover have "\ A =\<^bsub>\\<^esub> B" by fact ultimately have "\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> (A =\<^bsub>\\<^esub> B)" by (rule rule_R[where p = "[\,\]"]) force+ have "\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> (C =\<^bsub>\\<^esub> C)" by (fact prop_5210[OF \C \ wffs\<^bsub>\\<^esub>\]) moreover have "\ C =\<^bsub>\\<^esub> D" by fact ultimately have "\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> (C =\<^bsub>\\<^esub> D)" by (rule rule_R[where p = "[\,\]"]) force+ then show ?thesis proof - have "\ T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>" by (fact prop_5212) from this and \\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> (A =\<^bsub>\\<^esub> B)\ have "\ (A =\<^bsub>\\<^esub> B) \\<^sup>\ T\<^bsub>o\<^esub>" by (rule rule_R[where p = "[\,\]"]) force+ from this and \\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> (C =\<^bsub>\\<^esub> D)\ show ?thesis by (rule rule_R[where p = "[\]"]) force+ qed qed subsection \Proposition 5214\ proposition prop_5214: shows "\ T\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" proof - have id_on_o_is_wff: "(\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>) \ wffs\<^bsub>o\o\<^esub>" by blast have "\
1": "\ (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>) \ T\<^bsub>o\<^esub> \\<^sup>\ (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>) \ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>) \ \\<^bsub>o\<^esub>" proof - have "\ \\<^bsub>o\o\<^esub> \ T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\o\<^esub> \ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. \\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub>" (is "\ ?B =\<^bsub>o\<^esub> ?C") using axiom_1[unfolded equivalence_def] by (rule axiom_is_derivable_from_no_hyps) moreover have "?B \ wffs\<^bsub>o\<^esub>" and "?C \ wffs\<^bsub>o\<^esub>" and "is_free_for (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>) (\, o\o) (?B =\<^bsub>o\<^esub> ?C)" by auto ultimately have "\ \<^bold>S {(\, o\o) \ (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>)} (?B =\<^bsub>o\<^esub> ?C)" using id_on_o_is_wff and prop_5209 by presburger then show ?thesis by simp qed then have "\ T\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>" proof - have id_\_redex: "\ (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>) \ A =\<^bsub>o\<^esub> A" if "A \ wffs\<^bsub>o\<^esub>" for A by (fact axiom_is_derivable_from_no_hyps[OF axiom_4_2[OF that]]) from "\
1" and id_\_redex[OF true_wff] have "\ T\<^bsub>o\<^esub> \\<^sup>\ (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>) \ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>) \ \\<^bsub>o\<^esub>" by (rule rule_R[where p = "[\,\,\,\]"]) force+ from this and id_\_redex[OF false_wff] have "\ T\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>) \ \\<^bsub>o\<^esub>" by (rule rule_R[where p = "[\,\,\]"]) force+ from this and id_\_redex[OF wffs_of_type_intros(1)] show ?thesis by (rule rule_R[where p = "[\,\,\]"]) force+ qed then show ?thesis by simp qed subsection \Proposition 5215 (Universal Instantiation)\ proposition prop_5215: assumes "\ \ \x\<^bsub>\\<^esub>. B" and "A \ wffs\<^bsub>\\<^esub>" and "is_free_for A (x, \) B" shows "\ \ \<^bold>S {(x, \) \ A} B" proof - from assms(1) have "is_hyps \" by (blast elim: is_derivable_from_hyps.cases) from assms(1) have "\ \ (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) =\<^bsub>\\o\<^esub> (\x\<^bsub>\\<^esub>. B)" by simp with assms(2) have "\ \ (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) \ A =\<^bsub>o\<^esub> (\x\<^bsub>\\<^esub>. B) \ A" by (intro Equality_Rules(5)) then have "\ \ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \<^bold>S {(x, \) \ A} B" proof - have "\ \ (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) \ A =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" proof - have "\ (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) \ A =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using prop_5207[OF assms(2) true_wff is_free_for_in_true] and derived_substitution_simps(1) by (simp only:) from this and \is_hyps \\ show ?thesis by (rule derivability_implies_hyp_derivability) qed moreover have "\ \ (\x\<^bsub>\\<^esub>. B) \ A =\<^bsub>o\<^esub> \<^bold>S {(x, \) \ A} B" proof - have "B \ wffs\<^bsub>o\<^esub>" using hyp_derivable_form_is_wffso[OF assms(1)] by (fastforce elim: wffs_from_forall) with assms(2,3) have "\ (\x\<^bsub>\\<^esub>. B) \ A =\<^bsub>o\<^esub> \<^bold>S {(x, \) \ A} B" using prop_5207 by (simp only:) from this and \is_hyps \\ show ?thesis by (rule derivability_implies_hyp_derivability) qed ultimately show ?thesis using \\ \ (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) \ A =\<^bsub>o\<^esub> (\x\<^bsub>\\<^esub>. B) \ A\ and Equality_Rules(2,3) by meson qed then show ?thesis proof - have "\ \ T\<^bsub>o\<^esub>" by (fact derivability_implies_hyp_derivability[OF true_is_derivable \is_hyps \\]) from this and \\ \ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \<^bold>S {(x, \) \ A} B\ show ?thesis by (rule Equality_Rules(1)[unfolded equivalence_def]) qed qed lemmas "\I" = prop_5215 (* synonym *) subsection \Proposition 5216\ proposition prop_5216: assumes "A \ wffs\<^bsub>o\<^esub>" shows "\ (T\<^bsub>o\<^esub> \\<^sup>\ A) =\<^bsub>o\<^esub> A" proof - let ?B = "\\\<^bsub>o\<^esub>. (T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>)" have B_is_wff: "?B \ wffs\<^bsub>o\o\<^esub>" by auto have "\
1": "\ ?B \ T\<^bsub>o\<^esub> \\<^sup>\ ?B \ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. ?B \ \\<^bsub>o\<^esub>" proof - have "\ \\<^bsub>o\o\<^esub> \ T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\o\<^esub> \ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. \\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub>" (is "\ ?C =\<^bsub>o\<^esub> ?D") using axiom_1[unfolded equivalence_def] by (rule axiom_is_derivable_from_no_hyps) moreover have "?C \ wffs\<^bsub>o\<^esub>" and "?D \ wffs\<^bsub>o\<^esub>" and "is_free_for ?B (\, o\o) (?C =\<^bsub>o\<^esub> ?D)" by auto ultimately have "\ \<^bold>S {(\, o\o) \ ?B} (?C =\<^bsub>o\<^esub> ?D)" using B_is_wff and prop_5209 by presburger then show ?thesis by simp qed have *: "is_free_for A (\, o) (T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>)" for A by (intro is_free_for_in_conj is_free_for_in_equality is_free_for_in_true is_free_for_in_var) have "\ (T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>) \\<^sup>\ (T\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>)" by (fact prop_5213[OF prop_5211 prop_5214]) moreover have "\ (T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>) \\<^sup>\ (T\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. (T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>)" proof - have B_\_redex: "\ ?B \ A =\<^bsub>o\<^esub> (T\<^bsub>o\<^esub> \\<^sup>\ A =\<^bsub>o\<^esub> A)" if "A \ wffs\<^bsub>o\<^esub>" for A proof - have "T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub> \ wffs\<^bsub>o\<^esub>" by blast moreover have "\<^bold>S {(\, o) \ A} (T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>) = (T\<^bsub>o\<^esub> \\<^sup>\ A =\<^bsub>o\<^esub> A)" by simp ultimately show ?thesis using * and prop_5207[OF that] by metis qed from "\
1" and B_\_redex[OF true_wff] have "\ (T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>) \\<^sup>\ ?B \ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. ?B \ \\<^bsub>o\<^esub>" by (rule rule_R[where p = "[\,\,\,\]"]) force+ from this and B_\_redex[OF false_wff] have "\ (T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>) \\<^sup>\ (T\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. ?B \ \\<^bsub>o\<^esub>" by (rule rule_R[where p = "[\,\,\]"]) force+ from this and B_\_redex[OF wffs_of_type_intros(1)] show ?thesis by (rule rule_R[where p = "[\,\,\]"]) force+ qed ultimately have "\ \\\<^bsub>o\<^esub>. (T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>)" by (rule rule_R[where p = "[]"]) fastforce+ show ?thesis using "\I"[OF \\ \\\<^bsub>o\<^esub>. (T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>)\ assms *] by simp qed subsection \Proposition 5217\ proposition prop_5217: shows "\ (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" proof - let ?B = "\\\<^bsub>o\<^esub>. (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>)" have B_is_wff: "?B \ wffs\<^bsub>o\o\<^esub>" by auto have *: "is_free_for A (\, o) (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>)" for A by (intro is_free_for_in_equality is_free_for_in_true is_free_for_in_var) have "\
1": "\ ?B \ T\<^bsub>o\<^esub> \\<^sup>\ ?B \ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. ?B \ \\<^bsub>o\<^esub>" proof - have "\ \\<^bsub>o\o\<^esub> \ T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\o\<^esub> \ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. \\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub>" (is "\ ?C =\<^bsub>o\<^esub> ?D") using axiom_1[unfolded equivalence_def] by (rule axiom_is_derivable_from_no_hyps) moreover have "?C \ wffs\<^bsub>o\<^esub>" and "?D \ wffs\<^bsub>o\<^esub>" and "is_free_for ?B (\, o\o) (?C =\<^bsub>o\<^esub> ?D)" by auto ultimately have "\ \<^bold>S {(\, o\o) \ ?B} (?C =\<^bsub>o\<^esub> ?D)" using B_is_wff and prop_5209 by presburger then show ?thesis by simp qed then have "\ (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>) \\<^sup>\ (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>)" (is "\ ?A") proof - have B_\_redex: "\ ?B \ A =\<^bsub>o\<^esub> (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> A)" if "A \ wffs\<^bsub>o\<^esub>" for A proof - have "T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub> \ wffs\<^bsub>o\<^esub>" by auto moreover have "\<^bold>S {(\, o) \ A} (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>) = (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> A)" by simp ultimately show ?thesis using * and prop_5207[OF that] by metis qed from "\
1" and B_\_redex[OF true_wff] have "\ (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>) \\<^sup>\ ?B \ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. ?B \ \\<^bsub>o\<^esub>" by (rule rule_R[where p = "[\,\,\,\]"]) force+ from this and B_\_redex[OF false_wff] have "\ (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>) \\<^sup>\ (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. ?B \ \\<^bsub>o\<^esub>" by (rule rule_R[where p = "[\,\,\]"]) force+ from this and B_\_redex[OF wffs_of_type_intros(1)] show ?thesis by (rule rule_R[where p = "[\,\,\]"]) force+ qed from prop_5210[OF true_wff] have "\ T\<^bsub>o\<^esub> \\<^sup>\ (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>)" by (rule rule_RR[OF disjI2, where p = "[\,\,\,\]" and C = ?A]) (force+, fact) from this and prop_5216[where A = "T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>"] have "\ (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>)" by (rule rule_R [where p = "[\,\]"]) force+ moreover have "\
5": " \ ((\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>) =\<^bsub>o\o\<^esub> (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>)) =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. ((\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>) \ \\<^bsub>o\<^esub> =\<^bsub>o\<^esub> (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>) \ \\<^bsub>o\<^esub>)" proof - have "\ (\\<^bsub>o\o\<^esub> =\<^bsub>o\o\<^esub> \\<^bsub>o\o\<^esub>) =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. (\\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub>)" (is "\ ?C =\<^bsub>o\<^esub> ?D") using axiom_3[unfolded equivalence_def] by (rule axiom_is_derivable_from_no_hyps) moreover have "is_free_for ((\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>)) (\, o\o) (?C =\<^bsub>o\<^esub> ?D)" by fastforce moreover have "(\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>) \ wffs\<^bsub>o\o\<^esub>" and "?C \ wffs\<^bsub>o\<^esub>" and "?D \ wffs\<^bsub>o\<^esub>" by auto ultimately have "\ \<^bold>S {(\, o\o) \ (\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>)} (?C =\<^bsub>o\<^esub> ?D)" using prop_5209 by presburger then have "\ ((\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>) =\<^bsub>o\o\<^esub> \\<^bsub>o\o\<^esub>) =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. ((\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>) \ \\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub>)" (is "\ ?C' =\<^bsub>o\<^esub> ?D'") by simp moreover have "is_free_for ((\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>)) (\, o\o) (?C' =\<^bsub>o\<^esub> ?D')" by fastforce moreover have "(\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>) \ wffs\<^bsub>o\o\<^esub>" and "?C' \ wffs\<^bsub>o\<^esub>" and "?D' \ wffs\<^bsub>o\<^esub>" using \(\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>) \ wffs\<^bsub>o\o\<^esub>\ by auto ultimately have "\ \<^bold>S {(\, o\o) \ (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>)} (?C' =\<^bsub>o\<^esub> ?D')" using prop_5209 by presburger then show ?thesis by simp qed then have "\ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>)" proof - have "\ (\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>) \ \\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using prop_5208[where vs = "[(\, o)]"] and true_wff by simp with "\
5" have *: " \ ((\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>) =\<^bsub>o\o\<^esub> (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>)) =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>) \ \\<^bsub>o\<^esub>)" by (rule rule_R[where p = "[\,\,\,\,\]"]) force+ have "\ (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>) \ \\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>" using prop_5208[where vs = "[(\, o)]"] by fastforce with * have "\ ((\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>) =\<^bsub>o\o\<^esub> (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>)) =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>)" by (rule rule_R[where p = "[\,\,\,\]"]) force+ then show ?thesis by simp qed ultimately show ?thesis using Equality_Rules(2,3) by blast qed subsection \Proposition 5218\ proposition prop_5218: assumes "A \ wffs\<^bsub>o\<^esub>" shows "\ (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> A) =\<^bsub>o\<^esub> A" proof - let ?B = "\\\<^bsub>o\<^esub>. ((T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>)" have B_is_wff: "?B \ wffs\<^bsub>o\o\<^esub>" by auto have "\
1": "\ ?B \ T\<^bsub>o\<^esub> \\<^sup>\ ?B \ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. ?B \ \\<^bsub>o\<^esub>" proof - have "\ \\<^bsub>o\o\<^esub> \ T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\o\<^esub> \ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. \\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub>" (is "\ ?C =\<^bsub>o\<^esub> ?D") using axiom_1[unfolded equivalence_def] by (rule axiom_is_derivable_from_no_hyps) moreover have "?C \ wffs\<^bsub>o\<^esub>" and "?D \ wffs\<^bsub>o\<^esub>" and "is_free_for ?B (\, o\o) (?C =\<^bsub>o\<^esub> ?D)" by auto ultimately have "\ \<^bold>S {(\, o\o) \ ?B} (?C =\<^bsub>o\<^esub> ?D)" using prop_5209[OF B_is_wff] by presburger then show ?thesis by simp qed have *: "is_free_for A (\, o) ((T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>)" for A by (intro is_free_for_in_equality is_free_for_in_true is_free_for_in_var) have "\
2": " \ ((T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>) \\<^sup>\ ((T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. ((T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>)" proof - have B_\_redex: "\ ?B \ A =\<^bsub>o\<^esub> ((T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> A) =\<^bsub>o\<^esub> A)" if "A \ wffs\<^bsub>o\<^esub>" for A proof - have "(T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> \\<^bsub>o\<^esub> \ wffs\<^bsub>o\<^esub>" by auto moreover have "\<^bold>S {(\, o) \ A} ((T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>) = ((T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> A) =\<^bsub>o\<^esub> A)" by simp ultimately show ?thesis using * and prop_5207[OF that] by metis qed from "\
1" and B_\_redex[OF true_wff] have "\ ((T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>) \\<^sup>\ ?B \ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. ?B \ \\<^bsub>o\<^esub>" by (rule rule_R[where p = "[\,\,\,\]"]) force+ from this and B_\_redex[OF false_wff] have "\ ((T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>) \\<^sup>\ ((T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. ?B \ \\<^bsub>o\<^esub>" by (rule rule_R[where p = "[\,\,\]"]) force+ from this and B_\_redex[OF wffs_of_type_intros(1)] show ?thesis by (rule rule_R[where p = "[\,\,\]"]) force+ qed have "\
3": "\ (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by (fact Equality_Rules(2)[OF prop_5210 [OF true_wff]]) have "\ ((T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>) \\<^sup>\ ((T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>)" by (fact prop_5213[OF "\
3" prop_5217]) from this and "\
2" have "\
4": "\ \\\<^bsub>o\<^esub>. ((T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>)" by (rule rule_R[where p = "[]"]) fastforce+ then show ?thesis using "\I"[OF "\
4" assms *] by simp qed subsection \Proposition 5219 (Rule T)\ proposition prop_5219_1: assumes "A \ wffs\<^bsub>o\<^esub>" shows "\ \ A \ \ \ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> A" proof safe assume "\ \ A" then have "is_hyps \" by (blast dest: is_derivable_from_hyps.cases) then have "\ \ (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> A) =\<^bsub>o\<^esub> A" by (fact derivability_implies_hyp_derivability[OF prop_5218[OF assms]]) with \\ \ A\ show "\ \ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> A" using Equality_Rules(1)[unfolded equivalence_def] and Equality_Rules(2) by blast next assume "\ \ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> A" then have "is_hyps \" by (blast dest: is_derivable_from_hyps.cases) then have "\ \ (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> A) =\<^bsub>o\<^esub> A" by (fact derivability_implies_hyp_derivability[OF prop_5218[OF assms]]) with \\ \ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> A\ show "\ \ A" by (rule Equality_Rules(1)[unfolded equivalence_def]) qed proposition prop_5219_2: assumes "A \ wffs\<^bsub>o\<^esub>" shows "\ \ A \ \ \ A =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using prop_5219_1[OF assms] and Equality_Rules(2) by blast lemmas rule_T = prop_5219_1 prop_5219_2 subsection \Proposition 5220 (Universal Generalization)\ context begin private lemma const_true_\_conversion: shows "\ (\x\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) =\<^bsub>\\o\<^esub> (\z\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>)" proof - have "(z, \) \ free_vars T\<^bsub>o\<^esub>" and "is_free_for (z\<^bsub>\\<^esub>) (x, \) T\<^bsub>o\<^esub>" by auto then have "\ (\x\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) =\<^bsub>\\o\<^esub> \z\<^bsub>\\<^esub>. \<^bold>S {(x, \) \ z\<^bsub>\\<^esub>} T\<^bsub>o\<^esub>" by (rule prop_5206[OF true_wff]) then show ?thesis by simp qed proposition prop_5220: assumes "\ \ A" and "(x, \) \ free_vars \" shows "\ \ \x\<^bsub>\\<^esub>. A" proof - from \\ \ A\ have "is_hyps \" by (blast dest: is_derivable_from_hyps.cases) have "\ \ A" by fact then have "\
2": "\ \ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> A" using rule_T(1)[OF hyp_derivable_form_is_wffso[OF \\ \ A\]] by simp have "\
3": "\ \ (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) =\<^bsub>\\o\<^esub> (\x\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>)" by (fact derivability_implies_hyp_derivability[OF const_true_\_conversion \is_hyps \\]) from "\
3" and "\
2" have "\ \ \\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub> =\<^bsub>\\o\<^esub> \x\<^bsub>\\<^esub>. A" proof (induction rule: rule_R'[where p = "[\, \]"]) case no_capture have *: "[\,\] \ positions (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub> =\<^bsub>\\o\<^esub> \x\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>)" by simp show ?case unfolding rule_R'_side_condition_def and capture_exposed_vars_at_alt_def[OF *] using assms(2) by simp qed force+ then show ?thesis unfolding forall_def[unfolded PI_def, folded equality_of_type_def] . qed end lemmas Gen = prop_5220 (* synonym *) proposition generalized_Gen: assumes "\ \ A" and "lset vs \ free_vars \ = {}" shows "\ \ \\<^sup>\\<^sub>\ vs A" using assms(2) proof (induction vs) case Nil then show ?case using assms(1) by simp next case (Cons v vs) obtain x and \ where "v = (x, \)" by fastforce with Cons.prems have "lset vs \ free_vars \ = {}" and "(x, \) \ free_vars \" by simp_all from \lset vs \ free_vars \ = {}\ have "\ \ \\<^sup>\\<^sub>\ vs A" by (fact Cons.IH) with \(x, \) \ free_vars \\ and \v = (x, \)\ show ?case using Gen by simp qed subsection \Proposition 5221 (Substitution)\ context begin private lemma prop_5221_aux: assumes "\ \ B" and "(x, \) \ free_vars \" and "is_free_for A (x, \) B" and "A \ wffs\<^bsub>\\<^esub>" shows "\ \ \<^bold>S {(x, \) \ A} B" proof - have "\ \ B" by fact from this and assms(2) have "\ \ \x\<^bsub>\\<^esub>. B" by (rule Gen) from this and assms(4,3) show ?thesis by (rule "\I") qed proposition prop_5221: assumes "\ \ B" and "is_substitution \" and "\v \ fmdom' \. var_name v \ free_var_names \ \ is_free_for (\ $$! v) v B" and "\ \ {$$}" shows "\ \ \<^bold>S \ B" proof - obtain xs and As where "lset xs = fmdom' \" \ \i.e., $x^1_{\alpha_1},\dots,x^n_{\alpha_n}$\ and "As = map (($$!) \) xs" \ \i.e., $A^1_{\alpha_1},\dots,A^n_{\alpha_n}$\ and "length xs = card (fmdom' \)" by (metis distinct_card finite_distinct_list finite_fmdom') then have "distinct xs" by (simp add: card_distinct) from \lset xs = fmdom' \\ and \As = map (($$!) \) xs\ have "lset As = fmran' \" by (intro subset_antisym subsetI) (force simp add: fmlookup_dom'_iff fmlookup_ran'_iff)+ from assms(1) have "finite (var_name ` (vars B \ vars (lset As) \ vars \))" by (cases rule: is_derivable_from_hyps.cases) (simp_all add: finite_Domain vars_form_finiteness) then obtain ys \ \i.e., $y^1_{\alpha_1},\dots,y^n_{\alpha_n}$\ where "length ys = length xs" and "distinct ys" and ys_fresh: " (var_name ` lset ys) \ (var_name ` (vars B \ vars (lset As) \ vars \ \ lset xs)) = {}" and "map var_type ys = map var_type xs" using fresh_var_list_existence by (metis image_Un) have "length xs = length As" by (simp add: \As = map (($$!) \) xs\) \ \$\mathcal{H}\;\vdash\;\d{\textsf{S}}\; ^{x^1_{\alpha_1}\;\dots\;x^k_{\alpha_k} x^{k+1}_{\alpha_{k+1}}\;\dots\;x^n_{\alpha_n}} _{A^1_{\alpha_1}\;\dots\;A^k_{\alpha_k} y^{k+1}_{\alpha_{k+1}}\;\dots\;y^n_{\alpha_n}} B$\ have "\ \ \<^bold>S (fmap_of_list (zip xs (take k As @ drop k (map FVar ys)))) B" if "k \ length xs" for k using that proof (induction k) case 0 have "\ \ \<^bold>S (fmap_of_list (zip xs (map FVar ys))) B" using \length ys = length xs\ and \length xs = length As\ and \(var_name ` lset ys) \ (var_name ` (vars B \ vars (lset As) \ vars \ \ lset xs)) = {}\ and \lset xs = fmdom' \\ and \distinct ys\ and assms(3) and \map var_type ys = map var_type xs\ and \distinct xs\ and \length xs = card (fmdom' \)\ proof (induction ys xs As arbitrary: \ rule: list_induct3) case Nil with assms(1) show ?case using empty_substitution_neutrality by auto next \ \In the following: \begin{itemize} \item $\vartheta = \{x^1_{\alpha_1} \rightarrowtail y^1_{\alpha_1},\dots,x^n_{\alpha_n} \rightarrowtail y^n_{\alpha_n}\}$ \item \emph?$\vartheta = \{x^2_{\alpha_2} \rightarrowtail y^2_{\alpha_2},\dots,x^n_{\alpha_n} \rightarrowtail y^n_{\alpha_n}\}$ \item $v_x = x^1_{\alpha_1}$, and $v_y = y^1_{\alpha_1}$ \end{itemize} \ case (Cons v\<^sub>y ys v\<^sub>x xs A' As') let ?\ = "fmap_of_list (zip xs (map FVar ys))" from Cons.hyps(1) have "lset xs = fmdom' ?\" by simp from Cons.hyps(1) and Cons.prems(6) have "fmran' ?\ = FVar ` lset ys" by force have "is_substitution ?\" unfolding is_substitution_def proof fix v assume "v \ fmdom' ?\" with \lset xs = fmdom' ?\\ obtain k where "v = xs ! k" and "k < length xs" by (metis in_set_conv_nth) moreover obtain \ where "var_type v = \" by blast moreover from \k < length xs\ and \v = xs ! k\ have "?\ $$! v = (map FVar ys) ! k" using Cons.hyps(1) and Cons.prems(6) by auto moreover from this and \k < length xs\ obtain y and \ where "?\ $$! v = y\<^bsub>\\<^esub>" using Cons.hyps(1) by force ultimately have "\ = \" using Cons.hyps(1) and Cons.prems(5) by (metis form.inject(1) list.inject list.simps(9) nth_map snd_conv) then show "case v of (x, \) \ ?\ $$! (x, \) \ wffs\<^bsub>\\<^esub>" using \?\ $$! v = y\<^bsub>\\<^esub>\ and \var_type v = \\ by fastforce qed have "v\<^sub>x \ fmdom' ?\" using Cons.prems(6) and \lset xs = fmdom' ?\\ by auto obtain x and \ where "v\<^sub>x = (x, \)" by fastforce have "FVar v\<^sub>y \ wffs\<^bsub>\\<^esub>" using Cons.prems(5) and surj_pair[of v\<^sub>y] unfolding \v\<^sub>x = (x, \)\ by fastforce have "distinct xs" using Cons.prems(6) by fastforce moreover have ys_fresh': " (var_name ` lset ys) \ (var_name ` (vars B \ vars (lset As') \ vars \ \ lset xs)) = {}" proof - have "vars (lset (A' # As')) = vars {A'} \ vars (lset As')" by simp moreover have "var_name ` (lset (v\<^sub>x # xs)) = {var_name v\<^sub>x} \ var_name ` (lset xs)" by simp moreover from Cons.prems(1) have " var_name ` lset ys \ ( var_name ` (vars B) \ var_name ` (vars (lset (A' # As'))) \ var_name ` (vars \) \ var_name ` (lset (v\<^sub>x # xs)) ) = {}" by (simp add: image_Un) ultimately have " var_name ` lset ys \ ( var_name ` (vars B) \ var_name ` (vars (lset As')) \ var_name ` (vars \) \ var_name ` (lset (v\<^sub>x # xs)) ) = {}" by fast then show ?thesis by (simp add: image_Un) qed moreover have "distinct ys" using Cons.prems(3) by auto moreover have "\v \ fmdom' ?\. var_name v \ free_var_names \ \ is_free_for (?\ $$! v) v B" proof fix v assume "v \ fmdom' ?\" with Cons.hyps(1) obtain y where "?\ $$! v = FVar y" and "y \ lset ys" by (metis (mono_tags, lifting) fmap_of_zipped_list_range image_iff length_map list.set_map) moreover from Cons.prems(2,4) have "var_name v \ free_var_names \" using \lset xs = fmdom' ?\\ and \v \ fmdom' ?\\ by auto moreover from \y \ lset ys\ have "y \ vars B" using ys_fresh' by blast then have "is_free_for (FVar y) v B" by (intro absent_var_is_free_for) ultimately show "var_name v \ free_var_names \ \ is_free_for (?\ $$! v) v B" by simp qed moreover have "map var_type ys = map var_type xs" using Cons.prems(5) by simp moreover have "length xs = card (fmdom' ?\)" by (fact distinct_card[OF \distinct xs\, unfolded \lset xs = fmdom' ?\\, symmetric]) \ \$\mathcal{H}\;\vdash\;\d{\textsf{S}}\; ^{x^2_{\alpha_2}\;\dots\;x^n_{\alpha_n}} _{y^2_{\alpha_2}\;\dots\;y^n_{\alpha_n}} B$\ ultimately have "\ \ \<^bold>S ?\ B" using Cons.IH and \lset xs = fmdom' ?\\ by blast moreover from Cons.prems(2,4) have "(x, \) \ free_vars \" using \v\<^sub>x = (x, \)\ by auto moreover have "is_free_for (FVar v\<^sub>y) (x, \) (\<^bold>S ?\ B)" proof - have "v\<^sub>y \ fmdom' ?\" using Cons.prems(1) and \lset xs = fmdom' ?\\ by force moreover have "fmran' ?\ = lset (map FVar ys)" using Cons.hyps(1) and \distinct xs\ by simp then have "v\<^sub>y \ vars (fmran' ?\)" using Cons.prems(3) by force moreover have "v\<^sub>y \ vars B" using Cons.prems(1) by fastforce ultimately have "v\<^sub>y \ vars (\<^bold>S ?\ B)" by (rule excluded_var_from_substitution[OF \is_substitution ?\\]) then show ?thesis by (fact absent_var_is_free_for) qed \ \$\mathcal{H}\;\vdash\;\d{\textsf{S}}\;^{x^1_{\alpha_1}}_{y^1_{\alpha_1}}\;\d{\textsf{S}}\; ^{x^2_{\alpha_2}\;\dots\;x^n_{\alpha_n}}_{y^2_{\alpha_2}\;\dots\;y^n_{\alpha_n}} B$\ ultimately have "\ \ \<^bold>S {(x, \) \ FVar v\<^sub>y} (\<^bold>S ?\ B)" using \FVar v\<^sub>y \ wffs\<^bsub>\\<^esub>\ by (rule prop_5221_aux) \ \$\d{\textsf{S}}\;^{x^1_{\alpha_1}}_{y^1_{\alpha_1}}\;\d{\textsf{S}}\; ^{x^2_{\alpha_2}\;\dots\;x^n_{\alpha_n}}_{y^2_{\alpha_2}\;\dots\;y^n_{\alpha_n}} B = \d{\textsf{S}}\;^{x^1_{\alpha_1}\;\dots\;x^n_{\alpha_n}}_{y^1_{\alpha_1}\;\dots\;y^n_{\alpha_n}} B$\ moreover have "\<^bold>S {v\<^sub>x \ FVar v\<^sub>y} \<^bold>S ?\ B = \<^bold>S ({v\<^sub>x \ FVar v\<^sub>y} ++\<^sub>f ?\) B" proof - have "v\<^sub>x \ lset ys" using Cons.prems(1) by fastforce then have "\<^bold>S {v\<^sub>x \ FVar v\<^sub>y} (FVar y) = FVar y" if "y \ lset ys" for y using that and free_var_singleton_substitution_neutrality and surj_pair[of y] by fastforce with \fmran' ?\ = FVar ` lset ys\ have "fmmap (\A'. \<^bold>S {v\<^sub>x \ FVar v\<^sub>y} A') ?\ = ?\" by (fastforce intro: fmap.map_ident_strong) with \v\<^sub>x \ fmdom' ?\\ show ?thesis using \\v \ fmdom' ?\. var_name v \ free_var_names \ \ is_free_for (?\ $$! v) v B\ and substitution_consolidation by auto qed \ \$\mathcal{H}\;\vdash\;\d{\textsf{S}}\; ^{x^1_{\alpha_1}\;\dots\;x^n_{\alpha_n}}_{y^1_{\alpha_1}\;\dots\;y^n_{\alpha_n}} B$\ ultimately show ?case using \v\<^sub>x = (x, \)\ and \v\<^sub>x \ fmdom' ?\\ and fmap_singleton_comm by fastforce qed with 0 and that show ?case by auto next case (Suc k) let ?ps = "\k. zip xs (take k As @ drop k (map FVar ys))" let ?y = "ys ! k" and ?A = "As ! k" let ?\ = "\k. fmap_of_list (?ps k)" let ?\' = "\k. fmap_of_list (map (\(v', A'). (v', \<^bold>S {?y \ ?A} A')) (?ps k))" have "fmdom' (?\ k') = lset xs" for k' by (simp add: \length xs = length As\ \length ys = length xs\) have "fmdom' (?\' k') = lset xs" for k' using \length xs = length As\ and \length ys = length xs\ and fmdom'_fmap_of_list by simp have "?y \ lset ys" using Suc.prems \length ys = length xs\ by simp have "\j < length ys. ys ! j \ vars (\::form set) \ ys ! j \ vars B" using \(var_name ` lset ys) \ (var_name ` (vars B \ vars (lset As) \ vars \ \ lset xs)) = {}\ by force obtain n\<^sub>y and \\<^sub>y where "(n\<^sub>y, \\<^sub>y) = ?y" using surj_pair[of ?y] by fastforce moreover have "?A \ wffs\<^bsub>\\<^sub>y\<^esub>" proof - from Suc.prems and \(n\<^sub>y, \\<^sub>y) = ?y\ have "var_type (xs ! k) = \\<^sub>y" using \length ys = length xs\ and \map var_type ys = map var_type xs\ and Suc_le_lessD by (metis nth_map snd_conv) with Suc.prems and assms(2) and \lset xs = fmdom' \\ and \As = map (($$!) \) xs\ show ?thesis using less_eq_Suc_le and nth_mem by fastforce qed ultimately have "is_substitution {?y \ ?A}" by auto have wfs: "is_substitution (?\ k)" for k unfolding is_substitution_def proof fix v assume "v \ fmdom' (?\ k)" with \fmdom' (?\ k) = lset xs\ obtain j where "v = xs ! j" and "j < length xs" by (fastforce simp add: in_set_conv_nth) obtain \ where "var_type v = \" by blast show "case v of (x, \) \ (?\ k) $$! (x, \) \ wffs\<^bsub>\\<^esub>" proof (cases "j < k") case True with \j < length xs\ and \v = xs ! j\ have "(?\ k) $$! v = As ! j" using \distinct xs\ and \length xs = length As\ and \length ys = length xs\ by force with assms(2) \v = xs ! j\ and \v \ fmdom' (?\ k)\ and \var_type v = \\ and \j < length xs\ have "(?\ k) $$! v \ wffs\<^bsub>\\<^esub>" using \As = map (($$!) \) xs\ and \fmdom' (?\ k) = lset xs\ and \lset xs = fmdom' \\ by auto then show ?thesis using \var_type v = \\ by force next case False with \j < length xs\ and \v = xs ! j\ have "(?\ k) $$! v = FVar (ys ! j)" using \distinct xs\ and \length xs = length As\ and \length ys = length xs\ by force with \j < length xs\ and \v = xs ! j\ and \var_type v = \\ and \length ys = length xs\ have "(?\ k) $$! v \ wffs\<^bsub>\\<^esub>" using \map var_type ys = map var_type xs\ and surj_pair[of "ys ! j"] by (metis nth_map snd_conv wffs_of_type_intros(1)) then show ?thesis using \var_type v = \\ by force qed qed have \'_alt_def: "?\' k = fmap_of_list (zip xs (take k (map (\A'. \<^bold>S {?y \ ?A} A') As) @ (drop k (map (\A'. \<^bold>S {?y \ ?A} A') (map FVar ys)))))" proof - have " fmap_of_list (zip xs (map (\A'. \<^bold>S {?y \ ?A} A') (take k As @ drop k (map FVar ys)))) = fmap_of_list (zip xs (map (\A'. \<^bold>S {?y \ ?A} A') (take k As) @ (drop k (map (\A'. \<^bold>S {?y \ ?A} A') (map FVar ys)))))" by (simp add: drop_map) then show ?thesis by (metis take_map zip_map2) qed \ \$\mathcal{H}\;\vdash\;\d{\textsf{S}}\; ^{x^1_{\alpha_1}\;\dots\;x^k_{\alpha_k} x^{k+1}_{\alpha_{k+1}}\;\dots\;x^n_{\alpha_n}} _{A^1_{\alpha_1}\;\dots\;A^k_{\alpha_k} y^{k+1}_{\alpha_{k+1}}\;\dots\;y^n_{\alpha_n}} B$\ have "\ \ \<^bold>S (?\ k) B" by (fact Suc.IH[OF Suc_leD[OF Suc.prems]]) \ \$\mathcal{H}\;\vdash\;\d{\textsf{S}}\;^{y^{k+1}_{\alpha_{k+1}}}_{A^{k+1}_{\alpha_{k+1}}}\; \d{\textsf{S}}\; ^{x^1_{\alpha_1}\;\dots\;x^k_{\alpha_k} x^{k+1}_{\alpha_{k+1}}\;\dots\;x^n_{\alpha_n}} _{A^1_{\alpha_1}\;\dots\;A^k_{\alpha_k} y^{k+1}_{\alpha_{k+1}}\;\dots\;y^n_{\alpha_n}} B$\ then have "\ \ \<^bold>S {?y \ ?A} \<^bold>S (?\ k) B" proof - from \(n\<^sub>y, \\<^sub>y) = ?y\ and \length ys = length xs\ have "(n\<^sub>y, \\<^sub>y) \ free_vars \" using \\j < length ys. ys ! j \ vars (\::form set) \ ys ! j \ vars B\ and free_vars_in_all_vars_set and Suc_le_lessD[OF Suc.prems] by fastforce moreover have "is_free_for ?A (n\<^sub>y, \\<^sub>y) (\<^bold>S (?\ k) B)" proof - have "is_substitution (fmdrop (xs ! k) (?\ k))" using wfs and \fmdom' (?\ k) = lset xs\ by force moreover from Suc_le_lessD[OF Suc.prems] have "var_type (xs ! k) = var_type (ys ! k)" using \length ys = length xs\ and \map var_type ys = map var_type xs\ by (metis nth_map) then have "is_substitution {xs ! k \ FVar ?y}" unfolding is_substitution_def using \(n\<^sub>y, \\<^sub>y) = ?y\ by (intro ballI) (clarsimp, metis snd_eqD wffs_of_type_intros(1)) moreover have "(xs ! k) \ fmdom' (fmdrop (xs ! k) (?\ k))" by simp moreover have " \v \ fmdom' (fmdrop (xs ! k) (?\ k)). ?y \ vars (fmdrop (xs ! k) (?\ k) $$! v)" proof fix v assume "v \ fmdom' (fmdrop (xs ! k) (?\ k))" then have "v \ fmdom' (?\ k)" by simp with \fmdom' (?\ k) = lset xs\ obtain j where "v = xs ! j" and "j < length xs" and "j \ k" using \v \ fmdom' (fmdrop (xs ! k) (?\ k))\ and \(xs ! k) \ fmdom' (fmdrop (xs ! k) (?\ k))\ by (metis in_set_conv_nth) then show "?y \ vars ((fmdrop (xs ! k) (?\ k)) $$! v)" proof (cases "j < k") case True with \j < length xs\ and \v = xs ! j\ have "(?\ k) $$! v = As ! j" using \distinct xs\ and \length xs = length As\ and \length ys = length xs\ by force moreover from \j < length xs\ and \length xs = length As\ have "?y \ vars (As ! j)" using \?y \ lset ys\ and ys_fresh by fastforce ultimately show ?thesis using \v \ fmdom' (fmdrop (xs ! k) (?\ k))\ by auto next case False with \j < length xs\ and \v = xs ! j\ have "(?\ k) $$! v = FVar (ys ! j)" using \distinct xs\ and \length xs = length As\ and \length ys = length xs\ by force moreover from Suc_le_lessD[OF Suc.prems] and \j \ k\ have "?y \ ys ! j" by (simp add: \distinct ys\ \j < length xs\ \length ys = length xs\ nth_eq_iff_index_eq) ultimately show ?thesis using \v \ fmdom' (fmdrop (xs ! k) (?\ k))\ and \xs ! k \ fmdom' (fmdrop (xs ! k) (?\ k))\ and surj_pair[of "ys ! j"] by fastforce qed qed moreover from \k < length xs\ and \length ys = length xs\ have "?y \ vars B" by (simp add: \\j < length ys. ys ! j \ vars \ \ ys ! j \ vars B\) moreover have "is_free_for ?A (xs ! k) B" proof - from Suc_le_lessD[OF Suc.prems] and \lset xs = fmdom' \\ have "xs ! k \ fmdom' \" using nth_mem by blast moreover from Suc.prems and \As = map (($$!) \) xs\ have "\ $$! (xs ! k) = ?A" by fastforce ultimately show ?thesis using assms(3) by simp qed moreover have "\v \ fmdom' (fmdrop (xs ! k) (?\ k)). is_free_for (fmdrop (xs ! k) (?\ k) $$! v) v B" proof fix v assume "v \ fmdom' (fmdrop (xs ! k) (?\ k))" then have "v \ fmdom' (?\ k)" by simp with \fmdom' (?\ k) = lset xs\ obtain j where "v = xs ! j" and "j < length xs" and "j \ k" using \v \ fmdom' (fmdrop (xs ! k) (?\ k))\ and \(xs ! k) \ fmdom' (fmdrop (xs ! k) (?\ k))\ by (metis in_set_conv_nth) then show "is_free_for (fmdrop (xs ! k) (?\ k) $$! v) v B" proof (cases "j < k") case True with \j < length xs\ and \v = xs ! j\ have "(?\ k) $$! v = As ! j" using \distinct xs\ and \length xs = length As\ and \length ys = length xs\ by force moreover have "is_free_for (As ! j) v B" proof - from \j < length xs\ and \lset xs = fmdom' \\ and \v = xs ! j\ have "v \ fmdom' \" using nth_mem by blast moreover have "\ $$! v = As ! j" by (simp add: \As = map (($$!) \) xs\ \j < length xs\ \v = xs ! j\) ultimately show ?thesis using assms(3) by simp qed ultimately show ?thesis using \v \ fmdom' (fmdrop (xs ! k) (?\ k))\ by auto next case False with \j < length xs\ and \v = xs ! j\ have "(?\ k) $$! v = FVar (ys ! j)" using \distinct xs\ and \length xs = length As\ and \length ys = length xs\ by force moreover from \j < length xs\ and \length ys = length xs\ have "ys ! j \ vars B" using \\j < length ys. ys ! j \ vars \ \ ys ! j \ vars B\ by simp then have "is_free_for (FVar (ys ! j)) v B" by (fact absent_var_is_free_for) ultimately show ?thesis using \v \ fmdom' (fmdrop (xs ! k) (?\ k))\ by auto qed qed ultimately have "is_free_for ?A (ys ! k) \<^bold>S ({xs ! k \ FVar ?y} ++\<^sub>f fmdrop (xs ! k) (?\ k)) B" using is_free_for_with_renaming_substitution by presburger moreover have "\<^bold>S ({xs ! k \ FVar ?y} ++\<^sub>f fmdrop (xs ! k) (?\ k)) B = \<^bold>S (?\ k) B" using \length xs = length As\ and \length ys = length xs\ and Suc_le_eq and Suc.prems and \distinct xs\ by simp ultimately show ?thesis unfolding \(n\<^sub>y, \\<^sub>y) = ?y\ by simp qed ultimately show ?thesis using prop_5221_aux[OF \\ \ \<^bold>S (?\ k) B\] and \?A \ wffs\<^bsub>\\<^sub>y\<^esub>\ and \(n\<^sub>y, \\<^sub>y) = ?y\ by metis qed \ \$\d{\textsf{S}}\;^{y^{k+1}_{\alpha_{k+1}}}_{A^{k+1}_{\alpha_{k+1}}}\; \d{\textsf{S}}\; ^{x^1_{\alpha_1}\;\dots\;x^k_{\alpha_k} x^{k+1}_{\alpha_{k+1}}\;\dots\;x^n_{\alpha_n}} _{A^1_{\alpha_1}\;\dots\;A^k_{\alpha_k} y^{k+1}_{\alpha_{k+1}}\;\dots\;y^n_{\alpha_n}} B = \d{\textsf{S}}\; ^{x^1_{\alpha_1}\;\dots\;x^k_{\alpha_k} x^{k+1}_{\alpha_{k+1}} x^{k+2}_{\alpha_{k+2}} \;\dots\;x^n_{\alpha_n}} _{A^1_{\alpha_1}\;\dots\;A^k_{\alpha_k} A^{k+1}_{\alpha_{k+1}} y^{k+2}_{\alpha_{k+2}} \;\dots\;y^n_{\alpha_n}} B$\ moreover have "\<^bold>S {?y \ ?A} \<^bold>S (?\ k) B = \<^bold>S (?\ (Suc k)) B" proof - have "\<^bold>S {?y \ ?A} \<^bold>S (?\ k) B = \<^bold>S {?y \ ?A} ++\<^sub>f (?\' k) B" proof - have "?y \ fmdom' (?\ k)" using \?y \ lset ys\ and \fmdom' (?\ k) = lset xs\ and ys_fresh by blast moreover have "(?\' k) = fmmap (\A'. \<^bold>S {?y \ ?A} A') (?\ k)" using \length xs = length As\ and \length ys = length xs\ by simp moreover have "\v' \ fmdom' (?\ k). is_free_for (?\ k $$! v') v' B" proof fix v' assume "v' \ fmdom' (?\ k)" with \fmdom' (?\ k) = lset xs\ obtain j where "v' = xs ! j" and "j < length xs" by (metis in_set_conv_nth) obtain \ where "var_type v' = \" by blast show "is_free_for (?\ k $$! v') v' B" proof (cases "j < k") case True with \j < length xs\ and \v' = xs ! j\ have "(?\ k) $$! v' = As ! j" using \distinct xs\ and \length xs = length As\ and \length ys = length xs\ by force moreover from \lset xs = fmdom' \\ and assms(3) have "is_free_for (As ! j) (xs ! j) B" by (metis \As = map (($$!) \) xs\ \j < length xs\ nth_map nth_mem) ultimately show ?thesis using \v' = xs ! j\ by (simp only:) next case False with \j < length xs\ and \v' = xs ! j\ have "(?\ k) $$! v' = FVar (ys ! j)" using \distinct xs\ and \length xs = length As\ and \length ys = length xs\ by force moreover from \j < length xs\ have "is_free_for (FVar (ys ! j)) (xs ! j) B" using \\j < length ys. ys ! j \ vars \ \ ys ! j \ vars B\ and \length ys = length xs\ and absent_var_is_free_for by presburger ultimately show ?thesis using \v' = xs ! j\ by (simp only:) qed qed ultimately show ?thesis using substitution_consolidation by simp qed also have "\ = \<^bold>S {?y \ ?A} ++\<^sub>f (?\ (Suc k)) B" proof - have "?\' k = ?\ (Suc k)" proof (intro fsubset_antisym[unfolded fmsubset_alt_def] fmpredI) { fix v' and A' assume "?\' k $$ v' = Some A'" then have "v' \ fmdom' (?\' k)" by (intro fmdom'I) then obtain j where "j < length xs" and "xs ! j = v'" using \fmdom' (?\' k) = lset xs\ by (metis in_set_conv_nth) then consider (a) "j < k" | (b) "j = k" | (c) "j \ {k<..< length xs}" by fastforce then show "?\ (Suc k) $$ v' = Some A'" proof cases case a with \'_alt_def and \distinct xs\ and \j < length xs\ have "?\' k $$ (xs ! j) = Some (take k (map (\A'. \<^bold>S {?y \ ?A} A') As) ! j)" using \length xs = length As\ and \length ys = length xs\ by auto also from a and Suc.prems have "\ = Some (\<^bold>S {?y \ ?A} (As ! j))" using \length xs = length As\ by auto also have "\ = Some (As ! j)" proof - from Suc.prems and \length ys = length xs\ have "Suc k \ length ys" by (simp only:) moreover have "j < length As" using \j < length xs\ and \length xs = length As\ by (simp only:) ultimately have "?y \ vars (As ! j)" using ys_fresh by force then show ?thesis using free_var_singleton_substitution_neutrality and free_vars_in_all_vars by blast qed also from a and \xs ! j = v'\ and \distinct xs\ have "\ = ?\ (Suc k) $$ v'" using \j < length xs\ and \length xs = length As\ and \length ys = length xs\ by fastforce finally show ?thesis using \?\' k $$ v' = Some A'\ and \xs ! j = v'\ by simp next case b then have " ?\' k $$ (xs ! k) = Some (drop k (map (\A'. \<^bold>S {?y \ ?A} A') (map FVar ys)) ! 0)" using \distinct xs\ and \j < length xs\ and \length xs = length As\ and \length ys = length xs\ and fmap_of_list_nth_split by simp also from Suc.prems have "\ = Some (\<^bold>S {?y \ ?A} (FVar ?y))" using \length ys = length xs\ by simp also from \(n\<^sub>y, \\<^sub>y) = ys ! k\ have "\ = Some ?A" by (metis singleton_substitution_simps(1)) also from b and \xs ! j = v'\ and \distinct xs\ have "\ = ?\ (Suc k) $$ v'" using \j < length xs\ and \length xs = length As\ and \length ys = length xs\ by fastforce finally show ?thesis using b and \?\' k $$ v' = Some A'\ and \xs ! j = v'\ by force next case c then have "j > k" by simp with \'_alt_def and \distinct xs\ and \j < length xs\ have " ?\' k $$ (xs ! j) = Some (drop k (map (\A'. \<^bold>S {?y \ ?A} A') (map FVar ys)) ! (j - k))" using fmap_of_list_nth_split and \length xs = length As\ and \length ys = length xs\ by simp also from Suc.prems and c have "\ = Some (\<^bold>S {?y \ ?A} (FVar (ys ! j)))" using \length ys = length xs\ by simp also from Suc_le_lessD[OF Suc.prems] and \distinct ys\ have "\ = Some (FVar (ys ! j))" using \j < length xs\ and \k < j\ and \length ys = length xs\ by (metis nless_le nth_eq_iff_index_eq prod.exhaust_sel singleton_substitution_simps(1)) also from c and \distinct xs\ have "\ = ?\ (Suc k) $$ v'" using \xs ! j = v'\ and \length xs = length As\ and \length ys = length xs\ by force finally show ?thesis using \?\' k $$ v' = Some A'\ and \xs ! j = v'\ by force qed } note \_k_in_Sub_k = this { fix v' and A' assume "?\ (Suc k) $$ v' = Some A'" then have "v' \ fmdom' (?\ (Suc k))" by (intro fmdom'I) then obtain j where "j < length xs" and "xs ! j = v'" using \fmdom' (?\ (Suc k)) = lset xs\ by (metis in_set_conv_nth) then consider (a) "j < k" | (b) "j = k" | (c) "j \ {k<..< length xs}" by fastforce with \j < length xs\ and \xs ! j = v'\ and \_k_in_Sub_k show "?\' k $$ v' = Some A'" using \\k'. fmdom' (?\' k') = lset xs\ and \?\ (Suc k) $$ v' = Some A'\ by (metis (mono_tags, lifting) fmlookup_dom'_iff nth_mem)+ } qed then show ?thesis by presburger qed also have "\ = \<^bold>S (?\ (Suc k)) B" proof - have "?\ (Suc k) $$ ?y = None" using \?y \ lset ys\ \\k'. fmdom' (?\ k') = lset xs\ and ys_fresh by blast moreover from Suc_le_lessD[OF Suc.prems] have "?y \ vars B" using \\j < length ys. ys ! j \ vars \ \ ys ! j \ vars B\ and \length ys = length xs\ by auto ultimately show ?thesis by (rule substitution_absorption) qed finally show ?thesis . qed \ \$\mathcal{H}\;\vdash\;\d{\textsf{S}}\; ^{x^1_{\alpha_1}\;\dots\;x^k_{\alpha_k} x^{k+1}_{\alpha_{k+1}} x^{k+2}_{\alpha_{k+2}} \;\dots\;x^n_{\alpha_n}} _{A^1_{\alpha_1}\;\dots\;A^k_{\alpha_k} A^{k+1}_{\alpha_{k+1}} y^{k+2}_{\alpha_{k+2}} \;\dots\;y^n_{\alpha_n}} B$\ ultimately show ?case by (simp only:) qed \ \$\mathcal{H}\;\vdash\;\d{\textsf{S}}\; ^{x^1_{\alpha_1}\;\dots\;x^n_{\alpha_n}} _{A^1_{\alpha_1}\;\dots\;A^n_{\alpha_n}} B$\ then have "\ \ \<^bold>S (fmap_of_list (zip xs As)) B" using \length xs = length As\ and \length ys = length xs\ by force moreover have "fmap_of_list (zip xs As) = \" proof (intro fsubset_antisym[unfolded fmsubset_alt_def] fmpredI) fix v and A assume "fmap_of_list (zip xs As) $$ v = Some A" with \lset xs = fmdom' \\ have "v \ fmdom' \" by (fast dest: fmap_of_list_SomeD set_zip_leftD) with \fmap_of_list (zip xs As) $$ v = Some A\ and \As = map (($$!) \) xs\ show "\ $$ v = Some A" by (simp add: map_of_zip_map fmap_of_list.rep_eq split: if_splits) (meson fmdom'_notI option.exhaust_sel) next fix v and A assume "\ $$ v = Some A" with \As = map (($$!) \) xs\ show "fmap_of_list (zip xs As) $$ v = Some A" using \lset xs = fmdom' \\ by (simp add: fmap_of_list.rep_eq fmdom'I map_of_zip_map) qed ultimately show ?thesis by (simp only:) qed end lemmas Sub = prop_5221 (* synonym *) subsection \Proposition 5222 (Rule of Cases)\ lemma forall_\_conversion: assumes "A \ wffs\<^bsub>o\<^esub>" and "(z, \) \ free_vars A" and "is_free_for (z\<^bsub>\\<^esub>) (x, \) A" shows "\ \x\<^bsub>\\<^esub>. A =\<^bsub>o\<^esub> \z\<^bsub>\\<^esub>. \<^bold>S {(x, \) \ z\<^bsub>\\<^esub>} A" proof - from assms(1) have "\x\<^bsub>\\<^esub>. A \ wffs\<^bsub>o\<^esub>" by (intro forall_wff) then have "\ \x\<^bsub>\\<^esub>. A =\<^bsub>o\<^esub> \x\<^bsub>\\<^esub>. A" by (fact prop_5200) moreover from assms have "\ (\x\<^bsub>\\<^esub>. A) =\<^bsub>\\o\<^esub> (\z\<^bsub>\\<^esub>. \<^bold>S {(x, \) \ z\<^bsub>\\<^esub>} A)" by (rule prop_5206) ultimately show ?thesis unfolding forall_def and PI_def by (rule rule_R [where p = "[\,\]"]) force+ qed proposition prop_5222: assumes "\ \ \<^bold>S {(x, o) \ T\<^bsub>o\<^esub>} A" and "\ \ \<^bold>S {(x, o) \ F\<^bsub>o\<^esub>} A" and "A \ wffs\<^bsub>o\<^esub>" shows "\ \ A" proof - from assms(1) have "is_hyps \" by (blast elim: is_derivable_from_hyps.cases) have "\
1": "\ \ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> (\x\<^bsub>o\<^esub>. A) \ T\<^bsub>o\<^esub>" proof - have "\ (\x\<^bsub>o\<^esub>. A) \ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \<^bold>S {(x, o) \ T\<^bsub>o\<^esub>} A" using prop_5207[OF true_wff assms(3) closed_is_free_for] by simp from this and assms(1) have "\ \ (\x\<^bsub>o\<^esub>. A) \ T\<^bsub>o\<^esub>" using rule_RR[OF disjI2, where p = "[]"] by fastforce moreover have "(\x\<^bsub>o\<^esub>. A) \ T\<^bsub>o\<^esub> \ wffs\<^bsub>o\<^esub>" by (fact hyp_derivable_form_is_wffso[OF \\ \ (\x\<^bsub>o\<^esub>. A) \ T\<^bsub>o\<^esub>\]) ultimately show ?thesis using rule_T(1) by blast qed moreover have "\
2": "\ \ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> (\x\<^bsub>o\<^esub>. A) \ F\<^bsub>o\<^esub>" proof - have "\ (\x\<^bsub>o\<^esub>. A) \ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \<^bold>S {(x, o) \ F\<^bsub>o\<^esub>} A" using prop_5207[OF false_wff assms(3) closed_is_free_for] by simp from this and assms(2) have "\ \ (\x\<^bsub>o\<^esub>. A) \ F\<^bsub>o\<^esub>" using rule_RR[OF disjI2, where p = "[]"] by fastforce moreover have "(\x\<^bsub>o\<^esub>. A) \ F\<^bsub>o\<^esub> \ wffs\<^bsub>o\<^esub>" by (fact hyp_derivable_form_is_wffso[OF \\ \ (\x\<^bsub>o\<^esub>. A) \ F\<^bsub>o\<^esub>\]) ultimately show ?thesis using rule_T(1) by blast qed moreover from prop_5212 and \is_hyps \\ have "\
3": "\ \ T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>" by (rule derivability_implies_hyp_derivability) ultimately have "\ \ (\x\<^bsub>o\<^esub>. A) \ T\<^bsub>o\<^esub> \\<^sup>\ (\x\<^bsub>o\<^esub>. A) \ F\<^bsub>o\<^esub>" proof - from "\
3" and "\
1" have "\ \ (\x\<^bsub>o\<^esub>. A) \ T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>" by (rule rule_R'[where p = "[\,\]"]) (force+, fastforce dest: subforms_from_app) from this and "\
2" show ?thesis by (rule rule_R'[where p = "[\]"]) (force+, fastforce dest: subforms_from_app) qed moreover have "\ (\x\<^bsub>o\<^esub>. A) \ T\<^bsub>o\<^esub> \\<^sup>\ (\x\<^bsub>o\<^esub>. A) \ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \x\<^bsub>o\<^esub>. A" proof - have "\\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub> \ wffs\<^bsub>o\<^esub>" by blast have "\ \\<^bsub>o\o\<^esub> \ T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\o\<^esub> \ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. \\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub>" using axiom_1[unfolded equivalence_def] by (rule axiom_is_derivable_from_no_hyps) \ \By $\alpha$-conversion\ then have "\ \\<^bsub>o\o\<^esub> \ T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\o\<^esub> \ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \x\<^bsub>o\<^esub>. \\<^bsub>o\o\<^esub> \ x\<^bsub>o\<^esub>" (is "\ ?B =\<^bsub>o\<^esub> ?C") proof - have "\ \\\<^bsub>o\<^esub>. \\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \x\<^bsub>o\<^esub>. \\<^bsub>o\o\<^esub> \ x\<^bsub>o\<^esub>" proof (cases "x = \") case True from \\\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub> \ wffs\<^bsub>o\<^esub>\ have "\ \\\<^bsub>o\<^esub>. \\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. \\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub>" by (fact prop_5200[OF forall_wff]) with True show ?thesis using identity_singleton_substitution_neutrality by simp next case False from \\\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub> \ wffs\<^bsub>o\<^esub>\ have "\ \\\<^bsub>o\<^esub>. \\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \x\<^bsub>o\<^esub>. \<^bold>S {(\, o) \ x\<^bsub>o\<^esub>} (\\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub>)" by (rule forall_\_conversion) (simp add: False, intro is_free_for_to_app is_free_for_in_var) then show ?thesis by force qed with \\ \\<^bsub>o\o\<^esub> \ T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\o\<^esub> \ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\\<^bsub>o\<^esub>. \\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub>\ show ?thesis using Equality_Rules(3) by blast qed \ \By Sub\ then have *: "\ (\x\<^bsub>o\<^esub>. A) \ T\<^bsub>o\<^esub> \\<^sup>\ (\x\<^bsub>o\<^esub>. A) \ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \x\<^bsub>o\<^esub>. (\x\<^bsub>o\<^esub>. A) \ x\<^bsub>o\<^esub>" proof - let ?\ = "{(\, o\o) \ \x\<^bsub>o\<^esub>. A}" from assms(3) have "is_substitution ?\" by auto moreover have " \v \ fmdom' ?\. var_name v \ free_var_names ({}::form set) \ is_free_for (?\ $$! v) v (?B =\<^bsub>o\<^esub> ?C)" by (code_simp, (unfold atomize_conj[symmetric])?, simp)+ blast moreover have "?\ \ {$$}" by simp ultimately have "\ \<^bold>S ?\ (?B =\<^bsub>o\<^esub> ?C)" by (rule Sub [OF \\ ?B =\<^bsub>o\<^esub> ?C\]) then show ?thesis by simp qed \ \By $\lambda$-conversion\ then show ?thesis proof - have "\ (\x\<^bsub>o\<^esub>. A) \ x\<^bsub>o\<^esub> =\<^bsub>o\<^esub> A" using prop_5208[where vs = "[(x, o)]"] and assms(3) by simp from * and this show ?thesis by (rule rule_R[where p = "[\,\,\]"]) force+ qed qed ultimately have "\ \ \x\<^bsub>o\<^esub>. A" using rule_RR and is_subform_at.simps(1) by (blast intro: empty_is_position) then show ?thesis proof - have "is_free_for (x\<^bsub>o\<^esub>) (x, o) A" by fastforce from \\ \ \x\<^bsub>o\<^esub>. A\ and wffs_of_type_intros(1) and this show ?thesis by (rule "\I"[of \ x o A "x\<^bsub>o\<^esub>", unfolded identity_singleton_substitution_neutrality]) qed qed lemmas Cases = prop_5222 (* synonym *) subsection \Proposition 5223\ proposition prop_5223: shows "\ (T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>" proof - have "\ (T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> (T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>))" proof - let ?A = "(\\\<^bsub>o\<^esub>. \\\<^bsub>o\<^esub>. (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)) \ T\<^bsub>o\<^esub> \ \\<^bsub>o\<^esub>" have "?A \ wffs\<^bsub>o\<^esub>" by force then have "\ ?A =\<^bsub>o\<^esub> ?A" by (fact prop_5200) then have "\ (T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> ?A" unfolding imp_fun_def and imp_op_def . moreover have "\ (\\\<^bsub>o\<^esub>. \\\<^bsub>o\<^esub>. (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)) \ T\<^bsub>o\<^esub> =\<^bsub>o\o\<^esub> \\\<^bsub>o\<^esub>. (T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)" proof - have "\\\<^bsub>o\<^esub>. (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>) \ wffs\<^bsub>o\o\<^esub>" by auto moreover have "is_free_for T\<^bsub>o\<^esub> (\, o) (\\\<^bsub>o\<^esub>. (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>))" by fastforce moreover have "\<^bold>S {(\, o) \ T\<^bsub>o\<^esub>} (\\\<^bsub>o\<^esub>. (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)) = (\\\<^bsub>o\<^esub>. (T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>))" by simp ultimately show ?thesis using prop_5207[OF true_wff] by metis qed ultimately have *: "\ (T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> (\\\<^bsub>o\<^esub>. (T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)) \ \\<^bsub>o\<^esub>" by (rule rule_R [where p = "[\,\]"]) force+ have "T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> \ wffs\<^bsub>o\<^esub>" by auto then have "\ (\\\<^bsub>o\<^esub>. (T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)) \ \\<^bsub>o\<^esub> =\<^bsub>o\<^esub> (T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)" using prop_5208[where vs = "[(\, o)]"] by simp from * and this show ?thesis by (rule rule_R[where p = "[\]"]) force+ qed with prop_5218 have "\ (T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> (T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)" using rule_R and Equality_Rules(3) by (meson conj_op_wff true_wff wffs_of_type_intros(1)) with prop_5216 show ?thesis using rule_R and Equality_Rules(3) by (meson conj_op_wff true_wff wffs_of_type_intros(1)) qed corollary generalized_prop_5223: assumes "A \ wffs\<^bsub>o\<^esub>" shows "\ (T\<^bsub>o\<^esub> \\<^sup>\ A) =\<^bsub>o\<^esub> A" proof - have "T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> \ wffs\<^bsub>o\<^esub>" and "is_free_for A (\, o) ((T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>)" by (blast, intro is_free_for_in_equality is_free_for_in_imp is_free_for_in_true is_free_for_in_var) from this(2) have "\ \<^bold>S {(\, o) \ A} ((T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>)" by (rule prop_5209[OF assms \T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> \ wffs\<^bsub>o\<^esub>\ wffs_of_type_intros(1) prop_5223]) then show ?thesis by simp qed subsection \Proposition 5224 (Modus Ponens)\ proposition prop_5224: assumes "\ \ A" and "\ \ A \\<^sup>\ B" shows "\ \ B" proof - have "\ \ A \\<^sup>\ B" by fact moreover from assms(1) have "A \ wffs\<^bsub>o\<^esub>" by (fact hyp_derivable_form_is_wffso) from this and assms(1) have "\ \ A =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using rule_T(2) by blast ultimately have "\ \ T\<^bsub>o\<^esub> \\<^sup>\ B" by (rule rule_R'[where p = "[\,\]"]) (force+, fastforce dest: subforms_from_app) have "\ (T\<^bsub>o\<^esub> \\<^sup>\ B) =\<^bsub>o\<^esub> B" proof - let ?\ = "{(\, o) \ B}" have "B \ wffs\<^bsub>o\<^esub>" by (fact hyp_derivable_form_is_wffso[OF assms(2), THEN wffs_from_imp_op(2)]) then have "is_substitution ?\" by simp moreover have " \v \ fmdom' ?\. var_name v \ free_var_names ({}::form set) \ is_free_for (?\ $$! v) v ((T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>)" by (code_simp, (unfold atomize_conj[symmetric])?, simp)+ blast moreover have "?\ \ {$$}" by simp ultimately have "\ \<^bold>S ?\ ((T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>)" by (rule Sub[OF prop_5223]) then show ?thesis by simp qed then show ?thesis by (rule rule_RR[OF disjI1, where p = "[]"]) (use \\ \ T\<^bsub>o\<^esub> \\<^sup>\ B\ in \force+\) qed lemmas MP = prop_5224 (* synonym *) corollary generalized_modus_ponens: assumes "\ \ hs \\<^sup>\\<^sub>\ B" and "\H \ lset hs. \ \ H" shows "\ \ B" using assms proof (induction hs arbitrary: B rule: rev_induct) case Nil then show ?case by simp next case (snoc H' hs) from \\H \ lset (hs @ [H']). \ \ H\ have "\ \ H'" by simp moreover have "\ \ H' \\<^sup>\ B" proof - from \\ \ (hs @ [H']) \\<^sup>\\<^sub>\ B\ have "\ \ hs \\<^sup>\\<^sub>\ (H' \\<^sup>\ B)" by simp moreover from \\H \ lset (hs @ [H']). \ \ H\ have "\H \ lset hs. \ \ H" by simp ultimately show ?thesis by (elim snoc.IH) qed ultimately show ?case by (rule MP) qed subsection \Proposition 5225\ proposition prop_5225: shows "\ \\<^bsub>\\<^esub> \ \\<^bsub>\\o\<^esub> \\<^sup>\ \\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub>" proof - have "\\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub> \ wffs\<^bsub>o\<^esub>" by blast have "\
1": " \ \\<^bsub>\\<^esub> \ \\<^bsub>\\o\<^esub> \\<^sup>\ (((\\\<^bsub>\\o\<^esub>. \\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub>) \ (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>)) =\<^bsub>o\<^esub> ((\\\<^bsub>\\o\<^esub>. \\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub>) \ \\<^bsub>\\o\<^esub>))" proof - let ?\ = "{(\, (\\o)\o) \ \\\<^bsub>\\o\<^esub>. \\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub>, (\, \\o) \ \\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>, (\, \\o) \ \\<^bsub>\\o\<^esub>}" and ?A = "(\\<^bsub>\\o\<^esub> =\<^bsub>\\o\<^esub> \\<^bsub>\\o\<^esub>) \\<^sup>\ (\\<^bsub>(\\o)\o\<^esub> \ \\<^bsub>\\o\<^esub> \\<^sup>\ \\<^bsub>(\\o)\o\<^esub> \ \\<^bsub>\\o\<^esub>)" have "\ ?A" by (fact axiom_is_derivable_from_no_hyps[OF axiom_2]) moreover have "\\\<^bsub>\\o\<^esub>. \\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub> \ wffs\<^bsub>(\\o)\o\<^esub>" and "\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub> \ wffs\<^bsub>\\o\<^esub>" and "\\<^bsub>\\o\<^esub> \ wffs\<^bsub>\\o\<^esub>" by blast+ then have "is_substitution ?\" by simp moreover have " \v \ fmdom' ?\. var_name v \ free_var_names ({}::form set) \ is_free_for (?\ $$! v) v ?A" by (code_simp, (unfold atomize_conj[symmetric])?, simp)+ blast moreover have "?\ \ {$$}" by simp ultimately have "\ \<^bold>S ?\ ?A" by (rule Sub) then show ?thesis by simp qed have "\ \\<^bsub>\\<^esub> \ \\<^bsub>\\o\<^esub> \\<^sup>\ (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub>)" proof - have " \ (\\\<^bsub>\\o\<^esub>. \\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub>) \ (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) \ \\<^bsub>\\<^esub>" (is "\ (\?x\<^bsub>?\\<^esub>. ?B) \ ?A =\<^bsub>o\<^esub> ?C") proof - have "\ (\?x\<^bsub>?\\<^esub>. ?B) \ ?A =\<^bsub>o\<^esub> \<^bold>S {(?x, ?\) \ ?A} ?B" using prop_5207[OF wffs_of_type_intros(4)[OF true_wff] \?B \ wffs\<^bsub>o\<^esub>\] by fastforce then show ?thesis by simp qed moreover have "\ (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) \ \\<^bsub>\\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using prop_5208[where vs = "[(\, \)]"] and true_wff by simp ultimately have "\ (\\\<^bsub>\\o\<^esub>. \\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub>) \ (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by (rule Equality_Rules(3)) from "\
1" and this have "\ \\<^bsub>\\<^esub> \ \\<^bsub>\\o\<^esub> \\<^sup>\ (T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> ((\\\<^bsub>\\o\<^esub>. \\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub>) \ \\<^bsub>\\o\<^esub>))" by (rule rule_R[where p = "[\,\,\]"]) force+ moreover have "\ (\\\<^bsub>\\o\<^esub>. \\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub>) \ \\<^bsub>\\o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub>" using prop_5208[where vs = "[(\, \\o)]"] by force ultimately show ?thesis by (rule rule_R[where p = "[\,\]"]) force+ qed from this and prop_5218[OF \\\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub> \ wffs\<^bsub>o\<^esub>\] show ?thesis by (rule rule_R[where p = "[\]"]) auto qed subsection \Proposition 5226\ proposition prop_5226: assumes "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" and "is_free_for A (x, \) B" shows "\ \x\<^bsub>\\<^esub>. B \\<^sup>\ \<^bold>S {(x, \) \ A} B" proof - have "\ \\<^bsub>\\<^esub> \ (\x\<^bsub>\\<^esub>. B) \\<^sup>\ (\x\<^bsub>\\<^esub>. B) \ A" proof - let ?\ = "{(\, \\o) \ \x\<^bsub>\\<^esub>. B, (\, \) \ A}" have "\ \\<^bsub>\\<^esub> \ \\<^bsub>\\o\<^esub> \\<^sup>\ \\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub>" (is "\ ?C") by (fact prop_5225) moreover from assms have "is_substitution ?\" by auto moreover have " \v \ fmdom' ?\. var_name v \ free_var_names ({}::form set) \ is_free_for (?\ $$! v) v ?C" by (code_simp, (unfold atomize_conj[symmetric])?, fastforce)+ blast moreover have "?\ \ {$$}" by simp ultimately have "\ \<^bold>S ?\ ?C" by (rule Sub) moreover have "\<^bold>S ?\ ?C = \\<^bsub>\\<^esub> \ (\x\<^bsub>\\<^esub>. B) \\<^sup>\ (\x\<^bsub>\\<^esub>. B) \ A" by simp ultimately show ?thesis by (simp only:) qed moreover from assms have "\ (\x\<^bsub>\\<^esub>. B) \ A =\<^bsub>o\<^esub> \<^bold>S {(x, \) \ A} B" by (rule prop_5207) ultimately show ?thesis by (rule rule_R [where p = "[\]"]) force+ qed subsection \Proposition 5227\ corollary prop_5227: shows "\ F\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>" proof - have "\ \\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub> \\<^sup>\ \<^bold>S {(\, o) \ \\<^bsub>o\<^esub>} (\\<^bsub>o\<^esub>)" by (rule prop_5226) auto then show ?thesis using identity_singleton_substitution_neutrality by simp qed corollary generalized_prop_5227: assumes "A \ wffs\<^bsub>o\<^esub>" shows "\ F\<^bsub>o\<^esub> \\<^sup>\ A" proof - let ?\ = "{(\, o) \ A}" and ?B = "F\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>" from assms have "is_substitution ?\" by simp moreover have "is_free_for A (\, o) ?B" by (intro is_free_for_in_false is_free_for_in_imp is_free_for_in_var) then have " \v \ fmdom' ?\. var_name v \ free_var_names ({}::form set) \ is_free_for (?\ $$! v) v ?B" by force ultimately have "\ \<^bold>S {(\, o) \ A} (F\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)" using Sub[OF prop_5227, where \ = ?\] by fastforce then show ?thesis by simp qed subsection \Proposition 5228\ proposition prop_5228: shows "\ (T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" and "\ (T\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" and "\ (F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" and "\ (F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" proof - show "\ (T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" and "\ (T\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" using generalized_prop_5223 by blast+ next have "\ F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>" and "\ F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>" using generalized_prop_5227 by blast+ then show "\ (F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" and "\ (F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using rule_T(2) by blast+ qed subsection \Proposition 5229\ lemma false_in_conj_provability: assumes "A \ wffs\<^bsub>o\<^esub>" shows "\ F\<^bsub>o\<^esub> \\<^sup>\ A \\<^sup>\ F\<^bsub>o\<^esub>" proof - have "\ (\\\<^bsub>o\<^esub>. \\\<^bsub>o\<^esub>. (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)) \ F\<^bsub>o\<^esub> \ A" by (intro generalized_prop_5227[OF assms, unfolded imp_op_def imp_fun_def]) moreover have " \ (\\\<^bsub>o\<^esub>. \\\<^bsub>o\<^esub>. (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)) \ F\<^bsub>o\<^esub> =\<^bsub>o\o\<^esub> \\\<^bsub>o\<^esub>. (F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)" (is "\ (\?x\<^bsub>?\\<^esub>. ?A) \ ?B =\<^bsub>?\\<^esub> ?C") proof - have "?B \ wffs\<^bsub>?\\<^esub>" and "?A \ wffs\<^bsub>?\\<^esub>" and "is_free_for ?B (?x, ?\) ?A" by auto then have "\ (\?x\<^bsub>?\\<^esub>. ?A) \ ?B =\<^bsub>?\\<^esub> \<^bold>S {(?x, ?\) \ ?B} ?A" by (rule prop_5207) moreover have "\<^bold>S {(?x, ?\) \ ?B} ?A = ?C" by simp ultimately show ?thesis by (simp only:) qed ultimately have "\ (\\\<^bsub>o\<^esub>. (F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)) \ A" by (rule rule_R[where p = "[\]"]) auto moreover have " \ (\\\<^bsub>o\<^esub>. (F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)) \ A =\<^bsub>o\<^esub> (F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> \\<^sup>\ A)" (is "\ (\?x\<^bsub>?\\<^esub>. ?A) \ ?B =\<^bsub>o\<^esub> ?C") proof - have "?B \ wffs\<^bsub>?\\<^esub>" and "?A \ wffs\<^bsub>o\<^esub>" using assms by auto moreover have "is_free_for ?B (?x, ?\) ?A" by (intro is_free_for_in_equivalence is_free_for_in_conj is_free_for_in_false) fastforce ultimately have "\ (\?x\<^bsub>?\\<^esub>. ?A) \ ?B =\<^bsub>o\<^esub> \<^bold>S {(?x, ?\) \ ?B} ?A" by (rule prop_5207) moreover have "\<^bold>S {(?x, ?\) \ ?B} ?A = ?C" by simp ultimately show ?thesis by (simp only:) qed ultimately have "\ F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> \\<^sup>\ A" by (rule rule_R[where p = "[]"]) auto then show ?thesis unfolding equivalence_def by (rule Equality_Rules(2)) qed proposition prop_5229: shows "\ (T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" and "\ (T\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" and "\ (F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" and "\ (F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" proof - show "\ (T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" and "\ (T\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" using prop_5216 by blast+ next show "\ (F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" and "\ (F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" using false_in_conj_provability and true_wff and false_wff by simp_all qed subsection \Proposition 5230\ proposition prop_5230: shows "\ (T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" and "\ (T\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" and "\ (F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" and "\ (F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" proof - show "\ (T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" and "\ (T\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" unfolding equivalence_def using prop_5218 by blast+ next show "\ (F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" unfolding equivalence_def by (rule Equality_Rules(2)[OF prop_5210[OF false_wff]]) next have "\
1": "\ (F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) \\<^sup>\ ((\\\<^bsub>o\<^esub>. (\\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>)) \ F\<^bsub>o\<^esub> \\<^sup>\ (\\\<^bsub>o\<^esub>. (\\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>)) \ T\<^bsub>o\<^esub>)" proof - let ?\ = "{(\, o\o) \ \\\<^bsub>o\<^esub>. (\\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>), (\, o) \ F\<^bsub>o\<^esub>, (\, o) \ T\<^bsub>o\<^esub>}" and ?A = "(\\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \\<^bsub>o\<^esub>) \\<^sup>\ (\\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub>)" have "\ ?A" by (fact axiom_is_derivable_from_no_hyps[OF axiom_2]) moreover have "is_substitution ?\" by auto moreover have " \v \ fmdom' ?\. var_name v \ free_var_names ({}::form set) \ is_free_for (?\ $$! v) v ?A" by (code_simp, unfold atomize_conj[symmetric], simp, force)+ blast moreover have "?\ \ {$$}" by simp ultimately have "\ \<^bold>S ?\ ?A" by (rule Sub) then show ?thesis by simp qed then have "\
2": "\ (F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) \\<^sup>\ ((F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>) \\<^sup>\ (T\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>))" (is "\ ?A2") proof - have "is_free_for A (\, o) (\\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>)" for A by code_simp blast have \_reduction: "\ (\\\<^bsub>o\<^esub>. (\\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>)) \ A =\<^bsub>o\<^esub> (A \\<^sup>\ F\<^bsub>o\<^esub>)" if "A \ wffs\<^bsub>o\<^esub>" for A using prop_5207 [ OF that equivalence_wff[OF wffs_of_type_intros(1) false_wff] \is_free_for A (\, o) (\\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>)\ ] by simp from "\
1" and \_reduction[OF false_wff] have " \ (F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>) \\<^sup>\ ((F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>) \\<^sup>\ (\\\<^bsub>o\<^esub>. (\\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>)) \ T\<^bsub>o\<^esub>)" by (rule rule_R[where p = "[\,\,\]"]) force+ from this and \_reduction[OF true_wff] show ?thesis by (rule rule_R[where p = "[\,\]"]) force+ qed then have "\
3": "\ (F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) \\<^sup>\ F\<^bsub>o\<^esub>" proof - note r1 = rule_RR[OF disjI1] and r2 = rule_RR[OF disjI2] have "\
3\<^sub>1": "\ (F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) \\<^sup>\ ((F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>) \\<^sup>\ F\<^bsub>o\<^esub>)" (is \\ ?A3\<^sub>1\) by (rule r1[OF prop_5218[OF false_wff], where p = "[\,\]" and C = ?A2]) (use "\
2" in \force+\) have "\
3\<^sub>2": "\ (F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) \\<^sup>\ (T\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>)" (is \\ ?A3\<^sub>2\) by (rule r2[OF prop_5210[OF false_wff], where p = "[\,\,\]" and C = ?A3\<^sub>1]) (use "\
3\<^sub>1" in \force+\) show ?thesis by (rule r1[OF prop_5218[OF false_wff], where p = "[\]" and C = ?A3\<^sub>2]) (use "\
3\<^sub>2" in \force+\) qed then have "\ (F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) \\<^sup>\ ((F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) \\<^sup>\ F\<^bsub>o\<^esub>)" proof - have " \ (\\\<^bsub>o\<^esub>. \\\<^bsub>o\<^esub>. (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)) \ (F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\o\<^esub> \<^bold>S {(\, o) \ F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>} (\\\<^bsub>o\<^esub>. (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>))" by (rule prop_5207) auto from "\
3"[unfolded imp_op_def imp_fun_def] and this have "\ (\\\<^bsub>o\<^esub>. ((F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) \\<^sup>\ \\<^bsub>o\<^esub>)) \ F\<^bsub>o\<^esub>" by (rule rule_R[where p = "[\]"]) force+ moreover have " \ (\\\<^bsub>o\<^esub>. ((F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) \\<^sup>\ \\<^bsub>o\<^esub>)) \ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> \<^bold>S {(\, o) \ F\<^bsub>o\<^esub>} ((F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) \\<^sup>\ \\<^bsub>o\<^esub>)" by (rule prop_5207) auto ultimately show ?thesis by (rule rule_R[where p = "[]"]) force+ qed moreover have "\
5": "\ \\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>" proof - from prop_5229(2,4) have " \ \<^bold>S {(\, o) \ T\<^bsub>o\<^esub>} (\\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>)" and "\ \<^bold>S {(\, o) \ F\<^bsub>o\<^esub>} (\\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>)" by simp_all moreover have "\\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> \ wffs\<^bsub>o\<^esub>" by auto ultimately show ?thesis by (rule Cases) qed then have "\ (F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) \\<^sup>\ F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>" proof - let ?\ = "{(\, o) \ F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>}" have "is_substitution ?\" by auto moreover have "\v \ fmdom' ?\. var_name v \ free_var_names ({}::form set) \ is_free_for (?\ $$! v) v (\\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>)" by simp moreover have "?\ \ {$$}" by simp ultimately have "\ \<^bold>S ?\ (\\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>)" by (rule Sub[OF \\ \\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>\]) then show ?thesis by simp qed ultimately show "\ (F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" unfolding equivalence_def by (rule Equality_Rules(3)) qed subsection \Proposition 5231\ proposition prop_5231: shows "\ \\<^sup>\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" and "\ \\<^sup>\ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using prop_5230(3,4) unfolding neg_def and equivalence_def and equality_of_type_def . subsection \Proposition 5232\ lemma disj_op_alt_def_provability: assumes "A \ wffs\<^bsub>o\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" shows "\ A \\<^sup>\ B =\<^bsub>o\<^esub> \\<^sup>\ (\\<^sup>\ A \\<^sup>\ \\<^sup>\ B)" proof - let ?C = "(\\\<^bsub>o\<^esub>. \\\<^bsub>o\<^esub>. \\<^sup>\ (\\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^sup>\ \\<^bsub>o\<^esub>)) \ A \ B" from assms have "?C \ wffs\<^bsub>o\<^esub>" by blast have "(\\<^sup>\ (\\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^sup>\ \\<^bsub>o\<^esub>)) \ wffs\<^bsub>o\<^esub>" by auto moreover obtain z where "(z, o) \ {(\, o), (\, o)}" and "(z, o) \ free_vars A" using free_vars_form_finiteness and fresh_var_existence by (metis Un_iff Un_insert_right free_vars_form.simps(1,3) inf_sup_aci(5) sup_bot_left) then have "(z, o) \ free_vars (\\<^sup>\ (\\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^sup>\ \\<^bsub>o\<^esub>))" by simp moreover have "is_free_for (z\<^bsub>o\<^esub>) (\, o) (\\<^sup>\ (\\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^sup>\ \\<^bsub>o\<^esub>))" by (intro is_free_for_in_conj is_free_for_in_neg is_free_for_in_var) ultimately have " \ (\\\<^bsub>o\<^esub>. \\<^sup>\ (\\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^sup>\ \\<^bsub>o\<^esub>)) =\<^bsub>o\o\<^esub> (\z\<^bsub>o\<^esub>. \<^bold>S {(\, o) \ z\<^bsub>o\<^esub>} \\<^sup>\ (\\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^sup>\ \\<^bsub>o\<^esub>))" by (rule "\") then have "\ (\\\<^bsub>o\<^esub>. \\<^sup>\ (\\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^sup>\ \\<^bsub>o\<^esub>)) =\<^bsub>o\o\<^esub> (\z\<^bsub>o\<^esub>. \\<^sup>\ (\\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^sup>\ z\<^bsub>o\<^esub>))" by simp from prop_5200[OF \?C \ wffs\<^bsub>o\<^esub>\] and this have " \ (\\\<^bsub>o\<^esub>. \z\<^bsub>o\<^esub>. \\<^sup>\ (\\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^sup>\ z\<^bsub>o\<^esub>)) \ A \ B =\<^bsub>o\<^esub> (\\\<^bsub>o\<^esub>. \\\<^bsub>o\<^esub>. \\<^sup>\ (\\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^sup>\ \\<^bsub>o\<^esub>)) \ A \ B" by (rule rule_R[where p = "[\,\,\,\,\]"]) force+ moreover have "\z\<^bsub>o\<^esub>. \\<^sup>\ (\\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^sup>\ z\<^bsub>o\<^esub>) \ wffs\<^bsub>o\o\<^esub>" by blast have " \ (\\\<^bsub>o\<^esub>. \z\<^bsub>o\<^esub>. \\<^sup>\ (\\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^sup>\ z\<^bsub>o\<^esub>)) \ A =\<^bsub>o\o\<^esub> \<^bold>S {(\, o) \ A} (\z\<^bsub>o\<^esub>. \\<^sup>\ (\\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^sup>\ z\<^bsub>o\<^esub>))" by (rule prop_5207) ( fact, blast, intro is_free_for_in_neg is_free_for_in_conj is_free_for_to_abs, (fastforce simp add: \(z, o) \ free_vars A\)+ ) then have "\ (\\\<^bsub>o\<^esub>. \z\<^bsub>o\<^esub>. \\<^sup>\ (\\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^sup>\ z\<^bsub>o\<^esub>)) \ A =\<^bsub>o\o\<^esub> (\z\<^bsub>o\<^esub>. \\<^sup>\ (\\<^sup>\ A \\<^sup>\ \\<^sup>\ z\<^bsub>o\<^esub>))" using \(z, o) \ free_vars (\\<^sup>\ (\\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^sup>\ \\<^bsub>o\<^esub>))\ by simp ultimately have " \ (\z\<^bsub>o\<^esub>. \\<^sup>\ (\\<^sup>\ A \\<^sup>\ \\<^sup>\ z\<^bsub>o\<^esub>)) \ B =\<^bsub>o\<^esub> (\\\<^bsub>o\<^esub>. \\\<^bsub>o\<^esub>. \\<^sup>\ (\\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^sup>\ \\<^bsub>o\<^esub>)) \ A \ B" by (rule rule_R[where p = "[\,\,\]"]) force+ moreover have "\ (\z\<^bsub>o\<^esub>. \\<^sup>\ (\\<^sup>\ A \\<^sup>\ \\<^sup>\ z\<^bsub>o\<^esub>)) \ B =\<^bsub>o\<^esub> \<^bold>S {(z, o) \ B} (\\<^sup>\ (\\<^sup>\ A \\<^sup>\ \\<^sup>\ z\<^bsub>o\<^esub>))" by (rule prop_5207) ( fact, blast intro: assms(1), intro is_free_for_in_neg is_free_for_in_conj, use \(z, o) \ free_vars A\ is_free_at_in_free_vars in \fastforce+\ ) moreover have "\<^bold>S {(z, o) \ B} (\\<^sup>\ (\\<^sup>\ A \\<^sup>\ \\<^sup>\ z\<^bsub>o\<^esub>)) = \\<^sup>\ (\\<^sup>\ A \\<^sup>\ \\<^sup>\ B)" using free_var_singleton_substitution_neutrality[OF \(z, o) \ free_vars A\] by simp ultimately have "\ (\\\<^bsub>o\<^esub>. \\\<^bsub>o\<^esub>. \\<^sup>\ (\\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^sup>\ \\<^bsub>o\<^esub>)) \ A \ B =\<^bsub>o\<^esub> \\<^sup>\ (\\<^sup>\ A \\<^sup>\ \\<^sup>\ B)" using Equality_Rules(2,3) by metis then show ?thesis by simp qed context begin private lemma prop_5232_aux: assumes "\ \\<^sup>\ (A \\<^sup>\ B) =\<^bsub>o\<^esub> C" and "\ \\<^sup>\ A' =\<^bsub>o\<^esub> A" and "\ \\<^sup>\ B' =\<^bsub>o\<^esub> B" shows "\ A' \\<^sup>\ B' =\<^bsub>o\<^esub> C" proof - let ?D = "\\<^sup>\ (A \\<^sup>\ B) =\<^bsub>o\<^esub> C" from assms(2) have "\ \\<^sup>\ (\\<^sup>\ A' \\<^sup>\ B) =\<^bsub>o\<^esub> C" (is \\ ?A1\) by (rule rule_RR[OF disjI2, where p = "[\,\,\,\,\]" and C = ?D]) (use assms(1) in \force+\) from assms(3) have "\ \\<^sup>\ (\\<^sup>\ A' \\<^sup>\ \\<^sup>\ B') =\<^bsub>o\<^esub> C" by (rule rule_RR[OF disjI2, where p = "[\,\,\,\]" and C = ?A1]) (use \\ ?A1\ in \force+\) moreover from assms(2,3) have "A' \ wffs\<^bsub>o\<^esub>" and "B' \ wffs\<^bsub>o\<^esub>" using hyp_derivable_form_is_wffso by (blast dest: wffs_from_equality wffs_from_neg)+ then have "\ A' \\<^sup>\ B' =\<^bsub>o\<^esub> \\<^sup>\ (\\<^sup>\ A' \\<^sup>\ \\<^sup>\ B')" by (rule disj_op_alt_def_provability) ultimately show ?thesis using prop_5201_3 by blast qed proposition prop_5232: shows "\ (T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" and "\ (T\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" and "\ (F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" and "\ (F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" proof - from prop_5231(2) have "\ \\<^sup>\ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" (is \\ ?A\) . from prop_5229(4) have "\ \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by (rule rule_RR[OF disjI2, where p = "[\,\,\]" and C = ?A]) (use \\ ?A\ in \force+\) from prop_5232_aux[OF this prop_5231(1) prop_5231(1)] show "\ (T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" . from prop_5229(3) have "\ \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by (rule rule_RR[OF disjI2, where p = "[\,\,\]" and C = ?A]) (use \\ ?A\ in \force+\) from prop_5232_aux[OF this prop_5231(1) prop_5231(2)] show "\ (T\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" . from prop_5229(2) have "\ \\<^sup>\ (T\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by (rule rule_RR[OF disjI2, where p = "[\,\,\]" and C = ?A]) (use \\ ?A\ in \force+\) from prop_5232_aux[OF this prop_5231(2) prop_5231(1)] show "\ (F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" . next from prop_5231(1) have "\ \\<^sup>\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" (is \\ ?A\) . from prop_5229(1) have "\ \\<^sup>\ (T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" by (rule rule_RR[OF disjI2, where p = "[\,\,\]" and C = ?A]) (use \\ ?A\ in \force+\) from prop_5232_aux[OF this prop_5231(2) prop_5231(2)] show "\ (F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" . qed end subsection \Proposition 5233\ context begin private lemma lem_prop_5233_no_free_vars: assumes "A \ pwffs" and "free_vars A = {}" shows "(\\. is_tv_assignment \ \ \\<^sub>B \ A = \<^bold>T) \ \ A =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" (is "?A\<^sub>T \ _") and "(\\. is_tv_assignment \ \ \\<^sub>B \ A = \<^bold>F) \ \ A =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" (is "?A\<^sub>F \ _") proof - from assms have "(?A\<^sub>T \ \ A =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>) \ (?A\<^sub>F \ \ A =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>)" proof induction case T_pwff have "\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by (rule prop_5200[OF true_wff]) moreover have "\\. is_tv_assignment \ \ \\<^sub>B \ T\<^bsub>o\<^esub> = \<^bold>T" using \\<^sub>B_T by blast then have "\ (\\. is_tv_assignment \ \ \\<^sub>B \ T\<^bsub>o\<^esub> = \<^bold>F)" by (auto simp: inj_eq) ultimately show ?case by blast next case F_pwff have "\ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" by (rule prop_5200[OF false_wff]) moreover have "\\. is_tv_assignment \ \ \\<^sub>B \ F\<^bsub>o\<^esub> = \<^bold>F" using \\<^sub>B_F by blast then have "\ (\\. is_tv_assignment \ \ \\<^sub>B \ F\<^bsub>o\<^esub> = \<^bold>T)" by (auto simp: inj_eq) ultimately show ?case by blast next case (var_pwff p) \ \impossible case\ then show ?case by simp next case (neg_pwff B) from neg_pwff.hyps have "\\<^sup>\ B \ pwffs" and "free_vars B = {}" using neg_pwff.prems by (force, auto elim: free_vars_form.elims) consider (a) "\\. is_tv_assignment \ \ \\<^sub>B \ B = \<^bold>T" | (b) "\\. is_tv_assignment \ \ \\<^sub>B \ B = \<^bold>F" using closed_pwff_denotation_uniqueness[OF neg_pwff.hyps \free_vars B = {}\] and neg_pwff.hyps[THEN \\<^sub>B_graph_denotation_is_truth_value[OF \\<^sub>B_graph_\\<^sub>B]] by (auto dest: tv_cases) metis then show ?case proof cases case a with \free_vars B = {}\ have "\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> B" using neg_pwff.IH and Equality_Rules(2) by blast from prop_5231(1)[unfolded neg_def, folded equality_of_type_def] and this have "\ \\<^sup>\ B =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" unfolding neg_def[folded equality_of_type_def] by (rule rule_R[where p = "[\,\,\]"]) force+ moreover from a have "\\. is_tv_assignment \ \ \\<^sub>B \ (\\<^sup>\ B) = \<^bold>F" using \\<^sub>B_neg[OF neg_pwff.hyps] by simp ultimately show ?thesis by (auto simp: inj_eq) next case b with \free_vars B = {}\ have "\ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> B" using neg_pwff.IH and Equality_Rules(2) by blast then have "\ \\<^sup>\ B =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" unfolding neg_def[folded equality_of_type_def] using rule_T(2)[OF hyp_derivable_form_is_wffso] by blast moreover from b have "\\. is_tv_assignment \ \ \\<^sub>B \ (\\<^sup>\ B) = \<^bold>T" using \\<^sub>B_neg[OF neg_pwff.hyps] by simp ultimately show ?thesis by (auto simp: inj_eq) qed next case (conj_pwff B C) from conj_pwff.prems have "free_vars B = {}" and "free_vars C = {}" by simp_all with conj_pwff.hyps obtain b and b' where B_den: "\\. is_tv_assignment \ \ \\<^sub>B \ B = b" and C_den: "\\. is_tv_assignment \ \ \\<^sub>B \ C = b'" using closed_pwff_denotation_uniqueness by metis then have "b \ elts \" and "b' \ elts \" using closed_pwff_denotation_uniqueness[OF conj_pwff.hyps(1) \free_vars B = {}\] and closed_pwff_denotation_uniqueness[OF conj_pwff.hyps(2) \free_vars C = {}\] and conj_pwff.hyps[THEN \\<^sub>B_graph_denotation_is_truth_value[OF \\<^sub>B_graph_\\<^sub>B]] by force+ with conj_pwff.hyps consider (a) "b = \<^bold>T" and "b' = \<^bold>T" | (b) "b = \<^bold>T" and "b' = \<^bold>F" | (c) "b = \<^bold>F" and "b' = \<^bold>T" | (d) "b = \<^bold>F" and "b' = \<^bold>F" by auto then show ?case proof cases case a from prop_5229(1) have "\ T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" (is \\ ?A1\) . from B_den[unfolded a(1)] and \free_vars B = {}\ have "\ B =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using conj_pwff.IH(1) by simp then have "\ B \\<^sup>\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" (is \\ ?A2\) by (rule rule_RR[OF disjI2, where p = "[\,\,\,\]" and C = ?A1]) (use \\ ?A1\ in \force+\) from C_den[unfolded a(2)] and \free_vars C = {}\ have "\ C =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using conj_pwff.IH(2) by simp then have "\ B \\<^sup>\ C =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by (rule rule_RR[OF disjI2, where p = "[\,\,\]" and C = ?A2]) (use \\ ?A2\ in \force+\) then have "(\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) = \<^bold>T) \ \ B \\<^sup>\ C =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by blast moreover have "\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) \ \<^bold>F" using \\<^sub>B_conj[OF conj_pwff.hyps] and B_den[unfolded a(1)] and C_den[unfolded a(2)] by (auto simp: inj_eq) ultimately show ?thesis by force next case b from prop_5229(2) have "\ T\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" (is \\ ?A1\) . from B_den[unfolded b(1)] and \free_vars B = {}\ have "\ B =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using conj_pwff.IH(1) by simp then have "\ B \\<^sup>\ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" (is \\ ?A2\) by (rule rule_RR[OF disjI2, where p = "[\,\,\,\]" and C = ?A1]) (use \\ ?A1\ in \force+\) from C_den[unfolded b(2)] and \free_vars C = {}\ have "\ C =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" using conj_pwff.IH(2) by simp then have "\ B \\<^sup>\ C =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" by (rule rule_RR[OF disjI2, where p = "[\,\,\]" and C = ?A2]) (use \\ ?A2\ in \force+\) then have "(\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) = \<^bold>F) \ \ B \\<^sup>\ C =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" by blast moreover have "\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) \ \<^bold>T" using \\<^sub>B_conj[OF conj_pwff.hyps] and B_den[unfolded b(1)] and C_den[unfolded b(2)] by (auto simp: inj_eq) ultimately show ?thesis by force next case c from prop_5229(3) have "\ F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" (is \\ ?A1\) . from B_den[unfolded c(1)] and \free_vars B = {}\ have "\ B =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" using conj_pwff.IH(1) by simp then have "\ B \\<^sup>\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" (is \\ ?A2\) by (rule rule_RR[OF disjI2, where p = "[\,\,\,\]" and C = ?A1]) (use \\ ?A1\ in \force+\) from C_den[unfolded c(2)] and \free_vars C = {}\ have "\ C =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using conj_pwff.IH(2) by simp then have "\ B \\<^sup>\ C =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" by (rule rule_RR[OF disjI2, where p = "[\,\,\]" and C = ?A2]) (use \\ ?A2\ in \force+\) then have "(\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) = \<^bold>F) \ \ B \\<^sup>\ C =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" by blast moreover have "\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) \ \<^bold>T" using \\<^sub>B_conj[OF conj_pwff.hyps] and B_den[unfolded c(1)] and C_den[unfolded c(2)] by (auto simp: inj_eq) ultimately show ?thesis by force next case d from prop_5229(4) have "\ F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" (is \\ ?A1\) . from B_den[unfolded d(1)] and \free_vars B = {}\ have "\ B =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" using conj_pwff.IH(1) by simp then have "\ B \\<^sup>\ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" (is \\ ?A2\) by (rule rule_RR[OF disjI2, where p = "[\,\,\,\]" and C = ?A1]) (use \\ ?A1\ in \force+\) from C_den[unfolded d(2)] and \free_vars C = {}\ have "\ C =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" using conj_pwff.IH(2) by simp then have "\ B \\<^sup>\ C =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" by (rule rule_RR[OF disjI2, where p = "[\,\,\]" and C = ?A2]) (use \\ ?A2\ in \force+\) then have "(\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) = \<^bold>F) \ \ B \\<^sup>\ C =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" by blast moreover have "\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) \ \<^bold>T" using \\<^sub>B_conj[OF conj_pwff.hyps] and B_den[unfolded d(1)] and C_den[unfolded d(2)] by (auto simp: inj_eq) ultimately show ?thesis by force qed next case (disj_pwff B C) from disj_pwff.prems have "free_vars B = {}" and "free_vars C = {}" by simp_all with disj_pwff.hyps obtain b and b' where B_den: "\\. is_tv_assignment \ \ \\<^sub>B \ B = b" and C_den: "\\. is_tv_assignment \ \ \\<^sub>B \ C = b'" using closed_pwff_denotation_uniqueness by metis then have "b \ elts \" and "b' \ elts \" using closed_pwff_denotation_uniqueness[OF disj_pwff.hyps(1) \free_vars B = {}\] and closed_pwff_denotation_uniqueness[OF disj_pwff.hyps(2) \free_vars C = {}\] and disj_pwff.hyps[THEN \\<^sub>B_graph_denotation_is_truth_value[OF \\<^sub>B_graph_\\<^sub>B]] by force+ with disj_pwff.hyps consider (a) "b = \<^bold>T" and "b' = \<^bold>T" | (b) "b = \<^bold>T" and "b' = \<^bold>F" | (c) "b = \<^bold>F" and "b' = \<^bold>T" | (d) "b = \<^bold>F" and "b' = \<^bold>F" by auto then show ?case proof cases case a from prop_5232(1) have "\ T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" (is \\ ?A1\) . from B_den[unfolded a(1)] and \free_vars B = {}\ have "\ B =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using disj_pwff.IH(1) by simp then have "\ B \\<^sup>\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" (is \\ ?A2\) by (rule rule_RR[OF disjI2, where p = "[\,\,\,\]" and C = ?A1]) (use \\ ?A1\ in \force+\) from C_den[unfolded a(2)] and \free_vars C = {}\ have "\ C =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using disj_pwff.IH(2) by simp then have "\ B \\<^sup>\ C =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by (rule rule_RR[OF disjI2, where p = "[\,\,\]" and C = ?A2]) (use \\ ?A2\ in \force+\) then have "(\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) = \<^bold>T) \ \ B \\<^sup>\ C =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by blast moreover have "\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) \ \<^bold>F" using \\<^sub>B_disj[OF disj_pwff.hyps] and B_den[unfolded a(1)] and C_den[unfolded a(2)] by (auto simp: inj_eq) ultimately show ?thesis by force next case b from prop_5232(2) have "\ T\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" (is \\ ?A1\) . from B_den[unfolded b(1)] and \free_vars B = {}\ have "\ B =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using disj_pwff.IH(1) by simp then have "\ B \\<^sup>\ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" (is \\ ?A2\) by (rule rule_RR[OF disjI2, where p = "[\,\,\,\]" and C = ?A1]) (use \\ ?A1\ in \force+\) from C_den[unfolded b(2)] and \free_vars C = {}\ have "\ C =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" using disj_pwff.IH(2) by simp then have "\ B \\<^sup>\ C =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by (rule rule_RR[OF disjI2, where p = "[\,\,\]" and C = ?A2]) (use \\ ?A2\ in \force+\) then have "(\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) = \<^bold>T) \ \ B \\<^sup>\ C =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by blast moreover have "\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) \ \<^bold>F" using \\<^sub>B_disj[OF disj_pwff.hyps] and B_den[unfolded b(1)] and C_den[unfolded b(2)] by (auto simp: inj_eq) ultimately show ?thesis by force next case c from prop_5232(3) have "\ F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" (is \\ ?A1\) . from B_den[unfolded c(1)] and \free_vars B = {}\ have "\ B =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" using disj_pwff.IH(1) by simp then have "\ B \\<^sup>\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" (is \\ ?A2\) by (rule rule_RR[OF disjI2, where p = "[\,\,\,\]" and C = ?A1]) (use \\ ?A1\ in \force+\) from C_den[unfolded c(2)] and \free_vars C = {}\ have "\ C =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using disj_pwff.IH(2) by simp then have "\ B \\<^sup>\ C =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by (rule rule_RR[OF disjI2, where p = "[\,\,\]" and C = ?A2]) (use \\ ?A2\ in \force+\) then have "(\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) = \<^bold>T) \ \ B \\<^sup>\ C =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by blast moreover have "\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) \ \<^bold>F" using \\<^sub>B_disj[OF disj_pwff.hyps] and B_den[unfolded c(1)] and C_den[unfolded c(2)] by (auto simp: inj_eq) ultimately show ?thesis by force next case d from prop_5232(4) have "\ F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" (is \\ ?A1\) . from B_den[unfolded d(1)] and \free_vars B = {}\ have "\ B =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" using disj_pwff.IH(1) by simp then have "\ B \\<^sup>\ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" (is \\ ?A2\) by (rule rule_RR[OF disjI2, where p = "[\,\,\,\]" and C = ?A1]) (use \\ ?A1\ in \force+\) from C_den[unfolded d(2)] and \free_vars C = {}\ have "\ C =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" using disj_pwff.IH(2) by simp then have "\ B \\<^sup>\ C =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" by (rule rule_RR[OF disjI2, where p = "[\,\,\]" and C = ?A2]) (use \\ ?A2\ in \force+\) then have "(\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) = \<^bold>T) \ \ B \\<^sup>\ C =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" by blast moreover have "\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) \ \<^bold>T" using \\<^sub>B_disj[OF disj_pwff.hyps] and B_den[unfolded d(1)] and C_den[unfolded d(2)] by (auto simp: inj_eq) ultimately show ?thesis using \\ B \\<^sup>\ C =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>\ by auto qed next case (imp_pwff B C) from imp_pwff.prems have "free_vars B = {}" and "free_vars C = {}" by simp_all with imp_pwff.hyps obtain b and b' where B_den: "\\. is_tv_assignment \ \ \\<^sub>B \ B = b" and C_den: "\\. is_tv_assignment \ \ \\<^sub>B \ C = b'" using closed_pwff_denotation_uniqueness by metis then have "b \ elts \" and "b' \ elts \" using closed_pwff_denotation_uniqueness[OF imp_pwff.hyps(1) \free_vars B = {}\] and closed_pwff_denotation_uniqueness[OF imp_pwff.hyps(2) \free_vars C = {}\] and imp_pwff.hyps[THEN \\<^sub>B_graph_denotation_is_truth_value[OF \\<^sub>B_graph_\\<^sub>B]] by force+ with imp_pwff.hyps consider (a) "b = \<^bold>T" and "b' = \<^bold>T" | (b) "b = \<^bold>T" and "b' = \<^bold>F" | (c) "b = \<^bold>F" and "b' = \<^bold>T" | (d) "b = \<^bold>F" and "b' = \<^bold>F" by auto then show ?case proof cases case a from prop_5228(1) have "\ T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" (is \\ ?A1\) . from B_den[unfolded a(1)] and \free_vars B = {}\ have "\ B =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using imp_pwff.IH(1) by simp then have "\ B \\<^sup>\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" (is \\ ?A2\) by (rule rule_RR[OF disjI2, where p = "[\,\,\,\]" and C = ?A1]) (use \\ ?A1\ in \force+\) from C_den[unfolded a(2)] and \free_vars C = {}\ have "\ C =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using imp_pwff.IH(2) by simp then have "\ B \\<^sup>\ C =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by (rule rule_RR[OF disjI2, where p = "[\,\,\]" and C = ?A2]) (use \\ ?A2\ in \force+\) then have "(\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) = \<^bold>T) \ \ B \\<^sup>\ C =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by blast moreover have "\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) \ \<^bold>F" using \\<^sub>B_imp[OF imp_pwff.hyps] and B_den[unfolded a(1)] and C_den[unfolded a(2)] by (auto simp: inj_eq) ultimately show ?thesis by force next case b from prop_5228(2) have "\ T\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" (is \\ ?A1\) . from B_den[unfolded b(1)] and \free_vars B = {}\ have "\ B =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using imp_pwff.IH(1) by simp then have "\ B \\<^sup>\ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" (is \\ ?A2\) by (rule rule_RR[OF disjI2, where p = "[\,\,\,\]" and C = ?A1]) (use \\ ?A1\ in \force+\) from C_den[unfolded b(2)] and \free_vars C = {}\ have "\ C =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" using imp_pwff.IH(2) by simp then have "\ B \\<^sup>\ C =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" by (rule rule_RR[OF disjI2, where p = "[\,\,\]" and C = ?A2]) (use \\ ?A2\ in \force+\) then have "(\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) = \<^bold>F) \ \ B \\<^sup>\ C =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" by blast moreover have "\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) \ \<^bold>T" using \\<^sub>B_imp[OF imp_pwff.hyps] and B_den[unfolded b(1)] and C_den[unfolded b(2)] by (auto simp: inj_eq) ultimately show ?thesis by force next case c from prop_5228(3) have "\ F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" (is \\ ?A1\) . from B_den[unfolded c(1)] and \free_vars B = {}\ have "\ B =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" using imp_pwff.IH(1) by simp then have "\ B \\<^sup>\ T\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" (is \\ ?A2\) by (rule rule_RR[OF disjI2, where p = "[\,\,\,\]" and C = ?A1]) (use \\ ?A1\ in \force+\) from C_den[unfolded c(2)] and \free_vars C = {}\ have "\ C =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using imp_pwff.IH(2) by simp then have "\ B \\<^sup>\ C =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by (rule rule_RR[OF disjI2, where p = "[\,\,\]" and C = ?A2]) (use \\ ?A2\ in \force+\) then have "(\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) = \<^bold>T) \ \ B \\<^sup>\ C =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by blast moreover have "\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) \ \<^bold>F" using \\<^sub>B_imp[OF imp_pwff.hyps] and B_den[unfolded c(1)] and C_den[unfolded c(2)] by (auto simp: inj_eq) ultimately show ?thesis by force next case d from prop_5228(4) have "\ F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" (is \\ ?A1\) . from B_den[unfolded d(1)] and \free_vars B = {}\ have "\ B =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" using imp_pwff.IH(1) by simp then have "\ B \\<^sup>\ F\<^bsub>o\<^esub> =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" (is \\ ?A2\) by (rule rule_RR[OF disjI2, where p = "[\,\,\,\]" and C = ?A1]) (use \\ ?A1\ in \force+\) from C_den[unfolded d(2)] and \free_vars C = {}\ have "\ C =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" using imp_pwff.IH(2) by simp then have "\ B \\<^sup>\ C =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by (rule rule_RR[OF disjI2, where p = "[\,\,\]" and C = ?A2]) (use \\ ?A2\ in \force+\) then have "(\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) = \<^bold>T) \ \ B \\<^sup>\ C =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by blast moreover have "\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) \ \<^bold>F" using \\<^sub>B_imp[OF imp_pwff.hyps] and B_den[unfolded d(1)] and C_den[unfolded d(2)] by (auto simp: inj_eq) ultimately show ?thesis by force qed next case (eqv_pwff B C) from eqv_pwff.prems have "free_vars B = {}" and "free_vars C = {}" by simp_all with eqv_pwff.hyps obtain b and b' where B_den: "\\. is_tv_assignment \ \ \\<^sub>B \ B = b" and C_den: "\\. is_tv_assignment \ \ \\<^sub>B \ C = b'" using closed_pwff_denotation_uniqueness by metis then have "b \ elts \" and "b' \ elts \" using closed_pwff_denotation_uniqueness[OF eqv_pwff.hyps(1) \free_vars B = {}\] and closed_pwff_denotation_uniqueness[OF eqv_pwff.hyps(2) \free_vars C = {}\] and eqv_pwff.hyps[THEN \\<^sub>B_graph_denotation_is_truth_value[OF \\<^sub>B_graph_\\<^sub>B]] by force+ with eqv_pwff.hyps consider (a) "b = \<^bold>T" and "b' = \<^bold>T" | (b) "b = \<^bold>T" and "b' = \<^bold>F" | (c) "b = \<^bold>F" and "b' = \<^bold>T" | (d) "b = \<^bold>F" and "b' = \<^bold>F" by auto then show ?case proof cases case a from prop_5230(1) have "\ (T\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" (is \\ ?A1\) . from B_den[unfolded a(1)] and \free_vars B = {}\ have "\ B =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using eqv_pwff.IH(1) by simp then have "\ (B \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" (is \\ ?A2\) by (rule rule_RR[OF disjI2, where p = "[\,\,\,\]" and C = ?A1]) (use \\ ?A1\ in \force+\) from C_den[unfolded a(2)] and \free_vars C = {}\ have "\ C =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using eqv_pwff.IH(2) by simp then have "\ (B \\<^sup>\ C) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by (rule rule_RR[OF disjI2, where p = "[\,\,\]" and C = ?A2]) (use \\ ?A2\ in \force+\) then have "(\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) = \<^bold>T) \ \ (B \\<^sup>\ C) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by blast moreover have "\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) \ \<^bold>F" using \\<^sub>B_eqv[OF eqv_pwff.hyps] and B_den[unfolded a(1)] and C_den[unfolded a(2)] by (auto simp: inj_eq) ultimately show ?thesis by force next case b from prop_5230(2) have "\ (T\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" (is \\ ?A1\) . from B_den[unfolded b(1)] and \free_vars B = {}\ have "\ B =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using eqv_pwff.IH(1) by simp then have "\ (B \\<^sup>\ F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" (is \\ ?A2\) by (rule rule_RR[OF disjI2, where p = "[\,\,\,\]" and C = ?A1]) (use \\ ?A1\ in \force+\) from C_den[unfolded b(2)] and \free_vars C = {}\ have "\ C =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" using eqv_pwff.IH(2) by simp then have "\ (B \\<^sup>\ C) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" by (rule rule_RR[OF disjI2, where p = "[\,\,\]" and C = ?A2]) (use \\ ?A2\ in \force+\) then have "(\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) = \<^bold>F) \ \ (B \\<^sup>\ C) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" by blast moreover have "\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) \ \<^bold>T" using \\<^sub>B_eqv[OF eqv_pwff.hyps] and B_den[unfolded b(1)] and C_den[unfolded b(2)] by (auto simp: inj_eq) ultimately show ?thesis by force next case c from prop_5230(3) have "\ (F\<^bsub>o\<^esub> \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" (is \\ ?A1\) . from B_den[unfolded c(1)] and \free_vars B = {}\ have "\ B =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" using eqv_pwff.IH(1) by simp then have "\ (B \\<^sup>\ T\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" (is \\ ?A2\) by (rule rule_RR[OF disjI2, where p = "[\,\,\,\]" and C = ?A1]) (use \\ ?A1\ in \force+\) from C_den[unfolded c(2)] and \free_vars C = {}\ have "\ C =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using eqv_pwff.IH(2) by simp then have "\ (B \\<^sup>\ C) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" by (rule rule_RR[OF disjI2, where p = "[\,\,\]" and C = ?A2]) (use \\ ?A2\ in \force+\) then have "(\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) = \<^bold>F) \ \ (B \\<^sup>\ C) =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" by blast moreover have "\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) \ \<^bold>T" using \\<^sub>B_eqv[OF eqv_pwff.hyps] and B_den[unfolded c(1)] and C_den[unfolded c(2)] by (auto simp: inj_eq) ultimately show ?thesis by force next case d from prop_5230(4) have "\ (F\<^bsub>o\<^esub> \\<^sup>\ F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" (is \\ ?A1\) . from B_den[unfolded d(1)] and \free_vars B = {}\ have "\ B =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" using eqv_pwff.IH(1) by simp then have "\ (B \\<^sup>\ F\<^bsub>o\<^esub>) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" (is \\ ?A2\) by (rule rule_RR[OF disjI2, where p = "[\,\,\,\]" and C = ?A1]) (use \\ ?A1\ in \force+\) from C_den[unfolded d(2)] and \free_vars C = {}\ have "\ C =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" using eqv_pwff.IH(2) by simp then have "\ (B \\<^sup>\ C) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by (rule rule_RR[OF disjI2, where p = "[\,\,\]" and C = ?A2]) (use \\ ?A2\ in \force+\) then have "(\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) = \<^bold>T) \ \ (B \\<^sup>\ C) =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" by blast moreover have "\\. is_tv_assignment \ \ \\<^sub>B \ (B \\<^sup>\ C) \ \<^bold>F" using \\<^sub>B_eqv[OF eqv_pwff.hyps] and B_den[unfolded d(1)] and C_den[unfolded d(2)] by (auto simp: inj_eq) ultimately show ?thesis by force qed qed then show "?A\<^sub>T \ \ A =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" and "?A\<^sub>F \ \ A =\<^bsub>o\<^esub> F\<^bsub>o\<^esub>" by blast+ qed proposition prop_5233: assumes "is_tautology A" shows "\ A" proof - have "finite (free_vars A)" using free_vars_form_finiteness by presburger from this and assms show ?thesis proof (induction "free_vars A" arbitrary: A) case empty from empty(2) have "A \ pwffs" and "\\. is_tv_assignment \ \ \\<^sub>B \ A = \<^bold>T" unfolding is_tautology_def by blast+ with empty(1) have "\ A =\<^bsub>o\<^esub> T\<^bsub>o\<^esub>" using lem_prop_5233_no_free_vars(1) by (simp only:) then show ?case using rule_T(2)[OF tautology_is_wffo[OF empty(2)]] by (simp only:) next case (insert v F) from insert.prems have "A \ pwffs" by blast with insert.hyps(4) obtain p where "v = (p, o)" using pwffs_free_vars_are_propositional by blast from \v = (p, o)\ and insert.hyps(4) have " is_tautology (\<^bold>S {(p, o) \ T\<^bsub>o\<^esub>} A)" and "is_tautology (\<^bold>S {(p, o) \ F\<^bsub>o\<^esub>} A)" using pwff_substitution_tautology_preservation [OF insert.prems] by blast+ moreover from insert.hyps(2,4) and \v = (p, o)\ and \A \ pwffs\ have "free_vars (\<^bold>S {(p, o) \ T\<^bsub>o\<^esub>} A) = F" and "free_vars (\<^bold>S {(p, o) \ F\<^bsub>o\<^esub>} A) = F" using closed_pwff_substitution_free_vars and T_pwff and F_pwff and T_fv and F_fv by (metis Diff_insert_absorb insertI1)+ ultimately have "\ \<^bold>S {(p, o) \ T\<^bsub>o\<^esub>} A" and "\ \<^bold>S {(p, o) \ F\<^bsub>o\<^esub>} A" using insert.hyps(3) by (simp_all only:) from this and tautology_is_wffo[OF insert.prems] show ?case by (rule Cases) qed qed end subsection \Proposition 5234 (Rule P)\ text \ - According to the proof in @{cite "andrews:2002"}, if $[A^1 \wedge \dots \wedge A^n] \supset B$ is + According to the proof in \<^cite>\"andrews:2002"\, if $[A^1 \wedge \dots \wedge A^n] \supset B$ is tautologous, then clearly $A^1 \supset (\dots (A^n \supset B) \dots)$ is also tautologous. - Since this is not clear to us, we prove instead the version of Rule P found in @{cite "andrews:1965"}: + Since this is not clear to us, we prove instead the version of Rule P found in \<^cite>\"andrews:1965"\: \ proposition tautologous_horn_clause_is_hyp_derivable: assumes "is_hyps \" and "is_hyps \" and "\A \ \. \ \ A" and "lset hs = \" and "is_tautologous (hs \\<^sup>\\<^sub>\ B)" shows "\ \ B" proof - from assms(5) obtain \ and C where "is_tautology C" and "is_substitution \" and "\(x, \) \ fmdom' \. \ = o" and "hs \\<^sup>\\<^sub>\ B = \<^bold>S \ C" by blast then have "\ hs \\<^sup>\\<^sub>\ B" proof (cases "\ = {$$}") case True with \hs \\<^sup>\\<^sub>\ B = \<^bold>S \ C\ have "C = hs \\<^sup>\\<^sub>\ B" using empty_substitution_neutrality by simp with \hs \\<^sup>\\<^sub>\ B = \<^bold>S \ C\ and \is_tautology C\ show ?thesis using prop_5233 by (simp only:) next case False from \is_tautology C\ have "\ C" and "C \ pwffs" using prop_5233 by simp_all moreover have " \v \ fmdom' \. var_name v \ free_var_names ({}::form set) \ is_free_for (\ $$! v) v C" proof fix v assume "v \ fmdom' \" then show "var_name v \ free_var_names ({}::form set) \ is_free_for (\ $$! v) v C" proof (cases "v \ free_vars C") case True with \C \ pwffs\ show ?thesis using is_free_for_in_pwff by simp next case False then have "is_free_for (\ $$! v) v C" unfolding is_free_for_def using is_free_at_in_free_vars by blast then show ?thesis by simp qed qed ultimately show ?thesis using False and \is_substitution \\ and Sub by (simp add: \hs \\<^sup>\\<^sub>\ B = \<^bold>S \ C\[unfolded generalized_imp_op_def]) qed from this and assms(1) have "\ \ hs \\<^sup>\\<^sub>\ B" by (rule derivability_implies_hyp_derivability) with assms(3,4) show ?thesis using generalized_modus_ponens by blast qed corollary tautologous_is_hyp_derivable: assumes "is_hyps \" and "is_tautologous B" shows "\ \ B" using assms and tautologous_horn_clause_is_hyp_derivable[where \ = "{}"] by simp lemmas prop_5234 = tautologous_horn_clause_is_hyp_derivable tautologous_is_hyp_derivable lemmas rule_P = prop_5234 (* synonym *) subsection \Proposition 5235\ proposition prop_5235: assumes "A \ pwffs" and "B \ pwffs" and "(x, \) \ free_vars A" shows "\ \x\<^bsub>\\<^esub>. (A \\<^sup>\ B) \\<^sup>\ (A \\<^sup>\ \x\<^bsub>\\<^esub>. B)" proof - have "\
1": "\ \x\<^bsub>\\<^esub>. (T\<^bsub>o\<^esub> \\<^sup>\ B) \\<^sup>\ (T\<^bsub>o\<^esub> \\<^sup>\ \x\<^bsub>\\<^esub>. B)" proof (intro rule_P(2)) show "is_tautologous (\x\<^bsub>\\<^esub>. (T\<^bsub>o\<^esub> \\<^sup>\ B) \\<^sup>\ T\<^bsub>o\<^esub> \\<^sup>\ \x\<^bsub>\\<^esub>. B)" proof - let ?\ = "{(\, o) \ \x\<^bsub>\\<^esub>. (T\<^bsub>o\<^esub> \\<^sup>\ B), (\, o) \ \x\<^bsub>\\<^esub>. B}" and ?C = "\\<^bsub>o\<^esub> \\<^sup>\ (T\<^bsub>o\<^esub> \\<^sup>\ (\\<^bsub>o\<^esub>))" have "is_tautology ?C" using \\<^sub>B_simps by simp moreover from assms(2) have "is_pwff_substitution ?\" using pwffs_subset_of_wffso by fastforce moreover have "\x\<^bsub>\\<^esub>. (T\<^bsub>o\<^esub> \\<^sup>\ B) \\<^sup>\ T\<^bsub>o\<^esub> \\<^sup>\ \x\<^bsub>\\<^esub>. B = \<^bold>S ?\ ?C" by simp ultimately show ?thesis by blast qed qed simp have "\
2": "\ \x\<^bsub>\\<^esub>. B \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ \x\<^bsub>\\<^esub>. B)" proof (intro rule_P(2)) show "is_tautologous (\x\<^bsub>\\<^esub>. B \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ \x\<^bsub>\\<^esub>. B))" proof - let ?\ = "{(\, o) \ \x\<^bsub>\\<^esub>. B}" and ?C = "\\<^bsub>o\<^esub> \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ (\\<^bsub>o\<^esub>))" have "is_tautology (\\<^bsub>o\<^esub> \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ (\\<^bsub>o\<^esub>)))" (is \is_tautology ?C\) using \\<^sub>B_simps by simp moreover from assms(2) have "is_pwff_substitution ?\" using pwffs_subset_of_wffso by auto moreover have "\x\<^bsub>\\<^esub>. B \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ \x\<^bsub>\\<^esub>. B) = \<^bold>S ?\ ?C" by simp ultimately show ?thesis by blast qed qed simp have "\
3": "\ B \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ B)" proof (intro rule_P(2)) show "is_tautologous (B \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ B))" proof - let ?\ = "{(\, o) \ B}" and ?C = "\\<^bsub>o\<^esub> \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ (\\<^bsub>o\<^esub>))" have "is_tautology ?C" using \\<^sub>B_simps by simp moreover from assms(2) have "is_pwff_substitution ?\" using pwffs_subset_of_wffso by auto moreover have "B \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ B) = \<^bold>S ?\ ?C" by simp ultimately show ?thesis by blast qed qed simp from "\
2" and "\
3"[unfolded equivalence_def] have "\
4": "\ \x\<^bsub>\\<^esub>. (F\<^bsub>o\<^esub> \\<^sup>\ B) \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ \x\<^bsub>\\<^esub>. B)" by (rule rule_R[where p = "[\,\,\,\]"]) force+ obtain p where "(p, o) \ vars (\x\<^bsub>\\<^esub>. (A \\<^sup>\ B) \\<^sup>\ (A \\<^sup>\ \x\<^bsub>\\<^esub>. B))" by (meson fresh_var_existence vars_form_finiteness) then have "(p, o) \ (x, \)" and "(p, o) \ vars A" and "(p, o) \ vars B" by simp_all from \(p, o) \ vars B\ have sub: "\<^bold>S {(p, o) \ C} B = B" for C using free_var_singleton_substitution_neutrality and free_vars_in_all_vars by blast have "\
5": "\ \x\<^bsub>\\<^esub>. (p\<^bsub>o\<^esub> \\<^sup>\ B) \\<^sup>\ (p\<^bsub>o\<^esub> \\<^sup>\ \x\<^bsub>\\<^esub>. B)" (is \\ ?C\) proof - from sub and "\
1" have "\ \<^bold>S {(p, o) \ T\<^bsub>o\<^esub>} ?C" using \(p, o) \ (x, \)\ by auto moreover from sub and "\
4" have "\ \<^bold>S {(p, o) \ F\<^bsub>o\<^esub>} ?C" using \(p, o) \ (x, \)\ by auto moreover from assms(2) have "?C \ wffs\<^bsub>o\<^esub>" using pwffs_subset_of_wffso by auto ultimately show ?thesis by (rule Cases) qed then show ?thesis proof - let ?\ = "{(p, o) \ A}" from assms(1) have "is_substitution ?\" using pwffs_subset_of_wffso by auto moreover have " \v \ fmdom' ?\. var_name v \ free_var_names ({}::form set) \ is_free_for (?\ $$! v) v ?C" proof fix v assume "v \ fmdom' ?\" then have "v = (p, o)" by simp with assms(3) and \(p, o) \ vars B\ have "is_free_for (?\ $$! v) v ?C" using occurs_in_vars by (intro is_free_for_in_imp is_free_for_in_forall is_free_for_in_disj) auto moreover have "var_name v \ free_var_names ({}::form set)" by simp ultimately show "var_name v \ free_var_names ({}::form set) \ is_free_for (?\ $$! v) v ?C" unfolding \v = (p, o)\ by blast qed moreover have "?\ \ {$$}" by simp ultimately have "\ \<^bold>S ?\ ?C" by (rule Sub[OF "\
5"]) moreover have "\<^bold>S ?\ ?C = \x\<^bsub>\\<^esub>. (A \\<^sup>\ B) \\<^sup>\ (A \\<^sup>\ \x\<^bsub>\\<^esub>. B)" using \(p, o) \ (x, \)\ and sub[of A] by simp fast ultimately show ?thesis by (simp only:) qed qed subsection \Proposition 5237 ($\supset \forall$ Rule)\ text \ - The proof in @{cite "andrews:2002"} uses the pseudo-rule Q and the axiom 5 of \\\. Therefore, we - prove such axiom, following the proof of Theorem 143 in @{cite "andrews:1965"}: + The proof in \<^cite>\"andrews:2002"\ uses the pseudo-rule Q and the axiom 5 of \\\. Therefore, we + prove such axiom, following the proof of Theorem 143 in \<^cite>\"andrews:1965"\: \ context begin private lemma prop_5237_aux: assumes "A \ wffs\<^bsub>o\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" and "(x, \) \ free_vars A" shows "\ \x\<^bsub>\\<^esub>. (A \\<^sup>\ B) \\<^sup>\ (A \\<^sup>\ (\x\<^bsub>\\<^esub>. B))" proof - have "is_tautology (\\<^bsub>o\<^esub> \\<^sup>\ (T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>))" (is \is_tautology ?C\<^sub>1\) using \\<^sub>B_simps by simp have "is_tautology (\\<^bsub>o\<^esub> \\<^sup>\ (\\<^bsub>o\<^esub> \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)))" (is \is_tautology ?C\<^sub>2\) using \\<^sub>B_simps by simp have "\
1": "\ \x\<^bsub>\\<^esub>. B \\<^sup>\ (T\<^bsub>o\<^esub> \\<^sup>\ \x\<^bsub>\\<^esub>. B)" proof (intro rule_P(2)) show "is_tautologous (\x\<^bsub>\\<^esub>. B \\<^sup>\ (T\<^bsub>o\<^esub> \\<^sup>\ \x\<^bsub>\\<^esub>. B))" proof - let ?\ = "{(\, o) \ \x\<^bsub>\\<^esub>. B}" from assms(2) have "is_pwff_substitution ?\" using pwffs_subset_of_wffso by auto moreover have "\x\<^bsub>\\<^esub>. B \\<^sup>\ (T\<^bsub>o\<^esub> \\<^sup>\ \x\<^bsub>\\<^esub>. B) = \<^bold>S ?\ ?C\<^sub>1" by simp ultimately show ?thesis using \is_tautology ?C\<^sub>1\ by blast qed qed simp have "\
2": "\ B \\<^sup>\ (T\<^bsub>o\<^esub> \\<^sup>\ B)" proof (intro rule_P(2)) show "is_tautologous (B \\<^sup>\ T\<^bsub>o\<^esub> \\<^sup>\ B)" proof - let ?\ = "{(\, o) \ B}" from assms(2) have "is_pwff_substitution ?\" using pwffs_subset_of_wffso by auto moreover have "B \\<^sup>\ T\<^bsub>o\<^esub> \\<^sup>\ B = \<^bold>S ?\ ?C\<^sub>1" by simp ultimately show ?thesis using \is_tautology ?C\<^sub>1\ by blast qed qed simp have "\ T\<^bsub>o\<^esub>" by (fact true_is_derivable) then have "\
3": "\ \x\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>" using Gen by simp have "\
4": "\ \x\<^bsub>\\<^esub>. T\<^bsub>o\<^esub> \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ \x\<^bsub>\\<^esub>. B)" proof (intro rule_P(1)[where \ = "{\x\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>}"]) show "is_tautologous ([\x\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>] \\<^sup>\\<^sub>\ (\x\<^bsub>\\<^esub>. T\<^bsub>o\<^esub> \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ \x\<^bsub>\\<^esub>. B)))" proof - let ?\ = "{(\, o) \ \x\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>, (\, o) \ \x\<^bsub>\\<^esub>. B}" from assms(2) have "is_pwff_substitution ?\" using pwffs_subset_of_wffso by auto moreover have "[\x\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>] \\<^sup>\\<^sub>\ (\x\<^bsub>\\<^esub>. T\<^bsub>o\<^esub> \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ \x\<^bsub>\\<^esub>. B)) = \<^bold>S ?\ ?C\<^sub>2" by simp ultimately show ?thesis using \is_tautology ?C\<^sub>2\ by blast qed qed (use "\
3" in fastforce)+ have "\
5": "\ T\<^bsub>o\<^esub> \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ B)" proof (intro rule_P(2)) show "is_tautologous (T\<^bsub>o\<^esub> \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ B))" proof - let ?\ = "{(\, o) \ B}" and ?C = "T\<^bsub>o\<^esub> \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)" have "is_tautology ?C" using \\<^sub>B_simps by simp moreover from assms(2) have "is_pwff_substitution ?\" using pwffs_subset_of_wffso by auto moreover have "T\<^bsub>o\<^esub> \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ B) = \<^bold>S ?\ ?C" by simp ultimately show ?thesis by blast qed qed simp from "\
4" and "\
5" have "\
6": "\ \x\<^bsub>\\<^esub>. (F\<^bsub>o\<^esub> \\<^sup>\ B) \\<^sup>\ (F\<^bsub>o\<^esub> \\<^sup>\ \x\<^bsub>\\<^esub>. B)" unfolding equivalence_def by (rule rule_R[where p = "[\,\,\,\]"]) force+ from "\
1" and "\
2" have "\
7": "\ \x\<^bsub>\\<^esub>. (T\<^bsub>o\<^esub> \\<^sup>\ B) \\<^sup>\ (T\<^bsub>o\<^esub> \\<^sup>\ \x\<^bsub>\\<^esub>. B)" unfolding equivalence_def by (rule rule_R[where p = "[\,\,\,\]"]) force+ obtain p where "(p, o) \ vars B" and "p \ x" using fresh_var_existence and vars_form_finiteness by (metis finite_insert insert_iff) from \(p, o) \ vars B\ have sub: "\<^bold>S {(p, o) \ C} B = B" for C using free_var_singleton_substitution_neutrality and free_vars_in_all_vars by blast have "\
8": "\ \x\<^bsub>\\<^esub>. (p\<^bsub>o\<^esub> \\<^sup>\ B) \\<^sup>\ (p\<^bsub>o\<^esub> \\<^sup>\ \x\<^bsub>\\<^esub>. B)" (is \\ ?C\<^sub>3\) proof - from sub and "\
7" have "\ \<^bold>S {(p, o) \ T\<^bsub>o\<^esub>} ?C\<^sub>3" using \p \ x\ by auto moreover from sub and "\
6" have "\ \<^bold>S {(p, o) \ F\<^bsub>o\<^esub>} ?C\<^sub>3" using \p \ x\ by auto moreover from assms(2) have "?C\<^sub>3 \ wffs\<^bsub>o\<^esub>" using pwffs_subset_of_wffso by auto ultimately show ?thesis by (rule Cases) qed then show ?thesis proof - let ?\ = "{(p, o) \ A}" from assms(1) have "is_substitution ?\" using pwffs_subset_of_wffso by auto moreover have " \v \ fmdom' ?\. var_name v \ free_var_names ({}::form set) \ is_free_for (?\ $$! v) v ?C\<^sub>3" proof fix v assume "v \ fmdom' ?\" then have "v = (p, o)" by simp with assms(3) and \(p, o) \ vars B\ have "is_free_for (?\ $$! v) v ?C\<^sub>3" using occurs_in_vars by (intro is_free_for_in_imp is_free_for_in_forall is_free_for_in_equivalence) auto moreover have "var_name v \ free_var_names ({}::form set)" by simp ultimately show "var_name v \ free_var_names ({}::form set) \ is_free_for (?\ $$! v) v ?C\<^sub>3" unfolding \v = (p, o)\ by blast qed moreover have "?\ \ {$$}" by simp ultimately have "\ \<^bold>S ?\ ?C\<^sub>3" by (rule Sub[OF "\
8"]) moreover have "\<^bold>S ?\ ?C\<^sub>3 = \x\<^bsub>\\<^esub>. (A \\<^sup>\ B) \\<^sup>\ (A \\<^sup>\ \x\<^bsub>\\<^esub>. B)" using \p \ x\ and sub[of A] by simp ultimately show ?thesis by (simp only:) qed qed proposition prop_5237: assumes "is_hyps \" and "\ \ A \\<^sup>\ B" and "(x, \) \ free_vars ({A} \ \)" shows "\ \ A \\<^sup>\ (\x\<^bsub>\\<^esub>. B)" proof - have "\ \ A \\<^sup>\ B" by fact with assms(3) have "\ \ \x\<^bsub>\\<^esub>. (A \\<^sup>\ B)" using Gen by simp moreover have "\ \ \x\<^bsub>\\<^esub>. (A \\<^sup>\ B) \\<^sup>\ (A \\<^sup>\ (\x\<^bsub>\\<^esub>. B))" proof - from assms(2) have "A \ wffs\<^bsub>o\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" using hyp_derivable_form_is_wffso by (blast dest: wffs_from_imp_op)+ with assms(1,3) show ?thesis using prop_5237_aux and derivability_implies_hyp_derivability by simp qed ultimately show ?thesis by (rule Equality_Rules(1)) qed lemmas "\\" = prop_5237 (* synonym *) corollary generalized_prop_5237: assumes "is_hyps \" and "\ \ A \\<^sup>\ B" and "\v \ S. v \ free_vars ({A} \ \)" and "lset vs = S" shows "\ \ A \\<^sup>\ (\\<^sup>\\<^sub>\ vs B)" using assms proof (induction vs arbitrary: S) case Nil then show ?case by simp next case (Cons v vs) obtain x and \ where "v = (x, \)" by fastforce from Cons.prems(3) have *: "\v' \ S. v' \ free_vars ({A} \ \)" by blast then show ?case proof (cases "v \ lset vs") case True with Cons.prems(4) have "lset vs = S" by auto with assms(1,2) and * have "\ \ A \\<^sup>\ \\<^sup>\\<^sub>\ vs B" by (fact Cons.IH) with True and \lset vs = S\ and \v = (x, \)\ and * have "\ \ A \\<^sup>\ (\x\<^bsub>\\<^esub>. \\<^sup>\\<^sub>\ vs B)" using prop_5237[OF assms(1)] by simp with \v = (x, \)\ show ?thesis by simp next case False with \lset (v # vs) = S\ have "lset vs = S - {v}" by auto moreover from * have "\v' \ S - {v}. v' \ free_vars ({A} \ \)" by blast ultimately have "\ \ A \\<^sup>\ \\<^sup>\\<^sub>\ vs B" using assms(1,2) by (intro Cons.IH) moreover from Cons.prems(4) and \v = (x, \)\ and * have "(x, \) \ free_vars ({A} \ \)" by auto ultimately have "\ \ A \\<^sup>\ (\x\<^bsub>\\<^esub>. \\<^sup>\\<^sub>\ vs B)" using assms(1) by (intro prop_5237) with \v = (x, \)\ show ?thesis by simp qed qed end subsection \Proposition 5238\ context begin private lemma prop_5238_aux: assumes "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" shows "\ ((\x\<^bsub>\\<^esub>. A) =\<^bsub>\\\\<^esub> (\x\<^bsub>\\<^esub>. B)) \\<^sup>\ \x\<^bsub>\\<^esub>. (A =\<^bsub>\\<^esub> B)" proof - have "\
1": " \ (\\<^bsub>\\\\<^esub> =\<^bsub>\\\\<^esub> \\<^bsub>\\\\<^esub>) \\<^sup>\ \\\<^bsub>\\<^esub>. (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>)" (is \\ _ \\<^sup>\ \\\<^bsub>\\<^esub>. ?C\<^sub>1\) by (fact axiom_is_derivable_from_no_hyps[OF axiom_3]) then have "\
2": " \ (\\<^bsub>\\\\<^esub> =\<^bsub>\\\\<^esub> \\<^bsub>\\\\<^esub>) \\<^sup>\ \x\<^bsub>\\<^esub>. (\\<^bsub>\\\\<^esub> \ x\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ x\<^bsub>\\<^esub>)" (is \\ ?C\<^sub>2\) proof (cases "x = \") case True with "\
1" show ?thesis by (simp only:) next case False have "?C\<^sub>1 \ wffs\<^bsub>o\<^esub>" by blast moreover from False have "(x, \) \ free_vars ?C\<^sub>1" by simp moreover have "is_free_for (x\<^bsub>\\<^esub>) (\, \) ?C\<^sub>1" by (intro is_free_for_in_equality is_free_for_to_app) simp_all ultimately have "\ \\\<^bsub>\\<^esub>. ?C\<^sub>1 =\<^bsub>\\o\<^esub> \x\<^bsub>\\<^esub>. (\<^bold>S {(\, \) \ x\<^bsub>\\<^esub>} ?C\<^sub>1)" by (rule "\") from "\
1" and this show ?thesis by (rule rule_R[where p = "[\,\]"]) force+ qed then have "\
3": " \ ((\x\<^bsub>\\<^esub>. A) =\<^bsub>\\\\<^esub> (\x\<^bsub>\\<^esub>. B)) \\<^sup>\ \x\<^bsub>\\<^esub>. ((\x\<^bsub>\\<^esub>. A) \ x\<^bsub>\\<^esub> =\<^bsub>\\<^esub> (\x\<^bsub>\\<^esub>. B) \ x\<^bsub>\\<^esub>)" proof - let ?\ = "{(\, \\\) \ \x\<^bsub>\\<^esub>. A, (\, \\\) \ \x\<^bsub>\\<^esub>. B}" have "\x\<^bsub>\\<^esub>. A \ wffs\<^bsub>\\\\<^esub>" and "\x\<^bsub>\\<^esub>. B \ wffs\<^bsub>\\\\<^esub>" by (blast intro: assms(1,2))+ then have "is_substitution ?\" by simp moreover have " \v \ fmdom' ?\. var_name v \ free_var_names ({}::form set) \ is_free_for (?\ $$! v) v ?C\<^sub>2" proof fix v assume "v \ fmdom' ?\" then consider (a) "v = (\, \\\)" | (b) "v = (\, \\\)" by fastforce then show "var_name v \ free_var_names ({}::form set) \ is_free_for (?\ $$! v) v ?C\<^sub>2" proof cases case a have "(x, \) \ free_vars (\x\<^bsub>\\<^esub>. A)" by simp then have "is_free_for (\x\<^bsub>\\<^esub>. A) (\, \\\) ?C\<^sub>2" unfolding equivalence_def by (intro is_free_for_in_equality is_free_for_in_forall is_free_for_to_app, simp_all) with a show ?thesis by force next case b have "(x, \) \ free_vars (\x\<^bsub>\\<^esub>. B)" by simp then have "is_free_for (\x\<^bsub>\\<^esub>. B) (\, \\\) ?C\<^sub>2" unfolding equivalence_def by (intro is_free_for_in_equality is_free_for_in_forall is_free_for_to_app, simp_all) with b show ?thesis by force qed qed moreover have "?\ \ {$$}" by simp ultimately have "\ \<^bold>S ?\ ?C\<^sub>2" by (rule Sub[OF "\
2"]) then show ?thesis by simp qed then have "\
4": "\ ((\x\<^bsub>\\<^esub>. A) =\<^bsub>\\\\<^esub> (\x\<^bsub>\\<^esub>. B)) \\<^sup>\ \x\<^bsub>\\<^esub>. (A =\<^bsub>\\<^esub> (\x\<^bsub>\\<^esub>. B) \ x\<^bsub>\\<^esub>)" proof - have "\ (\x\<^bsub>\\<^esub>. A) \ x\<^bsub>\\<^esub> =\<^bsub>\\<^esub> A" using prop_5208[where vs = "[(x, \)]"] and assms(1) by simp from "\
3" and this show ?thesis by (rule rule_R[where p = "[\,\,\,\,\]"]) force+ qed then show ?thesis proof - have "\ (\x\<^bsub>\\<^esub>. B) \ x\<^bsub>\\<^esub> =\<^bsub>\\<^esub> B" using prop_5208[where vs = "[(x, \)]"] and assms(2) by simp from "\
4" and this show ?thesis by (rule rule_R[where p = "[\,\,\,\]"]) force+ qed qed proposition prop_5238: assumes "vs \ []" and "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" shows "\ \\<^sup>\\<^sub>\ vs A =\<^bsub>foldr (\) (map var_type vs) \\<^esub> \\<^sup>\\<^sub>\ vs B \\<^sup>\ \\<^sup>\\<^sub>\ vs (A =\<^bsub>\\<^esub> B)" using assms proof (induction vs arbitrary: A B \ rule: rev_nonempty_induct) case (single v) obtain x and \ where "v = (x, \)" by fastforce from single.prems have " \\<^sup>\\<^sub>\ vs A =\<^bsub>foldr (\) (map var_type vs) \\<^esub> \\<^sup>\\<^sub>\ vs B \\<^sup>\ \\<^sup>\\<^sub>\ vs (A =\<^bsub>\\<^esub> B) \ wffs\<^bsub>o\<^esub>" by blast with single.prems and \v = (x, \)\ show ?case using prop_5238_aux by simp next case (snoc v vs) obtain x and \ where "v = (x, \)" by fastforce from snoc.prems have "\x\<^bsub>\\<^esub>. A \ wffs\<^bsub>\\\\<^esub>" and "\x\<^bsub>\\<^esub>. B \ wffs\<^bsub>\\\\<^esub>" by auto then have " \ \\<^sup>\\<^sub>\ vs (\x\<^bsub>\\<^esub>. A) =\<^bsub>foldr (\) (map var_type vs) (\\\)\<^esub> \\<^sup>\\<^sub>\ vs (\x\<^bsub>\\<^esub>. B) \\<^sup>\ \\<^sup>\\<^sub>\ vs ((\x\<^bsub>\\<^esub>. A) =\<^bsub>\\\\<^esub> (\x\<^bsub>\\<^esub>. B))" by (fact snoc.IH) moreover from snoc.prems have "\ \x\<^bsub>\\<^esub>. A =\<^bsub>\\\\<^esub> \x\<^bsub>\\<^esub>. B \\<^sup>\ \x\<^bsub>\\<^esub>. (A =\<^bsub>\\<^esub> B)" by (fact prop_5238_aux) ultimately have " \ \\<^sup>\\<^sub>\ vs (\x\<^bsub>\\<^esub>. A) =\<^bsub>foldr (\) (map var_type vs) (\\\)\<^esub> \\<^sup>\\<^sub>\ vs (\x\<^bsub>\\<^esub>. B) \\<^sup>\ \\<^sup>\\<^sub>\ vs \x\<^bsub>\\<^esub>. (A =\<^bsub>\\<^esub> B)" unfolding equivalence_def proof (induction rule: rule_R[where p = "\ # foldr (\_. (@) [\,\]) vs []"]) case occ_subform then show ?case using innermost_subform_in_generalized_forall[OF snoc.hyps] and is_subform_at.simps(3) by fastforce next case replacement then show ?case using innermost_replacement_in_generalized_forall[OF snoc.hyps] and is_replacement_at_implies_in_positions and replace_right_app by force qed with \v = (x, \)\ show ?case by simp qed end subsection \Proposition 5239\ lemma replacement_derivability: assumes "C \ wffs\<^bsub>\\<^esub>" and "A \\<^bsub>p\<^esub> C" and "\ A =\<^bsub>\\<^esub> B" and "C\p \ B\ \ D" shows "\ C =\<^bsub>\\<^esub> D" using assms proof (induction arbitrary: D p) case (var_is_wff \ x) from var_is_wff.prems(1) have "p = []" and "A = x\<^bsub>\\<^esub>" by (auto elim: is_subform_at.elims(2)) with var_is_wff.prems(2) have "\ = \" using hyp_derivable_form_is_wffso and wff_has_unique_type and wffs_from_equality by blast moreover from \p = []\ and var_is_wff.prems(3) have "D = B" using is_replacement_at_minimal_change(1) and is_subform_at.simps(1) by iprover ultimately show ?case using \A = x\<^bsub>\\<^esub>\ and var_is_wff.prems(2) by (simp only:) next case (con_is_wff \ c) from con_is_wff.prems(1) have "p = []" and "A = \c\\<^bsub>\\<^esub>" by (auto elim: is_subform_at.elims(2)) with con_is_wff.prems(2) have "\ = \" using hyp_derivable_form_is_wffso and wff_has_unique_type by (meson wffs_from_equality wffs_of_type_intros(2)) moreover from \p = []\ and con_is_wff.prems(3) have "D = B" using is_replacement_at_minimal_change(1) and is_subform_at.simps(1) by iprover ultimately show ?case using \A = \c\\<^bsub>\\<^esub>\ and con_is_wff.prems(2) by (simp only:) next case (app_is_wff \ \ C\<^sub>1 C\<^sub>2) from app_is_wff.prems(1) consider (a) "p = []" | (b) "\p'. p = \ # p' \ A \\<^bsub>p'\<^esub> C\<^sub>1" | (c) "\p'. p = \ # p' \ A \\<^bsub>p'\<^esub> C\<^sub>2" using subforms_from_app by blast then show ?case proof cases case a with app_is_wff.prems(1) have "A = C\<^sub>1 \ C\<^sub>2" by simp moreover from a and app_is_wff.prems(3) have "D = B" using is_replacement_at_minimal_change(1) and at_top_is_self_subform by blast moreover from \A = C\<^sub>1 \ C\<^sub>2\ and \D = B\ and app_is_wff.hyps(1,2) and assms(3) have "\ = \" using hyp_derivable_form_is_wffso and wff_has_unique_type by (blast dest: wffs_from_equality) ultimately show ?thesis using assms(3) by (simp only:) next case b then obtain p' where "p = \ # p'" and "A \\<^bsub>p'\<^esub> C\<^sub>1" by blast moreover obtain D\<^sub>1 where "D = D\<^sub>1 \ C\<^sub>2" and "C\<^sub>1\p' \ B\ \ D\<^sub>1" using app_is_wff.prems(3) and \p = \ # p'\ by (force dest: is_replacement_at.cases) ultimately have "\ C\<^sub>1 =\<^bsub>\\\\<^esub> D\<^sub>1" using app_is_wff.IH(1) and assms(3) by blast moreover have "\ C\<^sub>2 =\<^bsub>\\<^esub> C\<^sub>2" by (fact prop_5200[OF app_is_wff.hyps(2)]) ultimately have "\ C\<^sub>1 \ C\<^sub>2 =\<^bsub>\\<^esub> D\<^sub>1 \ C\<^sub>2" using Equality_Rules(4) by (simp only:) with \D = D\<^sub>1 \ C\<^sub>2\ show ?thesis by (simp only:) next case c then obtain p' where "p = \ # p'" and "A \\<^bsub>p'\<^esub> C\<^sub>2" by blast moreover obtain D\<^sub>2 where "D = C\<^sub>1 \ D\<^sub>2" and "C\<^sub>2\p' \ B\ \ D\<^sub>2" using app_is_wff.prems(3) and \p = \ # p'\ by (force dest: is_replacement_at.cases) ultimately have "\ C\<^sub>2 =\<^bsub>\\<^esub> D\<^sub>2" using app_is_wff.IH(2) and assms(3) by blast moreover have "\ C\<^sub>1 =\<^bsub>\\\\<^esub> C\<^sub>1" by (fact prop_5200[OF app_is_wff.hyps(1)]) ultimately have "\ C\<^sub>1 \ C\<^sub>2 =\<^bsub>\\<^esub> C\<^sub>1 \ D\<^sub>2" using Equality_Rules(4) by (simp only:) with \D = C\<^sub>1 \ D\<^sub>2\ show ?thesis by (simp only:) qed next case (abs_is_wff \ C' \ x) from abs_is_wff.prems(1) consider (a) "p = []" | (b) "\p'. p = \ # p' \ A \\<^bsub>p'\<^esub> C'" using subforms_from_abs by blast then show ?case proof cases case a with abs_is_wff.prems(1) have "A = \x\<^bsub>\\<^esub>. C'" by simp moreover from a and abs_is_wff.prems(3) have "D = B" using is_replacement_at_minimal_change(1) and at_top_is_self_subform by blast moreover from \A = \x\<^bsub>\\<^esub>. C'\ and \D = B\ and abs_is_wff.hyps(1) and assms(3) have "\ = \\\" using hyp_derivable_form_is_wffso and wff_has_unique_type by (blast dest: wffs_from_abs wffs_from_equality) ultimately show ?thesis using assms(3) by (simp only:) next case b then obtain p' where "p = \ # p'" and "A \\<^bsub>p'\<^esub> C'" by blast moreover obtain D' where "D = \x\<^bsub>\\<^esub>. D'" and "C'\p' \ B\ \ D'" using abs_is_wff.prems(3) and \p = \ # p'\ by (force dest: is_replacement_at.cases) ultimately have "\ C' =\<^bsub>\\<^esub> D'" using abs_is_wff.IH and assms(3) by blast then have "\ \x\<^bsub>\\<^esub>. C' =\<^bsub>\\\\<^esub> \x\<^bsub>\\<^esub>. D'" proof - from \\ C' =\<^bsub>\\<^esub> D'\ have "\ \x\<^bsub>\\<^esub>. (C' =\<^bsub>\\<^esub> D')" using Gen by simp moreover from \\ C' =\<^bsub>\\<^esub> D'\ and abs_is_wff.hyps have "D' \ wffs\<^bsub>\\<^esub>" using hyp_derivable_form_is_wffso by (blast dest: wffs_from_equality) with abs_is_wff.hyps have "\ (\x\<^bsub>\\<^esub>. C' =\<^bsub>\\\\<^esub> \x\<^bsub>\\<^esub>. D') \\<^sup>\ \x\<^bsub>\\<^esub>. (C' =\<^bsub>\\<^esub> D')" using prop_5238[where vs = "[(x, \)]"] by simp ultimately show ?thesis using Equality_Rules(1,2) unfolding equivalence_def by blast qed with \D = \x\<^bsub>\\<^esub>. D'\ show ?thesis by (simp only:) qed qed context begin private lemma prop_5239_aux_1: assumes "p \ positions (\\<^sup>\\<^sub>\ (FVar v) (map FVar vs))" and "p \ replicate (length vs) \" shows " (\A B. A \ B \\<^bsub>p\<^esub> (\\<^sup>\\<^sub>\ (FVar v) (map FVar vs))) \ (\v \ lset vs. occurs_at v p (\\<^sup>\\<^sub>\ (FVar v) (map FVar vs)))" using assms proof (induction vs arbitrary: p rule: rev_induct) case Nil then show ?case using surj_pair[of v] by fastforce next case (snoc v' vs) from snoc.prems(1) consider (a) "p = []" | (b) "p = [\]" | (c) "\p' \ positions (\\<^sup>\\<^sub>\ (FVar v) (map FVar vs)). p = \ # p'" using surj_pair[of v'] by fastforce then show ?case proof cases case c then obtain p' where "p' \ positions (\\<^sup>\\<^sub>\ (FVar v) (map FVar vs))" and "p = \ # p'" by blast from \p = \ # p'\ and snoc.prems(2) have "p' \ replicate (length vs) \" by force then have " (\A B. A \ B \\<^bsub>p'\<^esub> \\<^sup>\\<^sub>\ (FVar v) (map FVar vs)) \ (\v \ lset vs. occurs_at v p' (\\<^sup>\\<^sub>\ (FVar v) (map FVar vs)))" using \p' \ positions (\\<^sup>\\<^sub>\ (FVar v) (map FVar vs))\ and snoc.IH by simp with \p = \ # p'\ show ?thesis by auto qed simp_all qed private lemma prop_5239_aux_2: assumes "t \ lset vs \ vars C" and "C\p \ (\\<^sup>\\<^sub>\ (FVar t) (map FVar vs))\ \ G" and "C\p \ (\\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs A) (map FVar vs))\ \ G'" shows "\<^bold>S {t \ \\<^sup>\\<^sub>\ vs A} G = G'" (is \\<^bold>S ?\ G = G'\) proof - have "\<^bold>S ?\ (\\<^sup>\\<^sub>\ (FVar t) (map FVar vs)) = \\<^sup>\\<^sub>\ (\<^bold>S ?\ (FVar t)) (map (\v'. \<^bold>S ?\ v') (map FVar vs))" using generalized_app_substitution by blast moreover have "\<^bold>S ?\ (FVar t) = \\<^sup>\\<^sub>\ vs A" using surj_pair[of t] by fastforce moreover from assms(1) have "map (\v'. \<^bold>S ?\ v') (map FVar vs) = map FVar vs" by (induction vs) auto ultimately show ?thesis using assms proof (induction C arbitrary: G G' p) case (FVar v) from FVar.prems(5) have "p = []" and "G = \\<^sup>\\<^sub>\ (FVar t) (map FVar vs)" by (blast dest: is_replacement_at.cases)+ moreover from FVar.prems(6) and \p = []\ have "G' = \\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs A) (map FVar vs)" by (blast dest: is_replacement_at.cases) ultimately show ?case using FVar.prems(1-3) by (simp only:) next case (FCon k) from FCon.prems(5) have "p = []" and "G = \\<^sup>\\<^sub>\ (FVar t) (map FVar vs)" by (blast dest: is_replacement_at.cases)+ moreover from FCon.prems(6) and \p = []\ have "G' = \\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs A) (map FVar vs)" by (blast dest: is_replacement_at.cases) ultimately show ?case using FCon.prems(1-3) by (simp only:) next case (FApp C\<^sub>1 C\<^sub>2) from FApp.prems(4) have "t \ lset vs \ vars C\<^sub>1" and "t \ lset vs \ vars C\<^sub>2" by auto consider (a) "p = []" | (b) "\p'. p = \ # p'" | (c) "\p'. p = \ # p'" by (metis direction.exhaust list.exhaust) then show ?case proof cases case a with FApp.prems(5) have "G = \\<^sup>\\<^sub>\ (FVar t) (map FVar vs)" by (blast dest: is_replacement_at.cases) moreover from FApp.prems(6) and \p = []\ have "G' = \\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs A) (map FVar vs)" by (blast dest: is_replacement_at.cases) ultimately show ?thesis using FApp.prems(1-3) by (simp only:) next case b then obtain p' where "p = \ # p'" by blast with FApp.prems(5) obtain G\<^sub>1 where "G = G\<^sub>1 \ C\<^sub>2" and "C\<^sub>1\p' \ (\\<^sup>\\<^sub>\ (FVar t) (map FVar vs))\ \ G\<^sub>1" by (blast elim: is_replacement_at.cases) moreover from \p = \ # p'\ and FApp.prems(6) obtain G'\<^sub>1 where "G' = G'\<^sub>1 \ C\<^sub>2" and "C\<^sub>1\p' \ (\\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs A) (map FVar vs))\ \ G'\<^sub>1" by (blast elim: is_replacement_at.cases) moreover from \t \ lset vs \ vars C\<^sub>2\ have "\<^bold>S {t \ \\<^sup>\\<^sub>\ vs A} C\<^sub>2 = C\<^sub>2" using surj_pair[of t] and free_var_singleton_substitution_neutrality by (simp add: vars_is_free_and_bound_vars) ultimately show ?thesis using FApp.IH(1)[OF FApp.prems(1-3) \t \ lset vs \ vars C\<^sub>1\] by simp next case c then obtain p' where "p = \ # p'" by blast with FApp.prems(5) obtain G\<^sub>2 where "G = C\<^sub>1 \ G\<^sub>2" and "C\<^sub>2\p' \ (\\<^sup>\\<^sub>\ (FVar t) (map FVar vs))\ \ G\<^sub>2" by (blast elim: is_replacement_at.cases) moreover from \p = \ # p'\ and FApp.prems(6) obtain G'\<^sub>2 where "G' = C\<^sub>1 \ G'\<^sub>2" and "C\<^sub>2\p' \ (\\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs A) (map FVar vs))\ \ G'\<^sub>2" by (blast elim: is_replacement_at.cases) moreover from \t \ lset vs \ vars C\<^sub>1\ have "\<^bold>S {t \ \\<^sup>\\<^sub>\ vs A} C\<^sub>1 = C\<^sub>1" using surj_pair[of t] and free_var_singleton_substitution_neutrality by (simp add: vars_is_free_and_bound_vars) ultimately show ?thesis using FApp.IH(2)[OF FApp.prems(1-3) \t \ lset vs \ vars C\<^sub>2\] by simp qed next case (FAbs v C') from FAbs.prems(4) have "t \ lset vs \ vars C'" and "t \ v" using vars_form.elims by blast+ from FAbs.prems(5) consider (a) "p = []" | (b) "\p'. p = \ # p'" using is_replacement_at.simps by blast then show ?case proof cases case a with FAbs.prems(5) have "G = \\<^sup>\\<^sub>\ (FVar t) (map FVar vs)" by (blast dest: is_replacement_at.cases) moreover from FAbs.prems(6) and \p = []\ have "G' = \\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs A) (map FVar vs)" by (blast dest: is_replacement_at.cases) ultimately show ?thesis using FAbs.prems(1-3) by (simp only:) next case b then obtain p' where "p = \ # p'" by blast then obtain G\<^sub>1 where "G = FAbs v G\<^sub>1" and "C'\p' \ (\\<^sup>\\<^sub>\ (FVar t) (map FVar vs))\ \ G\<^sub>1" using FAbs.prems(5) by (blast elim: is_replacement_at.cases) moreover from \p = \ # p'\ and FAbs.prems(6) obtain G'\<^sub>1 where "G' = FAbs v G'\<^sub>1" and "C'\p' \ (\\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs A) (map FVar vs))\ \ G'\<^sub>1" by (blast elim: is_replacement_at.cases) ultimately have "\<^bold>S {t \ \\<^sup>\\<^sub>\ vs A} G\<^sub>1 = G'\<^sub>1" using FAbs.IH[OF FAbs.prems(1-3) \t \ lset vs \ vars C'\] by simp with \G = FAbs v G\<^sub>1\ and \G' = FAbs v G'\<^sub>1\ and \t \ v\ show ?thesis using surj_pair[of v] by fastforce qed qed qed private lemma prop_5239_aux_3: assumes "t \ lset vs \ vars {A, C}" and "C\p \ (\\<^sup>\\<^sub>\ (FVar t) (map FVar vs))\ \ G" and "occurs_at t p' G" shows "p' = p @ replicate (length vs) \" (is \p' = ?p\<^sub>t\) proof (cases "vs = []") case True then have "t \ vars C" using assms(1) by auto moreover from True and assms(2) have "C\p \ FVar t\ \ G" by force ultimately show ?thesis using assms(3) and True and fresh_var_replacement_position_uniqueness by simp next case False show ?thesis proof (rule ccontr) assume "p' \ ?p\<^sub>t" have "\ prefix ?p\<^sub>t p" by (simp add: False) from assms(3) have "p' \ positions G" using is_subform_implies_in_positions by fastforce from assms(2) have "?p\<^sub>t \ positions G" using is_replacement_at_minimal_change(1) and is_subform_at_transitivity and is_subform_implies_in_positions and leftmost_subform_in_generalized_app by (metis length_map) from assms(2) have "occurs_at t ?p\<^sub>t G" unfolding occurs_at_def using is_replacement_at_minimal_change(1) and is_subform_at_transitivity and leftmost_subform_in_generalized_app by (metis length_map) moreover from assms(2) and \p' \ positions G\ have *: " subform_at C p' = subform_at G p'" if "\ prefix p' p" and "\ prefix p p'" using is_replacement_at_minimal_change(2) by (simp add: that(1,2)) ultimately show False proof (cases "\ prefix p' p \ \ prefix p p'") case True with assms(3) and * have "occurs_at t p' C" using is_replacement_at_occurs[OF assms(2)] by blast then have "t \ vars C" using is_subform_implies_in_positions and occurs_in_vars by fastforce with assms(1) show ?thesis by simp next case False then consider (a) "prefix p' p" | (b) "prefix p p'" by blast then show ?thesis proof cases case a with \occurs_at t ?p\<^sub>t G\ and \p' \ ?p\<^sub>t\ and assms(3) show ?thesis unfolding occurs_at_def using loop_subform_impossibility by (metis prefix_order.dual_order.order_iff_strict prefix_prefix) next case b have "strict_prefix p' ?p\<^sub>t" proof (rule ccontr) assume "\ strict_prefix p' ?p\<^sub>t" then consider (b\<^sub>1) "p' = ?p\<^sub>t" | (b\<^sub>2) "strict_prefix ?p\<^sub>t p'" | (b\<^sub>3) "\ prefix p' ?p\<^sub>t" and "\ prefix ?p\<^sub>t p'" by fastforce then show False proof cases case b\<^sub>1 with \p' \ ?p\<^sub>t\ show ?thesis by contradiction next case b\<^sub>2 with \occurs_at t ?p\<^sub>t G\ and assms(3) show ?thesis using loop_subform_impossibility by blast next case b\<^sub>3 from b obtain p'' where "p' = p @ p''" and "p'' \ positions (\\<^sup>\\<^sub>\ (FVar t) (map FVar vs))" using is_replacement_at_new_positions and \p' \ positions G\ and assms(2) by blast moreover have "p'' \ replicate (length vs) \" using \p' = p @ p''\ and \p' \ ?p\<^sub>t\ by blast ultimately consider (b\<^sub>3\<^sub>_\<^sub>1) "\F\<^sub>1 F\<^sub>2. F\<^sub>1 \ F\<^sub>2 \\<^bsub>p''\<^esub> (\\<^sup>\\<^sub>\ (FVar t) (map FVar vs))" | (b\<^sub>3\<^sub>_\<^sub>2) "\v \ lset vs. occurs_at v p'' (\\<^sup>\\<^sub>\ (FVar t) (map FVar vs))" using prop_5239_aux_1 and b\<^sub>3(1,2) and is_replacement_at_occurs and leftmost_subform_in_generalized_app_replacement by (metis (no_types, opaque_lifting) length_map prefix_append) then show ?thesis proof cases case b\<^sub>3\<^sub>_\<^sub>1 with assms(2) and \p' = p @ p''\ have "\F\<^sub>1 F\<^sub>2. F\<^sub>1 \ F\<^sub>2 \\<^bsub>p'\<^esub> G" using is_replacement_at_minimal_change(1) and is_subform_at_transitivity by meson with \occurs_at t p' G\ show ?thesis using is_subform_at_uniqueness by fastforce next case b\<^sub>3\<^sub>_\<^sub>2 with assms(2) and \p' = p @ p''\ have "\v \ lset vs. occurs_at v p' G" unfolding occurs_at_def using is_replacement_at_minimal_change(1) and is_subform_at_transitivity by meson with assms(1,3) show ?thesis using is_subform_at_uniqueness by fastforce qed qed qed with \occurs_at t ?p\<^sub>t G\ and assms(3) show ?thesis using loop_subform_impossibility by blast qed qed qed qed private lemma prop_5239_aux_4: assumes "t \ lset vs \ vars {A, C}" and "A \\<^bsub>p\<^esub> C" and "lset vs \ capture_exposed_vars_at p C A" and "C\p \ (\\<^sup>\\<^sub>\ (FVar t) (map FVar vs))\ \ G" shows "is_free_for (\\<^sup>\\<^sub>\ vs A) t G" unfolding is_free_for_def proof (intro ballI impI) let ?p\<^sub>t = "p @ replicate (length vs) \" from assms(4) have "FVar t \\<^bsub>?p\<^sub>t\<^esub> G" using is_replacement_at_minimal_change(1) and is_subform_at_transitivity and leftmost_subform_in_generalized_app by (metis length_map) fix v' and p' assume "v' \ free_vars (\\<^sup>\\<^sub>\ vs A)" and "p' \ positions G" and "is_free_at t p' G" have "v' \ binders_at G ?p\<^sub>t" proof - have "free_vars (\\<^sup>\\<^sub>\ vs A) = free_vars A - lset vs" by (fact free_vars_of_generalized_abs) also from assms(2,3) have "\ \ free_vars A - (binders_at C p \ free_vars A)" using capture_exposed_vars_at_alt_def and is_subform_implies_in_positions by fastforce also have "\ = free_vars A - (binders_at G p \ free_vars A)" using assms(2,4) is_replacement_at_binders is_subform_implies_in_positions by blast finally have "free_vars (\\<^sup>\\<^sub>\ vs A) \ free_vars A - (binders_at G p \ free_vars A)" . moreover have "binders_at (\\<^sup>\\<^sub>\ (FVar t) (map FVar vs)) (replicate (length vs) \) = {}" by (induction vs rule: rev_induct) simp_all with assms(4) have "binders_at G ?p\<^sub>t = binders_at G p" using binders_at_concat and is_replacement_at_minimal_change(1) by blast ultimately show ?thesis using \v' \ free_vars (\\<^sup>\\<^sub>\ vs A)\ by blast qed moreover have "p' = ?p\<^sub>t" by ( fact prop_5239_aux_3 [OF assms(1,4) \is_free_at t p' G\[unfolded is_free_at_def, THEN conjunct1]] ) ultimately show "\ in_scope_of_abs v' p' G" using binders_at_alt_def[OF \p' \ positions G\] and in_scope_of_abs_alt_def by auto qed proposition prop_5239: assumes "is_rule_R_app p D C (A =\<^bsub>\\<^esub> B)" and "lset vs = {(x, \) | x \ p' E. strict_prefix p' p \ \x\<^bsub>\\<^esub>. E \\<^bsub>p'\<^esub> C \ (x, \) \ free_vars (A =\<^bsub>\\<^esub> B)}" shows "\ \\<^sup>\\<^sub>\ vs (A =\<^bsub>\\<^esub> B) \\<^sup>\ (C \\<^sup>\ D)" proof - let ?\ = "foldr (\) (map var_type vs) \" obtain t where "(t, ?\) \ lset vs \ vars {A,B,C,D}" using fresh_var_existence and vars_form_set_finiteness by (metis List.finite_set finite.simps finite_UnI) from assms(1) have "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" and "A \\<^bsub>p\<^esub> C" using wffs_from_equality[OF equality_wff] by simp_all from assms(1) have "C \ wffs\<^bsub>o\<^esub>" and "D \ wffs\<^bsub>o\<^esub>" using replacement_preserves_typing by fastforce+ have "\\<^sup>\\<^sub>\ t\<^bsub>?\\<^esub> (map FVar vs) \ wffs\<^bsub>\\<^esub>" using generalized_app_wff[where As = "map FVar vs" and ts = "map var_type vs"] by (metis eq_snd_iff length_map nth_map wffs_of_type_intros(1)) from assms(1) have "p \ positions C" using is_subform_implies_in_positions by fastforce then obtain G where "C\p \ (\\<^sup>\\<^sub>\ t\<^bsub>?\\<^esub> (map FVar vs))\ \ G" using is_replacement_at_existence by blast with \A \\<^bsub>p\<^esub> C\ and \\\<^sup>\\<^sub>\ t\<^bsub>?\\<^esub> (map FVar vs) \ wffs\<^bsub>\\<^esub>\ have "G \ wffs\<^bsub>o\<^esub>" using \A \ wffs\<^bsub>\\<^esub>\ and \C \ wffs\<^bsub>o\<^esub>\ and replacement_preserves_typing by blast let ?\ = "{(\, ?\\o) \ \t\<^bsub>?\\<^esub>. G, (\, ?\) \ \\<^sup>\\<^sub>\ vs A, (\, ?\) \ \\<^sup>\\<^sub>\ vs B}" and ?A = "(\\<^bsub>?\\<^esub> =\<^bsub>?\\<^esub> \\<^bsub>?\\<^esub>) \\<^sup>\ (\\<^bsub>?\\o\<^esub> \ \\<^bsub>?\\<^esub> \\<^sup>\ \\<^bsub>?\\o\<^esub> \ \\<^bsub>?\\<^esub>)" have "\ ?A" by (fact axiom_is_derivable_from_no_hyps[OF axiom_2]) moreover have "\t\<^bsub>?\\<^esub>. G \ wffs\<^bsub>?\\o\<^esub>" and "\\<^sup>\\<^sub>\ vs A \ wffs\<^bsub>?\\<^esub>" and "\\<^sup>\\<^sub>\ vs B \ wffs\<^bsub>?\\<^esub>" by (blast intro: \G \ wffs\<^bsub>o\<^esub>\ \A \ wffs\<^bsub>\\<^esub>\ \B \ wffs\<^bsub>\\<^esub>\)+ then have "is_substitution ?\" by simp moreover have " \v \ fmdom' ?\. var_name v \ free_var_names ({}::form set) \ is_free_for (?\ $$! v) v ?A" by ( ( code_simp, unfold atomize_conj[symmetric], simp, use is_free_for_in_equality is_free_for_in_equivalence is_free_for_in_imp is_free_for_in_var is_free_for_to_app in presburger )+, blast ) moreover have "?\ \ {$$}" by simp ultimately have "\ \<^bold>S ?\ ?A" by (rule Sub) moreover have " \<^bold>S ?\ ?A = (\\<^sup>\\<^sub>\ vs A =\<^bsub>?\\<^esub> \\<^sup>\\<^sub>\ vs B) \\<^sup>\ ((\t\<^bsub>?\\<^esub>. G) \ (\\<^sup>\\<^sub>\ vs A) \\<^sup>\ (\t\<^bsub>?\\<^esub>. G) \ (\\<^sup>\\<^sub>\ vs B))" by simp ultimately have "\
1": " \ (\\<^sup>\\<^sub>\ vs A =\<^bsub>?\\<^esub> \\<^sup>\\<^sub>\ vs B) \\<^sup>\ ((\t\<^bsub>?\\<^esub>. G) \ (\\<^sup>\\<^sub>\ vs A) \\<^sup>\ (\t\<^bsub>?\\<^esub>. G) \ (\\<^sup>\\<^sub>\ vs B))" by (simp only:) then have "\
2": "\ (\\<^sup>\\<^sub>\ vs (A =\<^bsub>\\<^esub> B)) \\<^sup>\ ((\t\<^bsub>?\\<^esub>. G) \ (\\<^sup>\\<^sub>\ vs A) \\<^sup>\ (\t\<^bsub>?\\<^esub>. G) \ (\\<^sup>\\<^sub>\ vs B))" proof (cases "vs = []") case True with "\
1" show ?thesis by simp next case False from "\
1" and prop_5238[OF False \A \ wffs\<^bsub>\\<^esub>\ \B \ wffs\<^bsub>\\<^esub>\] show ?thesis unfolding equivalence_def by (rule rule_R[where p = "[\,\]"]) force+ qed moreover have "\ (\t\<^bsub>?\\<^esub>. G) \ (\\<^sup>\\<^sub>\ vs A) =\<^bsub>o\<^esub> C" and "\ (\t\<^bsub>?\\<^esub>. G) \ (\\<^sup>\\<^sub>\ vs B) =\<^bsub>o\<^esub> D" proof - from assms(1) have "B \\<^bsub>p\<^esub> D" using is_replacement_at_minimal_change(1) by force from assms(1) have "D\p \ (\\<^sup>\\<^sub>\ t\<^bsub>?\\<^esub> (map FVar vs))\ \ G" using \C\p \ (\\<^sup>\\<^sub>\ t\<^bsub>?\\<^esub> (map FVar vs))\ \ G\ and replacement_override by (meson is_rule_R_app_def) from \B \\<^bsub>p\<^esub> D\ have "p \ positions D" using is_subform_implies_in_positions by auto from assms(1) have "binders_at D p = binders_at C p" using is_replacement_at_binders by fastforce then have "binders_at D p \ free_vars B = binders_at C p \ free_vars B" by simp with assms(2) [ folded capture_exposed_vars_at_def, unfolded capture_exposed_vars_at_alt_def[OF \p \ positions C\] ] have "lset vs \ capture_exposed_vars_at p D B" unfolding capture_exposed_vars_at_alt_def[OF \p \ positions D\] by auto have "is_free_for (\\<^sup>\\<^sub>\ vs A) (t, ?\) G" and "is_free_for (\\<^sup>\\<^sub>\ vs B) (t, ?\) G" proof - have "(t, ?\) \ lset vs \ vars {A, C}" and "(t, ?\) \ lset vs \ vars {B, D}" using \(t, ?\) \ lset vs \ vars {A, B, C, D}\ by simp_all moreover from assms(2) have " lset vs \ capture_exposed_vars_at p C A" and "lset vs \ capture_exposed_vars_at p D B" by fastforce fact ultimately show "is_free_for (\\<^sup>\\<^sub>\ vs A) (t, ?\) G" and "is_free_for (\\<^sup>\\<^sub>\ vs B) (t, ?\) G" using prop_5239_aux_4 and \B \\<^bsub>p\<^esub> D\ and \A \\<^bsub>p\<^esub> C\ and \C\p \ (\\<^sup>\\<^sub>\ t\<^bsub>?\\<^esub> (map FVar vs))\ \ G\ and \D\p \ (\\<^sup>\\<^sub>\ t\<^bsub>?\\<^esub> (map FVar vs))\ \ G\ by meson+ qed then have "\ (\t\<^bsub>?\\<^esub>. G) \ (\\<^sup>\\<^sub>\ vs A) =\<^bsub>o\<^esub> \<^bold>S {(t, ?\) \ \\<^sup>\\<^sub>\ vs A} G" and "\ (\t\<^bsub>?\\<^esub>. G) \ (\\<^sup>\\<^sub>\ vs B) =\<^bsub>o\<^esub> \<^bold>S {(t, ?\) \ \\<^sup>\\<^sub>\ vs B} G" using prop_5207[OF \\\<^sup>\\<^sub>\ vs A \ wffs\<^bsub>?\\<^esub>\ \G \ wffs\<^bsub>o\<^esub>\] and prop_5207[OF \\\<^sup>\\<^sub>\ vs B \ wffs\<^bsub>?\\<^esub>\ \G \ wffs\<^bsub>o\<^esub>\] by auto moreover obtain G'\<^sub>1 and G'\<^sub>2 where "C\p \ (\\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs A) (map FVar vs))\ \ G'\<^sub>1" and "D\p \ (\\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs B) (map FVar vs))\ \ G'\<^sub>2" using is_replacement_at_existence[OF \p \ positions C\] and is_replacement_at_existence[OF \p \ positions D\] by metis then have "\<^bold>S {(t, ?\) \ \\<^sup>\\<^sub>\ vs A} G = G'\<^sub>1" and "\<^bold>S {(t, ?\) \ \\<^sup>\\<^sub>\ vs B} G = G'\<^sub>2" proof - have "(t, ?\) \ lset vs \ vars C" and "(t, ?\) \ lset vs \ vars D" using \(t, ?\) \ lset vs \ vars {A, B, C, D}\ by simp_all then show "\<^bold>S {(t, ?\) \ \\<^sup>\\<^sub>\ vs A} G = G'\<^sub>1" and "\<^bold>S {(t, ?\) \ \\<^sup>\\<^sub>\ vs B} G = G'\<^sub>2" using \C\p \ (\\<^sup>\\<^sub>\ t\<^bsub>?\\<^esub> (map FVar vs))\ \ G\ and \D\p \ (\\<^sup>\\<^sub>\ t\<^bsub>?\\<^esub> map FVar vs)\ \ G\ and \C\p \ (\\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs A) (map FVar vs))\ \ G'\<^sub>1\ and \D\p \ (\\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs B) (map FVar vs))\ \ G'\<^sub>2\ and prop_5239_aux_2 by blast+ qed ultimately have "\ (\t\<^bsub>?\\<^esub>. G) \ (\\<^sup>\\<^sub>\ vs A) =\<^bsub>o\<^esub> G'\<^sub>1" and "\ (\t\<^bsub>?\\<^esub>. G) \ (\\<^sup>\\<^sub>\ vs B) =\<^bsub>o\<^esub> G'\<^sub>2" by (simp_all only:) moreover have "\ A =\<^bsub>\\<^esub> (\\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs A) (map FVar vs))" and "\ B =\<^bsub>\\<^esub> (\\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs B) (map FVar vs))" unfolding atomize_conj proof (cases "vs = []") assume "vs = []" show "\ A =\<^bsub>\\<^esub> \\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs A) (map FVar vs) \ \ B =\<^bsub>\\<^esub> \\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs B) (map FVar vs)" unfolding \vs = []\ using prop_5200 and \A \ wffs\<^bsub>\\<^esub>\ and \B \ wffs\<^bsub>\\<^esub>\ by simp next assume "vs \ []" show "\ A =\<^bsub>\\<^esub> \\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs A) (map FVar vs) \ \ B =\<^bsub>\\<^esub> \\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs B) (map FVar vs)" using Equality_Rules(2)[OF prop_5208[OF \vs \ []\]] and \A \ wffs\<^bsub>\\<^esub>\ and \B \ wffs\<^bsub>\\<^esub>\ by blast+ qed with \C\p \ (\\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs A) (map FVar vs))\ \ G'\<^sub>1\ and \D\p \ (\\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs B) (map FVar vs))\ \ G'\<^sub>2\ have "\ G'\<^sub>1 =\<^bsub>o\<^esub> C" and "\ G'\<^sub>2 =\<^bsub>o\<^esub> D" using Equality_Rules(2)[OF replacement_derivability] and \C \ wffs\<^bsub>o\<^esub>\ and \D \ wffs\<^bsub>o\<^esub>\ and \A \\<^bsub>p\<^esub> C\ and \B \\<^bsub>p\<^esub> D\ by blast+ ultimately show "\ (\t\<^bsub>?\\<^esub>. G) \ (\\<^sup>\\<^sub>\ vs A) =\<^bsub>o\<^esub> C" and "\ (\t\<^bsub>?\\<^esub>. G) \ (\\<^sup>\\<^sub>\ vs B) =\<^bsub>o\<^esub> D" using Equality_Rules(3) by blast+ qed ultimately show ?thesis proof - from "\
2" and \\ (\t\<^bsub>?\\<^esub>. G) \ (\\<^sup>\\<^sub>\ vs A) =\<^bsub>o\<^esub> C\ have " \ (\\<^sup>\\<^sub>\ vs (A =\<^bsub>\\<^esub> B)) \\<^sup>\ (C \\<^sup>\ (\t\<^bsub>?\\<^esub>. G) \ (\\<^sup>\\<^sub>\ vs B))" by (rule rule_R[where p = "[\,\,\]"]) force+ from this and \\ (\t\<^bsub>?\\<^esub>. G) \ (\\<^sup>\\<^sub>\ vs B) =\<^bsub>o\<^esub> D\ show ?thesis by (rule rule_R[where p = "[\,\]"]) force+ qed qed end subsection \Theorem 5240 (Deduction Theorem)\ lemma pseudo_rule_R_is_tautologous: assumes "C \ wffs\<^bsub>o\<^esub>" and "D \ wffs\<^bsub>o\<^esub>" and "E \ wffs\<^bsub>o\<^esub>" and "H \ wffs\<^bsub>o\<^esub>" shows "is_tautologous (((H \\<^sup>\ C) \\<^sup>\ ((H \\<^sup>\ E) \\<^sup>\ ((E \\<^sup>\ (C \\<^sup>\ D)) \\<^sup>\ (H \\<^sup>\ D)))))" proof - let ?\ = "{(\, o) \ C, (\, o) \ D, (\, o) \ E, (\, o) \ H}" have " is_tautology (((\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>) \\<^sup>\ ((\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>) \\<^sup>\ ((\\<^bsub>o\<^esub> \\<^sup>\ (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)) \\<^sup>\ (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)))))" using \\<^sub>B_simps by simp moreover have "is_substitution ?\" using assms by auto moreover have "\(x, \) \ fmdom' ?\. \ = o" by simp moreover have " ((H \\<^sup>\ C) \\<^sup>\ ((H \\<^sup>\ E) \\<^sup>\ ((E \\<^sup>\ (C \\<^sup>\ D)) \\<^sup>\ (H \\<^sup>\ D)))) = \<^bold>S ?\ (((\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>) \\<^sup>\ ((\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>) \\<^sup>\ ((\\<^bsub>o\<^esub> \\<^sup>\ (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)) \\<^sup>\ (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)))))" by simp ultimately show ?thesis by blast qed syntax "_HypDer" :: "form \ form set \ form \ bool" ("_,_ \ _" [50, 50, 50] 50) translations "\, H \ P" \ "\ \ {H} \ P" theorem thm_5240: assumes "finite \" and "\, H \ P" shows "\ \ H \\<^sup>\ P" proof - from \\, H \ P\ obtain \\<^sub>1 and \\<^sub>2 where *: "is_hyp_proof_of (\ \ {H}) \\<^sub>1 \\<^sub>2 P" using hyp_derivability_implies_hyp_proof_existence by blast have "\ \ H \\<^sup>\ (\\<^sub>2 ! i')" if "i' \<^sub>2" for i' using that proof (induction i' rule: less_induct) case (less i') let ?R = "\\<^sub>2 ! i'" from less.prems(1) and * have "is_hyps \" by fastforce from less.prems and * have "?R \ wffs\<^bsub>o\<^esub>" using elem_of_proof_is_wffo[simplified] by auto from less.prems and * have "is_hyp_proof_step (\ \ {H}) \\<^sub>1 \\<^sub>2 i'" by simp then consider (hyp) "?R \ \ \ {H}" | (seq) "?R \ lset \\<^sub>1" | (rule_R') "\j k p. {j, k} \ {0.. is_rule_R'_app (\ \ {H}) p ?R (\\<^sub>2 ! j) (\\<^sub>2 ! k)" by force then show ?case proof cases case hyp then show ?thesis proof (cases "?R = H") case True with \?R \ wffs\<^bsub>o\<^esub>\ have "is_tautologous (H \\<^sup>\ ?R)" using implication_reflexivity_is_tautologous by (simp only:) with \is_hyps \\ show ?thesis by (rule rule_P(2)) next case False with hyp have "?R \ \" by blast with \is_hyps \\ have "\ \ ?R" by (intro dv_hyp) moreover from less.prems(1) and * have "is_tautologous (?R \\<^sup>\ (H \\<^sup>\ ?R))" using principle_of_simplification_is_tautologous[OF \?R \ wffs\<^bsub>o\<^esub>\] by force moreover from \?R \ wffs\<^bsub>o\<^esub>\ have "is_hyps {?R}" by simp ultimately show ?thesis using rule_P(1)[where \ = "{?R}" and hs = "[?R]", OF \is_hyps \\] by simp qed next case seq then have "\\<^sub>1 \ []" by force moreover from less.prems(1) and * have "is_proof \\<^sub>1" by fastforce moreover from seq obtain i'' where "i'' < length \\<^sub>1" and "?R = \\<^sub>1 ! i''" by (metis in_set_conv_nth) ultimately have "is_theorem ?R" using proof_form_is_theorem by fastforce with \is_hyps \\ have "\ \ ?R" by (intro dv_thm) moreover from \?R \ wffs\<^bsub>o\<^esub>\ and less.prems(1) and * have "is_tautologous (?R \\<^sup>\ (H \\<^sup>\ ?R))" using principle_of_simplification_is_tautologous by force moreover from \?R \ wffs\<^bsub>o\<^esub>\ have "is_hyps {?R}" by simp ultimately show ?thesis using rule_P(1)[where \ = "{?R}" and hs = "[?R]", OF \is_hyps \\] by simp next case rule_R' then obtain j and k and p where "{j, k} \ {0.. \ {H}) p ?R (\\<^sub>2 ! j) (\\<^sub>2 ! k)" by auto then obtain A and B and C and \ where "C = \\<^sub>2 ! j" and "\\<^sub>2 ! k = A =\<^bsub>\\<^esub> B" by fastforce with \{j, k} \ {0.. have "\ \ H \\<^sup>\ C" and "\ \ H \\<^sup>\ (A =\<^bsub>\\<^esub> B)" using less.IH and less.prems(1) by (simp, force) define S where "S \ {(x, \) | x \ p' E. strict_prefix p' p \ \x\<^bsub>\\<^esub>. E \\<^bsub>p'\<^esub> C \ (x, \) \ free_vars (A =\<^bsub>\\<^esub> B)}" with \C = \\<^sub>2 ! j\ and \\\<^sub>2 ! k = A =\<^bsub>\\<^esub> B\ have "\v \ S. v \ free_vars (\ \ {H})" using rule_R'_app by fastforce moreover have "S \ free_vars (A =\<^bsub>\\<^esub> B)" unfolding S_def by blast then have "finite S" by (fact rev_finite_subset[OF free_vars_form_finiteness]) then obtain vs where "lset vs = S" using finite_list by blast ultimately have "\ \ H \\<^sup>\ \\<^sup>\\<^sub>\ vs (A =\<^bsub>\\<^esub> B)" using generalized_prop_5237[OF \is_hyps \\ \\ \ H \\<^sup>\ (A =\<^bsub>\\<^esub> B)\] by simp moreover have rule_R_app: "is_rule_R_app p ?R (\\<^sub>2 ! j) (\\<^sub>2 ! k)" using rule_R'_app by fastforce with S_def and \lset vs = S\ have "\ \\<^sup>\\<^sub>\ vs (A =\<^bsub>\\<^esub> B) \\<^sup>\ (C \\<^sup>\ ?R)" unfolding \C = \\<^sub>2 ! j\ and \\\<^sub>2 ! k = A =\<^bsub>\\<^esub> B\ using prop_5239 by (simp only:) with \is_hyps \\ have "\ \ \\<^sup>\\<^sub>\ vs (A =\<^bsub>\\<^esub> B) \\<^sup>\ (C \\<^sup>\ ?R)" by (elim derivability_implies_hyp_derivability) ultimately show ?thesis proof - let ?A\<^sub>1 = "H \\<^sup>\ C" and ?A\<^sub>2 = "H \\<^sup>\ \\<^sup>\\<^sub>\ vs (A =\<^bsub>\\<^esub> B)" and ?A\<^sub>3 = "\\<^sup>\\<^sub>\ vs (A =\<^bsub>\\<^esub> B) \\<^sup>\ (C \\<^sup>\ ?R)" let ?hs = "[?A\<^sub>1, ?A\<^sub>2, ?A\<^sub>3]" let ?\ = "lset ?hs" from \\ \ ?A\<^sub>1\ have "H \ wffs\<^bsub>o\<^esub>" using hyp_derivable_form_is_wffso by (blast dest: wffs_from_imp_op(1)) moreover from \\ \ ?A\<^sub>2\ have "\\<^sup>\\<^sub>\ vs (A =\<^bsub>\\<^esub> B) \ wffs\<^bsub>o\<^esub>" using hyp_derivable_form_is_wffso by (blast dest: wffs_from_imp_op(2)) moreover from \C = \\<^sub>2 ! j\ and rule_R_app have "C \ wffs\<^bsub>o\<^esub>" using replacement_preserves_typing by fastforce ultimately have *: "is_tautologous (?A\<^sub>1 \\<^sup>\ (?A\<^sub>2 \\<^sup>\ (?A\<^sub>3 \\<^sup>\ (H \\<^sup>\ ?R))))" using \?R \ wffs\<^bsub>o\<^esub>\ by (intro pseudo_rule_R_is_tautologous) moreover from \\ \ ?A\<^sub>1\ and \\ \ ?A\<^sub>2\ and \\ \ ?A\<^sub>3\ have "is_hyps ?\" using hyp_derivable_form_is_wffso by simp moreover from \\ \ ?A\<^sub>1\ and \\ \ ?A\<^sub>2\ and \\ \ ?A\<^sub>3\ have "\A \ ?\. \ \ A" by force ultimately show ?thesis using rule_P(1)[where \ = ?\ and hs = ?hs and B = "H \\<^sup>\ ?R", OF \is_hyps \\] by simp qed qed qed moreover from \is_hyp_proof_of (\ \ {H}) \\<^sub>1 \\<^sub>2 P\ have "\\<^sub>2 ! (length \\<^sub>2 - 1) = P" using last_conv_nth by fastforce ultimately show ?thesis using \is_hyp_proof_of (\ \ {H}) \\<^sub>1 \\<^sub>2 P\ by force qed lemmas Deduction_Theorem = thm_5240 (* synonym *) text \ We prove a generalization of the Deduction Theorem, namely that if \\ \ {H\<^sub>1, \ ,H\<^sub>n} \ P\ then \\ \ H\<^sub>1 \\<^sup>\ (\ \\<^sup>\ (H\<^sub>n \\<^sup>\ P) \)\: \ corollary generalized_deduction_theorem: assumes "finite \" and "finite \'" and "\ \ \' \ P" and "lset hs = \'" shows "\ \ hs \\<^sup>\\<^sub>\ P" using assms proof (induction hs arbitrary: \' P rule: rev_induct) case Nil then show ?case by simp next case (snoc H hs) from \lset (hs @ [H]) = \'\ have "H \ \'" by fastforce from \lset (hs @ [H]) = \'\ obtain \'' where "\'' \ {H} = \'" and "\'' = lset hs" by simp from \\'' \ {H} = \'\ and \\ \ \' \ P\ have "\ \ \'' \ {H} \ P" by fastforce with \finite \\ and \finite \'\ and \\'' = lset hs\ have "\ \ \'' \ H \\<^sup>\ P" using Deduction_Theorem by simp with \\'' = lset hs\ and \finite \\ have "\ \ foldr (\\<^sup>\) hs (H \\<^sup>\ P)" using snoc.IH by fastforce moreover have "(hs @ [H]) \\<^sup>\\<^sub>\ P = hs \\<^sup>\\<^sub>\ (H \\<^sup>\ P)" by simp ultimately show ?case by auto qed subsection \Proposition 5241\ proposition prop_5241: assumes "is_hyps \" and "\ \ A" and "\ \ \" shows "\ \ A" proof (cases "\ = {}") case True show ?thesis by (fact derivability_implies_hyp_derivability[OF assms(2)[unfolded True] assms(1)]) next case False then obtain hs where "lset hs = \" and "hs \ []" using hyp_derivability_implies_hyp_proof_existence[OF assms(2)] unfolding is_hyp_proof_of_def by (metis empty_set finite_list) with assms(2) have "\ hs \\<^sup>\\<^sub>\ A" using generalized_deduction_theorem by force moreover from \lset hs = \\ and assms(1,3) have "\ \ H" if "H \ lset hs" for H using that by (blast intro: dv_hyp) ultimately show ?thesis using assms(1) and generalized_modus_ponens and derivability_implies_hyp_derivability by meson qed subsection \Proposition 5242 (Rule of Existential Generalization)\ proposition prop_5242: assumes "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" and "\ \ \<^bold>S {(x, \) \ A} B" and "is_free_for A (x, \) B" shows "\ \ \x\<^bsub>\\<^esub>. B" proof - from assms(3) have "is_hyps \" by (blast dest: is_derivable_from_hyps.cases) then have "\ \ \x\<^bsub>\\<^esub>. \\<^sup>\ B \\<^sup>\ \\<^sup>\ \<^bold>S {(x, \) \ A} B" (is \\ \ ?C \\<^sup>\ \\<^sup>\ ?D\) using prop_5226[OF assms(1) neg_wff[OF assms(2)] is_free_for_in_neg[OF assms(4)]] unfolding derived_substitution_simps(4) using derivability_implies_hyp_derivability by (simp only:) moreover have *: "is_tautologous ((?C \\<^sup>\ \\<^sup>\ ?D) \\<^sup>\ (?D \\<^sup>\ \\<^sup>\ ?C))" proof - have "?C \ wffs\<^bsub>o\<^esub>" and "?D \ wffs\<^bsub>o\<^esub>" using assms(2) and hyp_derivable_form_is_wffso[OF assms(3)] by auto then show ?thesis by (fact pseudo_modus_tollens_is_tautologous) qed moreover from assms(3) and \\ \ ?C \\<^sup>\ \\<^sup>\ ?D\ have "is_hyps {?C \\<^sup>\ \\<^sup>\ ?D, ?D}" using hyp_derivable_form_is_wffso by force ultimately show ?thesis unfolding exists_def using assms(3) and rule_P(1) [ where \ = "{?C \\<^sup>\ \\<^sup>\ ?D, ?D}" and hs = "[?C \\<^sup>\ \\<^sup>\ ?D, ?D]" and B = "\\<^sup>\ ?C", OF \is_hyps \\ ] by simp qed lemmas "\Gen" = prop_5242 (* synonym *) subsection \Proposition 5243 (Comprehension Theorem)\ context begin private lemma prop_5243_aux: assumes "\\<^sup>\\<^sub>\ B (map FVar vs) \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" and "k < length vs" shows "\ \ var_type (vs ! k)" proof - from assms(1) obtain ts where "length ts = length (map FVar vs)" and *: "\k < length (map FVar vs). (map FVar vs) ! k \ wffs\<^bsub>ts ! k\<^esub>" and "B \ wffs\<^bsub>foldr (\) ts \\<^esub>" using wffs_from_generalized_app by force have "\ = foldr (\) ts \" by (fact wff_has_unique_type[OF assms(2) \B \ wffs\<^bsub>foldr (\) ts \\<^esub>\]) have "ts = map var_type vs" proof - have "length ts = length (map var_type vs)" by (simp add: \length ts = length (map FVar vs)\) moreover have "\k < length ts. ts ! k = (map var_type vs) ! k" proof (intro allI impI) fix k assume "k < length ts" with * have "(map FVar vs) ! k \ wffs\<^bsub>ts ! k\<^esub>" by (simp add: \length ts = length (map FVar vs)\) with \k < length ts\ and \length ts = length (map var_type vs)\ show "ts ! k = (map var_type vs) ! k" using surj_pair[of "vs ! k"] and wff_has_unique_type and wffs_of_type_intros(1) by force qed ultimately show ?thesis using list_eq_iff_nth_eq by blast qed with \\ = foldr (\) ts \\ and assms(3) show ?thesis using fun_type_atoms_neq_fun_type by (metis length_map nth_map) qed proposition prop_5243: assumes "B \ wffs\<^bsub>\\<^esub>" and "\ = foldr (\) (map var_type vs) \" and "(u, \) \ free_vars B" shows "\ \u\<^bsub>\\<^esub>. \\<^sup>\\<^sub>\ vs ((\\<^sup>\\<^sub>\ u\<^bsub>\\<^esub> (map FVar vs)) =\<^bsub>\\<^esub> B)" proof (cases "vs = []") case True with assms(2) have "\ = \" by simp from assms(1) have "u\<^bsub>\\<^esub> =\<^bsub>\\<^esub> B \ wffs\<^bsub>o\<^esub>" by blast moreover have "\ B =\<^bsub>\\<^esub> B" by (fact prop_5200[OF assms(1)]) then have "\ \<^bold>S {(u, \) \ B} (u\<^bsub>\\<^esub> =\<^bsub>\\<^esub> B)" using free_var_singleton_substitution_neutrality[OF assms(3)] unfolding \\ = \\ by simp moreover from assms(3)[unfolded \\ = \\] have "is_free_for B (u, \) (u\<^bsub>\\<^esub> =\<^bsub>\\<^esub> B)" by (intro is_free_for_in_equality) (use is_free_at_in_free_vars in auto) ultimately have "\ \u\<^bsub>\\<^esub>. (u\<^bsub>\\<^esub> =\<^bsub>\\<^esub> B)" by (rule "\Gen"[OF assms(1)]) with \\ = \\ and True show ?thesis by simp next case False let ?\ = "{(u, \) \ \\<^sup>\\<^sub>\ vs B}" from assms(2) have *: "(u, \) \ v" if "v \ lset vs" for v using that and fun_type_atoms_neq_fun_type by (metis in_set_conv_nth length_map nth_map snd_conv) from False and assms(1) have "\ \\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs B) (map FVar vs) =\<^bsub>\\<^esub> B" by (fact prop_5208) then have "\ \\<^sup>\\<^sub>\ vs (\\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs B) (map FVar vs) =\<^bsub>\\<^esub> B)" using generalized_Gen by simp moreover have "\<^bold>S ?\ (\\<^sup>\\<^sub>\ vs ((\\<^sup>\\<^sub>\ u\<^bsub>\\<^esub> (map FVar vs)) =\<^bsub>\\<^esub> B)) = \\<^sup>\\<^sub>\ vs (\\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs B) (map FVar vs) =\<^bsub>\\<^esub> B)" proof - from * have **: "map (\A. \<^bold>S {(u, \) \ B} A) (map FVar vs) = map FVar vs" for B by (induction vs) fastforce+ from * have " \<^bold>S ?\ (\\<^sup>\\<^sub>\ vs ((\\<^sup>\\<^sub>\ u\<^bsub>\\<^esub> (map FVar vs)) =\<^bsub>\\<^esub> B)) = \\<^sup>\\<^sub>\ vs (\<^bold>S ?\ ((\\<^sup>\\<^sub>\ u\<^bsub>\\<^esub> (map FVar vs)) =\<^bsub>\\<^esub> B))" using generalized_forall_substitution by force also have "\ = \\<^sup>\\<^sub>\ vs ((\<^bold>S ?\ (\\<^sup>\\<^sub>\ u\<^bsub>\\<^esub> (map FVar vs))) =\<^bsub>\\<^esub> \<^bold>S {(u, \) \ \\<^sup>\\<^sub>\ vs B} B)" by simp also from assms(3) have "\ = \\<^sup>\\<^sub>\ vs ((\<^bold>S ?\ (\\<^sup>\\<^sub>\ u\<^bsub>\\<^esub> (map FVar vs))) =\<^bsub>\\<^esub> B)" using free_var_singleton_substitution_neutrality by simp also have "\ = \\<^sup>\\<^sub>\ vs (\\<^sup>\\<^sub>\ \<^bold>S ?\ (u\<^bsub>\\<^esub>) (map (\A. \<^bold>S ?\ A) (map FVar vs)) =\<^bsub>\\<^esub> B)" using generalized_app_substitution by simp also have "\ = \\<^sup>\\<^sub>\ vs (\\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs B) (map (\A. \<^bold>S ?\ A) (map FVar vs)) =\<^bsub>\\<^esub> B)" by simp also from ** have "\ = \\<^sup>\\<^sub>\ vs (\\<^sup>\\<^sub>\ (\\<^sup>\\<^sub>\ vs B) (map FVar vs) =\<^bsub>\\<^esub> B)" by presburger finally show ?thesis . qed ultimately have "\ \<^bold>S ?\ (\\<^sup>\\<^sub>\ vs (\\<^sup>\\<^sub>\ u\<^bsub>\\<^esub> (map FVar vs) =\<^bsub>\\<^esub> B))" by simp moreover from assms(3) have "is_free_for (\\<^sup>\\<^sub>\ vs B) (u, \) (\\<^sup>\\<^sub>\ vs (\\<^sup>\\<^sub>\ u\<^bsub>\\<^esub> (map FVar vs) =\<^bsub>\\<^esub> B))" by (intro is_free_for_in_generalized_forall is_free_for_in_equality is_free_for_in_generalized_app) (use free_vars_of_generalized_abs is_free_at_in_free_vars in \fastforce+\) moreover have "\\<^sup>\\<^sub>\ vs B \ wffs\<^bsub>\\<^esub>" and "\\<^sup>\\<^sub>\ vs (\\<^sup>\\<^sub>\ u\<^bsub>\\<^esub> (map FVar vs) =\<^bsub>\\<^esub> B) \ wffs\<^bsub>o\<^esub>" proof - have "FVar (vs ! k) \ wffs\<^bsub>var_type (vs ! k)\<^esub>" if "k < length vs" for k using that and surj_pair[of "vs ! k"] by fastforce with assms(2) have "\\<^sup>\\<^sub>\ u\<^bsub>\\<^esub> (map FVar vs) \ wffs\<^bsub>\\<^esub>" using generalized_app_wff[where ts = "map var_type vs"] by force with assms(1) show "\\<^sup>\\<^sub>\ vs (\\<^sup>\\<^sub>\ u\<^bsub>\\<^esub> (map FVar vs) =\<^bsub>\\<^esub> B) \ wffs\<^bsub>o\<^esub>" by (auto simp only:) qed (use assms(1,2) in blast) ultimately show ?thesis using "\Gen" by (simp only:) qed end subsection \Proposition 5244 (Existential Rule)\ text \ - The proof in @{cite "andrews:2002"} uses the pseudo-rule Q and 2123 of \\\. Therefore, we instead - base our proof on the proof of Theorem 170 in @{cite "andrews:1965"}: + The proof in \<^cite>\"andrews:2002"\ uses the pseudo-rule Q and 2123 of \\\. Therefore, we instead + base our proof on the proof of Theorem 170 in \<^cite>\"andrews:1965"\: \ lemma prop_5244_aux: assumes "A \ wffs\<^bsub>o\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" and "(x, \) \ free_vars A" shows "\ \x\<^bsub>\\<^esub>. (B \\<^sup>\ A) \\<^sup>\ (\x\<^bsub>\\<^esub>. B \\<^sup>\ A)" proof - have "B \\<^sup>\ A \ wffs\<^bsub>o\<^esub>" using assms by blast moreover have "is_free_for (x\<^bsub>\\<^esub>) (x, \) (B \\<^sup>\ A)" by simp ultimately have "\ \x\<^bsub>\\<^esub>. (B \\<^sup>\ A) \\<^sup>\ (B \\<^sup>\ A)" using prop_5226[where A = "x\<^bsub>\\<^esub>" and B = "B \\<^sup>\ A", OF wffs_of_type_intros(1)] and identity_singleton_substitution_neutrality by metis moreover have "is_hyps {\x\<^bsub>\\<^esub>. (B \\<^sup>\ A)}" using \B \\<^sup>\ A \ wffs\<^bsub>o\<^esub>\ by blast ultimately have "\
1": "{\x\<^bsub>\\<^esub>. (B \\<^sup>\ A)} \ \x\<^bsub>\\<^esub>. (B \\<^sup>\ A) \\<^sup>\ (B \\<^sup>\ A)" by (fact derivability_implies_hyp_derivability) have "\
2": "{\x\<^bsub>\\<^esub>. (B \\<^sup>\ A)} \ \x\<^bsub>\\<^esub>. (B \\<^sup>\ A)" using \B \\<^sup>\ A \ wffs\<^bsub>o\<^esub>\ by (blast intro: dv_hyp) have "\
3": "{\x\<^bsub>\\<^esub>. (B \\<^sup>\ A)} \ \\<^sup>\ A \\<^sup>\ \\<^sup>\ B" proof (intro rule_P(1) [where \ = "{\x\<^bsub>\\<^esub>. (B \\<^sup>\ A)}" and \ = "{\x\<^bsub>\\<^esub>. (B \\<^sup>\ A) \\<^sup>\ (B \\<^sup>\ A), \x\<^bsub>\\<^esub>. (B \\<^sup>\ A)}"]) have "is_tautologous ([C \\<^sup>\ (B \\<^sup>\ A), C] \\<^sup>\\<^sub>\ (\\<^sup>\ A \\<^sup>\ \\<^sup>\ B))" if "C \ wffs\<^bsub>o\<^esub>" for C proof - let ?\ = "{(\, o) \ A, (\, o) \ B, (\, o) \ C}" have "is_tautology ((\\<^bsub>o\<^esub> \\<^sup>\ (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)) \\<^sup>\ (\\<^bsub>o\<^esub> \\<^sup>\ (\\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^sup>\ \\<^bsub>o\<^esub>)))" (is "is_tautology ?A") using \\<^sub>B_simps by (auto simp add: inj_eq) moreover have "is_pwff_substitution ?\" using assms(1,2) and that by auto moreover have "[C \\<^sup>\ (B \\<^sup>\ A), C] \\<^sup>\\<^sub>\ (\\<^sup>\ A \\<^sup>\ \\<^sup>\ B) = \<^bold>S ?\ ?A" by simp ultimately show ?thesis by blast qed then show "is_tautologous ([\x\<^bsub>\\<^esub>. (B \\<^sup>\ A) \\<^sup>\ (B \\<^sup>\ A), \x\<^bsub>\\<^esub>. (B \\<^sup>\ A)] \\<^sup>\\<^sub>\ (\\<^sup>\ A \\<^sup>\ \\<^sup>\ B))" using \B \\<^sup>\ A \ wffs\<^bsub>o\<^esub>\ and forall_wff by simp qed (use "\
1" "\
2" \is_hyps {\x\<^bsub>\\<^esub>. (B \\<^sup>\ A)}\ hyp_derivable_form_is_wffso[OF "\
1"] in force)+ have "\
4": "{\x\<^bsub>\\<^esub>. (B \\<^sup>\ A)} \ \\<^sup>\ A \\<^sup>\ \x\<^bsub>\\<^esub>. \\<^sup>\ B" using prop_5237[OF \is_hyps {\x\<^bsub>\\<^esub>. (B \\<^sup>\ A)}\ "\
3"] and assms(3) by auto have "\
5": "{\x\<^bsub>\\<^esub>. (B \\<^sup>\ A)} \ \x\<^bsub>\\<^esub>. B \\<^sup>\ A" unfolding exists_def proof (intro rule_P(1)[where \ = "{\x\<^bsub>\\<^esub>. (B \\<^sup>\ A)}" and \ = "{\\<^sup>\ A \\<^sup>\ \x\<^bsub>\\<^esub>. \\<^sup>\ B}"]) have "is_tautologous ([\\<^sup>\ A \\<^sup>\ C] \\<^sup>\\<^sub>\ (\\<^sup>\ C \\<^sup>\ A))" if "C \ wffs\<^bsub>o\<^esub>" for C proof - let ?\ = "{(\, o) \ A, (\, o) \ C}" have "is_tautology ((\\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>) \\<^sup>\ (\\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>))" (is "is_tautology ?A") using \\<^sub>B_simps by (auto simp add: inj_eq) moreover have "is_pwff_substitution ?\" using assms(1) and that by auto moreover have "[\\<^sup>\ A \\<^sup>\ C] \\<^sup>\\<^sub>\ (\\<^sup>\ C \\<^sup>\ A) = \<^bold>S ?\ ?A" by simp ultimately show ?thesis by blast qed then show "is_tautologous ([\\<^sup>\ A \\<^sup>\ \x\<^bsub>\\<^esub>. \\<^sup>\ B] \\<^sup>\\<^sub>\ (\\<^sup>\ \x\<^bsub>\\<^esub>. \\<^sup>\ B \\<^sup>\ A))" using forall_wff[OF neg_wff[OF assms(2)]] by (simp only:) qed (use "\
4" \is_hyps {\x\<^bsub>\\<^esub>. (B \\<^sup>\ A)}\ hyp_derivable_form_is_wffso[OF "\
4"] in force)+ then show ?thesis using Deduction_Theorem by simp qed proposition prop_5244: assumes "\, B \ A" and "(x, \) \ free_vars (\ \ {A})" shows "\, \x\<^bsub>\\<^esub>. B \ A" proof - from assms(1) have "is_hyps \" using hyp_derivability_implies_hyp_proof_existence by force then have "\ \ B \\<^sup>\ A" using assms(1) and Deduction_Theorem by simp then have "\ \ \x\<^bsub>\\<^esub>. (B \\<^sup>\ A)" using Gen and assms(2) by simp moreover have "A \ wffs\<^bsub>o\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" by ( fact hyp_derivable_form_is_wffso[OF assms(1)], fact hyp_derivable_form_is_wffso[OF \\ \ B \\<^sup>\ A\, THEN wffs_from_imp_op(1)] ) with assms(2) and \is_hyps \\ have "\ \ \x\<^bsub>\\<^esub>. (B \\<^sup>\ A) \\<^sup>\ (\x\<^bsub>\\<^esub>. B \\<^sup>\ A)" using prop_5244_aux[THEN derivability_implies_hyp_derivability] by simp ultimately have "\ \ \x\<^bsub>\\<^esub>. B \\<^sup>\ A" by (rule MP) then have "\, \x\<^bsub>\\<^esub>. B \ \x\<^bsub>\\<^esub>. B \\<^sup>\ A" using prop_5241 and exists_wff[OF \B \ wffs\<^bsub>o\<^esub>\] and \is_hyps \\ by (meson Un_subset_iff empty_subsetI finite.simps finite_Un inf_sup_ord(3) insert_subsetI) moreover from \is_hyps \\ and \B \ wffs\<^bsub>o\<^esub>\ have "is_hyps (\ \ {\x\<^bsub>\\<^esub>. B})" by auto then have "\, \x\<^bsub>\\<^esub>. B \ \x\<^bsub>\\<^esub>. B" using dv_hyp by simp ultimately show ?thesis using MP by blast qed lemmas "\_Rule" = prop_5244 (* synonym *) subsection \Proposition 5245 (Rule C)\ lemma prop_5245_aux: assumes "x \ y" and "(y, \) \ free_vars (\x\<^bsub>\\<^esub>. B)" and "is_free_for (y\<^bsub>\\<^esub>) (x, \) B" shows "is_free_for (x\<^bsub>\\<^esub>) (y, \) \<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} B" using assms(2,3) proof (induction B) case (FVar v) then show ?case using surj_pair[of v] by fastforce next case (FCon k) then show ?case using surj_pair[of k] by fastforce next case (FApp B\<^sub>1 B\<^sub>2) from FApp.prems(1) have "(y, \) \ free_vars (\x\<^bsub>\\<^esub>. B\<^sub>1)" and "(y, \) \ free_vars (\x\<^bsub>\\<^esub>. B\<^sub>2)" by force+ moreover from FApp.prems(2) have "is_free_for (y\<^bsub>\\<^esub>) (x, \) B\<^sub>1" and "is_free_for (y\<^bsub>\\<^esub>) (x, \) B\<^sub>2" using is_free_for_from_app by iprover+ ultimately have "is_free_for (x\<^bsub>\\<^esub>) (y, \) \<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} B\<^sub>1" and "is_free_for (x\<^bsub>\\<^esub>) (y, \) \<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} B\<^sub>2" using FApp.IH by simp_all then have "is_free_for (x\<^bsub>\\<^esub>) (y, \) ((\<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} B\<^sub>1) \ (\<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} B\<^sub>2))" by (intro is_free_for_to_app) then show ?case unfolding singleton_substitution_simps(3) . next case (FAbs v B') obtain z and \ where "v = (z, \)" by fastforce then show ?case proof (cases "v = (x, \)") case True with FAbs.prems(1) have "(y, \) \ free_vars (\x\<^bsub>\\<^esub>. B')" by simp moreover from assms(1) have "(y, \) \ (x, \)" by blast ultimately have "(y, \) \ free_vars B'" using FAbs.prems(1) by simp with \(y, \) \ (x, \)\ have "(y, \) \ free_vars (\x\<^bsub>\\<^esub>. B')" by simp then have "is_free_for (x\<^bsub>\\<^esub>) (y, \) (\x\<^bsub>\\<^esub>. B')" unfolding is_free_for_def using is_free_at_in_free_vars by blast then have "is_free_for (x\<^bsub>\\<^esub>) (y, \) \<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} (\x\<^bsub>\\<^esub>. B')" using singleton_substitution_simps(4) by presburger then show ?thesis unfolding True . next case False from assms(1) have "(y, \) \ (x, \)" by blast with FAbs.prems(1) have *: "(y, \) \ free_vars (\x\<^bsub>\\<^esub>. (\z\<^bsub>\\<^esub>. B'))" using \v = (z, \)\ by fastforce then show ?thesis proof (cases "(y, \) \ v") case True from True[unfolded \v = (z, \)\] and * have "(y, \) \ free_vars (\x\<^bsub>\\<^esub>. B')" by simp moreover from False[unfolded \v = (z, \)\] have "is_free_for (y\<^bsub>\\<^esub>) (x, \) B'" using is_free_for_from_abs[OF FAbs.prems(2)[unfolded \v = (z, \)\]] by blast ultimately have "is_free_for (x\<^bsub>\\<^esub>) (y, \) (\<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} B')" by (fact FAbs.IH) then have "is_free_for (x\<^bsub>\\<^esub>) (y, \) (\z\<^bsub>\\<^esub>. (\<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} B'))" using False[unfolded \v = (z, \)\] by (intro is_free_for_to_abs, fastforce+) then show ?thesis unfolding singleton_substitution_simps(4) and \v = (z, \)\ using \(z, \) \ (x, \)\ by auto next case False then have "v = (y, \)" by simp have "is_free_for (x\<^bsub>\\<^esub>) (y, \) (\y\<^bsub>\\<^esub>. \<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} B')" proof- have "(y, \) \ free_vars (\y\<^bsub>\\<^esub>. \<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} B')" by simp then show ?thesis using is_free_at_in_free_vars by blast qed with\v = (y, \)\ and \(y, \) \ (x, \)\ show ?thesis using singleton_substitution_simps(4) by presburger qed qed qed proposition prop_5245: assumes "\ \ \x\<^bsub>\\<^esub>. B" and "\, \<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} B \ A" and "is_free_for (y\<^bsub>\\<^esub>) (x, \) B" and "(y, \) \ free_vars (\ \ {\x\<^bsub>\\<^esub>. B, A})" shows "\ \ A" proof - from assms(1) have "is_hyps \" by (blast elim: is_derivable_from_hyps.cases) from assms(2,4) have "\, \y\<^bsub>\\<^esub>. \<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} B \ A" using "\_Rule" by simp then have *: "\ \ (\y\<^bsub>\\<^esub>. \<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} B) \\<^sup>\ A" (is \_ \ ?F\) using Deduction_Theorem and \is_hyps \\ by blast then have "\ \ \x\<^bsub>\\<^esub>. B \\<^sup>\ A" proof (cases "x = y") case True with * show ?thesis using identity_singleton_substitution_neutrality by force next case False from assms(4) have "(y, \) \ free_vars (\x\<^bsub>\\<^esub>. B)" using free_vars_in_all_vars by auto have "\\<^sup>\ \<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} B \ wffs\<^bsub>o\<^esub>" by ( fact hyp_derivable_form_is_wffso [OF *, THEN wffs_from_imp_op(1), THEN wffs_from_exists, THEN neg_wff] ) moreover from False have "(x, \) \ free_vars (\\<^sup>\ \<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} B)" using free_var_in_renaming_substitution by simp moreover have "is_free_for (x\<^bsub>\\<^esub>) (y, \) (\\<^sup>\ \<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} B)" by (intro is_free_for_in_neg prop_5245_aux[OF False \(y, \) \ free_vars (\x\<^bsub>\\<^esub>. B)\ assms(3)]) moreover from assms(3,4) have "\<^bold>S {(y, \) \ x\<^bsub>\\<^esub>} \<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} B = B" using identity_singleton_substitution_neutrality and renaming_substitution_composability by force ultimately have "\ (\y\<^bsub>\\<^esub>. \\<^sup>\ \<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} B) =\<^bsub>\\o\<^esub> (\x\<^bsub>\\<^esub>. \\<^sup>\ B)" using "\"[where A = "\\<^sup>\ \<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} B"] by (metis derived_substitution_simps(4)) then show ?thesis by (rule rule_RR[OF disjI1, where p = "[\,\,\,\]" and C = "?F"]) (use * in force)+ qed with assms(1) show ?thesis by (rule MP) qed lemmas Rule_C = prop_5245 (* synonym *) end diff --git a/thys/Q0_Metatheory/Semantics.thy b/thys/Q0_Metatheory/Semantics.thy --- a/thys/Q0_Metatheory/Semantics.thy +++ b/thys/Q0_Metatheory/Semantics.thy @@ -1,329 +1,329 @@ section \Semantics\ theory Semantics imports "ZFC_in_HOL.ZFC_Typeclasses" Syntax Boolean_Algebra begin no_notation funcset (infixr "\" 60) notation funcset (infixr "\" 60) abbreviation vfuncset :: "V \ V \ V" (infixr "\" 60) where "A \ B \ VPi A (\_. B)" notation app (infixl "\" 300) syntax "_vlambda" :: "pttrn \ V \ (V \ V) \ V" ("(3\<^bold>\_\<^bold>:_ \<^bold>./ _)" [0, 0, 3] 3) translations "\<^bold>\x \<^bold>: A\<^bold>. f" \ "CONST VLambda A (\x. f)" lemma vlambda_extensionality: assumes "\x. x \ elts A \ f x = g x" shows "(\<^bold>\x \<^bold>: A\<^bold>. f x) = (\<^bold>\x \<^bold>: A\<^bold>. g x)" unfolding VLambda_def using assms by auto subsection \Frames\ locale frame = fixes \ :: "type \ V" assumes truth_values_domain_def: "\ o = \" and function_domain_def: "\\ \. \ (\\\) \ \ \ \ \ \" and domain_nonemptiness: "\\. \ \ \ 0" begin lemma function_domainD: assumes "f \ elts (\ (\\\))" shows "f \ elts (\ \ \ \ \)" using assms and function_domain_def by blast lemma vlambda_from_function_domain: assumes "f \ elts (\ (\\\))" obtains b where "f = (\<^bold>\x \<^bold>: \ \\<^bold>. b x)" and "\x \ elts (\ \). b x \ elts (\ \)" using function_domainD[OF assms] by (metis VPi_D eta) lemma app_is_domain_respecting: assumes "f \ elts (\ (\\\))" and "x \ elts (\ \)" shows "f \ x \ elts (\ \)" by (fact VPi_D[OF function_domainD[OF assms(1)] assms(2)]) text \One-element function on @{term \\ \\}:\ definition one_element_function :: "V \ type \ V" ("{_}\<^bsub>_\<^esub>" [901, 0] 900) where [simp]: "{x}\<^bsub>\\<^esub> = (\<^bold>\y \<^bold>: \ \\<^bold>. bool_to_V (y = x))" lemma one_element_function_is_domain_respecting: shows "{x}\<^bsub>\\<^esub> \ elts (\ \ \ \ o)" unfolding one_element_function_def and truth_values_domain_def by (intro VPi_I) (simp, metis) lemma one_element_function_simps: shows "x \ elts (\ \) \ {x}\<^bsub>\\<^esub> \ x = \<^bold>T" and "\{x, y} \ elts (\ \); y \ x\ \ {x}\<^bsub>\\<^esub> \ y = \<^bold>F" by simp_all lemma one_element_function_injectivity: assumes "{x, x'} \ elts (\ i)" and "{x}\<^bsub>i\<^esub> = {x'}\<^bsub>i\<^esub>" shows "x = x'" using assms(1) and VLambda_eq_D2[OF assms(2)[unfolded one_element_function_def]] and injD[OF bool_to_V_injectivity] by blast lemma one_element_function_uniqueness: assumes "x \ elts (\ i)" shows "(SOME x'. x' \ elts (\ i) \ {x}\<^bsub>i\<^esub> = {x'}\<^bsub>i\<^esub>) = x" by (auto simp add: assms one_element_function_injectivity) text \Identity relation on @{term \\ \\}:\ definition identity_relation :: "type \ V" ("q\<^bsub>_\<^esub>" [0] 100) where [simp]: "q\<^bsub>\\<^esub> = (\<^bold>\x \<^bold>: \ \\<^bold>. {x}\<^bsub>\\<^esub>)" lemma identity_relation_is_domain_respecting: shows "q\<^bsub>\\<^esub> \ elts (\ \ \ \ \ \ \ o)" using VPi_I and one_element_function_is_domain_respecting by simp lemma q_is_equality: assumes "{x, y} \ elts (\ \)" shows "(q\<^bsub>\\<^esub>) \ x \ y = \<^bold>T \ x = y" unfolding identity_relation_def using assms and injD[OF bool_to_V_injectivity] by fastforce text \Unique member selector:\ definition is_unique_member_selector :: "V \ bool" where [iff]: "is_unique_member_selector f \ (\x \ elts (\ i). f \ {x}\<^bsub>i\<^esub> = x)" text \Assignment:\ definition is_assignment :: "(var \ V) \ bool" where [iff]: "is_assignment \ \ (\x \. \ (x, \) \ elts (\ \))" end abbreviation one_element_function_in ("{_}\<^bsub>_\<^esub>\<^bsup>_\<^esup>" [901, 0, 0] 900) where "{x}\<^bsub>\\<^esub>\<^bsup>\\<^esup> \ frame.one_element_function \ x \" abbreviation identity_relation_in ("q\<^bsub>_\<^esub>\<^bsup>_\<^esup>" [0, 0] 100) where "q\<^bsub>\\<^esub>\<^bsup>\\<^esup> \ frame.identity_relation \ \" text \ \\\ is a ``\v\-variant'' of \\\ if \\\ is an assignment that agrees with \\\ except possibly on \v\: \ definition is_variant_of :: "(var \ V) \ var \ (var \ V) \ bool" ("_ \\<^bsub>_\<^esub> _" [51, 0, 51] 50) where [iff]: "\ \\<^bsub>v\<^esub> \ \ (\v'. v' \ v \ \ v' = \ v')" subsection \Pre-models (interpretations)\ text \We use the term ``pre-model'' instead of ``interpretation'' since the latter is already a keyword:\ locale premodel = frame + fixes \ :: "con \ V" assumes Q_denotation: "\\. \ (Q_constant_of_type \) = q\<^bsub>\\<^esub>" and \_denotation: "is_unique_member_selector (\ iota_constant)" and non_logical_constant_denotation: "\c \. \ is_logical_constant (c, \) \ \ (c, \) \ elts (\ \)" begin text \Wff denotation function:\ definition is_wff_denotation_function :: "((var \ V) \ form \ V) \ bool" where [iff]: "is_wff_denotation_function \ \ ( \\. is_assignment \ \ (\A \. A \ wffs\<^bsub>\\<^esub> \ \ \ A \ elts (\ \)) \ \ \closure condition, see note in page 186\ (\x \. \ \ (x\<^bsub>\\<^esub>) = \ (x, \)) \ (\c \. \ \ (\c\\<^bsub>\\<^esub>) = \ (c, \)) \ (\A B \ \. A \ wffs\<^bsub>\\\\<^esub> \ B \ wffs\<^bsub>\\<^esub> \ \ \ (A \ B) = (\ \ A) \ (\ \ B)) \ (\x B \ \. B \ wffs\<^bsub>\\<^esub> \ \ \ (\x\<^bsub>\\<^esub>. B) = (\<^bold>\z \<^bold>: \ \\<^bold>. \ (\((x, \) := z)) B)) )" lemma wff_denotation_function_is_domain_respecting: assumes "is_wff_denotation_function \" and "A \ wffs\<^bsub>\\<^esub>" and "is_assignment \" shows "\ \ A \ elts (\ \)" using assms by force lemma wff_var_denotation: assumes "is_wff_denotation_function \" and "is_assignment \" shows "\ \ (x\<^bsub>\\<^esub>) = \ (x, \)" using assms by force lemma wff_Q_denotation: assumes "is_wff_denotation_function \" and "is_assignment \" shows "\ \ (Q\<^bsub>\\<^esub>) = q\<^bsub>\\<^esub>" using assms and Q_denotation by force lemma wff_iota_denotation: assumes "is_wff_denotation_function \" and "is_assignment \" shows "is_unique_member_selector (\ \ \)" using assms and \_denotation by fastforce lemma wff_non_logical_constant_denotation: assumes "is_wff_denotation_function \" and "is_assignment \" and "\ is_logical_constant (c, \)" shows "\ \ (\c\\<^bsub>\\<^esub>) = \ (c, \)" using assms by auto lemma wff_app_denotation: assumes "is_wff_denotation_function \" and "is_assignment \" and "A \ wffs\<^bsub>\\\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" shows "\ \ (A \ B) = \ \ A \ \ \ B" using assms by blast lemma wff_abs_denotation: assumes "is_wff_denotation_function \" and "is_assignment \" and "B \ wffs\<^bsub>\\<^esub>" shows "\ \ (\x\<^bsub>\\<^esub>. B) = (\<^bold>\z \<^bold>: \ \\<^bold>. \ (\((x, \) := z)) B)" using assms unfolding is_wff_denotation_function_def by metis lemma wff_denotation_function_is_uniquely_determined: assumes "is_wff_denotation_function \" and "is_wff_denotation_function \'" and "is_assignment \" and "A \ wffs" shows "\ \ A = \' \ A" proof - obtain \ where "A \ wffs\<^bsub>\\<^esub>" using assms(4) by blast then show ?thesis using assms(3) proof (induction A arbitrary: \) case var_is_wff with assms(1,2) show ?case by auto next case con_is_wff with assms(1,2) show ?case by auto next case app_is_wff with assms(1,2) show ?case using wff_app_denotation by metis next case (abs_is_wff \ A \ x) have "is_assignment (\((x, \) := z))" if "z \ elts (\ \)" for z using that and abs_is_wff.prems by simp then have *: "\ (\((x, \) := z)) A = \' (\((x, \) := z)) A" if "z \ elts (\ \)" for z using abs_is_wff.IH and that by blast have "\ \ (\x\<^bsub>\\<^esub>. A) = (\<^bold>\z \<^bold>: \ \\<^bold>. \ (\((x, \) := z)) A)" by (fact wff_abs_denotation[OF assms(1) abs_is_wff.prems abs_is_wff.hyps]) also have "\ = (\<^bold>\z \<^bold>: \ \\<^bold>. \' (\((x, \) := z)) A)" using * and vlambda_extensionality by fastforce also have "\ = \' \ (\x\<^bsub>\\<^esub>. A)" by (fact wff_abs_denotation[OF assms(2) abs_is_wff.prems abs_is_wff.hyps, symmetric]) finally show ?case . qed qed end subsection \General models\ type_synonym model_structure = "(type \ V) \ (con \ V) \ ((var \ V) \ form \ V)" text \ The assumption in the following locale implies that there must exist a function that is a wff denotation function for the pre-model, which is a requirement in the definition of general model - in @{cite "andrews:2002"}: + in \<^cite>\"andrews:2002"\: \ locale general_model = premodel + fixes \ :: "(var \ V) \ form \ V" assumes \_is_wff_denotation_function: "is_wff_denotation_function \" begin lemma mixed_beta_conversion: assumes "is_assignment \" and "y \ elts (\ \)" and "B \ wffs\<^bsub>\\<^esub>" shows "\ \ (\x\<^bsub>\\<^esub>. B) \ y = \ (\((x, \) := y)) B" using wff_abs_denotation[OF \_is_wff_denotation_function assms(1,3)] and beta[OF assms(2)] by simp lemma conj_fun_is_domain_respecting: assumes "is_assignment \" shows "\ \ (\\<^bsub>o\o\o\<^esub>) \ elts (\ (o\o\o))" using assms and conj_fun_wff and \_is_wff_denotation_function by auto lemma fully_applied_conj_fun_is_domain_respecting: assumes "is_assignment \" and "{x, y} \ elts (\ o)" shows "\ \ (\\<^bsub>o\o\o\<^esub>) \ x \ y \ elts (\ o)" using assms and conj_fun_is_domain_respecting and app_is_domain_respecting by (meson insert_subset) lemma imp_fun_denotation_is_domain_respecting: assumes "is_assignment \" shows "\ \ (\\<^bsub>o\o\o\<^esub>) \ elts (\ (o\o\o))" using assms and imp_fun_wff and \_is_wff_denotation_function by simp lemma fully_applied_imp_fun_denotation_is_domain_respecting: assumes "is_assignment \" and "{x, y} \ elts (\ o)" shows "\ \ (\\<^bsub>o\o\o\<^esub>) \ x \ y \ elts (\ o)" using assms and imp_fun_denotation_is_domain_respecting and app_is_domain_respecting by (meson insert_subset) end abbreviation is_general_model :: "model_structure \ bool" where "is_general_model \ \ case \ of (\, \, \) \ general_model \ \ \" subsection \Standard models\ locale standard_model = general_model + assumes full_function_domain_def: "\\ \. \ (\\\) = \ \ \ \ \" abbreviation is_standard_model :: "model_structure \ bool" where "is_standard_model \ \ case \ of (\, \, \) \ standard_model \ \ \" lemma standard_model_is_general_model: assumes "is_standard_model \" shows "is_general_model \" using assms and standard_model.axioms(1) by force subsection \Validity\ abbreviation is_assignment_into_frame ("_ \ _" [51, 51] 50) where "\ \ \ \ frame.is_assignment \ \" abbreviation is_assignment_into_model ("_ \\<^sub>M _" [51, 51] 50) where "\ \\<^sub>M \ \ (case \ of (\, \, \) \ \ \ \)" abbreviation satisfies ("_ \\<^bsub>_\<^esub> _" [50, 50, 50] 50) where "\ \\<^bsub>\\<^esub> A \ case \ of (\, \, \) \ \ \ A = \<^bold>T" abbreviation is_satisfiable_in where "is_satisfiable_in A \ \ \\. \ \\<^sub>M \ \ \ \\<^bsub>\\<^esub> A" abbreviation is_valid_in ("_ \ _" [50, 50] 50) where "\ \ A \ \\. \ \\<^sub>M \ \ \ \\<^bsub>\\<^esub> A" abbreviation is_valid_in_the_general_sense ("\ _" [50] 50) where "\ A \ \\. is_general_model \ \ \ \ A" abbreviation is_valid_in_the_standard_sense ("\\<^sub>S _" [50] 50) where "\\<^sub>S A \ \\. is_standard_model \ \ \ \ A" abbreviation is_true_sentence_in where "is_true_sentence_in A \ \ is_sentence A \ \ \\<^bsub>undefined\<^esub> A" \ \assignments are not meaningful\ abbreviation is_false_sentence_in where "is_false_sentence_in A \ \ is_sentence A \ \ \ \\<^bsub>undefined\<^esub> A" \ \assignments are not meaningful\ abbreviation is_model_for where "is_model_for \ \ \ \A \ \. \ \ A" lemma general_validity_in_standard_validity: assumes "\ A" shows "\\<^sub>S A" using assms and standard_model_is_general_model by blast end diff --git a/thys/Q0_Metatheory/Syntax.thy b/thys/Q0_Metatheory/Syntax.thy --- a/thys/Q0_Metatheory/Syntax.thy +++ b/thys/Q0_Metatheory/Syntax.thy @@ -1,3852 +1,3852 @@ section \Syntax\ theory Syntax imports "HOL-Library.Sublist" Utilities begin subsection \Type symbols\ datatype type = TInd ("i") | TBool ("o") | TFun type type (infixr "\" 101) primrec type_size :: "type \ nat" where "type_size i = 1" | "type_size o = 1" | "type_size (\ \ \) = Suc (type_size \ + type_size \)" primrec subtypes :: "type \ type set" where "subtypes i = {}" | "subtypes o = {}" | "subtypes (\ \ \) = {\, \} \ subtypes \ \ subtypes \" lemma subtype_size_decrease: assumes "\ \ subtypes \" shows "type_size \ < type_size \" using assms by (induction rule: type.induct) force+ lemma subtype_is_not_type: assumes "\ \ subtypes \" shows "\ \ \" using assms and subtype_size_decrease by blast lemma fun_type_atoms_in_subtypes: assumes "k < length ts" shows "ts ! k \ subtypes (foldr (\) ts \)" using assms by (induction ts arbitrary: k) (cases k, use less_Suc_eq_0_disj in \fastforce+\) lemma fun_type_atoms_neq_fun_type: assumes "k < length ts" shows "ts ! k \ foldr (\) ts \" by (fact fun_type_atoms_in_subtypes[OF assms, THEN subtype_is_not_type]) subsection \Variables\ text \ Unfortunately, the Nominal package does not support multi-sort atoms yet; therefore, we need to implement this support from scratch. \ type_synonym var = "nat \ type" abbreviation var_name :: "var \ nat" where "var_name \ fst" abbreviation var_type :: "var \ type" where "var_type \ snd" lemma fresh_var_existence: assumes "finite (vs :: var set)" obtains x where "(x, \) \ vs" using ex_new_if_finite[OF infinite_UNIV_nat] proof - from assms obtain x where "x \ var_name ` vs" using ex_new_if_finite[OF infinite_UNIV_nat] by fastforce with that show ?thesis by force qed lemma fresh_var_name_list_existence: assumes "finite (ns :: nat set)" obtains ns' where "length ns' = n" and "distinct ns'" and "lset ns' \ ns = {}" using assms proof (induction n arbitrary: thesis) case 0 then show ?case by simp next case (Suc n) from assms obtain ns' where "length ns' = n" and "distinct ns'" and "lset ns' \ ns = {}" using Suc.IH by blast moreover from assms obtain n' where "n' \ lset ns' \ ns" using ex_new_if_finite[OF infinite_UNIV_nat] by blast ultimately have "length (n' # ns') = Suc n" and "distinct (n' # ns')" and "lset (n' # ns') \ ns = {}" by simp_all with Suc.prems(1) show ?case by blast qed lemma fresh_var_list_existence: fixes xs :: "var list" and ns :: "nat set" assumes "finite ns" obtains vs' :: "var list" where "length vs' = length xs" and "distinct vs'" and "var_name ` lset vs' \ (ns \ var_name ` lset xs) = {}" and "map var_type vs' = map var_type xs" proof - from assms(1) have "finite (ns \ var_name ` lset xs)" by blast then obtain ns' where "length ns' = length xs" and "distinct ns'" and "lset ns' \ (ns \ var_name ` lset xs) = {}" by (rule fresh_var_name_list_existence) define vs'' where "vs'' = zip ns' (map var_type xs)" from vs''_def and \length ns' = length xs\ have "length vs'' = length xs" by simp moreover from vs''_def and \distinct ns'\ have "distinct vs''" by (simp add: distinct_zipI1) moreover have "var_name ` lset vs'' \ (ns \ var_name ` lset xs) = {}" unfolding vs''_def using \length ns' = length xs\ and \lset ns' \ (ns \ var_name ` lset xs) = {}\ by (metis length_map set_map map_fst_zip) moreover from vs''_def have "map var_type vs'' = map var_type xs" by (simp add: \length ns' = length xs\) ultimately show ?thesis by (fact that) qed subsection \Constants\ type_synonym con = "nat \ type" subsection \Formulas\ datatype form = FVar var | FCon con | FApp form form (infixl "\" 200) | FAbs var form syntax "_FVar" :: "nat \ type \ form" ("_\<^bsub>_\<^esub>" [899, 0] 900) "_FCon" :: "nat \ type \ form" ("\_\\<^bsub>_\<^esub>" [899, 0] 900) "_FAbs" :: "nat \ type \ form \ form" ("(4\_\<^bsub>_\<^esub>./ _)" [0, 0, 104] 104) translations "x\<^bsub>\\<^esub>" \ "CONST FVar (x, \)" "\c\\<^bsub>\\<^esub>" \ "CONST FCon (c, \)" "\x\<^bsub>\\<^esub>. A" \ "CONST FAbs (x, \) A" subsection \Generalized operators\ text \Generalized application. We define \\\<^sup>\\<^sub>\ A [B\<^sub>1, B\<^sub>2, \, B\<^sub>n]\ as \A \ B\<^sub>1 \ B\<^sub>2 \ \ \ B\<^sub>n\:\ definition generalized_app :: "form \ form list \ form" ("\\<^sup>\\<^sub>\ _ _" [241, 241] 241) where [simp]: "\\<^sup>\\<^sub>\ A Bs = foldl (\) A Bs" text \Generalized abstraction. We define \\\<^sup>\\<^sub>\ [x\<^sub>1, \, x\<^sub>n] A\ as \\x\<^sub>1. \ \x\<^sub>n. A\:\ definition generalized_abs :: "var list \ form \ form" ("\\<^sup>\\<^sub>\ _ _" [141, 141] 141) where [simp]: "\\<^sup>\\<^sub>\ vs A = foldr (\(x, \) B. \x\<^bsub>\\<^esub>. B) vs A" fun form_size :: "form \ nat" where "form_size (x\<^bsub>\\<^esub>) = 1" | "form_size (\c\\<^bsub>\\<^esub>) = 1" | "form_size (A \ B) = Suc (form_size A + form_size B)" | "form_size (\x\<^bsub>\\<^esub>. A) = Suc (form_size A)" fun form_depth :: "form \ nat" where "form_depth (x\<^bsub>\\<^esub>) = 0" | "form_depth (\c\\<^bsub>\\<^esub>) = 0" | "form_depth (A \ B) = Suc (max (form_depth A) (form_depth B))" | "form_depth (\x\<^bsub>\\<^esub>. A) = Suc (form_depth A)" subsection \Subformulas\ fun subforms :: "form \ form set" where "subforms (x\<^bsub>\\<^esub>) = {}" | "subforms (\c\\<^bsub>\\<^esub>) = {}" | "subforms (A \ B) = {A, B}" | "subforms (\x\<^bsub>\\<^esub>. A) = {A}" datatype direction = Left ("\") | Right ("\") type_synonym position = "direction list" fun positions :: "form \ position set" where "positions (x\<^bsub>\\<^esub>) = {[]}" | "positions (\c\\<^bsub>\\<^esub>) = {[]}" | "positions (A \ B) = {[]} \ {\ # p | p. p \ positions A} \ {\ # p | p. p \ positions B}" | "positions (\x\<^bsub>\\<^esub>. A) = {[]} \ {\ # p | p. p \ positions A}" lemma empty_is_position [simp]: shows "[] \ positions A" by (cases A rule: positions.cases) simp_all fun subform_at :: "form \ position \ form" where "subform_at A [] = Some A" | "subform_at (A \ B) (\ # p) = subform_at A p" | "subform_at (A \ B) (\ # p) = subform_at B p" | "subform_at (\x\<^bsub>\\<^esub>. A) (\ # p) = subform_at A p" | "subform_at _ _ = None" fun is_subform_at :: "form \ position \ form \ bool" ("(_ \\<^bsub>_\<^esub>/ _)" [51,0,51] 50) where "is_subform_at A [] A' = (A = A')" | "is_subform_at C (\ # p) (A \ B) = is_subform_at C p A" | "is_subform_at C (\ # p) (A \ B) = is_subform_at C p B" | "is_subform_at C (\ # p) (\x\<^bsub>\\<^esub>. A) = is_subform_at C p A" | "is_subform_at _ _ _ = False" lemma is_subform_at_alt_def: shows "A' \\<^bsub>p\<^esub> A = (case subform_at A p of Some B \ B = A' | None \ False)" by (induction A' p A rule: is_subform_at.induct) auto lemma superform_existence: assumes "B \\<^bsub>p @ [d]\<^esub> C" obtains A where "B \\<^bsub>[d]\<^esub> A" and "A \\<^bsub>p\<^esub> C" using assms by (induction B p C rule: is_subform_at.induct) auto lemma subform_at_subforms_con: assumes "\c\\<^bsub>\\<^esub> \\<^bsub>p\<^esub> C" shows "\A. A \\<^bsub>p @ [d]\<^esub> C" using assms by (induction "\c\\<^bsub>\\<^esub>" p C rule: is_subform_at.induct) auto lemma subform_at_subforms_var: assumes "x\<^bsub>\\<^esub> \\<^bsub>p\<^esub> C" shows "\A. A \\<^bsub>p @ [d]\<^esub> C" using assms by (induction "x\<^bsub>\\<^esub>" p C rule: is_subform_at.induct) auto lemma subform_at_subforms_app: assumes "A \ B \\<^bsub>p\<^esub> C" shows "A \\<^bsub>p @ [\]\<^esub> C" and "B \\<^bsub>p @ [\]\<^esub> C" using assms by (induction "A \ B" p C rule: is_subform_at.induct) auto lemma subform_at_subforms_abs: assumes "\x\<^bsub>\\<^esub>. A \\<^bsub>p\<^esub> C" shows "A \\<^bsub>p @ [\]\<^esub> C" using assms by (induction "\x\<^bsub>\\<^esub>. A" p C rule: is_subform_at.induct) auto lemma is_subform_implies_in_positions: assumes "B \\<^bsub>p\<^esub> A" shows "p \ positions A" using assms by (induction rule: is_subform_at.induct) simp_all lemma subform_size_decrease: assumes "A \\<^bsub>p\<^esub> B" and "p \ []" shows "form_size A < form_size B" using assms by (induction A p B rule: is_subform_at.induct) force+ lemma strict_subform_is_not_form: assumes "p \ []" and "A' \\<^bsub>p\<^esub> A" shows "A' \ A" using assms and subform_size_decrease by blast lemma no_right_subform_of_abs: shows "\B. B \\<^bsub>\ # p\<^esub> \x\<^bsub>\\<^esub>. A" by simp lemma subforms_from_var: assumes "A \\<^bsub>p\<^esub> x\<^bsub>\\<^esub>" shows "A = x\<^bsub>\\<^esub>" and "p = []" using assms by (auto elim: is_subform_at.elims) lemma subforms_from_con: assumes "A \\<^bsub>p\<^esub> \c\\<^bsub>\\<^esub>" shows "A = \c\\<^bsub>\\<^esub>" and "p = []" using assms by (auto elim: is_subform_at.elims) lemma subforms_from_app: assumes "A \\<^bsub>p\<^esub> B \ C" shows " (A = B \ C \ p = []) \ (A \ B \ C \ (\p' \ positions B. p = \ # p' \ A \\<^bsub>p'\<^esub> B) \ (\p' \ positions C. p = \ # p' \ A \\<^bsub>p'\<^esub> C))" using assms and strict_subform_is_not_form by (auto simp add: is_subform_implies_in_positions elim: is_subform_at.elims) lemma subforms_from_abs: assumes "A \\<^bsub>p\<^esub> \x\<^bsub>\\<^esub>. B" shows "(A = \x\<^bsub>\\<^esub>. B \ p = []) \ (A \ \x\<^bsub>\\<^esub>. B \ (\p' \ positions B. p = \ # p' \ A \\<^bsub>p'\<^esub> B))" using assms and strict_subform_is_not_form by (auto simp add: is_subform_implies_in_positions elim: is_subform_at.elims) lemma leftmost_subform_in_generalized_app: shows "B \\<^bsub>replicate (length As) \\<^esub> \\<^sup>\\<^sub>\ B As" by (induction As arbitrary: B) (simp_all, metis replicate_append_same subform_at_subforms_app(1)) lemma self_subform_is_at_top: assumes "A \\<^bsub>p\<^esub> A" shows "p = []" using assms and strict_subform_is_not_form by blast lemma at_top_is_self_subform: assumes "A \\<^bsub>[]\<^esub> B" shows "A = B" using assms by (auto elim: is_subform_at.elims) lemma is_subform_at_uniqueness: assumes "B \\<^bsub>p\<^esub> A" and "C \\<^bsub>p\<^esub> A" shows "B = C" using assms by (induction A arbitrary: p B C) (auto elim: is_subform_at.elims) lemma is_subform_at_existence: assumes "p \ positions A" obtains B where "B \\<^bsub>p\<^esub> A" using assms by (induction A arbitrary: p) (auto elim: is_subform_at.elims, blast+) lemma is_subform_at_transitivity: assumes "A \\<^bsub>p\<^sub>1\<^esub> B" and "B \\<^bsub>p\<^sub>2\<^esub> C" shows "A \\<^bsub>p\<^sub>2 @ p\<^sub>1\<^esub> C" using assms by (induction B p\<^sub>2 C arbitrary: A p\<^sub>1 rule: is_subform_at.induct) simp_all lemma subform_nesting: assumes "strict_prefix p' p" and "B \\<^bsub>p'\<^esub> A" and "C \\<^bsub>p\<^esub> A" shows "C \\<^bsub>drop (length p') p\<^esub> B" proof - from assms(1) have "p \ []" using strict_prefix_simps(1) by blast with assms(1,3) show ?thesis proof (induction p arbitrary: C rule: rev_induct) case Nil then show ?case by blast next case (snoc d p'') then show ?case proof (cases "p'' = p'") case True obtain A' where "C \\<^bsub>[d]\<^esub> A'" and "A' \\<^bsub>p'\<^esub> A" by (fact superform_existence[OF snoc.prems(2)[unfolded True]]) from \A' \\<^bsub>p'\<^esub> A\ and assms(2) have "A' = B" by (rule is_subform_at_uniqueness) with \C \\<^bsub>[d]\<^esub> A'\ have "C \\<^bsub>[d]\<^esub> B" by (simp only:) with True show ?thesis by auto next case False with snoc.prems(1) have "strict_prefix p' p''" using prefix_order.dual_order.strict_implies_order by fastforce then have "p'' \ []" by force moreover from snoc.prems(2) obtain A' where "C \\<^bsub>[d]\<^esub> A'" and "A' \\<^bsub>p''\<^esub> A" using superform_existence by blast ultimately have "A' \\<^bsub>drop (length p') p''\<^esub> B" using snoc.IH and \strict_prefix p' p''\ by blast with \C \\<^bsub>[d]\<^esub> A'\ and snoc.prems(1) show ?thesis using is_subform_at_transitivity and prefix_length_less by fastforce qed qed qed lemma loop_subform_impossibility: assumes "B \\<^bsub>p\<^esub> A" and "strict_prefix p' p" shows "\ B \\<^bsub>p'\<^esub> A" using assms and prefix_length_less and self_subform_is_at_top and subform_nesting by fastforce lemma nested_subform_size_decreases: assumes "strict_prefix p' p" and "B \\<^bsub>p'\<^esub> A" and "C \\<^bsub>p\<^esub> A" shows "form_size C < form_size B" proof - from assms(1) have "p \ []" by force have "C \\<^bsub>drop (length p') p\<^esub> B" by (fact subform_nesting[OF assms]) moreover have "drop (length p') p \ []" using prefix_length_less[OF assms(1)] by force ultimately show ?thesis using subform_size_decrease by simp qed definition is_subform :: "form \ form \ bool" (infix "\" 50) where [simp]: "A \ B = (\p. A \\<^bsub>p\<^esub> B)" instantiation form :: ord begin definition "A \ B \ A \ B" definition "A < B \ A \ B \ A \ B" instance .. end instance form :: preorder proof (standard, unfold less_eq_form_def less_form_def) fix A show "A \ A" unfolding is_subform_def using is_subform_at.simps(1) by blast next fix A and B and C assume "A \ B" and "B \ C" then show "A \ C" unfolding is_subform_def using is_subform_at_transitivity by blast next fix A and B show "A \ B \ A \ B \ A \ B \ \ B \ A" unfolding is_subform_def by (metis is_subform_at.simps(1) not_less_iff_gr_or_eq subform_size_decrease) qed lemma position_subform_existence_equivalence: shows "p \ positions A \ (\A'. A' \\<^bsub>p\<^esub> A)" by (meson is_subform_at_existence is_subform_implies_in_positions) lemma position_prefix_is_position: assumes "p \ positions A" and "prefix p' p" shows "p' \ positions A" using assms proof (induction p rule: rev_induct) case Nil then show ?case by simp next case (snoc d p'') from snoc.prems(1) have "p'' \ positions A" by (meson position_subform_existence_equivalence superform_existence) with snoc.prems(1,2) show ?case using snoc.IH by fastforce qed subsection \Free and bound variables\ consts vars :: "'a \ var set" overloading "vars_form" \ "vars :: form \ var set" "vars_form_set" \ "vars :: form set \ var set" (* abuse of notation *) begin fun vars_form :: "form \ var set" where "vars_form (x\<^bsub>\\<^esub>) = {(x, \)}" | "vars_form (\c\\<^bsub>\\<^esub>) = {}" | "vars_form (A \ B) = vars_form A \ vars_form B" | "vars_form (\x\<^bsub>\\<^esub>. A) = vars_form A \ {(x, \)}" fun vars_form_set :: "form set \ var set" where "vars_form_set S = (\A \ S. vars A)" end abbreviation var_names :: "'a \ nat set" where "var_names \ \ var_name ` (vars \)" lemma vars_form_finiteness: fixes A :: form shows "finite (vars A)" by (induction rule: vars_form.induct) simp_all lemma vars_form_set_finiteness: fixes S :: "form set" assumes "finite S" shows "finite (vars S)" using assms unfolding vars_form_set.simps using vars_form_finiteness by blast lemma form_var_names_finiteness: fixes A :: form shows "finite (var_names A)" using vars_form_finiteness by blast lemma form_set_var_names_finiteness: fixes S :: "form set" assumes "finite S" shows "finite (var_names S)" using assms and vars_form_set_finiteness by blast consts free_vars :: "'a \ var set" overloading "free_vars_form" \ "free_vars :: form \ var set" "free_vars_form_set" \ "free_vars :: form set \ var set" (* abuse of notation *) begin fun free_vars_form :: "form \ var set" where "free_vars_form (x\<^bsub>\\<^esub>) = {(x, \)}" | "free_vars_form (\c\\<^bsub>\\<^esub>) = {}" | "free_vars_form (A \ B) = free_vars_form A \ free_vars_form B" | "free_vars_form (\x\<^bsub>\\<^esub>. A) = free_vars_form A - {(x, \)}" fun free_vars_form_set :: "form set \ var set" where "free_vars_form_set S = (\A \ S. free_vars A)" end abbreviation free_var_names :: "'a \ nat set" where "free_var_names \ \ var_name ` (free_vars \)" lemma free_vars_form_finiteness: fixes A :: form shows "finite (free_vars A)" by (induction rule: free_vars_form.induct) simp_all lemma free_vars_of_generalized_app: shows "free_vars (\\<^sup>\\<^sub>\ A Bs) = free_vars A \ free_vars (lset Bs)" by (induction Bs arbitrary: A) auto lemma free_vars_of_generalized_abs: shows "free_vars (\\<^sup>\\<^sub>\ vs A) = free_vars A - lset vs" by (induction vs arbitrary: A) auto lemma free_vars_in_all_vars: fixes A :: form shows "free_vars A \ vars A" proof (induction A) case (FVar v) then show ?case using surj_pair[of v] by force next case (FCon k) then show ?case using surj_pair[of k] by force next case (FApp A B) have "free_vars (A \ B) = free_vars A \ free_vars B" using free_vars_form.simps(3) . also from FApp.IH have "\ \ vars A \ vars B" by blast also have "\ = vars (A \ B)" using vars_form.simps(3)[symmetric] . finally show ?case by (simp only:) next case (FAbs v A) then show ?case using surj_pair[of v] by force qed lemma free_vars_in_all_vars_set: fixes S :: "form set" shows "free_vars S \ vars S" using free_vars_in_all_vars by fastforce lemma singleton_form_set_vars: shows "vars {FVar y} = {y}" using surj_pair[of y] by force fun bound_vars where "bound_vars (x\<^bsub>\\<^esub>) = {}" | "bound_vars (\c\\<^bsub>\\<^esub>) = {}" | "bound_vars (B \ C) = bound_vars B \ bound_vars C" | "bound_vars (\x\<^bsub>\\<^esub>. B) = {(x, \)} \ bound_vars B" lemma vars_is_free_and_bound_vars: shows "vars A = free_vars A \ bound_vars A" by (induction A) auto fun binders_at :: "form \ position \ var set" where "binders_at (A \ B) (\ # p) = binders_at A p" | "binders_at (A \ B) (\ # p) = binders_at B p" | "binders_at (\x\<^bsub>\\<^esub>. A) (\ # p) = {(x, \)} \ binders_at A p" | "binders_at A [] = {}" | "binders_at A p = {}" lemma binders_at_concat: assumes "A' \\<^bsub>p\<^esub> A" shows "binders_at A (p @ p') = binders_at A p \ binders_at A' p'" using assms by (induction p A rule: is_subform_at.induct) auto subsection \Free and bound occurrences\ definition occurs_at :: "var \ position \ form \ bool" where [iff]: "occurs_at v p B \ (FVar v \\<^bsub>p\<^esub> B)" lemma occurs_at_alt_def: shows "occurs_at v [] (FVar v') \ (v = v')" and "occurs_at v p (\c\\<^bsub>\\<^esub>) \ False" and "occurs_at v (\ # p) (A \ B) \ occurs_at v p A" and "occurs_at v (\ # p) (A \ B) \ occurs_at v p B" and "occurs_at v (\ # p) (\x\<^bsub>\\<^esub>. A) \ occurs_at v p A" and "occurs_at v (d # p) (FVar v') \ False" and "occurs_at v (\ # p) (\x\<^bsub>\\<^esub>. A) \ False" and "occurs_at v [] (A \ B) \ False" and "occurs_at v [] (\x\<^bsub>\\<^esub>. A) \ False" by (fastforce elim: is_subform_at.elims)+ definition occurs :: "var \ form \ bool" where [iff]: "occurs v B \ (\p \ positions B. occurs_at v p B)" lemma occurs_in_vars: assumes "occurs v A" shows "v \ vars A" using assms by (induction A) force+ abbreviation strict_prefixes where "strict_prefixes xs \ [ys \ prefixes xs. ys \ xs]" definition in_scope_of_abs :: "var \ position \ form \ bool" where [iff]: "in_scope_of_abs v p B \ ( p \ [] \ ( \p' \ lset (strict_prefixes p). case (subform_at B p') of Some (FAbs v' _) \ v = v' | _ \ False ) )" lemma in_scope_of_abs_alt_def: shows " in_scope_of_abs v p B \ p \ [] \ (\p' \ positions B. \C. strict_prefix p' p \ FAbs v C \\<^bsub>p'\<^esub> B)" proof assume "in_scope_of_abs v p B" then show "p \ [] \ (\p' \ positions B. \C. strict_prefix p' p \ FAbs v C \\<^bsub>p'\<^esub> B)" by (induction rule: subform_at.induct) force+ next assume "p \ [] \ (\p' \ positions B. \C. strict_prefix p' p \ FAbs v C \\<^bsub>p'\<^esub> B)" then show "in_scope_of_abs v p B" by (induction rule: subform_at.induct) fastforce+ qed lemma in_scope_of_abs_in_left_app: shows "in_scope_of_abs v (\ # p) (A \ B) \ in_scope_of_abs v p A" by force lemma in_scope_of_abs_in_right_app: shows "in_scope_of_abs v (\ # p) (A \ B) \ in_scope_of_abs v p B" by force lemma in_scope_of_abs_in_app: assumes "in_scope_of_abs v p (A \ B)" obtains p' where "(p = \ # p' \ in_scope_of_abs v p' A) \ (p = \ # p' \ in_scope_of_abs v p' B)" proof - from assms obtain d and p' where "p = d # p'" unfolding in_scope_of_abs_def by (meson list.exhaust) then show ?thesis proof (cases d) case Left with assms and \p = d # p'\ show ?thesis using that and in_scope_of_abs_in_left_app by simp next case Right with assms and \p = d # p'\ show ?thesis using that and in_scope_of_abs_in_right_app by simp qed qed lemma not_in_scope_of_abs_in_app: assumes " \p'. (p = \ # p' \ \ in_scope_of_abs v' p' A) \ (p = \ # p' \ \ in_scope_of_abs v' p' B)" shows "\ in_scope_of_abs v' p (A \ B)" using assms and in_scope_of_abs_in_app by metis lemma in_scope_of_abs_in_abs: shows "in_scope_of_abs v (\ # p) (FAbs v' B) \ v = v' \ in_scope_of_abs v p B" proof assume "in_scope_of_abs v (\ # p) (FAbs v' B)" then obtain p' and C where "p' \ positions (FAbs v' B)" and "strict_prefix p' (\ # p)" and "FAbs v C \\<^bsub>p'\<^esub> FAbs v' B" unfolding in_scope_of_abs_alt_def by blast then show "v = v' \ in_scope_of_abs v p B" proof (cases p') case Nil with \FAbs v C \\<^bsub>p'\<^esub> FAbs v' B\ have "v = v'" by auto then show ?thesis by simp next case (Cons d p'') with \strict_prefix p' (\ # p)\ have "d = \" by simp from \FAbs v C \\<^bsub>p'\<^esub> FAbs v' B\ and Cons have "p'' \ positions B" by (cases "(FAbs v C, p', FAbs v' B)" rule: is_subform_at.cases) (simp_all add: is_subform_implies_in_positions) moreover from \FAbs v C \\<^bsub>p'\<^esub> FAbs v' B\ and Cons and \d = \\ have "FAbs v C \\<^bsub>p''\<^esub> B" by (metis is_subform_at.simps(4) old.prod.exhaust) moreover from \strict_prefix p' (\ # p)\ and Cons have "strict_prefix p'' p" by auto ultimately have "in_scope_of_abs v p B" using in_scope_of_abs_alt_def by auto then show ?thesis by simp qed next assume "v = v' \ in_scope_of_abs v p B" then show "in_scope_of_abs v (\ # p) (FAbs v' B)" unfolding in_scope_of_abs_alt_def using position_subform_existence_equivalence and surj_pair[of v'] by force qed lemma not_in_scope_of_abs_in_var: shows "\ in_scope_of_abs v p (FVar v')" unfolding in_scope_of_abs_def by (cases p) simp_all lemma in_scope_of_abs_in_vars: assumes "in_scope_of_abs v p A" shows "v \ vars A" using assms proof (induction A arbitrary: p) case (FVar v') then show ?case using not_in_scope_of_abs_in_var by blast next case (FCon k) then show ?case using in_scope_of_abs_alt_def by (blast elim: is_subform_at.elims(2)) next case (FApp B C) from FApp.prems obtain d and p' where "p = d # p'" unfolding in_scope_of_abs_def by (meson neq_Nil_conv) then show ?case proof (cases d) case Left with FApp.prems and \p = d # p'\ have "in_scope_of_abs v p' B" using in_scope_of_abs_in_left_app by blast then have "v \ vars B" by (fact FApp.IH(1)) then show ?thesis by simp next case Right with FApp.prems and \p = d # p'\ have "in_scope_of_abs v p' C" using in_scope_of_abs_in_right_app by blast then have "v \ vars C" by (fact FApp.IH(2)) then show ?thesis by simp qed next case (FAbs v' B) then show ?case proof (cases "v = v'") case True then show ?thesis using surj_pair[of v] by force next case False with FAbs.prems obtain p' and d where "p = d # p'" unfolding in_scope_of_abs_def by (meson neq_Nil_conv) then show ?thesis proof (cases d) case Left with FAbs.prems and False and \p = d # p'\ have "in_scope_of_abs v p' B" using in_scope_of_abs_in_abs by blast then have "v \ vars B" by (fact FAbs.IH) then show ?thesis using surj_pair[of v'] by force next case Right with FAbs.prems and \p = d # p'\ and False show ?thesis by (cases rule: is_subform_at.cases) auto qed qed qed lemma binders_at_alt_def: assumes "p \ positions A" shows "binders_at A p = {v | v. in_scope_of_abs v p A}" using assms and in_set_prefixes by (induction rule: binders_at.induct) auto definition is_bound_at :: "var \ position \ form \ bool" where [iff]: "is_bound_at v p B \ occurs_at v p B \ in_scope_of_abs v p B" lemma not_is_bound_at_in_var: shows "\ is_bound_at v p (FVar v')" by (fastforce elim: is_subform_at.elims(2)) lemma not_is_bound_at_in_con: shows "\ is_bound_at v p (FCon k)" by (fastforce elim: is_subform_at.elims(2)) lemma is_bound_at_in_left_app: shows "is_bound_at v (\ # p) (B \ C) \ is_bound_at v p B" by auto lemma is_bound_at_in_right_app: shows "is_bound_at v (\ # p) (B \ C) \ is_bound_at v p C" by auto lemma is_bound_at_from_app: assumes "is_bound_at v p (B \ C)" obtains p' where "(p = \ # p' \ is_bound_at v p' B) \ (p = \ # p' \ is_bound_at v p' C)" proof - from assms obtain d and p' where "p = d # p'" using subforms_from_app by blast then show ?thesis proof (cases d) case Left with assms and that and \p = d # p'\ show ?thesis using is_bound_at_in_left_app by simp next case Right with assms and that and \p = d # p'\ show ?thesis using is_bound_at_in_right_app by simp qed qed lemma is_bound_at_from_abs: assumes "is_bound_at v (\ # p) (FAbs v' B)" shows "v = v' \ is_bound_at v p B" using assms by (fastforce elim: is_subform_at.elims) lemma is_bound_at_from_absE: assumes "is_bound_at v p (FAbs v' B)" obtains p' where "p = \ # p'" and "v = v' \ is_bound_at v p' B" proof - obtain x and \ where "v' = (x, \)" by fastforce with assms obtain p' where "p = \ # p'" using subforms_from_abs by blast with assms and that show ?thesis using is_bound_at_from_abs by simp qed lemma is_bound_at_to_abs: assumes "(v = v' \ occurs_at v p B) \ is_bound_at v p B" shows "is_bound_at v (\ # p) (FAbs v' B)" unfolding is_bound_at_def proof from assms(1) show "occurs_at v (\ # p) (FAbs v' B)" using surj_pair[of v'] by force from assms show "in_scope_of_abs v (\ # p) (FAbs v' B)" using in_scope_of_abs_in_abs by auto qed lemma is_bound_at_in_bound_vars: assumes "p \ positions A" and "is_bound_at v p A \ v \ binders_at A p" shows "v \ bound_vars A" using assms proof (induction A arbitrary: p) case (FApp B C) from FApp.prems(2) consider (a) "is_bound_at v p (B \ C)" | (b) "v \ binders_at (B \ C) p" by blast then show ?case proof cases case a then have "p \ []" using occurs_at_alt_def(8) by blast then obtain d and p' where "p = d # p'" by (meson list.exhaust) with \p \ positions (B \ C)\ consider (a\<^sub>1) "p = \ # p'" and "p' \ positions B" | (a\<^sub>2) "p = \ # p'" and "p' \ positions C" by force then show ?thesis proof cases case a\<^sub>1 from a\<^sub>1(1) and \is_bound_at v p (B \ C)\ have "is_bound_at v p' B" using is_bound_at_in_left_app by blast with a\<^sub>1(2) have "v \ bound_vars B" using FApp.IH(1) by blast then show ?thesis by simp next case a\<^sub>2 from a\<^sub>2(1) and \is_bound_at v p (B \ C)\ have "is_bound_at v p' C" using is_bound_at_in_right_app by blast with a\<^sub>2(2) have "v \ bound_vars C" using FApp.IH(2) by blast then show ?thesis by simp qed next case b then have "p \ []" by force then obtain d and p' where "p = d # p'" by (meson list.exhaust) with \p \ positions (B \ C)\ consider (b\<^sub>1) "p = \ # p'" and "p' \ positions B" | (b\<^sub>2) "p = \ # p'" and "p' \ positions C" by force then show ?thesis proof cases case b\<^sub>1 from b\<^sub>1(1) and \v \ binders_at (B \ C) p\ have "v \ binders_at B p'" by force with b\<^sub>1(2) have "v \ bound_vars B" using FApp.IH(1) by blast then show ?thesis by simp next case b\<^sub>2 from b\<^sub>2(1) and \v \ binders_at (B \ C) p\ have "v \ binders_at C p'" by force with b\<^sub>2(2) have "v \ bound_vars C" using FApp.IH(2) by blast then show ?thesis by simp qed qed next case (FAbs v' B) from FAbs.prems(2) consider (a) "is_bound_at v p (FAbs v' B)" | (b) "v \ binders_at (FAbs v' B) p" by blast then show ?case proof cases case a then have "p \ []" using occurs_at_alt_def(9) by force with \p \ positions (FAbs v' B)\ obtain p' where "p = \ # p'" and "p' \ positions B" by (cases "FAbs v' B" rule: positions.cases) fastforce+ from \p = \ # p'\ and \is_bound_at v p (FAbs v' B)\ have "v = v' \ is_bound_at v p' B" using is_bound_at_from_abs by blast then consider (a\<^sub>1) "v = v'" | (a\<^sub>2) "is_bound_at v p' B" by blast then show ?thesis proof cases case a\<^sub>1 then show ?thesis using surj_pair[of v'] by fastforce next case a\<^sub>2 then have "v \ bound_vars B" using \p' \ positions B\ and FAbs.IH by blast then show ?thesis using surj_pair[of v'] by fastforce qed next case b then have "p \ []" by force with FAbs.prems(1) obtain p' where "p = \ # p'" and "p' \ positions B" by (cases "FAbs v' B" rule: positions.cases) fastforce+ with b consider (b\<^sub>1) "v = v'" | (b\<^sub>2) "v \ binders_at B p'" by (cases "FAbs v' B" rule: positions.cases) fastforce+ then show ?thesis proof cases case b\<^sub>1 then show ?thesis using surj_pair[of v'] by fastforce next case b\<^sub>2 then have "v \ bound_vars B" using \p' \ positions B\ and FAbs.IH by blast then show ?thesis using surj_pair[of v'] by fastforce qed qed qed fastforce+ lemma bound_vars_in_is_bound_at: assumes "v \ bound_vars A" obtains p where "p \ positions A" and "is_bound_at v p A \ v \ binders_at A p" using assms proof (induction A arbitrary: thesis rule: bound_vars.induct) case (3 B C) from \v \ bound_vars (B \ C)\ consider (a) "v \ bound_vars B" | (b) "v \ bound_vars C" by fastforce then show ?case proof cases case a with "3.IH"(1) obtain p where "p \ positions B" and "is_bound_at v p B \ v \ binders_at B p" by blast from \p \ positions B\ have "\ # p \ positions (B \ C)" by simp from \is_bound_at v p B \ v \ binders_at B p\ consider (a\<^sub>1) "is_bound_at v p B" | (a\<^sub>2) "v \ binders_at B p" by blast then show ?thesis proof cases case a\<^sub>1 then have "is_bound_at v (\ # p) (B \ C)" using is_bound_at_in_left_app by blast then show ?thesis using "3.prems"(1) and is_subform_implies_in_positions by blast next case a\<^sub>2 then have "v \ binders_at (B \ C) (\ # p)" by simp then show ?thesis using "3.prems"(1) and \\ # p \ positions (B \ C)\ by blast qed next case b with "3.IH"(2) obtain p where "p \ positions C" and "is_bound_at v p C \ v \ binders_at C p" by blast from \p \ positions C\ have "\ # p \ positions (B \ C)" by simp from \is_bound_at v p C \ v \ binders_at C p\ consider (b\<^sub>1) "is_bound_at v p C" | (b\<^sub>2) "v \ binders_at C p" by blast then show ?thesis proof cases case b\<^sub>1 then have "is_bound_at v (\ # p) (B \ C)" using is_bound_at_in_right_app by blast then show ?thesis using "3.prems"(1) and is_subform_implies_in_positions by blast next case b\<^sub>2 then have "v \ binders_at (B \ C) (\ # p)" by simp then show ?thesis using "3.prems"(1) and \\ # p \ positions (B \ C)\ by blast qed qed next case (4 x \ B) from \v \ bound_vars (\x\<^bsub>\\<^esub>. B)\ consider (a) "v = (x, \)" | (b) "v \ bound_vars B" by force then show ?case proof cases case a then have "v \ binders_at (\x\<^bsub>\\<^esub>. B) [\]" by simp then show ?thesis using "4.prems"(1) and is_subform_implies_in_positions by fastforce next case b with "4.IH"(1) obtain p where "p \ positions B" and "is_bound_at v p B \ v \ binders_at B p" by blast from \p \ positions B\ have "\ # p \ positions (\x\<^bsub>\\<^esub>. B)" by simp from \is_bound_at v p B \ v \ binders_at B p\ consider (b\<^sub>1) "is_bound_at v p B" | (b\<^sub>2) "v \ binders_at B p" by blast then show ?thesis proof cases case b\<^sub>1 then have "is_bound_at v (\ # p) (\x\<^bsub>\\<^esub>. B)" using is_bound_at_to_abs by blast then show ?thesis using "4.prems"(1) and \\ # p \ positions (\x\<^bsub>\\<^esub>. B)\ by blast next case b\<^sub>2 then have "v \ binders_at (\x\<^bsub>\\<^esub>. B) (\ # p)" by simp then show ?thesis using "4.prems"(1) and \\ # p \ positions (\x\<^bsub>\\<^esub>. B)\ by blast qed qed qed simp_all lemma bound_vars_alt_def: shows "bound_vars A = {v | v p. p \ positions A \ (is_bound_at v p A \ v \ binders_at A p)}" using bound_vars_in_is_bound_at and is_bound_at_in_bound_vars by (intro subset_antisym subsetI CollectI, metis) blast definition is_free_at :: "var \ position \ form \ bool" where [iff]: "is_free_at v p B \ occurs_at v p B \ \ in_scope_of_abs v p B" lemma is_free_at_in_var: shows "is_free_at v [] (FVar v') \ v = v'" by simp lemma not_is_free_at_in_con: shows "\ is_free_at v [] (\c\\<^bsub>\\<^esub>)" by simp lemma is_free_at_in_left_app: shows "is_free_at v (\ # p) (B \ C) \ is_free_at v p B" by auto lemma is_free_at_in_right_app: shows "is_free_at v (\ # p) (B \ C) \ is_free_at v p C" by auto lemma is_free_at_from_app: assumes "is_free_at v p (B \ C)" obtains p' where "(p = \ # p' \ is_free_at v p' B) \ (p = \ # p' \ is_free_at v p' C)" proof - from assms obtain d and p' where "p = d # p'" using subforms_from_app by blast then show ?thesis proof (cases d) case Left with assms and that and \p = d # p'\ show ?thesis using is_free_at_in_left_app by blast next case Right with assms and that and \p = d # p'\ show ?thesis using is_free_at_in_right_app by blast qed qed lemma is_free_at_from_abs: assumes "is_free_at v (\ # p) (FAbs v' B)" shows "is_free_at v p B" using assms by (fastforce elim: is_subform_at.elims) lemma is_free_at_from_absE: assumes "is_free_at v p (FAbs v' B)" obtains p' where "p = \ # p'" and "is_free_at v p' B" proof - obtain x and \ where "v' = (x, \)" by fastforce with assms obtain p' where "p = \ # p'" using subforms_from_abs by blast with assms and that show ?thesis using is_free_at_from_abs by blast qed lemma is_free_at_to_abs: assumes "is_free_at v p B" and "v \ v'" shows "is_free_at v (\ # p) (FAbs v' B)" unfolding is_free_at_def proof from assms(1) show "occurs_at v (\ # p) (FAbs v' B)" using surj_pair[of v'] by fastforce from assms show "\ in_scope_of_abs v (\ # p) (FAbs v' B)" unfolding is_free_at_def using in_scope_of_abs_in_abs by presburger qed lemma is_free_at_in_free_vars: assumes "p \ positions A" and "is_free_at v p A" shows "v \ free_vars A" using assms proof (induction A arbitrary: p) case (FApp B C) from \is_free_at v p (B \ C)\ have "p \ []" using occurs_at_alt_def(8) by blast then obtain d and p' where "p = d # p'" by (meson list.exhaust) with \p \ positions (B \ C)\ consider (a) "p = \ # p'" and "p' \ positions B" | (b) "p = \ # p'" and "p' \ positions C" by force then show ?case proof cases case a from a(1) and \is_free_at v p (B \ C)\ have "is_free_at v p' B" using is_free_at_in_left_app by blast with a(2) have "v \ free_vars B" using FApp.IH(1) by blast then show ?thesis by simp next case b from b(1) and \is_free_at v p (B \ C)\ have "is_free_at v p' C" using is_free_at_in_right_app by blast with b(2) have "v \ free_vars C" using FApp.IH(2) by blast then show ?thesis by simp qed next case (FAbs v' B) from \is_free_at v p (FAbs v' B)\ have "p \ []" using occurs_at_alt_def(9) by force with \p \ positions (FAbs v' B)\ obtain p' where "p = \ # p'" and "p' \ positions B" by (cases "FAbs v' B" rule: positions.cases) fastforce+ moreover from \p = \ # p'\ and \is_free_at v p (FAbs v' B)\ have "is_free_at v p' B" using is_free_at_from_abs by blast ultimately have "v \ free_vars B" using FAbs.IH by simp moreover from \p = \ # p'\ and \is_free_at v p (FAbs v' B)\ have "v \ v'" using in_scope_of_abs_in_abs by blast ultimately show ?case using surj_pair[of v'] by force qed fastforce+ lemma free_vars_in_is_free_at: assumes "v \ free_vars A" obtains p where "p \ positions A" and "is_free_at v p A" using assms proof (induction A arbitrary: thesis rule: free_vars_form.induct) case (3 A B) from \v \ free_vars (A \ B)\ consider (a) "v \ free_vars A" | (b) "v \ free_vars B" by fastforce then show ?case proof cases case a with "3.IH"(1) obtain p where "p \ positions A" and "is_free_at v p A" by blast from \p \ positions A\ have "\ # p \ positions (A \ B)" by simp moreover from \is_free_at v p A\ have "is_free_at v (\ # p) (A \ B)" using is_free_at_in_left_app by blast ultimately show ?thesis using "3.prems"(1) by presburger next case b with "3.IH"(2) obtain p where "p \ positions B" and "is_free_at v p B" by blast from \p \ positions B\ have "\ # p \ positions (A \ B)" by simp moreover from \is_free_at v p B\ have "is_free_at v (\ # p) (A \ B)" using is_free_at_in_right_app by blast ultimately show ?thesis using "3.prems"(1) by presburger qed next case (4 x \ A) from \v \ free_vars (\x\<^bsub>\\<^esub>. A)\ have "v \ free_vars A - {(x, \)}" and "v \ (x, \)" by simp_all then have "v \ free_vars A" by blast with "4.IH" obtain p where "p \ positions A" and "is_free_at v p A" by blast from \p \ positions A\ have "\ # p \ positions (\x\<^bsub>\\<^esub>. A)" by simp moreover from \is_free_at v p A\ and \v \ (x, \)\ have "is_free_at v (\ # p) (\x\<^bsub>\\<^esub>. A)" using is_free_at_to_abs by blast ultimately show ?case using "4.prems"(1) by presburger qed simp_all lemma free_vars_alt_def: shows "free_vars A = {v | v p. p \ positions A \ is_free_at v p A}" using free_vars_in_is_free_at and is_free_at_in_free_vars by (intro subset_antisym subsetI CollectI, metis) blast text \ In the following definition, note that the variable immeditately preceded by \\\ counts as a bound variable: \ definition is_bound :: "var \ form \ bool" where [iff]: "is_bound v B \ (\p \ positions B. is_bound_at v p B \ v \ binders_at B p)" lemma is_bound_in_app_homomorphism: shows "is_bound v (A \ B) \ is_bound v A \ is_bound v B" proof assume "is_bound v (A \ B)" then obtain p where "p \ positions (A \ B)" and "is_bound_at v p (A \ B) \ v \ binders_at (A \ B) p" by auto then have "p \ []" by fastforce with \p \ positions (A \ B)\ obtain p' and d where "p = d # p'" by auto from \is_bound_at v p (A \ B) \ v \ binders_at (A \ B) p\ consider (a) "is_bound_at v p (A \ B)" | (b) "v \ binders_at (A \ B) p" by blast then show "is_bound v A \ is_bound v B" proof cases case a then show ?thesis proof (cases d) case Left then have "p' \ positions A" using \p = d # p'\ and \p \ positions (A \ B)\ by fastforce moreover from \is_bound_at v p (A \ B)\ have "occurs_at v p' A" using Left and \p = d # p'\ and is_subform_at.simps(2) by force moreover from \is_bound_at v p (A \ B)\ have "in_scope_of_abs v p' A" using Left and \p = d # p'\ by fastforce ultimately show ?thesis by auto next case Right then have "p' \ positions B" using \p = d # p'\ and \p \ positions (A \ B)\ by fastforce moreover from \is_bound_at v p (A \ B)\ have "occurs_at v p' B" using Right and \p = d # p'\ and is_subform_at.simps(3) by force moreover from \is_bound_at v p (A \ B)\ have "in_scope_of_abs v p' B" using Right and \p = d # p'\ by fastforce ultimately show ?thesis by auto qed next case b then show ?thesis proof (cases d) case Left then have "p' \ positions A" using \p = d # p'\ and \p \ positions (A \ B)\ by fastforce moreover from \v \ binders_at (A \ B) p\ have "v \ binders_at A p'" using Left and \p = d # p'\ by force ultimately show ?thesis by auto next case Right then have "p' \ positions B" using \p = d # p'\ and \p \ positions (A \ B)\ by fastforce moreover from \v \ binders_at (A \ B) p\ have "v \ binders_at B p'" using Right and \p = d # p'\ by force ultimately show ?thesis by auto qed qed next assume "is_bound v A \ is_bound v B" then show "is_bound v (A \ B)" proof (rule disjE) assume "is_bound v A" then obtain p where "p \ positions A" and "is_bound_at v p A \ v \ binders_at A p" by auto from \p \ positions A\ have "\ # p \ positions (A \ B)" by auto from \is_bound_at v p A \ v \ binders_at A p\ consider (a) "is_bound_at v p A" | (b) "v \ binders_at A p" by blast then show "is_bound v (A \ B)" proof cases case a then have "occurs_at v (\ # p) (A \ B)" by auto moreover from a have "is_bound_at v (\ # p) (A \ B)" by force ultimately show "is_bound v (A \ B)" using \\ # p \ positions (A \ B)\ by blast next case b then have "v \ binders_at (A \ B) (\ # p)" by auto then show "is_bound v (A \ B)" using \\ # p \ positions (A \ B)\ by blast qed next assume "is_bound v B" then obtain p where "p \ positions B" and "is_bound_at v p B \ v \ binders_at B p" by auto from \p \ positions B\ have "\ # p \ positions (A \ B)" by auto from \is_bound_at v p B \ v \ binders_at B p\ consider (a) "is_bound_at v p B" | (b) "v \ binders_at B p" by blast then show "is_bound v (A \ B)" proof cases case a then have "occurs_at v (\ # p) (A \ B)" by auto moreover from a have "is_bound_at v (\ # p) (A \ B)" by force ultimately show "is_bound v (A \ B)" using \\ # p \ positions (A \ B)\ by blast next case b then have "v \ binders_at (A \ B) (\ # p)" by auto then show "is_bound v (A \ B)" using \\ # p \ positions (A \ B)\ by blast qed qed qed lemma is_bound_in_abs_body: assumes "is_bound v A" shows "is_bound v (\x\<^bsub>\\<^esub>. A)" using assms unfolding is_bound_def proof fix p assume "p \ positions A" and "is_bound_at v p A \ v \ binders_at A p" moreover from \p \ positions A\ have "\ # p \ positions (\x\<^bsub>\\<^esub>. A)" by simp ultimately consider (a) "is_bound_at v p A" | (b) "v \ binders_at A p" by blast then show "\p \ positions (\x\<^bsub>\\<^esub>. A). is_bound_at v p (\x\<^bsub>\\<^esub>. A) \ v \ binders_at (\x\<^bsub>\\<^esub>. A) p" proof cases case a then have "is_bound_at v (\ # p) (\x\<^bsub>\\<^esub>. A)" by force with \\ # p \ positions (\x\<^bsub>\\<^esub>. A)\ show ?thesis by blast next case b then have "v \ binders_at (\x\<^bsub>\\<^esub>. A) (\ # p)" by simp with \\ # p \ positions (\x\<^bsub>\\<^esub>. A)\ show ?thesis by blast qed qed lemma absent_var_is_not_bound: assumes "v \ vars A" shows "\ is_bound v A" using assms and binders_at_alt_def and in_scope_of_abs_in_vars by blast lemma bound_vars_alt_def2: shows "bound_vars A = {v \ vars A. is_bound v A}" unfolding bound_vars_alt_def using absent_var_is_not_bound by fastforce definition is_free :: "var \ form \ bool" where [iff]: "is_free v B \ (\p \ positions B. is_free_at v p B)" subsection \Free variables for a formula in another formula\ definition is_free_for :: "form \ var \ form \ bool" where [iff]: "is_free_for A v B \ ( \v' \ free_vars A. \p \ positions B. is_free_at v p B \ \ in_scope_of_abs v' p B )" lemma is_free_for_absent_var [intro]: assumes "v \ vars B" shows "is_free_for A v B" using assms and occurs_def and is_free_at_def and occurs_in_vars by blast lemma is_free_for_in_var [intro]: shows "is_free_for A v (x\<^bsub>\\<^esub>)" using subforms_from_var(2) by force lemma is_free_for_in_con [intro]: shows "is_free_for A v (\c\\<^bsub>\\<^esub>)" using subforms_from_con(2) by force lemma is_free_for_from_app: assumes "is_free_for A v (B \ C)" shows "is_free_for A v B" and "is_free_for A v C" proof - { fix v' assume "v' \ free_vars A" then have "\p \ positions B. is_free_at v p B \ \ in_scope_of_abs v' p B" proof (intro ballI impI) fix p assume "v' \ free_vars A" and "p \ positions B" and "is_free_at v p B" from \p \ positions B\ have "\ # p \ positions (B \ C)" by simp moreover from \is_free_at v p B\ have "is_free_at v (\ # p) (B \ C)" using is_free_at_in_left_app by blast ultimately have "\ in_scope_of_abs v' (\ # p) (B \ C)" using assms and \v' \ free_vars A\ by blast then show "\ in_scope_of_abs v' p B" by simp qed } then show "is_free_for A v B" by force next { fix v' assume "v' \ free_vars A" then have "\p \ positions C. is_free_at v p C \ \ in_scope_of_abs v' p C" proof (intro ballI impI) fix p assume "v' \ free_vars A" and "p \ positions C" and "is_free_at v p C" from \p \ positions C\ have "\ # p \ positions (B \ C)" by simp moreover from \is_free_at v p C\ have "is_free_at v (\ # p) (B \ C)" using is_free_at_in_right_app by blast ultimately have "\ in_scope_of_abs v' (\ # p) (B \ C)" using assms and \v' \ free_vars A\ by blast then show "\ in_scope_of_abs v' p C" by simp qed } then show "is_free_for A v C" by force qed lemma is_free_for_to_app [intro]: assumes "is_free_for A v B" and "is_free_for A v C" shows "is_free_for A v (B \ C)" unfolding is_free_for_def proof (intro ballI impI) fix v' and p assume "v' \ free_vars A" and "p \ positions (B \ C)" and "is_free_at v p (B \ C)" from \is_free_at v p (B \ C)\ have "p \ []" using occurs_at_alt_def(8) by force then obtain d and p' where "p = d # p'" by (meson list.exhaust) with \p \ positions (B \ C)\ consider (b) "p = \ # p'" and "p' \ positions B" | (c) "p = \ # p'" and "p' \ positions C" by force then show "\ in_scope_of_abs v' p (B \ C)" proof cases case b from b(1) and \is_free_at v p (B \ C)\ have "is_free_at v p' B" using is_free_at_in_left_app by blast with assms(1) and \v' \ free_vars A\ and \p' \ positions B\ have "\ in_scope_of_abs v' p' B" by simp with b(1) show ?thesis using in_scope_of_abs_in_left_app by simp next case c from c(1) and \is_free_at v p (B \ C)\ have "is_free_at v p' C" using is_free_at_in_right_app by blast with assms(2) and \v' \ free_vars A\ and \p' \ positions C\ have "\ in_scope_of_abs v' p' C" by simp with c(1) show ?thesis using in_scope_of_abs_in_right_app by simp qed qed lemma is_free_for_in_app: shows "is_free_for A v (B \ C) \ is_free_for A v B \ is_free_for A v C" using is_free_for_from_app and is_free_for_to_app by iprover lemma is_free_for_to_abs [intro]: assumes "is_free_for A v B" and "(x, \) \ free_vars A" shows "is_free_for A v (\x\<^bsub>\\<^esub>. B)" unfolding is_free_for_def proof (intro ballI impI) fix v' and p assume "v' \ free_vars A" and "p \ positions (\x\<^bsub>\\<^esub>. B)" and "is_free_at v p (\x\<^bsub>\\<^esub>. B)" from \is_free_at v p (\x\<^bsub>\\<^esub>. B)\ have "p \ []" using occurs_at_alt_def(9) by force with \p \ positions (\x\<^bsub>\\<^esub>. B)\ obtain p' where "p = \ # p'" and "p' \ positions B" by force then show "\ in_scope_of_abs v' p (\x\<^bsub>\\<^esub>. B)" proof - from \p = \ # p'\ and \is_free_at v p (\x\<^bsub>\\<^esub>. B)\ have "is_free_at v p' B" using is_free_at_from_abs by blast with assms(1) and \v' \ free_vars A\ and \p' \ positions B\ have "\ in_scope_of_abs v' p' B" by force moreover from \v' \ free_vars A\ and assms(2) have "v' \ (x, \)" by blast ultimately show ?thesis using \p = \ # p'\ and in_scope_of_abs_in_abs by auto qed qed lemma is_free_for_from_abs: assumes "is_free_for A v (\x\<^bsub>\\<^esub>. B)" and "v \ (x, \)" shows "is_free_for A v B" unfolding is_free_for_def proof (intro ballI impI) fix v' and p assume "v' \ free_vars A" and "p \ positions B" and "is_free_at v p B" then show "\ in_scope_of_abs v' p B" proof - from \is_free_at v p B\ and assms(2) have "is_free_at v (\ # p) (\x\<^bsub>\\<^esub>. B)" by (rule is_free_at_to_abs) moreover from \p \ positions B\ have "\ # p \ positions (\x\<^bsub>\\<^esub>. B)" by simp ultimately have "\ in_scope_of_abs v' (\ # p) (\x\<^bsub>\\<^esub>. B)" using assms and \v' \ free_vars A\ by blast then show ?thesis using in_scope_of_abs_in_abs by blast qed qed lemma closed_is_free_for [intro]: assumes "free_vars A = {}" shows "is_free_for A v B" using assms by force lemma is_free_for_closed_form [intro]: assumes "free_vars B = {}" shows "is_free_for A v B" using assms and is_free_at_in_free_vars by blast lemma is_free_for_alt_def: shows " is_free_for A v B \ ( \p. ( p \ positions B \ is_free_at v p B \ p \ [] \ (\v' \ free_vars A. \p' C. strict_prefix p' p \ FAbs v' C \\<^bsub>p'\<^esub> B) ) )" unfolding is_free_for_def using in_scope_of_abs_alt_def and is_subform_implies_in_positions by meson lemma binding_var_not_free_for_in_abs: assumes "is_free x B" and "x \ w" shows "\ is_free_for (FVar w) x (FAbs w B)" proof (rule ccontr) assume "\ \ is_free_for (FVar w) x (FAbs w B)" then have " \v' \ free_vars (FVar w). \p \ positions (FAbs w B). is_free_at x p (FAbs w B) \ \ in_scope_of_abs v' p (FAbs w B)" by force moreover have "free_vars (FVar w) = {w}" using surj_pair[of w] by force ultimately have " \p \ positions (FAbs w B). is_free_at x p (FAbs w B) \ \ in_scope_of_abs w p (FAbs w B)" by blast moreover from assms(1) obtain p where "is_free_at x p B" by fastforce from this and assms(2) have "is_free_at x (\ # p) (FAbs w B)" by (rule is_free_at_to_abs) moreover from this have "\ # p \ positions (FAbs w B)" using is_subform_implies_in_positions by force ultimately have "\ in_scope_of_abs w (\ # p) (FAbs w B)" by blast moreover have "in_scope_of_abs w (\ # p) (FAbs w B)" using in_scope_of_abs_in_abs by blast ultimately show False by contradiction qed lemma absent_var_is_free_for [intro]: assumes "x \ vars A" shows "is_free_for (FVar x) y A" using in_scope_of_abs_in_vars and assms and surj_pair[of x] by fastforce lemma form_is_free_for_absent_var [intro]: assumes "x \ vars A" shows "is_free_for B x A" using assms and occurs_in_vars by fastforce lemma form_with_free_binder_not_free_for: assumes "v \ v'" and "v' \ free_vars A" and "v \ free_vars B" shows "\ is_free_for A v (FAbs v' B)" proof - from assms(3) obtain p where "p \ positions B" and "is_free_at v p B" using free_vars_in_is_free_at by blast then have "\ # p \ positions (FAbs v' B)" and "is_free_at v (\ # p) (FAbs v' B)" using surj_pair[of v'] and is_free_at_to_abs[OF \is_free_at v p B\ assms(1)] by force+ moreover have "in_scope_of_abs v' (\ # p) (FAbs v' B)" using in_scope_of_abs_in_abs by blast ultimately show ?thesis using assms(2) by blast qed subsection \Replacement of subformulas\ inductive is_replacement_at :: "form \ position \ form \ form \ bool" ("(4_\_ \ _\ \ _)" [1000, 0, 0, 0] 900) where pos_found: "A\p \ C\ \ C'" if "p = []" and "C = C'" | replace_left_app: "(G \ H)\\ # p \ C\ \ (G' \ H)" if "p \ positions G" and "G\p \ C\ \ G'" | replace_right_app: "(G \ H)\\ # p \ C\ \ (G \ H')" if "p \ positions H" and "H\p \ C\ \ H'" | replace_abs: "(\x\<^bsub>\\<^esub>. E)\\ # p \ C\ \ (\x\<^bsub>\\<^esub>. E')" if "p \ positions E" and "E\p \ C\ \ E'" lemma is_replacement_at_implies_in_positions: assumes "C\p \ A\ \ D" shows "p \ positions C" using assms by (induction rule: is_replacement_at.induct) auto declare is_replacement_at.intros [intro!] lemma is_replacement_at_existence: assumes "p \ positions C" obtains D where "C\p \ A\ \ D" using assms proof (induction C arbitrary: p thesis) case (FApp C\<^sub>1 C\<^sub>2) from FApp.prems(2) consider (a) "p = []" | (b) "\p'. p = \ # p' \ p' \ positions C\<^sub>1" | (c) "\p'. p = \ # p' \ p' \ positions C\<^sub>2" by fastforce then show ?case proof cases case a with FApp.prems(1) show ?thesis by blast next case b with FApp.prems(1) show ?thesis using FApp.IH(1) and replace_left_app by meson next case c with FApp.prems(1) show ?thesis using FApp.IH(2) and replace_right_app by meson qed next case (FAbs v C') from FAbs.prems(2) consider (a) "p = []" | (b) "\p'. p = \ # p' \ p' \ positions C'" using surj_pair[of v] by fastforce then show ?case proof cases case a with FAbs.prems(1) show ?thesis by blast next case b with FAbs.prems(1,2) show ?thesis using FAbs.IH and surj_pair[of v] by blast qed qed force+ lemma is_replacement_at_minimal_change: assumes "C\p \ A\ \ D" shows "A \\<^bsub>p\<^esub> D" and "\p' \ positions D. \ prefix p' p \ \ prefix p p' \ subform_at D p' = subform_at C p'" using assms by (induction rule: is_replacement_at.induct) auto lemma is_replacement_at_binders: assumes "C\p \ A\ \ D" shows "binders_at D p = binders_at C p" using assms by (induction rule: is_replacement_at.induct) simp_all lemma is_replacement_at_occurs: assumes "C\p \ A\ \ D" and "\ prefix p' p" and "\ prefix p p'" shows "occurs_at v p' C \ occurs_at v p' D" using assms proof (induction arbitrary: p' rule: is_replacement_at.induct) case pos_found then show ?case by simp next case replace_left_app then show ?case proof (cases p') case (Cons d p'') with replace_left_app.prems(1,2) show ?thesis by (cases d) (use replace_left_app.IH in force)+ qed force next case replace_right_app then show ?case proof (cases p') case (Cons d p'') with replace_right_app.prems(1,2) show ?thesis by (cases d) (use replace_right_app.IH in force)+ qed force next case replace_abs then show ?case proof (cases p') case (Cons d p'') with replace_abs.prems(1,2) show ?thesis by (cases d) (use replace_abs.IH in force)+ qed force qed lemma fresh_var_replacement_position_uniqueness: assumes "v \ vars C" and "C\p \ FVar v\ \ G" and "occurs_at v p' G" shows "p' = p" proof (rule ccontr) assume "p' \ p" from assms(2) have "occurs_at v p G" by (simp add: is_replacement_at_minimal_change(1)) moreover have *: "occurs_at v p' C \ occurs_at v p' G" if "\ prefix p' p" and "\ prefix p p'" using assms(2) and that and is_replacement_at_occurs by blast ultimately show False proof (cases "\ prefix p' p \ \ prefix p p'") case True with assms(3) and * have "occurs_at v p' C" by simp then have "v \ vars C" using is_subform_implies_in_positions and occurs_in_vars by fastforce with assms(1) show ?thesis by contradiction next case False have "FVar v \\<^bsub>p\<^esub> G" by (fact is_replacement_at_minimal_change(1)[OF assms(2)]) moreover from assms(3) have "FVar v \\<^bsub>p'\<^esub> G" by simp ultimately show ?thesis using \p' \ p\ and False and loop_subform_impossibility by (blast dest: prefix_order.antisym_conv2) qed qed lemma is_replacement_at_new_positions: assumes "C\p \ A\ \ D" and "prefix p p'" and "p' \ positions D" obtains p'' where "p' = p @ p''" and "p'' \ positions A" using assms by (induction arbitrary: thesis p' rule: is_replacement_at.induct, auto) blast+ lemma replacement_override: assumes "C\p \ B\ \ D" and "C\p \ A\ \ F" shows "D\p \ A\ \ F" using assms proof (induction arbitrary: F rule: is_replacement_at.induct) case pos_found from pos_found.hyps(1) and pos_found.prems have "A = F" using is_replacement_at.simps by blast with pos_found.hyps(1) show ?case by blast next case (replace_left_app p G C G' H) have "p \ positions G'" by ( fact is_subform_implies_in_positions [OF is_replacement_at_minimal_change(1)[OF replace_left_app.hyps(2)]] ) from replace_left_app.prems obtain F' where "F = F' \ H" and "G\p \ A\ \ F'" by (fastforce elim: is_replacement_at.cases) from \G\p \ A\ \ F'\ have "G'\p \ A\ \ F'" by (fact replace_left_app.IH) with \p \ positions G'\ show ?case unfolding \F = F' \ H\ by blast next case (replace_right_app p H C H' G) have "p \ positions H'" by ( fact is_subform_implies_in_positions [OF is_replacement_at_minimal_change(1)[OF replace_right_app.hyps(2)]] ) from replace_right_app.prems obtain F' where "F = G \ F'" and "H\p \ A\ \ F'" by (fastforce elim: is_replacement_at.cases) from \H\p \ A\ \ F'\ have "H'\p \ A\ \ F'" by (fact replace_right_app.IH) with \p \ positions H'\ show ?case unfolding \F = G \ F'\ by blast next case (replace_abs p E C E' x \) have "p \ positions E'" by ( fact is_subform_implies_in_positions [OF is_replacement_at_minimal_change(1)[OF replace_abs.hyps(2)]] ) from replace_abs.prems obtain F' where "F = \x\<^bsub>\\<^esub>. F'" and "E\p \ A\ \ F'" by (fastforce elim: is_replacement_at.cases) from \E\p \ A\ \ F'\ have "E'\p \ A\ \ F'" by (fact replace_abs.IH) with \p \ positions E'\ show ?case unfolding \F = \x\<^bsub>\\<^esub>. F'\ by blast qed lemma leftmost_subform_in_generalized_app_replacement: shows "(\\<^sup>\\<^sub>\ C As)\replicate (length As) \ \ D\ \ (\\<^sup>\\<^sub>\ D As)" using is_replacement_at_implies_in_positions and replace_left_app by (induction As arbitrary: D rule: rev_induct) auto subsection \Logical constants\ abbreviation (input) \ where "\ \ 0" abbreviation (input) \ where "\ \ Suc \" abbreviation (input) \ where "\ \ Suc \" abbreviation (input) \ where "\ \ Suc \" abbreviation (input) \ where "\ \ Suc \" abbreviation (input) \ where "\ \ Suc \" abbreviation (input) \ where "\ \ Suc \" abbreviation (input) \\<^sub>Q where "\\<^sub>Q \ Suc \" abbreviation (input) \\<^sub>\ where "\\<^sub>\ \ Suc \\<^sub>Q" definition Q_constant_of_type :: "type \ con" where [simp]: "Q_constant_of_type \ = (\\<^sub>Q, \\\\o)" definition iota_constant :: con where [simp]: "iota_constant \ (\\<^sub>\, (i\o)\i)" definition Q :: "type \ form" ("Q\<^bsub>_\<^esub>") where [simp]: "Q\<^bsub>\\<^esub> = FCon (Q_constant_of_type \)" definition iota :: form ("\") where [simp]: "\ = FCon iota_constant" definition is_Q_constant_of_type :: "con \ type \ bool" where [iff]: "is_Q_constant_of_type p \ \ p = Q_constant_of_type \" definition is_iota_constant :: "con \ bool" where [iff]: "is_iota_constant p \ p = iota_constant" definition is_logical_constant :: "con \ bool" where [iff]: "is_logical_constant p \ (\\. is_Q_constant_of_type p \) \ is_iota_constant p" definition type_of_Q_constant :: "con \ type" where [simp]: "type_of_Q_constant p = (THE \. is_Q_constant_of_type p \)" lemma constant_cases[case_names non_logical Q_constant \_constant, cases type: con]: assumes "\ is_logical_constant p \ P" and "\\. is_Q_constant_of_type p \ \ P" and "is_iota_constant p \ P" shows "P" using assms by blast subsection \Definitions and abbreviations\ definition equality_of_type :: "form \ type \ form \ form" ("(_ =\<^bsub>_\<^esub>/ _)" [103, 0, 103] 102) where [simp]: "A =\<^bsub>\\<^esub> B = Q\<^bsub>\\<^esub> \ A \ B" definition equivalence :: "form \ form \ form" (infixl "\\<^sup>\" 102) where - [simp]: "A \\<^sup>\ B = A =\<^bsub>o\<^esub> B" \ \more modular than the definition in @{cite "andrews:2002"}\ + [simp]: "A \\<^sup>\ B = A =\<^bsub>o\<^esub> B" \ \more modular than the definition in \<^cite>\"andrews:2002"\\ definition true :: form ("T\<^bsub>o\<^esub>") where [simp]: "T\<^bsub>o\<^esub> = Q\<^bsub>o\<^esub> =\<^bsub>o\o\o\<^esub> Q\<^bsub>o\<^esub>" definition false :: form ("F\<^bsub>o\<^esub>") where [simp]: "F\<^bsub>o\<^esub> = \\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub> =\<^bsub>o\o\<^esub> \\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>" definition PI :: "type \ form" ("\\<^bsub>_\<^esub>") where [simp]: "\\<^bsub>\\<^esub> = Q\<^bsub>\\o\<^esub> \ (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>)" definition forall :: "nat \ type \ form \ form" ("(4\_\<^bsub>_\<^esub>./ _)" [0, 0, 141] 141) where [simp]: "\x\<^bsub>\\<^esub>. A = \\<^bsub>\\<^esub> \ (\x\<^bsub>\\<^esub>. A)" text \Generalized universal quantification. We define \\\<^sup>\\<^sub>\ [x\<^sub>1, \, x\<^sub>n] A\ as \\x\<^sub>1. \ \x\<^sub>n. A\:\ definition generalized_forall :: "var list \ form \ form" ("\\<^sup>\\<^sub>\ _ _" [141, 141] 141) where [simp]: "\\<^sup>\\<^sub>\ vs A = foldr (\(x, \) B. \x\<^bsub>\\<^esub>. B) vs A" lemma innermost_subform_in_generalized_forall: assumes "vs \ []" shows "A \\<^bsub>foldr (\_ p. [\,\] @ p) vs []\<^esub> \\<^sup>\\<^sub>\ vs A" using assms by (induction vs) fastforce+ lemma innermost_replacement_in_generalized_forall: assumes "vs \ []" shows "(\\<^sup>\\<^sub>\ vs C)\foldr (\_. (@) [\,\]) vs [] \ B\ \ (\\<^sup>\\<^sub>\ vs B)" using assms proof (induction vs) case Nil then show ?case by blast next case (Cons v vs) obtain x and \ where "v = (x, \)" by fastforce then show ?case proof (cases "vs = []") case True with \v = (x, \)\ show ?thesis unfolding True by force next case False then have "foldr (\_. (@) [\,\]) vs [] \ positions (\\<^sup>\\<^sub>\ vs C)" using innermost_subform_in_generalized_forall and is_subform_implies_in_positions by blast moreover from False have "(\\<^sup>\\<^sub>\ vs C)\foldr (\_. (@) [\,\]) vs [] \ B\ \ (\\<^sup>\\<^sub>\ vs B)" by (fact Cons.IH) ultimately have "(\x\<^bsub>\\<^esub>. \\<^sup>\\<^sub>\ vs C)\\ # foldr (\_. (@) [\,\]) vs [] \ B\ \ (\x\<^bsub>\\<^esub>. \\<^sup>\\<^sub>\ vs B)" by (rule replace_abs) moreover have "\ # foldr (\_. (@) [\,\]) vs [] \ positions (\x\<^bsub>\\<^esub>. \\<^sup>\\<^sub>\ vs C)" using \foldr (\_. (@) [\,\]) vs [] \ positions (\\<^sup>\\<^sub>\ vs C)\ by simp ultimately have " (\\<^bsub>\\<^esub> \ (\x\<^bsub>\\<^esub>. \\<^sup>\\<^sub>\ vs C))\\ # \ # foldr (\_. (@) [\,\]) vs [] \ B\ \ (\\<^bsub>\\<^esub> \ (\x\<^bsub>\\<^esub>. \\<^sup>\\<^sub>\ vs B))" by blast then have "(\x\<^bsub>\\<^esub>. \\<^sup>\\<^sub>\ vs C)\[\,\] @ foldr (\_. (@) [\,\]) vs [] \ B\ \ (\x\<^bsub>\\<^esub>. \\<^sup>\\<^sub>\ vs B)" by simp then show ?thesis unfolding \v = (x, \)\ and generalized_forall_def and foldr.simps(2) and o_apply and case_prod_conv . qed qed lemma false_is_forall: shows "F\<^bsub>o\<^esub> = \\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>" unfolding false_def and forall_def and PI_def and equality_of_type_def .. definition conj_fun :: form ("\\<^bsub>o\o\o\<^esub>") where [simp]: "\\<^bsub>o\o\o\<^esub> = \\\<^bsub>o\<^esub>. \\\<^bsub>o\<^esub>. ( (\\\<^bsub>o\o\o\<^esub>. \\<^bsub>o\o\o\<^esub> \ T\<^bsub>o\<^esub> \ T\<^bsub>o\<^esub>) =\<^bsub>(o\o\o)\o\<^esub> (\\\<^bsub>o\o\o\<^esub>. \\<^bsub>o\o\o\<^esub> \ \\<^bsub>o\<^esub> \ \\<^bsub>o\<^esub>) )" definition conj_op :: "form \ form \ form" (infixl "\\<^sup>\" 131) where [simp]: "A \\<^sup>\ B = \\<^bsub>o\o\o\<^esub> \ A \ B" text \Generalized conjunction. We define \\\<^sup>\\<^sub>\ [A\<^sub>1, \, A\<^sub>n]\ as \A\<^sub>1 \\<^sup>\ (\ \\<^sup>\ (A\<^sub>n\<^sub>-\<^sub>1 \\<^sup>\ A\<^sub>n) \)\:\ definition generalized_conj_op :: "form list \ form" ("\\<^sup>\\<^sub>\ _" [0] 131) where [simp]: "\\<^sup>\\<^sub>\ As = foldr1 (\\<^sup>\) As" -definition imp_fun :: form ("\\<^bsub>o\o\o\<^esub>") where \ \\\\ used instead of \=\, see @{cite "andrews:2002"}\ +definition imp_fun :: form ("\\<^bsub>o\o\o\<^esub>") where \ \\\\ used instead of \=\, see \<^cite>\"andrews:2002"\\ [simp]: "\\<^bsub>o\o\o\<^esub> = \\\<^bsub>o\<^esub>. \\\<^bsub>o\<^esub>. (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)" definition imp_op :: "form \ form \ form" (infixl "\\<^sup>\" 111) where [simp]: "A \\<^sup>\ B = \\<^bsub>o\o\o\<^esub> \ A \ B" text \Generalized implication. We define \[A\<^sub>1, \, A\<^sub>n] \\<^sup>\\<^sub>\ B\ as \A\<^sub>1 \\<^sup>\ (\ \\<^sup>\ (A\<^sub>n \\<^sup>\ B) \)\:\ definition generalized_imp_op :: "form list \ form \ form" (infixl "\\<^sup>\\<^sub>\" 111) where [simp]: "As \\<^sup>\\<^sub>\ B = foldr (\\<^sup>\) As B" text \ Given the definition below, it is interesting to note that \\\<^sup>\ A\ and \F\<^bsub>o\<^esub> \\<^sup>\ A\ are exactly the same formula, namely \Q\<^bsub>o\<^esub> \ F\<^bsub>o\<^esub> \ A\: \ definition neg :: "form \ form" ("\\<^sup>\ _" [141] 141) where [simp]: "\\<^sup>\ A = Q\<^bsub>o\<^esub> \ F\<^bsub>o\<^esub> \ A" definition disj_fun :: form ("\\<^bsub>o\o\o\<^esub>") where [simp]: "\\<^bsub>o\o\o\<^esub> = \\\<^bsub>o\<^esub>. \\\<^bsub>o\<^esub>. \\<^sup>\ (\\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^sup>\ \\<^bsub>o\<^esub>)" definition disj_op :: "form \ form \ form" (infixl "\\<^sup>\" 126) where [simp]: "A \\<^sup>\ B = \\<^bsub>o\o\o\<^esub> \ A \ B" definition exists :: "nat \ type \ form \ form" ("(4\_\<^bsub>_\<^esub>./ _)" [0, 0, 141] 141) where [simp]: "\x\<^bsub>\\<^esub>. A = \\<^sup>\ (\x\<^bsub>\\<^esub>. \\<^sup>\ A)" lemma exists_fv: shows "free_vars (\x\<^bsub>\\<^esub>. A) = free_vars A - {(x, \)}" by simp definition inequality_of_type :: "form \ type \ form \ form" ("(_ \\<^bsub>_\<^esub>/ _)" [103, 0, 103] 102) where [simp]: "A \\<^bsub>\\<^esub> B = \\<^sup>\ (A =\<^bsub>\\<^esub> B)" subsection \Well-formed formulas\ inductive is_wff_of_type :: "type \ form \ bool" where var_is_wff: "is_wff_of_type \ (x\<^bsub>\\<^esub>)" | con_is_wff: "is_wff_of_type \ (\c\\<^bsub>\\<^esub>)" | app_is_wff: "is_wff_of_type \ (A \ B)" if "is_wff_of_type (\\\) A" and "is_wff_of_type \ B" | abs_is_wff: "is_wff_of_type (\\\) (\x\<^bsub>\\<^esub>. A)" if "is_wff_of_type \ A" definition wffs_of_type :: "type \ form set" ("wffs\<^bsub>_\<^esub>" [0]) where "wffs\<^bsub>\\<^esub> = {f :: form. is_wff_of_type \ f}" abbreviation wffs :: "form set" where "wffs \ \\. wffs\<^bsub>\\<^esub>" lemma is_wff_of_type_wffs_of_type_eq [pred_set_conv]: shows "is_wff_of_type \ = (\f. f \ wffs\<^bsub>\\<^esub>)" unfolding wffs_of_type_def by simp lemmas wffs_of_type_intros [intro!] = is_wff_of_type.intros[to_set] lemmas wffs_of_type_induct [consumes 1, induct set: wffs_of_type] = is_wff_of_type.induct[to_set] lemmas wffs_of_type_cases [consumes 1, cases set: wffs_of_type] = is_wff_of_type.cases[to_set] lemmas wffs_of_type_simps = is_wff_of_type.simps[to_set] lemma generalized_app_wff [intro]: assumes "length As = length ts" and "\k < length As. As ! k \ wffs\<^bsub>ts ! k\<^esub>" and "B \ wffs\<^bsub>foldr (\) ts \\<^esub>" shows "\\<^sup>\\<^sub>\ B As \ wffs\<^bsub>\\<^esub>" using assms proof (induction As ts arbitrary: B rule: list_induct2) case Nil then show ?case by simp next case (Cons A As t ts) from Cons.prems(1) have "A \ wffs\<^bsub>t\<^esub>" by fastforce moreover from Cons.prems(2) have "B \ wffs\<^bsub>t\foldr (\) ts \\<^esub>" by auto ultimately have "B \ A \ wffs\<^bsub>foldr (\) ts \\<^esub>" by blast moreover have "\k < length As. (A # As) ! (Suc k) = As ! k \ (t # ts) ! (Suc k) = ts ! k" by force with Cons.prems(1) have "\k < length As. As ! k \ wffs\<^bsub>ts ! k\<^esub>" by fastforce ultimately have "\\<^sup>\\<^sub>\ (B \ A) As \ wffs\<^bsub>\\<^esub>" using Cons.IH by (simp only:) moreover have "\\<^sup>\\<^sub>\ B (A # As) = \\<^sup>\\<^sub>\ (B \ A) As" by simp ultimately show ?case by (simp only:) qed lemma generalized_abs_wff [intro]: assumes "B \ wffs\<^bsub>\\<^esub>" shows "\\<^sup>\\<^sub>\ vs B \ wffs\<^bsub>foldr (\) (map snd vs) \\<^esub>" using assms proof (induction vs) case Nil then show ?case by simp next case (Cons v vs) let ?\ = "foldr (\) (map snd vs) \" obtain x and \ where "v = (x, \)" by fastforce then have "FVar v \ wffs\<^bsub>\\<^esub>" by auto from Cons.prems have "\\<^sup>\\<^sub>\ vs B \ wffs\<^bsub>?\\<^esub>" by (fact Cons.IH) with \v = (x, \)\ have "FAbs v (\\<^sup>\\<^sub>\ vs B) \ wffs\<^bsub>\\?\\<^esub>" by blast moreover from \v = (x, \)\ have "foldr (\) (map snd (v # vs)) \ = \\?\" by simp moreover have "\\<^sup>\\<^sub>\ (v # vs) B = FAbs v (\\<^sup>\\<^sub>\ vs B)" by simp ultimately show ?case by (simp only:) qed lemma Q_wff [intro]: shows "Q\<^bsub>\\<^esub> \ wffs\<^bsub>\\\\o\<^esub>" by auto lemma iota_wff [intro]: shows "\ \ wffs\<^bsub>(i\o)\i\<^esub>" by auto lemma equality_wff [intro]: assumes "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" shows "A =\<^bsub>\\<^esub> B \ wffs\<^bsub>o\<^esub>" using assms by auto lemma equivalence_wff [intro]: assumes "A \ wffs\<^bsub>o\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" shows "A \\<^sup>\ B \ wffs\<^bsub>o\<^esub>" using assms unfolding equivalence_def by blast lemma true_wff [intro]: shows "T\<^bsub>o\<^esub> \ wffs\<^bsub>o\<^esub>" by force lemma false_wff [intro]: shows "F\<^bsub>o\<^esub> \ wffs\<^bsub>o\<^esub>" by auto lemma pi_wff [intro]: shows "\\<^bsub>\\<^esub> \ wffs\<^bsub>(\\o)\o\<^esub>" using PI_def by fastforce lemma forall_wff [intro]: assumes "A \ wffs\<^bsub>o\<^esub>" shows "\x\<^bsub>\\<^esub>. A \ wffs\<^bsub>o\<^esub>" using assms and pi_wff unfolding forall_def by blast lemma generalized_forall_wff [intro]: assumes "B \ wffs\<^bsub>o\<^esub>" shows "\\<^sup>\\<^sub>\ vs B \ wffs\<^bsub>o\<^esub>" using assms proof (induction vs) case (Cons v vs) then show ?case using surj_pair[of v] by force qed simp lemma conj_fun_wff [intro]: shows "\\<^bsub>o\o\o\<^esub> \ wffs\<^bsub>o\o\o\<^esub>" by auto lemma conj_op_wff [intro]: assumes "A \ wffs\<^bsub>o\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" shows "A \\<^sup>\ B \ wffs\<^bsub>o\<^esub>" using assms unfolding conj_op_def by blast lemma imp_fun_wff [intro]: shows "\\<^bsub>o\o\o\<^esub> \ wffs\<^bsub>o\o\o\<^esub>" by auto lemma imp_op_wff [intro]: assumes "A \ wffs\<^bsub>o\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" shows "A \\<^sup>\ B \ wffs\<^bsub>o\<^esub>" using assms unfolding imp_op_def by blast lemma neg_wff [intro]: assumes "A \ wffs\<^bsub>o\<^esub>" shows " \\<^sup>\ A \ wffs\<^bsub>o\<^esub>" using assms by fastforce lemma disj_fun_wff [intro]: shows "\\<^bsub>o\o\o\<^esub> \ wffs\<^bsub>o\o\o\<^esub>" by auto lemma disj_op_wff [intro]: assumes "A \ wffs\<^bsub>o\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" shows "A \\<^sup>\ B \ wffs\<^bsub>o\<^esub>" using assms by auto lemma exists_wff [intro]: assumes "A \ wffs\<^bsub>o\<^esub>" shows "\x\<^bsub>\\<^esub>. A \ wffs\<^bsub>o\<^esub>" using assms by fastforce lemma inequality_wff [intro]: assumes "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" shows "A \\<^bsub>\\<^esub> B \ wffs\<^bsub>o\<^esub>" using assms by fastforce lemma wffs_from_app: assumes "A \ B \ wffs\<^bsub>\\<^esub>" obtains \ where "A \ wffs\<^bsub>\\\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" using assms by (blast elim: wffs_of_type_cases) lemma wffs_from_generalized_app: assumes "\\<^sup>\\<^sub>\ B As \ wffs\<^bsub>\\<^esub>" obtains ts where "length ts = length As" and "\k < length As. As ! k \ wffs\<^bsub>ts ! k\<^esub>" and "B \ wffs\<^bsub>foldr (\) ts \\<^esub>" using assms proof (induction As arbitrary: B thesis) case Nil then show ?case by simp next case (Cons A As) from Cons.prems have "\\<^sup>\\<^sub>\ (B \ A) As \ wffs\<^bsub>\\<^esub>" by auto then obtain ts where "length ts = length As" and "\k < length As. As ! k \ wffs\<^bsub>ts ! k\<^esub>" and "B \ A \ wffs\<^bsub>foldr (\) ts \\<^esub>" using Cons.IH by blast moreover from \B \ A \ wffs\<^bsub>foldr (\) ts \\<^esub>\ obtain t where "B \ wffs\<^bsub>t\foldr (\) ts \\<^esub>" and "A \ wffs\<^bsub>t\<^esub>" by (elim wffs_from_app) moreover from \length ts = length As\ have "length (t # ts) = length (A # As)" by simp moreover from \A \ wffs\<^bsub>t\<^esub>\ and \\k < length As. As ! k \ wffs\<^bsub>ts ! k\<^esub>\ have "\k < length (A # As). (A # As) ! k \ wffs\<^bsub>(t # ts) ! k\<^esub>" by (simp add: nth_Cons') moreover from \B \ wffs\<^bsub>t\foldr (\) ts \\<^esub>\ have "B \ wffs\<^bsub>foldr (\) (t # ts) \\<^esub>" by simp ultimately show ?case using Cons.prems(1) by blast qed lemma wffs_from_abs: assumes "\x\<^bsub>\\<^esub>. A \ wffs\<^bsub>\\<^esub>" obtains \ where "\ = \\\" and "A \ wffs\<^bsub>\\<^esub>" using assms by (blast elim: wffs_of_type_cases) lemma wffs_from_equality: assumes "A =\<^bsub>\\<^esub> B \ wffs\<^bsub>o\<^esub>" shows "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" using assms by (fastforce elim: wffs_of_type_cases)+ lemma wffs_from_equivalence: assumes "A \\<^sup>\ B \ wffs\<^bsub>o\<^esub>" shows "A \ wffs\<^bsub>o\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" using assms unfolding equivalence_def by (fact wffs_from_equality)+ lemma wffs_from_forall: assumes "\x\<^bsub>\\<^esub>. A \ wffs\<^bsub>o\<^esub>" shows "A \ wffs\<^bsub>o\<^esub>" using assms unfolding forall_def and PI_def by (fold equality_of_type_def) (drule wffs_from_equality, blast elim: wffs_from_abs) lemma wffs_from_conj_fun: assumes "\\<^bsub>o\o\o\<^esub> \ A \ B \ wffs\<^bsub>o\<^esub>" shows "A \ wffs\<^bsub>o\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" using assms by (auto elim: wffs_from_app wffs_from_abs) lemma wffs_from_conj_op: assumes "A \\<^sup>\ B \ wffs\<^bsub>o\<^esub>" shows "A \ wffs\<^bsub>o\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" using assms unfolding conj_op_def by (elim wffs_from_conj_fun)+ lemma wffs_from_imp_fun: assumes "\\<^bsub>o\o\o\<^esub> \ A \ B \ wffs\<^bsub>o\<^esub>" shows "A \ wffs\<^bsub>o\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" using assms by (auto elim: wffs_from_app wffs_from_abs) lemma wffs_from_imp_op: assumes "A \\<^sup>\ B \ wffs\<^bsub>o\<^esub>" shows "A \ wffs\<^bsub>o\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" using assms unfolding imp_op_def by (elim wffs_from_imp_fun)+ lemma wffs_from_neg: assumes "\\<^sup>\ A \ wffs\<^bsub>o\<^esub>" shows "A \ wffs\<^bsub>o\<^esub>" using assms unfolding neg_def by (fold equality_of_type_def) (drule wffs_from_equality, blast) lemma wffs_from_disj_fun: assumes "\\<^bsub>o\o\o\<^esub> \ A \ B \ wffs\<^bsub>o\<^esub>" shows "A \ wffs\<^bsub>o\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" using assms by (auto elim: wffs_from_app wffs_from_abs) lemma wffs_from_disj_op: assumes "A \\<^sup>\ B \ wffs\<^bsub>o\<^esub>" shows "A \ wffs\<^bsub>o\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" using assms and wffs_from_disj_fun unfolding disj_op_def by blast+ lemma wffs_from_exists: assumes "\x\<^bsub>\\<^esub>. A \ wffs\<^bsub>o\<^esub>" shows "A \ wffs\<^bsub>o\<^esub>" using assms unfolding exists_def using wffs_from_neg and wffs_from_forall by blast lemma wffs_from_inequality: assumes "A \\<^bsub>\\<^esub> B \ wffs\<^bsub>o\<^esub>" shows "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" using assms unfolding inequality_of_type_def using wffs_from_equality and wffs_from_neg by meson+ lemma wff_has_unique_type: assumes "A \ wffs\<^bsub>\\<^esub>" and "A \ wffs\<^bsub>\\<^esub>" shows "\ = \" using assms proof (induction arbitrary: \ \ rule: form.induct) case (FVar v) obtain x and \ where "v = (x, \)" by fastforce with FVar.prems have "\ = \" and "\ = \" by (blast elim: wffs_of_type_cases)+ then show ?case .. next case (FCon k) obtain x and \ where "k = (x, \)" by fastforce with FCon.prems have "\ = \" and "\ = \" by (blast elim: wffs_of_type_cases)+ then show ?case .. next case (FApp A B) from FApp.prems obtain \' and \' where "A \ wffs\<^bsub>\'\\\<^esub>" and "A \ wffs\<^bsub>\'\\\<^esub>" by (blast elim: wffs_from_app) with FApp.IH(1) show ?case by blast next case (FAbs v A) obtain x and \ where "v = (x, \)" by fastforce with FAbs.prems obtain \' and \' where "\ = \\\'" and "\ = \\\'" and "A \ wffs\<^bsub>\'\<^esub>" and "A \ wffs\<^bsub>\'\<^esub>" by (blast elim: wffs_from_abs) with FAbs.IH show ?case by simp qed lemma wffs_of_type_o_induct [consumes 1, case_names Var Con App]: assumes "A \ wffs\<^bsub>o\<^esub>" and "\x. \

(x\<^bsub>o\<^esub>)" and "\c. \

(\c\\<^bsub>o\<^esub>)" and "\A B \. A \ wffs\<^bsub>\\o\<^esub> \ B \ wffs\<^bsub>\\<^esub> \ \

(A \ B)" shows "\

A" using assms by (cases rule: wffs_of_type_cases) simp_all lemma diff_types_implies_diff_wffs: assumes "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" and "\ \ \" shows "A \ B" using assms and wff_has_unique_type by blast lemma is_free_for_in_generalized_app [intro]: assumes "is_free_for A v B" and "\C \ lset Cs. is_free_for A v C" shows "is_free_for A v (\\<^sup>\\<^sub>\ B Cs)" using assms proof (induction Cs rule: rev_induct) case Nil then show ?case by simp next case (snoc C Cs) from snoc.prems(2) have "is_free_for A v C" and "\C \ lset Cs. is_free_for A v C" by simp_all with snoc.prems(1) have "is_free_for A v (\\<^sup>\\<^sub>\ B Cs)" using snoc.IH by simp with \is_free_for A v C\ show ?case using is_free_for_to_app by simp qed lemma is_free_for_in_equality [intro]: assumes "is_free_for A v B" and "is_free_for A v C" shows "is_free_for A v (B =\<^bsub>\\<^esub> C)" using assms unfolding equality_of_type_def and Q_def and Q_constant_of_type_def by (intro is_free_for_to_app is_free_for_in_con) lemma is_free_for_in_equivalence [intro]: assumes "is_free_for A v B" and "is_free_for A v C" shows "is_free_for A v (B \\<^sup>\ C)" using assms unfolding equivalence_def by (rule is_free_for_in_equality) lemma is_free_for_in_true [intro]: shows "is_free_for A v (T\<^bsub>o\<^esub>)" by force lemma is_free_for_in_false [intro]: shows "is_free_for A v (F\<^bsub>o\<^esub>)" unfolding false_def by (intro is_free_for_in_equality is_free_for_closed_form) simp_all lemma is_free_for_in_forall [intro]: assumes "is_free_for A v B" and "(x, \) \ free_vars A" shows "is_free_for A v (\x\<^bsub>\\<^esub>. B)" unfolding forall_def and PI_def proof (fold equality_of_type_def) have "is_free_for A v (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>)" using is_free_for_to_abs[OF is_free_for_in_true assms(2)] by fastforce moreover have "is_free_for A v (\x\<^bsub>\\<^esub>. B)" by (fact is_free_for_to_abs[OF assms]) ultimately show "is_free_for A v (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub> =\<^bsub>\\o\<^esub> \x\<^bsub>\\<^esub>. B)" by (iprover intro: assms(1) is_free_for_in_equality is_free_for_in_true is_free_for_to_abs) qed lemma is_free_for_in_generalized_forall [intro]: assumes "is_free_for A v B" and "lset vs \ free_vars A = {}" shows "is_free_for A v (\\<^sup>\\<^sub>\ vs B)" using assms proof (induction vs) case Nil then show ?case by simp next case (Cons v' vs) obtain x and \ where "v' = (x, \)" by fastforce from Cons.prems(2) have "v' \ free_vars A" and "lset vs \ free_vars A = {}" by simp_all from Cons.prems(1) and \lset vs \ free_vars A = {}\ have "is_free_for A v (\\<^sup>\\<^sub>\ vs B)" by (fact Cons.IH) from this and \v' \ free_vars A\[unfolded \v' = (x, \)\] have "is_free_for A v (\x\<^bsub>\\<^esub>. \\<^sup>\\<^sub>\ vs B)" by (intro is_free_for_in_forall) with \v' = (x, \)\ show ?case by simp qed lemma is_free_for_in_conj [intro]: assumes "is_free_for A v B" and "is_free_for A v C" shows "is_free_for A v (B \\<^sup>\ C)" proof - have "free_vars \\<^bsub>o\o\o\<^esub> = {}" by force then have "is_free_for A v (\\<^bsub>o\o\o\<^esub>)" using is_free_for_closed_form by fast with assms have "is_free_for A v (\\<^bsub>o\o\o\<^esub> \ B \ C)" by (intro is_free_for_to_app) then show ?thesis by (fold conj_op_def) qed lemma is_free_for_in_imp [intro]: assumes "is_free_for A v B" and "is_free_for A v C" shows "is_free_for A v (B \\<^sup>\ C)" proof - have "free_vars \\<^bsub>o\o\o\<^esub> = {}" by force then have "is_free_for A v (\\<^bsub>o\o\o\<^esub>)" using is_free_for_closed_form by fast with assms have "is_free_for A v (\\<^bsub>o\o\o\<^esub> \ B \ C)" by (intro is_free_for_to_app) then show ?thesis by (fold imp_op_def) qed lemma is_free_for_in_neg [intro]: assumes "is_free_for A v B" shows "is_free_for A v (\\<^sup>\ B)" using assms unfolding neg_def and Q_def and Q_constant_of_type_def by (intro is_free_for_to_app is_free_for_in_false is_free_for_in_con) lemma is_free_for_in_disj [intro]: assumes "is_free_for A v B" and "is_free_for A v C" shows "is_free_for A v (B \\<^sup>\ C)" proof - have "free_vars \\<^bsub>o\o\o\<^esub> = {}" by force then have "is_free_for A v (\\<^bsub>o\o\o\<^esub>)" using is_free_for_closed_form by fast with assms have "is_free_for A v (\\<^bsub>o\o\o\<^esub> \ B \ C)" by (intro is_free_for_to_app) then show ?thesis by (fold disj_op_def) qed lemma replacement_preserves_typing: assumes "C\p \ B\ \ D" and "A \\<^bsub>p\<^esub> C" and "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" shows "C \ wffs\<^bsub>\\<^esub> \ D \ wffs\<^bsub>\\<^esub>" using assms proof (induction arbitrary: \ rule: is_replacement_at.induct) case (pos_found p C C' A) then show ?case using diff_types_implies_diff_wffs by auto qed (metis is_subform_at.simps(2,3,4) wffs_from_app wffs_from_abs wffs_of_type_simps)+ corollary replacement_preserves_typing': assumes "C\p \ B\ \ D" and "A \\<^bsub>p\<^esub> C" and "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" and "C \ wffs\<^bsub>\\<^esub>" and "D \ wffs\<^bsub>\\<^esub>" shows "\ = \" using assms and replacement_preserves_typing and wff_has_unique_type by simp text \Closed formulas and sentences:\ definition is_closed_wff_of_type :: "form \ type \ bool" where [iff]: "is_closed_wff_of_type A \ \ A \ wffs\<^bsub>\\<^esub> \ free_vars A = {}" definition is_sentence :: "form \ bool" where [iff]: "is_sentence A \ is_closed_wff_of_type A o" subsection \Substitutions\ type_synonym substitution = "(var, form) fmap" definition is_substitution :: "substitution \ bool" where [iff]: "is_substitution \ \ (\(x, \) \ fmdom' \. \ $$! (x, \) \ wffs\<^bsub>\\<^esub>)" fun substitute :: "substitution \ form \ form" ("\<^bold>S _ _" [51, 51]) where "\<^bold>S \ (x\<^bsub>\\<^esub>) = (case \ $$ (x, \) of None \ x\<^bsub>\\<^esub> | Some A \ A)" | "\<^bold>S \ (\c\\<^bsub>\\<^esub>) = \c\\<^bsub>\\<^esub>" | "\<^bold>S \ (A \ B) = (\<^bold>S \ A) \ (\<^bold>S \ B)" | "\<^bold>S \ (\x\<^bsub>\\<^esub>. A) = (if (x, \) \ fmdom' \ then \x\<^bsub>\\<^esub>. \<^bold>S \ A else \x\<^bsub>\\<^esub>. \<^bold>S (fmdrop (x, \) \) A)" lemma empty_substitution_neutrality: shows "\<^bold>S {$$} A = A" by (induction A) auto lemma substitution_preserves_typing: assumes "is_substitution \" and "A \ wffs\<^bsub>\\<^esub>" shows "\<^bold>S \ A \ wffs\<^bsub>\\<^esub>" using assms(2) and assms(1)[unfolded is_substitution_def] proof (induction arbitrary: \) case (var_is_wff \ x) then show ?case by (cases "(x, \) \ fmdom' \") (use fmdom'_notI in \force+\) next case (abs_is_wff \ A \ x) then show ?case proof (cases "(x, \) \ fmdom' \") case True then have "\<^bold>S \ (\x\<^bsub>\\<^esub>. A) = \x\<^bsub>\\<^esub>. \<^bold>S (fmdrop (x, \) \) A" by simp moreover from abs_is_wff.prems have "is_substitution (fmdrop (x, \) \)" by fastforce with abs_is_wff.IH have "\<^bold>S (fmdrop (x, \) \) A \ wffs\<^bsub>\\<^esub>" by simp ultimately show ?thesis by auto next case False then have "\<^bold>S \ (\x\<^bsub>\\<^esub>. A) = \x\<^bsub>\\<^esub>. \<^bold>S \ A" by simp moreover from abs_is_wff.IH have "\<^bold>S \ A \ wffs\<^bsub>\\<^esub>" using abs_is_wff.prems by blast ultimately show ?thesis by fastforce qed qed force+ lemma derived_substitution_simps: shows "\<^bold>S \ T\<^bsub>o\<^esub> = T\<^bsub>o\<^esub>" and "\<^bold>S \ F\<^bsub>o\<^esub> = F\<^bsub>o\<^esub>" and "\<^bold>S \ (\\<^bsub>\\<^esub>) = \\<^bsub>\\<^esub>" and "\<^bold>S \ (\\<^sup>\ B) = \\<^sup>\ (\<^bold>S \ B)" and "\<^bold>S \ (B =\<^bsub>\\<^esub> C) = (\<^bold>S \ B) =\<^bsub>\\<^esub> (\<^bold>S \ C)" and "\<^bold>S \ (B \\<^sup>\ C) = (\<^bold>S \ B) \\<^sup>\ (\<^bold>S \ C)" and "\<^bold>S \ (B \\<^sup>\ C) = (\<^bold>S \ B) \\<^sup>\ (\<^bold>S \ C)" and "\<^bold>S \ (B \\<^sup>\ C) = (\<^bold>S \ B) \\<^sup>\ (\<^bold>S \ C)" and "\<^bold>S \ (B \\<^sup>\ C) = (\<^bold>S \ B) \\<^sup>\ (\<^bold>S \ C)" and "\<^bold>S \ (B \\<^bsub>\\<^esub> C) = (\<^bold>S \ B) \\<^bsub>\\<^esub> (\<^bold>S \ C)" and "\<^bold>S \ (\x\<^bsub>\\<^esub>. B) = (if (x, \) \ fmdom' \ then \x\<^bsub>\\<^esub>. \<^bold>S \ B else \x\<^bsub>\\<^esub>. \<^bold>S (fmdrop (x, \) \) B)" and "\<^bold>S \ (\x\<^bsub>\\<^esub>. B) = (if (x, \) \ fmdom' \ then \x\<^bsub>\\<^esub>. \<^bold>S \ B else \x\<^bsub>\\<^esub>. \<^bold>S (fmdrop (x, \) \) B)" by auto lemma generalized_app_substitution: shows "\<^bold>S \ (\\<^sup>\\<^sub>\ A Bs) = \\<^sup>\\<^sub>\ (\<^bold>S \ A) (map (\B. \<^bold>S \ B) Bs)" by (induction Bs arbitrary: A) simp_all lemma generalized_abs_substitution: shows "\<^bold>S \ (\\<^sup>\\<^sub>\ vs A) = \\<^sup>\\<^sub>\ vs (\<^bold>S (fmdrop_set (fmdom' \ \ lset vs) \) A)" proof (induction vs arbitrary: \) case Nil then show ?case by simp next case (Cons v vs) obtain x and \ where "v = (x, \)" by fastforce then show ?case proof (cases "v \ fmdom' \") case True then have *: "fmdom' \ \ lset (v # vs) = fmdom' \ \ lset vs" by simp from True have "\<^bold>S \ (\\<^sup>\\<^sub>\ (v # vs) A) = \x\<^bsub>\\<^esub>. \<^bold>S \ (\\<^sup>\\<^sub>\ vs A)" using \v = (x, \)\ by auto also have "\ = \x\<^bsub>\\<^esub>. \\<^sup>\\<^sub>\ vs (\<^bold>S (fmdrop_set (fmdom' \ \ lset vs) \) A)" using Cons.IH by (simp only:) also have "\ = \\<^sup>\\<^sub>\ (v # vs) (\<^bold>S (fmdrop_set (fmdom' \ \ lset (v # vs)) \) A)" using \v = (x, \)\ and * by auto finally show ?thesis . next case False let ?\' = "fmdrop v \" have *: "fmdrop_set (fmdom' \ \ lset (v # vs)) \ = fmdrop_set (fmdom' ?\' \ lset vs) ?\'" using False by clarsimp (metis Int_Diff Int_commute fmdrop_set_insert insert_Diff_single) from False have "\<^bold>S \ (\\<^sup>\\<^sub>\ (v # vs) A) = \x\<^bsub>\\<^esub>. \<^bold>S ?\' (\\<^sup>\\<^sub>\ vs A)" using \v = (x, \)\ by auto also have "\ = \x\<^bsub>\\<^esub>. \\<^sup>\\<^sub>\ vs (\<^bold>S (fmdrop_set (fmdom' ?\' \ lset vs) ?\') A)" using Cons.IH by (simp only:) also have "\ = \\<^sup>\\<^sub>\ (v # vs) (\<^bold>S (fmdrop_set (fmdom' \ \ lset (v # vs)) \) A)" using \v = (x, \)\ and * by auto finally show ?thesis . qed qed lemma generalized_forall_substitution: shows "\<^bold>S \ (\\<^sup>\\<^sub>\ vs A) = \\<^sup>\\<^sub>\ vs (\<^bold>S (fmdrop_set (fmdom' \ \ lset vs) \) A)" proof (induction vs arbitrary: \) case Nil then show ?case by simp next case (Cons v vs) obtain x and \ where "v = (x, \)" by fastforce then show ?case proof (cases "v \ fmdom' \") case True then have *: "fmdom' \ \ lset (v # vs) = fmdom' \ \ lset vs" by simp from True have "\<^bold>S \ (\\<^sup>\\<^sub>\ (v # vs) A) = \x\<^bsub>\\<^esub>. \<^bold>S \ (\\<^sup>\\<^sub>\ vs A)" using \v = (x, \)\ by auto also have "\ = \x\<^bsub>\\<^esub>. \\<^sup>\\<^sub>\ vs (\<^bold>S (fmdrop_set (fmdom' \ \ lset vs) \) A)" using Cons.IH by (simp only:) also have "\ = \\<^sup>\\<^sub>\ (v # vs) (\<^bold>S (fmdrop_set (fmdom' \ \ lset (v # vs)) \) A)" using \v = (x, \)\ and * by auto finally show ?thesis . next case False let ?\' = "fmdrop v \" have *: "fmdrop_set (fmdom' \ \ lset (v # vs)) \ = fmdrop_set (fmdom' ?\' \ lset vs) ?\'" using False by clarsimp (metis Int_Diff Int_commute fmdrop_set_insert insert_Diff_single) from False have "\<^bold>S \ (\\<^sup>\\<^sub>\ (v # vs) A) = \x\<^bsub>\\<^esub>. \<^bold>S ?\' (\\<^sup>\\<^sub>\ vs A)" using \v = (x, \)\ by auto also have "\ = \x\<^bsub>\\<^esub>. \\<^sup>\\<^sub>\ vs (\<^bold>S (fmdrop_set (fmdom' ?\' \ lset vs) ?\') A)" using Cons.IH by (simp only:) also have "\ = \\<^sup>\\<^sub>\ (v # vs) (\<^bold>S (fmdrop_set (fmdom' \ \ lset (v # vs)) \) A)" using \v = (x, \)\ and * by auto finally show ?thesis . qed qed lemma singleton_substitution_simps: shows "\<^bold>S {(x, \) \ A} (y\<^bsub>\\<^esub>) = (if (x, \) \ (y, \) then y\<^bsub>\\<^esub> else A)" and "\<^bold>S {(x, \) \ A} (\c\\<^bsub>\\<^esub>) = \c\\<^bsub>\\<^esub>" and "\<^bold>S {(x, \) \ A} (B \ C) = (\<^bold>S {(x, \) \ A} B) \ (\<^bold>S {(x, \) \ A} C)" and "\<^bold>S {(x, \) \ A} (\y\<^bsub>\\<^esub>. B) = \y\<^bsub>\\<^esub>. (if (x, \) = (y, \) then B else \<^bold>S {(x, \) \ A} B)" by (simp_all add: empty_substitution_neutrality fmdrop_fmupd_same) lemma substitution_preserves_freeness: assumes "y \ free_vars A" and "y \ z" shows "y \ free_vars \<^bold>S {x \ FVar z} A" using assms(1) proof (induction A rule: free_vars_form.induct) case (1 x' \) with assms(2) show ?case using surj_pair[of z] by (cases "x = (x', \)") force+ next case (4 x' \ A) then show ?case using surj_pair[of z] by (cases "x = (x', \)") (use singleton_substitution_simps(4) in presburger, auto) qed auto lemma renaming_substitution_minimal_change: assumes "y \ vars A" and "y \ z" shows "y \ vars (\<^bold>S {x \ FVar z} A)" using assms(1) proof (induction A rule: vars_form.induct) case (1 x' \) with assms(2) show ?case using surj_pair[of z] by (cases "x = (x', \)") force+ next case (4 x' \ A) then show ?case using surj_pair[of z] by (cases "x = (x', \)") (use singleton_substitution_simps(4) in presburger, auto) qed auto lemma free_var_singleton_substitution_neutrality: assumes "v \ free_vars A" shows "\<^bold>S {v \ B} A = A" using assms by (induction A rule: free_vars_form.induct) (simp_all, metis empty_substitution_neutrality fmdrop_empty fmdrop_fmupd_same) lemma identity_singleton_substitution_neutrality: shows "\<^bold>S {v \ FVar v} A = A" by (induction A rule: free_vars_form.induct) (simp_all add: empty_substitution_neutrality fmdrop_fmupd_same) lemma free_var_in_renaming_substitution: assumes "x \ y" shows "(x, \) \ free_vars (\<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} B)" using assms by (induction B rule: free_vars_form.induct) simp_all lemma renaming_substitution_preserves_form_size: shows "form_size (\<^bold>S {v \ FVar v'} A) = form_size A" proof (induction A rule: form_size.induct) case (1 x \) then show ?case using form_size.elims by auto next case (4 x \ A) then show ?case by (cases "v = (x, \)") (use singleton_substitution_simps(4) in presburger, auto) qed simp_all -text \The following lemma corresponds to X5100 in @{cite "andrews:2002"}:\ +text \The following lemma corresponds to X5100 in \<^cite>\"andrews:2002"\:\ lemma substitution_composability: assumes "v' \ vars B" shows "\<^bold>S {v' \ A} \<^bold>S {v \ FVar v'} B = \<^bold>S {v \ A} B" using assms proof (induction B arbitrary: v') case (FAbs w C) then show ?case proof (cases "v = w") case True from \v' \ vars (FAbs w C)\ have "v' \ free_vars (FAbs w C)" using free_vars_in_all_vars by blast then have "\<^bold>S {v' \ A} (FAbs w C) = FAbs w C" by (rule free_var_singleton_substitution_neutrality) from \v = w\ have "v \ free_vars (FAbs w C)" using surj_pair[of w] by fastforce then have "\<^bold>S {v \ A} (FAbs w C) = FAbs w C" by (fact free_var_singleton_substitution_neutrality) also from \\<^bold>S {v' \ A} (FAbs w C) = FAbs w C\ have "\ = \<^bold>S {v' \ A} (FAbs w C)" by (simp only:) also from \v = w\ have "\ = \<^bold>S {v' \ A} \<^bold>S {v \ FVar v'} (FAbs w C)" using free_var_singleton_substitution_neutrality[OF \v \ free_vars (FAbs w C)\] by (simp only:) finally show ?thesis .. next case False from FAbs.prems have "v' \ vars C" using surj_pair[of w] by fastforce then show ?thesis proof (cases "v' = w") case True with FAbs.prems show ?thesis using vars_form.elims by auto next case False from \v \ w\ have "\<^bold>S {v \ A} (FAbs w C) = FAbs w (\<^bold>S {v \ A} C)" using surj_pair[of w] by fastforce also from FAbs.IH have "\ = FAbs w (\<^bold>S {v' \ A} \<^bold>S {v \ FVar v'} C)" using \v' \ vars C\ by simp also from \v' \ w\ have "\ = \<^bold>S {v' \ A} (FAbs w (\<^bold>S {v \ FVar v'} C))" using surj_pair[of w] by fastforce also from \v \ w\ have "\ = \<^bold>S {v' \ A} \<^bold>S {v \ FVar v'} (FAbs w C)" using surj_pair[of w] by fastforce finally show ?thesis .. qed qed qed auto -text \The following lemma corresponds to X5101 in @{cite "andrews:2002"}:\ +text \The following lemma corresponds to X5101 in \<^cite>\"andrews:2002"\:\ lemma renaming_substitution_composability: assumes "z \ free_vars A" and "is_free_for (FVar z) x A" shows "\<^bold>S {z \ FVar y} \<^bold>S {x \ FVar z} A = \<^bold>S {x \ FVar y} A" using assms proof (induction A arbitrary: z) case (FVar v) then show ?case using surj_pair[of v] and surj_pair[of z] by fastforce next case (FCon k) then show ?case using surj_pair[of k] by fastforce next case (FApp B C) let ?\\<^sub>z\<^sub>y = "{z \ FVar y}" and ?\\<^sub>x\<^sub>z = "{x \ FVar z}" and ?\\<^sub>x\<^sub>y = "{x \ FVar y}" from \is_free_for (FVar z) x (B \ C)\ have "is_free_for (FVar z) x B" and "is_free_for (FVar z) x C" using is_free_for_from_app by iprover+ moreover from \z \ free_vars (B \ C)\ have "z \ free_vars B" and "z \ free_vars C" by simp_all ultimately have *: "\<^bold>S ?\\<^sub>z\<^sub>y \<^bold>S ?\\<^sub>x\<^sub>z B = \<^bold>S ?\\<^sub>x\<^sub>y B" and **: "\<^bold>S ?\\<^sub>z\<^sub>y \<^bold>S ?\\<^sub>x\<^sub>z C = \<^bold>S ?\\<^sub>x\<^sub>y C" using FApp.IH by simp_all have "\<^bold>S ?\\<^sub>z\<^sub>y \<^bold>S ?\\<^sub>x\<^sub>z (B \ C) = (\<^bold>S ?\\<^sub>z\<^sub>y \<^bold>S ?\\<^sub>x\<^sub>z B) \ (\<^bold>S ?\\<^sub>z\<^sub>y \<^bold>S ?\\<^sub>x\<^sub>z C)" by simp also from * and ** have "\ = (\<^bold>S ?\\<^sub>x\<^sub>y B) \ (\<^bold>S ?\\<^sub>x\<^sub>y C)" by (simp only:) also have "\ = \<^bold>S ?\\<^sub>x\<^sub>y (B \ C)" by simp finally show ?case . next case (FAbs w B) let ?\\<^sub>z\<^sub>y = "{z \ FVar y}" and ?\\<^sub>x\<^sub>z = "{x \ FVar z}" and ?\\<^sub>x\<^sub>y = "{x \ FVar y}" show ?case proof (cases "x = w") case True then show ?thesis proof (cases "z = w") case True with \x = w\ have "x \ free_vars (FAbs w B)" and "z \ free_vars (FAbs w B)" using surj_pair[of w] by fastforce+ from \x \ free_vars (FAbs w B)\ have "\<^bold>S ?\\<^sub>x\<^sub>y (FAbs w B) = FAbs w B" by (fact free_var_singleton_substitution_neutrality) also from \z \ free_vars (FAbs w B)\ have "\ = \<^bold>S ?\\<^sub>z\<^sub>y (FAbs w B)" by (fact free_var_singleton_substitution_neutrality[symmetric]) also from \x \ free_vars (FAbs w B)\ have "\ = \<^bold>S ?\\<^sub>z\<^sub>y \<^bold>S ?\\<^sub>x\<^sub>z (FAbs w B)" using free_var_singleton_substitution_neutrality by simp finally show ?thesis .. next case False with \x = w\ have "z \ free_vars B" and "x \ free_vars (FAbs w B)" using \z \ free_vars (FAbs w B)\ and surj_pair[of w] by fastforce+ from \z \ free_vars B\ have "\<^bold>S ?\\<^sub>z\<^sub>y B = B" by (fact free_var_singleton_substitution_neutrality) from \x \ free_vars (FAbs w B)\ have "\<^bold>S ?\\<^sub>x\<^sub>y (FAbs w B) = FAbs w B" by (fact free_var_singleton_substitution_neutrality) also from \\<^bold>S ?\\<^sub>z\<^sub>y B = B\ have "\ = FAbs w (\<^bold>S ?\\<^sub>z\<^sub>y B)" by (simp only:) also from \z \ free_vars (FAbs w B)\ have "\ = \<^bold>S ?\\<^sub>z\<^sub>y (FAbs w B)" by (simp add: \FAbs w B = FAbs w (\<^bold>S ?\\<^sub>z\<^sub>y B)\ free_var_singleton_substitution_neutrality) also from \x \ free_vars (FAbs w B)\ have "\ = \<^bold>S ?\\<^sub>z\<^sub>y \<^bold>S ?\\<^sub>x\<^sub>z (FAbs w B)" using free_var_singleton_substitution_neutrality by simp finally show ?thesis .. qed next case False then show ?thesis proof (cases "z = w") case True have "x \ free_vars B" proof (rule ccontr) assume "\ x \ free_vars B" with \x \ w\ have "x \ free_vars (FAbs w B)" using surj_pair[of w] by fastforce then obtain p where "p \ positions (FAbs w B)" and "is_free_at x p (FAbs w B)" using free_vars_in_is_free_at by blast with \is_free_for (FVar z) x (FAbs w B)\ have "\ in_scope_of_abs z p (FAbs w B)" by (meson empty_is_position is_free_at_in_free_vars is_free_at_in_var is_free_for_def) moreover obtain p' where "p = \ # p'" using is_free_at_from_absE[OF \is_free_at x p (FAbs w B)\] by blast ultimately have "z \ w" using in_scope_of_abs_in_abs by blast with \z = w\ show False by contradiction qed then have *: "\<^bold>S ?\\<^sub>x\<^sub>y B = \<^bold>S ?\\<^sub>x\<^sub>z B" using free_var_singleton_substitution_neutrality by auto from \x \ w\ have "\<^bold>S ?\\<^sub>x\<^sub>y (FAbs w B) = FAbs w (\<^bold>S ?\\<^sub>x\<^sub>y B)" using surj_pair[of w] by fastforce also from * have "\ = FAbs w (\<^bold>S ?\\<^sub>x\<^sub>z B)" by (simp only:) also from FAbs.prems(1) have "\ = \<^bold>S ?\\<^sub>z\<^sub>y (FAbs w (\<^bold>S ?\\<^sub>x\<^sub>z B))" using \x \ free_vars B\ and free_var_singleton_substitution_neutrality by auto also from \x \ w\ have "\ = \<^bold>S ?\\<^sub>z\<^sub>y \<^bold>S ?\\<^sub>x\<^sub>z (FAbs w B)" using surj_pair[of w] by fastforce finally show ?thesis .. next case False obtain v\<^sub>w and \ where "w = (v\<^sub>w, \)" by fastforce with \is_free_for (FVar z) x (FAbs w B)\ and \x \ w\ have "is_free_for (FVar z) x B" using is_free_for_from_abs by iprover moreover from \z \ free_vars (FAbs w B)\ and \z \ w\ and \w = (v\<^sub>w, \)\ have "z \ free_vars B" by simp ultimately have *: "\<^bold>S ?\\<^sub>z\<^sub>y \<^bold>S ?\\<^sub>x\<^sub>z B = \<^bold>S ?\\<^sub>x\<^sub>y B" using FAbs.IH by simp from \x \ w\ have "\<^bold>S ?\\<^sub>x\<^sub>y (FAbs w B) = FAbs w (\<^bold>S ?\\<^sub>x\<^sub>y B)" using \w = (v\<^sub>w, \)\ and free_var_singleton_substitution_neutrality by simp also from * have "\ = FAbs w (\<^bold>S ?\\<^sub>z\<^sub>y \<^bold>S ?\\<^sub>x\<^sub>z B)" by (simp only:) also from \z \ w\ have "\ = \<^bold>S ?\\<^sub>z\<^sub>y (FAbs w (\<^bold>S ?\\<^sub>x\<^sub>z B))" using \w = (v\<^sub>w, \)\ and free_var_singleton_substitution_neutrality by simp also from \x \ w\ have "\ = \<^bold>S ?\\<^sub>z\<^sub>y \<^bold>S ?\\<^sub>x\<^sub>z (FAbs w B)" using \w = (v\<^sub>w, \)\ and free_var_singleton_substitution_neutrality by simp finally show ?thesis .. qed qed qed lemma absent_vars_substitution_preservation: assumes "v \ vars A" and "\v' \ fmdom' \. v \ vars (\ $$! v')" shows "v \ vars (\<^bold>S \ A)" using assms proof (induction A arbitrary: \) case (FVar v') then show ?case using surj_pair[of v'] by (cases "v' \ fmdom' \") (use fmlookup_dom'_iff in force)+ next case (FCon k) then show ?case using surj_pair[of k] by fastforce next case FApp then show ?case by simp next case (FAbs w B) from FAbs.prems(1) have "v \ vars B" using vars_form.elims by auto then show ?case proof (cases "w \ fmdom' \") case True from FAbs.prems(2) have "\v' \ fmdom' (fmdrop w \). v \ vars ((fmdrop w \) $$! v')" by auto with \v \ vars B\ have "v \ vars (\<^bold>S (fmdrop w \) B)" by (fact FAbs.IH) with FAbs.prems(1) have "v \ vars (FAbs w (\<^bold>S (fmdrop w \) B))" using surj_pair[of w] by fastforce moreover from True have "\<^bold>S \ (FAbs w B) = FAbs w (\<^bold>S (fmdrop w \) B)" using surj_pair[of w] by fastforce ultimately show ?thesis by simp next case False then show ?thesis using FAbs.IH and FAbs.prems and surj_pair[of w] by fastforce qed qed lemma substitution_free_absorption: assumes "\ $$ v = None" and "v \ free_vars B" shows "\<^bold>S ({v \ A} ++\<^sub>f \) B = \<^bold>S \ B" using assms proof (induction B arbitrary: \) case (FAbs w B) show ?case proof (cases "v \ w") case True with FAbs.prems(2) have "v \ free_vars B" using surj_pair[of w] by fastforce then show ?thesis proof (cases "w \ fmdom' \") case True then have "\<^bold>S ({v \ A} ++\<^sub>f \) (FAbs w B) = FAbs w (\<^bold>S (fmdrop w ({v \ A} ++\<^sub>f \)) B)" using surj_pair[of w] by fastforce also from \v \ w\ and True have "\ = FAbs w (\<^bold>S ({v \ A} ++\<^sub>f fmdrop w \) B)" by (simp add: fmdrop_fmupd) also from FAbs.prems(1) and \v \ free_vars B\ have "\ = FAbs w (\<^bold>S (fmdrop w \) B)" using FAbs.IH by simp also from True have "\ = \<^bold>S \ (FAbs w B)" using surj_pair[of w] by fastforce finally show ?thesis . next case False with FAbs.prems(1) have "\<^bold>S ({v \ A} ++\<^sub>f \) (FAbs w B) = FAbs w (\<^bold>S ({v \ A} ++\<^sub>f \) B)" using \v \ w\ and surj_pair[of w] by fastforce also from FAbs.prems(1) and \v \ free_vars B\ have "\ = FAbs w (\<^bold>S \ B)" using FAbs.IH by simp also from False have "\ = \<^bold>S \ (FAbs w B)" using surj_pair[of w] by fastforce finally show ?thesis . qed next case False then have "fmdrop w ({v \ A} ++\<^sub>f \) = fmdrop w \" by (simp add: fmdrop_fmupd_same) then show ?thesis using surj_pair[of w] by (metis (no_types, lifting) fmdrop_idle' substitute.simps(4)) qed qed fastforce+ lemma substitution_absorption: assumes "\ $$ v = None" and "v \ vars B" shows "\<^bold>S ({v \ A} ++\<^sub>f \) B = \<^bold>S \ B" using assms by (meson free_vars_in_all_vars in_mono substitution_free_absorption) lemma is_free_for_with_renaming_substitution: assumes "is_free_for A x B" and "y \ vars B" and "x \ fmdom' \" and "\v \ fmdom' \. y \ vars (\ $$! v)" and "\v \ fmdom' \. is_free_for (\ $$! v) v B" shows "is_free_for A y (\<^bold>S ({x \ FVar y} ++\<^sub>f \) B)" using assms proof (induction B arbitrary: \) case (FVar w) then show ?case proof (cases "w = x") case True with FVar.prems(3) have "\<^bold>S ({x \ FVar y} ++\<^sub>f \) (FVar w) = FVar y" using surj_pair[of w] by fastforce then show ?thesis using self_subform_is_at_top by fastforce next case False then show ?thesis proof (cases "w \ fmdom' \") case True from False have "\<^bold>S ({x \ FVar y} ++\<^sub>f \) (FVar w) = \<^bold>S \ (FVar w)" using substitution_absorption and surj_pair[of w] by force also from True have "\ = \ $$! w" using surj_pair[of w] by (metis fmdom'_notI option.case_eq_if substitute.simps(1)) finally have "\<^bold>S ({x \ FVar y} ++\<^sub>f \) (FVar w) = \ $$! w" . moreover from True and FVar.prems(4) have "y \ vars (\ $$! w)" by blast ultimately show ?thesis using form_is_free_for_absent_var by presburger next case False with FVar.prems(3) and \w \ x\ have "\<^bold>S ({x \ FVar y} ++\<^sub>f \) (FVar w) = FVar w" using surj_pair[of w] by fastforce with FVar.prems(2) show ?thesis using form_is_free_for_absent_var by presburger qed qed next case (FCon k) then show ?case using surj_pair[of k] by fastforce next case (FApp C D) from FApp.prems(2) have "y \ vars C" and "y \ vars D" by simp_all from FApp.prems(1) have "is_free_for A x C" and "is_free_for A x D" using is_free_for_from_app by iprover+ have "\v \ fmdom' \. is_free_for (\ $$! v) v C \ is_free_for (\ $$! v) v D" proof (rule ballI) fix v assume "v \ fmdom' \" with FApp.prems(5) have "is_free_for (\ $$! v) v (C \ D)" by blast then show "is_free_for (\ $$! v) v C \ is_free_for (\ $$! v) v D" using is_free_for_from_app by iprover+ qed then have *: "\v \ fmdom' \. is_free_for (\ $$! v) v C" and **: "\v \ fmdom' \. is_free_for (\ $$! v) v D" by auto have "\<^bold>S ({x \ FVar y} ++\<^sub>f \) (C \ D) = (\<^bold>S ({x \ FVar y} ++\<^sub>f \) C) \ (\<^bold>S ({x \ FVar y} ++\<^sub>f \) D)" by simp moreover have "is_free_for A y (\<^bold>S ({x \ FVar y} ++\<^sub>f \) C)" by (rule FApp.IH(1)[OF \is_free_for A x C\ \y \ vars C\ FApp.prems(3,4) *]) moreover have "is_free_for A y (\<^bold>S ({x \ FVar y} ++\<^sub>f \) D)" by (rule FApp.IH(2)[OF \is_free_for A x D\ \y \ vars D\ FApp.prems(3,4) **]) ultimately show ?case using is_free_for_in_app by simp next case (FAbs w B) obtain x\<^sub>w and \\<^sub>w where "w = (x\<^sub>w, \\<^sub>w)" by fastforce from FAbs.prems(2) have "y \ vars B" using vars_form.elims by auto then show ?case proof (cases "w = x") case True from True and \x \ fmdom' \\ have "w \ fmdom' \" and "x \ free_vars (FAbs w B)" using \w = (x\<^sub>w, \\<^sub>w)\ by fastforce+ with True have "\<^bold>S ({x \ FVar y} ++\<^sub>f \) (FAbs w B) = \<^bold>S \ (FAbs w B)" using substitution_free_absorption by blast also have "\ = FAbs w (\<^bold>S \ B)" using \w = (x\<^sub>w, \\<^sub>w)\ \w \ fmdom' \\ substitute.simps(4) by presburger finally have "\<^bold>S ({x \ FVar y} ++\<^sub>f \) (FAbs w B) = FAbs w (\<^bold>S \ B)" . moreover from \\<^bold>S \ (FAbs w B) = FAbs w (\<^bold>S \ B)\ have "y \ vars (FAbs w (\<^bold>S \ B))" using absent_vars_substitution_preservation[OF FAbs.prems(2,4)] by simp ultimately show ?thesis using is_free_for_absent_var by (simp only:) next case False obtain v\<^sub>w and \\<^sub>w where "w = (v\<^sub>w, \\<^sub>w)" by fastforce from FAbs.prems(1) and \w \ x\ and \w = (v\<^sub>w, \\<^sub>w)\ have "is_free_for A x B" using is_free_for_from_abs by iprover then show ?thesis proof (cases "w \ fmdom' \") case True then have "\<^bold>S ({x \ FVar y} ++\<^sub>f \) (FAbs w B) = FAbs w (\<^bold>S (fmdrop w ({x \ FVar y} ++\<^sub>f \)) B)" using \w = (v\<^sub>w, \\<^sub>w)\ by (simp add: fmdrop_idle') also from \w \ x\ and True have "\ = FAbs w (\<^bold>S ({x \ FVar y} ++\<^sub>f fmdrop w \) B)" by (simp add: fmdrop_fmupd) finally have *: "\<^bold>S ({x \ FVar y} ++\<^sub>f \) (FAbs w B) = FAbs w (\<^bold>S ({x \ FVar y} ++\<^sub>f fmdrop w \) B)" . have "\v \ fmdom' (fmdrop w \). is_free_for (fmdrop w \ $$! v) v B" proof fix v assume "v \ fmdom' (fmdrop w \)" with FAbs.prems(5) have "is_free_for (fmdrop w \ $$! v) v (FAbs w B)" by auto moreover from \v \ fmdom' (fmdrop w \)\ have "v \ w" by auto ultimately show "is_free_for (fmdrop w \ $$! v) v B" unfolding \w = (v\<^sub>w, \\<^sub>w)\ using is_free_for_from_abs by iprover qed moreover from FAbs.prems(3) have "x \ fmdom' (fmdrop w \)" by simp moreover from FAbs.prems(4) have "\v \ fmdom' (fmdrop w \). y \ vars (fmdrop w \ $$! v)" by simp ultimately have "is_free_for A y (\<^bold>S ({x \ FVar y} ++\<^sub>f fmdrop w \) B)" using \is_free_for A x B\ and \y \ vars B\ and FAbs.IH by iprover then show ?thesis proof (cases "x \ free_vars B") case True have "y \ vars (\<^bold>S ({x \ FVar y} ++\<^sub>f \) (FAbs w B))" proof - have "\<^bold>S ({x \ FVar y} ++\<^sub>f \) (FAbs w B) = FAbs w (\<^bold>S ({x \ FVar y} ++\<^sub>f fmdrop w \) B)" using * . also from \x \ free_vars B\ and FAbs.prems(3) have "\ = FAbs w (\<^bold>S (fmdrop w \) B)" using substitution_free_absorption by (simp add: fmdom'_notD) finally have "\<^bold>S ({x \ FVar y} ++\<^sub>f \) (FAbs w B) = FAbs w (\<^bold>S (fmdrop w \) B)" . with FAbs.prems(2) and \w = (v\<^sub>w, \\<^sub>w)\ and FAbs.prems(4) show ?thesis using absent_vars_substitution_preservation by auto qed then show ?thesis using is_free_for_absent_var by simp next case False have "w \ free_vars A" proof (rule ccontr) assume "\ w \ free_vars A" with False and \w \ x\ have "\ is_free_for A x (FAbs w B)" using form_with_free_binder_not_free_for by simp with FAbs.prems(1) show False by contradiction qed with \is_free_for A y (\<^bold>S ({x \ FVar y} ++\<^sub>f fmdrop w \) B)\ have "is_free_for A y (FAbs w (\<^bold>S ({x \ FVar y} ++\<^sub>f fmdrop w \) B))" unfolding \w = (v\<^sub>w, \\<^sub>w)\ using is_free_for_to_abs by iprover with * show ?thesis by (simp only:) qed next case False have "\v \ fmdom' \. is_free_for (\ $$! v) v B" proof (rule ballI) fix v assume "v \ fmdom' \" with FAbs.prems(5) have "is_free_for (\ $$! v) v (FAbs w B)" by blast moreover from \v \ fmdom' \\ and \w \ fmdom' \\ have "v \ w" by blast ultimately show "is_free_for (\ $$! v) v B" unfolding \w = (v\<^sub>w, \\<^sub>w)\ using is_free_for_from_abs by iprover qed with \is_free_for A x B\ and \y \ vars B\ and FAbs.prems(3,4) have "is_free_for A y (\<^bold>S ({x \ FVar y} ++\<^sub>f \) B)" using FAbs.IH by iprover then show ?thesis proof (cases "x \ free_vars B") case True have "y \ vars (\<^bold>S ({x \ FVar y} ++\<^sub>f \) (FAbs w B))" proof - from False and \w = (v\<^sub>w, \\<^sub>w)\ and \w \ x\ have "\<^bold>S ({x \ FVar y} ++\<^sub>f \) (FAbs w B) = FAbs w (\<^bold>S ({x \ FVar y} ++\<^sub>f \) B)" by auto also from \x \ free_vars B\ and FAbs.prems(3) have "\ = FAbs w (\<^bold>S \ B)" using substitution_free_absorption by (simp add: fmdom'_notD) finally have "\<^bold>S ({x \ FVar y} ++\<^sub>f \) (FAbs w B) = FAbs w (\<^bold>S \ B)" . with FAbs.prems(2,4) and \w = (v\<^sub>w, \\<^sub>w)\ show ?thesis using absent_vars_substitution_preservation by auto qed then show ?thesis using is_free_for_absent_var by simp next case False have "w \ free_vars A" proof (rule ccontr) assume "\ w \ free_vars A" with False and \w \ x\ have "\ is_free_for A x (FAbs w B)" using form_with_free_binder_not_free_for by simp with FAbs.prems(1) show False by contradiction qed with \is_free_for A y (\<^bold>S ({x \ FVar y} ++\<^sub>f \) B)\ have "is_free_for A y (FAbs w (\<^bold>S ({x \ FVar y} ++\<^sub>f \) B))" unfolding \w = (v\<^sub>w, \\<^sub>w)\ using is_free_for_to_abs by iprover moreover from \w \ fmdom' \\ and \w \ x\ and FAbs.prems(3) have "\<^bold>S ({x \ FVar y} ++\<^sub>f \) (FAbs w B) = FAbs w (\<^bold>S ({x \ FVar y} ++\<^sub>f \) B)" using surj_pair[of w] by fastforce ultimately show ?thesis by (simp only:) qed qed qed qed text \ The following lemma allows us to fuse a singleton substitution and a simultaneous substitution, as long as the variable of the former does not occur anywhere in the latter: \ lemma substitution_fusion: assumes "is_substitution \" and "is_substitution {v \ A}" and "\ $$ v = None" and "\v' \ fmdom' \. v \ vars (\ $$! v')" shows "\<^bold>S {v \ A} \<^bold>S \ B = \<^bold>S ({v \ A} ++\<^sub>f \) B" using assms(1,3,4) proof (induction B arbitrary: \) case (FVar v') then show ?case proof (cases "v' \ fmdom' \") case True then show ?thesis using surj_pair[of v'] by fastforce next case False then obtain A' where "\ $$ v' = Some A'" by (meson fmlookup_dom'_iff) with False and FVar.prems(3) have "v \ vars A'" by fastforce then have "\<^bold>S {v \ A} A' = A'" using free_var_singleton_substitution_neutrality and free_vars_in_all_vars by blast from \\ $$ v' = Some A'\ have "\<^bold>S {v \ A} \<^bold>S \ (FVar v') = \<^bold>S {v \ A} A'" using surj_pair[of v'] by fastforce also from \\<^bold>S {v \ A} A' = A'\ have "\ = A'" by (simp only:) also from \\ $$ v' = Some A'\ and \\ $$ v = None\ have "\ = \<^bold>S ({v \ A} ++\<^sub>f \) (FVar v')" using surj_pair[of v'] by fastforce finally show ?thesis . qed next case (FCon k) then show ?case using surj_pair[of k] by fastforce next case (FApp C D) have "\<^bold>S {v \ A} \<^bold>S \ (C \ D) = \<^bold>S {v \ A} ((\<^bold>S \ C) \ (\<^bold>S \ D))" by auto also have "\ = (\<^bold>S {v \ A} \<^bold>S \ C) \ (\<^bold>S {v \ A} \<^bold>S \ D)" by simp also from FApp.IH have "\ = (\<^bold>S ({v \ A} ++\<^sub>f \) C) \ (\<^bold>S ({v \ A} ++\<^sub>f \) D)" using FApp.prems(1,2,3) by presburger also have "\ = \<^bold>S ({v \ A} ++\<^sub>f \) (C \ D)" by simp finally show ?case . next case (FAbs w C) obtain v\<^sub>w and \ where "w = (v\<^sub>w, \)" by fastforce then show ?case proof (cases "v \ w") case True show ?thesis proof (cases "w \ fmdom' \") case True then have "\<^bold>S {v \ A} \<^bold>S \ (FAbs w C) = \<^bold>S {v \ A} (FAbs w (\<^bold>S \ C))" by (simp add: \w = (v\<^sub>w, \)\) also from \v \ w\ have "\ = FAbs w (\<^bold>S {v \ A} \<^bold>S \ C)" by (simp add: \w = (v\<^sub>w, \)\) also from FAbs.IH have "\ = FAbs w (\<^bold>S ({v \ A} ++\<^sub>f \) C)" using FAbs.prems(1,2,3) by blast also from \v \ w\ and True have "\ = \<^bold>S ({v \ A} ++\<^sub>f \) (FAbs w C)" by (simp add: \w = (v\<^sub>w, \)\) finally show ?thesis . next case False then have "\<^bold>S {v \ A} \<^bold>S \ (FAbs w C) = \<^bold>S {v \ A} (FAbs w (\<^bold>S (fmdrop w \) C))" by (simp add: \w = (v\<^sub>w, \)\) also from \v \ w\ have "\ = FAbs w (\<^bold>S {v \ A} \<^bold>S (fmdrop w \) C)" by (simp add: \w = (v\<^sub>w, \)\) also have "\ = FAbs w (\<^bold>S ({v \ A} ++\<^sub>f fmdrop w \) C)" proof - from \is_substitution \\ have "is_substitution (fmdrop w \)" by fastforce moreover from \\ $$ v = None\ have "(fmdrop w \) $$ v = None" by force moreover from FAbs.prems(3) have "\v' \ fmdom' (fmdrop w \). v \ vars ((fmdrop w \) $$! v')" by force ultimately show ?thesis using FAbs.IH by blast qed also from \v \ w\ have "\ = \<^bold>S ({v \ A} ++\<^sub>f \) (FAbs w C)" by (simp add: \w = (v\<^sub>w, \)\ fmdrop_idle') finally show ?thesis . qed next case False then show ?thesis proof (cases "w \ fmdom' \") case True then have "\<^bold>S {v \ A} \<^bold>S \ (FAbs w C) = \<^bold>S {v \ A} (FAbs w (\<^bold>S \ C))" by (simp add: \w = (v\<^sub>w, \)\) also from \\ v \ w\ have "\ = FAbs w (\<^bold>S \ C)" using \w = (v\<^sub>w, \)\ and singleton_substitution_simps(4) by presburger also from \\ v \ w\ and True have "\ = FAbs w (\<^bold>S (fmdrop w ({v \ A} ++\<^sub>f \)) C)" by (simp add: fmdrop_fmupd_same fmdrop_idle') also from \\ v \ w\ have "\ = \<^bold>S ({v \ A} ++\<^sub>f \) (FAbs w C)" by (simp add: \w = (v\<^sub>w, \)\) finally show ?thesis . next case False then have "\<^bold>S {v \ A} \<^bold>S \ (FAbs w C) = \<^bold>S {v \ A} (FAbs w (\<^bold>S (fmdrop w \) C))" by (simp add: \w = (v\<^sub>w, \)\) also from \\ v \ w\ have "\ = FAbs w (\<^bold>S (fmdrop w \) C)" using \\ $$ v = None\ and False by (simp add: fmdom'_notI) also from \\ v \ w\ have "\ = FAbs w (\<^bold>S (fmdrop w ({v \ A} ++\<^sub>f \)) C)" by (simp add: fmdrop_fmupd_same) also from \\ v \ w\ and False and \\ $$ v = None\ have "\ = \<^bold>S ({v \ A} ++\<^sub>f \) (FAbs w C)" by (simp add: fmdom'_notI) finally show ?thesis . qed qed qed lemma updated_substitution_is_substitution: assumes "v \ fmdom' \" and "is_substitution (\(v \ A))" shows "is_substitution \" unfolding is_substitution_def proof (intro ballI) fix v' :: var obtain x and \ where "v' = (x, \)" by fastforce assume "v' \ fmdom' \" with assms(2)[unfolded is_substitution_def] have "v' \ fmdom' (\(v \ A))" by simp with assms(2)[unfolded is_substitution_def] have "\(v \ A) $$! (x, \) \ wffs\<^bsub>\\<^esub>" using \v' = (x, \)\ by fastforce with assms(1) and \v' \ fmdom' \\ and \v' = (x, \)\ have "\ $$! (x, \) \ wffs\<^bsub>\\<^esub>" by (metis fmupd_lookup) then show "case v' of (x, \) \ \ $$! (x, \) \ wffs\<^bsub>\\<^esub>" by (simp add: \v' = (x, \)\) qed definition is_renaming_substitution where [iff]: "is_renaming_substitution \ \ is_substitution \ \ fmpred (\_ A. \v. A = FVar v) \" text \ The following lemma proves that $ \d{\textsf{S}}\;^{x^1_{\alpha_1}\;\dots\;x^n_{\alpha_n}}_{y^1_{\alpha_1}\;\dots\;y^n_{\alpha_n}} B = \d{\textsf{S}}\;^{x^1_{\alpha_1}}_{y^1_{\alpha_1}}\;\cdots\; \d{\textsf{S}}\;^{x^n_{\alpha_n}}_{y^n_{\alpha_n}} B$ provided that \<^item> $x^1_{\alpha_1}\;\dots\;x^n_{\alpha_n}$ are distinct variables \<^item> $y^1_{\alpha_1}\;\dots\;y^n_{\alpha_n}$ are distinct variables, distinct from $x^1_{\alpha_1}\;\dots\;x^n_{\alpha_n}$ and from all variables in \B\ (i.e., they are fresh variables) In other words, simultaneously renaming distinct variables with fresh ones is equivalent to renaming each variable one at a time. \ lemma fresh_vars_substitution_unfolding: fixes ps :: "(var \ form) list" assumes "\ = fmap_of_list ps" and "is_renaming_substitution \" and "distinct (map fst ps)" and "distinct (map snd ps)" and "vars (fmran' \) \ (fmdom' \ \ vars B) = {}" shows "\<^bold>S \ B = foldr (\(x, y) C. \<^bold>S {x \ y} C) ps B" using assms proof (induction ps arbitrary: \) case Nil then have "\ = {$$}" by simp then have "\<^bold>S \ B = B" using empty_substitution_neutrality by (simp only:) with Nil show ?case by simp next case (Cons p ps) from Cons.prems(1,2) obtain x and y where "\ $$ (fst p) = Some (FVar y)" and "p = (x, FVar y)" using surj_pair[of p] by fastforce let ?\' = "fmap_of_list ps" from Cons.prems(1) and \p = (x, FVar y)\ have "\ = fmupd x (FVar y) ?\'" by simp moreover from Cons.prems(3) and \p = (x, FVar y)\ have "x \ fmdom' ?\'" by simp ultimately have "\ = {x \ FVar y} ++\<^sub>f ?\'" using fmap_singleton_comm by fastforce with Cons.prems(2) and \x \ fmdom' ?\'\ have "is_renaming_substitution ?\'" unfolding is_renaming_substitution_def and \\ = fmupd x (FVar y) ?\'\ using updated_substitution_is_substitution by (metis fmdiff_fmupd fmdom'_notD fmpred_filter) from Cons.prems(2) and \\ = fmupd x (FVar y) ?\'\ have "is_renaming_substitution {x \ FVar y}" by auto have " foldr (\(x, y) C. \<^bold>S {x \ y} C) (p # ps) B = \<^bold>S {x \ FVar y} (foldr (\(x, y) C. \<^bold>S {x \ y} C) ps B)" by (simp add: \p = (x, FVar y)\) also have "\ = \<^bold>S {x \ FVar y} \<^bold>S ?\' B" proof - from Cons.prems(3,4) have "distinct (map fst ps)" and "distinct (map snd ps)" by fastforce+ moreover have "vars (fmran' ?\') \ (fmdom' ?\' \ vars B) = {}" proof - have "vars (fmran' \) = vars ({FVar y} \ fmran' ?\')" using \\ = fmupd x (FVar y) ?\'\ and \x \ fmdom' ?\'\ by (metis fmdom'_notD fmran'_fmupd) then have "vars (fmran' \) = {y} \ vars (fmran' ?\')" using singleton_form_set_vars by auto moreover have "fmdom' \ = {x} \ fmdom' ?\'" by (simp add: \\ = {x \ FVar y} ++\<^sub>f ?\'\) ultimately show ?thesis using Cons.prems(5) by auto qed ultimately show ?thesis using Cons.IH and \is_renaming_substitution ?\'\ by simp qed also have "\ = \<^bold>S ({x \ FVar y} ++\<^sub>f ?\') B" proof (rule substitution_fusion) show "is_substitution ?\'" using \is_renaming_substitution ?\'\ by simp show "is_substitution {x \ FVar y}" using \is_renaming_substitution {x \ FVar y}\ by simp show "?\' $$ x = None" using \x \ fmdom' ?\'\ by blast show "\v' \ fmdom' ?\'. x \ vars (?\' $$! v')" proof - have "x \ fmdom' \" using \\ = {x \ FVar y} ++\<^sub>f ?\'\ by simp then have "x \ vars (fmran' \)" using Cons.prems(5) by blast moreover have "{?\' $$! v' | v'. v' \ fmdom' ?\'} \ fmran' \" unfolding \\ = ?\'(x \ FVar y)\ using \?\' $$ x = None\ by (auto simp add: fmlookup_of_list fmlookup_dom'_iff fmran'I weak_map_of_SomeI) ultimately show ?thesis by force qed qed also from \\ = {x \ FVar y} ++\<^sub>f ?\'\ have "\ = \<^bold>S \ B" by (simp only:) finally show ?case .. qed lemma free_vars_agreement_substitution_equality: assumes "fmdom' \ = fmdom' \'" and "\v \ free_vars A \ fmdom' \. \ $$! v = \' $$! v" shows "\<^bold>S \ A = \<^bold>S \' A" using assms proof (induction A arbitrary: \ \') case (FVar v) have "free_vars (FVar v) = {v}" using surj_pair[of v] by fastforce with FVar have "\ $$! v = \' $$! v" by force with FVar.prems(1) show ?case using surj_pair[of v] by (metis fmdom'_notD fmdom'_notI option.collapse substitute.simps(1)) next case FCon then show ?case by (metis prod.exhaust_sel substitute.simps(2)) next case (FApp B C) have "\<^bold>S \ (B \ C) = (\<^bold>S \ B) \ (\<^bold>S \ C)" by simp also have "\ = (\<^bold>S \' B) \ (\<^bold>S \' C)" proof - have "\v \ free_vars B \ fmdom' \. \ $$! v = \' $$! v" and "\v \ free_vars C \ fmdom' \. \ $$! v = \' $$! v" using FApp.prems(2) by auto with FApp.IH(1,2) and FApp.prems(1) show ?thesis by blast qed finally show ?case by simp next case (FAbs w B) from FAbs.prems(1,2) have *: "\v \ free_vars B - {w} \ fmdom' \. \ $$! v = \' $$! v" using surj_pair[of w] by fastforce show ?case proof (cases "w \ fmdom' \") case True then have "\<^bold>S \ (FAbs w B) = FAbs w (\<^bold>S (fmdrop w \) B)" using surj_pair[of w] by fastforce also have "\ = FAbs w (\<^bold>S (fmdrop w \') B)" proof - from * have "\v \ free_vars B \ fmdom' (fmdrop w \). (fmdrop w \) $$! v = (fmdrop w \') $$! v" by simp moreover have "fmdom' (fmdrop w \) = fmdom' (fmdrop w \')" by (simp add: FAbs.prems(1)) ultimately show ?thesis using FAbs.IH by blast qed finally show ?thesis using FAbs.prems(1) and True and surj_pair[of w] by fastforce next case False then have "\<^bold>S \ (FAbs w B) = FAbs w (\<^bold>S \ B)" using surj_pair[of w] by fastforce also have "\ = FAbs w (\<^bold>S \' B)" proof - from * have "\v \ free_vars B \ fmdom' \. \ $$! v = \' $$! v" using False by blast with FAbs.prems(1) show ?thesis using FAbs.IH by blast qed finally show ?thesis using FAbs.prems(1) and False and surj_pair[of w] by fastforce qed qed text \ The following lemma proves that $ \d{\textsf{S}}\;^{x_\alpha}_{A_\alpha}\; \d{\textsf{S}}\;^{x^1_{\alpha_1}\;\dots\;x^n_{\alpha_n}}_{A^1_{\alpha_1}\;\dots\;A^n_{\alpha_n}} B = \d{\textsf{S}}\;{ \begingroup \setlength\arraycolsep{2pt} \begin{matrix} \scriptstyle{x_\alpha} & \scriptstyle{x^1_{\alpha_1}} & \scriptstyle{\dots} & \scriptstyle{x^n_{\alpha_n}} \\ \scriptstyle{A_\alpha} & \scriptstyle{\d{\small{\textsf{S}}}\;^{x_\alpha}_{A_\alpha} A^1_{\alpha_1}} & \scriptstyle{\dots} & \scriptstyle{\d{\small{\textsf{S}}}\;^{x_\alpha}_{A_\alpha} A^n_{\alpha_n}} \end{matrix} \endgroup} B$ provided that $x_\alpha$ is distinct from $x^1_{\alpha_1}, \dots, x^n_{\alpha_n}$ and $A^i_{\alpha_i}$ is free for $x^i_{\alpha_i}$ in $B$: \ lemma substitution_consolidation: assumes "v \ fmdom' \" and "\v' \ fmdom' \. is_free_for (\ $$! v') v' B" shows "\<^bold>S {v \ A} \<^bold>S \ B = \<^bold>S ({v \ A} ++\<^sub>f fmmap (\A'. \<^bold>S {v \ A} A') \) B" using assms proof (induction B arbitrary: \) case (FApp B C) have "\v' \ fmdom' \. is_free_for (\ $$! v') v' B \ is_free_for (\ $$! v') v' C" proof fix v' assume "v' \ fmdom' \" with FApp.prems(2) have "is_free_for (\ $$! v') v' (B \ C)" by blast then show "is_free_for (\ $$! v') v' B \ is_free_for (\ $$! v') v' C" using is_free_for_from_app by iprover qed with FApp.IH and FApp.prems(1) show ?case by simp next case (FAbs w B) let ?\' = "fmmap (\A'. \<^bold>S {v \ A} A') \" show ?case proof (cases "w \ fmdom' \") case True then have "w \ fmdom' ?\'" by simp with True and FAbs.prems have "v \ w" by blast from True have "\<^bold>S {v \ A} \<^bold>S \ (FAbs w B) = \<^bold>S {v \ A} (FAbs w (\<^bold>S (fmdrop w \) B))" using surj_pair[of w] by fastforce also from \v \ w\ have "\ = FAbs w (\<^bold>S {v \ A} \<^bold>S (fmdrop w \) B)" using surj_pair[of w] by fastforce also have "\ = FAbs w (\<^bold>S (fmdrop w ({v \ A} ++\<^sub>f ?\')) B)" proof - obtain x\<^sub>w and \\<^sub>w where "w = (x\<^sub>w, \\<^sub>w)" by fastforce have "\v' \ fmdom' (fmdrop w \). is_free_for ((fmdrop w \) $$! v') v' B" proof fix v' assume "v' \ fmdom' (fmdrop w \)" with FAbs.prems(2) have "is_free_for (\ $$! v') v' (FAbs w B)" by auto with \w = (x\<^sub>w, \\<^sub>w)\ and \v' \ fmdom' (fmdrop w \)\ have "is_free_for (\ $$! v') v' (\x\<^sub>w\<^bsub>\\<^sub>w\<^esub>. B)" and "v' \ (x\<^sub>w, \\<^sub>w)" by auto then have "is_free_for (\ $$! v') v' B" using is_free_for_from_abs by presburger with \v' \ (x\<^sub>w, \\<^sub>w)\ and \w = (x\<^sub>w, \\<^sub>w)\ show "is_free_for (fmdrop w \ $$! v') v' B" by simp qed moreover have "v \ fmdom' (fmdrop w \)" by (simp add: FAbs.prems(1)) ultimately show ?thesis using FAbs.IH and \v \ w\ by (simp add: fmdrop_fmupd) qed finally show ?thesis using \w \ fmdom' ?\'\ and surj_pair[of w] by fastforce next case False then have "w \ fmdom' ?\'" by simp from FAbs.prems have "v \ fmdom' ?\'" by simp from False have *: "\<^bold>S {v \ A} \<^bold>S \ (FAbs w B) = \<^bold>S {v \ A} (FAbs w (\<^bold>S \ B))" using surj_pair[of w] by fastforce then show ?thesis proof (cases "v \ w") case True then have "\<^bold>S {v \ A} (FAbs w (\<^bold>S \ B)) = FAbs w (\<^bold>S {v \ A} (\<^bold>S \ B))" using surj_pair[of w] by fastforce also have "\ = FAbs w (\<^bold>S ({v \ A} ++\<^sub>f ?\') B)" proof - obtain x\<^sub>w and \\<^sub>w where "w = (x\<^sub>w, \\<^sub>w)" by fastforce have "\v' \ fmdom' \. is_free_for (\ $$! v') v' B" proof fix v' assume "v' \ fmdom' \" with FAbs.prems(2) have "is_free_for (\ $$! v') v' (FAbs w B)" by auto with \w = (x\<^sub>w, \\<^sub>w)\ and \v' \ fmdom' \\ and False have "is_free_for (\ $$! v') v' (\x\<^sub>w\<^bsub>\\<^sub>w\<^esub>. B)" and "v' \ (x\<^sub>w, \\<^sub>w)" by fastforce+ then have "is_free_for (\ $$! v') v' B" using is_free_for_from_abs by presburger with \v' \ (x\<^sub>w, \\<^sub>w)\ and \w = (x\<^sub>w, \\<^sub>w)\ show "is_free_for (\ $$! v') v' B" by simp qed with FAbs.IH show ?thesis using FAbs.prems(1) by blast qed finally show ?thesis proof - assume " \<^bold>S {v \ A} (FAbs w (\<^bold>S \ B)) = FAbs w (\<^bold>S ({v \ A} ++\<^sub>f fmmap (substitute {v \ A}) \) B)" moreover have "w \ fmdom' ({v \ A} ++\<^sub>f fmmap (substitute {v \ A}) \)" using False and True by auto ultimately show ?thesis using * and surj_pair[of w] by fastforce qed next case False then have "v \ free_vars (FAbs w (\<^bold>S \ B))" using surj_pair[of w] by fastforce then have **: "\<^bold>S {v \ A} (FAbs w (\<^bold>S \ B)) = FAbs w (\<^bold>S \ B)" using free_var_singleton_substitution_neutrality by blast also have "\ = FAbs w (\<^bold>S ?\' B)" proof - { fix v' assume "v' \ fmdom' \" with FAbs.prems(1) have "v' \ v" by blast assume "v \ free_vars (\ $$! v')" and "v' \ free_vars B" with \v' \ v\ have "\ is_free_for (\ $$! v') v' (FAbs v B)" using form_with_free_binder_not_free_for by blast with FAbs.prems(2) and \v' \ fmdom' \\ and False have False by blast } then have "\v' \ fmdom' \. v \ free_vars (\ $$! v') \ v' \ free_vars B" by blast then have "\v' \ fmdom' \. v' \ free_vars B \ \<^bold>S {v \ A} (\ $$! v') = \ $$! v'" using free_var_singleton_substitution_neutrality by blast then have "\v' \ free_vars B. \ $$! v' = ?\' $$! v'" by (metis fmdom'_map fmdom'_notD fmdom'_notI fmlookup_map option.map_sel) then have "\<^bold>S \ B = \<^bold>S ?\' B" using free_vars_agreement_substitution_equality by (metis IntD1 fmdom'_map) then show ?thesis by simp qed also from False and FAbs.prems(1) have "\ = FAbs w (\<^bold>S (fmdrop w ({v \ A} ++\<^sub>f ?\')) B)" by (simp add: fmdrop_fmupd_same fmdrop_idle') also from False have "\ = \<^bold>S ({v \ A} ++\<^sub>f ?\') (FAbs w B)" using surj_pair[of w] by fastforce finally show ?thesis using * and ** by (simp only:) qed qed qed force+ lemma vars_range_substitution: assumes "is_substitution \" and "v \ vars (fmran' \)" shows "v \ vars (fmran' (fmdrop w \))" using assms proof (induction \) case fmempty then show ?case by simp next case (fmupd v' A \) from fmdom'_notI[OF fmupd.hyps] and fmupd.prems(1) have "is_substitution \" by (rule updated_substitution_is_substitution) moreover from fmupd.prems(2) and fmupd.hyps have "v \ vars (fmran' \)" by simp ultimately have "v \ vars (fmran' (fmdrop w \))" by (rule fmupd.IH) with fmupd.hyps and fmupd.prems(2) show ?case by (simp add: fmdrop_fmupd) qed lemma excluded_var_from_substitution: assumes "is_substitution \" and "v \ fmdom' \" and "v \ vars (fmran' \)" and "v \ vars A" shows "v \ vars (\<^bold>S \ A)" using assms proof (induction A arbitrary: \) case (FVar v') then show ?case proof (cases "v' \ fmdom' \") case True then have "\ $$! v' \ fmran' \" by (simp add: fmlookup_dom'_iff fmran'I) with FVar(3) have "v \ vars (\ $$! v')" by simp with True show ?thesis using surj_pair[of v'] and fmdom'_notI by force next case False with FVar.prems(4) show ?thesis using surj_pair[of v'] by force qed next case (FCon k) then show ?case using surj_pair[of k] by force next case (FApp B C) then show ?case by auto next case (FAbs w B) have "v \ vars B" and "v \ w" using surj_pair[of w] and FAbs.prems(4) by fastforce+ then show ?case proof (cases "w \ fmdom' \") case True then have "\<^bold>S \ (FAbs w B) = FAbs w (\<^bold>S \ B)" using surj_pair[of w] by fastforce moreover from FAbs.IH have "v \ vars (\<^bold>S \ B)" using FAbs.prems(1-3) and \v \ vars B\ by blast ultimately show ?thesis using \v \ w\ and surj_pair[of w] by fastforce next case False then have "\<^bold>S \ (FAbs w B) = FAbs w (\<^bold>S (fmdrop w \) B)" using surj_pair[of w] by fastforce moreover have "v \ vars (\<^bold>S (fmdrop w \) B)" proof - from FAbs.prems(1) have "is_substitution (fmdrop w \)" by fastforce moreover from FAbs.prems(2) have "v \ fmdom' (fmdrop w \)" by simp moreover from FAbs.prems(1,3) have "v \ vars (fmran' (fmdrop w \))" by (fact vars_range_substitution) ultimately show ?thesis using FAbs.IH and \v \ vars B\ by simp qed ultimately show ?thesis using \v \ w\ and surj_pair[of w] by fastforce qed qed subsection \Renaming of bound variables\ fun rename_bound_var :: "var \ nat \ form \ form" where "rename_bound_var v y (x\<^bsub>\\<^esub>) = x\<^bsub>\\<^esub>" | "rename_bound_var v y (\c\\<^bsub>\\<^esub>) = \c\\<^bsub>\\<^esub>" | "rename_bound_var v y (B \ C) = rename_bound_var v y B \ rename_bound_var v y C" | "rename_bound_var v y (\x\<^bsub>\\<^esub>. B) = ( if (x, \) = v then \y\<^bsub>\\<^esub>. \<^bold>S {(x, \) \ y\<^bsub>\\<^esub>} (rename_bound_var v y B) else \x\<^bsub>\\<^esub>. (rename_bound_var v y B) )" lemma rename_bound_var_preserves_typing: assumes "A \ wffs\<^bsub>\\<^esub>" shows "rename_bound_var (y, \) z A \ wffs\<^bsub>\\<^esub>" using assms proof (induction A) case (abs_is_wff \ A \ x) then show ?case proof (cases "(x, \) = (y, \)") case True from abs_is_wff.IH have "\<^bold>S {(y, \) \ z\<^bsub>\\<^esub>} (rename_bound_var (y, \) z A) \ wffs\<^bsub>\\<^esub>" using substitution_preserves_typing by (simp add: wffs_of_type_intros(1)) then have "\z\<^bsub>\\<^esub>. \<^bold>S {(y, \) \ z\<^bsub>\\<^esub>} (rename_bound_var (y, \) z A) \ wffs\<^bsub>\\\\<^esub>" by blast with True show ?thesis by simp next case False from abs_is_wff.IH have "\x\<^bsub>\\<^esub>. rename_bound_var (y, \) z A \ wffs\<^bsub>\\\\<^esub>" by blast with False show ?thesis by auto qed qed auto lemma old_bound_var_not_free_in_abs_after_renaming: assumes "A \ wffs\<^bsub>\\<^esub>" and "z\<^bsub>\\<^esub> \ y\<^bsub>\\<^esub>" and "(z, \) \ vars A" shows "(y, \) \ free_vars (rename_bound_var (y, \) z (\y\<^bsub>\\<^esub>. A))" using assms and free_var_in_renaming_substitution by (induction A) auto lemma rename_bound_var_free_vars: assumes "A \ wffs\<^bsub>\\<^esub>" and "z\<^bsub>\\<^esub> \ y\<^bsub>\\<^esub>" and "(z, \) \ vars A" shows "(z, \) \ free_vars (rename_bound_var (y, \) z A)" using assms by (induction A) auto lemma old_bound_var_not_free_after_renaming: assumes "A \ wffs\<^bsub>\\<^esub>" and "z\<^bsub>\\<^esub> \ y\<^bsub>\\<^esub>" and "(z, \) \ vars A" and "(y, \) \ free_vars A" shows "(y, \) \ free_vars (rename_bound_var (y, \) z A)" using assms proof induction case (abs_is_wff \ A \ x) then show ?case proof (cases "(x, \) = (y, \)") case True with abs_is_wff.hyps and abs_is_wff.prems(2) show ?thesis using old_bound_var_not_free_in_abs_after_renaming by auto next case False with abs_is_wff.prems(2,3) and assms(2) show ?thesis using abs_is_wff.IH by force qed qed fastforce+ lemma old_bound_var_not_ocurring_after_renaming: assumes "A \ wffs\<^bsub>\\<^esub>" and "z\<^bsub>\\<^esub> \ y\<^bsub>\\<^esub>" shows "\ occurs_at (y, \) p (\<^bold>S {(y, \) \ z\<^bsub>\\<^esub>} (rename_bound_var (y, \) z A))" using assms(1) proof (induction A arbitrary: p) case (var_is_wff \ x) from assms(2) show ?case using subform_size_decrease by (cases "(x, \) = (y, \)") fastforce+ next case (con_is_wff \ c) then show ?case using occurs_at_alt_def(2) by auto next case (app_is_wff \ \ A B) then show ?case proof (cases p) case (Cons d p') then show ?thesis by (cases d) (use app_is_wff.IH in auto) qed simp next case (abs_is_wff \ A \ x) then show ?case proof (cases p) case (Cons d p') then show ?thesis proof (cases d) case Left have *: "\ occurs_at (y, \) p (\x\<^bsub>\\<^esub>. \<^bold>S {(y, \) \ z\<^bsub>\\<^esub>} (rename_bound_var (y, \) z A))" for x and \ using Left and Cons and abs_is_wff.IH by simp then show ?thesis proof (cases "(x, \) = (y, \)") case True with assms(2) have " \<^bold>S {(y, \) \ z\<^bsub>\\<^esub>} (rename_bound_var (y, \) z (\x\<^bsub>\\<^esub>. A)) = \z\<^bsub>\\<^esub>. \<^bold>S {(y, \) \ z\<^bsub>\\<^esub>} (rename_bound_var (y, \) z A)" using free_var_in_renaming_substitution and free_var_singleton_substitution_neutrality by simp moreover have "\ occurs_at (y, \) p (\z\<^bsub>\\<^esub>. \<^bold>S {(y, \) \ z\<^bsub>\\<^esub>} (rename_bound_var (y, \) z A))" using Left and Cons and * by simp ultimately show ?thesis by simp next case False with assms(2) have " \<^bold>S {(y, \) \ z\<^bsub>\\<^esub>} (rename_bound_var (y, \) z (\x\<^bsub>\\<^esub>. A)) = \x\<^bsub>\\<^esub>. \<^bold>S {(y, \) \ z\<^bsub>\\<^esub>} (rename_bound_var (y, \) z A)" by simp moreover have "\ occurs_at (y, \) p (\x\<^bsub>\\<^esub>. \<^bold>S {(y, \) \ z\<^bsub>\\<^esub>} (rename_bound_var (y, \) z A))" using Left and Cons and * by simp ultimately show ?thesis by simp qed qed (simp add: Cons) qed simp qed text \ The following lemma states that the result of \<^term>\rename_bound_var\ does not contain bound occurrences of the renamed variable: \ lemma rename_bound_var_not_bound_occurrences: assumes "A \ wffs\<^bsub>\\<^esub>" and "z\<^bsub>\\<^esub> \ y\<^bsub>\\<^esub>" and "(z, \) \ vars A" and "occurs_at (y, \) p (rename_bound_var (y, \) z A)" shows "\ in_scope_of_abs (z, \) p (rename_bound_var (y, \) z A)" using assms(1,3,4) proof (induction arbitrary: p) case (var_is_wff \ x) then show ?case by (simp add: subforms_from_var(2)) next case (con_is_wff \ c) then show ?case using occurs_at_alt_def(2) by auto next case (app_is_wff \ \ B C) from app_is_wff.prems(1) have "(z, \) \ vars B" and "(z, \) \ vars C" by simp_all from app_is_wff.prems(2) have "occurs_at (y, \) p (rename_bound_var (y, \) z B \ rename_bound_var (y, \) z C)" by simp then consider (a) "\p'. p = \ # p' \ occurs_at (y, \) p' (rename_bound_var (y, \) z B)" | (b) "\p'. p = \ # p' \ occurs_at (y, \) p' (rename_bound_var (y, \) z C)" using subforms_from_app by force then show ?case proof cases case a then obtain p' where "p = \ # p'" and "occurs_at (y, \) p' (rename_bound_var (y, \) z B)" by blast then have "\ in_scope_of_abs (z, \) p' (rename_bound_var (y, \) z B)" using app_is_wff.IH(1)[OF \(z, \) \ vars B\] by blast then have "\ in_scope_of_abs (z, \) p (rename_bound_var (y, \) z (B \ C))" for C using \p = \ # p'\ and in_scope_of_abs_in_left_app by simp then show ?thesis by blast next case b then obtain p' where "p = \ # p'" and "occurs_at (y, \) p' (rename_bound_var (y, \) z C)" by blast then have "\ in_scope_of_abs (z, \) p' (rename_bound_var (y, \) z C)" using app_is_wff.IH(2)[OF \(z, \) \ vars C\] by blast then have "\ in_scope_of_abs (z, \) p (rename_bound_var (y, \) z (B \ C))" for B using \p = \ # p'\ and in_scope_of_abs_in_right_app by simp then show ?thesis by blast qed next case (abs_is_wff \ A \ x) from abs_is_wff.prems(1) have "(z, \) \ vars A" and "(z, \) \ (x, \)" by fastforce+ then show ?case proof (cases "(y, \) = (x, \)") case True then have "occurs_at (y, \) p (\z\<^bsub>\\<^esub>. \<^bold>S {(y, \) \ z\<^bsub>\\<^esub>} (rename_bound_var (y, \) z A))" using abs_is_wff.prems(2) by simp moreover have "\ occurs_at (y, \) p (\z\<^bsub>\\<^esub>. \<^bold>S {(y, \) \ z\<^bsub>\\<^esub>} (rename_bound_var (y, \) z A))" using old_bound_var_not_ocurring_after_renaming[OF abs_is_wff.hyps assms(2)] and subforms_from_abs by fastforce ultimately show ?thesis by contradiction next case False then have *: "rename_bound_var (y, \) z (\x\<^bsub>\\<^esub>. A) = \x\<^bsub>\\<^esub>. rename_bound_var (y, \) z A" by auto with abs_is_wff.prems(2) have "occurs_at (y, \) p (\x\<^bsub>\\<^esub>. rename_bound_var (y, \) z A)" by auto then obtain p' where "p = \ # p'" and "occurs_at (y, \) p' (rename_bound_var (y, \) z A)" using subforms_from_abs by fastforce then have "\ in_scope_of_abs (z, \) p' (rename_bound_var (y, \) z A)" using abs_is_wff.IH[OF \(z, \) \ vars A\] by blast then have "\ in_scope_of_abs (z, \) (\ # p') (\x\<^bsub>\\<^esub>. rename_bound_var (y, \) z A)" using \p = \ # p'\ and in_scope_of_abs_in_abs and \(z, \) \ (x, \)\ by auto then show ?thesis using * and \p = \ # p'\ by simp qed qed lemma is_free_for_in_rename_bound_var: assumes "A \ wffs\<^bsub>\\<^esub>" and "z\<^bsub>\\<^esub> \ y\<^bsub>\\<^esub>" and "(z, \) \ vars A" shows "is_free_for (z\<^bsub>\\<^esub>) (y, \) (rename_bound_var (y, \) z A)" proof (rule ccontr) assume "\ is_free_for (z\<^bsub>\\<^esub>) (y, \) (rename_bound_var (y, \) z A)" then obtain p where "is_free_at (y, \) p (rename_bound_var (y, \) z A)" and "in_scope_of_abs (z, \) p (rename_bound_var (y, \) z A)" by force then show False using rename_bound_var_not_bound_occurrences[OF assms] by fastforce qed lemma renaming_substitution_preserves_bound_vars: shows "bound_vars (\<^bold>S {(y, \) \ z\<^bsub>\\<^esub>} A) = bound_vars A" proof (induction A) case (FAbs v A) then show ?case using singleton_substitution_simps(4) and surj_pair[of v] by (cases "v = (y, \)") (presburger, force) qed force+ lemma rename_bound_var_bound_vars: assumes "A \ wffs\<^bsub>\\<^esub>" and "z\<^bsub>\\<^esub> \ y\<^bsub>\\<^esub>" shows "(y, \) \ bound_vars (rename_bound_var (y, \) z A)" using assms and renaming_substitution_preserves_bound_vars by (induction A) auto lemma old_var_not_free_not_occurring_after_rename: assumes "A \ wffs\<^bsub>\\<^esub>" and "z\<^bsub>\\<^esub> \ y\<^bsub>\\<^esub>" and "(y, \) \ free_vars A" and "(z, \) \ vars A" shows "(y, \) \ vars (rename_bound_var (y, \) z A)" using assms and rename_bound_var_bound_vars[OF assms(1,2)] and old_bound_var_not_free_after_renaming and vars_is_free_and_bound_vars by blast end