diff --git a/thys/Q0_Metatheory/Boolean_Algebra.thy b/thys/Q0_Metatheory/Boolean_Algebra.thy new file mode 100644 --- /dev/null +++ b/thys/Q0_Metatheory/Boolean_Algebra.thy @@ -0,0 +1,115 @@ +section \Boolean Algebra\ + +theory Boolean_Algebra + imports + "ZFC_in_HOL.ZFC_Typeclasses" +begin + +text \This theory contains an embedding of two-valued boolean algebra into \<^term>\V\.\ + +hide_const (open) List.set + +definition bool_to_V :: "bool \ V" where + "bool_to_V = (SOME f. inj f)" + +lemma bool_to_V_injectivity [simp]: + shows "inj bool_to_V" + unfolding bool_to_V_def by (fact someI_ex[OF embeddable_class.ex_inj]) + +definition bool_from_V :: "V \ bool" where + [simp]: "bool_from_V = inv bool_to_V" + +definition top :: V ("\<^bold>T") where + [simp]: "\<^bold>T = bool_to_V True" + +definition bottom :: V ("\<^bold>F") where + [simp]: "\<^bold>F = bool_to_V False" + +definition two_valued_boolean_algebra_universe :: V ("\") where + [simp]: "\ = set {\<^bold>T, \<^bold>F}" + +definition negation :: "V \ V" ("\ _" [141] 141) where + [simp]: "\ p = bool_to_V (\ bool_from_V p)" + +definition conjunction :: "V \ V \ V" (infixr "\<^bold>\" 136) where + [simp]: "p \<^bold>\ q = bool_to_V (bool_from_V p \ bool_from_V q)" + +definition disjunction :: "V \ V \ V" (infixr "\<^bold>\" 131) where + [simp]: "p \<^bold>\ q = \ (\ p \<^bold>\ \ q)" + +definition implication :: "V \ V \ V" (infixr "\<^bold>\" 121) where + [simp]: "p \<^bold>\ q = \ p \<^bold>\ q" + +definition iff :: "V \ V \ V" (infixl "\<^bold>\" 150) where + [simp]: "p \<^bold>\ q = (p \<^bold>\ q) \<^bold>\ (q \<^bold>\ p)" + +lemma boolean_algebra_simps [simp]: + assumes "p \ elts \" and "q \ elts \" and "r \ elts \" + shows "\ \ p = p" + and "((\ p) \<^bold>\ (\ q)) = (p \<^bold>\ q)" + and "\ (p \<^bold>\ q) = (p \<^bold>\ (\ q))" + and "(p \<^bold>\ \ p) = \<^bold>T" + and "(\ p \<^bold>\ p) = \<^bold>T" + and "(p \<^bold>\ p) = \<^bold>T" + and "(\ p) \ p" + and "p \ (\ p)" + and "(\<^bold>T \<^bold>\ p) = p" + and "(p \<^bold>\ \<^bold>T) = p" + and "(\<^bold>F \<^bold>\ p) = (\ p)" + and "(p \<^bold>\ \<^bold>F) = (\ p)" + and "(\<^bold>T \<^bold>\ p) = p" + and "(\<^bold>F \<^bold>\ p) = \<^bold>T" + and "(p \<^bold>\ \<^bold>T) = \<^bold>T" + and "(p \<^bold>\ p) = \<^bold>T" + and "(p \<^bold>\ \<^bold>F) = (\ p)" + and "(p \<^bold>\ \ p) = (\ p)" + and "(p \<^bold>\ \<^bold>T) = p" + and "(\<^bold>T \<^bold>\ p) = p" + and "(p \<^bold>\ \<^bold>F) = \<^bold>F" + and "(\<^bold>F \<^bold>\ p) = \<^bold>F" + and "(p \<^bold>\ p) = p" + and "(p \<^bold>\ (p \<^bold>\ q)) = (p \<^bold>\ q)" + and "(p \<^bold>\ \ p) = \<^bold>F" + and "(\ p \<^bold>\ p) = \<^bold>F" + and "(p \<^bold>\ \<^bold>T) = \<^bold>T" + and "(\<^bold>T \<^bold>\ p) = \<^bold>T" + and "(p \<^bold>\ \<^bold>F) = p" + and "(\<^bold>F \<^bold>\ p) = p" + and "(p \<^bold>\ p) = p" + and "(p \<^bold>\ (p \<^bold>\ q)) = (p \<^bold>\ q)" + and "p \<^bold>\ q = q \<^bold>\ p" + and "p \<^bold>\ (q \<^bold>\ r) = q \<^bold>\ (p \<^bold>\ r)" + and "p \<^bold>\ q = q \<^bold>\ p" + and "p \<^bold>\ (q \<^bold>\ r) = q \<^bold>\ (p \<^bold>\ r)" + and "(p \<^bold>\ q) \<^bold>\ r = p \<^bold>\ (q \<^bold>\ r)" + and "p \<^bold>\ (q \<^bold>\ r) = p \<^bold>\ q \<^bold>\ p \<^bold>\ r" + and "(p \<^bold>\ q) \<^bold>\ r = p \<^bold>\ r \<^bold>\ q \<^bold>\ r" + and "p \<^bold>\ (q \<^bold>\ r) = (p \<^bold>\ q) \<^bold>\ (p \<^bold>\ r)" + and "(p \<^bold>\ q) \<^bold>\ r = (p \<^bold>\ r) \<^bold>\ (q \<^bold>\ r)" + and "(p \<^bold>\ (q \<^bold>\ r)) = ((p \<^bold>\ q) \<^bold>\ (p \<^bold>\ r))" + and "((p \<^bold>\ q) \<^bold>\ r) = (p \<^bold>\ (q \<^bold>\ r))" + and "((p \<^bold>\ q) \<^bold>\ r) = ((p \<^bold>\ r) \<^bold>\ (q \<^bold>\ r))" + and "((p \<^bold>\ q) \<^bold>\ r) = (p \<^bold>\ q \<^bold>\ r)" + and "(q \<^bold>\ (p \<^bold>\ r)) = (p \<^bold>\ q \<^bold>\ r)" + and "\ (p \<^bold>\ q) = \ p \<^bold>\ \ q" + and "\ (p \<^bold>\ q) = \ p \<^bold>\ \ q" + and "\ (p \<^bold>\ q) = p \<^bold>\ \ q" + and "\ p \<^bold>\ q = (p \<^bold>\ q)" + and "p \<^bold>\ \ q = (q \<^bold>\ p)" + and "(p \<^bold>\ q) = (\ p) \<^bold>\ q" + and "p \<^bold>\ q = \ p \<^bold>\ q" + and "(p \<^bold>\ q) = (p \<^bold>\ q) \<^bold>\ (q \<^bold>\ p)" + and "(p \<^bold>\ q) \<^bold>\ (\ p \<^bold>\ q) = q" + and "p = \<^bold>T \ \ (p = \<^bold>F)" + and "p = \<^bold>F \ \ (p = \<^bold>T)" + and "p = \<^bold>T \ p = \<^bold>F" + using assms by (auto simp add: inj_eq) + +lemma tv_cases [consumes 1, case_names top bottom, cases type: V]: + assumes "p \ elts \" + and "p = \<^bold>T \ P" + and "p = \<^bold>F \ P" + shows P + using assms by auto + +end diff --git a/thys/Q0_Metatheory/Consistency.thy b/thys/Q0_Metatheory/Consistency.thy new file mode 100644 --- /dev/null +++ b/thys/Q0_Metatheory/Consistency.thy @@ -0,0 +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 + 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 new file mode 100644 --- /dev/null +++ b/thys/Q0_Metatheory/Elementary_Logic.thy @@ -0,0 +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 + 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"}: +\ + +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"}: +\ + +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"}: +\ + +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/Proof_System.thy b/thys/Q0_Metatheory/Proof_System.thy new file mode 100644 --- /dev/null +++ b/thys/Q0_Metatheory/Proof_System.thy @@ -0,0 +1,1114 @@ +section \Proof System\ + +theory Proof_System + imports + Syntax +begin + +subsection \Axioms\ + +inductive_set + axioms :: "form set" +where + axiom_1: + "\\<^bsub>o\o\<^esub> \ T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\o\<^esub> \ F\<^bsub>o\<^esub> \\<^sup>\ \\\<^bsub>o\<^esub>. \\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub> \ axioms" +| axiom_2: + "(\\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\<^esub>) \\<^sup>\ (\\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub> \\<^sup>\ \\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub>) \ axioms" +| axiom_3: + "(\\<^bsub>\\\\<^esub> =\<^bsub>\\\\<^esub> \\<^bsub>\\\\<^esub>) \\<^sup>\ \\\<^bsub>\\<^esub>. (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>) \ axioms" +| axiom_4_1_con: + "(\x\<^bsub>\\<^esub>. \c\\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> \c\\<^bsub>\\<^esub> \ axioms" if "A \ wffs\<^bsub>\\<^esub>" +| axiom_4_1_var: + "(\x\<^bsub>\\<^esub>. y\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> y\<^bsub>\\<^esub> \ axioms" if "A \ wffs\<^bsub>\\<^esub>" and "y\<^bsub>\\<^esub> \ x\<^bsub>\\<^esub>" +| axiom_4_2: + "(\x\<^bsub>\\<^esub>. x\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> A \ axioms" if "A \ wffs\<^bsub>\\<^esub>" +| axiom_4_3: + "(\x\<^bsub>\\<^esub>. B \ C) \ A =\<^bsub>\\<^esub> ((\x\<^bsub>\\<^esub>. B) \ A) \ ((\x\<^bsub>\\<^esub>. C) \ A) \ axioms" + if "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\\\<^esub>" and "C \ wffs\<^bsub>\\<^esub>" +| axiom_4_4: + "(\x\<^bsub>\\<^esub>. \y\<^bsub>\\<^esub>. B) \ A =\<^bsub>\\\\<^esub> (\y\<^bsub>\\<^esub>. (\x\<^bsub>\\<^esub>. B) \ A) \ axioms" + if "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" and "(y, \) \ {(x, \)} \ vars A" +| axiom_4_5: + "(\x\<^bsub>\\<^esub>. \x\<^bsub>\\<^esub>. B) \ A =\<^bsub>\\\\<^esub> (\x\<^bsub>\\<^esub>. B) \ axioms" if "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" +| axiom_5: + "\ \ (Q\<^bsub>i\<^esub> \ \\<^bsub>i\<^esub>) =\<^bsub>i\<^esub> \\<^bsub>i\<^esub> \ axioms" + +lemma axioms_are_wffs_of_type_o: + shows "axioms \ wffs\<^bsub>o\<^esub>" + by (intro subsetI, cases rule: axioms.cases) auto + +subsection \Inference rule R\ + +definition is_rule_R_app :: "position \ form \ form \ form \ bool" where + [iff]: "is_rule_R_app p D C E \ + ( + \\ A B. + E = A =\<^bsub>\\<^esub> B \ A \ wffs\<^bsub>\\<^esub> \ B \ wffs\<^bsub>\\<^esub> \ \ \\<^term>\E\ is a well-formed equality\ + A \\<^bsub>p\<^esub> C \ + D \ wffs\<^bsub>o\<^esub> \ + C\p \ B\ \ D + )" + +lemma rule_R_original_form_is_wffo: + assumes "is_rule_R_app p D C E" + shows "C \ wffs\<^bsub>o\<^esub>" + using assms and replacement_preserves_typing by fastforce + +subsection \Proof and derivability\ + +inductive is_derivable :: "form \ bool" where + dv_axiom: "is_derivable A" if "A \ axioms" +| dv_rule_R: "is_derivable D" if "is_derivable C" and "is_derivable E" and "is_rule_R_app p D C E" + +lemma derivable_form_is_wffso: + assumes "is_derivable A" + shows "A \ wffs\<^bsub>o\<^esub>" + using assms and axioms_are_wffs_of_type_o by (fastforce elim: is_derivable.cases) + +definition is_proof_step :: "form list \ nat \ bool" where + [iff]: "is_proof_step \ i' \ + \ ! i' \ axioms \ + (\p j k. {j, k} \ {0.. is_rule_R_app p (\ ! i') (\ ! j) (\ ! k))" + +definition is_proof :: "form list \ bool" where + [iff]: "is_proof \ \ (\i' < length \. is_proof_step \ i')" + +lemma common_prefix_is_subproof: + assumes "is_proof (\ @ \\<^sub>1)" + and "i' < length \" + shows "is_proof_step (\ @ \\<^sub>2) i'" +proof - + from assms(2) have *: "(\ @ \\<^sub>1) ! i' = (\ @ \\<^sub>2) ! i'" + by (simp add: nth_append) + moreover from assms(2) have "i' < length (\ @ \\<^sub>1)" + by simp + ultimately obtain p and j and k where **: + "(\ @ \\<^sub>1) ! i' \ axioms \ + {j, k} \ {0.. is_rule_R_app p ((\ @ \\<^sub>1) ! i') ((\ @ \\<^sub>1) ! j) ((\ @ \\<^sub>1) ! k)" + using assms(1) by fastforce + then consider + (axiom) "(\ @ \\<^sub>1) ! i' \ axioms" + | (rule_R) "{j, k} \ {0.. is_rule_R_app p ((\ @ \\<^sub>1) ! i') ((\ @ \\<^sub>1) ! j) ((\ @ \\<^sub>1) ! k)" + by blast + then have " + (\ @ \\<^sub>2) ! i' \ axioms \ + ({j, k} \ {0.. is_rule_R_app p ((\ @ \\<^sub>2) ! i') ((\ @ \\<^sub>2) ! j) ((\ @ \\<^sub>2) ! k))" + proof cases + case axiom + with * have "(\ @ \\<^sub>2) ! i' \ axioms" + by (simp only:) + then show ?thesis .. + next + case rule_R + with assms(2) have "(\ @ \\<^sub>1) ! j = (\ @ \\<^sub>2) ! j" and "(\ @ \\<^sub>1) ! k = (\ @ \\<^sub>2) ! k" + by (simp_all add: nth_append) + then have "{j, k} \ {0.. is_rule_R_app p ((\ @ \\<^sub>2) ! i') ((\ @ \\<^sub>2) ! j) ((\ @ \\<^sub>2) ! k)" + using * and rule_R by simp + then show ?thesis .. + qed + with ** show ?thesis + by fastforce +qed + +lemma added_suffix_proof_preservation: + assumes "is_proof \" + and "i' < length (\ @ \') - length \'" + shows "is_proof_step (\ @ \') i'" + using assms and common_prefix_is_subproof[where \\<^sub>1 = "[]"] by simp + +lemma append_proof_step_is_proof: + assumes "is_proof \" + and "is_proof_step (\ @ [A]) (length (\ @ [A]) - 1)" + shows "is_proof (\ @ [A])" + using assms and added_suffix_proof_preservation by (simp add: All_less_Suc) + +lemma added_prefix_proof_preservation: + assumes "is_proof \'" + and "i' \ {length \.. @ \')}" + shows "is_proof_step (\ @ \') i'" +proof - + let ?\ = "\ @ \'" + let ?i = "i' - length \" + from assms(2) have "?\ ! i' = \' ! ?i" and "?i < length \'" + by (simp_all add: nth_append less_diff_conv2) + then have "is_proof_step ?\ i' = is_proof_step \' ?i" + proof - + from assms(1) and \?i < length \'\ obtain j and k and p where *: + "\' ! ?i \ axioms \ ({j, k} \ {0.. is_rule_R_app p (\' ! ?i) (\' ! j) (\' ! k))" + by fastforce + then consider + (axiom) "\' ! ?i \ axioms" + | (rule_R) "{j, k} \ {0.. is_rule_R_app p (\' ! ?i) (\' ! j) (\' ! k)" + by blast + then have " + ?\ ! i' \ axioms \ + ( + {j + length \, k + length \} \ {0.. + is_rule_R_app p (?\ ! i') (?\ ! (j + length \)) (?\ ! (k + length \)) + )" + proof cases + case axiom + with \?\ ! i' = \' ! ?i\ have "?\ ! i' \ axioms" + by (simp only:) + then show ?thesis .. + next + case rule_R + with assms(2) have "?\ ! (j + length \) = \' ! j" and "?\ ! (k + length \) = \' ! k" + by (simp_all add: nth_append) + with \?\ ! i' = \' ! ?i\ and rule_R have " + {j + length \, k + length \} \ {0.. + is_rule_R_app p (?\ ! i') (?\ ! (j + length \)) (?\ ! (k + length \))" + by auto + then show ?thesis .. + qed + with * show ?thesis + by fastforce + qed + with assms(1) and \?i < length \'\ show ?thesis + by simp +qed + +lemma proof_but_last_is_proof: + assumes "is_proof (\ @ [A])" + shows "is_proof \" + using assms and common_prefix_is_subproof[where \\<^sub>1 = "[A]" and \\<^sub>2 = "[]"] by simp + +lemma proof_prefix_is_proof: + assumes "is_proof (\\<^sub>1 @ \\<^sub>2)" + shows "is_proof \\<^sub>1" + using assms and proof_but_last_is_proof + by (induction \\<^sub>2 arbitrary: \\<^sub>1 rule: rev_induct) (simp, metis append.assoc) + +lemma single_axiom_is_proof: + assumes "A \ axioms" + shows "is_proof [A]" + using assms by fastforce + +lemma proofs_concatenation_is_proof: + assumes "is_proof \\<^sub>1" and "is_proof \\<^sub>2" + shows "is_proof (\\<^sub>1 @ \\<^sub>2)" +proof - + from assms(1) have "\i' < length \\<^sub>1. is_proof_step (\\<^sub>1 @ \\<^sub>2) i'" + using added_suffix_proof_preservation by auto + moreover from assms(2) have "\i' \ {length \\<^sub>1..\<^sub>1 @ \\<^sub>2)}. is_proof_step (\\<^sub>1 @ \\<^sub>2) i'" + using added_prefix_proof_preservation by auto + ultimately show ?thesis + unfolding is_proof_def by (meson atLeastLessThan_iff linorder_not_le) +qed + +lemma elem_of_proof_is_wffo: + assumes "is_proof \" and "A \ lset \" + shows "A \ wffs\<^bsub>o\<^esub>" + using assms and axioms_are_wffs_of_type_o + unfolding is_rule_R_app_def and is_proof_step_def and is_proof_def + by (induction \) (simp, metis (full_types) in_mono in_set_conv_nth) + +lemma axiom_prepended_to_proof_is_proof: + assumes "is_proof \" + and "A \ axioms" + shows "is_proof ([A] @ \)" + using proofs_concatenation_is_proof[OF single_axiom_is_proof[OF assms(2)] assms(1)] . + +lemma axiom_appended_to_proof_is_proof: + assumes "is_proof \" + and "A \ axioms" + shows "is_proof (\ @ [A])" + using proofs_concatenation_is_proof[OF assms(1) single_axiom_is_proof[OF assms(2)]] . + +lemma rule_R_app_appended_to_proof_is_proof: + assumes "is_proof \" + and "i\<^sub>C < length \" and "\ ! i\<^sub>C = C" + and "i\<^sub>E < length \" and "\ ! i\<^sub>E = E" + and "is_rule_R_app p D C E" + shows "is_proof (\ @ [D])" +proof - + let ?\ = "\ @ [D]" + let ?i\<^sub>D = "length \" + from assms(2,4) have "i\<^sub>C < ?i\<^sub>D" and "i\<^sub>E < ?i\<^sub>D" + by fastforce+ + with assms(3,5,6) have "is_rule_R_app p (?\ ! ?i\<^sub>D) (?\ ! i\<^sub>C) (?\ ! i\<^sub>E)" + by (simp add: nth_append) + with assms(2,4) have "\p j k. {j, k} \ {0..D} \ is_rule_R_app p (?\ ! ?i\<^sub>D) (?\ ! j) (?\ ! k)" + by fastforce + then have "is_proof_step ?\ (length ?\ - 1)" + by simp + moreover from assms(1) have "\i' < length ?\ - 1. is_proof_step ?\ i'" + using added_suffix_proof_preservation by auto + ultimately show ?thesis + using less_Suc_eq by auto +qed + +definition is_proof_of :: "form list \ form \ bool" where + [iff]: "is_proof_of \ A \ \ \ [] \ is_proof \ \ last \ = A" + +lemma proof_prefix_is_proof_of_last: + assumes "is_proof (\ @ \')" and "\ \ []" + shows "is_proof_of \ (last \)" +proof - + from assms(1) have "is_proof \" + by (fact proof_prefix_is_proof) + with assms(2) show ?thesis + by fastforce +qed + +definition is_theorem :: "form \ bool" where + [iff]: "is_theorem A \ (\\. is_proof_of \ A)" + +lemma proof_form_is_wffo: + assumes "is_proof_of \ A" + and "B \ lset \" + shows "B \ wffs\<^bsub>o\<^esub>" + using assms and elem_of_proof_is_wffo by blast + +lemma proof_form_is_theorem: + assumes "is_proof \" and "\ \ []" + and "i' < length \" + shows "is_theorem (\ ! i')" +proof - + let ?\\<^sub>1 = "take (Suc i') \" + from assms(1) obtain \\<^sub>2 where "is_proof (?\\<^sub>1 @ \\<^sub>2)" + by (metis append_take_drop_id) + then have "is_proof ?\\<^sub>1" + by (fact proof_prefix_is_proof) + moreover from assms(3) have "last ?\\<^sub>1 = \ ! i'" + by (simp add: take_Suc_conv_app_nth) + ultimately show ?thesis + using assms(2) unfolding is_proof_of_def and is_theorem_def by (metis Zero_neq_Suc take_eq_Nil2) +qed + +theorem derivable_form_is_theorem: + assumes "is_derivable A" + shows "is_theorem A" +using assms proof (induction rule: is_derivable.induct) + case (dv_axiom A) + then have "is_proof [A]" + by (fact single_axiom_is_proof) + moreover have "last [A] = A" + by simp + ultimately show ?case + by blast +next + case (dv_rule_R C E p D) + obtain \\<^sub>C and \\<^sub>E where + "is_proof \\<^sub>C" and "\\<^sub>C \ []" and "last \\<^sub>C = C" and + "is_proof \\<^sub>E" and "\\<^sub>E \ []" and "last \\<^sub>E = E" + using dv_rule_R.IH by fastforce + let ?i\<^sub>C = "length \\<^sub>C - 1" and ?i\<^sub>E = "length \\<^sub>C + length \\<^sub>E - 1" and ?i\<^sub>D = "length \\<^sub>C + length \\<^sub>E" + let ?\ = "\\<^sub>C @ \\<^sub>E @ [D]" + from \\\<^sub>C \ []\ have "?i\<^sub>C < length (\\<^sub>C @ \\<^sub>E)" and "?i\<^sub>E < length (\\<^sub>C @ \\<^sub>E)" + using linorder_not_le by fastforce+ + moreover have "(\\<^sub>C @ \\<^sub>E) ! ?i\<^sub>C = C" and "(\\<^sub>C @ \\<^sub>E) ! ?i\<^sub>E = E" + using \\\<^sub>C \ []\ and \last \\<^sub>C = C\ + by + ( + simp add: last_conv_nth nth_append, + metis \last \\<^sub>E = E\ \\\<^sub>E \ []\ append_is_Nil_conv last_appendR last_conv_nth length_append + ) + with \is_rule_R_app p D C E\ have "is_rule_R_app p D ((\\<^sub>C @ \\<^sub>E) ! ?i\<^sub>C) ((\\<^sub>C @ \\<^sub>E) ! ?i\<^sub>E)" + using \(\\<^sub>C @ \\<^sub>E) ! ?i\<^sub>C = C\ by fastforce + moreover from \is_proof \\<^sub>C\ and \is_proof \\<^sub>E\ have "is_proof (\\<^sub>C @ \\<^sub>E)" + by (fact proofs_concatenation_is_proof) + ultimately have "is_proof ((\\<^sub>C @ \\<^sub>E) @ [D])" + using rule_R_app_appended_to_proof_is_proof by presburger + with \\\<^sub>C \ []\ show ?case + unfolding is_proof_of_def and is_theorem_def by (metis snoc_eq_iff_butlast) +qed + +theorem theorem_is_derivable_form: + assumes "is_theorem A" + shows "is_derivable A" +proof - + from assms obtain \ where "is_proof \" and "\ \ []" and "last \ = A" + by fastforce + then show ?thesis + proof (induction "length \" arbitrary: \ A rule: less_induct) + case less + let ?i' = "length \ - 1" + from \\ \ []\ and \last \ = A\ have "\ ! ?i' = A" + by (simp add: last_conv_nth) + from \is_proof \\ and \\ \ []\ and \last \ = A\ have "is_proof_step \ ?i'" + using added_suffix_proof_preservation[where \' = "[]"] by simp + then consider + (axiom) "\ ! ?i' \ axioms" + | (rule_R) "\p j k. {j, k} \ {0.. is_rule_R_app p (\ ! ?i') (\ ! j) (\ ! k)" + by fastforce + then show ?case + proof cases + case axiom + with \\ ! ?i' = A\ show ?thesis + by (fastforce intro: dv_axiom) + next + case rule_R + then obtain p and j and k + where "{j, k} \ {0.. ! ?i') (\ ! j) (\ ! k)" + by force + let ?\\<^sub>j = "take (Suc j) \" + let ?\\<^sub>k = "take (Suc k) \" + obtain \\<^sub>j' and \\<^sub>k' where "\ = ?\\<^sub>j @ \\<^sub>j'" and "\ = ?\\<^sub>k @ \\<^sub>k'" + by (metis append_take_drop_id) + with \is_proof \\ have "is_proof (?\\<^sub>j @ \\<^sub>j')" and "is_proof (?\\<^sub>k @ \\<^sub>k')" + by (simp_all only:) + moreover + from \\ = ?\\<^sub>j @ \\<^sub>j'\ and \\ = ?\\<^sub>k @ \\<^sub>k'\ and \last \ = A\ and \{j, k} \ {0.. - 1}\ + have "last \\<^sub>j' = A" and "last \\<^sub>k' = A" + using length_Cons and less_le_not_le and take_Suc and take_tl and append.right_neutral + by (metis atLeastLessThan_iff diff_Suc_1 insert_subset last_appendR take_all_iff)+ + moreover from \\ \ []\ have "?\\<^sub>j \ []" and "?\\<^sub>k \ []" + by simp_all + ultimately have "is_proof_of ?\\<^sub>j (last ?\\<^sub>j)" and "is_proof_of ?\\<^sub>k (last ?\\<^sub>k)" + using proof_prefix_is_proof_of_last [where \ = ?\\<^sub>j and \' = \\<^sub>j'] + and proof_prefix_is_proof_of_last [where \ = ?\\<^sub>k and \' = \\<^sub>k'] + by fastforce+ + moreover from \last \\<^sub>j' = A\ and \last \\<^sub>k' = A\ + have "length ?\\<^sub>j < length \" and "length ?\\<^sub>k < length \" + using \{j, k} \ {0.. - 1}\ by force+ + moreover from calculation(3,4) have "last ?\\<^sub>j = \ ! j" and "last ?\\<^sub>k = \ ! k" + by (metis Suc_lessD last_snoc linorder_not_le nat_neq_iff take_Suc_conv_app_nth take_all_iff)+ + ultimately have "is_derivable (\ ! j)" and "is_derivable (\ ! k)" + using \?\\<^sub>j \ []\ and \?\\<^sub>k \ []\ and less(1) by blast+ + with \is_rule_R_app p (\ ! ?i') (\ ! j) (\ ! k)\ and \\ ! ?i' = A\ show ?thesis + by (blast intro: dv_rule_R) + qed + qed +qed + +theorem theoremhood_derivability_equivalence: + shows "is_theorem A \ is_derivable A" + using derivable_form_is_theorem and theorem_is_derivable_form by blast + +lemma theorem_is_wffo: + assumes "is_theorem A" + shows "A \ wffs\<^bsub>o\<^esub>" +proof - + from assms obtain \ where "is_proof_of \ A" + by blast + then have "A \ lset \" + by auto + with \is_proof_of \ A\ show ?thesis + using proof_form_is_wffo by blast +qed + +lemma equality_reflexivity: + assumes "A \ wffs\<^bsub>\\<^esub>" + shows "is_theorem (A =\<^bsub>\\<^esub> A)" (is "is_theorem ?A\<^sub>2") +proof - + let ?A\<^sub>1 = "(\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> A" + let ?\ = "[?A\<^sub>1, ?A\<^sub>2]" + \ \ (.1) Axiom 4.2 \ + have "is_proof_step ?\ 0" + proof - + from assms have "?A\<^sub>1 \ axioms" + by (intro axiom_4_2) + then show ?thesis + by simp + qed + \ \ (.2) Rule R: .1,.1 \ + moreover have "is_proof_step ?\ 1" + proof - + let ?p = "[\, \]" + have "\p j k. {j::nat, k} \ {0..<1} \ is_rule_R_app ?p ?A\<^sub>2 (?\ ! j) (?\ ! k)" + proof - + let ?D = "?A\<^sub>2" and ?j = "0::nat" and ?k = "0" + have "{?j, ?k} \ {0..<1}" + by simp + moreover have "is_rule_R_app ?p ?A\<^sub>2 (?\ ! ?j) (?\ ! ?k)" + proof - + have "(\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>) \ A \\<^bsub>?p\<^esub> (?\ ! ?j)" + by force + moreover have "(?\ ! ?j)\?p \ A\ \ ?D" + by force + moreover from \A \ wffs\<^bsub>\\<^esub>\ have "?D \ wffs\<^bsub>o\<^esub>" + by (intro equality_wff) + moreover from \A \ wffs\<^bsub>\\<^esub>\ have "(\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>) \ A \ wffs\<^bsub>\\<^esub>" + by (meson wffs_of_type_simps) + ultimately show ?thesis + using \A \ wffs\<^bsub>\\<^esub>\ by simp + qed + ultimately show ?thesis + by meson + qed + then show ?thesis + by auto + qed + moreover have "last ?\ = ?A\<^sub>2" + by simp + moreover have "{0..} = {0, 1}" + by (simp add: atLeast0_lessThan_Suc insert_commute) + ultimately show ?thesis + unfolding is_theorem_def and is_proof_def and is_proof_of_def + by (metis One_nat_def Suc_1 length_Cons less_2_cases list.distinct(1) list.size(3)) +qed + +lemma equality_reflexivity': + assumes "A \ wffs\<^bsub>\\<^esub>" + shows "is_theorem (A =\<^bsub>\\<^esub> A)" (is "is_theorem ?A\<^sub>2") +proof (intro derivable_form_is_theorem) + let ?A\<^sub>1 = "(\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> A" + \ \ (.1) Axiom 4.2 \ + from assms have "?A\<^sub>1 \ axioms" + by (intro axiom_4_2) + then have step_1: "is_derivable ?A\<^sub>1" + by (intro dv_axiom) + \ \ (.2) Rule R: .1,.1 \ + then show "is_derivable ?A\<^sub>2" + proof - + let ?p = "[\, \]" and ?C = "?A\<^sub>1" and ?E = "?A\<^sub>1" and ?D = "?A\<^sub>2" + have "is_rule_R_app ?p ?D ?C ?E" + proof - + have "(\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>) \ A \\<^bsub>?p\<^esub> ?C" + by force + moreover have "?C\?p \ A\ \ ?D" + by force + moreover from \A \ wffs\<^bsub>\\<^esub>\ have "?D \ wffs\<^bsub>o\<^esub>" + by (intro equality_wff) + moreover from \A \ wffs\<^bsub>\\<^esub>\ have "(\\\<^bsub>\\<^esub>. \\<^bsub>\\<^esub>) \ A \ wffs\<^bsub>\\<^esub>" + by (meson wffs_of_type_simps) + ultimately show ?thesis + using \A \ wffs\<^bsub>\\<^esub>\ by simp + qed + with step_1 show ?thesis + by (blast intro: dv_rule_R) + qed +qed + +subsection \Hypothetical proof and derivability\ + +text \The set of free variables in \\\ that are exposed to capture at position \p\ in \A\:\ + +definition capture_exposed_vars_at :: "position \ form \ 'a \ var set" where + [simp]: "capture_exposed_vars_at p A \ = + {(x, \) | x \ p' E. strict_prefix p' p \ \x\<^bsub>\\<^esub>. E \\<^bsub>p'\<^esub> A \ (x, \) \ free_vars \}" + +lemma capture_exposed_vars_at_alt_def: + assumes "p \ positions A" + shows "capture_exposed_vars_at p A \ = binders_at A p \ free_vars \" + unfolding binders_at_alt_def[OF assms] and in_scope_of_abs_alt_def + using is_subform_implies_in_positions by auto + +text \Inference rule R$'$:\ + +definition rule_R'_side_condition :: "form set \ position \ form \ form \ form \ bool" where + [iff]: "rule_R'_side_condition \ p D C E \ + capture_exposed_vars_at p C E \ capture_exposed_vars_at p C \ = {}" + +lemma rule_R'_side_condition_alt_def: + fixes \ :: "form set" + assumes "C \ wffs\<^bsub>\\<^esub>" + shows " + rule_R'_side_condition \ p D C (A =\<^bsub>\\<^esub> B) + \ + ( + \x \ E p'. + strict_prefix p' p \ + \x\<^bsub>\\<^esub>. E \\<^bsub>p'\<^esub> C \ + (x, \) \ free_vars (A =\<^bsub>\\<^esub> B) \ + (\H \ \. (x, \) \ free_vars H) + )" +proof - + have " + capture_exposed_vars_at p C (A =\<^bsub>\\<^esub> B) + = + {(x, \) | x \ p' E. strict_prefix p' p \ \x\<^bsub>\\<^esub>. E \\<^bsub>p'\<^esub> C \ (x, \) \ free_vars (A =\<^bsub>\\<^esub> B)}" + using assms and capture_exposed_vars_at_alt_def unfolding capture_exposed_vars_at_def by fast + moreover have " + capture_exposed_vars_at p C \ + = + {(x, \) | x \ p' E. strict_prefix p' p \ \x\<^bsub>\\<^esub>. E \\<^bsub>p'\<^esub> C \ (x, \) \ free_vars \}" + using assms and capture_exposed_vars_at_alt_def unfolding capture_exposed_vars_at_def by fast + ultimately have " + capture_exposed_vars_at p C (A =\<^bsub>\\<^esub> B) \ capture_exposed_vars_at p C \ + = + {(x, \) | x \ p' E. strict_prefix p' p \ \x\<^bsub>\\<^esub>. E \\<^bsub>p'\<^esub> C \ (x, \) \ free_vars (A =\<^bsub>\\<^esub> B) \ + (x, \) \ free_vars \}" + by auto + also have " + \ + = + {(x, \) | x \ p' E. strict_prefix p' p \ \x\<^bsub>\\<^esub>. E \\<^bsub>p'\<^esub> C \ (x, \) \ free_vars (A =\<^bsub>\\<^esub> B) \ + (\H \ \. (x, \) \ free_vars H)}" + by auto + finally show ?thesis + by fast +qed + +definition is_rule_R'_app :: "form set \ position \ form \ form \ form \ bool" where + [iff]: "is_rule_R'_app \ p D C E \ is_rule_R_app p D C E \ rule_R'_side_condition \ p D C E" + +lemma is_rule_R'_app_alt_def: + shows " + is_rule_R'_app \ p D C E + \ + ( + \\ A B. + E = A =\<^bsub>\\<^esub> B \ A \ wffs\<^bsub>\\<^esub> \ B \ wffs\<^bsub>\\<^esub> \ \ \\<^term>\E\ is a well-formed equality\ + A \\<^bsub>p\<^esub> C \ D \ wffs\<^bsub>o\<^esub> \ + C\p \ B\ \ D \ + ( + \x \ E p'. + strict_prefix p' p \ + \x\<^bsub>\\<^esub>. E \\<^bsub>p'\<^esub> C \ + (x, \) \ free_vars (A =\<^bsub>\\<^esub> B) \ + (\H \ \. (x, \) \ free_vars H) + ) + )" + using rule_R'_side_condition_alt_def by fastforce + +lemma rule_R'_preserves_typing: + assumes "is_rule_R'_app \ p D C E" + shows "C \ wffs\<^bsub>o\<^esub> \ D \ wffs\<^bsub>o\<^esub>" + using assms and replacement_preserves_typing unfolding is_rule_R_app_def and is_rule_R'_app_def + by meson + +abbreviation is_hyps :: "form set \ bool" where + "is_hyps \ \ \ \ wffs\<^bsub>o\<^esub> \ finite \" + +inductive is_derivable_from_hyps :: "form set \ form \ bool" ("_ \ _" [50, 50] 50) for \ where + dv_hyp: "\ \ A" if "A \ \" and "is_hyps \" +| dv_thm: "\ \ A" if "is_theorem A" and "is_hyps \" +| dv_rule_R': "\ \ D" if "\ \ C" and "\ \ E" and "is_rule_R'_app \ p D C E" and "is_hyps \" + +lemma hyp_derivable_form_is_wffso: + assumes "is_derivable_from_hyps \ A" + shows "A \ wffs\<^bsub>o\<^esub>" + using assms and theorem_is_wffo by (cases rule: is_derivable_from_hyps.cases) auto + +definition is_hyp_proof_step :: "form set \ form list \ form list \ nat \ bool" where + [iff]: "is_hyp_proof_step \ \\<^sub>1 \\<^sub>2 i' \ + \\<^sub>2 ! i' \ \ \ + \\<^sub>2 ! i' \ lset \\<^sub>1 \ + (\p j k. {j, k} \ {0.. is_rule_R'_app \ p (\\<^sub>2 ! i') (\\<^sub>2 ! j) (\\<^sub>2 ! k))" + +type_synonym hyp_proof = "form list \ form list" + +definition is_hyp_proof :: "form set \ form list \ form list \ bool" where + [iff]: "is_hyp_proof \ \\<^sub>1 \\<^sub>2 \ (\i' < length \\<^sub>2. is_hyp_proof_step \ \\<^sub>1 \\<^sub>2 i')" + +lemma common_prefix_is_hyp_subproof_from: + assumes "is_hyp_proof \ \\<^sub>1 (\\<^sub>2 @ \\<^sub>2')" + and "i' < length \\<^sub>2" + shows "is_hyp_proof_step \ \\<^sub>1 (\\<^sub>2 @ \\<^sub>2'') i'" +proof - + let ?\ = "\\<^sub>2 @ \\<^sub>2'" + from assms(2) have "?\ ! i' = (\\<^sub>2 @ \\<^sub>2'') ! i'" + by (simp add: nth_append) + moreover from assms(2) have "i' < length ?\" + by simp + ultimately obtain p and j and k where " + ?\ ! i' \ \ \ + ?\ ! i' \ lset \\<^sub>1 \ + {j, k} \ {0.. is_rule_R'_app \ p (?\ ! i') (?\ ! j) (?\ ! k)" + using assms(1) unfolding is_hyp_proof_def and is_hyp_proof_step_def by meson + then consider + (hyp) "?\ ! i' \ \" + | (seq) "?\ ! i' \ lset \\<^sub>1" + | (rule_R') "{j, k} \ {0.. is_rule_R'_app \ p (?\ ! i') (?\ ! j) (?\ ! k)" + by blast + then have " + (\\<^sub>2 @ \\<^sub>2'') ! i' \ \ \ + (\\<^sub>2 @ \\<^sub>2'') ! i' \ lset \\<^sub>1 \ + ({j, k} \ {0.. is_rule_R'_app \ p ((\\<^sub>2 @ \\<^sub>2'') ! i') ((\\<^sub>2 @ \\<^sub>2'') ! j) ((\\<^sub>2 @ \\<^sub>2'') ! k))" + proof cases + case hyp + with assms(2) have "(\\<^sub>2 @ \\<^sub>2'') ! i' \ \" + by (simp add: nth_append) + then show ?thesis .. + next + case seq + with assms(2) have "(\\<^sub>2 @ \\<^sub>2'') ! i' \ lset \\<^sub>1" + by (simp add: nth_append) + then show ?thesis + by (intro disjI1 disjI2) + next + case rule_R' + with assms(2) have "?\ ! j = (\\<^sub>2 @ \\<^sub>2'') ! j" and "?\ ! k = (\\<^sub>2 @ \\<^sub>2'') ! k" + by (simp_all add: nth_append) + with assms(2) and rule_R' have " + {j, k} \ {0.. is_rule_R'_app \ p ((\\<^sub>2 @ \\<^sub>2'') ! i') ((\\<^sub>2 @ \\<^sub>2'') ! j) ((\\<^sub>2 @ \\<^sub>2'') ! k)" + by (metis nth_append) + then show ?thesis + by (intro disjI2) + qed + then show ?thesis + unfolding is_hyp_proof_step_def by meson +qed + +lemma added_suffix_thms_hyp_proof_preservation: + assumes "is_hyp_proof \ \\<^sub>1 \\<^sub>2" + shows "is_hyp_proof \ (\\<^sub>1 @ \\<^sub>1') \\<^sub>2" + using assms by auto + +lemma added_suffix_hyp_proof_preservation: + assumes "is_hyp_proof \ \\<^sub>1 \\<^sub>2" + and "i' < length (\\<^sub>2 @ \\<^sub>2') - length \\<^sub>2'" + shows "is_hyp_proof_step \ \\<^sub>1 (\\<^sub>2 @ \\<^sub>2') i'" + using assms and common_prefix_is_hyp_subproof_from[where \\<^sub>2' = "[]"] by auto + +lemma appended_hyp_proof_step_is_hyp_proof: + assumes "is_hyp_proof \ \\<^sub>1 \\<^sub>2" + and "is_hyp_proof_step \ \\<^sub>1 (\\<^sub>2 @ [A]) (length (\\<^sub>2 @ [A]) - 1)" + shows "is_hyp_proof \ \\<^sub>1 (\\<^sub>2 @ [A])" +proof (standard, intro allI impI) + fix i' + assume "i' < length (\\<^sub>2 @ [A])" + then consider (a) "i' < length \\<^sub>2" | (b) "i' = length \\<^sub>2" + by fastforce + then show "is_hyp_proof_step \ \\<^sub>1 (\\<^sub>2 @ [A]) i'" + proof cases + case a + with assms(1) show ?thesis + using added_suffix_hyp_proof_preservation by simp + next + case b + with assms(2) show ?thesis + by simp + qed +qed + +lemma added_prefix_hyp_proof_preservation: + assumes "is_hyp_proof \ \\<^sub>1 \\<^sub>2'" + and "i' \ {length \\<^sub>2..\<^sub>2 @ \\<^sub>2')}" + shows "is_hyp_proof_step \ \\<^sub>1 (\\<^sub>2 @ \\<^sub>2') i'" +proof - + let ?\ = "\\<^sub>2 @ \\<^sub>2'" + let ?i = "i' - length \\<^sub>2" + from assms(2) have "?\ ! i' = \\<^sub>2' ! ?i" and "?i < length \\<^sub>2'" + by (simp_all add: nth_append less_diff_conv2) + then have "is_hyp_proof_step \ \\<^sub>1 ?\ i' = is_hyp_proof_step \ \\<^sub>1 \\<^sub>2' ?i" + proof - + from assms(1) and \?i < length \\<^sub>2'\ obtain j and k and p where " + \\<^sub>2' ! ?i \ \ \ + \\<^sub>2' ! ?i \ lset \\<^sub>1 \ + ({j, k} \ {0.. is_rule_R'_app \ p (\\<^sub>2' ! ?i) (\\<^sub>2' ! j) (\\<^sub>2' ! k))" + unfolding is_hyp_proof_def and is_hyp_proof_step_def by meson + then consider + (hyp) "\\<^sub>2' ! ?i \ \" + | (seq) "\\<^sub>2' ! ?i \ lset \\<^sub>1" + | (rule_R') "{j, k} \ {0.. is_rule_R'_app \ p (\\<^sub>2' ! ?i) (\\<^sub>2' ! j) (\\<^sub>2' ! k)" + by blast + then have " + ?\ ! i' \ \ \ + ?\ ! i' \ lset \\<^sub>1 \ + ({j + length \\<^sub>2, k + length \\<^sub>2} \ {0.. + is_rule_R'_app \ p (?\ ! i') (?\ ! (j + length \\<^sub>2)) (?\ ! (k + length \\<^sub>2)))" + proof cases + case hyp + with \?\ ! i' = \\<^sub>2' ! ?i\ have "?\ ! i' \ \" + by (simp only:) + then show ?thesis .. + next + case seq + with \?\ ! i' = \\<^sub>2' ! ?i\ have "?\ ! i' \ lset \\<^sub>1" + by (simp only:) + then show ?thesis + by (intro disjI1 disjI2) + next + case rule_R' + with assms(2) have "?\ ! (j + length \\<^sub>2) = \\<^sub>2' ! j" and "?\ ! (k + length \\<^sub>2) = \\<^sub>2' ! k" + by (simp_all add: nth_append) + with \?\ ! i' = \\<^sub>2' ! ?i\ and rule_R' have " + {j + length \\<^sub>2, k + length \\<^sub>2} \ {0.. + is_rule_R'_app \ p (?\ ! i') (?\ ! (j + length \\<^sub>2)) (?\ ! (k + length \\<^sub>2))" + by auto + then show ?thesis + by (intro disjI2) + qed + with assms(1) and \?i < length \\<^sub>2'\ show ?thesis + unfolding is_hyp_proof_def and is_hyp_proof_step_def by meson + qed + with assms(1) and \?i < length \\<^sub>2'\ show ?thesis + by simp +qed + +lemma hyp_proof_but_last_is_hyp_proof: + assumes "is_hyp_proof \ \\<^sub>1 (\\<^sub>2 @ [A])" + shows "is_hyp_proof \ \\<^sub>1 \\<^sub>2" + using assms and common_prefix_is_hyp_subproof_from[where \\<^sub>2' = "[A]" and \\<^sub>2'' = "[]"] + by simp + +lemma hyp_proof_prefix_is_hyp_proof: + assumes "is_hyp_proof \ \\<^sub>1 (\\<^sub>2 @ \\<^sub>2')" + shows "is_hyp_proof \ \\<^sub>1 \\<^sub>2" + using assms and hyp_proof_but_last_is_hyp_proof + by (induction \\<^sub>2' arbitrary: \\<^sub>2 rule: rev_induct) (simp, metis append.assoc) + +lemma single_hyp_is_hyp_proof: + assumes "A \ \" + shows "is_hyp_proof \ \\<^sub>1 [A]" + using assms by fastforce + +lemma single_thm_is_hyp_proof: + assumes "A \ lset \\<^sub>1" + shows "is_hyp_proof \ \\<^sub>1 [A]" + using assms by fastforce + +lemma hyp_proofs_from_concatenation_is_hyp_proof: + assumes "is_hyp_proof \ \\<^sub>1 \\<^sub>1'" and "is_hyp_proof \ \\<^sub>2 \\<^sub>2'" + shows "is_hyp_proof \ (\\<^sub>1 @ \\<^sub>2) (\\<^sub>1' @ \\<^sub>2')" +proof (standard, intro allI impI) + let ?\ = "\\<^sub>1 @ \\<^sub>2" and ?\' = "\\<^sub>1' @ \\<^sub>2'" + fix i' + assume "i' < length ?\'" + then consider (a) "i' < length \\<^sub>1'" | (b) "i' \ {length \\<^sub>1'..'}" + by fastforce + then show "is_hyp_proof_step \ ?\ ?\' i'" + proof cases + case a + from \is_hyp_proof \ \\<^sub>1 \\<^sub>1'\ have "is_hyp_proof \ (\\<^sub>1 @ \\<^sub>2) \\<^sub>1'" + by auto + with assms(1) and a show ?thesis + using added_suffix_hyp_proof_preservation[where \\<^sub>1 = "\\<^sub>1 @ \\<^sub>2"] by auto + next + case b + from assms(2) have "is_hyp_proof \ (\\<^sub>1 @ \\<^sub>2) \\<^sub>2'" + by auto + with b show ?thesis + using added_prefix_hyp_proof_preservation[where \\<^sub>1 = "\\<^sub>1 @ \\<^sub>2"] by auto + qed +qed + +lemma elem_of_hyp_proof_is_wffo: + assumes "is_hyps \" + and "lset \\<^sub>1 \ wffs\<^bsub>o\<^esub>" + and "is_hyp_proof \ \\<^sub>1 \\<^sub>2" + and "A \ lset \\<^sub>2" + shows "A \ wffs\<^bsub>o\<^esub>" +using assms proof (induction \\<^sub>2 rule: rev_induct) + case Nil + then show ?case + by simp +next + case (snoc A' \\<^sub>2) + from \is_hyp_proof \ \\<^sub>1 (\\<^sub>2 @ [A'])\ have "is_hyp_proof \ \\<^sub>1 \\<^sub>2" + using hyp_proof_prefix_is_hyp_proof[where \\<^sub>2' = "[A']"] by presburger + then show ?case + proof (cases "A \ lset \\<^sub>2") + case True + with snoc.prems(1,2) and \is_hyp_proof \ \\<^sub>1 \\<^sub>2\ show ?thesis + by (fact snoc.IH) + next + case False + with snoc.prems(4) have "A' = A" + by simp + with snoc.prems(3) have " + (\\<^sub>2 @ [A]) ! i' \ \ \ + (\\<^sub>2 @ [A]) ! i' \ lset \\<^sub>1 \ + (\\<^sub>2 @ [A]) ! i' \ wffs\<^bsub>o\<^esub>" if "i' \ {0..\<^sub>2 @ [A])}" for i' + using that by auto + then have "A \ wffs\<^bsub>o\<^esub> \ A \ \ \ A \ lset \\<^sub>1 \ length \\<^sub>2 \ {0..\<^sub>2)}" + by (metis (no_types) length_append_singleton nth_append_length) + with assms(1) and \lset \\<^sub>1 \ wffs\<^bsub>o\<^esub>\ show ?thesis + using atLeast0_lessThan_Suc by blast + qed +qed + +lemma hyp_prepended_to_hyp_proof_is_hyp_proof: + assumes "is_hyp_proof \ \\<^sub>1 \\<^sub>2" + and "A \ \" + shows "is_hyp_proof \ \\<^sub>1 ([A] @ \\<^sub>2)" + using + hyp_proofs_from_concatenation_is_hyp_proof + [ + OF single_hyp_is_hyp_proof[OF assms(2)] assms(1), + where \\<^sub>1 = "[]" + ] + by simp + +lemma hyp_appended_to_hyp_proof_is_hyp_proof: + assumes "is_hyp_proof \ \\<^sub>1 \\<^sub>2" + and "A \ \" + shows "is_hyp_proof \ \\<^sub>1 (\\<^sub>2 @ [A])" + using + hyp_proofs_from_concatenation_is_hyp_proof + [ + OF assms(1) single_hyp_is_hyp_proof[OF assms(2)], + where \\<^sub>2 = "[]" + ] + by simp + +lemma dropped_duplicated_thm_in_hyp_proof_is_hyp_proof: + assumes "is_hyp_proof \ (A # \\<^sub>1) \\<^sub>2" + and "A \ lset \\<^sub>1" + shows "is_hyp_proof \ \\<^sub>1 \\<^sub>2" + using assms by auto + +lemma thm_prepended_to_hyp_proof_is_hyp_proof: + assumes "is_hyp_proof \ \\<^sub>1 \\<^sub>2" + and "A \ lset \\<^sub>1" + shows "is_hyp_proof \ \\<^sub>1 ([A] @ \\<^sub>2)" + using hyp_proofs_from_concatenation_is_hyp_proof[OF single_thm_is_hyp_proof[OF assms(2)] assms(1)] + and dropped_duplicated_thm_in_hyp_proof_is_hyp_proof by simp + +lemma thm_appended_to_hyp_proof_is_hyp_proof: + assumes "is_hyp_proof \ \\<^sub>1 \\<^sub>2" + and "A \ lset \\<^sub>1" + shows "is_hyp_proof \ \\<^sub>1 (\\<^sub>2 @ [A])" + using hyp_proofs_from_concatenation_is_hyp_proof[OF assms(1) single_thm_is_hyp_proof[OF assms(2)]] + and dropped_duplicated_thm_in_hyp_proof_is_hyp_proof by simp + +lemma rule_R'_app_appended_to_hyp_proof_is_hyp_proof: + assumes "is_hyp_proof \ \' \" + and "i\<^sub>C < length \" and "\ ! i\<^sub>C = C" + and "i\<^sub>E < length \" and "\ ! i\<^sub>E = E" + and "is_rule_R'_app \ p D C E" + shows "is_hyp_proof \ \' (\ @ [D])" +proof (standard, intro allI impI) + let ?\ = "\ @ [D]" + fix i' + assume "i' < length ?\" + then consider (a) "i' < length \" | (b) "i' = length \" + by fastforce + then show "is_hyp_proof_step \ \' (\ @ [D]) i'" + proof cases + case a + with assms(1) show ?thesis + using added_suffix_hyp_proof_preservation by auto + next + case b + let ?i\<^sub>D = "length \" + from assms(2,4) have "i\<^sub>C < ?i\<^sub>D" and "i\<^sub>E < ?i\<^sub>D" + by fastforce+ + with assms(3,5,6) have "is_rule_R'_app \ p (?\ ! ?i\<^sub>D) (?\ ! i\<^sub>C) (?\ ! i\<^sub>E)" + by (simp add: nth_append) + with assms(2,4) have " + \p j k. {j, k} \ {0..D} \ is_rule_R'_app \ p (?\ ! ?i\<^sub>D) (?\ ! j) (?\ ! k)" + by (intro exI)+ auto + then have "is_hyp_proof_step \ \' ?\ (length ?\ - 1)" + by simp + moreover from b have "i' = length ?\ - 1" + by simp + ultimately show ?thesis + by fast + qed +qed + +definition is_hyp_proof_of :: "form set \ form list \ form list \ form \ bool" where + [iff]: "is_hyp_proof_of \ \\<^sub>1 \\<^sub>2 A \ + is_hyps \ \ + is_proof \\<^sub>1 \ + \\<^sub>2 \ [] \ + is_hyp_proof \ \\<^sub>1 \\<^sub>2 \ + last \\<^sub>2 = A" + +lemma hyp_proof_prefix_is_hyp_proof_of_last: + assumes "is_hyps \" + and "is_proof \''" + and "is_hyp_proof \ \'' (\ @ \')" and "\ \ []" + shows "is_hyp_proof_of \ \'' \ (last \)" + using assms and hyp_proof_prefix_is_hyp_proof by simp + +theorem hyp_derivability_implies_hyp_proof_existence: + assumes "\ \ A" + shows "\\\<^sub>1 \\<^sub>2. is_hyp_proof_of \ \\<^sub>1 \\<^sub>2 A" +using assms proof (induction rule: is_derivable_from_hyps.induct) + case (dv_hyp A) + from \A \ \\ have "is_hyp_proof \ [] [A]" + by (fact single_hyp_is_hyp_proof) + moreover have "last [A] = A" + by simp + moreover have "is_proof []" + by simp + ultimately show ?case + using \is_hyps \\ unfolding is_hyp_proof_of_def by (meson list.discI) +next + case (dv_thm A) + then obtain \ where "is_proof \" and "\ \ []" and "last \ = A" + by fastforce + then have "is_hyp_proof \ \ [A]" + using single_thm_is_hyp_proof by auto + with \is_hyps \\ and \is_proof \\ have "is_hyp_proof_of \ \ [A] A" + by fastforce + then show ?case + by (intro exI) +next + case (dv_rule_R' C E p D) + from dv_rule_R'.IH obtain \\<^sub>C and \\<^sub>C' and \\<^sub>E and \\<^sub>E' where + "is_hyp_proof \ \\<^sub>C' \\<^sub>C" and "is_proof \\<^sub>C'" and "\\<^sub>C \ []" and "last \\<^sub>C = C" and + "is_hyp_proof \ \\<^sub>E' \\<^sub>E" and "is_proof \\<^sub>E'" and "\\<^sub>E \ []" and "last \\<^sub>E = E" + by auto + let ?i\<^sub>C = "length \\<^sub>C - 1" and ?i\<^sub>E = "length \\<^sub>C + length \\<^sub>E - 1" and ?i\<^sub>D = "length \\<^sub>C + length \\<^sub>E" + let ?\ = "\\<^sub>C @ \\<^sub>E @ [D]" + from \\\<^sub>C \ []\ have "?i\<^sub>C < length (\\<^sub>C @ \\<^sub>E)" and "?i\<^sub>E < length (\\<^sub>C @ \\<^sub>E)" + using linorder_not_le by fastforce+ + moreover have "(\\<^sub>C @ \\<^sub>E) ! ?i\<^sub>C = C" and "(\\<^sub>C @ \\<^sub>E) ! ?i\<^sub>E = E" + using \\\<^sub>C \ []\ and \last \\<^sub>C = C\ and \\\<^sub>E \ []\ and \last \\<^sub>E = E\ + by + ( + simp add: last_conv_nth nth_append, + metis append_is_Nil_conv last_appendR last_conv_nth length_append + ) + with \is_rule_R'_app \ p D C E\ have "is_rule_R'_app \ p D ((\\<^sub>C @ \\<^sub>E) ! ?i\<^sub>C) ((\\<^sub>C @ \\<^sub>E) ! ?i\<^sub>E)" + by fastforce + moreover from \is_hyp_proof \ \\<^sub>C' \\<^sub>C\ and \is_hyp_proof \ \\<^sub>E' \\<^sub>E\ + have "is_hyp_proof \ (\\<^sub>C' @ \\<^sub>E') (\\<^sub>C @ \\<^sub>E)" + by (fact hyp_proofs_from_concatenation_is_hyp_proof) + ultimately have "is_hyp_proof \ (\\<^sub>C' @ \\<^sub>E') ((\\<^sub>C @ \\<^sub>E) @ [D])" + using rule_R'_app_appended_to_hyp_proof_is_hyp_proof + by presburger + moreover from \is_proof \\<^sub>C'\ and \is_proof \\<^sub>E'\ have "is_proof (\\<^sub>C' @ \\<^sub>E')" + by (fact proofs_concatenation_is_proof) + ultimately have "is_hyp_proof_of \ (\\<^sub>C' @ \\<^sub>E') ((\\<^sub>C @ \\<^sub>E) @ [D]) D" + using \is_hyps \\ by fastforce + then show ?case + by (intro exI) +qed + +theorem hyp_proof_existence_implies_hyp_derivability: + assumes "\\\<^sub>1 \\<^sub>2. is_hyp_proof_of \ \\<^sub>1 \\<^sub>2 A" + shows "\ \ A" +proof - + from assms obtain \\<^sub>1 and \\<^sub>2 + where "is_hyps \" and "is_proof \\<^sub>1" and "\\<^sub>2 \ []" and "is_hyp_proof \ \\<^sub>1 \\<^sub>2" and "last \\<^sub>2 = A" + by fastforce + then show ?thesis + proof (induction "length \\<^sub>2" arbitrary: \\<^sub>2 A rule: less_induct) + case less + let ?i' = "length \\<^sub>2 - 1" + from \\\<^sub>2 \ []\ and \last \\<^sub>2 = A\ have "\\<^sub>2 ! ?i' = A" + by (simp add: last_conv_nth) + from \is_hyp_proof \ \\<^sub>1 \\<^sub>2\ and \\\<^sub>2 \ []\ have "is_hyp_proof_step \ \\<^sub>1 \\<^sub>2 ?i'" + by simp + then consider + (hyp) "\\<^sub>2 ! ?i' \ \" + | (seq) "\\<^sub>2 ! ?i' \ lset \\<^sub>1" + | (rule_R') "\p j k. {j, k} \ {0.. is_rule_R'_app \ p (\\<^sub>2 ! ?i') (\\<^sub>2 ! j) (\\<^sub>2 ! k)" + by force + then show ?case + proof cases + case hyp + with \\\<^sub>2 ! ?i' = A\ and \is_hyps \\ show ?thesis + by (fastforce intro: dv_hyp) + next + case seq + from \\\<^sub>2 ! ?i' \ lset \\<^sub>1\ and \\\<^sub>2 ! ?i' = A\ + obtain j where "\\<^sub>1 ! j = A" and "\\<^sub>1 \ []" and "j < length \\<^sub>1" + by (metis empty_iff in_set_conv_nth list.set(1)) + with \is_proof \\<^sub>1\ have "is_proof (take (Suc j) \\<^sub>1)" and "take (Suc j) \\<^sub>1 \ []" + using proof_prefix_is_proof[where \\<^sub>1 = "take (Suc j) \\<^sub>1" and \\<^sub>2 = "drop (Suc j) \\<^sub>1"] + by simp_all + moreover from \\\<^sub>1 ! j = A\ and \j < length \\<^sub>1\ have "last (take (Suc j) \\<^sub>1) = A" + by (simp add: take_Suc_conv_app_nth) + ultimately have "is_proof_of (take (Suc j) \\<^sub>1) A" + by fastforce + then have "is_theorem A" + using is_theorem_def by blast + with \is_hyps \\ show ?thesis + by (intro dv_thm) + next + case rule_R' + then obtain p and j and k + where "{j, k} \ {0.. p (\\<^sub>2 ! ?i') (\\<^sub>2 ! j) (\\<^sub>2 ! k)" + by force + let ?\\<^sub>j = "take (Suc j) \\<^sub>2" and ?\\<^sub>k = "take (Suc k) \\<^sub>2" + obtain \\<^sub>j' and \\<^sub>k' where "\\<^sub>2 = ?\\<^sub>j @ \\<^sub>j'" and "\\<^sub>2 = ?\\<^sub>k @ \\<^sub>k'" + by (metis append_take_drop_id) + then have "is_hyp_proof \ \\<^sub>1 (?\\<^sub>j @ \\<^sub>j')" and "is_hyp_proof \ \\<^sub>1 (?\\<^sub>k @ \\<^sub>k')" + by (simp_all only: \is_hyp_proof \ \\<^sub>1 \\<^sub>2\) + moreover from \\\<^sub>2 \ []\ and \\\<^sub>2 = ?\\<^sub>j @ \\<^sub>j'\ and \\\<^sub>2 = ?\\<^sub>k @ \\<^sub>k'\ and \last \\<^sub>2 = A\ + have "last \\<^sub>j' = A" and "last \\<^sub>k' = A" + using \{j, k} \ {0..\<^sub>2 - 1}\ and take_tl and less_le_not_le and append.right_neutral + by (metis atLeastLessThan_iff insert_subset last_appendR length_tl take_all_iff)+ + moreover from \\\<^sub>2 \ []\ have "?\\<^sub>j \ []" and "?\\<^sub>k \ []" + by simp_all + ultimately have "is_hyp_proof_of \ \\<^sub>1 ?\\<^sub>j (last ?\\<^sub>j)" and "is_hyp_proof_of \ \\<^sub>1 ?\\<^sub>k (last ?\\<^sub>k)" + using hyp_proof_prefix_is_hyp_proof_of_last + [OF \is_hyps \\ \is_proof \\<^sub>1\ \is_hyp_proof \ \\<^sub>1 (?\\<^sub>j @ \\<^sub>j')\ \?\\<^sub>j \ []\] + and hyp_proof_prefix_is_hyp_proof_of_last + [OF \is_hyps \\ \is_proof \\<^sub>1\ \is_hyp_proof \ \\<^sub>1 (?\\<^sub>k @ \\<^sub>k')\ \?\\<^sub>k \ []\] + by fastforce+ + moreover from \last \\<^sub>j' = A\ and \last \\<^sub>k' = A\ + have "length ?\\<^sub>j < length \\<^sub>2" and "length ?\\<^sub>k < length \\<^sub>2" + using \{j, k} \ {0..\<^sub>2 - 1}\ by force+ + moreover from calculation(3,4) have "last ?\\<^sub>j = \\<^sub>2 ! j" and "last ?\\<^sub>k = \\<^sub>2 ! k" + by (metis Suc_lessD last_snoc linorder_not_le nat_neq_iff take_Suc_conv_app_nth take_all_iff)+ + ultimately have "\ \ \\<^sub>2 ! j" and "\ \ \\<^sub>2 ! k" + using \is_hyps \\ + and less(1)[OF \length ?\\<^sub>j < length \\<^sub>2\] and less(1)[OF \length ?\\<^sub>k < length \\<^sub>2\] + by fast+ + with \is_hyps \\ and \\\<^sub>2 ! ?i' = A\ show ?thesis + using \is_rule_R'_app \ p (\\<^sub>2 ! ?i') (\\<^sub>2 ! j) (\\<^sub>2 ! k)\ by (blast intro: dv_rule_R') + qed + qed +qed + +theorem hypothetical_derivability_proof_existence_equivalence: + shows "\ \ A \ (\\\<^sub>1 \\<^sub>2. is_hyp_proof_of \ \\<^sub>1 \\<^sub>2 A)" + using hyp_derivability_implies_hyp_proof_existence and hyp_proof_existence_implies_hyp_derivability .. + +proposition derivability_from_no_hyps_theoremhood_equivalence: + shows "{} \ A \ is_theorem A" +proof + assume "{} \ A" + then show "is_theorem A" + proof (induction rule: is_derivable_from_hyps.induct) + case (dv_rule_R' C E p D) + from \is_rule_R'_app {} p D C E\ have "is_rule_R_app p D C E" + by simp + moreover from \is_theorem C\ and \is_theorem E\ have "is_derivable C" and "is_derivable E" + using theoremhood_derivability_equivalence by (simp_all only:) + ultimately have "is_derivable D" + by (fastforce intro: dv_rule_R) + then show ?case + using theoremhood_derivability_equivalence by (simp only:) + qed simp +next + assume "is_theorem A" + then show "{} \ A" + by (blast intro: dv_thm) +qed + +abbreviation is_derivable_from_no_hyps ("\ _" [50] 50) where + "\ A \ {} \ A" + +corollary derivability_implies_hyp_derivability: + assumes "\ A" and "is_hyps \" + shows "\ \ A" + using assms and derivability_from_no_hyps_theoremhood_equivalence and dv_thm by simp + +lemma axiom_is_derivable_from_no_hyps: + assumes "A \ axioms" + shows "\ A" + using derivability_from_no_hyps_theoremhood_equivalence + and derivable_form_is_theorem[OF dv_axiom[OF assms]] by (simp only:) + +lemma axiom_is_derivable_from_hyps: + assumes "A \ axioms" and "is_hyps \" + shows "\ \ A" + using assms and axiom_is_derivable_from_no_hyps and derivability_implies_hyp_derivability by blast + +lemma rule_R [consumes 2, case_names occ_subform replacement]: + assumes "\ C" and "\ A =\<^bsub>\\<^esub> B" + and "A \\<^bsub>p\<^esub> C" and "C\p \ B\ \ D" + shows "\ D" +proof - + from assms(1,2) have "is_derivable C" and "is_derivable (A =\<^bsub>\\<^esub> B)" + using derivability_from_no_hyps_theoremhood_equivalence + and theoremhood_derivability_equivalence by blast+ + moreover have "is_rule_R_app p D C (A =\<^bsub>\\<^esub> B)" + proof - + from assms(1-4) have "D \ wffs\<^bsub>o\<^esub>" and "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" + by (meson hyp_derivable_form_is_wffso replacement_preserves_typing wffs_from_equality)+ + with assms(3,4) show ?thesis + by fastforce + qed + ultimately have "is_derivable D" + by (rule dv_rule_R) + then show ?thesis + using derivability_from_no_hyps_theoremhood_equivalence and derivable_form_is_theorem by simp +qed + +lemma rule_R' [consumes 2, case_names occ_subform replacement no_capture]: + assumes "\ \ C" and "\ \ A =\<^bsub>\\<^esub> B" + and "A \\<^bsub>p\<^esub> C" and "C\p \ B\ \ D" + and "rule_R'_side_condition \ p D C (A =\<^bsub>\\<^esub> B)" + shows "\ \ D" +using assms(1,2) proof (rule dv_rule_R') + from assms(1) show "is_hyps \" + by (blast elim: is_derivable_from_hyps.cases) + moreover from assms(1-4) have "D \ wffs\<^bsub>o\<^esub>" + by (meson hyp_derivable_form_is_wffso replacement_preserves_typing wffs_from_equality) + ultimately show "is_rule_R'_app \ p D C (A =\<^bsub>\\<^esub> B)" + using assms(2-5) and hyp_derivable_form_is_wffso and wffs_from_equality + unfolding is_rule_R_app_def and is_rule_R'_app_def by metis +qed + +end diff --git a/thys/Q0_Metatheory/Propositional_Wff.thy b/thys/Q0_Metatheory/Propositional_Wff.thy new file mode 100644 --- /dev/null +++ b/thys/Q0_Metatheory/Propositional_Wff.thy @@ -0,0 +1,874 @@ +section \Propositional Well-Formed Formulas\ + +theory Propositional_Wff + imports + Syntax + Boolean_Algebra +begin + +subsection \Syntax\ + +inductive_set pwffs :: "form set" where + T_pwff: "T\<^bsub>o\<^esub> \ pwffs" +| F_pwff: "F\<^bsub>o\<^esub> \ pwffs" +| var_pwff: "p\<^bsub>o\<^esub> \ pwffs" +| neg_pwff: "\\<^sup>\ A \ pwffs" if "A \ pwffs" +| conj_pwff: "A \\<^sup>\ B \ pwffs" if "A \ pwffs" and "B \ pwffs" +| disj_pwff: "A \\<^sup>\ B \ pwffs" if "A \ pwffs" and "B \ pwffs" +| imp_pwff: "A \\<^sup>\ B \ pwffs" if "A \ pwffs" and "B \ pwffs" +| eqv_pwff: "A \\<^sup>\ B \ pwffs" if "A \ pwffs" and "B \ pwffs" + +lemmas [intro!] = pwffs.intros + +lemma pwffs_distinctnesses [induct_simp]: + shows "T\<^bsub>o\<^esub> \ F\<^bsub>o\<^esub>" + and "T\<^bsub>o\<^esub> \ p\<^bsub>o\<^esub>" + and "T\<^bsub>o\<^esub> \ \\<^sup>\ A" + and "T\<^bsub>o\<^esub> \ A \\<^sup>\ B" + and "T\<^bsub>o\<^esub> \ A \\<^sup>\ B" + and "T\<^bsub>o\<^esub> \ A \\<^sup>\ B" + and "T\<^bsub>o\<^esub> \ A \\<^sup>\ B" + and "F\<^bsub>o\<^esub> \ p\<^bsub>o\<^esub>" + and "F\<^bsub>o\<^esub> \ \\<^sup>\ A" + and "F\<^bsub>o\<^esub> \ A \\<^sup>\ B" + and "F\<^bsub>o\<^esub> \ A \\<^sup>\ B" + and "F\<^bsub>o\<^esub> \ A \\<^sup>\ B" + and "F\<^bsub>o\<^esub> \ A \\<^sup>\ B" + and "p\<^bsub>o\<^esub> \ \\<^sup>\ A" + and "p\<^bsub>o\<^esub> \ A \\<^sup>\ B" + and "p\<^bsub>o\<^esub> \ A \\<^sup>\ B" + and "p\<^bsub>o\<^esub> \ A \\<^sup>\ B" + and "p\<^bsub>o\<^esub> \ A \\<^sup>\ B" + and "\\<^sup>\ A \ B \\<^sup>\ C" + and "\\<^sup>\ A \ B \\<^sup>\ C" + and "\\<^sup>\ A \ B \\<^sup>\ C" + and "\ (B = F\<^bsub>o\<^esub> \ A = C) \ \\<^sup>\ A \ B \\<^sup>\ C" \ \\<^term>\\\<^sup>\ A\ is the same as \<^term>\F\<^bsub>o\<^esub> \\<^sup>\ A\\ + and "A \\<^sup>\ B \ C \\<^sup>\ D" + and "A \\<^sup>\ B \ C \\<^sup>\ D" + and "A \\<^sup>\ B \ C \\<^sup>\ D" + and "A \\<^sup>\ B \ C \\<^sup>\ D" + and "A \\<^sup>\ B \ C \\<^sup>\ D" + and "A \\<^sup>\ B \ C \\<^sup>\ D" + by simp_all + +lemma pwffs_injectivities [induct_simp]: + shows "\\<^sup>\ A = \\<^sup>\ A' \ A = A'" + and "A \\<^sup>\ B = A' \\<^sup>\ B' \ A = A' \ B = B'" + and "A \\<^sup>\ B = A' \\<^sup>\ B' \ A = A' \ B = B'" + and "A \\<^sup>\ B = A' \\<^sup>\ B' \ A = A' \ B = B'" + and "A \\<^sup>\ B = A' \\<^sup>\ B' \ A = A' \ B = B'" + by simp_all + +lemma pwff_from_neg_pwff [elim!]: + assumes "\\<^sup>\ A \ pwffs" + shows "A \ pwffs" + using assms by cases simp_all + +lemma pwffs_from_conj_pwff [elim!]: + assumes "A \\<^sup>\ B \ pwffs" + shows "{A, B} \ pwffs" + using assms by cases simp_all + +lemma pwffs_from_disj_pwff [elim!]: + assumes "A \\<^sup>\ B \ pwffs" + shows "{A, B} \ pwffs" + using assms by cases simp_all + +lemma pwffs_from_imp_pwff [elim!]: + assumes "A \\<^sup>\ B \ pwffs" + shows "{A, B} \ pwffs" + using assms by cases simp_all + +lemma pwffs_from_eqv_pwff [elim!]: + assumes "A \\<^sup>\ B \ pwffs" + shows "{A, B} \ pwffs" + using assms by cases (simp_all, use F_pwff in fastforce) + +lemma pwffs_subset_of_wffso: + shows "pwffs \ wffs\<^bsub>o\<^esub>" +proof + fix A + assume "A \ pwffs" + then show "A \ wffs\<^bsub>o\<^esub>" + by induction auto +qed + +lemma pwff_free_vars_simps [simp]: + shows T_fv: "free_vars T\<^bsub>o\<^esub> = {}" + and F_fv: "free_vars F\<^bsub>o\<^esub> = {}" + and var_fv: "free_vars (p\<^bsub>o\<^esub>) = {(p, o)}" + and neg_fv: "free_vars (\\<^sup>\ A) = free_vars A" + and conj_fv: "free_vars (A \\<^sup>\ B) = free_vars A \ free_vars B" + and disj_fv: "free_vars (A \\<^sup>\ B) = free_vars A \ free_vars B" + and imp_fv: "free_vars (A \\<^sup>\ B) = free_vars A \ free_vars B" + and eqv_fv: "free_vars (A \\<^sup>\ B) = free_vars A \ free_vars B" + by force+ + +lemma pwffs_free_vars_are_propositional: + assumes "A \ pwffs" + and "v \ free_vars A" + obtains p where "v = (p, o)" +using assms by (induction A arbitrary: thesis) auto + +lemma is_free_for_in_pwff [intro]: + assumes "A \ pwffs" + and "v \ free_vars A" + shows "is_free_for B v A" +using assms proof (induction A) + case (neg_pwff C) + then show ?case + using is_free_for_in_neg by simp +next + case (conj_pwff C D) + from conj_pwff.prems consider + (a) "v \ free_vars C" and "v \ free_vars D" + | (b) "v \ free_vars C" and "v \ free_vars D" + | (c) "v \ free_vars C" and "v \ free_vars D" + by auto + then show ?case + proof cases + case a + then show ?thesis + using conj_pwff.IH by (intro is_free_for_in_conj) + next + case b + have "is_free_for B v C" + by (fact conj_pwff.IH(1)[OF b(1)]) + moreover from b(2) have "is_free_for B v D" + using is_free_at_in_free_vars by blast + ultimately show ?thesis + by (rule is_free_for_in_conj) + next + case c + from c(1) have "is_free_for B v C" + using is_free_at_in_free_vars by blast + moreover have "is_free_for B v D" + by (fact conj_pwff.IH(2)[OF c(2)]) + ultimately show ?thesis + by (rule is_free_for_in_conj) + qed +next + case (disj_pwff C D) + from disj_pwff.prems consider + (a) "v \ free_vars C" and "v \ free_vars D" + | (b) "v \ free_vars C" and "v \ free_vars D" + | (c) "v \ free_vars C" and "v \ free_vars D" + by auto + then show ?case + proof cases + case a + then show ?thesis + using disj_pwff.IH by (intro is_free_for_in_disj) + next + case b + have "is_free_for B v C" + by (fact disj_pwff.IH(1)[OF b(1)]) + moreover from b(2) have "is_free_for B v D" + using is_free_at_in_free_vars by blast + ultimately show ?thesis + by (rule is_free_for_in_disj) + next + case c + from c(1) have "is_free_for B v C" + using is_free_at_in_free_vars by blast + moreover have "is_free_for B v D" + by (fact disj_pwff.IH(2)[OF c(2)]) + ultimately show ?thesis + by (rule is_free_for_in_disj) + qed +next + case (imp_pwff C D) + from imp_pwff.prems consider + (a) "v \ free_vars C" and "v \ free_vars D" + | (b) "v \ free_vars C" and "v \ free_vars D" + | (c) "v \ free_vars C" and "v \ free_vars D" + by auto + then show ?case + proof cases + case a + then show ?thesis + using imp_pwff.IH by (intro is_free_for_in_imp) + next + case b + have "is_free_for B v C" + by (fact imp_pwff.IH(1)[OF b(1)]) + moreover from b(2) have "is_free_for B v D" + using is_free_at_in_free_vars by blast + ultimately show ?thesis + by (rule is_free_for_in_imp) + next + case c + from c(1) have "is_free_for B v C" + using is_free_at_in_free_vars by blast + moreover have "is_free_for B v D" + by (fact imp_pwff.IH(2)[OF c(2)]) + ultimately show ?thesis + by (rule is_free_for_in_imp) + qed +next + case (eqv_pwff C D) + from eqv_pwff.prems consider + (a) "v \ free_vars C" and "v \ free_vars D" + | (b) "v \ free_vars C" and "v \ free_vars D" + | (c) "v \ free_vars C" and "v \ free_vars D" + by auto + then show ?case + proof cases + case a + then show ?thesis + using eqv_pwff.IH by (intro is_free_for_in_equivalence) + next + case b + have "is_free_for B v C" + by (fact eqv_pwff.IH(1)[OF b(1)]) + moreover from b(2) have "is_free_for B v D" + using is_free_at_in_free_vars by blast + ultimately show ?thesis + by (rule is_free_for_in_equivalence) + next + case c + from c(1) have "is_free_for B v C" + using is_free_at_in_free_vars by blast + moreover have "is_free_for B v D" + by (fact eqv_pwff.IH(2)[OF c(2)]) + ultimately show ?thesis + by (rule is_free_for_in_equivalence) + qed +qed auto + +subsection \Semantics\ + +text \Assignment of truth values to propositional variables:\ + +definition is_tv_assignment :: "(nat \ V) \ bool" where + [iff]: "is_tv_assignment \ \ (\p. \ p \ elts \)" + +text \Denotation of a pwff:\ + +definition is_pwff_denotation_function where + [iff]: "is_pwff_denotation_function \ \ + ( + \\. is_tv_assignment \ \ + ( + \ \ T\<^bsub>o\<^esub> = \<^bold>T \ + \ \ F\<^bsub>o\<^esub> = \<^bold>F \ + (\p. \ \ (p\<^bsub>o\<^esub>) = \ p) \ + (\A. A \ pwffs \ \ \ (\\<^sup>\ A) = \ \ \ A) \ + (\A B. A \ pwffs \ B \ pwffs \ \ \ (A \\<^sup>\ B) = \ \ A \<^bold>\ \ \ B) \ + (\A B. A \ pwffs \ B \ pwffs \ \ \ (A \\<^sup>\ B) = \ \ A \<^bold>\ \ \ B) \ + (\A B. A \ pwffs \ B \ pwffs \ \ \ (A \\<^sup>\ B) = \ \ A \<^bold>\ \ \ B) \ + (\A B. A \ pwffs \ B \ pwffs \ \ \ (A \\<^sup>\ B) = \ \ A \<^bold>\ \ \ B) + ) + )" + +lemma pwff_denotation_is_truth_value: + assumes "A \ pwffs" + and "is_tv_assignment \" + and "is_pwff_denotation_function \" + shows "\ \ A \ elts \" +using assms(1) proof induction + case (neg_pwff A) + then have "\ \ (\\<^sup>\ A) = \ \ \ A" + using assms(2,3) by auto + then show ?case + using neg_pwff.IH by auto +next + case (conj_pwff A B) + then have "\ \ (A \\<^sup>\ B) = \ \ A \<^bold>\ \ \ B" + using assms(2,3) by auto + then show ?case + using conj_pwff.IH by auto +next + case (disj_pwff A B) + then have "\ \ (A \\<^sup>\ B) = \ \ A \<^bold>\ \ \ B" + using assms(2,3) by auto + then show ?case + using disj_pwff.IH by auto +next + case (imp_pwff A B) + then have "\ \ (A \\<^sup>\ B) = \ \ A \<^bold>\ \ \ B" + using assms(2,3) by blast + then show ?case + using imp_pwff.IH by auto +next + case (eqv_pwff A B) + then have "\ \ (A \\<^sup>\ B) = \ \ A \<^bold>\ \ \ B" + using assms(2,3) by blast + then show ?case + using eqv_pwff.IH by auto +qed (use assms(2,3) in auto) + +lemma closed_pwff_is_meaningful_regardless_of_assignment: + assumes "A \ pwffs" + and "free_vars A = {}" + and "is_tv_assignment \" + and "is_tv_assignment \" + and "is_pwff_denotation_function \" + shows "\ \ A = \ \ A" +using assms(1,2) proof induction + case T_pwff + have "\ \ T\<^bsub>o\<^esub> = \<^bold>T" + using assms(3,5) by blast + also have "\ = \ \ T\<^bsub>o\<^esub>" + using assms(4,5) by force + finally show ?case . +next + case F_pwff + have "\ \ F\<^bsub>o\<^esub> = \<^bold>F" + using assms(3,5) by blast + also have "\ = \ \ F\<^bsub>o\<^esub>" + using assms(4,5) by force + finally show ?case . +next + case (var_pwff p) \ \impossible case\ + then show ?case + by simp +next + case (neg_pwff A) + from \free_vars (\\<^sup>\ A) = {}\ have "free_vars A = {}" + by simp + have "\ \ (\\<^sup>\ A) = \ \ \ A" + using assms(3,5) and neg_pwff.hyps by auto + also from \free_vars A = {}\ have "\ = \ \ \ A" + using assms(3-5) and neg_pwff.IH by presburger + also have "\ = \ \ (\\<^sup>\ A)" + using assms(4,5) and neg_pwff.hyps by simp + finally show ?case . +next + case (conj_pwff A B) + from \free_vars (A \\<^sup>\ B) = {}\ have "free_vars A = {}" and "free_vars B = {}" + by simp_all + have "\ \ (A \\<^sup>\ B) = \ \ A \<^bold>\ \ \ B" + using assms(3,5) and conj_pwff.hyps(1,2) by auto + also from \free_vars A = {}\ and \free_vars B = {}\ have "\ = \ \ A \<^bold>\ \ \ B" + using conj_pwff.IH(1,2) by presburger + also have "\ = \ \ (A \\<^sup>\ B)" + using assms(4,5) and conj_pwff.hyps(1,2) by fastforce + finally show ?case . +next + case (disj_pwff A B) + from \free_vars (A \\<^sup>\ B) = {}\ have "free_vars A = {}" and "free_vars B = {}" + by simp_all + have "\ \ (A \\<^sup>\ B) = \ \ A \<^bold>\ \ \ B" + using assms(3,5) and disj_pwff.hyps(1,2) by auto + also from \free_vars A = {}\ and \free_vars B = {}\ have "\ = \ \ A \<^bold>\ \ \ B" + using disj_pwff.IH(1,2) by presburger + also have "\ = \ \ (A \\<^sup>\ B)" + using assms(4,5) and disj_pwff.hyps(1,2) by fastforce + finally show ?case . +next + case (imp_pwff A B) + from \free_vars (A \\<^sup>\ B) = {}\ have "free_vars A = {}" and "free_vars B = {}" + by simp_all + have "\ \ (A \\<^sup>\ B) = \ \ A \<^bold>\ \ \ B" + using assms(3,5) and imp_pwff.hyps(1,2) by auto + also from \free_vars A = {}\ and \free_vars B = {}\ have "\ = \ \ A \<^bold>\ \ \ B" + using imp_pwff.IH(1,2) by presburger + also have "\ = \ \ (A \\<^sup>\ B)" + using assms(4,5) and imp_pwff.hyps(1,2) by fastforce + finally show ?case . +next + case (eqv_pwff A B) + from \free_vars (A \\<^sup>\ B) = {}\ have "free_vars A = {}" and "free_vars B = {}" + by simp_all + have "\ \ (A \\<^sup>\ B) = \ \ A \<^bold>\ \ \ B" + using assms(3,5) and eqv_pwff.hyps(1,2) by auto + also from \free_vars A = {}\ and \free_vars B = {}\ have "\ = \ \ A \<^bold>\ \ \ B" + using eqv_pwff.IH(1,2) by presburger + also have "\ = \ \ (A \\<^sup>\ B)" + using assms(4,5) and eqv_pwff.hyps(1,2) by fastforce + finally show ?case . +qed + +inductive \\<^sub>B_graph for \ where + \\<^sub>B_graph_T: "\\<^sub>B_graph \ T\<^bsub>o\<^esub> \<^bold>T" +| \\<^sub>B_graph_F: "\\<^sub>B_graph \ F\<^bsub>o\<^esub> \<^bold>F" +| \\<^sub>B_graph_var: "\\<^sub>B_graph \ (p\<^bsub>o\<^esub>) (\ p)" +| \\<^sub>B_graph_neg: "\\<^sub>B_graph \ (\\<^sup>\ A) (\ b\<^sub>A)" if "\\<^sub>B_graph \ A b\<^sub>A" +| \\<^sub>B_graph_conj: "\\<^sub>B_graph \ (A \\<^sup>\ B) (b\<^sub>A \<^bold>\ b\<^sub>B)" if "\\<^sub>B_graph \ A b\<^sub>A" and "\\<^sub>B_graph \ B b\<^sub>B" +| \\<^sub>B_graph_disj: "\\<^sub>B_graph \ (A \\<^sup>\ B) (b\<^sub>A \<^bold>\ b\<^sub>B)" if "\\<^sub>B_graph \ A b\<^sub>A" and "\\<^sub>B_graph \ B b\<^sub>B" +| \\<^sub>B_graph_imp: "\\<^sub>B_graph \ (A \\<^sup>\ B) (b\<^sub>A \<^bold>\ b\<^sub>B)" if "\\<^sub>B_graph \ A b\<^sub>A" and "\\<^sub>B_graph \ B b\<^sub>B" +| \\<^sub>B_graph_eqv: "\\<^sub>B_graph \ (A \\<^sup>\ B) (b\<^sub>A \<^bold>\ b\<^sub>B)" if "\\<^sub>B_graph \ A b\<^sub>A" and "\\<^sub>B_graph \ B b\<^sub>B" and "A \ F\<^bsub>o\<^esub>" + +lemmas [intro!] = \\<^sub>B_graph.intros + +lemma \\<^sub>B_graph_denotation_is_truth_value [elim!]: + assumes "\\<^sub>B_graph \ A b" + and "is_tv_assignment \" + shows "b \ elts \" +using assms proof induction + case (\\<^sub>B_graph_neg A b\<^sub>A) + show ?case + using \\<^sub>B_graph_neg.IH[OF assms(2)] by force +next + case (\\<^sub>B_graph_conj A b\<^sub>A B b\<^sub>B) + then show ?case + using \\<^sub>B_graph_conj.IH and assms(2) by force +next + case (\\<^sub>B_graph_disj A b\<^sub>A B b\<^sub>B) + then show ?case + using \\<^sub>B_graph_disj.IH and assms(2) by force +next + case (\\<^sub>B_graph_imp A b\<^sub>A B b\<^sub>B) + then show ?case + using \\<^sub>B_graph_imp.IH and assms(2) by force +next + case (\\<^sub>B_graph_eqv A b\<^sub>A B b\<^sub>B) + then show ?case + using \\<^sub>B_graph_eqv.IH and assms(2) by force +qed simp_all + +lemma \\<^sub>B_graph_denotation_uniqueness: + assumes "A \ pwffs" + and "is_tv_assignment \" + and "\\<^sub>B_graph \ A b" and "\\<^sub>B_graph \ A b'" + shows "b = b'" +using assms(3,1,4) proof (induction arbitrary: b') + case \\<^sub>B_graph_T + from \\\<^sub>B_graph \ T\<^bsub>o\<^esub> b'\ show ?case + by (cases rule: \\<^sub>B_graph.cases) simp_all +next + case \\<^sub>B_graph_F + from \\\<^sub>B_graph \ F\<^bsub>o\<^esub> b'\ show ?case + by (cases rule: \\<^sub>B_graph.cases) simp_all +next + case (\\<^sub>B_graph_var p) + from \\\<^sub>B_graph \ (p\<^bsub>o\<^esub>) b'\ show ?case + by (cases rule: \\<^sub>B_graph.cases) simp_all +next + case (\\<^sub>B_graph_neg A b\<^sub>A) + with \\\<^sub>B_graph \ (\\<^sup>\ A) b'\ have "\\<^sub>B_graph \ A (\ b')" + proof (cases rule: \\<^sub>B_graph.cases) + case (\\<^sub>B_graph_neg A' b\<^sub>A) + from \\\<^sup>\ A = \\<^sup>\ A'\ have "A = A'" + by simp + with \\\<^sub>B_graph \ A' b\<^sub>A\ have "\\<^sub>B_graph \ A b\<^sub>A" + by simp + moreover have "b\<^sub>A = \ b'" + proof - + have "b\<^sub>A \ elts \" + by (fact \\<^sub>B_graph_denotation_is_truth_value[OF \\<^sub>B_graph_neg(3) assms(2)]) + moreover from \b\<^sub>A \ elts \\ and \\<^sub>B_graph_neg(2) have "\ b' \ elts \" + by fastforce + ultimately show ?thesis + using \\<^sub>B_graph_neg(2) by fastforce + qed + ultimately show ?thesis + by blast + qed simp_all + moreover from \\<^sub>B_graph_neg.prems(1) have "A \ pwffs" + by (force elim: pwffs.cases) + moreover have "b\<^sub>A \ elts \" and "b' \ elts \" and "b\<^sub>A = \ b'" + proof - + show "b\<^sub>A \ elts \" + by (fact \\<^sub>B_graph_denotation_is_truth_value[OF \\\<^sub>B_graph \ A b\<^sub>A\ assms(2)]) + show "b' \ elts \" + by (fact \\<^sub>B_graph_denotation_is_truth_value[OF \\\<^sub>B_graph \ (\\<^sup>\ A) b'\ assms(2)]) + show "b\<^sub>A = \ b'" + by (fact \\<^sub>B_graph_neg(2)[OF \A \ pwffs\ \\\<^sub>B_graph \ A (\ b')\]) + qed + ultimately show ?case + by force +next + case (\\<^sub>B_graph_conj A b\<^sub>A B b\<^sub>B) + with \\\<^sub>B_graph \ (A \\<^sup>\ B) b'\ obtain b\<^sub>A' and b\<^sub>B' + where "b' = b\<^sub>A' \<^bold>\ b\<^sub>B'" and "\\<^sub>B_graph \ A b\<^sub>A'" and "\\<^sub>B_graph \ B b\<^sub>B'" + by (cases rule: \\<^sub>B_graph.cases) simp_all + moreover have "A \ pwffs" and "B \ pwffs" + using pwffs_from_conj_pwff[OF \\<^sub>B_graph_conj.prems(1)] by blast+ + ultimately show ?case + using \\<^sub>B_graph_conj.IH and \\<^sub>B_graph_conj.prems(2) by blast +next + case (\\<^sub>B_graph_disj A b\<^sub>A B b\<^sub>B) + from \\\<^sub>B_graph \ (A \\<^sup>\ B) b'\ obtain b\<^sub>A' and b\<^sub>B' + where "b' = b\<^sub>A' \<^bold>\ b\<^sub>B'" and "\\<^sub>B_graph \ A b\<^sub>A'" and "\\<^sub>B_graph \ B b\<^sub>B'" + by (cases rule: \\<^sub>B_graph.cases) simp_all + moreover have "A \ pwffs" and "B \ pwffs" + using pwffs_from_disj_pwff[OF \\<^sub>B_graph_disj.prems(1)] by blast+ + ultimately show ?case + using \\<^sub>B_graph_disj.IH and \\<^sub>B_graph_disj.prems(2) by blast +next + case (\\<^sub>B_graph_imp A b\<^sub>A B b\<^sub>B) + from \\\<^sub>B_graph \ (A \\<^sup>\ B) b'\ obtain b\<^sub>A' and b\<^sub>B' + where "b' = b\<^sub>A' \<^bold>\ b\<^sub>B'" and "\\<^sub>B_graph \ A b\<^sub>A'" and "\\<^sub>B_graph \ B b\<^sub>B'" + by (cases rule: \\<^sub>B_graph.cases) simp_all + moreover have "A \ pwffs" and "B \ pwffs" + using pwffs_from_imp_pwff[OF \\<^sub>B_graph_imp.prems(1)] by blast+ + ultimately show ?case + using \\<^sub>B_graph_imp.IH and \\<^sub>B_graph_imp.prems(2) by blast +next + case (\\<^sub>B_graph_eqv A b\<^sub>A B b\<^sub>B) + with \\\<^sub>B_graph \ (A \\<^sup>\ B) b'\ obtain b\<^sub>A' and b\<^sub>B' + where "b' = b\<^sub>A' \<^bold>\ b\<^sub>B'" and "\\<^sub>B_graph \ A b\<^sub>A'" and "\\<^sub>B_graph \ B b\<^sub>B'" + by (cases rule: \\<^sub>B_graph.cases) simp_all + moreover have "A \ pwffs" and "B \ pwffs" + using pwffs_from_eqv_pwff[OF \\<^sub>B_graph_eqv.prems(1)] by blast+ + ultimately show ?case + using \\<^sub>B_graph_eqv.IH and \\<^sub>B_graph_eqv.prems(2) by blast +qed + +lemma \\<^sub>B_graph_denotation_existence: + assumes "A \ pwffs" + and "is_tv_assignment \" + shows "\b. \\<^sub>B_graph \ A b" +using assms proof induction + case (eqv_pwff A B) + then obtain b\<^sub>A and b\<^sub>B where "\\<^sub>B_graph \ A b\<^sub>A" and "\\<^sub>B_graph \ B b\<^sub>B" + by blast + then show ?case + proof (cases "A \ F\<^bsub>o\<^esub>") + case True + then show ?thesis + using eqv_pwff.IH and eqv_pwff.prems by blast + next + case False + then have "A = F\<^bsub>o\<^esub>" + by blast + then show ?thesis + using \\<^sub>B_graph_neg[OF \\\<^sub>B_graph \ B b\<^sub>B\] by auto + qed +qed blast+ + +lemma \\<^sub>B_graph_is_functional: + assumes "A \ pwffs" + and "is_tv_assignment \" + shows "\!b. \\<^sub>B_graph \ A b" + using assms and \\<^sub>B_graph_denotation_existence and \\<^sub>B_graph_denotation_uniqueness by blast + +definition \\<^sub>B :: "(nat \ V) \ form \ V" where + [simp]: "\\<^sub>B \ A = (THE b. \\<^sub>B_graph \ A b)" + +lemma \\<^sub>B_equality: + assumes "A \ pwffs" + and "is_tv_assignment \" + and "\\<^sub>B_graph \ A b" + shows "\\<^sub>B \ A = b" + unfolding \\<^sub>B_def using assms using \\<^sub>B_graph_denotation_uniqueness by blast + +lemma \\<^sub>B_graph_\\<^sub>B: + assumes "A \ pwffs" + and "is_tv_assignment \" + shows "\\<^sub>B_graph \ A (\\<^sub>B \ A)" + using \\<^sub>B_equality[OF assms] and \\<^sub>B_graph_is_functional[OF assms] by blast + +named_theorems \\<^sub>B_simps + +lemma \\<^sub>B_T [\\<^sub>B_simps]: + assumes "is_tv_assignment \" + shows "\\<^sub>B \ T\<^bsub>o\<^esub> = \<^bold>T" + by (rule \\<^sub>B_equality[OF T_pwff assms], intro \\<^sub>B_graph_T) + +lemma \\<^sub>B_F [\\<^sub>B_simps]: + assumes "is_tv_assignment \" + shows "\\<^sub>B \ F\<^bsub>o\<^esub> = \<^bold>F" + by (rule \\<^sub>B_equality[OF F_pwff assms], intro \\<^sub>B_graph_F) + +lemma \\<^sub>B_var [\\<^sub>B_simps]: + assumes "is_tv_assignment \" + shows "\\<^sub>B \ (p\<^bsub>o\<^esub>) = \ p" + by (rule \\<^sub>B_equality[OF var_pwff assms], intro \\<^sub>B_graph_var) + +lemma \\<^sub>B_neg [\\<^sub>B_simps]: + assumes "A \ pwffs" + and "is_tv_assignment \" + shows "\\<^sub>B \ (\\<^sup>\ A) = \ \\<^sub>B \ A" + by (rule \\<^sub>B_equality[OF neg_pwff[OF assms(1)] assms(2)], intro \\<^sub>B_graph_neg \\<^sub>B_graph_\\<^sub>B[OF assms]) + +lemma \\<^sub>B_disj [\\<^sub>B_simps]: + assumes "A \ pwffs" and "B \ pwffs" + and "is_tv_assignment \" + shows "\\<^sub>B \ (A \\<^sup>\ B) = \\<^sub>B \ A \<^bold>\ \\<^sub>B \ B" +proof - + from assms(1,3) have "\\<^sub>B_graph \ A (\\<^sub>B \ A)" + by (intro \\<^sub>B_graph_\\<^sub>B) + moreover from assms(2,3) have "\\<^sub>B_graph \ B (\\<^sub>B \ B)" + by (intro \\<^sub>B_graph_\\<^sub>B) + ultimately have "\\<^sub>B_graph \ (A \\<^sup>\ B) (\\<^sub>B \ A \<^bold>\ \\<^sub>B \ B)" + by (intro \\<^sub>B_graph_disj) + with assms show ?thesis + using disj_pwff by (intro \\<^sub>B_equality) +qed + +lemma \\<^sub>B_conj [\\<^sub>B_simps]: + assumes "A \ pwffs" and "B \ pwffs" + and "is_tv_assignment \" + shows "\\<^sub>B \ (A \\<^sup>\ B) = \\<^sub>B \ A \<^bold>\ \\<^sub>B \ B" +proof - + from assms(1,3) have "\\<^sub>B_graph \ A (\\<^sub>B \ A)" + by (intro \\<^sub>B_graph_\\<^sub>B) + moreover from assms(2,3) have "\\<^sub>B_graph \ B (\\<^sub>B \ B)" + by (intro \\<^sub>B_graph_\\<^sub>B) + ultimately have "\\<^sub>B_graph \ (A \\<^sup>\ B) (\\<^sub>B \ A \<^bold>\ \\<^sub>B \ B)" + by (intro \\<^sub>B_graph_conj) + with assms show ?thesis + using conj_pwff by (intro \\<^sub>B_equality) +qed + +lemma \\<^sub>B_imp [\\<^sub>B_simps]: + assumes "A \ pwffs" and "B \ pwffs" + and "is_tv_assignment \" + shows "\\<^sub>B \ (A \\<^sup>\ B) = \\<^sub>B \ A \<^bold>\ \\<^sub>B \ B" +proof - + from assms(1,3) have "\\<^sub>B_graph \ A (\\<^sub>B \ A)" + by (intro \\<^sub>B_graph_\\<^sub>B) + moreover from assms(2,3) have "\\<^sub>B_graph \ B (\\<^sub>B \ B)" + by (intro \\<^sub>B_graph_\\<^sub>B) + ultimately have "\\<^sub>B_graph \ (A \\<^sup>\ B) (\\<^sub>B \ A \<^bold>\ \\<^sub>B \ B)" + by (intro \\<^sub>B_graph_imp) + with assms show ?thesis + using imp_pwff by (intro \\<^sub>B_equality) +qed + +lemma \\<^sub>B_eqv [\\<^sub>B_simps]: + assumes "A \ pwffs" and "B \ pwffs" + and "is_tv_assignment \" + shows "\\<^sub>B \ (A \\<^sup>\ B) = \\<^sub>B \ A \<^bold>\ \\<^sub>B \ B" +proof (cases "A = F\<^bsub>o\<^esub>") + case True + then show ?thesis + using \\<^sub>B_F[OF assms(3)] and \\<^sub>B_neg[OF assms(2,3)] by force +next + case False + from assms(1,3) have "\\<^sub>B_graph \ A (\\<^sub>B \ A)" + by (intro \\<^sub>B_graph_\\<^sub>B) + moreover from assms(2,3) have "\\<^sub>B_graph \ B (\\<^sub>B \ B)" + by (intro \\<^sub>B_graph_\\<^sub>B) + ultimately have "\\<^sub>B_graph \ (A \\<^sup>\ B) (\\<^sub>B \ A \<^bold>\ \\<^sub>B \ B)" + using False by (intro \\<^sub>B_graph_eqv) + with assms show ?thesis + using eqv_pwff by (intro \\<^sub>B_equality) +qed + +declare pwffs.intros [\\<^sub>B_simps] + +lemma pwff_denotation_function_existence: + shows "is_pwff_denotation_function \\<^sub>B" + using \\<^sub>B_simps by simp + +text \Tautologies:\ + +definition is_tautology :: "form \ bool" where + [iff]: "is_tautology A \ A \ pwffs \ (\\. is_tv_assignment \ \ \\<^sub>B \ A = \<^bold>T)" + +lemma tautology_is_wffo: + assumes "is_tautology A" + shows "A \ wffs\<^bsub>o\<^esub>" + using assms and pwffs_subset_of_wffso by blast + +lemma propositional_implication_reflexivity_is_tautology: + shows "is_tautology (p\<^bsub>o\<^esub> \\<^sup>\ p\<^bsub>o\<^esub>)" + using \\<^sub>B_simps by simp + +lemma propositional_principle_of_simplification_is_tautology: + shows "is_tautology (p\<^bsub>o\<^esub> \\<^sup>\ (r\<^bsub>o\<^esub> \\<^sup>\ p\<^bsub>o\<^esub>))" + using \\<^sub>B_simps by simp + +lemma closed_pwff_denotation_uniqueness: + assumes "A \ pwffs" and "free_vars A = {}" + obtains b where "\\. is_tv_assignment \ \ \\<^sub>B \ A = b" + using assms + by (meson closed_pwff_is_meaningful_regardless_of_assignment pwff_denotation_function_existence) + +lemma pwff_substitution_simps: + shows "\<^bold>S {(p, o) \ A} T\<^bsub>o\<^esub> = T\<^bsub>o\<^esub>" + and "\<^bold>S {(p, o) \ A} F\<^bsub>o\<^esub> = F\<^bsub>o\<^esub>" + and "\<^bold>S {(p, o) \ A} (p'\<^bsub>o\<^esub>) = (if p = p' then A else (p'\<^bsub>o\<^esub>))" + and "\<^bold>S {(p, o) \ A} (\\<^sup>\ B) = \\<^sup>\ (\<^bold>S {(p, o) \ A} B)" + and "\<^bold>S {(p, o) \ A} (B \\<^sup>\ C) = (\<^bold>S {(p, o) \ A} B) \\<^sup>\ (\<^bold>S {(p, o) \ A} C)" + and "\<^bold>S {(p, o) \ A} (B \\<^sup>\ C) = (\<^bold>S {(p, o) \ A} B) \\<^sup>\ (\<^bold>S {(p, o) \ A} C)" + and "\<^bold>S {(p, o) \ A} (B \\<^sup>\ C) = (\<^bold>S {(p, o) \ A} B) \\<^sup>\ (\<^bold>S {(p, o) \ A} C)" + and "\<^bold>S {(p, o) \ A} (B \\<^sup>\ C) = (\<^bold>S {(p, o) \ A} B) \\<^sup>\ (\<^bold>S {(p, o) \ A} C)" + by simp_all + +lemma pwff_substitution_in_pwffs: + assumes "A \ pwffs" and "B \ pwffs" + shows "\<^bold>S {(p, o) \ A} B \ pwffs" +using assms(2) proof induction + case T_pwff + then show ?case + using pwffs.T_pwff by simp +next + case F_pwff + then show ?case + using pwffs.F_pwff by simp +next + case (var_pwff p) + from assms(1) show ?case + using pwffs.var_pwff by simp +next + case (neg_pwff A) + then show ?case + using pwff_substitution_simps(4) and pwffs.neg_pwff by simp +next + case (conj_pwff A B) + then show ?case + using pwff_substitution_simps(5) and pwffs.conj_pwff by simp +next + case (disj_pwff A B) + then show ?case + using pwff_substitution_simps(6) and pwffs.disj_pwff by simp +next + case (imp_pwff A B) + then show ?case + using pwff_substitution_simps(7) and pwffs.imp_pwff by simp +next + case (eqv_pwff A B) + then show ?case + using pwff_substitution_simps(8) and pwffs.eqv_pwff by simp +qed + +lemma pwff_substitution_denotation: + assumes "A \ pwffs" and "B \ pwffs" + and "is_tv_assignment \" + shows "\\<^sub>B \ (\<^bold>S {(p, o) \ A} B) = \\<^sub>B (\(p := \\<^sub>B \ A)) B" +proof - + from assms(1,3) have "is_tv_assignment (\(p := \\<^sub>B \ A))" + using \\<^sub>B_graph_denotation_is_truth_value[OF \\<^sub>B_graph_\\<^sub>B] by simp + with assms(2,1,3) show ?thesis + using \\<^sub>B_simps and pwff_substitution_in_pwffs by induction auto +qed + +lemma pwff_substitution_tautology_preservation: + assumes "is_tautology B" and "A \ pwffs" + and "(p, o) \ free_vars B" + shows "is_tautology (\<^bold>S {(p, o) \ A} B)" +proof (safe, fold is_tv_assignment_def) + from assms(1,2) show "\<^bold>S {(p, o) \ A} B \ pwffs" + using pwff_substitution_in_pwffs by blast +next + fix \ + assume "is_tv_assignment \" + with assms(1,2) have "\\<^sub>B \ (\<^bold>S {(p, o) \ A} B) = \\<^sub>B (\(p := \\<^sub>B \ A)) B" + using pwff_substitution_denotation by blast + moreover from \is_tv_assignment \\ and assms(2) have "is_tv_assignment (\(p := \\<^sub>B \ A))" + using \\<^sub>B_graph_denotation_is_truth_value[OF \\<^sub>B_graph_\\<^sub>B] by simp + with assms(1) have "\\<^sub>B (\(p := \\<^sub>B \ A)) B = \<^bold>T" + by fastforce + ultimately show "\\<^sub>B \ \<^bold>S {(p, o) \ A} B = \<^bold>T" + by (simp only:) +qed + +lemma closed_pwff_substitution_free_vars: + assumes "A \ pwffs" and "B \ pwffs" + and "free_vars A = {}" + and "(p, o) \ free_vars B" + shows "free_vars (\<^bold>S {(p, o) \ A} B) = free_vars B - {(p, o)}" (is \free_vars (\<^bold>S ?\ B) = _\) +using assms(2,4) proof induction + case (conj_pwff C D) + have "free_vars (\<^bold>S ?\ (C \\<^sup>\ D)) = free_vars ((\<^bold>S ?\ C) \\<^sup>\ (\<^bold>S ?\ D))" + by simp + also have "\ = free_vars (\<^bold>S ?\ C) \ free_vars (\<^bold>S ?\ D)" + by (fact conj_fv) + finally have *: "free_vars (\<^bold>S ?\ (C \\<^sup>\ D)) = free_vars (\<^bold>S ?\ C) \ free_vars (\<^bold>S ?\ D)" . + from conj_pwff.prems consider + (a) "(p, o) \ free_vars C" and "(p, o) \ free_vars D" + | (b) "(p, o) \ free_vars C" and "(p, o) \ free_vars D" + | (c) "(p, o) \ free_vars C" and "(p, o) \ free_vars D" + by auto + from this and * and conj_pwff.IH show ?case + using free_var_singleton_substitution_neutrality by cases auto +next + case (disj_pwff C D) + have "free_vars (\<^bold>S ?\ (C \\<^sup>\ D)) = free_vars ((\<^bold>S ?\ C) \\<^sup>\ (\<^bold>S ?\ D))" + by simp + also have "\ = free_vars (\<^bold>S ?\ C) \ free_vars (\<^bold>S ?\ D)" + by (fact disj_fv) + finally have *: "free_vars (\<^bold>S ?\ (C \\<^sup>\ D)) = free_vars (\<^bold>S ?\ C) \ free_vars (\<^bold>S ?\ D)" . + from disj_pwff.prems consider + (a) "(p, o) \ free_vars C" and "(p, o) \ free_vars D" + | (b) "(p, o) \ free_vars C" and "(p, o) \ free_vars D" + | (c) "(p, o) \ free_vars C" and "(p, o) \ free_vars D" + by auto + from this and * and disj_pwff.IH show ?case + using free_var_singleton_substitution_neutrality by cases auto +next + case (imp_pwff C D) + have "free_vars (\<^bold>S ?\ (C \\<^sup>\ D)) = free_vars ((\<^bold>S ?\ C) \\<^sup>\ (\<^bold>S ?\ D))" + by simp + also have "\ = free_vars (\<^bold>S ?\ C) \ free_vars (\<^bold>S ?\ D)" + by (fact imp_fv) + finally have *: "free_vars (\<^bold>S ?\ (C \\<^sup>\ D)) = free_vars (\<^bold>S ?\ C) \ free_vars (\<^bold>S ?\ D)" . + from imp_pwff.prems consider + (a) "(p, o) \ free_vars C" and "(p, o) \ free_vars D" + | (b) "(p, o) \ free_vars C" and "(p, o) \ free_vars D" + | (c) "(p, o) \ free_vars C" and "(p, o) \ free_vars D" + by auto + from this and * and imp_pwff.IH show ?case + using free_var_singleton_substitution_neutrality by cases auto +next + case (eqv_pwff C D) + have "free_vars (\<^bold>S ?\ (C \\<^sup>\ D)) = free_vars ((\<^bold>S ?\ C) \\<^sup>\ (\<^bold>S ?\ D))" + by simp + also have "\ = free_vars (\<^bold>S ?\ C) \ free_vars (\<^bold>S ?\ D)" + by (fact eqv_fv) + finally have *: "free_vars (\<^bold>S ?\ (C \\<^sup>\ D)) = free_vars (\<^bold>S ?\ C) \ free_vars (\<^bold>S ?\ D)" . + from eqv_pwff.prems consider + (a) "(p, o) \ free_vars C" and "(p, o) \ free_vars D" + | (b) "(p, o) \ free_vars C" and "(p, o) \ free_vars D" + | (c) "(p, o) \ free_vars C" and "(p, o) \ free_vars D" + by auto + from this and * and eqv_pwff.IH show ?case + using free_var_singleton_substitution_neutrality by cases auto +qed (use assms(3) in \force+\) + +text \Substitution in a pwff:\ + +definition is_pwff_substitution where + [iff]: "is_pwff_substitution \ \ is_substitution \ \ (\(x, \) \ fmdom' \. \ = o)" + +text \Tautologous pwff:\ + +definition is_tautologous :: "form \ bool" where + [iff]: "is_tautologous B \ (\\ A. is_tautology A \ is_pwff_substitution \ \ B = \<^bold>S \ A)" + +lemma tautologous_is_wffo: + assumes "is_tautologous A" + shows "A \ wffs\<^bsub>o\<^esub>" + using assms and substitution_preserves_typing and tautology_is_wffo by blast + +lemma implication_reflexivity_is_tautologous: + assumes "A \ wffs\<^bsub>o\<^esub>" + shows "is_tautologous (A \\<^sup>\ A)" +proof - + let ?\ = "{(\, o) \ A}" + have "is_tautology (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)" + by (fact propositional_implication_reflexivity_is_tautology) + moreover have "is_pwff_substitution ?\" + using assms by auto + moreover have "A \\<^sup>\ A = \<^bold>S ?\ (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)" + by simp + ultimately show ?thesis + by blast +qed + +lemma principle_of_simplification_is_tautologous: + assumes "A \ wffs\<^bsub>o\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" + shows "is_tautologous (A \\<^sup>\ (B \\<^sup>\ A))" +proof - + let ?\ = "{(\, o) \ A, (\, o) \ B}" + have "is_tautology (\\<^bsub>o\<^esub> \\<^sup>\ (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>))" + by (fact propositional_principle_of_simplification_is_tautology) + moreover have "is_pwff_substitution ?\" + using assms by auto + moreover have "A \\<^sup>\ (B \\<^sup>\ A) = \<^bold>S ?\ (\\<^bsub>o\<^esub> \\<^sup>\ (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>))" + by simp + ultimately show ?thesis + by blast +qed + +lemma pseudo_modus_tollens_is_tautologous: + assumes "A \ wffs\<^bsub>o\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" + shows "is_tautologous ((A \\<^sup>\ \\<^sup>\ B) \\<^sup>\ (B \\<^sup>\ \\<^sup>\ A))" +proof - + let ?\ = "{(\, o) \ A, (\, o) \ B}" + have "is_tautology ((\\<^bsub>o\<^esub> \\<^sup>\ \\<^sup>\ \\<^bsub>o\<^esub>) \\<^sup>\ (\\<^bsub>o\<^esub> \\<^sup>\ \\<^sup>\ \\<^bsub>o\<^esub>))" + using \\<^sub>B_simps by (safe, fold is_tv_assignment_def, simp only:) simp + moreover have "is_pwff_substitution ?\" + using assms by auto + moreover have "(A \\<^sup>\ \\<^sup>\ B) \\<^sup>\ (B \\<^sup>\ \\<^sup>\ A) = \<^bold>S ?\ ((\\<^bsub>o\<^esub> \\<^sup>\ \\<^sup>\ \\<^bsub>o\<^esub>) \\<^sup>\ (\\<^bsub>o\<^esub> \\<^sup>\ \\<^sup>\ \\<^bsub>o\<^esub>))" + by simp + ultimately show ?thesis + by blast +qed + +end diff --git a/thys/Q0_Metatheory/ROOT b/thys/Q0_Metatheory/ROOT new file mode 100644 --- /dev/null +++ b/thys/Q0_Metatheory/ROOT @@ -0,0 +1,21 @@ +chapter AFP + +session Q0_Metatheory (AFP) = "HOL-Library" + + options [timeout = 600] + sessions + ZFC_in_HOL + "Finite-Map-Extras" + theories + Utilities + Syntax + Boolean_Algebra + Propositional_Wff + Proof_System + Elementary_Logic + Semantics + Soundness + Consistency + document_files + "root.tex" + "root.bib" + diff --git a/thys/Q0_Metatheory/Semantics.thy b/thys/Q0_Metatheory/Semantics.thy new file mode 100644 --- /dev/null +++ b/thys/Q0_Metatheory/Semantics.thy @@ -0,0 +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"}: +\ + +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/Soundness.thy b/thys/Q0_Metatheory/Soundness.thy new file mode 100644 --- /dev/null +++ b/thys/Q0_Metatheory/Soundness.thy @@ -0,0 +1,1561 @@ +section \Soundness\ + +theory Soundness + imports + Elementary_Logic + Semantics +begin + +no_notation funcset (infixr "\" 60) +notation funcset (infixr "\" 60) + +subsection \Proposition 5400\ + +proposition (in general_model) prop_5400: + assumes "A \ wffs\<^bsub>\\<^esub>" + and "\ \ \" and "\ \ \" + and "\v \ free_vars A. \ v = \ v" + shows "\ \ A = \ \ A" +proof - + from assms(1) show ?thesis + using assms(2,3,4) proof (induction A arbitrary: \ \) + case (var_is_wff \ x) + have "(x, \) \ free_vars (x\<^bsub>\\<^esub>)" + by simp + from assms(1) and var_is_wff.prems(1) have "\ \ (x\<^bsub>\\<^esub>) = \ (x, \)" + using \_is_wff_denotation_function by fastforce + also from \(x, \) \ free_vars (x\<^bsub>\\<^esub>)\ and var_is_wff.prems(3) have "\ = \ (x, \)" + by (simp only:) + also from assms(1) and var_is_wff.prems(2) have "\ = \ \ (x\<^bsub>\\<^esub>)" + using \_is_wff_denotation_function by fastforce + finally show ?case . + next + case (con_is_wff \ c) + from assms(1) and con_is_wff.prems(1) have "\ \ (\c\\<^bsub>\\<^esub>) = \ (c, \)" + using \_is_wff_denotation_function by fastforce + also from assms(1) and con_is_wff.prems(2) have "\ = \ \ (\c\\<^bsub>\\<^esub>)" + using \_is_wff_denotation_function by fastforce + finally show ?case . + next + case (app_is_wff \ \ A B) + have "free_vars (A \ B) = free_vars A \ free_vars B" + by simp + with app_is_wff.prems(3) + have "\v \ free_vars A. \ v = \ v" and "\v \ free_vars B. \ v = \ v" + by blast+ + with app_is_wff.IH and app_is_wff.prems(1,2) have "\ \ A = \ \ A" and "\ \ B = \ \ B" + by blast+ + from assms(1) and app_is_wff.prems(1) and app_is_wff.hyps have "\ \ (A \ B) = \ \ A \ \ \ B" + using \_is_wff_denotation_function by fastforce + also from \\ \ A = \ \ A\ and \\ \ B = \ \ B\ have "\ = \ \ A \ \ \ B" + by (simp only:) + also from assms(1) and app_is_wff.prems(2) and app_is_wff.hyps have "\ = \ \ (A \ B)" + using \_is_wff_denotation_function by fastforce + finally show ?case . + next + case (abs_is_wff \ A \ x) + have "free_vars (\x\<^bsub>\\<^esub>. A) = free_vars A - {(x, \)}" + by simp + with abs_is_wff.prems(3) have "\v \ free_vars A. v \ (x, \)\ \ v = \ v" + by blast + then have "\v \ free_vars A. (\((x, \) := z)) v = (\((x, \) := z)) v" if "z \ elts (\ \)" for z + by simp + moreover from abs_is_wff.prems(1,2) + have "\x' \'. (\((x, \) := z)) (x', \') \ elts (\ \')" (* needs \-conversion *) + and "\x' \'. (\((x, \) := z)) (x', \') \ elts (\ \')" (* needs \-conversion *) + if "z \ elts (\ \)" for z + using that by force+ + ultimately have \_\_\_eq: "\ (\((x, \) := z)) A = \ (\((x, \) := z)) A" if "z \ elts (\ \)" for z + using abs_is_wff.IH and that by simp + from assms(1) and abs_is_wff.prems(1) and abs_is_wff.hyps + have "\ \ (\x\<^bsub>\\<^esub>. A) = (\<^bold>\z \<^bold>: \ \\<^bold>. \ (\((x, \) := z)) A)" + using wff_abs_denotation[OF \_is_wff_denotation_function] by simp + also from \_\_\_eq have "\ = (\<^bold>\z \<^bold>: \ \\<^bold>. \ (\((x, \) := z)) A)" + by (fact vlambda_extensionality) + also from assms(1) and abs_is_wff.hyps have "\ = \ \ (\x\<^bsub>\\<^esub>. A)" + using wff_abs_denotation[OF \_is_wff_denotation_function abs_is_wff.prems(2)] by simp + finally show ?case . + qed +qed + +corollary (in general_model) closed_wff_is_meaningful_regardless_of_assignment: + assumes "is_closed_wff_of_type A \" + and "\ \ \" and "\ \ \" + shows "\ \ A = \ \ A" + using assms and prop_5400 by blast + +subsection \Proposition 5401\ + +lemma (in general_model) prop_5401_a: + assumes "\ \ \" + and "A \ wffs\<^bsub>\\<^esub>" + and "B \ wffs\<^bsub>\\<^esub>" + shows "\ \ ((\x\<^bsub>\\<^esub>. B) \ A) = \ (\((x, \) := \ \ A)) B" +proof - + from assms(2,3) have "\x\<^bsub>\\<^esub>. B \ wffs\<^bsub>\\\\<^esub>" + by blast + with assms(1,2) have "\ \ ((\x\<^bsub>\\<^esub>. B) \ A) = \ \ (\x\<^bsub>\\<^esub>. B) \ \ \ A" + using \_is_wff_denotation_function by blast + also from assms(1,3) have "\ = app (\<^bold>\z \<^bold>: \ \\<^bold>. \ (\((x, \) := z)) B) (\ \ A)" + using wff_abs_denotation[OF \_is_wff_denotation_function] by simp + also from assms(1,2) have "\ = \ (\((x, \) := \ \ A)) B" + using \_is_wff_denotation_function by auto + finally show ?thesis . +qed + +lemma (in general_model) prop_5401_b: + assumes "\ \ \" + and "A \ wffs\<^bsub>\\<^esub>" + and "B \ wffs\<^bsub>\\<^esub>" + shows "\ \ (A =\<^bsub>\\<^esub> B) = \<^bold>T \ \ \ A = \ \ B" +proof - + from assms have "{\ \ A, \ \ B} \ elts (\ \)" + using \_is_wff_denotation_function by auto + have "\ \ (A =\<^bsub>\\<^esub> B) = \ \ (Q\<^bsub>\\<^esub> \ A \ B)" + by simp + also from assms have "\ = \ \ (Q\<^bsub>\\<^esub> \ A) \ \ \ B" + using \_is_wff_denotation_function by blast + also from assms have "\ = \ \ (Q\<^bsub>\\<^esub>) \ \ \ A \ \ \ B" + using Q_wff and wff_app_denotation[OF \_is_wff_denotation_function] by fastforce + also from assms(1) have "\ = (q\<^bsub>\\<^esub>) \ \ \ A \ \ \ B" + using Q_denotation and \_is_wff_denotation_function by fastforce + also from \{\ \ A, \ \ B} \ elts (\ \)\ have "\ = \<^bold>T \ \ \ A = \ \ B" + using q_is_equality by simp + finally show ?thesis . +qed + +corollary (in general_model) prop_5401_b': + assumes "\ \ \" + and "A \ wffs\<^bsub>o\<^esub>" + and "B \ wffs\<^bsub>o\<^esub>" + shows "\ \ (A \\<^sup>\ B) = \<^bold>T \ \ \ A = \ \ B" + using assms and prop_5401_b by auto + +lemma (in general_model) prop_5401_c: + assumes "\ \ \" + shows "\ \ T\<^bsub>o\<^esub> = \<^bold>T" +proof - + have "Q\<^bsub>o\<^esub> \ wffs\<^bsub>o\o\o\<^esub>" + by blast + moreover have "\ \ T\<^bsub>o\<^esub> = \ \ (Q\<^bsub>o\<^esub> =\<^bsub>o\o\o\<^esub> Q\<^bsub>o\<^esub>)" + unfolding true_def .. + ultimately have "\ = \<^bold>T \ \ \ (Q\<^bsub>o\<^esub>) = \ \ (Q\<^bsub>o\<^esub>)" + using prop_5401_b and assms by blast + then show ?thesis + by simp +qed + +lemma (in general_model) prop_5401_d: + assumes "\ \ \" + shows "\ \ F\<^bsub>o\<^esub> = \<^bold>F" +proof - + have "\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub> \ wffs\<^bsub>o\o\<^esub>" and "\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub> \ wffs\<^bsub>o\o\<^esub>" + by blast+ + moreover have "\ \ F\<^bsub>o\<^esub> = \ \ (\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub> =\<^bsub>o\o\<^esub> \\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>)" + unfolding false_def .. + ultimately have "\ \ F\<^bsub>o\<^esub> = \<^bold>T \ \ \ (\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>) = \ \ (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>)" + using prop_5401_b and assms by simp + moreover have "\ \ (\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>) \ \ \ (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>)" + proof - + have "\ \ (\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>) = (\<^bold>\z \<^bold>: \ o\<^bold>. \<^bold>T)" + proof - + from assms have T_denotation: "\ (\((\, o) := z)) T\<^bsub>o\<^esub> = \<^bold>T" if "z \ elts (\ o)" for z + using prop_5401_c and that by simp + from assms have "\ \ (\\\<^bsub>o\<^esub>. T\<^bsub>o\<^esub>) = (\<^bold>\z \<^bold>: \ o\<^bold>. \ (\((\, o) := z)) T\<^bsub>o\<^esub>)" + using wff_abs_denotation[OF \_is_wff_denotation_function] by blast + also from assms and T_denotation have "\ = (\<^bold>\z \<^bold>: \ o\<^bold>. \<^bold>T)" + using vlambda_extensionality by fastforce + finally show ?thesis . + qed + moreover have "\ \ (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>) = (\<^bold>\z \<^bold>: \ o\<^bold>. z)" + proof - + from assms have \_denotation: "\ (\((\, o) := z)) (\\<^bsub>o\<^esub>) = z" if "z \ elts (\ o)" for z + using that and \_is_wff_denotation_function by auto + from assms have "\ \ (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>) = (\<^bold>\z \<^bold>: \ o\<^bold>. \ (\((\, o) := z)) (\\<^bsub>o\<^esub>))" + using wff_abs_denotation[OF \_is_wff_denotation_function] by blast + also from \_denotation have "\ = (\<^bold>\z \<^bold>: (\ o)\<^bold>. z)" + using vlambda_extensionality by fastforce + finally show ?thesis . + qed + moreover have "(\<^bold>\z \<^bold>: \ o\<^bold>. \<^bold>T) \ (\<^bold>\z \<^bold>: \ o\<^bold>. z)" + proof - + from assms(1) have "(\<^bold>\z \<^bold>: \ o\<^bold>. \<^bold>T) \ \<^bold>F = \<^bold>T" + by (simp add: truth_values_domain_def) + moreover from assms(1) have "(\<^bold>\z \<^bold>: \ o\<^bold>. z) \ \<^bold>F = \<^bold>F" + by (simp add: truth_values_domain_def) + ultimately show ?thesis + by (auto simp add: inj_eq) + qed + ultimately show ?thesis + by simp + qed + moreover from assms have "\ \ F\<^bsub>o\<^esub> \ elts (\ o)" + using false_wff and \_is_wff_denotation_function by fast + ultimately show ?thesis + using assms(1) by (simp add: truth_values_domain_def) +qed + +lemma (in general_model) prop_5401_e: + assumes "\ \ \" + and "{x, y} \ elts (\ o)" + shows "\ \ (\\<^bsub>o\o\o\<^esub>) \ x \ y = (if x = \<^bold>T \ y = \<^bold>T then \<^bold>T else \<^bold>F)" +proof - + let ?B\<^sub>l\<^sub>e\<^sub>q = "\\\<^bsub>o\o\o\<^esub>. \\<^bsub>o\o\o\<^esub> \ T\<^bsub>o\<^esub> \ T\<^bsub>o\<^esub>" + let ?B\<^sub>r\<^sub>e\<^sub>q = "\\\<^bsub>o\o\o\<^esub>. \\<^bsub>o\o\o\<^esub> \ \\<^bsub>o\<^esub> \ \\<^bsub>o\<^esub>" + let ?B\<^sub>e\<^sub>q = "?B\<^sub>l\<^sub>e\<^sub>q =\<^bsub>(o\o\o)\o\<^esub> ?B\<^sub>r\<^sub>e\<^sub>q" + let ?B\<^sub>\ = "\\\<^bsub>o\<^esub>. ?B\<^sub>e\<^sub>q" + let ?B\<^sub>\ = "\\\<^bsub>o\<^esub>. ?B\<^sub>\" + let ?\' = "\((\, o) := x, (\, o) := y)" + let ?\'' = "\g. ?\'((\, o\o\o) := g)" + have "\\<^bsub>o\o\o\<^esub> \ T\<^bsub>o\<^esub> \ wffs\<^bsub>o\o\<^esub>" + by blast + have "\\<^bsub>o\o\o\<^esub> \ T\<^bsub>o\<^esub> \ T\<^bsub>o\<^esub> \ wffs\<^bsub>o\<^esub>" and "\\<^bsub>o\o\o\<^esub> \ \\<^bsub>o\<^esub> \ \\<^bsub>o\<^esub> \ wffs\<^bsub>o\<^esub>" + by blast+ + have "?B\<^sub>l\<^sub>e\<^sub>q \ wffs\<^bsub>(o\o\o)\o\<^esub>" and "?B\<^sub>r\<^sub>e\<^sub>q \ wffs\<^bsub>(o\o\o)\o\<^esub>" + by blast+ + then have "?B\<^sub>e\<^sub>q \ wffs\<^bsub>o\<^esub>" and "?B\<^sub>\ \ wffs\<^bsub>o\o\<^esub>" and "?B\<^sub>\ \ wffs\<^bsub>o\o\o\<^esub>" + by blast+ + have "\ \ (\\<^bsub>o\o\o\<^esub>) \ x \ y = \ \ ?B\<^sub>\ \ x \ y" + by simp + also from assms and \?B\<^sub>\ \ wffs\<^bsub>o\o\<^esub>\ have "\ = \ (\((\, o) := x)) ?B\<^sub>\ \ y" + using mixed_beta_conversion by simp + also from assms and \?B\<^sub>e\<^sub>q \ wffs\<^bsub>o\<^esub>\ have "\ = \ ?\' ?B\<^sub>e\<^sub>q" + using mixed_beta_conversion by simp + finally have "\ \ (\\<^bsub>o\o\o\<^esub>) \ x \ y = \<^bold>T \ \ ?\' ?B\<^sub>l\<^sub>e\<^sub>q = \ ?\' ?B\<^sub>r\<^sub>e\<^sub>q" + using assms and \?B\<^sub>l\<^sub>e\<^sub>q \ wffs\<^bsub>(o\o\o)\o\<^esub>\ and \?B\<^sub>r\<^sub>e\<^sub>q \ wffs\<^bsub>(o\o\o)\o\<^esub>\ and prop_5401_b + by simp + also have "\ \ (\<^bold>\g \<^bold>: \ (o\o\o)\<^bold>. g \ \<^bold>T \ \<^bold>T) = (\<^bold>\g \<^bold>: \ (o\o\o)\<^bold>. g \ x \ y)" + proof - + have leq: "\ ?\' ?B\<^sub>l\<^sub>e\<^sub>q = (\<^bold>\g \<^bold>: \ (o\o\o)\<^bold>. g \ \<^bold>T \ \<^bold>T)" + and req: "\ ?\' ?B\<^sub>r\<^sub>e\<^sub>q = (\<^bold>\g \<^bold>: \ (o\o\o)\<^bold>. g \ x \ y)" + proof - + from assms(1,2) have is_assg_\'': "?\'' g \ \" if "g \ elts (\ (o\o\o))" for g + using that by auto + have side_eq_denotation: + "\ ?\' (\\\<^bsub>o\o\o\<^esub>. \\<^bsub>o\o\o\<^esub> \ A \ B) = (\<^bold>\g \<^bold>: \ (o\o\o)\<^bold>. g \ \ (?\'' g) A \ \ (?\'' g) B)" + if "A \ wffs\<^bsub>o\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" for A and B + proof - + from that have "\\<^bsub>o\o\o\<^esub> \ A \ B \ wffs\<^bsub>o\<^esub>" + by blast + have "\ (?\'' g) (\\<^bsub>o\o\o\<^esub> \ A \ B) = g \ \ (?\'' g) A \ \ (?\'' g) B" + if "g \ elts (\ (o\o\o))" for g + proof - + from \A \ wffs\<^bsub>o\<^esub>\ have "\\<^bsub>o\o\o\<^esub> \ A \ wffs\<^bsub>o\o\<^esub>" + by blast + with that and is_assg_\'' and \B \ wffs\<^bsub>o\<^esub>\ have " + \ (?\'' g) (\\<^bsub>o\o\o\<^esub> \ A \ B) = \ (?\'' g) (\\<^bsub>o\o\o\<^esub> \ A) \ \ (?\'' g) B" + using wff_app_denotation[OF \_is_wff_denotation_function] by simp + also from that and \A \ wffs\<^bsub>o\<^esub>\ and is_assg_\'' have " + \ = \ (?\'' g) (\\<^bsub>o\o\o\<^esub>) \ \ (?\'' g) A \ \ (?\'' g) B" + by (metis \_is_wff_denotation_function wff_app_denotation wffs_of_type_intros(1)) + finally show ?thesis + using that and is_assg_\'' and \_is_wff_denotation_function by auto + qed + moreover from assms have "is_assignment ?\'" + by auto + with \\\<^bsub>o\o\o\<^esub> \ A \ B \ wffs\<^bsub>o\<^esub>\ have " + \ ?\' (\\\<^bsub>o\o\o\<^esub>. \\<^bsub>o\o\o\<^esub> \ A \ B) = (\<^bold>\g \<^bold>: \ (o\o\o)\<^bold>. \ (?\'' g) (\\<^bsub>o\o\o\<^esub> \ A \ B))" + using wff_abs_denotation[OF \_is_wff_denotation_function] by simp + ultimately show ?thesis + using vlambda_extensionality by fastforce + qed + \ \Proof of \leq\:\ + show "\ ?\' ?B\<^sub>l\<^sub>e\<^sub>q = (\<^bold>\g \<^bold>: \ (o\o\o)\<^bold>. g \ \<^bold>T \ \<^bold>T)" + proof - + have "\ (?\'' g) T\<^bsub>o\<^esub> = \<^bold>T" if "g \ elts (\ (o\o\o))" for g + using that and is_assg_\'' and prop_5401_c by simp + then show ?thesis + using side_eq_denotation and true_wff and vlambda_extensionality by fastforce + qed + \ \Proof of \req\:\ + show "\ ?\' ?B\<^sub>r\<^sub>e\<^sub>q = (\<^bold>\g \<^bold>: \ (o\o\o)\<^bold>. g \ x \ y)" + proof - + from is_assg_\'' have "\ (?\'' g) (\\<^bsub>o\<^esub>) = x" and "\ (?\'' g) (\\<^bsub>o\<^esub>) = y" + if "g \ elts (\ (o\o\o))" for g + using that and \_is_wff_denotation_function by auto + with side_eq_denotation show ?thesis + using wffs_of_type_intros(1) and vlambda_extensionality by fastforce + qed + qed + then show ?thesis + by auto + qed + also have "\ \ (\g \ elts (\ (o\o\o)). g \ \<^bold>T \ \<^bold>T = g \ x \ y)" + using vlambda_extensionality and VLambda_eq_D2 by fastforce + finally have and_eqv: " + \ \ (\\<^bsub>o\o\o\<^esub>) \ x \ y = \<^bold>T \ (\g \ elts (\ (o\o\o)). g \ \<^bold>T \ \<^bold>T = g \ x \ y)" + by blast + then show ?thesis + proof - + from assms(1,2) have is_assg_1: "\((\, o) := \<^bold>T) \ \" + by (simp add: truth_values_domain_def) + then have is_assg_2: "\((\, o) := \<^bold>T, (\, o) := \<^bold>T) \ \" + unfolding is_assignment_def by (metis fun_upd_apply prod.sel(2)) + from assms consider (a) "x = \<^bold>T \ y = \<^bold>T" | (b) "x \ \<^bold>T" | (c) "y \ \<^bold>T" + by blast + then show ?thesis + proof cases + case a + then have "g \ \<^bold>T \ \<^bold>T = g \ x \ y" if "g \ elts (\ (o\o\o))" for g + by simp + with a and and_eqv show ?thesis + by simp + next + case b + let ?g_witness = "\\\<^bsub>o\<^esub>. \\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>" + have "\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub> \ wffs\<^bsub>o\o\<^esub>" + by blast + then have "is_closed_wff_of_type ?g_witness (o\o\o)" + by force + moreover from assms have is_assg_\': "?\' \ \" + by simp + ultimately have "\ \ ?g_witness \ \<^bold>T \ \<^bold>T = \ ?\' ?g_witness \ \<^bold>T \ \<^bold>T" + using assms(1) and closed_wff_is_meaningful_regardless_of_assignment by metis + also from assms and \\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub> \ wffs\<^bsub>o\o\<^esub>\ have " + \ ?\' ?g_witness \ \<^bold>T \ \<^bold>T = \ (?\'((\, o) := \<^bold>T)) (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>) \ \<^bold>T" + using mixed_beta_conversion and truth_values_domain_def by auto + also from assms(1) and \\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub> \ wffs\<^bsub>o\o\<^esub>\ and is_assg_1 and calculation have " + \ = \ (?\'((\, o) := \<^bold>T, (\, o) := \<^bold>T)) (\\<^bsub>o\<^esub>)" + using mixed_beta_conversion and is_assignment_def + by (metis fun_upd_same fun_upd_twist fun_upd_upd wffs_of_type_intros(1)) + also have "\ = \<^bold>T" + using is_assg_2 and \_is_wff_denotation_function by fastforce + finally have "\ \ ?g_witness \ \<^bold>T \ \<^bold>T = \<^bold>T" . + with b have "\ \ ?g_witness \ \<^bold>T \ \<^bold>T \ x" + by blast + moreover have "x = \ \ ?g_witness \ x \ y" + proof - + from is_assg_\' have "x = \ ?\' (\\<^bsub>o\<^esub>)" + using \_is_wff_denotation_function by auto + also from assms(2) and is_assg_\' have "\ = \ ?\' (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>) \ y" + using wffs_of_type_intros(1)[where x = \ and \ = o] + by (simp add: mixed_beta_conversion \_is_wff_denotation_function) + also from assms(2) have "\ = \ ?\' ?g_witness \ x \ y" + using is_assg_\' and \\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub> \ wffs\<^bsub>o\o\<^esub>\ + by (simp add: mixed_beta_conversion fun_upd_twist) + also from assms(1,2) have "\ = \ \ ?g_witness \ x \ y" + using is_assg_\' and \is_closed_wff_of_type ?g_witness (o\o\o)\ + and closed_wff_is_meaningful_regardless_of_assignment by metis + finally show ?thesis . + qed + moreover from assms(1,2) have "\ \ ?g_witness \ elts (\ (o\o\o))" + using \is_closed_wff_of_type ?g_witness (o\o\o)\ and \_is_wff_denotation_function by simp + ultimately have "\g \ elts (\ (o\o\o)). g \ \<^bold>T \ \<^bold>T \ g \ x \ y" + by auto + moreover from assms have "\ \ (\\<^bsub>o\o\o\<^esub>) \ x \ y \ elts (\ o)" + by (rule fully_applied_conj_fun_is_domain_respecting) + ultimately have "\ \ (\\<^bsub>o\o\o\<^esub>) \ x \ y = \<^bold>F" + using and_eqv and truth_values_domain_def by fastforce + with b show ?thesis + by simp + next + case c + let ?g_witness = "\\\<^bsub>o\<^esub>. \\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>" + have "\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub> \ wffs\<^bsub>o\o\<^esub>" + by blast + then have "is_closed_wff_of_type ?g_witness (o\o\o)" + by force + moreover from assms(1,2) have is_assg_\': "?\' \ \" + by simp + ultimately have "\ \ ?g_witness \ \<^bold>T \ \<^bold>T = \ ?\' ?g_witness \ \<^bold>T \ \<^bold>T" + using assms(1) and closed_wff_is_meaningful_regardless_of_assignment by metis + also from is_assg_1 and is_assg_\' have "\ = \ (?\'((\, o) := \<^bold>T)) (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>) \ \<^bold>T" + using \\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub> \ wffs\<^bsub>o\o\<^esub>\ and mixed_beta_conversion and truth_values_domain_def by auto + also from assms(1) and \\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub> \ wffs\<^bsub>o\o\<^esub>\ and is_assg_1 and calculation have " + \ = \ (?\'((\, o) := \<^bold>T, (\, o) := \<^bold>T)) (\\<^bsub>o\<^esub>)" + using mixed_beta_conversion and is_assignment_def + by (metis fun_upd_same fun_upd_twist fun_upd_upd wffs_of_type_intros(1)) + also have "\ = \<^bold>T" + using is_assg_2 and \_is_wff_denotation_function by force + finally have "\ \ ?g_witness \ \<^bold>T \ \<^bold>T = \<^bold>T" . + with c have "\ \ ?g_witness \ \<^bold>T \ \<^bold>T \ y" + by blast + moreover have "y = \ \ ?g_witness \ x \ y" + proof - + from assms(2) and is_assg_\' have "y = \ ?\' (\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub>) \ y" + using wffs_of_type_intros(1)[where x = \ and \ = o] + and \_is_wff_denotation_function and mixed_beta_conversion by auto + also from assms(2) and \\\\<^bsub>o\<^esub>. \\<^bsub>o\<^esub> \ wffs\<^bsub>o\o\<^esub>\ have "\ = \ ?\' ?g_witness \ x \ y" + using is_assg_\' by (simp add: mixed_beta_conversion fun_upd_twist) + also from assms(1,2) have "\ = \ \ ?g_witness \ x \ y" + using is_assg_\' and \is_closed_wff_of_type ?g_witness (o\o\o)\ + and closed_wff_is_meaningful_regardless_of_assignment by metis + finally show ?thesis . + qed + moreover from assms(1) have "\ \ ?g_witness \ elts (\ (o\o\o))" + using \is_closed_wff_of_type ?g_witness (o\o\o)\ and \_is_wff_denotation_function by auto + ultimately have "\g \ elts (\ (o\o\o)). g \ \<^bold>T \ \<^bold>T \ g \ x \ y" + by auto + moreover from assms have "\ \ (\\<^bsub>o\o\o\<^esub>) \ x \ y \ elts (\ o)" + by (rule fully_applied_conj_fun_is_domain_respecting) + ultimately have "\ \ (\\<^bsub>o\o\o\<^esub>) \ x \ y = \<^bold>F" + using and_eqv and truth_values_domain_def by fastforce + with c show ?thesis + by simp + qed + qed +qed + +corollary (in general_model) prop_5401_e': + assumes "\ \ \" + and "A \ wffs\<^bsub>o\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" + shows "\ \ (A \\<^sup>\ B) = \ \ A \<^bold>\ \ \ B" +proof - + from assms have "{\ \ A, \ \ B} \ elts (\ o)" + using \_is_wff_denotation_function by simp + from assms(2) have "\\<^bsub>o\o\o\<^esub> \ A \ wffs\<^bsub>o\o\<^esub>" + by blast + have "\ \ (A \\<^sup>\ B) = \ \ (\\<^bsub>o\o\o\<^esub> \ A \ B)" + by simp + also from assms have "\ = \ \ (\\<^bsub>o\o\o\<^esub> \ A) \ \ \ B" + using \_is_wff_denotation_function and \\\<^bsub>o\o\o\<^esub> \ A \ wffs\<^bsub>o\o\<^esub>\ by blast + also from assms have "\ = \ \ (\\<^bsub>o\o\o\<^esub>) \ \ \ A \ \ \ B" + using \_is_wff_denotation_function and conj_fun_wff by fastforce + also from assms(1,2) have "\ = (if \ \ A = \<^bold>T \ \ \ B = \<^bold>T then \<^bold>T else \<^bold>F)" + using \{\ \ A, \ \ B} \ elts (\ o)\ and prop_5401_e by simp + also have "\ = \ \ A \<^bold>\ \ \ B" + using truth_values_domain_def and \{\ \ A, \ \ B} \ elts (\ o)\ by fastforce + finally show ?thesis . +qed + +lemma (in general_model) prop_5401_f: + assumes "\ \ \" + and "{x, y} \ elts (\ o)" + shows "\ \ (\\<^bsub>o\o\o\<^esub>) \ x \ y = (if x = \<^bold>T \ y = \<^bold>F then \<^bold>F else \<^bold>T)" +proof - + let ?\' = "\((\, o) := x, (\, o) := y)" + from assms(2) have "{x, y} \ elts \" + unfolding truth_values_domain_def . + have "(\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>) \ wffs\<^bsub>o\<^esub>" + by blast + then have "\\\<^bsub>o\<^esub>. (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>) \ wffs\<^bsub>o\o\<^esub>" + by blast + from assms have is_assg_\': "?\' \ \" + by simp + from assms(1) have "\ ?\' (\\<^bsub>o\<^esub>) = x" and "\ ?\' (\\<^bsub>o\<^esub>) = y" + using is_assg_\' and \_is_wff_denotation_function by force+ + have "\ \ (\\<^bsub>o\o\o\<^esub>) \ x \ y = \ \ (\\\<^bsub>o\<^esub>. \\\<^bsub>o\<^esub>. (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)) \ x \ y" + by simp + also from assms have "\ = \ (\((\, o) := x)) (\\\<^bsub>o\<^esub>. (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)) \ y" + using \\\\<^bsub>o\<^esub>. (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>) \ wffs\<^bsub>o\o\<^esub>\ and mixed_beta_conversion by simp + also from assms have "\ = \ ?\' (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)" + using mixed_beta_conversion and \(\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>) \ wffs\<^bsub>o\<^esub>\ by simp + finally have "\ \ (\\<^bsub>o\o\o\<^esub>) \ x \ y = \<^bold>T \ \ ?\' (\\<^bsub>o\<^esub>) = \ ?\' (\\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\<^esub>)" + using prop_5401_b'[OF is_assg_\'] and conj_op_wff and wffs_of_type_intros(1) by simp + also have "\ \ x = x \<^bold>\ y" + unfolding prop_5401_e'[OF is_assg_\' wffs_of_type_intros(1) wffs_of_type_intros(1)] + and \\ ?\' (\\<^bsub>o\<^esub>) = x\ and \\ ?\' (\\<^bsub>o\<^esub>) = y\ .. + also have "\ \ x = (if x = \<^bold>T \ y = \<^bold>T then \<^bold>T else \<^bold>F)" + using \{x, y} \ elts \\ by auto + also have "\ \ \<^bold>T = (if x = \<^bold>T \ y = \<^bold>F then \<^bold>F else \<^bold>T)" + using \{x, y} \ elts \\ by auto + finally show ?thesis + using assms and fully_applied_imp_fun_denotation_is_domain_respecting and tv_cases + and truth_values_domain_def by metis +qed + +corollary (in general_model) prop_5401_f': + assumes "\ \ \" + and "A \ wffs\<^bsub>o\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" + shows "\ \ (A \\<^sup>\ B) = \ \ A \<^bold>\ \ \ B" +proof - + from assms have "{\ \ A, \ \ B} \ elts (\ o)" + using \_is_wff_denotation_function by simp + from assms(2) have "\\<^bsub>o\o\o\<^esub> \ A \ wffs\<^bsub>o\o\<^esub>" + by blast + have "\ \ (A \\<^sup>\ B) = \ \ (\\<^bsub>o\o\o\<^esub> \ A \ B)" + by simp + also from assms(1,3) have "\ = \ \ (\\<^bsub>o\o\o\<^esub> \ A) \ \ \ B" + using \_is_wff_denotation_function and \\\<^bsub>o\o\o\<^esub> \ A \ wffs\<^bsub>o\o\<^esub>\ by blast + also from assms have "\ = \ \ (\\<^bsub>o\o\o\<^esub>) \ \ \ A \ \ \ B" + using \_is_wff_denotation_function and imp_fun_wff by fastforce + also from assms have "\ = (if \ \ A = \<^bold>T \ \ \ B = \<^bold>F then \<^bold>F else \<^bold>T)" + using \{\ \ A, \ \ B} \ elts (\ o)\ and prop_5401_f by simp + also have "\ = \ \ A \<^bold>\ \ \ B" + using truth_values_domain_def and \{\ \ A, \ \ B} \ elts (\ o)\ by auto + finally show ?thesis . +qed + +lemma (in general_model) forall_denotation: + assumes "\ \ \" + and "A \ wffs\<^bsub>o\<^esub>" + shows "\ \ (\x\<^bsub>\\<^esub>. A) = \<^bold>T \ (\z \ elts (\ \). \ (\((x, \) := z)) A = \<^bold>T)" +proof - + from assms(1) have lhs: "\ \ (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) \ z = \<^bold>T" if "z \ elts (\ \)" for z + using prop_5401_c and mixed_beta_conversion and that and true_wff by simp + from assms have rhs: "\ \ (\x\<^bsub>\\<^esub>. A) \ z = \ (\((x, \) := z)) A" if "z \ elts (\ \)" for z + using that by (simp add: mixed_beta_conversion) + from assms(2) have "\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub> \ wffs\<^bsub>\\o\<^esub>" and "\x\<^bsub>\\<^esub>. A \ wffs\<^bsub>\\o\<^esub>" + by auto + have "\ \ (\x\<^bsub>\\<^esub>. A) = \ \ (\\<^bsub>\\<^esub> \ (\x\<^bsub>\\<^esub>. A))" + unfolding forall_def .. + also have "\ = \ \ (Q\<^bsub>\\o\<^esub> \ (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) \ (\x\<^bsub>\\<^esub>. A))" + unfolding PI_def .. + also have "\ = \ \ ((\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) =\<^bsub>\\o\<^esub> (\x\<^bsub>\\<^esub>. A))" + unfolding equality_of_type_def .. + finally have "\ \ (\x\<^bsub>\\<^esub>. A) = \ \ ((\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) =\<^bsub>\\o\<^esub> (\x\<^bsub>\\<^esub>. A))" . + moreover from assms(1,2) have " + \ \ ((\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) =\<^bsub>\\o\<^esub> (\x\<^bsub>\\<^esub>. A)) = \<^bold>T \ \ \ (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) = \ \ (\x\<^bsub>\\<^esub>. A)" + using \\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub> \ wffs\<^bsub>\\o\<^esub>\ and \\x\<^bsub>\\<^esub>. A \ wffs\<^bsub>\\o\<^esub>\ and prop_5401_b by blast + moreover + have "(\ \ (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) = \ \ (\x\<^bsub>\\<^esub>. A)) \ (\z \ elts (\ \). \ (\((x, \) := z)) A = \<^bold>T)" + proof + assume "\ \ (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) = \ \ (\x\<^bsub>\\<^esub>. A)" + with lhs and rhs show "\z \ elts (\ \). \ (\((x, \) := z)) A = \<^bold>T" + by auto + next + assume "\z \ elts (\ \). \ (\((x, \) := z)) A = \<^bold>T" + moreover from assms have "\ \ (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) = (\<^bold>\z \<^bold>: \ \\<^bold>. \ (\((\, \) := z)) T\<^bsub>o\<^esub>)" + using wff_abs_denotation[OF \_is_wff_denotation_function] by blast + moreover from assms have "\ \ (\x\<^bsub>\\<^esub>. A) = (\<^bold>\z \<^bold>: \ \\<^bold>. \ (\((x, \) := z)) A)" + using wff_abs_denotation[OF \_is_wff_denotation_function] by blast + ultimately show "\ \ (\\\<^bsub>\\<^esub>. T\<^bsub>o\<^esub>) = \ \ (\x\<^bsub>\\<^esub>. A)" + using lhs and vlambda_extensionality by fastforce + qed + ultimately show ?thesis + by (simp only:) +qed + +lemma prop_5401_g: + assumes "is_general_model \" + and "\ \\<^sub>M \" + and "A \ wffs\<^bsub>o\<^esub>" + shows "\ \\<^bsub>\\<^esub> \x\<^bsub>\\<^esub>. A \ (\\. \ \\<^sub>M \ \ \ \\<^bsub>(x, \)\<^esub> \ \ \ \\<^bsub>\\<^esub> A)" +proof - + obtain \ and \ and \ where "\ = (\, \, \)" + using prod_cases3 by blast + with assms have " + \ \\<^bsub>\\<^esub> \x\<^bsub>\\<^esub>. A + \ + \x\<^bsub>\\<^esub>. A \ wffs\<^bsub>o\<^esub> \ is_general_model (\, \, \) \ \ \ \ \ \ \ (\x\<^bsub>\\<^esub>. A) = \<^bold>T" + by fastforce + also from assms and \\ = (\, \, \)\ have "\ \ (\z \ elts (\ \). \ (\((x, \) := z)) A = \<^bold>T)" + using general_model.forall_denotation by fastforce + also have "\ \ (\\. \ \ \ \ \ \\<^bsub>(x, \)\<^esub> \ \ \ \\<^bsub>\\<^esub> A)" + proof + assume *: "\z \ elts (\ \). \ (\((x, \) := z)) A = \<^bold>T" + { + fix \ + assume "\ \ \" and "\ \\<^bsub>(x, \)\<^esub> \" + have "\ \ A = \<^bold>T" + proof - + have "\z \ elts (\ \). \ = \((x, \) := z)" + proof (rule ccontr) + assume "\ (\z \ elts (\ \). \ = \((x, \) := z))" + with \\ \\<^bsub>(x, \)\<^esub> \\ have "\z \ elts (\ \). \ (x, \) \ z" + by fastforce + then have "\ (x, \) \ elts (\ \)" + by blast + moreover from assms(1) and \\ = (\, \, \)\ and \\ \ \\ have "\ (x, \) \ elts (\ \)" + using general_model_def and premodel_def and frame.is_assignment_def by auto + ultimately show False + by simp + qed + with * show ?thesis + by fastforce + qed + with assms(1) and \\ = (\, \, \)\ have "\ \\<^bsub>\\<^esub> A" + by simp + } + then show "\\. \ \ \ \ \ \\<^bsub>(x, \)\<^esub> \ \ \ \\<^bsub>\\<^esub> A" + by blast + next + assume *: "\\. \ \ \ \ \ \\<^bsub>(x, \)\<^esub> \ \ \ \\<^bsub>\\<^esub> A" + show "\z \ elts (\ \). \ (\((x, \) := z)) A = \<^bold>T" + proof + fix z + assume "z \ elts (\ \)" + with assms(1,2) and \\ = (\, \, \)\ have "\((x, \) := z) \ \" + using general_model_def and premodel_def and frame.is_assignment_def by auto + moreover have "\((x, \) := z) \\<^bsub>(x, \)\<^esub> \" + by simp + ultimately have "\ \\<^bsub>\((x, \) := z)\<^esub> A" + using * by blast + with assms(1) and \\ = (\, \, \)\ and \\((x, \) := z) \ \\ show "\ (\((x, \) := z)) A = \<^bold>T" + by simp + qed + qed + finally show ?thesis + using \\ = (\, \, \)\ + by simp +qed + +lemma (in general_model) axiom_1_validity_aux: + assumes "\ \ \" + shows "\ \ (\\<^bsub>o\o\<^esub> \ T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\o\<^esub> \ F\<^bsub>o\<^esub> \\<^sup>\ \\\<^bsub>o\<^esub>. \\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub>) = \<^bold>T" (is "\ \ (?A \\<^sup>\ ?B) = \<^bold>T") +proof - + let ?\ = "(\, \, \)" + from assms have *: "is_general_model ?\" "\ \\<^sub>M ?\" + using general_model_axioms by blast+ + have "?A \\<^sup>\ ?B \ wffs\<^bsub>o\<^esub>" + using axioms.axiom_1 and axioms_are_wffs_of_type_o by blast + have lhs: "\ \ ?A = \ (\, o\o) \ \<^bold>T \<^bold>\ \ (\, o\o) \ \<^bold>F" + proof - + have "\\<^bsub>o\o\<^esub> \ T\<^bsub>o\<^esub> \ wffs\<^bsub>o\<^esub>" and "\\<^bsub>o\o\<^esub> \ F\<^bsub>o\<^esub> \ wffs\<^bsub>o\<^esub>" + by blast+ + with assms have "\ \ ?A = \ \ (\\<^bsub>o\o\<^esub> \ T\<^bsub>o\<^esub>) \<^bold>\ \ \ (\\<^bsub>o\o\<^esub> \ F\<^bsub>o\<^esub>)" + using prop_5401_e' by simp + also from assms have "\ = \ (\, o\o) \ \ \ (T\<^bsub>o\<^esub>) \<^bold>\ \ (\, o\o) \ \ \ (F\<^bsub>o\<^esub>)" + using wff_app_denotation[OF \_is_wff_denotation_function] + and wff_var_denotation[OF \_is_wff_denotation_function] + by (metis false_wff true_wff wffs_of_type_intros(1)) + finally show ?thesis + using assms and prop_5401_c and prop_5401_d by simp + qed + have "\ \ (?A \\<^sup>\ ?B) = \<^bold>T" + proof (cases "\z \ elts (\ o). \ (\, o\o) \ z = \<^bold>T") + case True + with assms have "\ (\, o\o) \ \<^bold>T = \<^bold>T" and "\ (\, o\o) \ \<^bold>F = \<^bold>T" + using truth_values_domain_def by auto + with lhs have "\ \ ?A = \<^bold>T \<^bold>\ \<^bold>T" + by (simp only:) + also have "\ = \<^bold>T" + by simp + finally have "\ \ ?A = \<^bold>T" . + moreover have "\ \ ?B = \<^bold>T" + proof - + have "\\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub> \ wffs\<^bsub>o\<^esub>" + by blast + moreover + { + fix \ + assume "\ \ \" and "\ \\<^bsub>(\, o)\<^esub> \" + with assms have "\ \ (\\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub>) = \ \ (\\<^bsub>o\o\<^esub>) \ \ \ (\\<^bsub>o\<^esub>)" + using \_is_wff_denotation_function by blast + also from \\ \ \\ have "\ = \ (\, o\o) \ \ (\, o)" + using \_is_wff_denotation_function by auto + also from \\ \\<^bsub>(\, o)\<^esub> \\ have "\ = \ (\, o\o) \ \ (\, o)" + by simp + also from True and \\ \ \\ have "\ = \<^bold>T" + by blast + finally have "\ \ (\\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub>) = \<^bold>T" . + with assms and \\\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub> \ wffs\<^bsub>o\<^esub>\ have "?\ \\<^bsub>\\<^esub> \\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub>" + by simp + } + ultimately have "?\ \\<^bsub>\\<^esub> ?B" + using assms and * and prop_5401_g by auto + with *(2) show ?thesis + by simp + qed + ultimately show ?thesis + using assms and prop_5401_b' and wffs_from_equivalence[OF \?A \\<^sup>\ ?B \ wffs\<^bsub>o\<^esub>\] by simp + next + case False + then have "\z \ elts (\ o). \ (\, o\o) \ z \ \<^bold>T" + by blast + moreover from * have "\z \ elts (\ o). \ (\, o\o) \ z \ elts (\ o)" + using app_is_domain_respecting by blast + ultimately obtain z where "z \ elts (\ o)" and "\ (\, o\o) \ z = \<^bold>F" + using truth_values_domain_def by auto + define \ where \_def: "\ = \((\, o) := z)" + with * and \z \ elts (\ o)\ have "\ \ \" + by simp + then have "\ \ (\\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub>) = \ \ (\\<^bsub>o\o\<^esub>) \ \ \ (\\<^bsub>o\<^esub>)" + using \_is_wff_denotation_function by blast + also from \\ \ \\ have "\ = \ (\, o\o) \ \ (\, o)" + using \_is_wff_denotation_function by auto + also from \_def have "\ = \ (\, o\o) \ z" + by simp + also have "\ = \<^bold>F" + unfolding \\ (\, o\o) \ z = \<^bold>F\ .. + finally have "\ \ (\\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub>) = \<^bold>F" . + with \\ \ \\ have "\ ?\ \\<^bsub>\\<^esub> \\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub>" + by (auto simp add: inj_eq) + with \\ \ \\ and \_def have "\ (\\. \ \ \ \ \ \\<^bsub>(\, o)\<^esub> \ \ ?\ \\<^bsub>\\<^esub> \\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub>)" + using fun_upd_other by fastforce + with \\ ?\ \\<^bsub>\\<^esub> \\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub>\ have "\ ?\ \\<^bsub>\\<^esub> ?B" + using prop_5401_g[OF * wffs_from_forall[OF wffs_from_equivalence(2)[OF \?A \\<^sup>\ ?B \ wffs\<^bsub>o\<^esub>\]]] + by blast + then have "\ \ (\\\<^bsub>o\<^esub>. \\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub>) \ \<^bold>T" + by simp + moreover from assms have "\ \ ?B \ elts (\ o)" + using wffs_from_equivalence[OF \?A \\<^sup>\ ?B \ wffs\<^bsub>o\<^esub>\] and \_is_wff_denotation_function by auto + ultimately have "\ \ ?B = \<^bold>F" + by (simp add: truth_values_domain_def) + moreover have "\ \ (\\<^bsub>o\o\<^esub> \ T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\o\<^esub> \ F\<^bsub>o\<^esub>) = \<^bold>F" + proof - + from \z \ elts (\ o)\ and \\ (\, o\o) \ z = \<^bold>F\ + have "((\ (\, o\o)) \ \<^bold>T) = \<^bold>F \ ((\ (\, o\o)) \ \<^bold>F) = \<^bold>F" + using truth_values_domain_def by fastforce + moreover from \z \ elts (\ o)\ and \\ (\, o\o) \ z = \<^bold>F\ + and \\z \ elts (\ o). \ (\, o\o) \ z \ elts (\ o)\ + have "{(\ (\, o\o)) \ \<^bold>T, (\ (\, o\o)) \ \<^bold>F} \ elts (\ o)" + by (simp add: truth_values_domain_def) + ultimately have "((\ (\, o\o)) \ \<^bold>T) \<^bold>\ ((\ (\, o\o)) \ \<^bold>F) = \<^bold>F" + by auto + with lhs show ?thesis + by (simp only:) + qed + ultimately show ?thesis + using assms and prop_5401_b' and wffs_from_equivalence[OF \?A \\<^sup>\ ?B \ wffs\<^bsub>o\<^esub>\] by simp + qed + then show ?thesis . +qed + +lemma axiom_1_validity: + shows "\ \\<^bsub>o\o\<^esub> \ T\<^bsub>o\<^esub> \\<^sup>\ \\<^bsub>o\o\<^esub> \ F\<^bsub>o\<^esub> \\<^sup>\ \\\<^bsub>o\<^esub>. \\<^bsub>o\o\<^esub> \ \\<^bsub>o\<^esub>" (is "\ ?A \\<^sup>\ ?B") +proof (intro allI impI) + fix \ and \ + assume *: "is_general_model \" "\ \\<^sub>M \" + show "\ \\<^bsub>\\<^esub> ?A \\<^sup>\ ?B" + proof - + obtain \ and \ and \ where "\ = (\, \, \)" + using prod_cases3 by blast + moreover from * and \\ = (\, \, \)\ have "\ \ (?A \\<^sup>\ ?B) = \<^bold>T" + using general_model.axiom_1_validity_aux by simp + ultimately show ?thesis + by simp + qed +qed + +lemma (in general_model) axiom_2_validity_aux: + assumes "\ \ \" + shows "\ \ ((\\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\<^esub>) \\<^sup>\ (\\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub> \\<^sup>\ \\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub>)) = \<^bold>T" (is "\ \ (?A \\<^sup>\ ?B) = \<^bold>T") +proof - + have "?A \\<^sup>\ ?B \ wffs\<^bsub>o\<^esub>" + using axioms.axiom_2 and axioms_are_wffs_of_type_o by blast + from \?A \\<^sup>\ ?B \ wffs\<^bsub>o\<^esub>\ have "?A \ wffs\<^bsub>o\<^esub>" and "?B \ wffs\<^bsub>o\<^esub>" + using wffs_from_imp_op by blast+ + with assms have "\ \ (?A \\<^sup>\ ?B) = \ \ ?A \<^bold>\ \ \ ?B" + using prop_5401_f' by simp + moreover from assms and \?A \ wffs\<^bsub>o\<^esub>\ and \?B \ wffs\<^bsub>o\<^esub>\ have "{\ \ ?A, \ \ ?B} \ elts (\ o)" + using \_is_wff_denotation_function by simp + then have "{\ \ ?A, \ \ ?B} \ elts \" + by (simp add: truth_values_domain_def) + ultimately have \_imp_T: "\ \ (?A \\<^sup>\ ?B) = \<^bold>T \ \ \ ?A = \<^bold>F \ \ \ ?B = \<^bold>T" + by fastforce + then show ?thesis + proof (cases "\ (\, \) = \ (\, \)") + case True + from assms and \?B \ wffs\<^bsub>o\<^esub>\ have "\ \ ?B = \<^bold>T \ \ \ (\\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub>) = \ \ (\\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub>)" + using wffs_from_equivalence and prop_5401_b' by metis + moreover have "\ \ (\\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub>) = \ \ (\\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub>)" + proof - + from assms and \?B \ wffs\<^bsub>o\<^esub>\ have "\ \ (\\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub>) = \ \ (\\<^bsub>\\o\<^esub>) \ \ \ (\\<^bsub>\\<^esub>)" + using \_is_wff_denotation_function by blast + also from assms have "\ = \ (\, \\o) \ \ (\, \)" + using \_is_wff_denotation_function by auto + also from True have "\ = \ (\, \\o) \ \ (\, \)" + by (simp only:) + also from assms have "\ = \ \ (\\<^bsub>\\o\<^esub>) \ \ \ (\\<^bsub>\\<^esub>)" + using \_is_wff_denotation_function by auto + also from assms and \?B \ wffs\<^bsub>o\<^esub>\ have "\ = \ \ (\\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub>)" + using wff_app_denotation[OF \_is_wff_denotation_function] by (metis wffs_of_type_intros(1)) + finally show ?thesis . + qed + ultimately show ?thesis + using \_imp_T by simp + next + case False + from assms have "\ \ ?A = \<^bold>T \ \ \ (\\<^bsub>\\<^esub>) = \ \ (\\<^bsub>\\<^esub>)" + using prop_5401_b by blast + moreover from False and assms have "\ \ (\\<^bsub>\\<^esub>) \ \ \ (\\<^bsub>\\<^esub>)" + using \_is_wff_denotation_function by auto + ultimately have "\ \ ?A = \<^bold>F" + using assms and \{\ \ ?A, \ \ ?B} \ elts \\ by simp + then show ?thesis + using \_imp_T by simp + qed +qed + +lemma axiom_2_validity: + shows "\ (\\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\<^esub>) \\<^sup>\ (\\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub> \\<^sup>\ \\<^bsub>\\o\<^esub> \ \\<^bsub>\\<^esub>)" (is "\ ?A \\<^sup>\ ?B") +proof (intro allI impI) + fix \ and \ + assume *: "is_general_model \" "\ \\<^sub>M \" + show "\ \\<^bsub>\\<^esub> ?A \\<^sup>\ ?B" + proof - + obtain \ and \ and \ where "\ = (\, \, \)" + using prod_cases3 by blast + moreover from * and \\ = (\, \, \)\ have "\ \ (?A \\<^sup>\ ?B) = \<^bold>T" + using general_model.axiom_2_validity_aux by simp + ultimately show ?thesis + by force + qed +qed + +lemma (in general_model) axiom_3_validity_aux: + assumes "\ \ \" + shows "\ \ ((\\<^bsub>\\\\<^esub> =\<^bsub>\\\\<^esub> \\<^bsub>\\\\<^esub>) \\<^sup>\ \\\<^bsub>\\<^esub>. (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>)) = \<^bold>T" + (is "\ \ (?A \\<^sup>\ ?B) = \<^bold>T") +proof - + let ?\ = "(\, \, \)" + from assms have *: "is_general_model ?\" "\ \\<^sub>M ?\" + using general_model_axioms by blast+ + have B'_wffo: "\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub> \ wffs\<^bsub>o\<^esub>" + by blast + have "?A \\<^sup>\ ?B \ wffs\<^bsub>o\<^esub>" and "?A \ wffs\<^bsub>o\<^esub>" and "?B \ wffs\<^bsub>o\<^esub>" + proof - + show "?A \\<^sup>\ ?B \ wffs\<^bsub>o\<^esub>" + using axioms.axiom_3 and axioms_are_wffs_of_type_o + by blast + then show "?A \ wffs\<^bsub>o\<^esub>" and "?B \ wffs\<^bsub>o\<^esub>" + by (blast dest: wffs_from_equivalence)+ + qed + have "\ \ ?A = \ \ ?B" + proof (cases "\ (\, \\\) = \ (\, \\\)") + case True + have "\ \ ?A = \<^bold>T" + proof - + from assms have "\ \ (\\<^bsub>\\\\<^esub>) = \ (\, \\\)" + using \_is_wff_denotation_function by auto + also from True have "\ = \ (\, \\\)" + by (simp only:) + also from assms have "\ = \ \ (\\<^bsub>\\\\<^esub>)" + using \_is_wff_denotation_function by auto + finally have "\ \ (\\<^bsub>\\\\<^esub>) = \ \ (\\<^bsub>\\\\<^esub>)" . + with assms show ?thesis + using prop_5401_b by blast + qed + moreover have "\ \ ?B = \<^bold>T" + proof - + { + fix \ + assume "\ \ \" and "\ \\<^bsub>(\, \)\<^esub> \" + from assms and \\ \ \\ have "\ \ (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>) = \ \ (\\<^bsub>\\\\<^esub>) \ \ \ (\\<^bsub>\\<^esub>)" + using \_is_wff_denotation_function by blast + also from assms and \\ \ \\ have "\ = \ (\, \\\) \ \ (\, \)" + using \_is_wff_denotation_function by auto + also from \\ \\<^bsub>(\, \)\<^esub> \\ have "\ = \ (\, \\\) \ \ (\, \)" + by simp + also from True have "\ = \ (\, \\\) \ \ (\, \)" + by (simp only:) + also from \\ \\<^bsub>(\, \)\<^esub> \\ have "\ = \ (\, \\\) \ \ (\, \)" + by simp + also from assms and \\ \ \\ have "\ = \ \ (\\<^bsub>\\\\<^esub>) \ \ \ (\\<^bsub>\\<^esub>)" + using \_is_wff_denotation_function by auto + also from assms and \\ \ \\ have "\ = \ \ (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>)" + using wff_app_denotation[OF \_is_wff_denotation_function] by (metis wffs_of_type_intros(1)) + finally have "\ \ (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>) = \ \ (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>)" . + with B'_wffo and assms and \\ \ \\ have "\ \ (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>) = \<^bold>T" + using prop_5401_b and wffs_from_equality by blast + with *(2) have "?\ \\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>" + by simp + } + with * and B'_wffo have "?\ \\<^bsub>\\<^esub> ?B" + using prop_5401_g by force + with *(2) show ?thesis + by auto + qed + ultimately show ?thesis .. + next + case False + from * have "\ (\, \\\) \ elts (\ \ \ \ \)" and "\ (\, \\\) \ elts (\ \ \ \ \)" + by (simp_all add: function_domainD) + with False obtain z where "z \ elts (\ \)" and "\ (\, \\\) \ z \ \ (\, \\\) \ z" + by (blast dest: fun_ext) + define \ where "\ = \((\, \) := z)" + from * and \z \ elts (\ \)\ have "\ \ \" and "\ \\<^bsub>(\, \)\<^esub> \" + unfolding \_def by fastforce+ + have "\ \ (f\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>) = \ (f, \\\) \ z" for f + proof - + from \\ \ \\ have "\ \ (f\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>) = \ \ (f\<^bsub>\\\\<^esub>) \ \ \ (\\<^bsub>\\<^esub>)" + using \_is_wff_denotation_function by blast + also from \\ \ \\ have "\ = \ (f, \\\) \ \ (\, \)" + using \_is_wff_denotation_function by auto + finally show ?thesis + unfolding \_def by simp + qed + then have "\ \ (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>) = \ (\, \\\) \ z" and "\ \ (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>) = \ (\, \\\) \ z" + by (simp_all only:) + with \\ (\, \\\) \ z \ \ (\, \\\) \ z\ have "\ \ (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>) \ \ \ (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>)" + by simp + then have "\ \ (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>) = \<^bold>F" + proof - + from B'_wffo and \\ \ \\ and * have "\ \ (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>) \ elts (\ o)" + using \_is_wff_denotation_function by auto + moreover from B'_wffo have "{\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>, \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>} \ wffs\<^bsub>\\<^esub>" + by blast + with \\ \ \\ and \\ \ (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>) \ \ \ (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>)\ and B'_wffo + have "\ \ (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>) \ \<^bold>T" + using prop_5401_b by simp + ultimately show ?thesis + by (simp add: truth_values_domain_def) + qed + with \\ \ \\ have "\ ?\ \\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>" + by (auto simp add: inj_eq) + with \\ \ \\ and \\ \\<^bsub>(\, \)\<^esub> \\ + have "\\. \ \ \ \ \ \\<^bsub>(\, \)\<^esub> \ \ \ ?\ \\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>" + by blast + with * and B'_wffo have "\ ?\ \\<^bsub>\\<^esub> ?B" + using prop_5401_g by blast + then have "\ \ ?B = \<^bold>F" + proof - + from \?B \ wffs\<^bsub>o\<^esub>\ and * have "\ \ ?B \ elts (\ o)" + using \_is_wff_denotation_function by auto + with \\ ?\ \\<^bsub>\\<^esub> ?B\ and \?B \ wffs\<^bsub>o\<^esub>\ show ?thesis + using truth_values_domain_def by fastforce + qed + moreover have "\ \ (\\<^bsub>\\\\<^esub> =\<^bsub>\\\\<^esub> \\<^bsub>\\\\<^esub>) = \<^bold>F" + proof - + from * have "\ \ (\\<^bsub>\\\\<^esub>) = \ (\, \\\)" and "\ \ (\\<^bsub>\\\\<^esub>) = \ (\, \\\)" + using \_is_wff_denotation_function by auto + with False have "\ \ (\\<^bsub>\\\\<^esub>) \ \ \ (\\<^bsub>\\\\<^esub>)" + by simp + with * have "\ \ (\\<^bsub>\\\\<^esub> =\<^bsub>\\\\<^esub> \\<^bsub>\\\\<^esub>) \ \<^bold>T" + using prop_5401_b by blast + moreover from * and \?A \ wffs\<^bsub>o\<^esub>\ have "\ \ (\\<^bsub>\\\\<^esub> =\<^bsub>\\\\<^esub> \\<^bsub>\\\\<^esub>) \ elts (\ o)" + using \_is_wff_denotation_function by auto + ultimately show ?thesis + by (simp add: truth_values_domain_def) + qed + ultimately show ?thesis + by (simp only:) + qed + with * and \?A \ wffs\<^bsub>o\<^esub>\ and \?B \ wffs\<^bsub>o\<^esub>\ show ?thesis + using prop_5401_b' by simp +qed + +lemma axiom_3_validity: + shows "\ (\\<^bsub>\\\\<^esub> =\<^bsub>\\\\<^esub> \\<^bsub>\\\\<^esub>) \\<^sup>\ \\\<^bsub>\\<^esub>. (\\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub> =\<^bsub>\\<^esub> \\<^bsub>\\\\<^esub> \ \\<^bsub>\\<^esub>)" (is "\ ?A \\<^sup>\ ?B") +proof (intro allI impI) + fix \ and \ + assume *: "is_general_model \" "\ \\<^sub>M \" + show "\ \\<^bsub>\\<^esub> ?A \\<^sup>\ ?B" + proof - + obtain \ and \ and \ where "\ = (\, \, \)" + using prod_cases3 by blast + moreover from * and \\ = (\, \, \)\ have "\ \ (?A \\<^sup>\ ?B) = \<^bold>T" + using general_model.axiom_3_validity_aux by simp + ultimately show ?thesis + by simp + qed +qed + +lemma (in general_model) axiom_4_1_con_validity_aux: + assumes "\ \ \" + and "A \ wffs\<^bsub>\\<^esub>" + shows "\ \ ((\x\<^bsub>\\<^esub>. \c\\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> \c\\<^bsub>\\<^esub>) = \<^bold>T" +proof - + from assms(2) have "(\x\<^bsub>\\<^esub>. \c\\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> \c\\<^bsub>\\<^esub> \ wffs\<^bsub>o\<^esub>" + using axioms.axiom_4_1_con and axioms_are_wffs_of_type_o by blast + define \ where "\ = \((x, \) := \ \ A)" + from assms have "\ \ ((\x\<^bsub>\\<^esub>. \c\\<^bsub>\\<^esub>) \ A) = \ (\((x, \) := \ \ A)) (\c\\<^bsub>\\<^esub>)" + using prop_5401_a by blast + also have "\ = \ \ (\c\\<^bsub>\\<^esub>)" + unfolding \_def .. + also from assms and \_def have "\ = \ \ (\c\\<^bsub>\\<^esub>)" + using \_is_wff_denotation_function by auto + finally have "\ \ ((\x\<^bsub>\\<^esub>. \c\\<^bsub>\\<^esub>) \ A) = \ \ (\c\\<^bsub>\\<^esub>)" . + with assms(1) and \(\x\<^bsub>\\<^esub>. \c\\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> \c\\<^bsub>\\<^esub> \ wffs\<^bsub>o\<^esub>\ show ?thesis + using wffs_from_equality(1) and prop_5401_b by blast +qed + +lemma axiom_4_1_con_validity: + assumes "A \ wffs\<^bsub>\\<^esub>" + shows "\ (\x\<^bsub>\\<^esub>. \c\\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> \c\\<^bsub>\\<^esub>" +proof (intro allI impI) + fix \ and \ + assume *: "is_general_model \" "\ \\<^sub>M \" + show "\ \\<^bsub>\\<^esub> (\x\<^bsub>\\<^esub>. \c\\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> \c\\<^bsub>\\<^esub>" + proof - + obtain \ and \ and \ where "\ = (\, \, \)" + using prod_cases3 by blast + moreover from assms and * and \\ = (\, \, \)\ have "\ \ ((\x\<^bsub>\\<^esub>. \c\\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> \c\\<^bsub>\\<^esub>) = \<^bold>T" + using general_model.axiom_4_1_con_validity_aux by simp + ultimately show ?thesis + by simp + qed +qed + +lemma (in general_model) axiom_4_1_var_validity_aux: + assumes "\ \ \" + and "A \ wffs\<^bsub>\\<^esub>" + and "(y, \) \ (x, \)" + shows "\ \ ((\x\<^bsub>\\<^esub>. y\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> y\<^bsub>\\<^esub>) = \<^bold>T" +proof - + from assms(2) have "(\x\<^bsub>\\<^esub>. y\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> y\<^bsub>\\<^esub> \ wffs\<^bsub>o\<^esub>" + using axioms.axiom_4_1_var and axioms_are_wffs_of_type_o by blast + define \ where "\ = \((x, \) := \ \ A)" + with assms(1,2) have "\ \ ((\x\<^bsub>\\<^esub>. y\<^bsub>\\<^esub>) \ A) = \ (\((x, \) := \ \ A)) (y\<^bsub>\\<^esub>)" + using prop_5401_a by blast + also have "\ = \ \ (y\<^bsub>\\<^esub>)" + unfolding \_def .. + also have "\ = \ \ (y\<^bsub>\\<^esub>)" + proof - + from assms(1,2) have "\ \ A \ elts (\ \)" + using \_is_wff_denotation_function by auto + with \_def and assms(1) have "\ \ \" + by simp + moreover have "free_vars (y\<^bsub>\\<^esub>) = {(y, \)}" + by simp + with \_def and assms(3) have "\v \ free_vars (y\<^bsub>\\<^esub>). \ v = \ v" + by auto + ultimately show ?thesis + using prop_5400[OF wffs_of_type_intros(1) assms(1)] by simp + qed + finally have "\ \ ((\x\<^bsub>\\<^esub>. y\<^bsub>\\<^esub>) \ A) = \ \ (y\<^bsub>\\<^esub>)" . + with \(\x\<^bsub>\\<^esub>. y\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> y\<^bsub>\\<^esub> \ wffs\<^bsub>o\<^esub>\ show ?thesis + using wffs_from_equality(1) and prop_5401_b[OF assms(1)] by blast +qed + +lemma axiom_4_1_var_validity: + assumes "A \ wffs\<^bsub>\\<^esub>" + and "(y, \) \ (x, \)" + shows "\ (\x\<^bsub>\\<^esub>. y\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> y\<^bsub>\\<^esub>" +proof (intro allI impI) + fix \ and \ + assume *: "is_general_model \" "\ \\<^sub>M \" + show "\ \\<^bsub>\\<^esub> (\x\<^bsub>\\<^esub>. y\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> y\<^bsub>\\<^esub>" + proof - + obtain \ and \ and \ where "\ = (\, \, \)" + using prod_cases3 by blast + moreover from assms and * and \\ = (\, \, \)\ have "\ \ ((\x\<^bsub>\\<^esub>. y\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> y\<^bsub>\\<^esub>) = \<^bold>T" + using general_model.axiom_4_1_var_validity_aux by auto + ultimately show ?thesis + by simp + qed +qed + +lemma (in general_model) axiom_4_2_validity_aux: + assumes "\ \ \" + and "A \ wffs\<^bsub>\\<^esub>" + shows "\ \ ((\x\<^bsub>\\<^esub>. x\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> A) = \<^bold>T" +proof - + from assms(2) have "(\x\<^bsub>\\<^esub>. x\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> A \ wffs\<^bsub>o\<^esub>" + using axioms.axiom_4_2 and axioms_are_wffs_of_type_o by blast + define \ where "\ = \((x, \) := \ \ A)" + with assms have "\ \ ((\x\<^bsub>\\<^esub>. x\<^bsub>\\<^esub>) \ A) = \ \ (x\<^bsub>\\<^esub>)" + using prop_5401_a by blast + also from assms and \_def have "\ = \ (x, \)" + using \_is_wff_denotation_function by force + also from \_def have "\ = \ \ A" + by simp + finally have "\ \ ((\x\<^bsub>\\<^esub>. x\<^bsub>\\<^esub>) \ A) = \ \ A" . + with assms(1) and \(\x\<^bsub>\\<^esub>. x\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> A \ wffs\<^bsub>o\<^esub>\ show ?thesis + using wffs_from_equality and prop_5401_b by meson +qed + +lemma axiom_4_2_validity: + assumes "A \ wffs\<^bsub>\\<^esub>" + shows "\ (\x\<^bsub>\\<^esub>. x\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> A" +proof (intro allI impI) + fix \ and \ + assume *: "is_general_model \" "\ \\<^sub>M \" + show "\ \\<^bsub>\\<^esub> (\x\<^bsub>\\<^esub>. x\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> A" + proof - + obtain \ and \ and \ where "\ = (\, \, \)" + using prod_cases3 by blast + moreover from assms and * and \\ = (\, \, \)\ have "\ \ ((\x\<^bsub>\\<^esub>. x\<^bsub>\\<^esub>) \ A =\<^bsub>\\<^esub> A) = \<^bold>T" + using general_model.axiom_4_2_validity_aux by simp + ultimately show ?thesis + by simp + qed +qed + +lemma (in general_model) axiom_4_3_validity_aux: + assumes "\ \ \" + and "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\\\<^esub>" and "C \ wffs\<^bsub>\\<^esub>" + shows "\ \ ((\x\<^bsub>\\<^esub>. B \ C) \ A =\<^bsub>\\<^esub> ((\x\<^bsub>\\<^esub>. B) \ A) \ ((\x\<^bsub>\\<^esub>. C) \ A)) = \<^bold>T" + (is "\ \ (?A =\<^bsub>\\<^esub> ?B) = \<^bold>T") +proof - + from assms(2-4) have "?A =\<^bsub>\\<^esub> ?B \ wffs\<^bsub>o\<^esub>" + using axioms.axiom_4_3 and axioms_are_wffs_of_type_o by blast + define \ where "\ = \((x, \) := \ \ A)" + with assms(1,2) have "\ \ \" + using \_is_wff_denotation_function by auto + from assms and \_def have "\ \ ?A = \ \ (B \ C)" + using prop_5401_a by blast + also from assms(3,4) and \_def and \\ \ \\ have "\ = \ \ B \ \ \ C" + using \_is_wff_denotation_function by blast + also from assms(1-3) and \_def have "\ = \ \ ((\x\<^bsub>\\<^esub>. B) \ A) \ \ \ C" + using prop_5401_a by simp + also from assms(1,2,4) and \_def have "\ = \ \ ((\x\<^bsub>\\<^esub>. B) \ A) \ \ \ ((\x\<^bsub>\\<^esub>. C) \ A)" + using prop_5401_a by simp + also have "\ = \ \ ?B" + proof - + have "(\x\<^bsub>\\<^esub>. B) \ A \ wffs\<^bsub>\\\\<^esub>" and "(\x\<^bsub>\\<^esub>. C) \ A \ wffs\<^bsub>\\<^esub>" + using assms(2-4) by blast+ + with assms(1) show ?thesis + using wff_app_denotation[OF \_is_wff_denotation_function] by simp + qed + finally have "\ \ ?A = \ \ ?B" . + with assms(1) and \?A =\<^bsub>\\<^esub> ?B \ wffs\<^bsub>o\<^esub>\ show ?thesis + using prop_5401_b and wffs_from_equality by meson +qed + +lemma axiom_4_3_validity: + assumes "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\\\<^esub>" and "C \ wffs\<^bsub>\\<^esub>" + shows "\ (\x\<^bsub>\\<^esub>. B \ C) \ A =\<^bsub>\\<^esub> ((\x\<^bsub>\\<^esub>. B) \ A) \ ((\x\<^bsub>\\<^esub>. C) \ A)" (is "\ ?A =\<^bsub>\\<^esub> ?B") +proof (intro allI impI) + fix \ and \ + assume *: "is_general_model \" "\ \\<^sub>M \" + show "\ \\<^bsub>\\<^esub> ?A =\<^bsub>\\<^esub> ?B" + proof - + obtain \ and \ and \ where "\ = (\, \, \)" + using prod_cases3 by blast + moreover from assms and * and \\ = (\, \, \)\ have "\ \ (?A =\<^bsub>\\<^esub> ?B) = \<^bold>T" + using general_model.axiom_4_3_validity_aux by simp + ultimately show ?thesis + by simp + qed +qed + +lemma (in general_model) axiom_4_4_validity_aux: + assumes "\ \ \" + and "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" + and "(y, \) \ {(x, \)} \ vars A" + shows "\ \ ((\x\<^bsub>\\<^esub>. \y\<^bsub>\\<^esub>. B) \ A =\<^bsub>\\\\<^esub> (\y\<^bsub>\\<^esub>. (\x\<^bsub>\\<^esub>. B) \ A)) = \<^bold>T" + (is "\ \ (?A =\<^bsub>\\\\<^esub> ?B) = \<^bold>T") +proof - + from assms(2,3) have "?A =\<^bsub>\\\\<^esub> ?B \ wffs\<^bsub>o\<^esub>" + using axioms.axiom_4_4 and axioms_are_wffs_of_type_o by blast + let ?D = "\y\<^bsub>\\<^esub>. B" + define \ where "\ = \((x, \) := \ \ A)" + from assms(1,2) and \_def have "\ \ \" + using \_is_wff_denotation_function by simp + { + fix z + assume "z \ elts (\ \)" + define \' where "\' = \((y, \) := z)" + from assms(1) and \z \ elts (\ \)\ and \'_def have "\' \ \" + by simp + moreover from \'_def and assms(4) have "\v \ free_vars A. \ v = \' v" + using free_vars_in_all_vars by auto + ultimately have "\ \ A = \ \' A" + using assms(1,2) and prop_5400 by blast + with \_def and \'_def and assms(4) have "\'((x, \) := \ \' A) = \((y, \) := z)" + by auto + with \\ \ \\ and \z \ elts (\ \)\ and assms(3) have "\ \ ?D \ z = \ (\((y, \) := z)) B" + by (simp add: mixed_beta_conversion) + also from \\' \ \\ and assms(2,3) have "\ = \ \' ((\x\<^bsub>\\<^esub>. B) \ A)" + using prop_5401_a and \\'((x, \) := \ \' A) = \((y, \) := z)\ by simp + also from \'_def and assms(1) and \z \ elts (\ \)\ and \?A =\<^bsub>\\\\<^esub> ?B \ wffs\<^bsub>o\<^esub>\ + have "\ = \ \ ?B \ z" + by (metis mixed_beta_conversion wffs_from_abs wffs_from_equality(2)) + finally have "\ \ ?D \ z = \ \ ?B \ z" . + } + note * = this + then have "\ \ ?D = \ \ ?B" + proof - + from \\ \ \\ and assms(3) have "\ \ ?D = (\<^bold>\z \<^bold>: \ \\<^bold>. \ (\((y, \) := z)) B)" + using wff_abs_denotation[OF \_is_wff_denotation_function] by simp + moreover from assms(1) have "\ \ ?B = (\<^bold>\z \<^bold>: \ \\<^bold>. \ (\((y, \) := z)) ((\x\<^bsub>\\<^esub>. B) \ A))" + using wffs_from_abs[OF wffs_from_equality(2)[OF \?A =\<^bsub>\\\\<^esub> ?B \ wffs\<^bsub>o\<^esub>\]] + and wff_abs_denotation[OF \_is_wff_denotation_function] by meson + ultimately show ?thesis + using vlambda_extensionality and * by fastforce + qed + with assms(1-3) and \_def have "\ \ ?A = \ \ ?B" + using prop_5401_a and wffs_of_type_intros(4) by metis + with assms(1) show ?thesis + using prop_5401_b and wffs_from_equality[OF \?A =\<^bsub>\\\\<^esub> ?B \ wffs\<^bsub>o\<^esub>\] by blast +qed + +lemma axiom_4_4_validity: + assumes "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" + and "(y, \) \ {(x, \)} \ vars A" + shows "\ (\x\<^bsub>\\<^esub>. \y\<^bsub>\\<^esub>. B) \ A =\<^bsub>\\\\<^esub> (\y\<^bsub>\\<^esub>. (\x\<^bsub>\\<^esub>. B) \ A)" (is "\ ?A =\<^bsub>\\\\<^esub> ?B") +proof (intro allI impI) + fix \ and \ + assume *: "is_general_model \" "\ \\<^sub>M \" + show "\ \\<^bsub>\\<^esub> ?A =\<^bsub>\\\\<^esub> ?B" + proof - + obtain \ and \ and \ where "\ = (\, \, \)" + using prod_cases3 by blast + moreover from assms and * and \\ = (\, \, \)\ have "\ \ (?A =\<^bsub>\\\\<^esub> ?B) = \<^bold>T" + using general_model.axiom_4_4_validity_aux by blast + ultimately show ?thesis + by simp + qed +qed + +lemma (in general_model) axiom_4_5_validity_aux: + assumes "\ \ \" + and "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" + shows "\ \ ((\x\<^bsub>\\<^esub>. \x\<^bsub>\\<^esub>. B) \ A =\<^bsub>\\\\<^esub> (\x\<^bsub>\\<^esub>. B)) = \<^bold>T" +proof - + define \ where "\ = \((x, \) := \ \ A)" + from assms have wff: "(\x\<^bsub>\\<^esub>. \x\<^bsub>\\<^esub>. B) \ A =\<^bsub>\\\\<^esub> (\x\<^bsub>\\<^esub>. B) \ wffs\<^bsub>o\<^esub>" + using axioms.axiom_4_5 and axioms_are_wffs_of_type_o by blast + with assms(1,2) and \_def have "\ \ ((\x\<^bsub>\\<^esub>. \x\<^bsub>\\<^esub>. B) \ A) = \ \ (\x\<^bsub>\\<^esub>. B)" + using prop_5401_a and wffs_from_equality(2) by blast + also have "\ = \ \ (\x\<^bsub>\\<^esub>. B)" + proof - + have "(x, \) \ free_vars (\x\<^bsub>\\<^esub>. B)" + by simp + with \_def have "\v \ free_vars (\x\<^bsub>\\<^esub>. B). \ v = \ v" + by simp + moreover from \_def and assms(1,2) have "\ \ \" + using \_is_wff_denotation_function by simp + moreover from assms(2,3) have "(\x\<^bsub>\\<^esub>. B) \ wffs\<^bsub>\\\\<^esub>" + by fastforce + ultimately show ?thesis + using assms(1) and prop_5400 by metis + qed + finally have "\ \ ((\x\<^bsub>\\<^esub>. \x\<^bsub>\\<^esub>. B) \ A) = \ \ (\x\<^bsub>\\<^esub>. B)" . + with wff and assms(1) show ?thesis + using prop_5401_b and wffs_from_equality by meson +qed + +lemma axiom_4_5_validity: + assumes "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" + shows "\ (\x\<^bsub>\\<^esub>. \x\<^bsub>\\<^esub>. B) \ A =\<^bsub>\\\\<^esub> (\x\<^bsub>\\<^esub>. B)" +proof (intro allI impI) + fix \ and \ + assume *: "is_general_model \" "\ \\<^sub>M \" + show "\ \\<^bsub>\\<^esub> (\x\<^bsub>\\<^esub>. \x\<^bsub>\\<^esub>. B) \ A =\<^bsub>\\\\<^esub> (\x\<^bsub>\\<^esub>. B)" + proof - + obtain \ and \ and \ where "\ = (\, \, \)" + using prod_cases3 by blast + moreover + from assms and * and \\ = (\, \, \)\ have "\ \ ((\x\<^bsub>\\<^esub>. \x\<^bsub>\\<^esub>. B) \ A =\<^bsub>\\\\<^esub> (\x\<^bsub>\\<^esub>. B)) = \<^bold>T" + using general_model.axiom_4_5_validity_aux by blast + ultimately show ?thesis + by simp + qed +qed + +lemma (in general_model) axiom_5_validity_aux: + assumes "\ \ \" + shows "\ \ (\ \ (Q\<^bsub>i\<^esub> \ \\<^bsub>i\<^esub>) =\<^bsub>i\<^esub> \\<^bsub>i\<^esub>) = \<^bold>T" +proof - + have "\ \ (Q\<^bsub>i\<^esub> \ \\<^bsub>i\<^esub>) =\<^bsub>i\<^esub> \\<^bsub>i\<^esub> \ wffs\<^bsub>o\<^esub>" + using axioms.axiom_5 and axioms_are_wffs_of_type_o by blast + have "Q\<^bsub>i\<^esub> \ \\<^bsub>i\<^esub> \ wffs\<^bsub>i\o\<^esub>" + by blast + with assms have "\ \ (\ \ (Q\<^bsub>i\<^esub> \ \\<^bsub>i\<^esub>)) = \ \ \ \ \ \ (Q\<^bsub>i\<^esub> \ \\<^bsub>i\<^esub>)" + using \_is_wff_denotation_function by blast + also from assms have "\ = \ \ \ \ (\ \ (Q\<^bsub>i\<^esub>) \ \ \ (\\<^bsub>i\<^esub>))" + using wff_app_denotation[OF \_is_wff_denotation_function] by (metis Q_wff wffs_of_type_intros(1)) + also from assms have "\ = \ (\\<^sub>\, (i\o)\i) \ (\ (\\<^sub>Q, i\i\o) \ \ \ (\\<^bsub>i\<^esub>))" + using \_is_wff_denotation_function by auto + also from assms have "\ = \ (\\<^sub>\, (i\o)\i) \ ((q\<^bsub>i\<^esub>\<^bsup>\\<^esup>) \ \ \ (\\<^bsub>i\<^esub>))" + using Q_constant_of_type_def and Q_denotation by simp + also from assms have "\ = \ (\\<^sub>\, (i\o)\i) \ {\ \ (\\<^bsub>i\<^esub>)}\<^bsub>i\<^esub>\<^bsup>\\<^esup>" + using \_is_wff_denotation_function by auto + finally have "\ \ (\ \ (Q\<^bsub>i\<^esub> \ \\<^bsub>i\<^esub>)) = \ (\\<^sub>\, (i\o)\i) \ {\ \ (\\<^bsub>i\<^esub>)}\<^bsub>i\<^esub>\<^bsup>\\<^esup>" . + moreover from assms have "\ (\\<^sub>\, (i\o)\i) \ {\ \ (\\<^bsub>i\<^esub>)}\<^bsub>i\<^esub>\<^bsup>\\<^esup> = \ \ (\\<^bsub>i\<^esub>)" + using \_is_wff_denotation_function and \_denotation by force + ultimately have "\ \ (\ \ (Q\<^bsub>i\<^esub> \ \\<^bsub>i\<^esub>)) = \ \ (\\<^bsub>i\<^esub>)" + by (simp only:) + with assms and \Q\<^bsub>i\<^esub> \ \\<^bsub>i\<^esub> \ wffs\<^bsub>i\o\<^esub>\ show ?thesis + using prop_5401_b by blast +qed + +lemma axiom_5_validity: + shows "\ \ \ (Q\<^bsub>i\<^esub> \ \\<^bsub>i\<^esub>) =\<^bsub>i\<^esub> \\<^bsub>i\<^esub>" +proof (intro allI impI) + fix \ and \ + assume *: "is_general_model \" "\ \\<^sub>M \" + show "\ \\<^bsub>\\<^esub> \ \ (Q\<^bsub>i\<^esub> \ \\<^bsub>i\<^esub>) =\<^bsub>i\<^esub> \\<^bsub>i\<^esub>" + proof - + obtain \ and \ and \ where "\ = (\, \, \)" + using prod_cases3 by blast + moreover from * and \\ = (\, \, \)\ have "\ \ (\ \ (Q\<^bsub>i\<^esub> \ \\<^bsub>i\<^esub>) =\<^bsub>i\<^esub> \\<^bsub>i\<^esub>) = \<^bold>T" + using general_model.axiom_5_validity_aux by simp + ultimately show ?thesis + by simp + qed +qed + +lemma axioms_validity: + assumes "A \ axioms" + shows "\ A" + using assms + and axiom_1_validity + and axiom_2_validity + and axiom_3_validity + and axiom_4_1_con_validity + and axiom_4_1_var_validity + and axiom_4_2_validity + and axiom_4_3_validity + and axiom_4_4_validity + and axiom_4_5_validity + and axiom_5_validity + by cases auto + +lemma (in general_model) rule_R_validity_aux: + assumes "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" + and "\\. \ \ \ \ \ \ A = \ \ B" + and "C \ wffs\<^bsub>\\<^esub>" and "C' \ wffs\<^bsub>\\<^esub>" + and "p \ positions C" and "A \\<^bsub>p\<^esub> C" and "C\p \ B\ \ C'" + shows "\\. \ \ \ \ \ \ C = \ \ C'" +proof - + from assms(8,3-5,7) show ?thesis + proof (induction arbitrary: \) + case pos_found + then show ?case + by simp + next + case (replace_left_app p G B' G' H) + show ?case + proof (intro allI impI) + fix \ + assume "\ \ \" + from \G \ H \ wffs\<^bsub>\\<^esub>\ obtain \ where "G \ wffs\<^bsub>\\\\<^esub>" and "H \ wffs\<^bsub>\\<^esub>" + by (rule wffs_from_app) + with \G' \ H \ wffs\<^bsub>\\<^esub>\ have "G' \ wffs\<^bsub>\\\\<^esub>" + by (metis wff_has_unique_type wffs_from_app) + from assms(1) and \\ \ \\ and \G \ wffs\<^bsub>\\\\<^esub>\ and \H \ wffs\<^bsub>\\<^esub>\ + have "\ \ (G \ H) = \ \ G \ \ \ H" + using \_is_wff_denotation_function by blast + also from \\ \ \\ and \G \ wffs\<^bsub>\\\\<^esub>\ and \G' \ wffs\<^bsub>\\\\<^esub>\ have "\ = \ \ G' \ \ \ H" + using replace_left_app.IH and replace_left_app.prems(1,4) by simp + also from assms(1) and \\ \ \\ and \G' \ wffs\<^bsub>\\\\<^esub>\ and \H \ wffs\<^bsub>\\<^esub>\ + have "\ = \ \ (G' \ H)" + using \_is_wff_denotation_function by fastforce + finally show "\ \ (G \ H) = \ \ (G' \ H)" . + qed + next + case (replace_right_app p H B' H' G) + show ?case + proof (intro allI impI) + fix \ + assume "\ \ \" + from \G \ H \ wffs\<^bsub>\\<^esub>\ obtain \ where "G \ wffs\<^bsub>\\\\<^esub>" and "H \ wffs\<^bsub>\\<^esub>" + by (rule wffs_from_app) + with \G \ H' \ wffs\<^bsub>\\<^esub>\ have "H' \ wffs\<^bsub>\\<^esub>" + using wff_has_unique_type and wffs_from_app by (metis type.inject) + from assms(1) and \\ \ \\ and \G \ wffs\<^bsub>\\\\<^esub>\ and \H \ wffs\<^bsub>\\<^esub>\ + have "\ \ (G \ H) = \ \ G \ \ \ H" + using \_is_wff_denotation_function by blast + also from \\ \ \\ and \H \ wffs\<^bsub>\\<^esub>\ and \H' \ wffs\<^bsub>\\<^esub>\ have "\ = \ \ G \ \ \ H'" + using replace_right_app.IH and replace_right_app.prems(1,4) by force + also from assms(1) and \\ \ \\ and \G \ wffs\<^bsub>\\\\<^esub>\ and \H' \ wffs\<^bsub>\\<^esub>\ + have "\ = \ \ (G \ H')" + using \_is_wff_denotation_function by fastforce + finally show "\ \ (G \ H) = \ \ (G \ H')" . + qed + next + case (replace_abs p E B' E' x \) + show ?case + proof (intro allI impI) + fix \ + assume "\ \ \" + define \ where "\ z = \((x, \) := z)" for z + with \\ \ \\ have \_assg: "\ z \ \" if "z \ elts (\ \)" for z + by (simp add: that) + from \\x\<^bsub>\\<^esub>. E \ wffs\<^bsub>\\<^esub>\ obtain \ where "\ = \\\" and "E \ wffs\<^bsub>\\<^esub>" + by (rule wffs_from_abs) + with \\x\<^bsub>\\<^esub>. E' \ wffs\<^bsub>\\<^esub>\ have "E' \ wffs\<^bsub>\\<^esub>" + using wffs_from_abs by blast + from assms(1) and \\ \ \\ and \E \ wffs\<^bsub>\\<^esub>\ and \_def + have "\ \ (\x\<^bsub>\\<^esub>. E) = (\<^bold>\z \<^bold>: \ \\<^bold>. \ (\ z) E)" + using wff_abs_denotation[OF \_is_wff_denotation_function] by simp + also have "\ = (\<^bold>\z \<^bold>: \ \\<^bold>. \ (\ z) E')" + proof (intro vlambda_extensionality) + fix z + assume "z \ elts (\ \)" + from \E \ wffs\<^bsub>\\<^esub>\ and \E' \ wffs\<^bsub>\\<^esub>\ have "\\. \ \ \ \ \ \ E = \ \ E'" + using replace_abs.prems(1,4) and replace_abs.IH by simp + with \_assg and \z \ elts (\ \)\ show "\ (\ z) E = \ (\ z) E'" + by simp + qed + also from assms(1) and \\ \ \\ and \E' \ wffs\<^bsub>\\<^esub>\ and \_def + have "\ = \ \ (\x\<^bsub>\\<^esub>. E')" + using wff_abs_denotation[OF \_is_wff_denotation_function] by simp + finally show "\ \ (\x\<^bsub>\\<^esub>. E) = \ \ (\x\<^bsub>\\<^esub>. E')" . + qed + qed +qed + +lemma rule_R_validity: + assumes "C \ wffs\<^bsub>o\<^esub>" and "C' \ wffs\<^bsub>o\<^esub>" and "E \ wffs\<^bsub>o\<^esub>" + and "\ C" and "\ E" + and "is_rule_R_app p C' C E" + shows "\ C'" +proof (intro allI impI) + fix \ and \ + assume "is_general_model \" and "\ \\<^sub>M \" + show "\ \\<^bsub>\\<^esub> C'" + proof - + have "\ \ C'" + proof - + obtain \ and \ and \ where "\ = (\, \, \)" + using prod_cases3 by blast + from assms(6) obtain A and B and \ where "A \ wffs\<^bsub>\\<^esub>" and "B \ wffs\<^bsub>\\<^esub>" and "E = A =\<^bsub>\\<^esub> B" + using wffs_from_equality by (meson is_rule_R_app_def) + note * = \is_general_model \\ \\ = (\, \, \)\ \\ \\<^sub>M \\ + have "\ \' C = \ \' C'" if "\' \ \" for \' + proof - + from assms(5) and *(1,2) and \A \ wffs\<^bsub>\\<^esub>\ and \B \ wffs\<^bsub>\\<^esub>\ and \E = A =\<^bsub>\\<^esub> B\ and that + have "\\'. \' \ \ \\ \' A = \ \' B" + using general_model.prop_5401_b by blast + moreover + from \E = A =\<^bsub>\\<^esub> B\ and assms(6) have "p \ positions C" and "A \\<^bsub>p\<^esub> C" and "C\p \ B\ \ C'" + using is_subform_implies_in_positions by auto + ultimately show ?thesis + using \A \ wffs\<^bsub>\\<^esub>\ and \B \ wffs\<^bsub>\\<^esub>\ and \C \ wffs\<^bsub>o\<^esub>\ and assms(2) and that and *(1,2) + and general_model.rule_R_validity_aux by blast + qed + with assms(4) and *(1,2) show ?thesis + by simp + qed + with \\ \\<^sub>M \\ show ?thesis + by blast + qed +qed + +lemma individual_proof_step_validity: + assumes "is_proof \" and "A \ lset \" + shows "\ A" +using assms proof (induction "length \" arbitrary: \ A rule: less_induct) + case less + from \A \ lset \\ obtain i' where "\ ! i' = A" and "\ \ []" and "i' < length \" + by (metis empty_iff empty_set in_set_conv_nth) + with \is_proof \\ have "is_proof (take (Suc i') \)" and "take (Suc i') \ \ []" + using proof_prefix_is_proof[where \\<^sub>1 = "take (Suc i') \" and \\<^sub>2 = "drop (Suc i') \"] + and append_take_drop_id by simp_all + from \i' < length \\ consider (a) "i' < length \ - 1" | (b) "i' = length \ - 1" + by fastforce + then show ?case + proof cases + case a + then have "length (take (Suc i') \) < length \" + by simp + with \\ ! i' = A\ and \take (Suc i') \ \ []\ have "A \ lset (take (Suc i') \)" + by (simp add: take_Suc_conv_app_nth) + with \length (take (Suc i') \) < length \\ and \is_proof (take (Suc i') \)\ show ?thesis + using less(1) by blast + next + case b + with \\ ! i' = A\ and \\ \ []\ have "last \ = A" + using last_conv_nth by blast + with \is_proof \\ and \\ \ []\ and b have "is_proof_step \ i'" + using added_suffix_proof_preservation[where \' = "[]"] by simp + then consider + (axiom) "\ ! i' \ axioms" + | (rule_R) "\p j k. {j, k} \ {0.. is_rule_R_app p (\ ! i') (\ ! j) (\ ! k)" + by fastforce + then show ?thesis + proof cases + case axiom + with \\ ! i' = A\ show ?thesis + by (blast dest: axioms_validity) + next + case rule_R + then obtain p and j and k + where "{j, k} \ {0.. ! i') (\ ! j) (\ ! k)" + by blast + let ?\\<^sub>j = "take (Suc j) \" and ?\\<^sub>k = "take (Suc k) \" + obtain \\<^sub>j' and \\<^sub>k' where "\ = ?\\<^sub>j @ \\<^sub>j'" and "\ = ?\\<^sub>k @ \\<^sub>k'" + by (metis append_take_drop_id) + with \is_proof \\ have "is_proof (?\\<^sub>j @ \\<^sub>j')" and "is_proof (?\\<^sub>k @ \\<^sub>k')" + by (simp_all only:) + moreover from \\ \ []\ have "?\\<^sub>j \ []" and "?\\<^sub>k \ []" + by simp_all + ultimately have "is_proof_of ?\\<^sub>j (last ?\\<^sub>j)" and "is_proof_of ?\\<^sub>k (last ?\\<^sub>k)" + using proof_prefix_is_proof_of_last[where \ = ?\\<^sub>j and \' = \\<^sub>j'] + and proof_prefix_is_proof_of_last[where \ = ?\\<^sub>k and \' = \\<^sub>k'] + by fastforce+ + moreover + from \{j, k} \ {0.. and b have "length ?\\<^sub>j < length \" and "length ?\\<^sub>k < length \" + by force+ + moreover from calculation(3,4) have "\ ! j \ lset ?\\<^sub>j" and "\ ! k \ lset ?\\<^sub>k" + by (simp_all add: take_Suc_conv_app_nth) + ultimately have "\ \ ! j" and "\ \ ! k" + using \?\\<^sub>j \ []\ and \?\\<^sub>k \ []\ and less(1) unfolding is_proof_of_def by presburger+ + moreover have "\ ! i' \ wffs\<^bsub>o\<^esub>" and "\ ! j \ wffs\<^bsub>o\<^esub>" and "\ ! k \ wffs\<^bsub>o\<^esub>" + using \is_rule_R_app p (\ ! i') (\ ! j) (\ ! k)\ and replacement_preserves_typing + by force+ + ultimately show ?thesis + using \is_rule_R_app p (\ ! i') (\ ! j) (\ ! k)\ and \\ ! i' = A\ + and rule_R_validity[where C' = A] by blast + qed + qed +qed + +lemma semantic_modus_ponens: + assumes "is_general_model \" + and "A \ wffs\<^bsub>o\<^esub>" and "B \ wffs\<^bsub>o\<^esub>" + and "\ \ A \\<^sup>\ B" + and "\ \ A" + shows "\ \ B" +proof (intro allI impI) + fix \ + assume "\ \\<^sub>M \" + moreover obtain \ and \ and \ where "\ = (\, \, \)" + using prod_cases3 by blast + ultimately have "\ \ \" + by simp + show "\ \\<^bsub>\\<^esub> B" + proof - + from assms(4) have "\ \ (A \\<^sup>\ B) = \<^bold>T" + using \\ = (\, \, \)\ and \\ \\<^sub>M \\ by auto + with assms(1-3) have "\ \ A \<^bold>\ \ \ B = \<^bold>T" + using \\ = (\, \, \)\ and \\ \\<^sub>M \\ and general_model.prop_5401_f' by simp + moreover from assms(5) have "\ \ A = \<^bold>T" + using \\ = (\, \, \)\ and \\ \ \\ by auto + moreover from \\ = (\, \, \)\ and assms(1) have "elts (\ o) = elts \" + using frame.truth_values_domain_def and general_model_def and premodel_def by fastforce + with assms and \\ = (\, \, \)\ and \\ \ \\ and \\ \ A = \<^bold>T\ have "{\ \ A, \ \ B} \ elts \" + using general_model.\_is_wff_denotation_function + and premodel.wff_denotation_function_is_domain_respecting and general_model.axioms(1) by blast + ultimately have "\ \ B = \<^bold>T" + by fastforce + with \\ = (\, \, \)\ and assms(1) and \\ \ \\ show ?thesis + by simp + qed +qed + +lemma generalized_semantic_modus_ponens: + assumes "is_general_model \" + and "lset hs \ wffs\<^bsub>o\<^esub>" + and "\H \ lset hs. \ \ H" + and "P \ wffs\<^bsub>o\<^esub>" + and "\ \ hs \\<^sup>\\<^sub>\ P" + shows "\ \ P" +using assms(2-5) proof (induction hs arbitrary: P rule: rev_induct) + case Nil + then show ?case by simp +next + case (snoc H' hs) + from \\ \ (hs @ [H']) \\<^sup>\\<^sub>\ P\ have "\ \ hs \\<^sup>\\<^sub>\ (H' \\<^sup>\ P)" + by simp + moreover from \\H \ lset (hs @ [H']). \ \ H\ and \lset (hs @ [H']) \ wffs\<^bsub>o\<^esub>\ + have "\H \ lset hs. \ \ H" and "lset hs \ wffs\<^bsub>o\<^esub>" + by simp_all + moreover from \lset (hs @ [H']) \ wffs\<^bsub>o\<^esub>\ and \P \ wffs\<^bsub>o\<^esub>\ have "H' \\<^sup>\ P \ wffs\<^bsub>o\<^esub>" + by auto + ultimately have "\ \ H' \\<^sup>\ P" + by (elim snoc.IH) + moreover from \\H \ lset (hs @ [H']). \ \ H\ have "\ \ H'" + by simp + moreover from \H' \\<^sup>\ P \ wffs\<^bsub>o\<^esub>\ have "H' \ wffs\<^bsub>o\<^esub>" + using wffs_from_imp_op(1) by blast + ultimately show ?case + using assms(1) and \P \ wffs\<^bsub>o\<^esub>\ and semantic_modus_ponens by simp +qed + +subsection \Proposition 5402(a)\ + +proposition theoremhood_implies_validity: + assumes "is_theorem A" + shows "\ A" + using assms and individual_proof_step_validity by force + +subsection \Proposition 5402(b)\ + +proposition hyp_derivability_implies_validity: + assumes "is_hyps \" + and "is_model_for \ \" + and "\ \ A" + and "is_general_model \" + shows "\ \ A" +proof - + from assms(3) have "A \ wffs\<^bsub>o\<^esub>" + by (fact hyp_derivable_form_is_wffso) + from \\ \ A\ and \is_hyps \\ obtain \ where "finite \" and "\ \ \" and "\ \ A" + by blast + moreover from \finite \\ obtain hs where "lset hs = \" + using finite_list by blast + ultimately have "\ hs \\<^sup>\\<^sub>\ A" + using generalized_deduction_theorem by simp + with assms(4) have "\ \ hs \\<^sup>\\<^sub>\ A" + using derivability_from_no_hyps_theoremhood_equivalence and theoremhood_implies_validity + by blast + moreover from \\ \ \\ and assms(2) have "\ \ H" if "H \ \" for H + using that by blast + moreover from \\ \ \\ and \lset hs = \\ and assms(1) have "lset hs \ wffs\<^bsub>o\<^esub>" + by blast + ultimately show ?thesis + using assms(1,4) and \A \ wffs\<^bsub>o\<^esub>\ and \lset hs = \\ and generalized_semantic_modus_ponens + by auto +qed + +subsection \Theorem 5402 (Soundness Theorem)\ + +lemmas thm_5402 = theoremhood_implies_validity hyp_derivability_implies_validity + +end diff --git a/thys/Q0_Metatheory/Syntax.thy b/thys/Q0_Metatheory/Syntax.thy new file mode 100644 --- /dev/null +++ b/thys/Q0_Metatheory/Syntax.thy @@ -0,0 +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"}\ + +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"}\ + [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"}:\ + +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"}:\ + +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 diff --git a/thys/Q0_Metatheory/Utilities.thy b/thys/Q0_Metatheory/Utilities.thy new file mode 100644 --- /dev/null +++ b/thys/Q0_Metatheory/Utilities.thy @@ -0,0 +1,265 @@ +section \Utilities\ + +theory Utilities + imports + "Finite-Map-Extras.Finite_Map_Extras" +begin + +subsection \Utilities for lists\ + +fun foldr1 :: "('a \ 'a \ 'a) \ 'a list \ 'a" where + "foldr1 f [x] = x" +| "foldr1 f (x # xs) = f x (foldr1 f xs)" +| "foldr1 f [] = undefined f" + +abbreviation lset where "lset \ List.set" + +lemma rev_induct2 [consumes 1, case_names Nil snoc]: + assumes "length xs = length ys" + and "P [] []" + and "\x xs y ys. length xs = length ys \ P xs ys \ P (xs @ [x]) (ys @ [y])" + shows "P xs ys" +using assms proof (induction xs arbitrary: ys rule: rev_induct) + case (snoc x xs) + then show ?case by (cases ys rule: rev_cases) simp_all +qed simp + +subsection \Utilities for finite maps\ + +no_syntax + "_fmaplet" :: "['a, 'a] \ fmaplet" ("_ /$$:=/ _") + "_fmaplets" :: "['a, 'a] \ fmaplet" ("_ /[$$:=]/ _") + +syntax + "_fmaplet" :: "['a, 'a] \ fmaplet" ("_ /\/ _") + "_fmaplets" :: "['a, 'a] \ fmaplet" ("_ /[\]/ _") + +lemma fmdom'_fmap_of_list [simp]: + shows "fmdom' (fmap_of_list ps) = lset (map fst ps)" + by (induction ps) force+ + +lemma fmran'_singleton [simp]: + shows "fmran' {k \ v} = {v}" +proof - + have "v' \ fmran' {k \ v} \ v' = v" for v' + proof - + assume "v' \ fmran' {k \ v}" + fix k' + have "fmdom' {k \ v} = {k}" + by simp + then show "v' = v" + proof (cases "k' = k") + case True + with \v' \ fmran' {k \ v}\ show ?thesis + using fmdom'I by fastforce + next + case False + with \fmdom' {k \ v} = {k}\ and \v' \ fmran' {k \ v}\ show ?thesis + using fmdom'I by fastforce + qed + qed + moreover have "v \ fmran' {k \ v}" + by (simp add: fmran'I) + ultimately show ?thesis + by blast +qed + +lemma fmran'_fmupd [simp]: + assumes "m $$ x = None" + shows "fmran' (m(x \ y)) = {y} \ fmran' m" +using assms proof (intro subset_antisym subsetI) + fix x' + assume "m $$ x = None" and "x' \ fmran' (m(x \ y))" + then show "x' \ {y} \ fmran' m" + by (auto simp add: fmlookup_ran'_iff, metis option.inject) +next + fix x' + assume "m $$ x = None" and "x' \ {y} \ fmran' m" + then show "x' \ fmran' (m(x \ y))" + by (force simp add: fmlookup_ran'_iff) +qed + +lemma fmran'_fmadd [simp]: + assumes "fmdom' m \ fmdom' m' = {}" + shows "fmran' (m ++\<^sub>f m') = fmran' m \ fmran' m'" +using assms proof (intro subset_antisym subsetI) + fix x + assume "fmdom' m \ fmdom' m' = {}" and "x \ fmran' (m ++\<^sub>f m')" + then show "x \ fmran' m \ fmran' m'" + by (auto simp add: fmlookup_ran'_iff) meson +next + fix x + assume "fmdom' m \ fmdom' m' = {}" and "x \ fmran' m \ fmran' m'" + then show "x \ fmran' (m ++\<^sub>f m')" + using fmap_disj_comm and fmlookup_ran'_iff by fastforce +qed + +lemma finite_fmran': + shows "finite (fmran' m)" + by (simp add: fmran'_alt_def) + +lemma fmap_of_zipped_list_range: + assumes "length ks = length vs" + and "m = fmap_of_list (zip ks vs)" + and "k \ fmdom' m" + shows "m $$! k \ lset vs" + using assms by (induction arbitrary: m rule: list_induct2) auto + +lemma fmap_of_zip_nth [simp]: + assumes "length ks = length vs" + and "distinct ks" + and "i < length ks" + shows "fmap_of_list (zip ks vs) $$! (ks ! i) = vs ! i" + using assms by (simp add: fmap_of_list.rep_eq map_of_zip_nth) + +lemma fmap_of_zipped_list_fmran' [simp]: + assumes "distinct (map fst ps)" + shows "fmran' (fmap_of_list ps) = lset (map snd ps)" +using assms proof (induction ps) + case Nil + then show ?case + by auto +next + case (Cons p ps) + then show ?case + proof (cases "p \ lset ps") + case True + then show ?thesis + using Cons.prems by auto + next + case False + obtain k and v where "p = (k, v)" + by fastforce + with Cons.prems have "k \ fmdom' (fmap_of_list ps)" + by auto + then have "fmap_of_list (p # ps) = {k \ v} ++\<^sub>f fmap_of_list ps" + using \p = (k, v)\ and fmap_singleton_comm by fastforce + with Cons.prems have "fmran' (fmap_of_list (p # ps)) = {v} \ fmran' (fmap_of_list ps)" + by (simp add: \p = (k, v)\) + then have "fmran' (fmap_of_list (p # ps)) = {v} \ lset (map snd ps)" + using Cons.IH and Cons.prems by force + then show ?thesis + by (simp add: \p = (k, v)\) + qed +qed + +lemma fmap_of_list_nth [simp]: + assumes "distinct (map fst ps)" + and "j < length ps" + shows "fmap_of_list ps $$ ((map fst ps) ! j) = Some (map snd ps ! j)" + using assms by (induction j) (simp_all add: fmap_of_list.rep_eq) + +lemma fmap_of_list_nth_split [simp]: + assumes "distinct xs" + and "j < length xs" + and "length ys = length xs" and "length zs = length xs" + shows "fmap_of_list (zip xs (take k ys @ drop k zs)) $$ (xs ! j) = + (if j < k then Some (take k ys ! j) else Some (drop k zs ! (j - k)))" +using assms proof (induction k arbitrary: xs ys zs j) + case 0 + then show ?case + by (simp add: fmap_of_list.rep_eq map_of_zip_nth) +next + case (Suc k) + then show ?case + proof (cases xs) + case Nil + with Suc.prems(2) show ?thesis + by auto + next + case (Cons x xs') + let ?ps = "zip xs (take (Suc k) ys @ drop (Suc k) zs)" + from Cons and Suc.prems(3,4) obtain y and z and ys' and zs' + where "ys = y # ys'" and "zs = z # zs'" + by (metis length_0_conv neq_Nil_conv) + let ?ps' = "zip xs' (take k ys' @ drop k zs')" + from Cons have *: "fmap_of_list ?ps = fmap_of_list ((x, y) # ?ps')" + using \ys = y # ys'\ and \zs = z # zs'\ by fastforce + also have "\ = {x \ y} ++\<^sub>f fmap_of_list ?ps'" + proof - + from \ys = y # ys'\ and \zs = z # zs'\ have "fmap_of_list ?ps' $$ x = None" + using Cons and Suc.prems(1,3,4) by (simp add: fmdom'_notD) + then show ?thesis + using fmap_singleton_comm by fastforce + qed + finally have "fmap_of_list ?ps = {x \ y} ++\<^sub>f fmap_of_list ?ps'" . + then show ?thesis + proof (cases "j = 0") + case True + with \ys = y # ys'\ and Cons show ?thesis + by simp + next + case False + then have "xs ! j = xs' ! (j - 1)" + by (simp add: Cons) + moreover from \ys = y # ys'\ and \zs = z # zs'\ have "fmdom' (fmap_of_list ?ps') = lset xs'" + using Cons and Suc.prems(3,4) by force + moreover from False and Suc.prems(2) and Cons have "j - 1 < length xs'" + using le_simps(2) by auto + ultimately have "fmap_of_list ?ps $$ (xs ! j) = fmap_of_list ?ps' $$ (xs' ! (j - 1))" + using Cons and * and Suc.prems(1) by auto + with Suc.IH and Suc.prems(1,3,4) and Cons have **: "fmap_of_list ?ps $$ (xs ! j) = + (if j - 1 < k then Some (take k ys' ! (j - 1)) else Some (drop k zs' ! ((j - 1) - k)))" + using \j - 1 < length xs'\ and \ys = y # ys'\ and \zs = z # zs'\ by simp + then show ?thesis + proof (cases "j - 1 < k") + case True + with False and ** show ?thesis + using \ys = y # ys'\ by auto + next + case False + from Suc.prems(1) and Cons and \j - 1 < length xs'\ and \xs ! j = xs' ! (j - 1)\ have "j > 0" + using nth_non_equal_first_eq by fastforce + with False have "j \ Suc k" + by simp + moreover have "fmap_of_list ?ps $$ (xs ! j) = Some (drop (Suc k) zs ! (j - Suc k))" + using ** and False and \zs = z # zs'\ by fastforce + ultimately show ?thesis + by simp + qed + qed + qed +qed + +lemma fmadd_drop_cancellation [simp]: + assumes "m $$ k = Some v" + shows "{k \ v} ++\<^sub>f fmdrop k m = m" +using assms proof (induction m) + case fmempty + then show ?case + by simp +next + case (fmupd k' v' m') + then show ?case + proof (cases "k' = k") + case True + with fmupd.prems have "v = v'" + by fastforce + have "fmdrop k' (m'(k' \ v')) = m'" + unfolding fmdrop_fmupd_same using fmdrop_idle'[OF fmdom'_notI[OF fmupd.hyps]] by (unfold True) + then have "{k \ v} ++\<^sub>f fmdrop k' (m'(k' \ v')) = {k \ v} ++\<^sub>f m'" + by simp + then show ?thesis + using fmap_singleton_comm[OF fmupd.hyps] by (simp add: True \v = v'\) + next + case False + with fmupd.prems have "m' $$ k = Some v" + by force + from False have "{k \ v} ++\<^sub>f fmdrop k (m'(k' \ v')) = {k \ v} ++\<^sub>f (fmdrop k m')(k' \ v')" + by (simp add: fmdrop_fmupd) + also have "\ = ({k \ v} ++\<^sub>f fmdrop k m')(k' \ v')" + by fastforce + also from fmupd.prems and fmupd.IH[OF \m' $$ k = Some v\] have "\ = m'(k' \ v')" + by force + finally show ?thesis . + qed +qed + +lemma fmap_of_list_fmmap [simp]: + shows "fmap_of_list (map2 (\v' A'. (v', f A')) xs ys) = fmmap f (fmap_of_list (zip xs ys))" + unfolding fmmap_of_list + using cond_case_prod_eta + [where f = "\v' A'.(v', f A')" and g = "apsnd f", unfolded apsnd_conv, simplified] + by (rule arg_cong) + +end diff --git a/thys/Q0_Metatheory/document/root.bib b/thys/Q0_Metatheory/document/root.bib new file mode 100644 --- /dev/null +++ b/thys/Q0_Metatheory/document/root.bib @@ -0,0 +1,21 @@ +@book{andrews:1965, + author = {Andrews, Peter B.}, + title = {A Transfinite Type Theory with Type Variables}, + year = {1965}, + series = {Studies in Logic and the Foundations of Mathematics}, + volume = {36}, + publisher = {North-Holland Publishing Company}, + isbn = {978-0-4445-3402-6} +} + +@book{andrews:2002, + author = {Andrews, Peter B.}, + title = {An Introduction to Mathematical Logic and Type Theory: To Truth Through Proof}, + year = {2002}, + series = {Applied Logic Series}, + volume = {27}, + publisher = {Springer Dordrecht}, + doi = {10.1007/978-94-015-9934-4}, + isbn = {978-1-4020-0763-7} +} + diff --git a/thys/Q0_Metatheory/document/root.tex b/thys/Q0_Metatheory/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/Q0_Metatheory/document/root.tex @@ -0,0 +1,38 @@ +\documentclass[11pt,a4paper]{article} +\usepackage[T1]{fontenc} +\usepackage{isabelle,isabellesym} + +\usepackage{latexsym} +\usepackage{amssymb} +\usepackage{amsmath} +\usepackage{pdfsetup} + +\addtolength{\hoffset}{-1,5cm} +\addtolength{\textwidth}{3cm} + +\urlstyle{rm} +\isabellestyle{it} + +\begin{document} + +\title{Metatheory of $\mathcal{Q}_0$} +\author{Javier D\'iaz\\\url{}} +\maketitle + +\begin{abstract} +This entry is a formalization of the metatheory of $\mathcal{Q}_0$ in Isabelle/HOL. $\mathcal{Q}_0$ \cite{andrews:2002} is a classical higher-order logic equivalent to Church's Simple Theory of Types. In this entry we formalize Chapter 5 of \cite{andrews:2002}, up to and including the proofs of soundness and consistency of $\mathcal{Q}_0$. These proof are, to the best of our knowledge, the first to be formalized in a proof assistant. +\end{abstract} + +\newpage +\tableofcontents + +\parindent 0pt\parskip 0.5ex + +\newpage +\input{session} + +\bibliographystyle{abbrv} +\bibliography{root.bib} + +\end{document} + diff --git a/thys/ROOTS b/thys/ROOTS --- a/thys/ROOTS +++ b/thys/ROOTS @@ -1,778 +1,779 @@ ABY3_Protocols ADS_Functor AI_Planning_Languages_Semantics AODV AOT AVL-Trees AWN Abortable_Linearizable_Modules Abs_Int_ITP2012 Abstract-Hoare-Logics Abstract-Rewriting Abstract_Completeness Abstract_Soundness Ackermanns_not_PR Actuarial_Mathematics Adaptive_State_Counting Affine_Arithmetic Aggregation_Algebras Akra_Bazzi Algebraic_Numbers Algebraic_VCs Allen_Calculus Amicable_Numbers Amortized_Complexity AnselmGod Applicative_Lifting Approximation_Algorithms Architectural_Design_Patterns Aristotles_Assertoric_Syllogistic Arith_Prog_Rel_Primes ArrowImpossibilityGS Attack_Trees Auto2_HOL Auto2_Imperative_HOL AutoFocus-Stream Automated_Stateful_Protocol_Verification Automatic_Refinement AxiomaticCategoryTheory BDD BD_Security_Compositional BNF_CC BNF_Operations BTree Balog_Szemeredi_Gowers Banach_Steinhaus Belief_Revision Bell_Numbers_Spivey BenOr_Kozen_Reif Berlekamp_Zassenhaus Bernoulli Bertrands_Postulate Bicategory BinarySearchTree Binary_Code_Imprimitive Binding_Syntax_Theory Binomial-Heaps Binomial-Queues BirdKMP Birkhoff_Finite_Distributive_Lattices Blue_Eyes Bondy Boolean_Expression_Checkers Boolos_Curious_Inference Boolos_Curious_Inference_Automated Bounded_Deducibility_Security Buchi_Complementation Budan_Fourier Buffons_Needle Buildings BytecodeLogicJmlTypes C2KA_DistributedSystems CAVA_Automata CAVA_LTL_Modelchecker CCS CHERI-C_Memory_Model CISC-Kernel CRDT CRYSTALS-Kyber CSP_RefTK CVP_Hardness CYK CZH_Elementary_Categories CZH_Foundations CZH_Universal_Constructions CakeML CakeML_Codegen Call_Arity Card_Equiv_Relations Card_Multisets Card_Number_Partitions Card_Partitions Cartan_FP Case_Labeling Catalan_Numbers Category Category2 Category3 Catoids Cauchy Cayley_Hamilton Certification_Monads Ceva Chandy_Lamport Chord_Segments Circus Clean Clique_and_Monotone_Circuits ClockSynchInst Closest_Pair_Points CoCon CoSMeDis CoSMed CofGroups Coinductive Coinductive_Languages Collections Combinable_Wands Combinatorial_Enumeration_Algorithms Combinatorics_Words Combinatorics_Words_Graph_Lemma Combinatorics_Words_Lyndon CommCSL Commuting_Hermitian Comparison_Sort_Lower_Bound Compiling-Exceptions-Correctly Complete_Non_Orders Completeness Complex_Bounded_Operators Complex_Geometry Complx ComponentDependencies ConcurrentGC ConcurrentIMP Concurrent_Ref_Alg Concurrent_Revisions Conditional_Simplification Conditional_Transfer_Rule Consensus_Refined Constructive_Cryptography Constructive_Cryptography_CM Constructor_Funs Containers Cook_Levin CoreC++ Core_DOM Core_SC_DOM Correctness_Algebras Cotangent_PFD_Formula Count_Complex_Roots Coupledsim_Contrasim CryptHOL CryptoBasedCompositionalProperties Crypto_Standards Cubic_Quartic_Equations DCR-ExecutionEquivalence DFS_Framework DOM_Components DPRM_Theorem DPT-SAT-Solver DataRefinementIBP Datatype_Order_Generator Decl_Sem_Fun_PL Decreasing-Diagrams Decreasing-Diagrams-II Dedekind_Real Deep_Learning Delta_System_Lemma Density_Compiler Dependent_SIFUM_Refinement Dependent_SIFUM_Type_Systems Depth-First-Search Derangements Deriving Descartes_Sign_Rule Design_Theory Dict_Construction Differential_Dynamic_Logic Differential_Game_Logic DigitsInBase Digit_Expansions Dijkstra_Shortest_Path Diophantine_Eqns_Lin_Hom Directed_Sets Dirichlet_L Dirichlet_Series DiscretePricing Discrete_Summation Disintegration DiskPaxos Distributed_Distinct_Elements Dominance_CHK DynamicArchitectures Dynamic_Tables E_Transcendental Earley_Parser Echelon_Form EdmondsKarp_Maxflow Edwards_Elliptic_Curves_Group Efficient-Mergesort Efficient_Weighted_Path_Order Elimination_Of_Repeated_Factors Elliptic_Curves_Group_Law Encodability_Process_Calculi Epistemic_Logic Equivalence_Relation_Enumeration Ergodic_Theory Error_Function Euler_MacLaurin Euler_Partition Euler_Polyhedron_Formula Eudoxus_Reals Eval_FO Example-Submission Executable_Randomized_Algorithms Expander_Graphs Extended_Finite_State_Machine_Inference Extended_Finite_State_Machines FFT FLP FOL-Fitting FOL_Axiomatic FOL_Harrison FOL_Seq_Calc1 FOL_Seq_Calc2 FOL_Seq_Calc3 FO_Theory_Rewriting FSM_Tests Factor_Algebraic_Polynomial Factored_Transition_System_Bounding Falling_Factorial_Sum Farkas FeatherweightJava Featherweight_OCL Fermat3_4 FileRefinement FinFun Finger-Trees Finite-Map-Extras Finite_Automata_HF Finite_Fields Finitely_Generated_Abelian_Groups First_Order_Terms First_Welfare_Theorem Fishburn_Impossibility Fisher_Yates Fishers_Inequality Fixed_Length_Vector Flow_Networks Floyd_Warshall Flyspeck-Tame FocusStreamsCaseStudies Forcing Formal_Puiseux_Series Formal_SSA Formula_Derivatives Foundation_of_geometry Fourier Free-Boolean-Algebra Free-Groups Frequency_Moments Fresh_Identifiers FunWithFunctions FunWithTilings Functional-Automata Functional_Ordered_Resolution_Prover Furstenberg_Topology GPU_Kernel_PL Gabow_SCC GaleStewart_Games Gale_Shapley Game_Based_Crypto Gauss-Jordan-Elim-Fun Gauss_Jordan Gauss_Sums Gaussian_Integers GenClock General-Triangle Generalized_Counting_Sort Generic_Deriving Generic_Join GewirthPGCProof Girth_Chromatic Given_Clause_Loops GoedelGod Goedel_HFSet_Semantic Goedel_HFSet_Semanticless Goedel_Incompleteness Goodstein_Lambda GraphMarkingIBP Graph_Saturation Graph_Theory Gray_Codes Green Groebner_Bases Groebner_Macaulay Gromov_Hyperbolicity Grothendieck_Schemes Group-Ring-Module HOL-CSP HOLCF-Prelude HRB-Slicing Hahn_Jordan_Decomposition Hales_Jewett Heard_Of Hello_World HereditarilyFinite Hermite Hermite_Lindemann Hidden_Markov_Models Higher_Order_Terms HoareForDivergence Hoare_Time Hood_Melville_Queue HotelKeyCards Huffman Hybrid_Logic Hybrid_Multi_Lane_Spatial_Logic Hybrid_Systems_VCs HyperCTL HyperHoareLogic Hyperdual Hypergraph_Basics Hypergraph_Colourings IEEE_Floating_Point IFC_Tracking IMAP-CRDT IMO2019 IMP2 IMP2_Binary_Heap IMP_Compiler IMP_Compiler_Reuse IO_Language_Conformance IP_Addresses Imperative_Insertion_Sort Implicational_Logic Impossible_Geometry Incompleteness Incredible_Proof_Machine Independence_CH Inductive_Confidentiality Inductive_Inference InfPathElimination InformationFlowSlicing InformationFlowSlicing_Inter Integration Interpolation_Polynomials_HOL_Algebra Interpreter_Optimizations Interval_Arithmetic_Word32 Intro_Dest_Elim Involutions2Squares Iptables_Semantics Irrational_Series_Erdos_Straus Irrationality_J_Hancl Irrationals_From_THEBOOK IsaGeoCoq IsaNet Isabelle_C Isabelle_Marries_Dirac Isabelle_Meta_Model Jacobson_Basic_Algebra Jinja JinjaDCI JinjaThreads JiveDataStoreModel Jordan_Hoelder Jordan_Normal_Form KAD KAT_and_DRA KBPs KD_Tree Key_Agreement_Strong_Adversaries Khovanskii_Theorem Kleene_Algebra Kneser_Cauchy_Davenport Knights_Tour Knot_Theory Knuth_Bendix_Order Knuth_Morris_Pratt Koenigsberg_Friendship Kruskal Kuratowski_Closure_Complement LLL_Basis_Reduction LLL_Factorization LOFT LP_Duality LTL LTL_Master_Theorem LTL_Normal_Form LTL_to_DRA LTL_to_GBA Lam-ml-Normalization LambdaAuth LambdaMu Lambda_Free_EPO Lambda_Free_KBOs Lambda_Free_RPOs Lambert_W Landau_Symbols Laplace_Transform Latin_Square LatticeProperties Launchbury Laws_of_Large_Numbers Lazy-Lists-II Lazy_Case Lehmer Lifting_Definition_Option Lifting_the_Exponent LightweightJava LinearQuantifierElim Linear_Inequalities Linear_Programming Linear_Recurrences Liouville_Numbers List-Index List-Infinite List_Interleaving List_Inversions List_Update LocalLexing Localization_Ring Locally-Nameless-Sigma Logging_Independent_Anonymity Lovasz_Local Lowe_Ontological_Argument Lower_Semicontinuous Lp Lucas_Theorem MDP-Algorithms MDP-Rewards MFMC_Countable MFODL_Monitor_Optimized MFOTL_Monitor MSO_Regex_Equivalence Markov_Models Marriage Mason_Stothers Matrices_for_ODEs Matrix Matrix_Tensor Matroids Max-Card-Matching Maximum_Segment_Sum Median_Method Median_Of_Medians_Selection Menger Mereology Mersenne_Primes Metalogic_ProofChecker MHComputation MiniML MiniSail Minimal_SSA Minkowskis_Theorem Minsky_Machines MLSS_Decision_Proc ML_Unification Modal_Logics_for_NTS Modular_Assembly_Kit_Security Modular_arithmetic_LLL_and_HNF_algorithms Monad_Memo_DP Monad_Normalisation MonoBoolTranAlgebra MonoidalCategory Monomorphic_Monad MuchAdoAboutTwo Multi_Party_Computation Multirelations Multirelations_Heterogeneous Multiset_Ordering_NPC Multitape_To_Singletape_TM Myhill-Nerode Name_Carrying_Type_Inference Nano_JSON Nash_Williams Nat-Interval-Logic Native_Word Nested_Multisets_Ordinals Network_Security_Policy_Verification Neumann_Morgenstern_Utility No_FTL_observers No_FTL_observers_Gen_Rel Nominal2 Noninterference_CSP Noninterference_Concurrent_Composition Noninterference_Generic_Unwinding Noninterference_Inductive_Unwinding Noninterference_Ipurge_Unwinding Noninterference_Sequential_Composition NormByEval Nullstellensatz Number_Theoretic_Transform Octonions OpSets Open_Induction Optics Optimal_BST Orbit_Stabiliser Order_Lattice_Props Ordered_Resolution_Prover Ordinal Ordinal_Partitions Ordinals_and_Cardinals Ordinary_Differential_Equations PAC_Checker PAL PAPP_Impossibility PCF PLM POPLmark-deBruijn PSemigroupsConvolution Package_logic Padic_Field Padic_Ints Pairing_Heap Paraconsistency Parity_Game Partial_Function_MR Partial_Order_Reduction Password_Authentication_Protocol Pell Perfect_Fields Perfect-Number-Thm Perron_Frobenius Physical_Quantities Pi_Calculus Pi_Transcendental Planarity_Certificates Pluennecke_Ruzsa_Inequality Poincare_Bendixson Poincare_Disc Polygonal_Number_Theorem Polynomial_Factorization Polynomial_Interpolation Polynomials Pop_Refinement Posix-Lexing Possibilistic_Noninterference Power_Sum_Polynomials Pratt_Certificate Prefix_Free_Code_Combinators Presburger-Automata Prim_Dijkstra_Simple Prime_Distribution_Elementary Prime_Harmonic_Series Prime_Number_Theorem Priority_Queue_Braun Priority_Search_Trees Probabilistic_Noninterference Probabilistic_Prime_Tests Probabilistic_System_Zoo Probabilistic_Timed_Automata Probabilistic_While Probability_Inequality_Completeness Program-Conflict-Analysis Progress_Tracking Projective_Geometry Projective_Measurements Promela Proof_Strategy_Language PropResPI Propositional_Logic_Class Propositional_Proof_Systems Prpu_Maxflow PseudoHoops Psi_Calculi Ptolemys_Theorem Public_Announcement_Logic +Q0_Metatheory QHLProver QR_Decomposition Quantales Quantales_Converse Quantifier_Elimination_Hybrid Quasi_Borel_Spaces Quaternions Query_Optimization Quick_Sort_Cost RIPEMD-160-SPARK ROBDD RSAPSS Ramsey-Infinite Random_BSTs Random_Graph_Subgraph_Threshold Randomised_BSTs Randomised_Social_Choice Rank_Nullity_Theorem Real_Impl Real_Power Real_Time_Deque Recursion-Addition Recursion-Theory-I Refine_Imperative_HOL Refine_Monadic RefinementReactive Regex_Equivalence Registers Regression_Test_Selection Regular-Sets Regular_Algebras Regular_Tree_Relations Relation_Algebra Relational-Incorrectness-Logic Relational_Cardinality Relational_Disjoint_Set_Forests Relational_Forests Relational_Method Relational_Minimum_Spanning_Trees Relational_Paths Rensets Rep_Fin_Groups ResiduatedTransitionSystem Residuated_Lattices Resolution_FOL Rewrite_Properties_Reduction Rewriting_Z Ribbon_Proofs Risk_Free_Lending Robbins-Conjecture Robinson_Arithmetic Root_Balanced_Tree Roth_Arithmetic_Progressions Routing Roy_Floyd_Warshall SATSolverVerification SCC_Bloemen_Sequential SC_DOM_Components SDS_Impossibility SIFPL SIFUM_Type_Systems SPARCv8 S_Finite_Measure_Monad Safe_Distance Safe_OCL Safe_Range_RC Saturation_Framework Saturation_Framework_Extensions Sauer_Shelah_Lemma Schutz_Spacetime Schwartz_Zippel Secondary_Sylow Security_Protocol_Refinement Selection_Heap_Sort SenSocialChoice Separata Separation_Algebra Separation_Logic_Imperative_HOL Separation_Logic_Unbounded SequentInvertibility Shadow_DOM Shadow_SC_DOM Shivers-CFA ShortestPath Show Sigma_Commit_Crypto Signature_Groebner Simpl Simple_Clause_Learning Simple_Firewall Simplex Simplicial_complexes_and_boolean_functions SimplifiedOntologicalArgument Skew_Heap Skip_Lists Slicing Sliding_Window_Algorithm Smith_Normal_Form Smooth_Manifolds Solidity Sophomores_Dream Sort_Encodings Source_Coding_Theorem SpecCheck Special_Function_Bounds Splay_Tree Sqrt_Babylonian Stable_Matching Stalnaker_Logic Standard_Borel_Spaces Statecharts Stateful_Protocol_Composition_and_Typing Stellar_Quorums Stern_Brocot Stewart_Apollonius Stirling_Formula Stochastic_Matrices Stone_Algebras Stone_Kleene_Relation_Algebras Stone_Relation_Algebras Store_Buffer_Reduction Stream-Fusion Stream_Fusion_Code StrictOmegaCategories Strong_Security Sturm_Sequences Sturm_Tarski Stuttering_Equivalence Subresultants Subset_Boolean_Algebras SumSquares Sunflowers SuperCalc Suppes_Theorem Surprise_Paradox Symmetric_Polynomials Syntax_Independent_Logic Synthetic_Completeness Szemeredi_Regularity Szpilrajn TESL_Language TLA Tail_Recursive_Functions Tarskis_Geometry Taylor_Models Three_Circles Three_Squares Timed_Automata Topological_Semantics Topology TortoiseHare TsirelsonBound Transcendence_Series_Hancl_Rucki Transformer_Semantics Transition_Systems_and_Automata Transitive-Closure Transitive-Closure-II Transitive_Models Transport Treaps Tree-Automata Tree_Decomposition Tree_Enumeration Triangle Trie Turans_Graph_Theorem Twelvefold_Way Two_Generated_Word_Monoids_Intersection Tycon Types_Tableaus_and_Goedels_God Types_To_Sets_Extension UPF UPF_Firewall UTP Undirected_Graph_Theory Universal_Hash_Families Universal_Turing_Machine UpDown_Scheme VYDRA_MDL Valuation Van_Emde_Boas_Trees Van_der_Waerden VectorSpace VeriComp Verified-Prover Verified_SAT_Based_AI_Planning VerifyThis2018 VerifyThis2019 Vickrey_Clarke_Groves Virtual_Substitution VolpanoSmith WHATandWHERE_Security WOOT_Strong_Eventual_Consistency WebAssembly Weight_Balanced_Trees Weighted_Arithmetic_Geometric_Mean Weighted_Path_Order Well_Quasi_Orders Wetzels_Problem Winding_Number_Eval Word_Lib WorkerWrapper X86_Semantics XML Youngs_Inequality ZFC_in_HOL Zeckendorf Zeta_3_Irrational Zeta_Function pGCL