diff --git a/thys/AOT/AOT_Axioms.thy b/thys/AOT/AOT_Axioms.thy new file mode 100644 --- /dev/null +++ b/thys/AOT/AOT_Axioms.thy @@ -0,0 +1,263 @@ +(*<*) +theory AOT_Axioms + imports AOT_Definitions +begin +(*>*) + +section\Axioms of PLM\ + +AOT_axiom "pl:1": \\ \ (\ \ \)\ + by (auto simp: AOT_sem_imp AOT_model_axiomI) +AOT_axiom "pl:2": \(\ \ (\ \ \)) \ ((\ \ \) \ (\ \ \))\ + by (auto simp: AOT_sem_imp AOT_model_axiomI) +AOT_axiom "pl:3": \(\\ \ \\) \ ((\\ \ \) \ \)\ + by (auto simp: AOT_sem_imp AOT_sem_not AOT_model_axiomI) + +AOT_axiom "cqt:1": \\\ \{\} \ (\\ \ \{\})\ + by (auto simp: AOT_sem_denotes AOT_sem_forall AOT_sem_imp AOT_model_axiomI) + +AOT_axiom "cqt:2[const_var]": \\\\ + using AOT_sem_vars_denote by (rule AOT_model_axiomI) +AOT_axiom "cqt:2[lambda]": + assumes \INSTANCE_OF_CQT_2(\)\ + shows \[\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\\ + by (auto intro!: AOT_model_axiomI AOT_sem_cqt_2[OF assms]) +AOT_axiom "cqt:2[lambda0]": + shows \[\ \]\\ + by (auto intro!: AOT_model_axiomI + simp: AOT_sem_lambda_denotes "existence:3"[unfolded AOT_model_equiv_def]) + +AOT_axiom "cqt:3": \\\ (\{\} \ \{\}) \ (\\ \{\} \ \\ \{\})\ + by (simp add: AOT_sem_forall AOT_sem_imp AOT_model_axiomI) +AOT_axiom "cqt:4": \\ \ \\ \\ + by (simp add: AOT_sem_forall AOT_sem_imp AOT_model_axiomI) +AOT_axiom "cqt:5:a": \[\]\\<^sub>1...\\<^sub>n \ (\\ & \\<^sub>1...\\<^sub>n\)\ + by (simp add: AOT_sem_conj AOT_sem_denotes AOT_sem_exe + AOT_sem_imp AOT_model_axiomI) +AOT_axiom "cqt:5:a[1]": \[\]\ \ (\\ & \\)\ + using "cqt:5:a" AOT_model_axiomI by blast +AOT_axiom "cqt:5:a[2]": \[\]\\<^sub>1\\<^sub>2 \ (\\ & \\<^sub>1\ & \\<^sub>2\)\ + by (rule AOT_model_axiomI) + (metis AOT_model_denotes_prod_def AOT_sem_conj AOT_sem_denotes AOT_sem_exe + AOT_sem_imp case_prodD) +AOT_axiom "cqt:5:a[3]": \[\]\\<^sub>1\\<^sub>2\\<^sub>3 \ (\\ & \\<^sub>1\ & \\<^sub>2\ & \\<^sub>3\)\ + by (rule AOT_model_axiomI) + (metis AOT_model_denotes_prod_def AOT_sem_conj AOT_sem_denotes AOT_sem_exe + AOT_sem_imp case_prodD) +AOT_axiom "cqt:5:a[4]": \[\]\\<^sub>1\\<^sub>2\\<^sub>3\\<^sub>4 \ (\\ & \\<^sub>1\ & \\<^sub>2\ & \\<^sub>3\ & \\<^sub>4\)\ + by (rule AOT_model_axiomI) + (metis AOT_model_denotes_prod_def AOT_sem_conj AOT_sem_denotes AOT_sem_exe + AOT_sem_imp case_prodD) +AOT_axiom "cqt:5:b": \\\<^sub>1...\\<^sub>n[\] \ (\\ & \\<^sub>1...\\<^sub>n\)\ + using AOT_sem_enc_denotes + by (auto intro!: AOT_model_axiomI simp: AOT_sem_conj AOT_sem_denotes AOT_sem_imp)+ +AOT_axiom "cqt:5:b[1]": \\[\] \ (\\ & \\)\ + using "cqt:5:b" AOT_model_axiomI by blast +AOT_axiom "cqt:5:b[2]": \\\<^sub>1\\<^sub>2[\] \ (\\ & \\<^sub>1\ & \\<^sub>2\)\ + by (rule AOT_model_axiomI) + (metis AOT_model_denotes_prod_def AOT_sem_conj AOT_sem_denotes + AOT_sem_enc_denotes AOT_sem_imp case_prodD) +AOT_axiom "cqt:5:b[3]": \\\<^sub>1\\<^sub>2\\<^sub>3[\] \ (\\ & \\<^sub>1\ & \\<^sub>2\ & \\<^sub>3\)\ + by (rule AOT_model_axiomI) + (metis AOT_model_denotes_prod_def AOT_sem_conj AOT_sem_denotes + AOT_sem_enc_denotes AOT_sem_imp case_prodD) +AOT_axiom "cqt:5:b[4]": \\\<^sub>1\\<^sub>2\\<^sub>3\\<^sub>4[\] \ (\\ & \\<^sub>1\ & \\<^sub>2\ & \\<^sub>3\ & \\<^sub>4\)\ + by (rule AOT_model_axiomI) + (metis AOT_model_denotes_prod_def AOT_sem_conj AOT_sem_denotes + AOT_sem_enc_denotes AOT_sem_imp case_prodD) + +AOT_axiom "l-identity": \\ = \ \ (\{\} \ \{\})\ + by (rule AOT_model_axiomI) + (simp add: AOT_sem_eq AOT_sem_imp) + +AOT_act_axiom "logic-actual": \\<^bold>\\ \ \\ + by (rule AOT_model_act_axiomI) + (simp add: AOT_sem_act AOT_sem_imp) + +AOT_axiom "logic-actual-nec:1": \\<^bold>\\\ \ \\<^bold>\\\ + by (rule AOT_model_axiomI) + (simp add: AOT_sem_act AOT_sem_equiv AOT_sem_not) +AOT_axiom "logic-actual-nec:2": \\<^bold>\(\ \ \) \ (\<^bold>\\ \ \<^bold>\\)\ + by (rule AOT_model_axiomI) + (simp add: AOT_sem_act AOT_sem_equiv AOT_sem_imp) + +AOT_axiom "logic-actual-nec:3": \\<^bold>\(\\ \{\}) \ \\ \<^bold>\\{\}\ + by (rule AOT_model_axiomI) + (simp add: AOT_sem_act AOT_sem_equiv AOT_sem_forall AOT_sem_denotes) +AOT_axiom "logic-actual-nec:4": \\<^bold>\\ \ \<^bold>\\<^bold>\\\ + by (rule AOT_model_axiomI) + (simp add: AOT_sem_act AOT_sem_equiv) + +AOT_axiom "qml:1": \\(\ \ \) \ (\\ \ \\)\ + by (rule AOT_model_axiomI) + (simp add: AOT_sem_box AOT_sem_imp) +AOT_axiom "qml:2": \\\ \ \\ + by (rule AOT_model_axiomI) + (simp add: AOT_sem_box AOT_sem_imp) +AOT_axiom "qml:3": \\\ \ \\\\ + by (rule AOT_model_axiomI) + (simp add: AOT_sem_box AOT_sem_dia AOT_sem_imp) + +AOT_axiom "qml:4": \\\x (E!x & \\<^bold>\E!x)\ + using AOT_sem_concrete AOT_model_contingent + by (auto intro!: AOT_model_axiomI + simp: AOT_sem_box AOT_sem_dia AOT_sem_imp AOT_sem_exists + AOT_sem_denotes AOT_sem_conj AOT_sem_not AOT_sem_act + AOT_sem_exe)+ + +AOT_axiom "qml-act:1": \\<^bold>\\ \ \\<^bold>\\\ + by (rule AOT_model_axiomI) + (simp add: AOT_sem_act AOT_sem_box AOT_sem_imp) +AOT_axiom "qml-act:2": \\\ \ \<^bold>\\\\ + by (rule AOT_model_axiomI) + (simp add: AOT_sem_act AOT_sem_box AOT_sem_equiv) + +AOT_axiom descriptions: \x = \<^bold>\x(\{x}) \ \z(\<^bold>\\{z} \ z = x)\ +proof (rule AOT_model_axiomI) + AOT_modally_strict { + AOT_show \x = \<^bold>\x(\{x}) \ \z(\<^bold>\\{z} \ z = x)\ + by (induct; simp add: AOT_sem_equiv AOT_sem_forall AOT_sem_act AOT_sem_eq) + (metis (no_types, opaque_lifting) AOT_sem_desc_denotes AOT_sem_desc_prop + AOT_sem_denotes) + } +qed + +AOT_axiom "lambda-predicates:1": + \[\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\ \ [\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}] = [\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\ + by (rule AOT_model_axiomI) + (simp add: AOT_sem_denotes AOT_sem_eq AOT_sem_imp) +AOT_axiom "lambda-predicates:1[zero]": \[\ p]\ \ [\ p] = [\ p]\ + by (rule AOT_model_axiomI) + (simp add: AOT_sem_denotes AOT_sem_eq AOT_sem_imp) +AOT_axiom "lambda-predicates:2": + \[\x\<^sub>1...x\<^sub>n \{x\<^sub>1...x\<^sub>n}]\ \ ([\x\<^sub>1...x\<^sub>n \{x\<^sub>1...x\<^sub>n}]x\<^sub>1...x\<^sub>n \ \{x\<^sub>1...x\<^sub>n})\ + by (rule AOT_model_axiomI) + (simp add: AOT_sem_equiv AOT_sem_imp AOT_sem_lambda_beta AOT_sem_vars_denote) +AOT_axiom "lambda-predicates:3": \[\x\<^sub>1...x\<^sub>n [F]x\<^sub>1...x\<^sub>n] = F\ + by (rule AOT_model_axiomI) + (simp add: AOT_sem_lambda_eta AOT_sem_vars_denote) +AOT_axiom "lambda-predicates:3[zero]": \[\ p] = p\ + by (rule AOT_model_axiomI) + (simp add: AOT_sem_eq AOT_sem_lambda0 AOT_sem_vars_denote) + +AOT_axiom "safe-ext": + \([\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\ & \\\\<^sub>1...\\\<^sub>n (\{\\<^sub>1...\\<^sub>n} \ \{\\<^sub>1...\\<^sub>n})) \ + [\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\\ + using AOT_sem_lambda_coex + by (auto intro!: AOT_model_axiomI simp: AOT_sem_imp AOT_sem_denotes AOT_sem_conj + AOT_sem_equiv AOT_sem_box AOT_sem_forall) +AOT_axiom "safe-ext[2]": + \([\\\<^sub>1\\<^sub>2 \{\\<^sub>1,\\<^sub>2}]\ & \\\\<^sub>1\\\<^sub>2 (\{\\<^sub>1, \\<^sub>2} \ \{\\<^sub>1, \\<^sub>2})) \ + [\\\<^sub>1\\<^sub>2 \{\\<^sub>1,\\<^sub>2}]\\ + using "safe-ext"[where \="\(x,y). \ x y"] + by (simp add: AOT_model_axiom_def AOT_sem_denotes AOT_model_denotes_prod_def + AOT_sem_forall AOT_sem_imp AOT_sem_conj AOT_sem_equiv AOT_sem_box) +AOT_axiom "safe-ext[3]": + \([\\\<^sub>1\\<^sub>2\\<^sub>3 \{\\<^sub>1,\\<^sub>2,\\<^sub>3}]\ & \\\\<^sub>1\\\<^sub>2\\\<^sub>3 (\{\\<^sub>1, \\<^sub>2, \\<^sub>3} \ \{\\<^sub>1, \\<^sub>2, \\<^sub>3})) \ + [\\\<^sub>1\\<^sub>2\\<^sub>3 \{\\<^sub>1,\\<^sub>2,\\<^sub>3}]\\ + using "safe-ext"[where \="\(x,y,z). \ x y z"] + by (simp add: AOT_model_axiom_def AOT_model_denotes_prod_def AOT_sem_forall + AOT_sem_denotes AOT_sem_imp AOT_sem_conj AOT_sem_equiv AOT_sem_box) +AOT_axiom "safe-ext[4]": + \([\\\<^sub>1\\<^sub>2\\<^sub>3\\<^sub>4 \{\\<^sub>1,\\<^sub>2,\\<^sub>3,\\<^sub>4}]\ & + \\\\<^sub>1\\\<^sub>2\\\<^sub>3\\\<^sub>4 (\{\\<^sub>1, \\<^sub>2, \\<^sub>3, \\<^sub>4} \ \{\\<^sub>1, \\<^sub>2, \\<^sub>3, \\<^sub>4})) \ + [\\\<^sub>1\\<^sub>2\\<^sub>3\\<^sub>4 \{\\<^sub>1,\\<^sub>2,\\<^sub>3,\\<^sub>4}]\\ + using "safe-ext"[where \="\(x,y,z,w). \ x y z w"] + by (simp add: AOT_model_axiom_def AOT_model_denotes_prod_def AOT_sem_forall + AOT_sem_denotes AOT_sem_imp AOT_sem_conj AOT_sem_equiv AOT_sem_box) + +AOT_axiom "nary-encoding[2]": + \x\<^sub>1x\<^sub>2[F] \ x\<^sub>1[\y [F]yx\<^sub>2] & x\<^sub>2[\y [F]x\<^sub>1y]\ + by (rule AOT_model_axiomI) + (simp add: AOT_sem_conj AOT_sem_equiv AOT_enc_prod_def AOT_proj_enc_prod_def + AOT_sem_unary_proj_enc AOT_sem_vars_denote) +AOT_axiom "nary-encoding[3]": + \x\<^sub>1x\<^sub>2x\<^sub>3[F] \ x\<^sub>1[\y [F]yx\<^sub>2x\<^sub>3] & x\<^sub>2[\y [F]x\<^sub>1yx\<^sub>3] & x\<^sub>3[\y [F]x\<^sub>1x\<^sub>2y]\ + by (rule AOT_model_axiomI) + (simp add: AOT_sem_conj AOT_sem_equiv AOT_enc_prod_def AOT_proj_enc_prod_def + AOT_sem_unary_proj_enc AOT_sem_vars_denote) +AOT_axiom "nary-encoding[4]": + \x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F] \ x\<^sub>1[\y [F]yx\<^sub>2x\<^sub>3x\<^sub>4] & + x\<^sub>2[\y [F]x\<^sub>1yx\<^sub>3x\<^sub>4] & + x\<^sub>3[\y [F]x\<^sub>1x\<^sub>2yx\<^sub>4] & + x\<^sub>4[\y [F]x\<^sub>1x\<^sub>2x\<^sub>3y]\ + by (rule AOT_model_axiomI) + (simp add: AOT_sem_conj AOT_sem_equiv AOT_enc_prod_def AOT_proj_enc_prod_def + AOT_sem_unary_proj_enc AOT_sem_vars_denote) + +AOT_axiom encoding: \x[F] \ \x[F]\ + using AOT_sem_enc_nec + by (auto intro!: AOT_model_axiomI simp: AOT_sem_imp AOT_sem_box) + +AOT_axiom nocoder: \O!x \ \\F x[F]\ + by (auto intro!: AOT_model_axiomI + simp: AOT_sem_imp AOT_sem_not AOT_sem_exists AOT_sem_ordinary + AOT_sem_dia + AOT_sem_lambda_beta[OF AOT_sem_ordinary_def_denotes, + OF AOT_sem_vars_denote]) + (metis AOT_sem_nocoder) + +AOT_axiom "A-objects": \\x (A!x & \F(x[F] \ \{F}))\ +proof(rule AOT_model_axiomI) + AOT_modally_strict { + AOT_obtain \ where \\\ & \\E!\ & \F (\[F] \ \{F})\ + using AOT_sem_A_objects[of _ \] + by (auto simp: AOT_sem_imp AOT_sem_box AOT_sem_forall AOT_sem_exists + AOT_sem_conj AOT_sem_not AOT_sem_dia AOT_sem_denotes + AOT_sem_equiv) blast + AOT_thus \\x (A!x & \F(x[F] \ \{F}))\ + unfolding AOT_sem_exists + by (auto intro!: exI[where x=\] + simp: AOT_sem_lambda_beta[OF AOT_sem_abstract_def_denotes] + AOT_sem_box AOT_sem_dia AOT_sem_not AOT_sem_denotes + AOT_var_of_term_inverse AOT_sem_conj + AOT_sem_equiv AOT_sem_forall AOT_sem_abstract) + } +qed + +AOT_theorem universal_closure: + assumes \for arbitrary \: \{\} \ \\<^sub>\\ + shows \\\ \{\} \ \\<^sub>\\ + using assms + by (metis AOT_term_of_var_cases AOT_model_axiom_def AOT_sem_denotes AOT_sem_forall) + +AOT_theorem act_closure: + assumes \\ \ \\<^sub>\\ + shows \\<^bold>\\ \ \\<^sub>\\ + using assms by (simp add: AOT_model_axiom_def AOT_sem_act) + +AOT_theorem nec_closure: + assumes \\ \ \\<^sub>\\ + shows \\\ \ \\<^sub>\\ + using assms by (simp add: AOT_model_axiom_def AOT_sem_box) + +AOT_theorem universal_closure_act: + assumes \for arbitrary \: \{\} \ \\ + shows \\\ \{\} \ \\ + using assms + by (metis AOT_term_of_var_cases AOT_model_act_axiom_def AOT_sem_denotes + AOT_sem_forall) + +text\The following are not part of PLM and only hold in the extended models. + They are a generalization of the predecessor axiom.\ +context AOT_ExtendedModel +begin +AOT_axiom indistinguishable_ord_enc_all: + \\\ & A!x & A!y & \F \([F]x \ [F]y) \ + ((\G(\z(O!z \ \([G]z \ [\]z)) \ x[G])) \ + \G(\z(O!z \ \([G]z \ [\]z)) \ y[G]))\ + by (rule AOT_model_axiomI) + (auto simp: AOT_sem_equiv AOT_sem_imp AOT_sem_conj + AOT_sem_indistinguishable_ord_enc_all) +AOT_axiom indistinguishable_ord_enc_ex: + \\\ & A!x & A!y & \F \([F]x \ [F]y) \ + ((\G(\z(O!z \ \([G]z \ [\]z)) & x[G])) \ + \G(\z(O!z \ \([G]z \ [\]z)) & y[G]))\ + by (rule AOT_model_axiomI) + (auto simp: AOT_sem_equiv AOT_sem_imp AOT_sem_conj + AOT_sem_indistinguishable_ord_enc_ex) +end + +(*<*) +end +(*>*) \ No newline at end of file diff --git a/thys/AOT/AOT_BasicLogicalObjects.thy b/thys/AOT/AOT_BasicLogicalObjects.thy new file mode 100644 --- /dev/null +++ b/thys/AOT/AOT_BasicLogicalObjects.thy @@ -0,0 +1,628 @@ +(*<*) +theory AOT_BasicLogicalObjects + imports AOT_PLM +begin +(*>*) + +section\Basic Logical Objects\ +(* Note: so far only the parts required for possible world theory are implemented *) + +AOT_define TruthValueOf :: \\ \ \ \ \\ (\TruthValueOf'(_,_')\) + "tv-p": \TruthValueOf(x,p) \\<^sub>d\<^sub>f A!x & \F (x[F] \ \q((q \ p) & F = [\y q]))\ + +AOT_theorem "p-has-!tv:1": \\x TruthValueOf(x,p)\ + using "tv-p"[THEN "\Df"] + by (AOT_subst \TruthValueOf(x,p)\ + \A!x & \F (x[F] \ \q((q \ p) & F = [\y q]))\ for: x) + (simp add: "A-objects"[axiom_inst]) + + +AOT_theorem "p-has-!tv:2": \\!x TruthValueOf(x,p)\ + using "tv-p"[THEN "\Df"] + by (AOT_subst \TruthValueOf(x,p)\ + \A!x & \F (x[F] \ \q((q \ p) & F = [\y q]))\ for: x) + (simp add: "A-objects!") + + +AOT_theorem "uni-tv": \\<^bold>\x TruthValueOf(x,p)\\ + using "A-Exists:2" "RA[2]" "\E"(2) "p-has-!tv:2" by blast + +AOT_define TheTruthValueOf :: \\ \ \\<^sub>s\ (\\_\ [100] 100) + "the-tv-p": \\p =\<^sub>d\<^sub>f \<^bold>\x TruthValueOf(x,p)\ + +AOT_define PropEnc :: \\ \ \ \ \\ (infixl \\<^bold>\\ 40) + "prop-enc": \x\<^bold>\p \\<^sub>d\<^sub>f x\ & x[\y p]\ + +AOT_theorem "tv-id:1": \\p = \<^bold>\x (A!x & \F (x[F] \ \q((q \ p) & F = [\y q])))\ +proof - + AOT_have \\\x(TruthValueOf(x,p) \ A!x & \F (x[F] \ \q((q \ p) & F = [\y q])))\ + by (rule RN; rule GEN; rule "tv-p"[THEN "\Df"]) + AOT_hence \\<^bold>\x TruthValueOf(x,p) = \<^bold>\x (A!x & \F (x[F] \ \q((q \ p) & F = [\y q])))\ + using "equiv-desc-eq:3"[THEN "\E", OF "&I", OF "uni-tv"] by simp + thus ?thesis + using "=\<^sub>d\<^sub>fI"(1)[OF "the-tv-p", OF "uni-tv"] by fast +qed + +AOT_theorem "tv-id:2": \\p\<^bold>\p\ +proof - + AOT_modally_strict { + AOT_have \(p \ p) & [\y p] = [\y p]\ + by (auto simp: "prop-prop2:2" "rule=I:1" intro!: "\I" "\I" "&I") + AOT_hence \\q ((q \ p) & [\y p] = [\y q])\ + using "\I" by fast + } + AOT_hence \\<^bold>\\q ((q \ p) & [\y p] = [\y q])\ + using "RA[2]" by blast + AOT_hence \\<^bold>\x(A!x & \F (x[F] \ \q ((q \ p) & F = [\y q])))[\y p]\ + by (safe intro!: "desc-nec-encode:1"[unvarify F, THEN "\E"(2)] "cqt:2") + AOT_hence \\<^bold>\x(A!x & \F (x[F] \ \q ((q \ p) & F = [\y q])))\<^bold>\p\ + by (safe intro!: "prop-enc"[THEN "\\<^sub>d\<^sub>fI"] "&I" "A-descriptions") + AOT_thus \\p\<^bold>\p\ + by (rule "rule=E"[rotated, OF "tv-id:1"[symmetric]]) +qed + +(* TODO more theorems *) + +AOT_theorem "TV-lem1:1": + \p \ \F(\q (q & F = [\y q]) \ \q((q \ p) & F = [\y q]))\ +proof(safe intro!: "\I" "\I" GEN) + fix F + AOT_assume \\q (q & F = [\y q])\ + then AOT_obtain q where \q & F = [\y q]\ using "\E"[rotated] by blast + moreover AOT_assume p + ultimately AOT_have \(q \ p) & F = [\y q]\ + by (metis "&I" "&E"(1) "&E"(2) "deduction-theorem" "\I") + AOT_thus \\q ((q \ p) & F = [\y q])\ by (rule "\I") +next + fix F + AOT_assume \\q ((q \ p) & F = [\y q])\ + then AOT_obtain q where \(q \ p) & F = [\y q]\ using "\E"[rotated] by blast + moreover AOT_assume p + ultimately AOT_have \q & F = [\y q]\ + by (metis "&I" "&E"(1) "&E"(2) "\E"(2)) + AOT_thus \\q (q & F = [\y q])\ by (rule "\I") +next + AOT_assume \\F (\q (q & F = [\y q]) \ \q ((q \ p) & F = [\y q]))\ + AOT_hence \\q (q & [\y p] = [\y q]) \ \q ((q \ p) & [\y p] = [\y q])\ + using "\E"(1)[rotated, OF "prop-prop2:2"] by blast + moreover AOT_have \\q ((q \ p) & [\y p] = [\y q])\ + by (rule "\I"(2)[where \=p]) + (simp add: "rule=I:1" "&I" "oth-class-taut:3:a" "prop-prop2:2") + ultimately AOT_have \\q (q & [\y p] = [\y q])\ using "\E"(2) by blast + then AOT_obtain q where \q & [\y p] = [\y q]\ using "\E"[rotated] by blast + AOT_thus \p\ + using "rule=E" "&E"(1) "&E"(2) id_sym "\E"(2) "p-identity-thm2:3" by fast +qed + +AOT_theorem "TV-lem1:2": + \\p \ \F(\q (\q & F = [\y q]) \ \q((q \ p) & F = [\y q]))\ +proof(safe intro!: "\I" "\I" GEN) + fix F + AOT_assume \\q (\q & F = [\y q])\ + then AOT_obtain q where \\q & F = [\y q]\ using "\E"[rotated] by blast + moreover AOT_assume \\p\ + ultimately AOT_have \(q \ p) & F = [\y q]\ + by (metis "&I" "&E"(1) "&E"(2) "deduction-theorem" "\I" "raa-cor:3") + AOT_thus \\q ((q \ p) & F = [\y q])\ by (rule "\I") +next + fix F + AOT_assume \\q ((q \ p) & F = [\y q])\ + then AOT_obtain q where \(q \ p) & F = [\y q]\ using "\E"[rotated] by blast + moreover AOT_assume \\p\ + ultimately AOT_have \\q & F = [\y q]\ + by (metis "&I" "&E"(1) "&E"(2) "\E"(1) "raa-cor:3") + AOT_thus \\q (\q & F = [\y q])\ by (rule "\I") +next + AOT_assume \\F (\q (\q & F = [\y q]) \ \q ((q \ p) & F = [\y q]))\ + AOT_hence \\q (\q & [\y p] = [\y q]) \ \q ((q \ p) & [\y p] = [\y q])\ + using "\E"(1)[rotated, OF "prop-prop2:2"] by blast + moreover AOT_have \\q ((q \ p) & [\y p] = [\y q])\ + by (rule "\I"(2)[where \=p]) + (simp add: "rule=I:1" "&I" "oth-class-taut:3:a" "prop-prop2:2") + ultimately AOT_have \\q (\q & [\y p] = [\y q])\ using "\E"(2) by blast + then AOT_obtain q where \\q & [\y p] = [\y q]\ using "\E"[rotated] by blast + AOT_thus \\p\ + using "rule=E" "&E"(1) "&E"(2) id_sym "\E"(2) "p-identity-thm2:3" by fast +qed + + +AOT_define TruthValue :: \\ \ \\ (\TruthValue'(_')\) + "T-value": \TruthValue(x) \\<^sub>d\<^sub>f \p (TruthValueOf(x,p))\ + +(* TODO more theorems *) + +AOT_act_theorem "T-lem:1": \TruthValueOf(\p, p)\ +proof - + AOT_have \: \\p = \<^bold>\x TruthValueOf(x, p)\ + using "rule-id-df:1" "the-tv-p" "uni-tv" by blast + moreover AOT_have \\p\\ + using "t=t-proper:1" calculation "vdash-properties:10" by blast + ultimately show ?thesis by (metis "rule=E" id_sym "vdash-properties:10" "y-in:3") +qed + +AOT_act_theorem "T-lem:2": \\F (\p[F] \ \q((q \ p) & F = [\y q]))\ + using "T-lem:1"[THEN "tv-p"[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(2)]. + +AOT_act_theorem "T-lem:3": \\p\<^bold>\r \ (r \ p)\ +proof - + AOT_have \: \\p[\y r] \ \q ((q \ p) & [\y r] = [\y q])\ + using "T-lem:2"[THEN "\E"(1), OF "prop-prop2:2"]. + show ?thesis + proof(rule "\I"; rule "\I") + AOT_assume \\p\<^bold>\r\ + AOT_hence \\p[\y r]\ by (metis "\\<^sub>d\<^sub>fE" "&E"(2) "prop-enc") + AOT_hence \\q ((q \ p) & [\y r] = [\y q])\ using \ "\E"(1) by blast + then AOT_obtain q where \(q \ p) & [\y r] = [\y q]\ using "\E"[rotated] by blast + moreover AOT_have \r = q\ using calculation + using "&E"(2) "\E"(2) "p-identity-thm2:3" by blast + ultimately AOT_show \r \ p\ + by (metis "rule=E" "&E"(1) "\E"(6) "oth-class-taut:3:a") + next + AOT_assume \r \ p\ + moreover AOT_have \[\y r] = [\y r]\ + by (simp add: "rule=I:1" "prop-prop2:2") + ultimately AOT_have \(r \ p) & [\y r] = [\y r]\ using "&I" by blast + AOT_hence \\q ((q \ p) & [\y r] = [\y q])\ by (rule "\I"(2)[where \=r]) + AOT_hence \\p[\y r]\ using \ "\E"(2) by blast + AOT_thus \\p\<^bold>\r\ + by (metis "\\<^sub>d\<^sub>fI" "&I" "prop-enc" "russell-axiom[enc,1].\_denotes_asm") + qed +qed + +AOT_act_theorem "T-lem:4": \TruthValueOf(x, p) \ x = \p\ +proof - + AOT_have \\x (x = \<^bold>\x TruthValueOf(x, p) \ \z (TruthValueOf(z, p) \ z = x))\ + by (simp add: "fund-cont-desc" GEN) + moreover AOT_have \\p\\ + using "\\<^sub>d\<^sub>fE" "tv-id:2" "&E"(1) "prop-enc" by blast + ultimately AOT_have + \(\p = \<^bold>\x TruthValueOf(x, p)) \ \z (TruthValueOf(z, p) \ z = \p)\ + using "\E"(1) by blast + AOT_hence \\z (TruthValueOf(z, p) \ z = \p)\ + using "\E"(1) "rule-id-df:1" "the-tv-p" "uni-tv" by blast + AOT_thus \TruthValueOf(x, p) \ x = \p\ using "\E"(2) by blast +qed + + +(* TODO more theorems *) + +AOT_theorem "TV-lem2:1": + \(A!x & \F (x[F] \ \q (q & F = [\y q]))) \ TruthValue(x)\ +proof(safe intro!: "\I" "T-value"[THEN "\\<^sub>d\<^sub>fI"] "tv-p"[THEN "\\<^sub>d\<^sub>fI"] + "\I"(1)[rotated, OF "log-prop-prop:2"]) + AOT_assume \[A!]x & \F (x[F] \ \q (q & F = [\y q]))\ + AOT_thus \[A!]x & \F (x[F] \ \q ((q \ (\p (p \ p))) & F = [\y q]))\ + apply (AOT_subst \\q ((q \ (\p (p \ p))) & F = [\y q])\ + \\q (q & F = [\y q])\ for: F :: \<\>\) + apply (AOT_subst \q \ \p (p \p)\ \q\ for: q) + apply (metis (no_types, lifting) "\I" "\I" "\E"(2) GEN) + by (auto simp: "cqt-further:7") +qed + +AOT_theorem "TV-lem2:2": + \(A!x & \F (x[F] \ \q (\q & F = [\y q]))) \ TruthValue(x)\ +proof(safe intro!: "\I" "T-value"[THEN "\\<^sub>d\<^sub>fI"] "tv-p"[THEN "\\<^sub>d\<^sub>fI"] + "\I"(1)[rotated, OF "log-prop-prop:2"]) + AOT_assume \[A!]x & \F (x[F] \ \q (\q & F = [\y q]))\ + AOT_thus \[A!]x & \F (x[F] \ \q ((q \ (\p (p & \p))) & F = [\y q]))\ + apply (AOT_subst \\q ((q \ (\p (p & \p))) & F = [\y q])\ + \\q (\q & F = [\y q])\ for: F :: \<\>\) + apply (AOT_subst \q \ \p (p & \p)\ \\q\ for: q) + apply (metis (no_types, lifting) + "\I" "\E" "\E"(1) "\I" "raa-cor:1" "raa-cor:3") + by (auto simp add: "cqt-further:7") +qed + +AOT_define TheTrue :: \\<^sub>s (\\\) + "the-true:1": \\ =\<^sub>d\<^sub>f \<^bold>\x (A!x & \F (x[F] \ \p(p & F = [\y p])))\ +AOT_define TheFalse :: \\<^sub>s (\\\) + "the-true:2": \\ =\<^sub>d\<^sub>f \<^bold>\x (A!x & \F (x[F] \ \p(\p & F = [\y p])))\ + + +AOT_theorem "the-true:3": \\ \ \\ +proof(safe intro!: "ab-obey:2"[unvarify x y, THEN "\E", rotated 2, OF "\I"(1)] + "\I"(1)[where \=\\[\x \q(q \ q)]\\] "&I" "prop-prop2:2") + AOT_have false_def: \\ = \<^bold>\x (A!x & \F (x[F] \ \p(\p & F = [\y p])))\ + by (simp add: "A-descriptions" "rule-id-df:1[zero]" "the-true:2") + moreover AOT_show false_den: \\\\ + by (meson "\E" "t=t-proper:1" "A-descriptions" + "rule-id-df:1[zero]" "the-true:2") + ultimately AOT_have false_prop: \\<^bold>\(A!\ & \F (\[F] \ \p(\p & F = [\y p])))\ + using "nec-hintikka-scheme"[unvarify x, THEN "\E"(1), THEN "&E"(1)] by blast + AOT_hence \\<^bold>\\F (\[F] \ \p(\p & F = [\y p]))\ + using "Act-Basic:2" "&E"(2) "\E"(1) by blast + AOT_hence \\F \<^bold>\(\[F] \ \p(\p & F = [\y p]))\ + using "\E"(1) "logic-actual-nec:3"[axiom_inst] by blast + AOT_hence false_enc_cond: + \\<^bold>\(\[\x \q(q \ q)] \ \p(\p & [\x \q(q \ q)] = [\y p]))\ + using "\E"(1)[rotated, OF "prop-prop2:2"] by blast + + AOT_have true_def: \\ = \<^bold>\x (A!x & \F (x[F] \ \p(p & F = [\y p])))\ + by (simp add: "A-descriptions" "rule-id-df:1[zero]" "the-true:1") + moreover AOT_show true_den: \\\\ + by (meson "t=t-proper:1" "A-descriptions" "rule-id-df:1[zero]" "the-true:1" "\E") + ultimately AOT_have true_prop: \\<^bold>\(A!\ & \F (\[F] \ \p(p & F = [\y p])))\ + using "nec-hintikka-scheme"[unvarify x, THEN "\E"(1), THEN "&E"(1)] by blast + AOT_hence \\<^bold>\\F (\[F] \ \p(p & F = [\y p]))\ + using "Act-Basic:2" "&E"(2) "\E"(1) by blast + AOT_hence \\F \<^bold>\(\[F] \ \p(p & F = [\y p]))\ + using "\E"(1) "logic-actual-nec:3"[axiom_inst] by blast + AOT_hence \\<^bold>\(\[\x \q(q \ q)] \ \p(p & [\x \q(q \ q)] = [\y p]))\ + using "\E"(1)[rotated, OF "prop-prop2:2"] by blast + moreover AOT_have \\<^bold>\\p(p & [\x \q(q \ q)] = [\y p])\ + by (safe intro!: "nec-imp-act"[THEN "\E"] RN "\I"(1)[where \="\\q(q \ q)\"] "&I" + GEN "\I" "log-prop-prop:2" "rule=I:1" "prop-prop2:2") + ultimately AOT_have \\<^bold>\(\[\x \q(q \ q)])\ + using "Act-Basic:5" "\E"(1,2) by blast + AOT_thus \\[\x \q(q \ q)]\ + using "en-eq:10[1]"[unvarify x\<^sub>1 F, THEN "\E"(1)] true_den "prop-prop2:2" by blast + + AOT_show \\\[\x \q(q \ q)]\ + proof(rule "raa-cor:2") + AOT_assume \\[\x \q(q \ q)]\ + AOT_hence \\<^bold>\\[\x \q(q \ q)]\ + using "en-eq:10[1]"[unvarify x\<^sub>1 F, THEN "\E"(2)] + false_den "prop-prop2:2" by blast + AOT_hence \\<^bold>\\p(\p & [\x \q(q \ q)] = [\y p])\ + using false_enc_cond "Act-Basic:5" "\E"(1) by blast + AOT_hence \\p \<^bold>\(\p & [\x \q(q \ q)] = [\y p])\ + using "Act-Basic:10" "\E"(1) by blast + then AOT_obtain p where p_prop: \\<^bold>\(\p & [\x \q(q \ q)] = [\y p])\ + using "\E"[rotated] by blast + AOT_hence \\<^bold>\[\x \q(q \ q)] = [\y p]\ + by (metis "Act-Basic:2" "&E"(2) "\E"(1)) + AOT_hence \[\x \q(q \ q)] = [\y p]\ + using "id-act:1"[unvarify \ \, THEN "\E"(2)] "prop-prop2:2" by blast + AOT_hence \(\q(q \ q)) = p\ + using "p-identity-thm2:3"[unvarify p, THEN "\E"(2)] + "log-prop-prop:2" by blast + moreover AOT_have \\<^bold>\\p\ using p_prop + using "Act-Basic:2" "&E"(1) "\E"(1) by blast + ultimately AOT_have \\<^bold>\\\q(q \ q)\ + by (metis "Act-Sub:1" "\E"(1,2) "raa-cor:3" "rule=E") + moreover AOT_have \\\<^bold>\\\q(q \ q)\ + by (meson "Act-Sub:1" "RA[2]" "if-p-then-p" "\E"(1) "universal-cor") + ultimately AOT_show \\<^bold>\\\q(q \ q) & \\<^bold>\\\q(q \ q)\ + using "&I" by blast + qed +qed + +AOT_act_theorem "T-T-value:1": \TruthValue(\)\ +proof - + AOT_have true_def: \\ = \<^bold>\x (A!x & \F (x[F] \ \p(p & F = [\y p])))\ + by (simp add: "A-descriptions" "rule-id-df:1[zero]" "the-true:1") + AOT_hence true_den: \\\\ + using "t=t-proper:1" "vdash-properties:6" by blast + AOT_show \TruthValue(\)\ + using "y-in:2"[unvarify z, OF true_den, THEN "\E", OF true_def] + "TV-lem2:1"[unvarify x, OF true_den, THEN "\E"] by blast +qed + +AOT_act_theorem "T-T-value:2": \TruthValue(\)\ +proof - + AOT_have false_def: \\ = \<^bold>\x (A!x & \F (x[F] \ \p(\p & F = [\y p])))\ + by (simp add: "A-descriptions" "rule-id-df:1[zero]" "the-true:2") + AOT_hence false_den: \\\\ + using "t=t-proper:1" "vdash-properties:6" by blast + AOT_show \TruthValue(\)\ + using "y-in:2"[unvarify z, OF false_den, THEN "\E", OF false_def] + "TV-lem2:2"[unvarify x, OF false_den, THEN "\E"] by blast +qed + +AOT_theorem "two-T": \\x\y(TruthValue(x) & TruthValue(y) & x \ y & + \z (TruthValue(z) \ z = x \ z = y))\ +proof - + AOT_obtain a where a_prop: \A!a & \F (a[F] \ \p (p & F = [\y p]))\ + using "A-objects"[axiom_inst] "\E"[rotated] by fast + AOT_obtain b where b_prop: \A!b & \F (b[F] \ \p (\p & F = [\y p]))\ + using "A-objects"[axiom_inst] "\E"[rotated] by fast + AOT_obtain p where p: p + by (metis "log-prop-prop:2" "raa-cor:3" "rule-ui:1" "universal-cor") + show ?thesis + proof(rule "\I"(2)[where \=a]; rule "\I"(2)[where \=b]; + safe intro!: "&I" GEN "\I") + AOT_show \TruthValue(a)\ + using "TV-lem2:1" a_prop "vdash-properties:10" by blast + next + AOT_show \TruthValue(b)\ + using "TV-lem2:2" b_prop "vdash-properties:10" by blast + next + AOT_show \a \ b\ + proof(rule "ab-obey:2"[THEN "\E", OF "\I"(1)]) + AOT_show \\F (a[F] & \b[F])\ + proof(rule "\I"(1)[where \="\[\y p]\"]; rule "&I" "prop-prop2:2") + AOT_show \a[\y p]\ + by(safe intro!: "\I"(2)[where \=p] "&I" p "rule=I:1"[OF "prop-prop2:2"] + a_prop[THEN "&E"(2), THEN "\E"(1), THEN "\E"(2), OF "prop-prop2:2"]) + next + AOT_show \\b[\y p]\ + proof (rule "raa-cor:2") + AOT_assume \b[\y p]\ + AOT_hence \\q (\q & [\y p] = [\y q])\ + using "\E"(1)[rotated, OF "prop-prop2:2", THEN "\E"(1)] + b_prop[THEN "&E"(2)] by fast + then AOT_obtain q where \\q & [\y p] = [\y q]\ + using "\E"[rotated] by blast + AOT_hence \\p\ + by (metis "rule=E" "&E"(1) "&E"(2) "deduction-theorem" "\I" + "\E"(2) "p-identity-thm2:3" "raa-cor:3") + AOT_thus \p & \p\ using p "&I" by blast + qed + qed + qed + next + fix z + AOT_assume \TruthValue(z)\ + AOT_hence \\p (TruthValueOf(z, p))\ + by (metis "\\<^sub>d\<^sub>fE" "T-value") + then AOT_obtain p where \TruthValueOf(z, p)\ using "\E"[rotated] by blast + AOT_hence z_prop: \A!z & \F (z[F] \ \q ((q \ p) & F = [\y q]))\ + using "\\<^sub>d\<^sub>fE" "tv-p" by blast + { + AOT_assume p: \p\ + AOT_have \z = a\ + proof(rule "ab-obey:1"[THEN "\E", THEN "\E", OF "&I", + OF z_prop[THEN "&E"(1)], OF a_prop[THEN "&E"(1)]]; + rule GEN) + fix G + AOT_have \z[G] \ \q ((q \ p) & G = [\y q])\ + using z_prop[THEN "&E"(2)] "\E"(2) by blast + also AOT_have \\q ((q \ p) & G = [\y q]) \ \q (q & G = [\y q])\ + using "TV-lem1:1"[THEN "\E"(1), OF p, THEN "\E"(2)[where \=G], symmetric]. + also AOT_have \\ \ a[G]\ + using a_prop[THEN "&E"(2), THEN "\E"(2)[where \=G], symmetric]. + finally AOT_show \z[G] \ a[G]\. + qed + AOT_hence \z = a \ z = b\ by (rule "\I") + } + moreover { + AOT_assume notp: \\p\ + AOT_have \z = b\ + proof(rule "ab-obey:1"[THEN "\E", THEN "\E", OF "&I", + OF z_prop[THEN "&E"(1)], OF b_prop[THEN "&E"(1)]]; + rule GEN) + fix G + AOT_have \z[G] \ \q ((q \ p) & G = [\y q])\ + using z_prop[THEN "&E"(2)] "\E"(2) by blast + also AOT_have \\q ((q \ p) & G = [\y q]) \ \q (\q & G = [\y q])\ + using "TV-lem1:2"[THEN "\E"(1), OF notp, THEN "\E"(2), symmetric]. + also AOT_have \\ \ b[G]\ + using b_prop[THEN "&E"(2), THEN "\E"(2), symmetric]. + finally AOT_show \z[G] \ b[G]\. + qed + AOT_hence \z = a \ z = b\ by (rule "\I") + } + ultimately AOT_show \z = a \ z = b\ + by (metis "reductio-aa:1") + qed +qed + +AOT_act_theorem "valueof-facts:1": \TruthValueOf(x, p) \ (p \ x = \)\ +proof(safe intro!: "\I" dest!: "tv-p"[THEN "\\<^sub>d\<^sub>fE"]) + AOT_assume \: \[A!]x & \F (x[F] \ \q ((q \ p) & F = [\y q]))\ + AOT_have a: \A!\\ + using "\E" "T-T-value:1" "T-value" "&E"(1) "\\<^sub>d\<^sub>fE" "tv-p" by blast + AOT_have true_def: \\ = \<^bold>\x (A!x & \F (x[F] \ \p(p & F = [\y p])))\ + by (simp add: "A-descriptions" "rule-id-df:1[zero]" "the-true:1") + AOT_hence true_den: \\\\ + using "t=t-proper:1" "vdash-properties:6" by blast + AOT_have b: \\F (\[F] \ \q (q & F = [\y q]))\ + using "y-in:2"[unvarify z, OF true_den, THEN "\E", OF true_def] "&E" by blast + AOT_show \p \ x = \\ + proof(safe intro!: "\I" "\I") + AOT_assume p + AOT_hence \\F (\q (q & F = [\y q]) \ \q ((q \ p) & F = [\y q]))\ + using "TV-lem1:1"[THEN "\E"(1)] by blast + AOT_hence \\F(\[F] \ \q ((q \ p) & F = [\y q]))\ + using b "cqt-basic:10"[THEN "\E", OF "&I", OF b] by fast + AOT_hence c: \\F(\q((q \ p) & F = [\y q]) \ \[F])\ + using "cqt-basic:11"[THEN "\E"(1)] by fast + AOT_hence \\F (x[F] \ \[F])\ + using "cqt-basic:10"[THEN "\E", OF "&I", OF \[THEN "&E"(2)]] by fast + AOT_thus \x = \\ + by (rule "ab-obey:1"[unvarify y, OF true_den, THEN "\E", THEN "\E", + OF "&I", OF \[THEN "&E"(1)], OF a]) + next + AOT_assume \x = \\ + AOT_hence d: \\F (\[F] \ \q ((q \ p) & F = [\y q]))\ + using "rule=E" \[THEN "&E"(2)] by fast + AOT_have \\F (\q (q & F = [\y q]) \ \q ((q \ p) & F = [\y q]))\ + using "cqt-basic:10"[THEN "\E", OF "&I", + OF b[THEN "cqt-basic:11"[THEN "\E"(1)]], OF d]. + AOT_thus p using "TV-lem1:1"[THEN "\E"(2)] by blast + qed +qed + +AOT_act_theorem "valueof-facts:2": \TruthValueOf(x, p) \ (\p \ x = \)\ +proof(safe intro!: "\I" dest!: "tv-p"[THEN "\\<^sub>d\<^sub>fE"]) + AOT_assume \: \[A!]x & \F (x[F] \ \q ((q \ p) & F = [\y q]))\ + AOT_have a: \A!\\ + using "\E" "T-T-value:2" "T-value" "&E"(1) "\\<^sub>d\<^sub>fE" "tv-p" by blast + AOT_have false_def: \\ = \<^bold>\x (A!x & \F (x[F] \ \p(\p & F = [\y p])))\ + by (simp add: "A-descriptions" "rule-id-df:1[zero]" "the-true:2") + AOT_hence false_den: \\\\ + using "t=t-proper:1" "vdash-properties:6" by blast + AOT_have b: \\F (\[F] \ \q (\q & F = [\y q]))\ + using "y-in:2"[unvarify z, OF false_den, THEN "\E", OF false_def] "&E" by blast + AOT_show \\p \ x = \\ + proof(safe intro!: "\I" "\I") + AOT_assume \\p\ + AOT_hence \\F (\q (\q & F = [\y q]) \ \q ((q \ p) & F = [\y q]))\ + using "TV-lem1:2"[THEN "\E"(1)] by blast + AOT_hence \\F(\[F] \ \q ((q \ p) & F = [\y q]))\ + using b "cqt-basic:10"[THEN "\E", OF "&I", OF b] by fast + AOT_hence c: \\F(\q((q \ p) & F = [\y q]) \ \[F])\ + using "cqt-basic:11"[THEN "\E"(1)] by fast + AOT_hence \\F (x[F] \ \[F])\ + using "cqt-basic:10"[THEN "\E", OF "&I", OF \[THEN "&E"(2)]] by fast + AOT_thus \x = \\ + by (rule "ab-obey:1"[unvarify y, OF false_den, THEN "\E", THEN "\E", + OF "&I", OF \[THEN "&E"(1)], OF a]) + next + AOT_assume \x = \\ + AOT_hence d: \\F (\[F] \ \q ((q \ p) & F = [\y q]))\ + using "rule=E" \[THEN "&E"(2)] by fast + AOT_have \\F (\q (\q & F = [\y q]) \ \q ((q \ p) & F = [\y q]))\ + using "cqt-basic:10"[THEN "\E", OF "&I", + OF b[THEN "cqt-basic:11"[THEN "\E"(1)]], OF d]. + AOT_thus \\p\ using "TV-lem1:2"[THEN "\E"(2)] by blast + qed +qed + +AOT_act_theorem "q-True:1": \p \ (\p = \)\ + apply (rule "valueof-facts:1"[unvarify x, THEN "\E", rotated, OF "T-lem:1"]) + using "\\<^sub>d\<^sub>fE" "tv-id:2" "&E"(1) "prop-enc" by blast + +AOT_act_theorem "q-True:2": \\p \ (\p = \)\ + apply (rule "valueof-facts:2"[unvarify x, THEN "\E", rotated, OF "T-lem:1"]) + using "\\<^sub>d\<^sub>fE" "tv-id:2" "&E"(1) "prop-enc" by blast + +AOT_act_theorem "q-True:3": \p \ \\<^bold>\p\ +proof(safe intro!: "\I" "\I") + AOT_assume p + AOT_hence \\p = \\ by (metis "\E"(1) "q-True:1") + moreover AOT_have \\p\<^bold>\p\ + by (simp add: "tv-id:2") + ultimately AOT_show \\\<^bold>\p\ + using "rule=E" "T-lem:4" by fast +next + AOT_have true_def: \\ = \<^bold>\x (A!x & \F (x[F] \ \p(p & F = [\y p])))\ + by (simp add: "A-descriptions" "rule-id-df:1[zero]" "the-true:1") + AOT_hence true_den: \\\\ + using "t=t-proper:1" "vdash-properties:6" by blast + AOT_have b: \\F (\[F] \ \q (q & F = [\y q]))\ + using "y-in:2"[unvarify z, OF true_den, THEN "\E", OF true_def] "&E" by blast + + AOT_assume \\\<^bold>\p\ + AOT_hence \\[\y p]\ by (metis "\\<^sub>d\<^sub>fE" "&E"(2) "prop-enc") + AOT_hence \\q (q & [\y p] = [\y q])\ + using b[THEN "\E"(1), OF "prop-prop2:2", THEN "\E"(1)] by blast + then AOT_obtain q where \q & [\y p] = [\y q]\ using "\E"[rotated] by blast + AOT_thus \p\ + using "rule=E" "&E"(1) "&E"(2) id_sym "\E"(2) "p-identity-thm2:3" by fast +qed + + +AOT_act_theorem "q-True:5": \\p \ \\<^bold>\p\ +proof(safe intro!: "\I" "\I") + AOT_assume \\p\ + AOT_hence \\p = \\ by (metis "\E"(1) "q-True:2") + moreover AOT_have \\p\<^bold>\p\ + by (simp add: "tv-id:2") + ultimately AOT_show \\\<^bold>\p\ + using "rule=E" "T-lem:4" by fast +next + AOT_have false_def: \\ = \<^bold>\x (A!x & \F (x[F] \ \p(\p & F = [\y p])))\ + by (simp add: "A-descriptions" "rule-id-df:1[zero]" "the-true:2") + AOT_hence false_den: \\\\ + using "t=t-proper:1" "vdash-properties:6" by blast + AOT_have b: \\F (\[F] \ \q (\q & F = [\y q]))\ + using "y-in:2"[unvarify z, OF false_den, THEN "\E", OF false_def] "&E" by blast + + AOT_assume \\\<^bold>\p\ + AOT_hence \\[\y p]\ by (metis "\\<^sub>d\<^sub>fE" "&E"(2) "prop-enc") + AOT_hence \\q (\q & [\y p] = [\y q])\ + using b[THEN "\E"(1), OF "prop-prop2:2", THEN "\E"(1)] by blast + then AOT_obtain q where \\q & [\y p] = [\y q]\ using "\E"[rotated] by blast + AOT_thus \\p\ + using "rule=E" "&E"(1) "&E"(2) id_sym "\E"(2) "p-identity-thm2:3" by fast +qed + +AOT_act_theorem "q-True:4": \p \ \(\\<^bold>\p)\ + using "q-True:5" + by (metis "deduction-theorem" "\I" "\E"(2) "\E"(4) "raa-cor:3") + +AOT_act_theorem "q-True:6": \\p \ \(\\<^bold>\p)\ + using "\E"(1) "oth-class-taut:4:b" "q-True:3" by blast + +AOT_define ExtensionOf :: \\ \ \ \ \\ (\ExtensionOf'(_,_')\) + "exten-p": \ExtensionOf(x,p) \\<^sub>d\<^sub>f A!x & + \F (x[F] \ Propositional([F])) & + \q ((x\<^bold>\q) \ (q \ p))\ + +AOT_theorem "extof-e": \ExtensionOf(x, p) \ TruthValueOf(x, p)\ +proof (safe intro!: "\I" "\I" "tv-p"[THEN "\\<^sub>d\<^sub>fI"] "exten-p"[THEN "\\<^sub>d\<^sub>fI"] + dest!: "tv-p"[THEN "\\<^sub>d\<^sub>fE"] "exten-p"[THEN "\\<^sub>d\<^sub>fE"]) + AOT_assume 1: \[A!]x & \F (x[F] \ Propositional([F])) & \q (x \<^bold>\ q \ (q \ p))\ + AOT_hence \: \[A!]x & \F (x[F] \ \q(F = [\y q])) & \q (x \<^bold>\ q \ (q \ p))\ + by (AOT_subst \\q(F = [\y q])\ \Propositional([F])\ for: F :: \<\>\) + (auto simp add: "df-rules-formulas[3]" "df-rules-formulas[4]" + "\I" "prop-prop1") + AOT_show \[A!]x & \F (x[F] \ \q ((q \ p) & F = [\y q]))\ + proof(safe intro!: "&I" GEN 1[THEN "&E"(1), THEN "&E"(1)] "\I" "\I") + fix F + AOT_assume 0: \x[F]\ + AOT_hence \\q (F = [\y q])\ + using \[THEN "&E"(1), THEN "&E"(2)] "\E"(2) "\E" by blast + then AOT_obtain q where q_prop: \F = [\y q]\ using "\E"[rotated] by blast + AOT_hence \x[\y q]\ using 0 "rule=E" by blast + AOT_hence \x\<^bold>\q\ by (metis "\\<^sub>d\<^sub>fI" "&I" "ex:1:a" "prop-enc" "rule-ui:3") + AOT_hence \q \ p\ using \[THEN "&E"(2)] "\E"(2) "\E"(1) by blast + AOT_hence \(q \ p) & F = [\y q]\ using q_prop "&I" by blast + AOT_thus \\q ((q \ p) & F = [\y q])\ by (rule "\I") + next + fix F + AOT_assume \\q ((q \ p) & F = [\y q])\ + then AOT_obtain q where q_prop: \(q \ p) & F = [\y q]\ + using "\E"[rotated] by blast + AOT_hence \x\<^bold>\q\ using \[THEN "&E"(2)] "\E"(2) "&E" "\E"(2) by blast + AOT_hence \x[\y q]\ by (metis "\\<^sub>d\<^sub>fE" "&E"(2) "prop-enc") + AOT_thus \x[F]\ using q_prop[THEN "&E"(2), symmetric] "rule=E" by blast + qed +next + AOT_assume 0: \[A!]x & \F (x[F] \ \q ((q \ p) & F = [\y q]))\ + AOT_show \[A!]x & \F (x[F] \ Propositional([F])) & \q (x \<^bold>\ q \ (q \ p))\ + proof(safe intro!: "&I" 0[THEN "&E"(1)] GEN "\I") + fix F + AOT_assume \x[F]\ + AOT_hence \\q ((q \ p) & F = [\y q])\ + using 0[THEN "&E"(2)] "\E"(2) "\E"(1) by blast + then AOT_obtain q where \(q \ p) & F = [\y q]\ + using "\E"[rotated] by blast + AOT_hence \F = [\y q]\ using "&E"(2) by blast + AOT_hence \\q F = [\y q]\ by (rule "\I") + AOT_thus \Propositional([F])\ by (metis "\\<^sub>d\<^sub>fI" "prop-prop1") + next + AOT_show \x\<^bold>\r \ (r \ p)\ for r + proof(rule "\I"; rule "\I") + AOT_assume \x\<^bold>\r\ + AOT_hence \x[\y r]\ by (metis "\\<^sub>d\<^sub>fE" "&E"(2) "prop-enc") + AOT_hence \\q ((q \ p) & [\y r] = [\y q])\ + using 0[THEN "&E"(2), THEN "\E"(1), OF "prop-prop2:2", THEN "\E"(1)] by blast + then AOT_obtain q where \(q \ p) & [\y r] = [\y q]\ + using "\E"[rotated] by blast + AOT_thus \r \ p\ + by (metis "rule=E" "&E"(1,2) id_sym "\E"(2) "Commutativity of \" + "p-identity-thm2:3") + next + AOT_assume \r \ p\ + AOT_hence \(r \ p) & [\y r] = [\y r]\ + by (metis "rule=I:1" "\S"(1) "\E"(2) "Commutativity of &" "prop-prop2:2") + AOT_hence \\q ((q \ p) & [\y r] = [\y q])\ by (rule "\I") + AOT_hence \x[\y r]\ + using 0[THEN "&E"(2), THEN "\E"(1), OF "prop-prop2:2", THEN "\E"(2)] by blast + AOT_thus \x\<^bold>\r\ by (metis "\\<^sub>d\<^sub>fI" "&I" "ex:1:a" "prop-enc" "rule-ui:3") + qed + qed +qed + +AOT_theorem "ext-p-tv:1": \\!x ExtensionOf(x, p)\ + by (AOT_subst \ExtensionOf(x, p)\ \TruthValueOf(x, p)\ for: x) + (auto simp: "extof-e" "p-has-!tv:2") + +AOT_theorem "ext-p-tv:2": \\<^bold>\x(ExtensionOf(x, p))\\ + using "A-Exists:2" "RA[2]" "ext-p-tv:1" "\E"(2) by blast + +AOT_theorem "ext-p-tv:3": \\<^bold>\x(ExtensionOf(x, p)) = \p\ +proof - + AOT_have 0: \\<^bold>\\x(ExtensionOf(x, p) \ TruthValueOf(x,p))\ + by (rule "RA[2]"; rule GEN; rule "extof-e") + AOT_have 1: \\p = \<^bold>\x TruthValueOf(x,p)\ + using "rule-id-df:1" "the-tv-p" "uni-tv" by blast + show ?thesis + apply (rule "equiv-desc-eq:1"[THEN "\E", OF 0, THEN "\E"(1)[where \=\\\p\\], + THEN "\E"(2), symmetric]) + using "1" "t=t-proper:1" "vdash-properties:10" apply blast + by (fact 1) +qed +(*<*)end(*>*) diff --git a/thys/AOT/AOT_Definitions.thy b/thys/AOT/AOT_Definitions.thy new file mode 100644 --- /dev/null +++ b/thys/AOT/AOT_Definitions.thy @@ -0,0 +1,153 @@ +theory AOT_Definitions + imports AOT_semantics +begin + +section\Definitions of AOT\ + +AOT_theorem "conventions:1": \\ & \ \\<^sub>d\<^sub>f \(\ \ \\)\ + using AOT_conj. +AOT_theorem "conventions:2": \\ \ \ \\<^sub>d\<^sub>f \\ \ \\ + using AOT_disj. +AOT_theorem "conventions:3": \\ \ \ \\<^sub>d\<^sub>f (\ \ \) & (\ \ \)\ + using AOT_equiv. +AOT_theorem "conventions:4": \\\ \{\} \\<^sub>d\<^sub>f \\\ \\{\}\ + using AOT_exists. +AOT_theorem "conventions:5": \\\ \\<^sub>d\<^sub>f \\\\\ + using AOT_dia. + +declare "conventions:1"[AOT_defs] "conventions:2"[AOT_defs] + "conventions:3"[AOT_defs] "conventions:4"[AOT_defs] + "conventions:5"[AOT_defs] + +notepad +begin + fix \ \ \ + text\\linelabel{precedence}\ + have "conventions3[1]": \\\ \ \ \ \\ \ \\\ = \(\ \ \) \ (\\ \ \\)\\ + by blast + have "conventions3[2]": \\\ & \ \ \\ = \(\ & \) \ \\\ + and \\\ \ \ \ \\ = \(\ \ \) \ \\\ + by blast+ + have "conventions3[3]": \\\ \ \ & \\ = \(\ \ \) & \\\ + and \\\ & \ \ \\ = \(\ & \) \ \\\ + by blast+ \ \Note that PLM instead generally uses parenthesis in these cases.\ +end + + +AOT_theorem "existence:1": \\\ \\<^sub>d\<^sub>f \F [F]\\ + by (simp add: AOT_sem_denotes AOT_sem_exists AOT_model_equiv_def) + (metis AOT_sem_denotes AOT_sem_exe AOT_sem_lambda_beta AOT_sem_lambda_denotes) +AOT_theorem "existence:2": \\\ \\<^sub>d\<^sub>f \x\<^sub>1...\x\<^sub>n x\<^sub>1...x\<^sub>n[\]\ + using AOT_sem_denotes AOT_sem_enc_denotes AOT_sem_universal_encoder + by (simp add: AOT_sem_denotes AOT_sem_exists AOT_model_equiv_def) blast +AOT_theorem "existence:2[1]": \\\ \\<^sub>d\<^sub>f \x x[\]\ + using "existence:2"[of \] by simp +AOT_theorem "existence:2[2]": \\\ \\<^sub>d\<^sub>f \x\y xy[\]\ + using "existence:2"[of \] + by (simp add: AOT_sem_denotes AOT_sem_exists AOT_model_equiv_def + AOT_model_denotes_prod_def) +AOT_theorem "existence:2[3]": \\\ \\<^sub>d\<^sub>f \x\y\z xyz[\]\ + using "existence:2"[of \] + by (simp add: AOT_sem_denotes AOT_sem_exists AOT_model_equiv_def + AOT_model_denotes_prod_def) +AOT_theorem "existence:2[4]": \\\ \\<^sub>d\<^sub>f \x\<^sub>1\x\<^sub>2\x\<^sub>3\x\<^sub>4 x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[\]\ + using "existence:2"[of \] + by (simp add: AOT_sem_denotes AOT_sem_exists AOT_model_equiv_def + AOT_model_denotes_prod_def) + +AOT_theorem "existence:3": \\\ \\<^sub>d\<^sub>f [\x \]\\ + by (simp add: AOT_sem_denotes AOT_model_denotes_\_def AOT_model_equiv_def + AOT_model_lambda_denotes) + +declare "existence:1"[AOT_defs] "existence:2"[AOT_defs] "existence:2[1]"[AOT_defs] + "existence:2[2]"[AOT_defs] "existence:2[3]"[AOT_defs] + "existence:2[4]"[AOT_defs] "existence:3"[AOT_defs] + + +AOT_theorem "oa:1": \O! =\<^sub>d\<^sub>f [\x \E!x]\ using AOT_ordinary . +AOT_theorem "oa:2": \A! =\<^sub>d\<^sub>f [\x \\E!x]\ using AOT_abstract . + +declare "oa:1"[AOT_defs] "oa:2"[AOT_defs] + +AOT_theorem "identity:1": + \x = y \\<^sub>d\<^sub>f ([O!]x & [O!]y & \\F ([F]x \ [F]y)) \ + ([A!]x & [A!]y & \\F (x[F] \ y[F]))\ + unfolding AOT_model_equiv_def + using AOT_sem_ind_eq[of _ x y] + by (simp add: AOT_sem_ordinary AOT_sem_abstract AOT_sem_conj + AOT_sem_box AOT_sem_equiv AOT_sem_forall AOT_sem_disj AOT_sem_eq + AOT_sem_denotes) + +AOT_theorem "identity:2": + \F = G \\<^sub>d\<^sub>f F\ & G\ & \\x(x[F] \ x[G])\ + using AOT_sem_enc_eq[of _ F G] + by (auto simp: AOT_model_equiv_def AOT_sem_imp AOT_sem_denotes AOT_sem_eq + AOT_sem_conj AOT_sem_forall AOT_sem_box AOT_sem_equiv) + +AOT_theorem "identity:3[2]": + \F = G \\<^sub>d\<^sub>f F\ & G\ & \y([\z [F]zy] = [\z [G]zy] & [\z [F]yz] = [\z [G]yz])\ + by (auto simp: AOT_model_equiv_def AOT_sem_proj_id_prop[of _ F G] + AOT_sem_proj_id_prod_def AOT_sem_conj AOT_sem_denotes + AOT_sem_forall AOT_sem_unary_proj_id AOT_model_denotes_prod_def) +AOT_theorem "identity:3[3]": + \F = G \\<^sub>d\<^sub>f F\ & G\ & \y\<^sub>1\y\<^sub>2([\z [F]zy\<^sub>1y\<^sub>2] = [\z [G]zy\<^sub>1y\<^sub>2] & + [\z [F]y\<^sub>1zy\<^sub>2] = [\z [G]y\<^sub>1zy\<^sub>2] & + [\z [F]y\<^sub>1y\<^sub>2z] = [\z [G]y\<^sub>1y\<^sub>2z])\ + by (auto simp: AOT_model_equiv_def AOT_sem_proj_id_prop[of _ F G] + AOT_sem_proj_id_prod_def AOT_sem_conj AOT_sem_denotes + AOT_sem_forall AOT_sem_unary_proj_id AOT_model_denotes_prod_def) +AOT_theorem "identity:3[4]": + \F = G \\<^sub>d\<^sub>f F\ & G\ & \y\<^sub>1\y\<^sub>2\y\<^sub>3([\z [F]zy\<^sub>1y\<^sub>2y\<^sub>3] = [\z [G]zy\<^sub>1y\<^sub>2y\<^sub>3] & + [\z [F]y\<^sub>1zy\<^sub>2y\<^sub>3] = [\z [G]y\<^sub>1zy\<^sub>2y\<^sub>3] & + [\z [F]y\<^sub>1y\<^sub>2zy\<^sub>3] = [\z [G]y\<^sub>1y\<^sub>2zy\<^sub>3] & + [\z [F]y\<^sub>1y\<^sub>2y\<^sub>3z] = [\z [G]y\<^sub>1y\<^sub>2y\<^sub>3z])\ + by (auto simp: AOT_model_equiv_def AOT_sem_proj_id_prop[of _ F G] + AOT_sem_proj_id_prod_def AOT_sem_conj AOT_sem_denotes + AOT_sem_forall AOT_sem_unary_proj_id AOT_model_denotes_prod_def) +AOT_theorem "identity:3": + \F = G \\<^sub>d\<^sub>f F\ & G\ & \x\<^sub>1...\x\<^sub>n \AOT_sem_proj_id x\<^sub>1x\<^sub>n (\ \ . AOT_exe F \) + (\ \ . AOT_exe G \)\\ + by (auto simp: AOT_model_equiv_def AOT_sem_proj_id_prop[of _ F G] + AOT_sem_proj_id_prod_def AOT_sem_conj AOT_sem_denotes + AOT_sem_forall AOT_sem_unary_proj_id AOT_model_denotes_prod_def) + +AOT_theorem "identity:4": + \p = q \\<^sub>d\<^sub>f p\ & q\ & [\x p] = [\x q]\ + by (auto simp: AOT_model_equiv_def AOT_sem_eq AOT_sem_denotes AOT_sem_conj + AOT_model_lambda_denotes AOT_sem_lambda_eq_prop_eq) + +declare "identity:1"[AOT_defs] "identity:2"[AOT_defs] "identity:3[2]"[AOT_defs] + "identity:3[3]"[AOT_defs] "identity:3[4]"[AOT_defs] "identity:3"[AOT_defs] + "identity:4"[AOT_defs] + +AOT_define AOT_nonidentical :: \\ \ \ \ \\ (infixl "\" 50) + "=-infix": \\ \ \ \\<^sub>d\<^sub>f \(\ = \)\ + +context AOT_meta_syntax +begin +notation AOT_nonidentical (infixl "\<^bold>\" 50) +end +context AOT_no_meta_syntax +begin +no_notation AOT_nonidentical (infixl "\<^bold>\" 50) +end + + +text\The following are purely technical pseudo-definitions required due to + our internal implementation of n-ary relations and ellipses using tuples.\ +AOT_theorem tuple_denotes: \\(\,\')\\ \\<^sub>d\<^sub>f \\ & \'\\ + by (simp add: AOT_model_denotes_prod_def AOT_model_equiv_def + AOT_sem_conj AOT_sem_denotes) +AOT_theorem tuple_identity_1: \\(\,\')\ = \(\, \')\ \\<^sub>d\<^sub>f (\ = \) & (\' = \')\ + by (auto simp: AOT_model_equiv_def AOT_sem_conj AOT_sem_eq + AOT_model_denotes_prod_def AOT_sem_denotes) +AOT_theorem tuple_forall: \\\\<^sub>1...\\\<^sub>n \{\\<^sub>1...\\<^sub>n} \\<^sub>d\<^sub>f \\\<^sub>1(\\\<^sub>2...\\\<^sub>n \{\(\\<^sub>1, \\<^sub>2\\<^sub>n)\})\ + by (auto simp: AOT_model_equiv_def AOT_sem_forall AOT_sem_denotes + AOT_model_denotes_prod_def) +AOT_theorem tuple_exists: \\\\<^sub>1...\\\<^sub>n \{\\<^sub>1...\\<^sub>n} \\<^sub>d\<^sub>f \\\<^sub>1(\\\<^sub>2...\\\<^sub>n \{\(\\<^sub>1, \\<^sub>2\\<^sub>n)\})\ + by (auto simp: AOT_model_equiv_def AOT_sem_exists AOT_sem_denotes + AOT_model_denotes_prod_def) +declare tuple_denotes[AOT_defs] tuple_identity_1[AOT_defs] tuple_forall[AOT_defs] + tuple_exists[AOT_defs] + +end diff --git a/thys/AOT/AOT_ExtendedRelationComprehension.thy b/thys/AOT/AOT_ExtendedRelationComprehension.thy new file mode 100644 --- /dev/null +++ b/thys/AOT/AOT_ExtendedRelationComprehension.thy @@ -0,0 +1,754 @@ +theory AOT_ExtendedRelationComprehension + imports AOT_RestrictedVariables +begin + +section\Extended Relation Comprehension\ + +text\This theory depends on choosing extended models.\ +interpretation AOT_ExtendedModel by (standard; auto) + +text\Auxiliary lemma: negations of denoting relations denote.\ +AOT_theorem negation_denotes: \[\x \{x}]\ \ [\x \\{x}]\\ +proof(rule "\I") + AOT_assume 0: \[\x \{x}]\\ + AOT_show \[\x \\{x}]\\ + proof (rule "safe-ext"[axiom_inst, THEN "\E", OF "&I"]) + AOT_show \[\x \[\x \{x}]x]\\ by "cqt:2" + next + AOT_have \\[\x \{x}]\\ + using 0 "exist-nec"[THEN "\E"] by blast + moreover AOT_have \\[\x \{x}]\ \ \\x (\[\x \{x}]x \ \\{x})\ + by(rule RM; safe intro!: GEN "\I" "\I" "\\C"(2) "\\C"(2) "cqt:2") + ultimately AOT_show \\\x (\[\x \{x}]x \ \\{x})\ + using "\E" by blast + qed +qed + +text\Auxiliary lemma: conjunctions of denoting relations denote.\ +AOT_theorem conjunction_denotes: \[\x \{x}]\ & [\x \{x}]\ \ [\x \{x} & \{x}]\\ +proof(rule "\I") + AOT_assume 0: \[\x \{x}]\ & [\x \{x}]\\ + AOT_show \[\x \{x} & \{x}]\\ + proof (rule "safe-ext"[axiom_inst, THEN "\E", OF "&I"]) + AOT_show \[\x [\x \{x}]x & [\x \{x}]x]\\ by "cqt:2" + next + AOT_have \\([\x \{x}]\ & [\x \{x}]\)\ + using 0 "exist-nec"[THEN "\E"] "&E" + "KBasic:3" "df-simplify:2" "intro-elim:3:b" by blast + moreover AOT_have + \\([\x \{x}]\ & [\x \{x}]\) \ \\x ([\x \{x}]x & [\x \{x}]x \ \{x} & \{x})\ + by(rule RM; auto intro!: GEN "\I" "\I" "cqt:2" "&I" + intro: "\\C" + dest: "&E" "\\C") + ultimately AOT_show \\\x ([\x \{x}]x & [\x \{x}]x \ \{x} & \{x})\ + using "\E" by blast + qed +qed + +AOT_register_rigid_restricted_type + Ordinary: \O!\\ +proof + AOT_modally_strict { + AOT_show \\x O!x\ + by (meson "B\" "T\" "o-objects-exist:1" "\E") + } +next + AOT_modally_strict { + AOT_show \O!\ \ \\\ for \ + by (simp add: "\I" "cqt:5:a[1]"[axiom_inst, THEN "\E", THEN "&E"(2)]) + } +next + AOT_modally_strict { + AOT_show \\\(O!\ \ \O!\)\ + by (simp add: GEN "oa-facts:1") + } +qed + +AOT_register_variable_names + Ordinary: u v r t s + +text\In PLM this is defined in the Natural Numbers chapter, + but since it is helpful for stating the comprehension principles, + we already define it here.\ +AOT_define eqE :: \\ \ \ \ \\ (infixl \\\<^sub>E\ 50) + eqE: \F \\<^sub>E G \\<^sub>d\<^sub>f F\ & G\ & \u ([F]u \ [G]u)\ + +text\Derive existence claims about relations from the axioms.\ +AOT_theorem denotes_all: \[\x \G (\G \\<^sub>E F \ x[G])]\\ + and denotes_all_neg: \[\x \G (\G \\<^sub>E F \ \x[G])]\\ +proof - + AOT_have Aux: \\F (\F \\<^sub>E G \ (x[F] \ x[G])), \(x[G] \ y[G]) + \<^bold>\\<^sub>\ \F([F]x & \[F]y)\ for x y G + proof - + AOT_modally_strict { + AOT_assume 0: \\F (\F \\<^sub>E G \ (x[F] \ x[G]))\ + AOT_assume G_prop: \\(x[G] \ y[G])\ + { + AOT_assume \\\F([F]x & \[F]y)\ + AOT_hence 0: \\F \([F]x & \[F]y)\ + by (metis "cqt-further:4" "\E") + AOT_have \\F ([F]x \ [F]y)\ + proof (rule GEN; rule "\I"; rule "\I") + fix F + AOT_assume \[F]x\ + moreover AOT_have \\([F]x & \[F]y)\ + using 0[THEN "\E"(2)] by blast + ultimately AOT_show \[F]y\ + by (metis "&I" "raa-cor:1") + next + fix F + AOT_assume \[F]y\ + AOT_hence \\[\x \[F]x]y\ + by (metis "\\I" "\\C"(2)) + moreover AOT_have \\([\x \[F]x]x & \[\x \[F]x]y)\ + apply (rule 0[THEN "\E"(1)]) by "cqt:2[lambda]" + ultimately AOT_have 1: \\[\x \[F]x]x\ + by (metis "&I" "raa-cor:3") + { + AOT_assume \\[F]x\ + AOT_hence \[\x \[F]x]x\ + by (auto intro!: "\\C"(1) "cqt:2") + AOT_hence \p & \p\ for p + using 1 by (metis "raa-cor:3") + } + AOT_thus \[F]x\ by (metis "raa-cor:1") + qed + AOT_hence \\\F ([F]x \ [F]y)\ + using "ind-nec"[THEN "\E"] by blast + AOT_hence \\F \([F]x \ [F]y)\ + by (metis "CBF" "\E") + } note indistI = this + { + AOT_assume G_prop: \x[G] & \y[G]\ + AOT_hence Ax: \A!x\ + using "&E"(1) "\I"(2) "\E" "encoders-are-abstract" by blast + + { + AOT_assume Ay: \A!y\ + { + fix F + { + AOT_assume \\u\([F]u \ [G]u)\ + AOT_hence \\\u([F]u \ [G]u)\ + using "Ordinary.res-var-bound-reas[BF]"[THEN "\E"] by simp + AOT_hence \\F \\<^sub>E G\ + by (AOT_subst \F \\<^sub>E G\ \\u ([F]u \ [G]u)\) + (auto intro!: "eqE"[THEN "\Df", THEN "\S"(1), OF "&I"] "cqt:2") + AOT_hence \x[F] \ x[G]\ + using 0[THEN "\E"(2)] "\E" "\E" by meson + AOT_hence \x[F]\ + using G_prop "&E" "\E" by blast + } + AOT_hence \\u\([F]u \ [G]u) \ x[F]\ + by (rule "\I") + } + AOT_hence xprop: \\F(\u\([F]u \ [G]u) \ x[F])\ + by (rule GEN) + moreover AOT_have yprop: \\\F(\u\([F]u \ [G]u) \ y[F])\ + proof (rule "raa-cor:2") + AOT_assume \\F(\u\([F]u \ [G]u) \ y[F])\ + AOT_hence \\F(\\u([F]u \ [G]u) \ y[F])\ + apply (AOT_subst \\\u([F]u \ [G]u)\ \\u\([F]u \ [G]u)\ for: F) + using "Ordinary.res-var-bound-reas[BF]" + "Ordinary.res-var-bound-reas[CBF]" + "intro-elim:2" apply presburger + by simp + AOT_hence A: \\F(\F \\<^sub>E G \ y[F])\ + by (AOT_subst \F \\<^sub>E G\ \\u ([F]u \ [G]u)\ for: F) + (auto intro!: "eqE"[THEN "\Df", THEN "\S"(1), OF "&I"] "cqt:2") + moreover AOT_have \\G \\<^sub>E G\ + by (auto intro!: "eqE"[THEN "\\<^sub>d\<^sub>fI"] "cqt:2" RN "&I" GEN "\I" "\I") + ultimately AOT_have \y[G]\ using "\E"(2) "\E" by blast + AOT_thus \p & \p\ for p using G_prop "&E" by (metis "raa-cor:3") + qed + AOT_have \\F([F]x & \[F]y)\ + proof(rule "raa-cor:1") + AOT_assume \\\F([F]x & \[F]y)\ + AOT_hence indist: \\F \([F]x \ [F]y)\ using indistI by blast + AOT_have \\F(\u\([F]u \ [G]u) \ y[F])\ + using indistinguishable_ord_enc_all[axiom_inst, THEN "\E", OF "&I", + OF "&I", OF "&I", OF "cqt:2[const_var]"[axiom_inst], + OF Ax, OF Ay, OF indist, THEN "\E"(1), OF xprop]. + AOT_thus \\F(\u\([F]u \ [G]u) \ y[F]) & \\F(\u\([F]u \ [G]u) \ y[F])\ + using yprop "&I" by blast + qed + } + moreover { + AOT_assume notAy: \\A!y\ + AOT_have \\F([F]x & \[F]y)\ + apply (rule "\I"(1)[where \=\\A!\\]) + using Ax notAy "&I" apply blast + by (simp add: "oa-exist:2") + } + ultimately AOT_have \\F([F]x & \[F]y)\ + by (metis "raa-cor:1") + } + moreover { + AOT_assume G_prop: \\x[G] & y[G]\ + AOT_hence Ay: \A!y\ + by (meson "&E"(2) "encoders-are-abstract" "existential:2[const_var]" "\E") + AOT_hence notOy: \\O!y\ + using "\E"(1) "oa-contingent:3" by blast + { + AOT_assume Ax: \A!x\ + { + fix F + { + AOT_assume \\\u([F]u \ [G]u)\ + AOT_hence \\F \\<^sub>E G\ + by (AOT_subst \F \\<^sub>E G\ \\u([F]u \ [G]u)\) + (auto intro!: "eqE"[THEN "\Df", THEN "\S"(1), OF "&I"] "cqt:2") + AOT_hence \x[F] \ x[G]\ + using 0[THEN "\E"(2)] "\E" "\E" by meson + AOT_hence \\x[F]\ + using G_prop "&E" "\E" by blast + } + AOT_hence \\\u([F]u \ [G]u) \ \x[F]\ + by (rule "\I") + } + AOT_hence x_prop: \\F(\\u([F]u \ [G]u) \ \x[F])\ + by (rule GEN) + AOT_have x_prop: \\\F(\u\([F]u \ [G]u) & x[F])\ + proof (rule "raa-cor:2") + AOT_assume \\F(\u \([F]u \ [G]u) & x[F])\ + then AOT_obtain F where F_prop: \\u \([F]u \ [G]u) & x[F]\ + using "\E"[rotated] by blast + AOT_have \\([F]u \ [G]u)\ for u + using F_prop[THEN "&E"(1), THEN "Ordinary.\E"]. + AOT_hence \\u \([F]u \ [G]u)\ + by (rule Ordinary.GEN) + AOT_hence \\\u([F]u \ [G]u)\ + by (metis "Ordinary.res-var-bound-reas[BF]" "\E") + AOT_hence \\x[F]\ + using x_prop[THEN "\E"(2), THEN "\E"] by blast + AOT_thus \p & \p\ for p + using F_prop[THEN "&E"(2)] by (metis "raa-cor:3") + qed + AOT_have y_prop: \\F(\u \([F]u \ [G]u) & y[F])\ + proof (rule "raa-cor:1") + AOT_assume \\\F (\u \([F]u \ [G]u) & y[F])\ + AOT_hence 0: \\F \(\u \([F]u \ [G]u) & y[F])\ + using "cqt-further:4"[THEN "\E"] by blast + { + fix F + { + AOT_assume \\u \([F]u \ [G]u)\ + AOT_hence \\y[F]\ + using 0[THEN "\E"(2)] "&I" "raa-cor:1" by meson + } + AOT_hence \(\u \([F]u \ [G]u) \ \y[F])\ + by (rule "\I") + } + AOT_hence A: \\F(\u \([F]u \ [G]u) \ \y[F])\ + by (rule GEN) + moreover AOT_have \\u \([G]u \ [G]u)\ + by (simp add: RN "oth-class-taut:3:a" "universal-cor" "\I") + ultimately AOT_have \\y[G]\ + using "\E"(2) "\E" by blast + AOT_thus \p & \p\ for p + using G_prop "&E" by (metis "raa-cor:3") + qed + AOT_have \\F([F]x & \[F]y)\ + proof(rule "raa-cor:1") + AOT_assume \\\F([F]x & \[F]y)\ + AOT_hence indist: \\F \([F]x \ [F]y)\ + using indistI by blast + AOT_thus \\F(\u \([F]u \ [G]u) & x[F]) & \\F(\u \([F]u \ [G]u) & x[F])\ + using indistinguishable_ord_enc_ex[axiom_inst, THEN "\E", OF "&I", + OF "&I", OF "&I", OF "cqt:2[const_var]"[axiom_inst], + OF Ax, OF Ay, OF indist, THEN "\E"(2), OF y_prop] + x_prop "&I" by blast + qed + } + moreover { + AOT_assume notAx: \\A!x\ + AOT_hence Ox: \O!x\ + using "\E"(3) "oa-exist:3" by blast + AOT_have \\F([F]x & \[F]y)\ + apply (rule "\I"(1)[where \=\\O!\\]) + using Ox notOy "&I" apply blast + by (simp add: "oa-exist:1") + } + ultimately AOT_have \\F([F]x & \[F]y)\ + by (metis "raa-cor:1") + } + ultimately AOT_show \\F([F]x & \[F]y)\ + using G_prop by (metis "&I" "\I" "\I" "raa-cor:1") + } + qed + + AOT_modally_strict { + fix x y + AOT_assume indist: \\F ([F]x \ [F]y)\ + AOT_hence nec_indist: \\\F ([F]x \ [F]y)\ + using "ind-nec" "vdash-properties:10" by blast + AOT_hence indist_nec: \\F \([F]x \ [F]y)\ + using "CBF" "vdash-properties:10" by blast + AOT_assume 0: \\G (\G \\<^sub>E F \ x[G])\ + AOT_hence 1: \\G (\\u ([G]u \ [F]u) \ x[G])\ + by (AOT_subst (reverse) \\u ([G]u \ [F]u)\ \G \\<^sub>E F\ for: G) + (auto intro!: "eqE"[THEN "\Df", THEN "\S"(1), OF "&I"] "cqt:2") + AOT_have \x[F]\ + by (safe intro!: 1[THEN "\E"(2), THEN "\E"] GEN "\I" RN "\I") + AOT_have \\G (\G \\<^sub>E F \ y[G])\ + proof(rule "raa-cor:1") + AOT_assume \\\G (\G \\<^sub>E F \ y[G])\ + AOT_hence \\G \(\G \\<^sub>E F \ y[G])\ + using "cqt-further:2" "\E" by blast + then AOT_obtain G where G_prop: \\(\G \\<^sub>E F \ y[G])\ + using "\E"[rotated] by blast + AOT_hence 1: \\G \\<^sub>E F & \y[G]\ + by (metis "\E"(1) "oth-class-taut:1:b") + AOT_have xG: \x[G]\ + using 0[THEN "\E"(2), THEN "\E"] 1[THEN "&E"(1)] by blast + AOT_hence \x[G] & \y[G]\ + using 1[THEN "&E"(2)] "&I" by blast + AOT_hence B: \\(x[G] \ y[G])\ + using "&E"(2) "\E"(1) "reductio-aa:1" xG by blast + { + fix H + { + AOT_assume \\H \\<^sub>E G\ + AOT_hence \\(H \\<^sub>E G & G \\<^sub>E F)\ + using 1 by (metis "KBasic:3" "con-dis-i-e:1" "con-dis-i-e:2:a" + "intro-elim:3:b") + moreover AOT_have \\(H \\<^sub>E G & G \\<^sub>E F) \ \(H \\<^sub>E F)\ + proof(rule RM) + AOT_modally_strict { + AOT_show \H \\<^sub>E G & G \\<^sub>E F \ H \\<^sub>E F\ + proof (safe intro!: "\I" "eqE"[THEN "\\<^sub>d\<^sub>fI"] "&I" "cqt:2" Ordinary.GEN) + fix u + AOT_assume \H \\<^sub>E G & G \\<^sub>E F\ + AOT_hence \\u ([H]u \ [G]u)\ and \\u ([G]u \ [F]u)\ + using "eqE"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast+ + AOT_thus \[H]u \ [F]u\ + by (auto dest!: "Ordinary.\E" dest: "\E") + qed + } + qed + ultimately AOT_have \\(H \\<^sub>E F)\ + using "\E" by blast + AOT_hence \x[H]\ + using 0[THEN "\E"(2)] "\E" by blast + AOT_hence \x[H] \ x[G]\ + using xG "\I" "\I" by blast + } + AOT_hence \\H \\<^sub>E G \ (x[H] \ x[G])\ by (rule "\I") + } + AOT_hence A: \\H(\H \\<^sub>E G \ (x[H] \ x[G]))\ + by (rule GEN) + then AOT_obtain F where F_prop: \[F]x & \[F]y\ + using Aux[OF A, OF B] "\E"[rotated] by blast + moreover AOT_have \[F]y\ + using indist[THEN "\E"(2), THEN "\E"(1), OF F_prop[THEN "&E"(1)]]. + AOT_thus \p & \p\ for p + using F_prop[THEN "&E"(2)] by (metis "raa-cor:3") + qed + } note 0 = this + AOT_modally_strict { + fix x y + AOT_assume \\F ([F]x \ [F]y)\ + moreover AOT_have \\F ([F]y \ [F]x)\ + by (metis calculation "cqt-basic:11" "\E"(2)) + ultimately AOT_have \\G (\G \\<^sub>E F \ x[G]) \ \G (\G \\<^sub>E F \ y[G])\ + using 0 "\I" "\I" by auto + } note 1 = this + AOT_show \[\x \G (\G \\<^sub>E F \ x[G])]\\ + by (safe intro!: RN GEN "\I" 1 "kirchner-thm:2"[THEN "\E"(2)]) + + AOT_modally_strict { + fix x y + AOT_assume indist: \\F ([F]x \ [F]y)\ + AOT_hence nec_indist: \\\F ([F]x \ [F]y)\ + using "ind-nec" "vdash-properties:10" by blast + AOT_hence indist_nec: \\F \([F]x \ [F]y)\ + using "CBF" "vdash-properties:10" by blast + AOT_assume 0: \\G (\G \\<^sub>E F \ \x[G])\ + AOT_hence 1: \\G (\\u ([G]u \ [F]u) \ \x[G])\ + by (AOT_subst (reverse) \\u ([G]u \ [F]u)\ \G \\<^sub>E F\ for: G) + (auto intro!: "eqE"[THEN "\Df", THEN "\S"(1), OF "&I"] "cqt:2") + AOT_have \\x[F]\ + by (safe intro!: 1[THEN "\E"(2), THEN "\E"] GEN "\I" RN "\I") + AOT_have \\G (\G \\<^sub>E F \ \y[G])\ + proof(rule "raa-cor:1") + AOT_assume \\\G (\G \\<^sub>E F \ \y[G])\ + AOT_hence \\G \(\G \\<^sub>E F \ \y[G])\ + using "cqt-further:2" "\E" by blast + then AOT_obtain G where G_prop: \\(\G \\<^sub>E F \ \y[G])\ + using "\E"[rotated] by blast + AOT_hence 1: \\G \\<^sub>E F & \\y[G]\ + by (metis "\E"(1) "oth-class-taut:1:b") + AOT_hence yG: \y[G]\ + using G_prop "\I" "raa-cor:3" by blast + moreover AOT_hence 12: \\x[G]\ + using 0[THEN "\E"(2), THEN "\E"] 1[THEN "&E"(1)] by blast + ultimately AOT_have \\x[G] & y[G]\ + using "&I" by blast + AOT_hence B: \\(x[G] \ y[G])\ + by (metis "12" "\E"(3) "raa-cor:3" yG) + { + fix H + { + AOT_assume 3: \\H \\<^sub>E G\ + AOT_hence \\(H \\<^sub>E G & G \\<^sub>E F)\ + using 1 + by (metis "KBasic:3" "con-dis-i-e:1" "\I" "intro-elim:3:b" + "reductio-aa:1" G_prop) + moreover AOT_have \\(H \\<^sub>E G & G \\<^sub>E F) \ \(H \\<^sub>E F)\ + proof (rule RM) + AOT_modally_strict { + AOT_show \H \\<^sub>E G & G \\<^sub>E F \ H \\<^sub>E F\ + proof (safe intro!: "\I" "eqE"[THEN "\\<^sub>d\<^sub>fI"] "&I" "cqt:2" Ordinary.GEN) + fix u + AOT_assume \H \\<^sub>E G & G \\<^sub>E F\ + AOT_hence \\u ([H]u \ [G]u)\ and \\u ([G]u \ [F]u)\ + using "eqE"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast+ + AOT_thus \[H]u \ [F]u\ + by (auto dest!: "Ordinary.\E" dest: "\E") + qed + } + qed + ultimately AOT_have \\(H \\<^sub>E F)\ + using "\E" by blast + AOT_hence \\x[H]\ + using 0[THEN "\E"(2)] "\E" by blast + AOT_hence \x[H] \ x[G]\ + using 12 "\I" "\I" by (metis "raa-cor:3") + } + AOT_hence \\H \\<^sub>E G \ (x[H] \ x[G])\ + by (rule "\I") + } + AOT_hence A: \\H(\H \\<^sub>E G \ (x[H] \ x[G]))\ + by (rule GEN) + then AOT_obtain F where F_prop: \[F]x & \[F]y\ + using Aux[OF A, OF B] "\E"[rotated] by blast + moreover AOT_have \[F]y\ + using indist[THEN "\E"(2), THEN "\E"(1), OF F_prop[THEN "&E"(1)]]. + AOT_thus \p & \p\ for p + using F_prop[THEN "&E"(2)] by (metis "raa-cor:3") + qed + } note 0 = this + AOT_modally_strict { + fix x y + AOT_assume \\F ([F]x \ [F]y)\ + moreover AOT_have \\F ([F]y \ [F]x)\ + by (metis calculation "cqt-basic:11" "\E"(2)) + ultimately AOT_have \\G (\G \\<^sub>E F \ \x[G]) \ \G (\G \\<^sub>E F \ \y[G])\ + using 0 "\I" "\I" by auto + } note 1 = this + AOT_show \[\x \G (\G \\<^sub>E F \ \x[G])]\\ + by (safe intro!: RN GEN "\I" 1 "kirchner-thm:2"[THEN "\E"(2)]) +qed + +text\Reformulate the existence claims in terms of their negations.\ + +AOT_theorem denotes_ex: \[\x \G (\G \\<^sub>E F & x[G])]\\ +proof (rule "safe-ext"[axiom_inst, THEN "\E", OF "&I"]) + AOT_show \[\x \\G (\G \\<^sub>E F \ \x[G])]\\ + using denotes_all_neg[THEN negation_denotes[THEN "\E"]]. +next + AOT_show \\\x (\\G (\G \\<^sub>E F \ \x[G]) \ \G (\G \\<^sub>E F & x[G]))\ + by (AOT_subst \\G \\<^sub>E F & x[G]\ \\(\G \\<^sub>E F \ \x[G])\ for: G x) + (auto simp: "conventions:1" "rule-eq-df:1" + intro: "oth-class-taut:4:b"[THEN "\E"(2)] + "intro-elim:3:f"[OF "cqt-further:3", OF "oth-class-taut:3:b"] + intro!: RN GEN) +qed + +AOT_theorem denotes_ex_neg: \[\x \G (\G \\<^sub>E F & \x[G])]\\ +proof (rule "safe-ext"[axiom_inst, THEN "\E", OF "&I"]) + AOT_show \[\x \\G (\G \\<^sub>E F \ x[G])]\\ + using denotes_all[THEN negation_denotes[THEN "\E"]]. +next + AOT_show \\\x (\\G (\G \\<^sub>E F \ x[G]) \ \G (\G \\<^sub>E F & \x[G]))\ + by (AOT_subst (reverse) \\G \\<^sub>E F & \x[G]\ \\(\G \\<^sub>E F \ x[G])\ for: G x) + (auto simp: "oth-class-taut:1:b" + intro: "oth-class-taut:4:b"[THEN "\E"(2)] + "intro-elim:3:f"[OF "cqt-further:3", OF "oth-class-taut:3:b"] + intro!: RN GEN) +qed + +text\Derive comprehension principles.\ + +AOT_theorem Comprehension_1: + shows \\\F\G(\G \\<^sub>E F \ (\{F} \ \{G})) \ [\x \F (\{F} & x[F])]\\ +proof(rule "\I") + AOT_assume assm: \\\F\G(\G \\<^sub>E F \ (\{F} \ \{G}))\ + AOT_modally_strict { + fix x y + AOT_assume 0: \\F\G (\G \\<^sub>E F \ (\{F} \ \{G}))\ + AOT_assume indist: \\F ([F]x \ [F]y)\ + AOT_assume x_prop: \\F (\{F} & x[F])\ + then AOT_obtain F where F_prop: \\{F} & x[F]\ + using "\E"[rotated] by blast + AOT_hence \\F \\<^sub>E F & x[F]\ + by (auto intro!: RN eqE[THEN "\\<^sub>d\<^sub>fI"] "&I" "cqt:2" GEN "\I" "\I" dest: "&E") + AOT_hence \\G(\G \\<^sub>E F & x[G])\ + by (rule "\I") + AOT_hence \[\x \G(\G \\<^sub>E F & x[G])]x\ + by (safe intro!: "\\C" denotes_ex "cqt:2") + AOT_hence \[\x \G(\G \\<^sub>E F & x[G])]y\ + using indist[THEN "\E"(1), OF denotes_ex, THEN "\E"(1)] by blast + AOT_hence \\G(\G \\<^sub>E F & y[G])\ + using "\\C" by blast + then AOT_obtain G where \\G \\<^sub>E F & y[G]\ + using "\E"[rotated] by blast + AOT_hence \\{G} & y[G]\ + using 0[THEN "\E"(2), THEN "\E"(2), THEN "\E", THEN "\E"(1)] + F_prop[THEN "&E"(1)] "&E" "&I" by blast + AOT_hence \\F (\{F} & y[F])\ + by (rule "\I") + } note 1 = this + AOT_modally_strict { + AOT_assume 0: \\F\G (\G \\<^sub>E F \ (\{F} \ \{G}))\ + { + fix x y + { + AOT_assume \\F ([F]x \ [F]y)\ + moreover AOT_have \\F ([F]y \ [F]x)\ + by (metis calculation "cqt-basic:11" "\E"(1)) + ultimately AOT_have \\F (\{F} & x[F]) \ \F (\{F} & y[F])\ + using 0 1[OF 0] "\I" "\I" by simp + } + AOT_hence \\F ([F]x \ [F]y) \ (\F (\{F} & x[F]) \ \F (\{F} & y[F]))\ + using "\I" by blast + } + AOT_hence \\x\y(\F ([F]x \ [F]y) \ (\F (\{F} & x[F]) \ \F (\{F} & y[F])))\ + by (auto intro!: GEN) + } note 1 = this + AOT_hence \\<^bold>\\<^sub>\ \F\G (\G \\<^sub>E F \ (\{F} \ \{G})) \ + \x\y(\F ([F]x \ [F]y) \ (\F (\{F} & x[F]) \ \F (\{F} & y[F])))\ + by (rule "\I") + AOT_hence \\\F\G (\G \\<^sub>E F \ (\{F} \ \{G})) \ + \\x\y(\F ([F]x \ [F]y) \ (\F (\{F} & x[F]) \ \F (\{F} & y[F])))\ + by (rule RM) + AOT_hence \\\x\y(\F ([F]x \ [F]y) \ (\F (\{F} & x[F]) \ \F (\{F} & y[F])))\ + using "\E" assm by blast + AOT_thus \[\x \F (\{F} & x[F])]\\ + by (safe intro!: "kirchner-thm:2"[THEN "\E"(2)]) +qed + +AOT_theorem Comprehension_2: + shows \\\F\G(\G \\<^sub>E F \ (\{F} \ \{G})) \ [\x \F (\{F} & \x[F])]\\ +proof(rule "\I") + AOT_assume assm: \\\F\G(\G \\<^sub>E F \ (\{F} \ \{G}))\ + AOT_modally_strict { + fix x y + AOT_assume 0: \\F\G (\G \\<^sub>E F \ (\{F} \ \{G}))\ + AOT_assume indist: \\F ([F]x \ [F]y)\ + AOT_assume x_prop: \\F (\{F} & \x[F])\ + then AOT_obtain F where F_prop: \\{F} & \x[F]\ + using "\E"[rotated] by blast + AOT_hence \\F \\<^sub>E F & \x[F]\ + by (auto intro!: RN eqE[THEN "\\<^sub>d\<^sub>fI"] "&I" "cqt:2" GEN "\I" "\I" dest: "&E") + AOT_hence \\G(\G \\<^sub>E F & \x[G])\ + by (rule "\I") + AOT_hence \[\x \G(\G \\<^sub>E F & \x[G])]x\ + by (safe intro!: "\\C" denotes_ex_neg "cqt:2") + AOT_hence \[\x \G(\G \\<^sub>E F & \x[G])]y\ + using indist[THEN "\E"(1), OF denotes_ex_neg, THEN "\E"(1)] by blast + AOT_hence \\G(\G \\<^sub>E F & \y[G])\ + using "\\C" by blast + then AOT_obtain G where \\G \\<^sub>E F & \y[G]\ + using "\E"[rotated] by blast + AOT_hence \\{G} & \y[G]\ + using 0[THEN "\E"(2), THEN "\E"(2), THEN "\E", THEN "\E"(1)] + F_prop[THEN "&E"(1)] "&E" "&I" by blast + AOT_hence \\F (\{F} & \y[F])\ + by (rule "\I") + } note 1 = this + AOT_modally_strict { + AOT_assume 0: \\F\G (\G \\<^sub>E F \ (\{F} \ \{G}))\ + { + fix x y + { + AOT_assume \\F ([F]x \ [F]y)\ + moreover AOT_have \\F ([F]y \ [F]x)\ + by (metis calculation "cqt-basic:11" "\E"(1)) + ultimately AOT_have \\F (\{F} & \x[F]) \ \F (\{F} & \y[F])\ + using 0 1[OF 0] "\I" "\I" by simp + } + AOT_hence \\F ([F]x \ [F]y) \ (\F (\{F} & \x[F]) \ \F (\{F} & \y[F]))\ + using "\I" by blast + } + AOT_hence \\x\y(\F ([F]x \ [F]y) \ (\F (\{F} & \x[F]) \ \F (\{F} & \y[F])))\ + by (auto intro!: GEN) + } note 1 = this + AOT_hence \\<^bold>\\<^sub>\ \F\G (\G \\<^sub>E F \ (\{F} \ \{G})) \ + \x\y(\F ([F]x \ [F]y) \ (\F (\{F} & \x[F]) \ \F (\{F} & \y[F])))\ + by (rule "\I") + AOT_hence \\\F\G (\G \\<^sub>E F \ (\{F} \ \{G})) \ + \\x\y(\F ([F]x \ [F]y) \ (\F (\{F} & \x[F]) \ \F (\{F} & \y[F])))\ + by (rule RM) + AOT_hence \\\x\y(\F ([F]x \ [F]y) \ (\F (\{F} & \x[F]) \ \F (\{F} & \y[F])))\ + using "\E" assm by blast + AOT_thus \[\x \F (\{F} & \x[F])]\\ + by (safe intro!: "kirchner-thm:2"[THEN "\E"(2)]) +qed + +text\Derived variants of the comprehension principles above.\ + +AOT_theorem Comprehension_1': + shows \\\F\G(\G \\<^sub>E F \ (\{F} \ \{G})) \ [\x \F (x[F] \ \{F})]\\ +proof(rule "\I") + AOT_assume \\\F\G(\G \\<^sub>E F \ (\{F} \ \{G}))\ + AOT_hence 0: \\\F\G(\G \\<^sub>E F \ (\\{F} \ \\{G}))\ + by (AOT_subst (reverse) \\\{F} \ \\{G}\ \\{F} \ \{G}\ for: F G) + (auto simp: "oth-class-taut:4:b") + AOT_show \[\x \F (x[F] \ \{F})]\\ + proof(rule "safe-ext"[axiom_inst, THEN "\E", OF "&I"]) + AOT_show \[\x \\F(\\{F} & x[F])]\\ + using Comprehension_1[THEN "\E", OF 0, THEN negation_denotes[THEN "\E"]]. + next + AOT_show \\\x (\\F (\\{F} & x[F]) \ \F (x[F] \ \{F}))\ + by (AOT_subst (reverse) \\\{F} & x[F]\ \\(x[F] \ \{F})\ for: F x) + (auto simp: "oth-class-taut:1:b"[THEN "intro-elim:3:e", + OF "oth-class-taut:2:a"] + intro: "intro-elim:3:f"[OF "cqt-further:3", OF "oth-class-taut:3:a", + symmetric] + intro!: RN GEN) + qed +qed + +AOT_theorem Comprehension_2': + shows \\\F\G(\G \\<^sub>E F \ (\{F} \ \{G})) \ [\x \F (\{F} \ x[F])]\\ +proof(rule "\I") + AOT_assume 0: \\\F\G(\G \\<^sub>E F \ (\{F} \ \{G}))\ + AOT_show \[\x \F (\{F} \ x[F])]\\ + proof(rule "safe-ext"[axiom_inst, THEN "\E", OF "&I"]) + AOT_show \[\x \\F(\{F} & \x[F])]\\ + using Comprehension_2[THEN "\E", OF 0, THEN negation_denotes[THEN "\E"]]. + next + AOT_show \\\x (\\F (\{F} & \x[F]) \ \F (\{F} \ x[F]))\ + by (AOT_subst (reverse) \\{F} & \x[F]\ \\(\{F} \ x[F])\ for: F x) + (auto simp: "oth-class-taut:1:b" + intro: "intro-elim:3:f"[OF "cqt-further:3", OF "oth-class-taut:3:a", + symmetric] + intro!: RN GEN) + qed +qed + +text\Derive a combined comprehension principles.\ + +AOT_theorem Comprehension_3: + \\\F\G(\G \\<^sub>E F \ (\{F} \ \{G})) \ [\x \F (x[F] \ \{F})]\\ +proof(rule "\I") + AOT_assume 0: \\\F\G(\G \\<^sub>E F \ (\{F} \ \{G}))\ + AOT_show \[\x \F (x[F] \ \{F})]\\ + proof(rule "safe-ext"[axiom_inst, THEN "\E", OF "&I"]) + AOT_show \[\x \F (x[F] \ \{F}) & \F (\{F} \ x[F])]\\ + by (safe intro!: conjunction_denotes[THEN "\E", OF "&I"] + Comprehension_1'[THEN "\E"] + Comprehension_2'[THEN "\E"] 0) + next + AOT_show \\\x (\F (x[F] \ \{F}) & \F (\{F} \ x[F]) \ \F (x[F] \ \{F}))\ + by (auto intro!: RN GEN "\I" "\I" "&I" dest: "&E" "\E"(2) "\E" "\E"(1,2)) + qed +qed + +notepad +begin +text\Verify that the original axioms are equivalent to @{thm denotes_ex} + and @{thm denotes_ex_neg}.\ +AOT_modally_strict { + fix x y H + AOT_have \A!x & A!y & \F \([F]x \ [F]y) \ + (\G (\z (O!z \ \([G]z \ [H]z)) \ x[G]) \ + \G (\z (O!z \ \([G]z \ [H]z)) \ y[G]))\ + proof(rule "\I") + { + fix x y + AOT_assume \A!x\ + AOT_assume \A!y\ + AOT_assume indist: \\F \([F]x \ [F]y)\ + AOT_assume \\G (\u \([G]u \ [H]u) \ x[G])\ + AOT_hence \\G (\\u ([G]u \ [H]u) \ x[G])\ + using "Ordinary.res-var-bound-reas[BF]" "Ordinary.res-var-bound-reas[CBF]" + "intro-elim:2" + by (AOT_subst \\\u ([G]u \ [H]u)\ \\u \([G]u \ [H]u)\ for: G) auto + AOT_hence \\G (\G \\<^sub>E H \ x[G])\ + by (AOT_subst \G \\<^sub>E H\ \\u ([G]u \ [H]u)\ for: G) + (safe intro!: "eqE"[THEN "\Df", THEN "\S"(1), OF "&I"] "cqt:2") + AOT_hence \\\G (\G \\<^sub>E H & \x[G])\ + by (AOT_subst (reverse) \(\G \\<^sub>E H & \x[G])\ \\(\G \\<^sub>E H \ x[G])\ for: G) + (auto simp: "oth-class-taut:1:b" "cqt-further:3"[THEN "\E"(1)]) + AOT_hence \\[\x \G (\G \\<^sub>E H & \x[G])]x\ + by (auto intro: "\\C") + AOT_hence \\[\x \G (\G \\<^sub>E H & \x[G])]y\ + using indist[THEN "\E"(1), OF denotes_ex_neg, + THEN "qml:2"[axiom_inst, THEN "\E"], + THEN "\E"(3)] by blast + AOT_hence \\\G (\G \\<^sub>E H & \y[G])\ + by (safe intro!: "\\C" denotes_ex_neg "cqt:2") + AOT_hence \\G \(\G \\<^sub>E H & \y[G])\ + using "cqt-further:4"[THEN "\E"] by blast + AOT_hence \\G(\G \\<^sub>E H \ y[G])\ + by (AOT_subst \\G \\<^sub>E H \ y[G]\ \\(\G \\<^sub>E H & \y[G])\ for: G) + (auto simp: "oth-class-taut:1:a") + AOT_hence \\G (\\u([G]u \ [H]u) \ y[G])\ + by (AOT_subst (reverse) \\u ([G]u \ [H]u)\ \G \\<^sub>E H\ for: G) + (safe intro!: "eqE"[THEN "\Df", THEN "\S"(1), OF "&I"] "cqt:2") + AOT_hence \\G (\u \([G]u \ [H]u) \ y[G])\ + using "Ordinary.res-var-bound-reas[BF]" "Ordinary.res-var-bound-reas[CBF]" + "intro-elim:2" + by (AOT_subst \\u \([G]u \ [H]u)\ \\\u ([G]u \ [H]u)\ for: G) auto + } note 0 = this + AOT_assume \A!x & A!y & \F \([F]x \ [F]y)\ + AOT_hence \A!x\ and \A!y\ and \\F \([F]x \ [F]y)\ + using "&E" by blast+ + moreover AOT_have \\F \([F]y \ [F]x)\ + using calculation(3) + apply (safe intro!: CBF[THEN "\E"] dest!: BF[THEN "\E"]) + using "RM:3" "cqt-basic:11" "intro-elim:3:b" by fast + ultimately AOT_show \\G (\u \([G]u \ [H]u) \ x[G]) \ + \G (\u \([G]u \ [H]u) \ y[G])\ + using 0 by (auto intro!: "\I" "\I") + qed + + AOT_have \A!x & A!y & \F \([F]x \ [F]y) \ + (\G (\z (O!z \ \([G]z \ [H]z)) & x[G]) \ \G (\z (O!z \ \([G]z \ [H]z)) & y[G]))\ + proof(rule "\I") + { + fix x y + AOT_assume \A!x\ + AOT_assume \A!y\ + AOT_assume indist: \\F \([F]x \ [F]y)\ + AOT_assume x_prop: \\G (\u \([G]u \ [H]u) & x[G])\ + AOT_hence \\G (\\u ([G]u \ [H]u) & x[G])\ + using "Ordinary.res-var-bound-reas[BF]" "Ordinary.res-var-bound-reas[CBF]" + "intro-elim:2" + by (AOT_subst \\\u ([G]u \ [H]u)\ \\u \([G]u \ [H]u)\ for: G) auto + AOT_hence \\G (\G \\<^sub>E H & x[G])\ + by (AOT_subst \G \\<^sub>E H\ \\u ([G]u \ [H]u)\ for: G) + (safe intro!: "eqE"[THEN "\Df", THEN "\S"(1), OF "&I"] "cqt:2") + AOT_hence \[\x \G (\G \\<^sub>E H & x[G])]x\ + by (safe intro!: "\\C" denotes_ex "cqt:2") + AOT_hence \[\x \G (\G \\<^sub>E H & x[G])]y\ + using indist[THEN "\E"(1), OF denotes_ex, + THEN "qml:2"[axiom_inst, THEN "\E"], + THEN "\E"(1)] by blast + AOT_hence \\G (\G \\<^sub>E H & y[G])\ + by (rule "\\C") + AOT_hence \\G (\\u ([G]u \ [H]u) & y[G])\ + by (AOT_subst (reverse) \\u ([G]u \ [H]u)\ \G \\<^sub>E H\ for: G) + (safe intro!: "eqE"[THEN "\Df", THEN "\S"(1), OF "&I"] "cqt:2") + AOT_hence \\G (\u \([G]u \ [H]u) & y[G])\ + using "Ordinary.res-var-bound-reas[BF]" + "Ordinary.res-var-bound-reas[CBF]" + "intro-elim:2" + by (AOT_subst \\u \([G]u \ [H]u)\ \\\u ([G]u \ [H]u)\ for: G) auto + } note 0 = this + AOT_assume \A!x & A!y & \F \([F]x \ [F]y)\ + AOT_hence \A!x\ and \A!y\ and \\F \([F]x \ [F]y)\ + using "&E" by blast+ + moreover AOT_have \\F \([F]y \ [F]x)\ + using calculation(3) + apply (safe intro!: CBF[THEN "\E"] dest!: BF[THEN "\E"]) + using "RM:3" "cqt-basic:11" "intro-elim:3:b" by fast + ultimately AOT_show \\G (\u \([G]u \ [H]u) & x[G]) \ + \G (\u \([G]u \ [H]u) & y[G])\ + using 0 by (auto intro!: "\I" "\I") + qed +} +end +end \ No newline at end of file diff --git a/thys/AOT/AOT_NaturalNumbers.thy b/thys/AOT/AOT_NaturalNumbers.thy new file mode 100644 --- /dev/null +++ b/thys/AOT/AOT_NaturalNumbers.thy @@ -0,0 +1,5508 @@ +(*<*) +theory AOT_NaturalNumbers + imports AOT_PossibleWorlds AOT_ExtendedRelationComprehension + abbrevs one-to-one = \\<^sub>1\<^sub>-\<^sub>1\ + and onto = \\<^sub>o\<^sub>n\<^sub>t\<^sub>o\ +begin +(*>*) + +section\Natural Numbers\ + +AOT_define CorrelatesOneToOne :: \\ \ \ \ \ \ \\ (\_ |: _ \<^sub>1\<^sub>-\<^sub>1\ _\) + "1-1-cor": \R |: F \<^sub>1\<^sub>-\<^sub>1\ G \\<^sub>d\<^sub>f R\ & F\ & G\ & + \x ([F]x \ \!y([G]y & [R]xy)) & + \y ([G]y \ \!x([F]x & [R]xy))\ + +AOT_define MapsTo :: \\ \ \ \ \ \ \\ (\_ |: _ \ _\) + "fFG:1": \R |: F \ G \\<^sub>d\<^sub>f R\ & F\ & G\ & \x ([F]x \ \!y([G]y & [R]xy))\ + +AOT_define MapsToOneToOne :: \\ \ \ \ \ \ \\ (\_ |: _ \<^sub>1\<^sub>-\<^sub>1\ _\) + "fFG:2": \R |: F \<^sub>1\<^sub>-\<^sub>1\ G \\<^sub>d\<^sub>f + R |: F \ G & \x\y\z (([F]x & [F]y & [G]z) \ ([R]xz & [R]yz \ x = y))\ + +AOT_define MapsOnto :: \\ \ \ \ \ \ \\ (\_ |: _ \\<^sub>o\<^sub>n\<^sub>t\<^sub>o _\) + "fFG:3": \R |: F \\<^sub>o\<^sub>n\<^sub>t\<^sub>o G \\<^sub>d\<^sub>f R |: F \ G & \y ([G]y \ \x([F]x & [R]xy))\ + +AOT_define MapsOneToOneOnto :: \\ \ \ \ \ \ \\ (\_ |: _ \<^sub>1\<^sub>-\<^sub>1\\<^sub>o\<^sub>n\<^sub>t\<^sub>o _\) + "fFG:4": \R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>o\<^sub>n\<^sub>t\<^sub>o G \\<^sub>d\<^sub>f R |: F \<^sub>1\<^sub>-\<^sub>1\ G & R |: F \\<^sub>o\<^sub>n\<^sub>t\<^sub>o G\ + +AOT_theorem "eq-1-1": \R |: F \<^sub>1\<^sub>-\<^sub>1\ G \ R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>o\<^sub>n\<^sub>t\<^sub>o G\ +proof(rule "\I"; rule "\I") + AOT_assume \R |: F \<^sub>1\<^sub>-\<^sub>1\ G\ + AOT_hence A: \\x ([F]x \ \!y([G]y & [R]xy))\ + and B: \\y ([G]y \ \!x([F]x & [R]xy))\ + using "\\<^sub>d\<^sub>fE"[OF "1-1-cor"] "&E" by blast+ + AOT_have C: \R |: F \ G\ + proof (rule "\\<^sub>d\<^sub>fI"[OF "fFG:1"]; rule "&I") + AOT_show \R\ & F\ & G\\ + using "cqt:2[const_var]"[axiom_inst] "&I" by metis + next + AOT_show \\x ([F]x \ \!y([G]y & [R]xy))\ by (rule A) + qed + AOT_show \R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>o\<^sub>n\<^sub>t\<^sub>o G\ + proof (rule "\\<^sub>d\<^sub>fI"[OF "fFG:4"]; rule "&I") + AOT_show \R |: F \<^sub>1\<^sub>-\<^sub>1\ G\ + proof (rule "\\<^sub>d\<^sub>fI"[OF "fFG:2"]; rule "&I") + AOT_show \R |: F \ G\ using C. + next + AOT_show \\x\y\z ([F]x & [F]y & [G]z \ ([R]xz & [R]yz \ x = y))\ + proof(rule GEN; rule GEN; rule GEN; rule "\I"; rule "\I") + fix x y z + AOT_assume 1: \[F]x & [F]y & [G]z\ + moreover AOT_assume 2: \[R]xz & [R]yz\ + ultimately AOT_have 3: \\!x ([F]x & [R]xz)\ + using B "&E" "\E" "\E" by fast + AOT_show \x = y\ + by (rule "uni-most"[THEN "\E", OF 3, THEN "\E"(2)[where \=x], + THEN "\E"(2)[where \=y], THEN "\E"]) + (metis "&I" "&E" 1 2) + qed + qed + next + AOT_show \R |: F \\<^sub>o\<^sub>n\<^sub>t\<^sub>o G\ + proof (rule "\\<^sub>d\<^sub>fI"[OF "fFG:3"]; rule "&I") + AOT_show \R |: F \ G\ using C. + next + AOT_show \\y ([G]y \ \x ([F]x & [R]xy))\ + proof(rule GEN; rule "\I") + fix y + AOT_assume \[G]y\ + AOT_hence \\!x ([F]x & [R]xy)\ + using B[THEN "\E"(2), THEN "\E"] by blast + AOT_hence \\x ([F]x & [R]xy & \\ (([F]\ & [R]\y) \ \ = x))\ + using "uniqueness:1"[THEN "\\<^sub>d\<^sub>fE"] by blast + then AOT_obtain x where \[F]x & [R]xy\ + using "\E"[rotated] "&E" by blast + AOT_thus \\x ([F]x & [R]xy)\ by (rule "\I") + qed + qed + qed +next + AOT_assume \R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>o\<^sub>n\<^sub>t\<^sub>o G\ + AOT_hence \R |: F \<^sub>1\<^sub>-\<^sub>1\ G\ and \R |: F \\<^sub>o\<^sub>n\<^sub>t\<^sub>o G\ + using "\\<^sub>d\<^sub>fE"[OF "fFG:4"] "&E" by blast+ + AOT_hence C: \R |: F \ G\ + and D: \\x\y\z ([F]x & [F]y & [G]z \ ([R]xz & [R]yz \ x = y))\ + and E: \\y ([G]y \ \x ([F]x & [R]xy))\ + using "\\<^sub>d\<^sub>fE"[OF "fFG:2"] "\\<^sub>d\<^sub>fE"[OF "fFG:3"] "&E" by blast+ + AOT_show \R |: F \<^sub>1\<^sub>-\<^sub>1\ G\ + proof(rule "1-1-cor"[THEN "\\<^sub>d\<^sub>fI"]; safe intro!: "&I" "cqt:2[const_var]"[axiom_inst]) + AOT_show \\x ([F]x \ \!y ([G]y & [R]xy))\ + using "\\<^sub>d\<^sub>fE"[OF "fFG:1", OF C] "&E" by blast + next + AOT_show \\y ([G]y \ \!x ([F]x & [R]xy))\ + proof (rule "GEN"; rule "\I") + fix y + AOT_assume 0: \[G]y\ + AOT_hence \\x ([F]x & [R]xy)\ + using E "\E" "\E" by fast + then AOT_obtain a where a_prop: \[F]a & [R]ay\ + using "\E"[rotated] by blast + moreover AOT_have \\z ([F]z & [R]zy \ z = a)\ + proof (rule GEN; rule "\I") + fix z + AOT_assume \[F]z & [R]zy\ + AOT_thus \z = a\ + using D[THEN "\E"(2)[where \=z], THEN "\E"(2)[where \=a], + THEN "\E"(2)[where \=y], THEN "\E", THEN "\E"] + a_prop 0 "&E" "&I" by metis + qed + ultimately AOT_have \\x ([F]x & [R]xy & \z ([F]z & [R]zy \ z = x))\ + using "&I" "\I"(2) by fast + AOT_thus \\!x ([F]x & [R]xy)\ + using "uniqueness:1"[THEN "\\<^sub>d\<^sub>fI"] by fast + qed + qed +qed + +text\We have already introduced the restricted type of Ordinary objects in the + Extended Relation Comprehension theory. However, make sure all variable names + are defined as expected (avoiding conflicts with situations + of possible world theory).\ +AOT_register_variable_names + Ordinary: u v r t s + +AOT_theorem "equi:1": \\!u \{u} \ \u (\{u} & \v (\{v} \ v =\<^sub>E u))\ +proof(rule "\I"; rule "\I") + AOT_assume \\!u \{u}\ + AOT_hence \\!x (O!x & \{x})\. + AOT_hence \\x (O!x & \{x} & \\ (O!\ & \{\} \ \ = x))\ + using "uniqueness:1"[THEN "\\<^sub>d\<^sub>fE"] by blast + then AOT_obtain x where x_prop: \O!x & \{x} & \\ (O!\ & \{\} \ \ = x)\ + using "\E"[rotated] by blast + { + fix \ + AOT_assume beta_ord: \O!\\ + moreover AOT_assume \\{\}\ + ultimately AOT_have \\ = x\ + using x_prop[THEN "&E"(2), THEN "\E"(2)[where \=\]] "&I" "\E" by blast + AOT_hence \\ =\<^sub>E x\ + using "ord-=E=:1"[THEN "\E", OF "\I"(1)[OF beta_ord], + THEN "qml:2"[axiom_inst, THEN "\E"], + THEN "\E"(1)] + by blast + } + AOT_hence \(O!\ \ (\{\} \ \ =\<^sub>E x))\ for \ + using "\I" by blast + AOT_hence \\\(O!\ \ (\{\} \ \ =\<^sub>E x))\ + by (rule GEN) + AOT_hence \O!x & \{x} & \y (O!y \ (\{y} \ y =\<^sub>E x))\ + using x_prop[THEN "&E"(1)] "&I" by blast + AOT_hence \O!x & (\{x} & \y (O!y \ (\{y} \ y =\<^sub>E x)))\ + using "&E" "&I" by meson + AOT_thus \\u (\{u} & \v (\{v} \ v =\<^sub>E u))\ + using "\I" by fast +next + AOT_assume \\u (\{u} & \v (\{v} \ v =\<^sub>E u))\ + AOT_hence \\x (O!x & (\{x} & \y (O!y \ (\{y} \ y =\<^sub>E x))))\ + by blast + then AOT_obtain x where x_prop: \O!x & (\{x} & \y (O!y \ (\{y} \ y =\<^sub>E x)))\ + using "\E"[rotated] by blast + AOT_have \\y ([O!]y & \{y} \ y = x)\ + proof(rule GEN; rule "\I") + fix y + AOT_assume \O!y & \{y}\ + AOT_hence \y =\<^sub>E x\ + using x_prop[THEN "&E"(2), THEN "&E"(2), THEN "\E"(2)[where \=y]] + "\E" "&E" by blast + AOT_thus \y = x\ + using "ord-=E=:1"[THEN "\E", OF "\I"(2)[OF x_prop[THEN "&E"(1)]], + THEN "qml:2"[axiom_inst, THEN "\E"], THEN "\E"(2)] by blast + qed + AOT_hence \[O!]x & \{x} & \y ([O!]y & \{y} \ y = x)\ + using x_prop "&E" "&I" by meson + AOT_hence \\x ([O!]x & \{x} & \y ([O!]y & \{y} \ y = x))\ + by (rule "\I") + AOT_hence \\!x (O!x & \{x})\ + by (rule "uniqueness:1"[THEN "\\<^sub>d\<^sub>fI"]) + AOT_thus \\!u \{u}\. +qed + +AOT_define CorrelatesEOneToOne :: \\ \ \ \ \ \ \\ (\_ |: _ \<^sub>1\<^sub>-\<^sub>1\\<^sub>E _\) + "equi:2": \R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E G \\<^sub>d\<^sub>f R\ & F\ & G\ & + \u ([F]u \ \!v([G]v & [R]uv)) & + \v ([G]v \ \!u([F]u & [R]uv))\ + +AOT_define EquinumerousE :: \\ \ \ \ \\ (infixl "\\<^sub>E" 50) + "equi:3": \F \\<^sub>E G \\<^sub>d\<^sub>f \R (R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E G)\ + +text\Note: not explicitly in PLM.\ +AOT_theorem eq_den_1: \\\\ if \\ \\<^sub>E \'\ +proof - + AOT_have \\R (R |: \ \<^sub>1\<^sub>-\<^sub>1\\<^sub>E \')\ + using "equi:3"[THEN "\\<^sub>d\<^sub>fE"] that by blast + then AOT_obtain R where \R |: \ \<^sub>1\<^sub>-\<^sub>1\\<^sub>E \'\ + using "\E"[rotated] by blast + AOT_thus \\\\ + using "equi:2"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast +qed + +text\Note: not explicitly in PLM.\ +AOT_theorem eq_den_2: \\'\\ if \\ \\<^sub>E \'\ +proof - + AOT_have \\R (R |: \ \<^sub>1\<^sub>-\<^sub>1\\<^sub>E \')\ + using "equi:3"[THEN "\\<^sub>d\<^sub>fE"] that by blast + then AOT_obtain R where \R |: \ \<^sub>1\<^sub>-\<^sub>1\\<^sub>E \'\ + using "\E"[rotated] by blast + AOT_thus \\'\\ + using "equi:2"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast+ +qed + +AOT_theorem "eq-part:1": \F \\<^sub>E F\ +proof (safe intro!: "&I" GEN "\I" "cqt:2[const_var]"[axiom_inst] + "\\<^sub>d\<^sub>fI"[OF "equi:3"] "\\<^sub>d\<^sub>fI"[OF "equi:2"] "\I"(1)) + fix x + AOT_assume 1: \O!x\ + AOT_assume 2: \[F]x\ + AOT_show \\!v ([F]v & x =\<^sub>E v)\ + proof(rule "equi:1"[THEN "\E"(2)]; + rule "\I"(2)[where \=x]; + safe dest!: "&E"(2) + intro!: "&I" "\I" 1 2 Ordinary.GEN "ord=Eequiv:1"[THEN "\E", OF 1]) + AOT_show \v =\<^sub>E x\ if \x =\<^sub>E v\ for v + by (metis that "ord=Eequiv:2"[THEN "\E"]) + qed +next + fix y + AOT_assume 1: \O!y\ + AOT_assume 2: \[F]y\ + AOT_show \\!u ([F]u & u =\<^sub>E y)\ + by(safe dest!: "&E"(2) + intro!: "equi:1"[THEN "\E"(2)] "\I"(2)[where \=y] + "&I" "\I" 1 2 GEN "ord=Eequiv:1"[THEN "\E", OF 1]) +qed(auto simp: "=E[denotes]") + + +AOT_theorem "eq-part:2": \F \\<^sub>E G \ G \\<^sub>E F\ +proof (rule "\I") + AOT_assume \F \\<^sub>E G\ + AOT_hence \\R R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E G\ + using "equi:3"[THEN "\\<^sub>d\<^sub>fE"] by blast + then AOT_obtain R where \R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E G\ + using "\E"[rotated] by blast + AOT_hence 0: \R\ & F\ & G\ & \u ([F]u \ \!v([G]v & [R]uv)) & + \v ([G]v \ \!u([F]u & [R]uv))\ + using "equi:2"[THEN "\\<^sub>d\<^sub>fE"] by blast + + AOT_have \[\xy [R]yx]\ & G\ & F\ & \u ([G]u \ \!v([F]v & [\xy [R]yx]uv)) & + \v ([F]v \ \!u([G]u & [\xy [R]yx]uv))\ + proof (AOT_subst \[\xy [R]yx]yx\ \[R]xy\ for: x y; + (safe intro!: "&I" "cqt:2[const_var]"[axiom_inst] 0[THEN "&E"(2)] + 0[THEN "&E"(1), THEN "&E"(2)]; "cqt:2[lambda]")?) + AOT_modally_strict { + AOT_have \[\xy [R]yx]xy\ if \[R]yx\ for y x + by (auto intro!: "\\C"(1) "cqt:2" + simp: "&I" "ex:1:a" prod_denotesI "rule-ui:3" that) + moreover AOT_have \[R]yx\ if \[\xy [R]yx]xy\ for y x + using "\\C"(1)[where \="\(x,y). _ (x,y)" and \\<^sub>1\\<^sub>n="(_,_)", + simplified, OF that, simplified]. + ultimately AOT_show \[\xy [R]yx]\\ \ [R]\\\ for \ \ + by (metis "deduction-theorem" "\I") + } + qed + AOT_hence \[\xy [R]yx] |: G \<^sub>1\<^sub>-\<^sub>1\\<^sub>E F\ + using "equi:2"[THEN "\\<^sub>d\<^sub>fI"] by blast + AOT_hence \\R R |: G \<^sub>1\<^sub>-\<^sub>1\\<^sub>E F\ + by (rule "\I"(1)) "cqt:2[lambda]" + AOT_thus \G \\<^sub>E F\ + using "equi:3"[THEN "\\<^sub>d\<^sub>fI"] by blast +qed + +text\Note: not explicitly in PLM.\ +AOT_theorem "eq-part:2[terms]": \\ \\<^sub>E \' \ \' \\<^sub>E \\ + using "eq-part:2"[unvarify F G] eq_den_1 eq_den_2 "\I" by meson +declare "eq-part:2[terms]"[THEN "\E", sym] + +AOT_theorem "eq-part:3": \(F \\<^sub>E G & G \\<^sub>E H) \ F \\<^sub>E H\ +proof (rule "\I") + AOT_assume \F \\<^sub>E G & G \\<^sub>E H\ + then AOT_obtain R\<^sub>1 and R\<^sub>2 where + \R\<^sub>1 |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E G\ + and \R\<^sub>2 |: G \<^sub>1\<^sub>-\<^sub>1\\<^sub>E H\ + using "equi:3"[THEN "\\<^sub>d\<^sub>fE"] "&E" "\E"[rotated] by metis + AOT_hence \: \\u ([F]u \ \!v([G]v & [R\<^sub>1]uv)) & \v ([G]v \ \!u([F]u & [R\<^sub>1]uv))\ + and \: \\u ([G]u \ \!v([H]v & [R\<^sub>2]uv)) & \v ([H]v \ \!u([G]u & [R\<^sub>2]uv))\ + using "equi:2"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2)] + "equi:2"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(1), THEN "&E"(2)] + "&I" by blast+ + AOT_have \\R R = [\xy O!x & O!y & \v ([G]v & [R\<^sub>1]xv & [R\<^sub>2]vy)]\ + by (rule "free-thms:3[lambda]") cqt_2_lambda_inst_prover + then AOT_obtain R where R_def: \R = [\xy O!x & O!y & \v ([G]v & [R\<^sub>1]xv & [R\<^sub>2]vy)]\ + using "\E"[rotated] by blast + AOT_have 1: \\!v (([H]v & [R]uv))\ if a: \[O!]u\ and b: \[F]u\ for u + proof (rule "\E"(2)[OF "equi:1"]) + AOT_obtain b where + b_prop: \[O!]b & ([G]b & [R\<^sub>1]ub & \v ([G]v & [R\<^sub>1]uv \ v =\<^sub>E b))\ + using \[THEN "&E"(1), THEN "\E"(2), THEN "\E", THEN "\E", + OF a b, THEN "\E"(1)[OF "equi:1"]] + "\E"[rotated] by blast + AOT_obtain c where + c_prop: "[O!]c & ([H]c & [R\<^sub>2]bc & \v ([H]v & [R\<^sub>2]bv \ v =\<^sub>E c))" + using \[THEN "&E"(1), THEN "\E"(2)[where \=b], THEN "\E", + OF b_prop[THEN "&E"(1)], THEN "\E", + OF b_prop[THEN "&E"(2), THEN "&E"(1), THEN "&E"(1)], + THEN "\E"(1)[OF "equi:1"]] + "\E"[rotated] by blast + AOT_show \\v ([H]v & [R]uv & \v' ([H]v' & [R]uv' \ v' =\<^sub>E v))\ + proof (safe intro!: "&I" GEN "\I" "\I"(2)[where \=c]) + AOT_show \O!c\ using c_prop "&E" by blast + next + AOT_show \[H]c\ using c_prop "&E" by blast + next + AOT_have 0: \[O!]u & [O!]c & \v ([G]v & [R\<^sub>1]uv & [R\<^sub>2]vc)\ + by (safe intro!: "&I" a c_prop[THEN "&E"(1)] "\I"(2)[where \=b] + b_prop[THEN "&E"(1)] b_prop[THEN "&E"(2), THEN "&E"(1)] + c_prop[THEN "&E"(2), THEN "&E"(1), THEN "&E"(2)]) + AOT_show \[R]uc\ + by (auto intro: "rule=E"[rotated, OF R_def[symmetric]] + intro!: "\\C"(1) "cqt:2" + simp: "&I" "ex:1:a" prod_denotesI "rule-ui:3" 0) + next + fix x + AOT_assume ordx: \O!x\ + AOT_assume \[H]x & [R]ux\ + AOT_hence hx: \[H]x\ and \[R]ux\ using "&E" by blast+ + AOT_hence \[\xy O!x & O!y & \v ([G]v & [R\<^sub>1]xv & [R\<^sub>2]vy)]ux\ + using "rule=E"[rotated, OF R_def] by fast + AOT_hence \O!u & O!x & \v ([G]v & [R\<^sub>1]uv & [R\<^sub>2]vx)\ + by (rule "\\C"(1)[where \="\(\,\'). _ \ \'" and \\<^sub>1\\<^sub>n="(_,_)", simplified]) + then AOT_obtain z where z_prop: \O!z & ([G]z & [R\<^sub>1]uz & [R\<^sub>2]zx)\ + using "&E" "\E"[rotated] by blast + AOT_hence \z =\<^sub>E b\ + using b_prop[THEN "&E"(2), THEN "&E"(2), THEN "\E"(2)[where \=z]] + using "&E" "\E" by metis + AOT_hence \z = b\ + by (metis "=E-simple:2"[THEN "\E"]) + AOT_hence \[R\<^sub>2]bx\ + using z_prop[THEN "&E"(2), THEN "&E"(2)] "rule=E" by fast + AOT_thus \x =\<^sub>E c\ + using c_prop[THEN "&E"(2), THEN "&E"(2), THEN "\E"(2)[where \=x], + THEN "\E", THEN "\E", OF ordx] + hx "&I" by blast + qed + qed + AOT_have 2: \\!u (([F]u & [R]uv))\ if a: \[O!]v\ and b: \[H]v\ for v + proof (rule "\E"(2)[OF "equi:1"]) + AOT_obtain b where + b_prop: \[O!]b & ([G]b & [R\<^sub>2]bv & \u ([G]u & [R\<^sub>2]uv \ u =\<^sub>E b))\ + using \[THEN "&E"(2), THEN "\E"(2), THEN "\E", THEN "\E", + OF a b, THEN "\E"(1)[OF "equi:1"]] + "\E"[rotated] by blast + AOT_obtain c where + c_prop: "[O!]c & ([F]c & [R\<^sub>1]cb & \v ([F]v & [R\<^sub>1]vb \ v =\<^sub>E c))" + using \[THEN "&E"(2), THEN "\E"(2)[where \=b], THEN "\E", + OF b_prop[THEN "&E"(1)], THEN "\E", + OF b_prop[THEN "&E"(2), THEN "&E"(1), THEN "&E"(1)], + THEN "\E"(1)[OF "equi:1"]] + "\E"[rotated] by blast + AOT_show \\u ([F]u & [R]uv & \v' ([F]v' & [R]v'v \ v' =\<^sub>E u))\ + proof (safe intro!: "&I" GEN "\I" "\I"(2)[where \=c]) + AOT_show \O!c\ using c_prop "&E" by blast + next + AOT_show \[F]c\ using c_prop "&E" by blast + next + AOT_have \[O!]c & [O!]v & \u ([G]u & [R\<^sub>1]cu & [R\<^sub>2]uv)\ + by (safe intro!: "&I" a "\I"(2)[where \=b] + c_prop[THEN "&E"(1)] b_prop[THEN "&E"(1)] + b_prop[THEN "&E"(2), THEN "&E"(1), THEN "&E"(1)] + b_prop[THEN "&E"(2), THEN "&E"(1), THEN "&E"(2)] + c_prop[THEN "&E"(2), THEN "&E"(1), THEN "&E"(2)]) + AOT_thus \[R]cv\ + by (auto intro: "rule=E"[rotated, OF R_def[symmetric]] + intro!: "\\C"(1) "cqt:2" + simp: "&I" "ex:1:a" prod_denotesI "rule-ui:3") + next + fix x + AOT_assume ordx: \O!x\ + AOT_assume \[F]x & [R]xv\ + AOT_hence hx: \[F]x\ and \[R]xv\ using "&E" by blast+ + AOT_hence \[\xy O!x & O!y & \v ([G]v & [R\<^sub>1]xv & [R\<^sub>2]vy)]xv\ + using "rule=E"[rotated, OF R_def] by fast + AOT_hence \O!x & O!v & \u ([G]u & [R\<^sub>1]xu & [R\<^sub>2]uv)\ + by (rule "\\C"(1)[where \="\(\,\'). _ \ \'" and \\<^sub>1\\<^sub>n="(_,_)", simplified]) + then AOT_obtain z where z_prop: \O!z & ([G]z & [R\<^sub>1]xz & [R\<^sub>2]zv)\ + using "&E" "\E"[rotated] by blast + AOT_hence \z =\<^sub>E b\ + using b_prop[THEN "&E"(2), THEN "&E"(2), THEN "\E"(2)[where \=z]] + using "&E" "\E" "&I" by metis + AOT_hence \z = b\ + by (metis "=E-simple:2"[THEN "\E"]) + AOT_hence \[R\<^sub>1]xb\ + using z_prop[THEN "&E"(2), THEN "&E"(1), THEN "&E"(2)] "rule=E" by fast + AOT_thus \x =\<^sub>E c\ + using c_prop[THEN "&E"(2), THEN "&E"(2), THEN "\E"(2)[where \=x], + THEN "\E", THEN "\E", OF ordx] + hx "&I" by blast + qed + qed + AOT_show \F \\<^sub>E H\ + apply (rule "equi:3"[THEN "\\<^sub>d\<^sub>fI"]) + apply (rule "\I"(2)[where \=R]) + by (auto intro!: 1 2 "equi:2"[THEN "\\<^sub>d\<^sub>fI"] "&I" "cqt:2[const_var]"[axiom_inst] + Ordinary.GEN "\I" Ordinary.\) +qed + +text\Note: not explicitly in PLM.\ +AOT_theorem "eq-part:3[terms]": \\ \\<^sub>E \''\ if \\ \\<^sub>E \'\ and \\' \\<^sub>E \''\ + using "eq-part:3"[unvarify F G H, THEN "\E"] eq_den_1 eq_den_2 "\I" "&I" + by (metis that(1) that(2)) +declare "eq-part:3[terms]"[trans] + +AOT_theorem "eq-part:4": \F \\<^sub>E G \ \H (H \\<^sub>E F \ H \\<^sub>E G)\ +proof(rule "\I"; rule "\I") + AOT_assume 0: \F \\<^sub>E G\ + AOT_hence 1: \G \\<^sub>E F\ using "eq-part:2"[THEN "\E"] by blast + AOT_show \\H (H \\<^sub>E F \ H \\<^sub>E G)\ + proof (rule GEN; rule "\I"; rule "\I") + AOT_show \H \\<^sub>E G\ if \H \\<^sub>E F\ for H using 0 + by (meson "&I" "eq-part:3" that "vdash-properties:6") + next + AOT_show \H \\<^sub>E F\ if \H \\<^sub>E G\ for H using 1 + by (metis "&I" "eq-part:3" that "vdash-properties:6") + qed +next + AOT_assume \\H (H \\<^sub>E F \ H \\<^sub>E G)\ + AOT_hence \F \\<^sub>E F \ F \\<^sub>E G\ using "\E" by blast + AOT_thus \F \\<^sub>E G\ using "eq-part:1" "\E" by blast +qed + +AOT_define MapsE :: \\ \ \ \ \ \ \\ ("_ |: _ \E _") + "equi-rem:1": + \R |: F \E G \\<^sub>d\<^sub>f R\ & F\ & G\ & \u ([F]u \ \!v ([G]v & [R]uv))\ + +AOT_define MapsEOneToOne :: \\ \ \ \ \ \ \\ ("_ |: _ \<^sub>1\<^sub>-\<^sub>1\E _") + "equi-rem:2": + \R |: F \<^sub>1\<^sub>-\<^sub>1\E G \\<^sub>d\<^sub>f + R |: F \E G & \t\u\v (([F]t & [F]u & [G]v) \ ([R]tv & [R]uv \ t =\<^sub>E u))\ + +AOT_define MapsEOnto :: \\ \ \ \ \ \ \\ ("_ |: _ \\<^sub>o\<^sub>n\<^sub>t\<^sub>oE _") + "equi-rem:3": + \R |: F \\<^sub>o\<^sub>n\<^sub>t\<^sub>oE G \\<^sub>d\<^sub>f R |: F \E G & \v ([G]v \ \u ([F]u & [R]uv))\ + +AOT_define MapsEOneToOneOnto :: \\ \ \ \ \ \ \\ ("_ |: _ \<^sub>1\<^sub>-\<^sub>1\\<^sub>o\<^sub>n\<^sub>t\<^sub>oE _") + "equi-rem:4": + \R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>o\<^sub>n\<^sub>t\<^sub>oE G \\<^sub>d\<^sub>f R |: F \<^sub>1\<^sub>-\<^sub>1\E G & R |: F \\<^sub>o\<^sub>n\<^sub>t\<^sub>oE G\ + +AOT_theorem "equi-rem-thm": + \R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E G \ R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>o\<^sub>n\<^sub>t\<^sub>oE G\ +proof - + AOT_have \R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E G \ R |: [\x O!x & [F]x] \<^sub>1\<^sub>-\<^sub>1\ [\x O!x & [G]x]\ + proof(safe intro!: "\I" "\I" "&I") + AOT_assume \R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E G\ + AOT_hence \\u ([F]u \ \!v ([G]v & [R]uv))\ + and \\v ([G]v \ \!u ([F]u & [R]uv))\ + using "equi:2"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast+ + AOT_hence a: \([F]u \ \!v ([G]v & [R]uv))\ + and b: \([G]v \ \!u ([F]u & [R]uv))\ for u v + using "Ordinary.\E" by fast+ + AOT_have \([\x [O!]x & [F]x]x \ \!y ([\x [O!]x & [G]x]y & [R]xy))\ for x + apply (AOT_subst \[\x [O!]x & [F]x]x\ \[O!]x & [F]x\) + apply (rule "beta-C-meta"[THEN "\E"]) + apply "cqt:2[lambda]" + apply (AOT_subst \[\x [O!]x & [G]x]x\ \[O!]x & [G]x\ for: x) + apply (rule "beta-C-meta"[THEN "\E"]) + apply "cqt:2[lambda]" + apply (AOT_subst \O!y & [G]y & [R]xy\ \O!y & ([G]y & [R]xy)\ for: y) + apply (meson "\E"(6) "Associativity of &" "oth-class-taut:3:a") + apply (rule "\I") apply (frule "&E"(1)) apply (drule "&E"(2)) + by (fact a[unconstrain u, THEN "\E", THEN "\E", of x]) + AOT_hence A: \\x ([\x [O!]x & [F]x]x \ \!y ([\x [O!]x & [G]x]y & [R]xy))\ + by (rule GEN) + AOT_have \([\x [O!]x & [G]x]y \ \!x ([\x [O!]x & [F]x]x & [R]xy))\ for y + apply (AOT_subst \[\x [O!]x & [G]x]y\ \[O!]y & [G]y\) + apply (rule "beta-C-meta"[THEN "\E"]) + apply "cqt:2[lambda]" + apply (AOT_subst \[\x [O!]x & [F]x]x\ \[O!]x & [F]x\ for: x) + apply (rule "beta-C-meta"[THEN "\E"]) + apply "cqt:2[lambda]" + apply (AOT_subst \O!x & [F]x & [R]xy\ \O!x & ([F]x & [R]xy)\ for: x) + apply (meson "\E"(6) "Associativity of &" "oth-class-taut:3:a") + apply (rule "\I") apply (frule "&E"(1)) apply (drule "&E"(2)) + by (fact b[unconstrain v, THEN "\E", THEN "\E", of y]) + AOT_hence B: \\y ([\x [O!]x & [G]x]y \ \!x ([\x [O!]x & [F]x]x & [R]xy))\ + by (rule GEN) + AOT_show \R |: [\x [O!]x & [F]x] \<^sub>1\<^sub>-\<^sub>1\ [\x [O!]x & [G]x]\ + by (safe intro!: "1-1-cor"[THEN "\\<^sub>d\<^sub>fI"] "&I" + "cqt:2[const_var]"[axiom_inst] A B) + "cqt:2[lambda]"+ + next + AOT_assume \R |: [\x [O!]x & [F]x] \<^sub>1\<^sub>-\<^sub>1\ [\x [O!]x & [G]x]\ + AOT_hence a: \([\x [O!]x & [F]x]x \ \!y ([\x [O!]x & [G]x]y & [R]xy))\ and + b: \([\x [O!]x & [G]x]y \ \!x ([\x [O!]x & [F]x]x & [R]xy))\ for x y + using "1-1-cor"[THEN "\\<^sub>d\<^sub>fE"] "&E" "\E"(2) by blast+ + AOT_have \[F]u \ \!v ([G]v & [R]uv)\ for u + proof (safe intro!: "\I") + AOT_assume fu: \[F]u\ + AOT_have 0: \[\x [O!]x & [F]x]u\ + by (auto intro!: "\\C"(1) "cqt:2" "cqt:2[const_var]"[axiom_inst] + Ordinary.\ fu "&I") + AOT_show \\!v ([G]v & [R]uv)\ + apply (AOT_subst \[O!]x & ([G]x & [R]ux)\ + \([O!]x & [G]x) & [R]ux\ for: x) + apply (simp add: "Associativity of &") + apply (AOT_subst (reverse) \[O!]x & [G]x\ + \[\x [O!]x & [G]x]x\ for: x) + apply (rule "beta-C-meta"[THEN "\E"]) + apply "cqt:2[lambda]" + using a[THEN "\E", OF 0] by blast + qed + AOT_hence A: \\u ([F]u \ \!v ([G]v & [R]uv))\ + by (rule Ordinary.GEN) + AOT_have \[G]v \ \!u ([F]u & [R]uv)\ for v + proof (safe intro!: "\I") + AOT_assume gu: \[G]v\ + AOT_have 0: \[\x [O!]x & [G]x]v\ + by (auto intro!: "\\C"(1) "cqt:2" "cqt:2[const_var]"[axiom_inst] + Ordinary.\ gu "&I") + AOT_show \\!u ([F]u & [R]uv)\ + apply (AOT_subst \[O!]x & ([F]x & [R]xv)\ \([O!]x & [F]x) & [R]xv\ for: x) + apply (simp add: "Associativity of &") + apply (AOT_subst (reverse) \[O!]x & [F]x\\[\x [O!]x & [F]x]x\ for: x) + apply (rule "beta-C-meta"[THEN "\E"]) + apply "cqt:2[lambda]" + using b[THEN "\E", OF 0] by blast + qed + AOT_hence B: \\v ([G]v \ \!u ([F]u & [R]uv))\ by (rule Ordinary.GEN) + AOT_show \R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E G\ + by (safe intro!: "equi:2"[THEN "\\<^sub>d\<^sub>fI"] "&I" A B "cqt:2[const_var]"[axiom_inst]) + qed + also AOT_have \\ \ R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>o\<^sub>n\<^sub>t\<^sub>oE G\ + proof(safe intro!: "\I" "\I" "&I") + AOT_assume \R |: [\x [O!]x & [F]x] \<^sub>1\<^sub>-\<^sub>1\ [\x [O!]x & [G]x]\ + AOT_hence a: \([\x [O!]x & [F]x]x \ \!y ([\x [O!]x & [G]x]y & [R]xy))\ and + b: \([\x [O!]x & [G]x]y \ \!x ([\x [O!]x & [F]x]x & [R]xy))\ for x y + using "1-1-cor"[THEN "\\<^sub>d\<^sub>fE"] "&E" "\E"(2) by blast+ + AOT_show \R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>o\<^sub>n\<^sub>t\<^sub>oE G\ + proof (safe intro!: "equi-rem:4"[THEN "\\<^sub>d\<^sub>fI"] "&I" "equi-rem:3"[THEN "\\<^sub>d\<^sub>fI"] + "equi-rem:2"[THEN "\\<^sub>d\<^sub>fI"] "equi-rem:1"[THEN "\\<^sub>d\<^sub>fI"] + "cqt:2[const_var]"[axiom_inst] Ordinary.GEN "\I") + fix u + AOT_assume fu: \[F]u\ + AOT_have 0: \[\x [O!]x & [F]x]u\ + by (auto intro!: "\\C"(1) "cqt:2" "cqt:2[const_var]"[axiom_inst] + Ordinary.\ fu "&I") + AOT_hence 1: \\!y ([\x [O!]x & [G]x]y & [R]uy)\ + using a[THEN "\E"] by blast + AOT_show \\!v ([G]v & [R]uv)\ + apply (AOT_subst \[O!]x & ([G]x & [R]ux)\ \([O!]x & [G]x) & [R]ux\ for: x) + apply (simp add: "Associativity of &") + apply (AOT_subst (reverse) \[O!]x & [G]x\ \[\x [O!]x & [G]x]x\ for: x) + apply (rule "beta-C-meta"[THEN "\E"]) + apply "cqt:2[lambda]" + by (fact 1) + next + fix t u v + AOT_assume \[F]t & [F]u & [G]v\ and rtv_tuv: \[R]tv & [R]uv\ + AOT_hence oft: \[\x O!x & [F]x]t\ and + ofu: \[\x O!x & [F]x]u\ and + ogv: \[\x O!x & [G]x]v\ + by (auto intro!: "\\C"(1) "cqt:2" "&I" + simp: Ordinary.\ dest: "&E") + AOT_hence \\!x ([\x [O!]x & [F]x]x & [R]xv)\ + using b[THEN "\E"] by blast + then AOT_obtain a where + a_prop: \[\x [O!]x & [F]x]a & [R]av & + \x (([\x [O!]x & [F]x]x & [R]xv) \ x = a)\ + using "uniqueness:1"[THEN "\\<^sub>d\<^sub>fE"] "\E"[rotated] by blast + AOT_hence ua: \u = a\ + using ofu rtv_tuv[THEN "&E"(2)] "\E"(2) "\E" "&I" "&E"(2) by blast + moreover AOT_have ta: \t = a\ + using a_prop oft rtv_tuv[THEN "&E"(1)] "\E"(2) "\E" "&I" "&E"(2) by blast + ultimately AOT_have \t = u\ by (metis "rule=E" id_sym) + AOT_thus \t =\<^sub>E u\ + using "rule=E" id_sym "ord=Eequiv:1" Ordinary.\ ta ua "\E" by fast + next + fix u + AOT_assume \[F]u\ + AOT_hence \[\x O!x & [F]x]u\ + by (auto intro!: "\\C"(1) "cqt:2" "&I" + simp: "cqt:2[const_var]"[axiom_inst] Ordinary.\) + AOT_hence \\!y ([\x [O!]x & [G]x]y & [R]uy)\ + using a[THEN "\E"] by blast + then AOT_obtain a where + a_prop: \[\x [O!]x & [G]x]a & [R]ua & + \x (([\x [O!]x & [G]x]x & [R]ux) \ x = a)\ + using "uniqueness:1"[THEN "\\<^sub>d\<^sub>fE"] "\E"[rotated] by blast + AOT_have \O!a & [G]a\ + by (rule "\\C"(1)) (auto simp: a_prop[THEN "&E"(1), THEN "&E"(1)]) + AOT_hence \O!a\ and \[G]a\ using "&E" by blast+ + moreover AOT_have \\v ([G]v & [R]uv \ v =\<^sub>E a)\ + proof(safe intro!: Ordinary.GEN "\I"; frule "&E"(1); drule "&E"(2)) + fix v + AOT_assume \[G]v\ and ruv: \[R]uv\ + AOT_hence \[\x [O!]x & [G]x]v\ + by (auto intro!: "\\C"(1) "cqt:2" "&I" simp: Ordinary.\) + AOT_hence \v = a\ + using a_prop[THEN "&E"(2), THEN "\E"(2), THEN "\E", OF "&I"] ruv by blast + AOT_thus \v =\<^sub>E a\ + using "rule=E" "ord=Eequiv:1" Ordinary.\ "\E" by fast + qed + ultimately AOT_have \O!a & ([G]a & [R]ua & \v' ([G]v' & [R]uv' \ v' =\<^sub>E a))\ + using "\I" "&I" a_prop[THEN "&E"(1), THEN "&E"(2)] by simp + AOT_hence \\v ([G]v & [R]uv & \v' ([G]v' & [R]uv' \ v' =\<^sub>E v))\ + by (rule "\I") + AOT_thus \\!v ([G]v & [R]uv)\ + by (rule "equi:1"[THEN "\E"(2)]) + next + fix v + AOT_assume \[G]v\ + AOT_hence \[\x O!x & [G]x]v\ + by (auto intro!: "\\C"(1) "cqt:2" "&I" Ordinary.\) + AOT_hence \\!x ([\x [O!]x & [F]x]x & [R]xv)\ + using b[THEN "\E"] by blast + then AOT_obtain a where + a_prop: \[\x [O!]x & [F]x]a & [R]av & + \y ([\x [O!]x & [F]x]y & [R]yv \ y = a)\ + using "uniqueness:1"[THEN "\\<^sub>d\<^sub>fE", THEN "\E"[rotated]] by blast + AOT_have \O!a & [F]a\ + by (rule "\\C"(1)) (auto simp: a_prop[THEN "&E"(1), THEN "&E"(1)]) + AOT_hence \O!a & ([F]a & [R]av)\ + using a_prop[THEN "&E"(1), THEN "&E"(2)] "&E" "&I" by metis + AOT_thus \\u ([F]u & [R]uv)\ + by (rule "\I") + qed + next + AOT_assume \R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>o\<^sub>n\<^sub>t\<^sub>oE G\ + AOT_hence 1: \R |: F \<^sub>1\<^sub>-\<^sub>1\E G\ + and 2: \R |: F \\<^sub>o\<^sub>n\<^sub>t\<^sub>oE G\ + using "equi-rem:4"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast+ + AOT_hence 3: \R |: F \E G\ + and A: \\t \u \v ([F]t & [F]u & [G]v \ ([R]tv & [R]uv \ t =\<^sub>E u))\ + using "equi-rem:2"[THEN "\\<^sub>d\<^sub>fE", OF 1] "&E" by blast+ + AOT_hence B: \\u ([F]u \ \!v ([G]v & [R]uv))\ + using "equi-rem:1"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + AOT_have C: \\v ([G]v \ \u ([F]u & [R]uv))\ + using "equi-rem:3"[THEN "\\<^sub>d\<^sub>fE", OF 2] "&E" by blast + AOT_show \R |: [\x [O!]x & [F]x] \<^sub>1\<^sub>-\<^sub>1\ [\x [O!]x & [G]x]\ + proof (rule "1-1-cor"[THEN "\\<^sub>d\<^sub>fI"]; + safe intro!: "&I" "cqt:2" GEN "\I") + fix x + AOT_assume 1: \[\x [O!]x & [F]x]x\ + AOT_have \O!x & [F]x\ + by (rule "\\C"(1)) (auto simp: 1) + AOT_hence \\!v ([G]v & [R]xv)\ + using B[THEN "\E"(2), THEN "\E", THEN "\E"] "&E" by blast + then AOT_obtain y where + y_prop: \O!y & ([G]y & [R]xy & \u ([G]u & [R]xu \ u =\<^sub>E y))\ + using "equi:1"[THEN "\E"(1)] "\E"[rotated] by fastforce + AOT_hence \[\x O!x & [G]x]y\ + by (auto intro!: "\\C"(1) "cqt:2" "&I" dest: "&E") + moreover AOT_have \\z ([\x O!x & [G]x]z & [R]xz \ z = y)\ + proof(safe intro!: GEN "\I"; frule "&E"(1); drule "&E"(2)) + fix z + AOT_assume 1: \[\x [O!]x & [G]x]z\ + AOT_have 2: \O!z & [G]z\ + by (rule "\\C"(1)) (auto simp: 1) + moreover AOT_assume \[R]xz\ + ultimately AOT_have \z =\<^sub>E y\ + using y_prop[THEN "&E"(2), THEN "&E"(2), THEN "\E"(2), + THEN "\E", THEN "\E", rotated, OF "&I"] "&E" + by blast + AOT_thus \z = y\ + using 2[THEN "&E"(1)] by (metis "=E-simple:2" "\E") + qed + ultimately AOT_have \[\x O!x & [G]x]y & [R]xy & + \z ([\x O!x & [G]x]z & [R]xz \ z = y)\ + using y_prop[THEN "&E"(2), THEN "&E"(1), THEN "&E"(2)] "&I" by auto + AOT_hence \\y ([\x O!x & [G]x]y & [R]xy & + \z ([\x O!x & [G]x]z & [R]xz \ z = y))\ + by (rule "\I") + AOT_thus \\!y ([\x [O!]x & [G]x]y & [R]xy)\ + using "uniqueness:1"[THEN "\\<^sub>d\<^sub>fI"] by fast + next + fix y + AOT_assume 1: \[\x [O!]x & [G]x]y\ + AOT_have oy_gy: \O!y & [G]y\ + by (rule "\\C"(1)) (auto simp: 1) + AOT_hence \\u ([F]u & [R]uy)\ + using C[THEN "\E"(2), THEN "\E", THEN "\E"] "&E" by blast + then AOT_obtain x where x_prop: \O!x & ([F]x & [R]xy)\ + using "\E"[rotated] by blast + AOT_hence ofx: \[\x O!x & [F]x]x\ + by (auto intro!: "\\C"(1) "cqt:2" "&I" dest: "&E") + AOT_have \\\ ([\x [O!]x & [F]x]\ & [R]\y & + \\ ([\x [O!]x & [F]x]\ & [R]\y \ \ = \))\ + proof (safe intro!: "\I"(2)[where \=x] "&I" GEN "\I") + AOT_show \[\x O!x & [F]x]x\ using ofx. + next + AOT_show \[R]xy\ using x_prop[THEN "&E"(2), THEN "&E"(2)]. + next + fix z + AOT_assume 1: \[\x [O!]x & [F]x]z & [R]zy\ + AOT_have oz_fz: \O!z & [F]z\ + by (rule "\\C"(1)) (auto simp: 1[THEN "&E"(1)]) + AOT_have \z =\<^sub>E x\ + using A[THEN "\E"(2)[where \=z], THEN "\E", THEN "\E"(2)[where \=x], + THEN "\E", THEN "\E"(2)[where \=y], THEN "\E", + THEN "\E", THEN "\E", OF oz_fz[THEN "&E"(1)], + OF x_prop[THEN "&E"(1)], OF oy_gy[THEN "&E"(1)], OF "&I", OF "&I", + OF oz_fz[THEN "&E"(2)], OF x_prop[THEN "&E"(2), THEN "&E"(1)], + OF oy_gy[THEN "&E"(2)], OF "&I", OF 1[THEN "&E"(2)], + OF x_prop[THEN "&E"(2), THEN "&E"(2)]]. + AOT_thus \z = x\ + by (metis "=E-simple:2" "vdash-properties:10") + qed + AOT_thus \\!x ([\x [O!]x & [F]x]x & [R]xy)\ + by (rule "uniqueness:1"[THEN "\\<^sub>d\<^sub>fI"]) + qed + qed + finally show ?thesis. +qed + +AOT_theorem "empty-approx:1": \(\\u [F]u & \\v [H]v) \ F \\<^sub>E H\ +proof(rule "\I"; frule "&E"(1); drule "&E"(2)) + AOT_assume 0: \\\u [F]u\ and 1: \\\v [H]v\ + AOT_have \\u ([F]u \ \!v ([H]v & [R]uv))\ for R + proof(rule Ordinary.GEN; rule "\I"; rule "raa-cor:1") + fix u + AOT_assume \[F]u\ + AOT_hence \\u [F]u\ using "Ordinary.\I" "&I" by fast + AOT_thus \\u [F]u & \\u [F]u\ using "&I" 0 by blast + qed + moreover AOT_have \\v ([H]v \ \!u ([F]u & [R]uv))\ for R + proof(rule Ordinary.GEN; rule "\I"; rule "raa-cor:1") + fix v + AOT_assume \[H]v\ + AOT_hence \\v [H]v\ using "Ordinary.\I" "&I" by fast + AOT_thus \\v [H]v & \\v [H]v\ using 1 "&I" by blast + qed + ultimately AOT_have \R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E H\ for R + apply (safe intro!: "equi:2"[THEN "\\<^sub>d\<^sub>fI"] "&I" GEN "cqt:2[const_var]"[axiom_inst]) + using "\E" by blast+ + AOT_hence \\R R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E H\ by (rule "\I") + AOT_thus \F \\<^sub>E H\ + by (rule "equi:3"[THEN "\\<^sub>d\<^sub>fI"]) +qed + +AOT_theorem "empty-approx:2": \(\u [F]u & \\v [H]v) \ \(F \\<^sub>E H)\ +proof(rule "\I"; frule "&E"(1); drule "&E"(2); rule "raa-cor:2") + AOT_assume 1: \\u [F]u\ and 2: \\\v [H]v\ + AOT_obtain b where b_prop: \O!b & [F]b\ + using 1 "\E"[rotated] by blast + AOT_assume \F \\<^sub>E H\ + AOT_hence \\R R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E H\ + by (rule "equi:3"[THEN "\\<^sub>d\<^sub>fE"]) + then AOT_obtain R where \R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E H\ + using "\E"[rotated] by blast + AOT_hence \: \\u ([F]u \ \!v ([H]v & [R]uv))\ + using "equi:2"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast+ + AOT_have \\!v ([H]v & [R]bv)\ for u + using \[THEN "\E"(2)[where \=b], THEN "\E", THEN "\E", + OF b_prop[THEN "&E"(1)], OF b_prop[THEN "&E"(2)]]. + AOT_hence \\v ([H]v & [R]bv & \u ([H]u & [R]bu \ u =\<^sub>E v))\ + by (rule "equi:1"[THEN "\E"(1)]) + then AOT_obtain x where \O!x & ([H]x & [R]bx & \u ([H]u & [R]bu \ u =\<^sub>E x))\ + using "\E"[rotated] by blast + AOT_hence \O!x & [H]x\ using "&E" "&I" by blast + AOT_hence \\v [H]v\ by (rule "\I") + AOT_thus \\v [H]v & \\v [H]v\ using 2 "&I" by blast +qed + + +AOT_define FminusU :: \\ \ \ \ \\ ("_\<^sup>-\<^sup>_") + "F-u": \[F]\<^sup>-\<^sup>x =\<^sub>d\<^sub>f [\z [F]z & z \\<^sub>E x]\ + +text\Note: not explicitly in PLM.\ +AOT_theorem "F-u[den]": \[F]\<^sup>-\<^sup>x\\ + by (rule "=\<^sub>d\<^sub>fI"(1)[OF "F-u", where \\<^sub>1\\<^sub>n="(_,_)", simplified]; "cqt:2[lambda]") +AOT_theorem "F-u[equiv]": \[[F]\<^sup>-\<^sup>x]y \ ([F]y & y \\<^sub>E x)\ + by (auto intro: "F-u"[THEN "=\<^sub>d\<^sub>fI"(1), where \\<^sub>1\\<^sub>n="(_,_)", simplified] + intro!: "cqt:2" "beta-C-cor:2"[THEN "\E", THEN "\E"(2)]) + +AOT_theorem eqP': \F \\<^sub>E G & [F]u & [G]v \ [F]\<^sup>-\<^sup>u \\<^sub>E [G]\<^sup>-\<^sup>v\ +proof (rule "\I"; frule "&E"(2); drule "&E"(1); frule "&E"(2); drule "&E"(1)) + AOT_assume \F \\<^sub>E G\ + AOT_hence \\R R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E G\ + using "equi:3"[THEN "\\<^sub>d\<^sub>fE"] by blast + then AOT_obtain R where R_prop: \R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E G\ + using "\E"[rotated] by blast + AOT_hence A: \\u ([F]u \ \!v ([G]v & [R]uv))\ + and B: \\v ([G]v \ \!u ([F]u & [R]uv))\ + using "equi:2"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast+ + AOT_have \R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>o\<^sub>n\<^sub>t\<^sub>oE G\ + using "equi-rem-thm"[THEN "\E"(1), OF R_prop]. + AOT_hence \R |: F \<^sub>1\<^sub>-\<^sub>1\E G & R |: F \\<^sub>o\<^sub>n\<^sub>t\<^sub>oE G\ + using "equi-rem:4"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_hence C: \\t\u\v (([F]t & [F]u & [G]v) \ ([R]tv & [R]uv \ t =\<^sub>E u))\ + using "equi-rem:2"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + AOT_assume fu: \[F]u\ + AOT_assume gv: \[G]v\ + AOT_have \[\z [\]z & z \\<^sub>E \]\\ for \ \ + by "cqt:2[lambda]" + note \_minus_\I = "rule-id-df:2:b[2]"[ + where \=\(\(\, \). \[\]\<^sup>-\<^sup>\\)\, simplified, OF "F-u", simplified, OF this] + and \_minus_\E = "rule-id-df:2:a[2]"[ + where \=\(\(\, \). \[\]\<^sup>-\<^sup>\\)\, simplified, OF "F-u", simplified, OF this] + AOT_have \_minus_\_den: \[\]\<^sup>-\<^sup>\\\ for \ \ + by (rule \_minus_\I) "cqt:2[lambda]"+ + { + fix R + AOT_assume R_prop: \R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E G\ + AOT_hence A: \\u ([F]u \ \!v ([G]v & [R]uv))\ + and B: \\v ([G]v \ \!u ([F]u & [R]uv))\ + using "equi:2"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast+ + AOT_have \R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>o\<^sub>n\<^sub>t\<^sub>oE G\ + using "equi-rem-thm"[THEN "\E"(1), OF R_prop]. + AOT_hence \R |: F \<^sub>1\<^sub>-\<^sub>1\E G & R |: F \\<^sub>o\<^sub>n\<^sub>t\<^sub>oE G\ + using "equi-rem:4"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_hence C: \\t\u\v (([F]t & [F]u & [G]v) \ ([R]tv & [R]uv \ t =\<^sub>E u))\ + using "equi-rem:2"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + + AOT_assume Ruv: \[R]uv\ + AOT_have \R |: [F]\<^sup>-\<^sup>u \<^sub>1\<^sub>-\<^sub>1\\<^sub>E [G]\<^sup>-\<^sup>v\ + proof(safe intro!: "equi:2"[THEN "\\<^sub>d\<^sub>fI"] "&I" "cqt:2[const_var]"[axiom_inst] + \_minus_\_den Ordinary.GEN "\I") + fix u' + AOT_assume \[[F]\<^sup>-\<^sup>u]u'\ + AOT_hence 0: \[\z [F]z & z \\<^sub>E u]u'\ + using \_minus_\E by fast + AOT_have 0: \[F]u' & u' \\<^sub>E u\ + by (rule "\\C"(1)[where \\<^sub>1\\<^sub>n="AOT_term_of_var (Ordinary.Rep u')"]) (fact 0) + AOT_have \\!v ([G]v & [R]u'v)\ + using A[THEN "Ordinary.\E"[where \=u'], THEN "\E", OF 0[THEN "&E"(1)]]. + then AOT_obtain v' where + v'_prop: \[G]v' & [R]u'v' & \ t ([G]t & [R]u't \ t =\<^sub>E v')\ + using "equi:1"[THEN "\E"(1)] "Ordinary.\E"[rotated] by fastforce + + AOT_show \\!v' ([[G]\<^sup>-\<^sup>v]v' & [R]u'v')\ + proof (safe intro!: "equi:1"[THEN "\E"(2)] "Ordinary.\I"[where \=v'] + "&I" Ordinary.GEN "\I") + AOT_show \[[G]\<^sup>-\<^sup>v]v'\ + proof (rule \_minus_\I; + safe intro!: "\\C"(1) "cqt:2" "&I" "thm-neg=E"[THEN "\E"(2)]) + AOT_show \[G]v'\ using v'_prop "&E" by blast + next + AOT_show \\v' =\<^sub>E v\ + proof (rule "raa-cor:2") + AOT_assume \v' =\<^sub>E v\ + AOT_hence \v' = v\ by (metis "=E-simple:2" "\E") + AOT_hence Ruv': \[R]uv'\ using "rule=E" Ruv id_sym by fast + AOT_have \u' =\<^sub>E u\ + by (rule C[THEN "Ordinary.\E", THEN "Ordinary.\E", + THEN "Ordinary.\E"[where \=v'], THEN "\E", THEN "\E"]) + (safe intro!: "&I" 0[THEN "&E"(1)] fu + v'_prop[THEN "&E"(1), THEN "&E"(1)] + Ruv' v'_prop[THEN "&E"(1), THEN "&E"(2)]) + moreover AOT_have \\(u' =\<^sub>E u)\ + using "0" "&E"(2) "\E"(1) "thm-neg=E" by blast + ultimately AOT_show \u' =\<^sub>E u & \u' =\<^sub>E u\ using "&I" by blast + qed + qed + next + AOT_show \[R]u'v'\ using v'_prop "&E" by blast + next + fix t + AOT_assume t_prop: \[[G]\<^sup>-\<^sup>v]t & [R]u't\ + AOT_have gt_t_noteq_v: \[G]t & t \\<^sub>E v\ + apply (rule "\\C"(1)[where \\<^sub>1\\<^sub>n="AOT_term_of_var (Ordinary.Rep t)"]) + apply (rule \_minus_\E) + by (fact t_prop[THEN "&E"(1)]) + AOT_show \t =\<^sub>E v'\ + using v'_prop[THEN "&E"(2), THEN "Ordinary.\E", THEN "\E", + OF "&I", OF gt_t_noteq_v[THEN "&E"(1)], + OF t_prop[THEN "&E"(2)]]. + qed + next + fix v' + AOT_assume G_minus_v_v': \[[G]\<^sup>-\<^sup>v]v'\ + AOT_have gt_t_noteq_v: \[G]v' & v' \\<^sub>E v\ + apply (rule "\\C"(1)[where \\<^sub>1\\<^sub>n="AOT_term_of_var (Ordinary.Rep v')"]) + apply (rule \_minus_\E) + by (fact G_minus_v_v') + AOT_have \\!u([F]u & [R]uv')\ + using B[THEN "Ordinary.\E", THEN "\E", OF gt_t_noteq_v[THEN "&E"(1)]]. + then AOT_obtain u' where + u'_prop: \[F]u' & [R]u'v' & \t ([F]t & [R]tv' \ t =\<^sub>E u')\ + using "equi:1"[THEN "\E"(1)] "Ordinary.\E"[rotated] by fastforce + AOT_show \\!u' ([[F]\<^sup>-\<^sup>u]u' & [R]u'v')\ + proof (safe intro!: "equi:1"[THEN "\E"(2)] "Ordinary.\I"[where \=u'] "&I" + u'_prop[THEN "&E"(1), THEN "&E"(2)] Ordinary.GEN "\I") + AOT_show \[[F]\<^sup>-\<^sup>u]u'\ + proof (rule \_minus_\I; + safe intro!: "\\C"(1) "cqt:2" "&I" "thm-neg=E"[THEN "\E"(2)] + u'_prop[THEN "&E"(1), THEN "&E"(1)]; rule "raa-cor:2") + AOT_assume u'_eq_u: \u' =\<^sub>E u\ + AOT_hence \u' = u\ + using "=E-simple:2" "vdash-properties:10" by blast + AOT_hence Ru'v: \[R]u'v\ using "rule=E" Ruv id_sym by fast + AOT_have \v' \\<^sub>E v\ + using "&E"(2) gt_t_noteq_v by blast + AOT_hence v'_noteq_v: \\(v' =\<^sub>E v)\ by (metis "\E"(1) "thm-neg=E") + AOT_have \\u ([G]u & [R]u'u & \v ([G]v & [R]u'v \ v =\<^sub>E u))\ + using A[THEN "Ordinary.\E", THEN "\E", + OF u'_prop[THEN "&E"(1), THEN "&E"(1)], + THEN "equi:1"[THEN "\E"(1)]]. + then AOT_obtain t where + t_prop: \[G]t & [R]u't & \v ([G]v & [R]u'v \ v =\<^sub>E t)\ + using "Ordinary.\E"[rotated] by meson + AOT_have \v =\<^sub>E t\ if \[G]v\ and \[R]u'v\ for v + using t_prop[THEN "&E"(2), THEN "Ordinary.\E", THEN "\E", + OF "&I", OF that]. + AOT_hence \v' =\<^sub>E t\ and \v =\<^sub>E t\ + by (auto simp: gt_t_noteq_v[THEN "&E"(1)] Ru'v gv + u'_prop[THEN "&E"(1), THEN "&E"(2)]) + AOT_hence \v' =\<^sub>E v\ + using "rule=E" "=E-simple:2" id_sym "\E" by fast + AOT_thus \v' =\<^sub>E v & \v' =\<^sub>E v\ + using v'_noteq_v "&I" by blast + qed + next + fix t + AOT_assume 0: \[[F]\<^sup>-\<^sup>u]t & [R]tv'\ + moreover AOT_have \[F]t & t \\<^sub>E u\ + apply (rule "\\C"(1)[where \\<^sub>1\\<^sub>n="AOT_term_of_var (Ordinary.Rep t)"]) + apply (rule \_minus_\E) + by (fact 0[THEN "&E"(1)]) + ultimately AOT_show \t =\<^sub>E u'\ + using u'_prop[THEN "&E"(2), THEN "Ordinary.\E", THEN "\E", OF "&I"] + "&E" by blast + qed + qed + AOT_hence \\R R |: [F]\<^sup>-\<^sup>u \<^sub>1\<^sub>-\<^sub>1\\<^sub>E [G]\<^sup>-\<^sup>v\ + by (rule "\I") + } note 1 = this + moreover { + AOT_assume not_Ruv: \\[R]uv\ + AOT_have \\!v ([G]v & [R]uv)\ + using A[THEN "Ordinary.\E", THEN "\E", OF fu]. + then AOT_obtain b where + b_prop: \O!b & ([G]b & [R]ub & \t([G]t & [R]ut \ t =\<^sub>E b))\ + using "equi:1"[THEN "\E"(1)] "\E"[rotated] by fastforce + AOT_hence ob: \O!b\ and gb: \[G]b\ and Rub: \[R]ub\ + using "&E" by blast+ + AOT_have \O!t \ ([G]t & [R]ut \ t =\<^sub>E b)\ for t + using b_prop "&E"(2) "\E"(2) by blast + AOT_hence b_unique: \t =\<^sub>E b\ if \O!t\ and \[G]t\ and \[R]ut\ for t + by (metis Adjunction "modus-tollens:1" "reductio-aa:1" that) + AOT_have not_v_eq_b: \\(v =\<^sub>E b)\ + proof(rule "raa-cor:2") + AOT_assume \v =\<^sub>E b\ + AOT_hence 0: \v = b\ + by (metis "=E-simple:2" "\E") + AOT_have \[R]uv\ + using b_prop[THEN "&E"(2), THEN "&E"(1), THEN "&E"(2)] + "rule=E"[rotated, OF 0[symmetric]] by fast + AOT_thus \[R]uv & \[R]uv\ + using not_Ruv "&I" by blast + qed + AOT_have not_b_eq_v: \\(b =\<^sub>E v)\ + using "modus-tollens:1" not_v_eq_b "ord=Eequiv:2" by blast + AOT_have \\!u ([F]u & [R]uv)\ + using B[THEN "Ordinary.\E", THEN "\E", OF gv]. + then AOT_obtain a where + a_prop: \O!a & ([F]a & [R]av & \t([F]t & [R]tv \ t =\<^sub>E a))\ + using "equi:1"[THEN "\E"(1)] "\E"[rotated] by fastforce + AOT_hence Oa: \O!a\ and fa: \[F]a\ and Rav: \[R]av\ + using "&E" by blast+ + AOT_have \O!t \ ([F]t & [R]tv \ t =\<^sub>E a)\ for t + using a_prop "&E" "\E"(2) by blast + AOT_hence a_unique: \t =\<^sub>E a\ if \O!t\ and \[F]t\ and \[R]tv\ for t + by (metis Adjunction "modus-tollens:1" "reductio-aa:1" that) + AOT_have not_u_eq_a: \\(u =\<^sub>E a)\ + proof(rule "raa-cor:2") + AOT_assume \u =\<^sub>E a\ + AOT_hence 0: \u = a\ + by (metis "=E-simple:2" "\E") + AOT_have \[R]uv\ + using a_prop[THEN "&E"(2), THEN "&E"(1), THEN "&E"(2)] + "rule=E"[rotated, OF 0[symmetric]] by fast + AOT_thus \[R]uv & \[R]uv\ + using not_Ruv "&I" by blast + qed + AOT_have not_a_eq_u: \\(a =\<^sub>E u)\ + using "modus-tollens:1" not_u_eq_a "ord=Eequiv:2" by blast + let ?R = \\[\u'v' (u' \\<^sub>E u & v' \\<^sub>E v & [R]u'v') \ + (u' =\<^sub>E a & v' =\<^sub>E b) \ + (u' =\<^sub>E u & v' =\<^sub>E v)]\\ + AOT_have \[\?R\]\\ by "cqt:2[lambda]" + AOT_hence \\ \ \ = [\?R\]\ + using "free-thms:1" "\E"(1) by fast + then AOT_obtain R\<^sub>1 where R\<^sub>1_def: \R\<^sub>1 = [\?R\]\ + using "\E"[rotated] by blast + AOT_have Rxy1: \[R]xy\ if \[R\<^sub>1]xy\ and \x \\<^sub>E u\ and \x \\<^sub>E a\ for x y + proof - + AOT_have 0: \[\?R\]xy\ + by (rule "rule=E"[rotated, OF R\<^sub>1_def]) (fact that(1)) + AOT_have \(x \\<^sub>E u & y \\<^sub>E v & [R]xy) \ (x =\<^sub>E a & y =\<^sub>E b) \ (x =\<^sub>E u & y =\<^sub>E v)\ + using "\\C"(1)[OF 0] by simp + AOT_hence \x \\<^sub>E u & y \\<^sub>E v & [R]xy\ using that(2,3) + by (metis "\E"(3) "Conjunction Simplification"(1) "\E"(1) + "modus-tollens:1" "thm-neg=E") + AOT_thus \[R]xy\ using "&E" by blast+ + qed + AOT_have Rxy2: \[R]xy\ if \[R\<^sub>1]xy\ and \y \\<^sub>E v\ and \y \\<^sub>E b\ for x y + proof - + AOT_have 0: \[\?R\]xy\ + by (rule "rule=E"[rotated, OF R\<^sub>1_def]) (fact that(1)) + AOT_have \(x \\<^sub>E u & y \\<^sub>E v & [R]xy) \ (x =\<^sub>E a & y =\<^sub>E b) \ (x =\<^sub>E u & y =\<^sub>E v)\ + using "\\C"(1)[OF 0] by simp + AOT_hence \x \\<^sub>E u & y \\<^sub>E v & [R]xy\ + using that(2,3) + by (metis "\E"(3) "Conjunction Simplification"(2) "\E"(1) + "modus-tollens:1" "thm-neg=E") + AOT_thus \[R]xy\ using "&E" by blast+ + qed + AOT_have R\<^sub>1xy: \[R\<^sub>1]xy\ if \[R]xy\ and \x \\<^sub>E u\ and \y \\<^sub>E v\ for x y + by (rule "rule=E"[rotated, OF R\<^sub>1_def[symmetric]]) + (auto intro!: "\\C"(1) "cqt:2" + simp: "&I" "ex:1:a" prod_denotesI "rule-ui:3" that "\I"(1)) + AOT_have R\<^sub>1ab: \[R\<^sub>1]ab\ + apply (rule "rule=E"[rotated, OF R\<^sub>1_def[symmetric]]) + apply (safe intro!: "\\C"(1) "cqt:2" prod_denotesI "&I") + by (meson a_prop b_prop "&I" "&E"(1) "\I"(1) "\I"(2) "ord=Eequiv:1" "\E") + AOT_have R\<^sub>1uv: \[R\<^sub>1]uv\ + apply (rule "rule=E"[rotated, OF R\<^sub>1_def[symmetric]]) + apply (safe intro!: "\\C"(1) "cqt:2" prod_denotesI "&I") + by (meson "&I" "\I"(2) "ord=Eequiv:1" Ordinary.\ "\E") + moreover AOT_have \R\<^sub>1 |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E G\ + proof (safe intro!: "equi:2"[THEN "\\<^sub>d\<^sub>fI"] "&I" "cqt:2" Ordinary.GEN "\I") + fix u' + AOT_assume fu': \[F]u'\ + { + AOT_assume not_u'_eq_u: \\(u' =\<^sub>E u)\ and not_u'_eq_a: \\(u' =\<^sub>E a)\ + AOT_hence u'_noteq_u: \u' \\<^sub>E u\ and u'_noteq_a: \u' \\<^sub>E a\ + by (metis "\E"(2) "thm-neg=E")+ + AOT_have \\!v ([G]v & [R]u'v)\ + using A[THEN "Ordinary.\E", THEN "\E", OF fu']. + AOT_hence \\v ([G]v & [R]u'v & \t ([G]t & [R]u't \ t =\<^sub>E v))\ + using "equi:1"[THEN "\E"(1)] by simp + then AOT_obtain v' where + v'_prop: \[G]v' & [R]u'v' & \t ([G]t & [R]u't \ t =\<^sub>E v')\ + using "Ordinary.\E"[rotated] by meson + AOT_hence gv': \[G]v'\ and Ru'v': \[R]u'v'\ + using "&E" by blast+ + AOT_have not_v'_eq_v: \\v' =\<^sub>E v\ + proof (rule "raa-cor:2") + AOT_assume \v' =\<^sub>E v\ + AOT_hence \v' = v\ + by (metis "=E-simple:2" "\E") + AOT_hence Ru'v: \[R]u'v\ + using "rule=E" Ru'v' by fast + AOT_have \u' =\<^sub>E a\ + using a_unique[OF Ordinary.\, OF fu', OF Ru'v]. + AOT_thus \u' =\<^sub>E a & \u' =\<^sub>E a\ + using not_u'_eq_a "&I" by blast + qed + AOT_hence v'_noteq_v: \v' \\<^sub>E v\ + using "\E"(2) "thm-neg=E" by blast + AOT_have \\t ([G]t & [R]u't \ t =\<^sub>E v')\ + using v'_prop "&E" by blast + AOT_hence \[G]t & [R]u't \ t =\<^sub>E v'\ for t + using "Ordinary.\E" by meson + AOT_hence v'_unique: \t =\<^sub>E v'\ if \[G]t\ and \[R]u't\ for t + by (metis "&I" that "\E") + + AOT_have \[G]v' & [R\<^sub>1]u'v' & \t ([G]t & [R\<^sub>1]u't \ t =\<^sub>E v')\ + proof (safe intro!: "&I" gv' R\<^sub>1xy Ru'v' u'_noteq_u u'_noteq_a "\I" + Ordinary.GEN "thm-neg=E"[THEN "\E"(2)] not_v'_eq_v) + fix t + AOT_assume 1: \[G]t & [R\<^sub>1]u't\ + AOT_have \[R]u't\ + using Rxy1[OF 1[THEN "&E"(2)], OF u'_noteq_u, OF u'_noteq_a]. + AOT_thus \t =\<^sub>E v'\ + using v'_unique 1[THEN "&E"(1)] by blast + qed + AOT_hence \\v ([G]v & [R\<^sub>1]u'v & \t ([G]t & [R\<^sub>1]u't \ t =\<^sub>E v))\ + by (rule "Ordinary.\I") + AOT_hence \\!v ([G]v & [R\<^sub>1]u'v)\ + by (rule "equi:1"[THEN "\E"(2)]) + } + moreover { + AOT_assume 0: \u' =\<^sub>E u\ + AOT_hence u'_eq_u: \u' = u\ + using "=E-simple:2" "\E" by blast + AOT_have \\!v ([G]v & [R\<^sub>1]u'v)\ + proof (safe intro!: "equi:1"[THEN "\E"(2)] "Ordinary.\I"[where \=v] + "&I" Ordinary.GEN "\I" gv) + AOT_show \[R\<^sub>1]u'v\ + apply (rule "rule=E"[rotated, OF R\<^sub>1_def[symmetric]]) + apply (safe intro!: "\\C"(1) "cqt:2" "&I" prod_denotesI) + by (safe intro!: "\I"(2) "&I" 0 "ord=Eequiv:1"[THEN "\E", OF Ordinary.\]) + next + fix v' + AOT_assume \[G]v' & [R\<^sub>1]u'v'\ + AOT_hence 0: \[R\<^sub>1]uv'\ + using "rule=E"[rotated, OF u'_eq_u] "&E"(2) by fast + AOT_have 1: \[\?R\]uv'\ + by (rule "rule=E"[rotated, OF R\<^sub>1_def]) (fact 0) + AOT_have 2: \(u \\<^sub>E u & v' \\<^sub>E v & [R]uv') \ + (u =\<^sub>E a & v' =\<^sub>E b) \ + (u =\<^sub>E u & v' =\<^sub>E v)\ + using "\\C"(1)[OF 1] by simp + AOT_have \\u \\<^sub>E u\ + using "\E"(4) "modus-tollens:1" "ord=Eequiv:1" Ordinary.\ + "reductio-aa:2" "thm-neg=E" by blast + AOT_hence \\((u \\<^sub>E u & v' \\<^sub>E v & [R]uv') \ (u =\<^sub>E a & v' =\<^sub>E b))\ + using not_u_eq_a + by (metis "\E"(2) "Conjunction Simplification"(1) + "modus-tollens:1" "reductio-aa:1") + AOT_hence \(u =\<^sub>E u & v' =\<^sub>E v)\ + using 2 by (metis "\E"(2)) + AOT_thus \v' =\<^sub>E v\ + using "&E" by blast + qed + } + moreover { + AOT_assume 0: \u' =\<^sub>E a\ + AOT_hence u'_eq_a: \u' = a\ + using "=E-simple:2" "\E" by blast + AOT_have \\!v ([G]v & [R\<^sub>1]u'v)\ + proof (safe intro!: "equi:1"[THEN "\E"(2)] "\I"(2)[where \=b] "&I" + Ordinary.GEN "\I" b_prop[THEN "&E"(1)] + b_prop[THEN "&E"(2), THEN "&E"(1), THEN "&E"(1)]) + AOT_show \[R\<^sub>1]u'b\ + apply (rule "rule=E"[rotated, OF R\<^sub>1_def[symmetric]]) + apply (safe intro!: "\\C"(1) "cqt:2" "&I" prod_denotesI) + apply (rule "\I"(1); rule "\I"(2); rule "&I") + apply (fact 0) + using b_prop "&E"(1) "ord=Eequiv:1" "\E" by blast + next + fix v' + AOT_assume gv'_R1u'v': \[G]v' & [R\<^sub>1]u'v'\ + AOT_hence 0: \[R\<^sub>1]av'\ + using u'_eq_a by (meson "rule=E" "&E"(2)) + AOT_have 1: \[\?R\]av'\ + by (rule "rule=E"[rotated, OF R\<^sub>1_def]) (fact 0) + AOT_have \(a \\<^sub>E u & v' \\<^sub>E v & [R]av') \ + (a =\<^sub>E a & v' =\<^sub>E b) \ + (a =\<^sub>E u & v' =\<^sub>E v)\ + using "\\C"(1)[OF 1] by simp + moreover { + AOT_assume 0: \a \\<^sub>E u & v' \\<^sub>E v & [R]av'\ + AOT_have \\!v ([G]v & [R]u'v)\ + using A[THEN "Ordinary.\E", THEN "\E", OF fu']. + AOT_hence \\!v ([G]v & [R]av)\ + using u'_eq_a "rule=E" by fast + AOT_hence \\v ([G]v & [R]av & \t ([G]t & [R]at \ t =\<^sub>E v))\ + using "equi:1"[THEN "\E"(1)] by fast + then AOT_obtain s where + s_prop: \[G]s & [R]as & \t ([G]t & [R]at \ t =\<^sub>E s)\ + using "Ordinary.\E"[rotated] by meson + AOT_have \v' =\<^sub>E s\ + using s_prop[THEN "&E"(2), THEN "Ordinary.\E"] + gv'_R1u'v'[THEN "&E"(1)] 0[THEN "&E"(2)] + by (metis "&I" "vdash-properties:10") + moreover AOT_have \v =\<^sub>E s\ + using s_prop[THEN "&E"(2), THEN "Ordinary.\E"] gv Rav + by (metis "&I" "\E") + ultimately AOT_have \v' =\<^sub>E v\ + by (metis "&I" "ord=Eequiv:2" "ord=Eequiv:3" "\E") + moreover AOT_have \\(v' =\<^sub>E v)\ + using 0[THEN "&E"(1), THEN "&E"(2)] + by (metis "\E"(1) "thm-neg=E") + ultimately AOT_have \v' =\<^sub>E b\ + by (metis "raa-cor:3") + } + moreover { + AOT_assume \a =\<^sub>E u & v' =\<^sub>E v\ + AOT_hence \v' =\<^sub>E b\ + by (metis "&E"(1) not_a_eq_u "reductio-aa:1") + } + ultimately AOT_show \v' =\<^sub>E b\ + by (metis "&E"(2) "\E"(3) "reductio-aa:1") + qed + } + ultimately AOT_show \\!v ([G]v & [R\<^sub>1]u'v)\ + by (metis "raa-cor:1") + next + fix v' + AOT_assume gv': \[G]v'\ + { + AOT_assume not_v'_eq_v: \\(v' =\<^sub>E v)\ + and not_v'_eq_b: \\(v' =\<^sub>E b)\ + AOT_hence v'_noteq_v: \v' \\<^sub>E v\ + and v'_noteq_b: \v' \\<^sub>E b\ + by (metis "\E"(2) "thm-neg=E")+ + AOT_have \\!u ([F]u & [R]uv')\ + using B[THEN "Ordinary.\E", THEN "\E", OF gv']. + AOT_hence \\u ([F]u & [R]uv' & \t ([F]t & [R]tv' \ t =\<^sub>E u))\ + using "equi:1"[THEN "\E"(1)] by simp + then AOT_obtain u' where + u'_prop: \[F]u' & [R]u'v' & \t ([F]t & [R]tv' \ t =\<^sub>E u')\ + using "Ordinary.\E"[rotated] by meson + AOT_hence fu': \[F]u'\ and Ru'v': \[R]u'v'\ + using "&E" by blast+ + AOT_have not_u'_eq_u: \\u' =\<^sub>E u\ + proof (rule "raa-cor:2") + AOT_assume \u' =\<^sub>E u\ + AOT_hence \u' = u\ + by (metis "=E-simple:2" "\E") + AOT_hence Ruv': \[R]uv'\ + using "rule=E" Ru'v' by fast + AOT_have \v' =\<^sub>E b\ + using b_unique[OF Ordinary.\, OF gv', OF Ruv']. + AOT_thus \v' =\<^sub>E b & \v' =\<^sub>E b\ + using not_v'_eq_b "&I" by blast + qed + AOT_hence u'_noteq_u: \u' \\<^sub>E u\ + using "\E"(2) "thm-neg=E" by blast + AOT_have \\t ([F]t & [R]tv' \ t =\<^sub>E u')\ + using u'_prop "&E" by blast + AOT_hence \[F]t & [R]tv' \ t =\<^sub>E u'\ for t + using "Ordinary.\E" by meson + AOT_hence u'_unique: \t =\<^sub>E u'\ if \[F]t\ and \[R]tv'\ for t + by (metis "&I" that "\E") + + AOT_have \[F]u' & [R\<^sub>1]u'v' & \t ([F]t & [R\<^sub>1]tv' \ t =\<^sub>E u')\ + proof (safe intro!: "&I" gv' R\<^sub>1xy Ru'v' u'_noteq_u Ordinary.GEN "\I" + "thm-neg=E"[THEN "\E"(2)] not_v'_eq_v fu') + fix t + AOT_assume 1: \[F]t & [R\<^sub>1]tv'\ + AOT_have \[R]tv'\ + using Rxy2[OF 1[THEN "&E"(2)], OF v'_noteq_v, OF v'_noteq_b]. + AOT_thus \t =\<^sub>E u'\ + using u'_unique 1[THEN "&E"(1)] by blast + qed + AOT_hence \\u ([F]u & [R\<^sub>1]uv' & \t ([F]t & [R\<^sub>1]tv' \ t =\<^sub>E u))\ + by (rule "Ordinary.\I") + AOT_hence \\!u ([F]u & [R\<^sub>1]uv')\ + by (rule "equi:1"[THEN "\E"(2)]) + } + moreover { + AOT_assume 0: \v' =\<^sub>E v\ + AOT_hence u'_eq_u: \v' = v\ + using "=E-simple:2" "\E" by blast + AOT_have \\!u ([F]u & [R\<^sub>1]uv')\ + proof (safe intro!: "equi:1"[THEN "\E"(2)] "Ordinary.\I"[where \=u] + "&I" Ordinary.GEN "\I" fu) + AOT_show \[R\<^sub>1]uv'\ + by (rule "rule=E"[rotated, OF R\<^sub>1_def[symmetric]]) + (safe intro!: "\\C"(1) "cqt:2" "&I" prod_denotesI Ordinary.\ + "\I"(2) 0 "ord=Eequiv:1"[THEN "\E"]) + next + fix u' + AOT_assume \[F]u' & [R\<^sub>1]u'v'\ + AOT_hence 0: \[R\<^sub>1]u'v\ + using "rule=E"[rotated, OF u'_eq_u] "&E"(2) by fast + AOT_have 1: \[\?R\]u'v\ + by (rule "rule=E"[rotated, OF R\<^sub>1_def]) (fact 0) + AOT_have 2: \(u' \\<^sub>E u & v \\<^sub>E v & [R]u'v) \ + (u' =\<^sub>E a & v =\<^sub>E b) \ + (u' =\<^sub>E u & v =\<^sub>E v)\ + using "\\C"(1)[OF 1, simplified] by simp + AOT_have \\v \\<^sub>E v\ + using "\E"(4) "modus-tollens:1" "ord=Eequiv:1" Ordinary.\ + "reductio-aa:2" "thm-neg=E" by blast + AOT_hence \\((u' \\<^sub>E u & v \\<^sub>E v & [R]u'v) \ (u' =\<^sub>E a & v =\<^sub>E b))\ + by (metis "&E"(1) "&E"(2) "\E"(3) not_v_eq_b "raa-cor:3") + AOT_hence \(u' =\<^sub>E u & v =\<^sub>E v)\ + using 2 by (metis "\E"(2)) + AOT_thus \u' =\<^sub>E u\ + using "&E" by blast + qed + } + moreover { + AOT_assume 0: \v' =\<^sub>E b\ + AOT_hence v'_eq_b: \v' = b\ + using "=E-simple:2" "\E" by blast + AOT_have \\!u ([F]u & [R\<^sub>1]uv')\ + proof (safe intro!: "equi:1"[THEN "\E"(2)] "\I"(2)[where \=a] "&I" + Ordinary.GEN "\I" b_prop[THEN "&E"(1)] Oa fa + b_prop[THEN "&E"(2), THEN "&E"(1), THEN "&E"(1)]) + AOT_show \[R\<^sub>1]av'\ + apply (rule "rule=E"[rotated, OF R\<^sub>1_def[symmetric]]) + apply (safe intro!: "\\C"(1) "cqt:2" "&I" prod_denotesI) + apply (rule "\I"(1); rule "\I"(2); rule "&I") + using Oa "ord=Eequiv:1" "\E" apply blast + using "0" by blast + next + fix u' + AOT_assume fu'_R1u'v': \[F]u' & [R\<^sub>1]u'v'\ + AOT_hence 0: \[R\<^sub>1]u'b\ + using v'_eq_b by (meson "rule=E" "&E"(2)) + AOT_have 1: \[\?R\]u'b\ + by (rule "rule=E"[rotated, OF R\<^sub>1_def]) (fact 0) + AOT_have \(u' \\<^sub>E u & b \\<^sub>E v & [R]u'b) \ + (u' =\<^sub>E a & b =\<^sub>E b) \ + (u' =\<^sub>E u & b =\<^sub>E v)\ + using "\\C"(1)[OF 1, simplified] by simp + moreover { + AOT_assume 0: \u' \\<^sub>E u & b \\<^sub>E v & [R]u'b\ + AOT_have \\!u ([F]u & [R]uv')\ + using B[THEN "Ordinary.\E", THEN "\E", OF gv']. + AOT_hence \\!u ([F]u & [R]ub)\ + using v'_eq_b "rule=E" by fast + AOT_hence \\u ([F]u & [R]ub & \t ([F]t & [R]tb \ t =\<^sub>E u))\ + using "equi:1"[THEN "\E"(1)] by fast + then AOT_obtain s where + s_prop: \[F]s & [R]sb & \t ([F]t & [R]tb \ t =\<^sub>E s)\ + using "Ordinary.\E"[rotated] by meson + AOT_have \u' =\<^sub>E s\ + using s_prop[THEN "&E"(2), THEN "Ordinary.\E"] + fu'_R1u'v'[THEN "&E"(1)] 0[THEN "&E"(2)] + by (metis "&I" "\E") + moreover AOT_have \u =\<^sub>E s\ + using s_prop[THEN "&E"(2), THEN "Ordinary.\E"] fu Rub + by (metis "&I" "\E") + ultimately AOT_have \u' =\<^sub>E u\ + by (metis "&I" "ord=Eequiv:2" "ord=Eequiv:3" "\E") + moreover AOT_have \\(u' =\<^sub>E u)\ + using 0[THEN "&E"(1), THEN "&E"(1)] by (metis "\E"(1) "thm-neg=E") + ultimately AOT_have \u' =\<^sub>E a\ + by (metis "raa-cor:3") + } + moreover { + AOT_assume \u' =\<^sub>E u & b =\<^sub>E v\ + AOT_hence \u' =\<^sub>E a\ + by (metis "&E"(2) not_b_eq_v "reductio-aa:1") + } + ultimately AOT_show \u' =\<^sub>E a\ + by (metis "&E"(1) "\E"(3) "reductio-aa:1") + qed + } + ultimately AOT_show \\!u ([F]u & [R\<^sub>1]uv')\ + by (metis "raa-cor:1") + qed + ultimately AOT_have \\R R |: [F]\<^sup>-\<^sup>u \<^sub>1\<^sub>-\<^sub>1\\<^sub>E [G]\<^sup>-\<^sup>v\ + using 1 by blast + } + ultimately AOT_have \\R R |: [F]\<^sup>-\<^sup>u \<^sub>1\<^sub>-\<^sub>1\\<^sub>E [G]\<^sup>-\<^sup>v\ + using R_prop by (metis "reductio-aa:2") + AOT_thus \[F]\<^sup>-\<^sup>u \\<^sub>E [G]\<^sup>-\<^sup>v\ + by (rule "equi:3"[THEN "\\<^sub>d\<^sub>fI"]) +qed + + +AOT_theorem "P'-eq": \[F]\<^sup>-\<^sup>u \\<^sub>E [G]\<^sup>-\<^sup>v & [F]u & [G]v \ F \\<^sub>E G\ +proof(safe intro!: "\I"; frule "&E"(1); drule "&E"(2); + frule "&E"(1); drule "&E"(2)) + AOT_have \[\z [\]z & z \\<^sub>E \]\\ for \ \ by "cqt:2[lambda]" + note \_minus_\I = "rule-id-df:2:b[2]"[ + where \=\(\(\, \). \[\]\<^sup>-\<^sup>\\)\, simplified, OF "F-u", simplified, OF this] + and \_minus_\E = "rule-id-df:2:a[2]"[ + where \=\(\(\, \). \[\]\<^sup>-\<^sup>\\)\, simplified, OF "F-u", simplified, OF this] + AOT_have \_minus_\_den: \[\]\<^sup>-\<^sup>\\\ for \ \ + by (rule \_minus_\I) "cqt:2[lambda]"+ + + AOT_have \_minus_\E1: \[\]\'\ + and \_minus_\E2: \\' \\<^sub>E \\ if \[[\]\<^sup>-\<^sup>\]\'\ for \ \ \' + proof - + AOT_have \[\z [\]z & z \\<^sub>E \]\'\ + using \_minus_\E that by fast + AOT_hence \[\]\' & \' \\<^sub>E \\ + by (rule "\\C"(1)) + AOT_thus \[\]\'\ and \\' \\<^sub>E \\ + using "&E" by blast+ + qed + AOT_have \_minus_\I': \[[\]\<^sup>-\<^sup>\]\'\ if \[\]\'\ and \\' \\<^sub>E \\ for \ \ \' + proof - + AOT_have \'_den: \\'\\ + by (metis "russell-axiom[exe,1].\_denotes_asm" that(1)) + AOT_have \[\z [\]z & z \\<^sub>E \]\'\ + by (safe intro!: "\\C"(1) "cqt:2" \'_den "&I" that) + AOT_thus \[[\]\<^sup>-\<^sup>\]\'\ + using \_minus_\I by fast + qed + + AOT_assume Gv: \[G]v\ + AOT_assume Fu: \[F]u\ + AOT_assume \[F]\<^sup>-\<^sup>u \\<^sub>E [G]\<^sup>-\<^sup>v\ + AOT_hence \\R R |: [F]\<^sup>-\<^sup>u \<^sub>1\<^sub>-\<^sub>1\\<^sub>E [G]\<^sup>-\<^sup>v\ + using "equi:3"[THEN "\\<^sub>d\<^sub>fE"] by blast + then AOT_obtain R where R_prop: \R |: [F]\<^sup>-\<^sup>u \<^sub>1\<^sub>-\<^sub>1\\<^sub>E [G]\<^sup>-\<^sup>v\ + using "\E"[rotated] by blast + AOT_hence Fact1: \\r([[F]\<^sup>-\<^sup>u]r \ \!s ([[G]\<^sup>-\<^sup>v]s & [R]rs))\ + and Fact1': \\s([[G]\<^sup>-\<^sup>v]s \ \!r ([[F]\<^sup>-\<^sup>u]r & [R]rs))\ + using "equi:2"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast+ + AOT_have \R |: [F]\<^sup>-\<^sup>u \<^sub>1\<^sub>-\<^sub>1\\<^sub>o\<^sub>n\<^sub>t\<^sub>oE [G]\<^sup>-\<^sup>v\ + using "equi-rem-thm"[unvarify F G, OF \_minus_\_den, OF \_minus_\_den, + THEN "\E"(1), OF R_prop]. + AOT_hence \R |: [F]\<^sup>-\<^sup>u \<^sub>1\<^sub>-\<^sub>1\E [G]\<^sup>-\<^sup>v & R |: [F]\<^sup>-\<^sup>u \\<^sub>o\<^sub>n\<^sub>t\<^sub>oE [G]\<^sup>-\<^sup>v\ + using "equi-rem:4"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_hence Fact2: + \\r\s\t(([[F]\<^sup>-\<^sup>u]r & [[F]\<^sup>-\<^sup>u]s & [[G]\<^sup>-\<^sup>v]t) \ ([R]rt & [R]st \ r =\<^sub>E s))\ + using "equi-rem:2"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + + let ?R = \\[\xy ([[F]\<^sup>-\<^sup>u]x & [[G]\<^sup>-\<^sup>v]y & [R]xy) \ (x =\<^sub>E u & y =\<^sub>E v)]\\ + AOT_have R_den: \\?R\\\ by "cqt:2[lambda]" + + AOT_show \F \\<^sub>E G\ + proof(safe intro!: "equi:3"[THEN "\\<^sub>d\<^sub>fI"] "\I"(1)[where \="?R"] R_den + "equi:2"[THEN "\\<^sub>d\<^sub>fI"] "&I" "cqt:2" Ordinary.GEN "\I") + fix r + AOT_assume Fr: \[F]r\ + { + AOT_assume not_r_eq_u: \\(r =\<^sub>E u)\ + AOT_hence r_noteq_u: \r \\<^sub>E u\ + using "\E"(2) "thm-neg=E" by blast + AOT_have \[[F]\<^sup>-\<^sup>u]r\ + by(rule \_minus_\I; safe intro!: "\\C"(1) "cqt:2" "&I" Fr r_noteq_u) + AOT_hence \\!s ([[G]\<^sup>-\<^sup>v]s & [R]rs)\ + using Fact1[THEN "\E"(2)] "\E" Ordinary.\ by blast + AOT_hence \\s ([[G]\<^sup>-\<^sup>v]s & [R]rs & \t ([[G]\<^sup>-\<^sup>v]t & [R]rt \ t =\<^sub>E s))\ + using "equi:1"[THEN "\E"(1)] by simp + then AOT_obtain s where s_prop: \[[G]\<^sup>-\<^sup>v]s & [R]rs & \t ([[G]\<^sup>-\<^sup>v]t & [R]rt \ t =\<^sub>E s)\ + using "Ordinary.\E"[rotated] by meson + AOT_hence G_minus_v_s: \[[G]\<^sup>-\<^sup>v]s\ and Rrs: \[R]rs\ + using "&E" by blast+ + AOT_have s_unique: \t =\<^sub>E s\ if \[[G]\<^sup>-\<^sup>v]t\ and \[R]rt\ for t + using s_prop[THEN "&E"(2), THEN "Ordinary.\E", THEN "\E", OF "&I", OF that]. + AOT_have Gs: \[G]s\ + using \_minus_\E1[OF G_minus_v_s]. + AOT_have s_noteq_v: \s \\<^sub>E v\ + using \_minus_\E2[OF G_minus_v_s]. + AOT_have \\s ([G]s & [\?R\]rs & (\t ([G]t & [\?R\]rt \ t =\<^sub>E s)))\ + proof(safe intro!: "Ordinary.\I"[where \=s] "&I" Gs Ordinary.GEN "\I") + AOT_show \[\?R\]rs\ + by (auto intro!: "\\C"(1) "cqt:2" "&I" "\I"(1) \_minus_\I' Fr Gs + s_noteq_v Rrs r_noteq_u + simp: "&I" "ex:1:a" prod_denotesI "rule-ui:3") + next + fix t + AOT_assume 0: \[G]t & [\?R\]rt\ + AOT_hence \([[F]\<^sup>-\<^sup>u]r & [[G]\<^sup>-\<^sup>v]t & [R]rt) \ (r =\<^sub>E u & t =\<^sub>E v)\ + using "\\C"(1)[OF 0[THEN "&E"(2)], simplified] by blast + AOT_hence 1: \[[F]\<^sup>-\<^sup>u]r & [[G]\<^sup>-\<^sup>v]t & [R]rt\ + using not_r_eq_u by (metis "&E"(1) "\E"(3) "reductio-aa:1") + AOT_show \t =\<^sub>E s\ using s_unique 1 "&E" by blast + qed + } + moreover { + AOT_assume r_eq_u: \r =\<^sub>E u\ + AOT_have \\s ([G]s & [\?R\]rs & (\t ([G]t & [\?R\]rt \ t =\<^sub>E s)))\ + proof(safe intro!: "Ordinary.\I"[where \=v] "&I" Gv Ordinary.GEN "\I") + AOT_show \[\?R\]rv\ + by (auto intro!: "\\C"(1) "cqt:2" "&I" "\I"(2) \_minus_\I' Fr r_eq_u + "ord=Eequiv:1"[THEN "\E"] Ordinary.\ + simp: "&I" "ex:1:a" prod_denotesI "rule-ui:3") + next + fix t + AOT_assume 0: \[G]t & [\?R\]rt\ + AOT_hence \([[F]\<^sup>-\<^sup>u]r & [[G]\<^sup>-\<^sup>v]t & [R]rt) \ (r =\<^sub>E u & t =\<^sub>E v)\ + using "\\C"(1)[OF 0[THEN "&E"(2)], simplified] by blast + AOT_hence \r =\<^sub>E u & t =\<^sub>E v\ + using r_eq_u \_minus_\E2 + by (metis "&E"(1) "\E"(2) "\E"(1) "reductio-aa:1" "thm-neg=E") + AOT_thus \t =\<^sub>E v\ using "&E" by blast + qed + } + ultimately AOT_show \\!s ([G]s & [\?R\]rs)\ + using "reductio-aa:2" "equi:1"[THEN "\E"(2)] by fast + next + fix s + AOT_assume Gs: \[G]s\ + + { + AOT_assume not_s_eq_v: \\(s =\<^sub>E v)\ + AOT_hence s_noteq_v: \s \\<^sub>E v\ + using "\E"(2) "thm-neg=E" by blast + AOT_have \[[G]\<^sup>-\<^sup>v]s\ + by (rule \_minus_\I; auto intro!: "\\C"(1) "cqt:2" "&I" Gs s_noteq_v) + AOT_hence \\!r ([[F]\<^sup>-\<^sup>u]r & [R]rs)\ + using Fact1'[THEN "Ordinary.\E"] "\E" by blast + AOT_hence \\r ([[F]\<^sup>-\<^sup>u]r & [R]rs & \t ([[F]\<^sup>-\<^sup>u]t & [R]ts \ t =\<^sub>E r))\ + using "equi:1"[THEN "\E"(1)] by simp + then AOT_obtain r where + r_prop: \[[F]\<^sup>-\<^sup>u]r & [R]rs & \t ([[F]\<^sup>-\<^sup>u]t & [R]ts \ t =\<^sub>E r)\ + using "Ordinary.\E"[rotated] by meson + AOT_hence F_minus_u_r: \[[F]\<^sup>-\<^sup>u]r\ and Rrs: \[R]rs\ + using "&E" by blast+ + AOT_have r_unique: \t =\<^sub>E r\ if \[[F]\<^sup>-\<^sup>u]t\ and \[R]ts\ for t + using r_prop[THEN "&E"(2), THEN "Ordinary.\E", + THEN "\E", OF "&I", OF that]. + AOT_have Fr: \[F]r\ + using \_minus_\E1[OF F_minus_u_r]. + AOT_have r_noteq_u: \r \\<^sub>E u\ + using \_minus_\E2[OF F_minus_u_r]. + AOT_have \\r ([F]r & [\?R\]rs & (\t ([F]t & [\?R\]ts \ t =\<^sub>E r)))\ + proof(safe intro!: "Ordinary.\I"[where \=r] "&I" Fr Ordinary.GEN "\I") + AOT_show \[\?R\]rs\ + by (auto intro!: "\\C"(1) "cqt:2" "&I" "\I"(1) \_minus_\I' Fr + Gs s_noteq_v Rrs r_noteq_u + simp: "&I" "ex:1:a" prod_denotesI "rule-ui:3") + next + fix t + AOT_assume 0: \[F]t & [\?R\]ts\ + AOT_hence \([[F]\<^sup>-\<^sup>u]t & [[G]\<^sup>-\<^sup>v]s & [R]ts) \ (t =\<^sub>E u & s =\<^sub>E v)\ + using "\\C"(1)[OF 0[THEN "&E"(2)], simplified] by blast + AOT_hence 1: \[[F]\<^sup>-\<^sup>u]t & [[G]\<^sup>-\<^sup>v]s & [R]ts\ + using not_s_eq_v by (metis "&E"(2) "\E"(3) "reductio-aa:1") + AOT_show \t =\<^sub>E r\ using r_unique 1 "&E" by blast + qed + } + moreover { + AOT_assume s_eq_v: \s =\<^sub>E v\ + AOT_have \\r ([F]r & [\?R\]rs & (\t ([F]t & [\?R\]ts \ t =\<^sub>E r)))\ + proof(safe intro!: "Ordinary.\I"[where \=u] "&I" Fu Ordinary.GEN "\I") + AOT_show \[\?R\]us\ + by (auto intro!: "\\C"(1) "cqt:2" "&I" prod_denotesI "\I"(2) + \_minus_\I' Gs s_eq_v Ordinary.\ + "ord=Eequiv:1"[THEN "\E"]) + next + fix t + AOT_assume 0: \[F]t & [\?R\]ts\ + AOT_hence 1: \([[F]\<^sup>-\<^sup>u]t & [[G]\<^sup>-\<^sup>v]s & [R]ts) \ (t =\<^sub>E u & s =\<^sub>E v)\ + using "\\C"(1)[OF 0[THEN "&E"(2)], simplified] by blast + moreover AOT_have \\([[F]\<^sup>-\<^sup>u]t & [[G]\<^sup>-\<^sup>v]s & [R]ts)\ + proof (rule "raa-cor:2") + AOT_assume \([[F]\<^sup>-\<^sup>u]t & [[G]\<^sup>-\<^sup>v]s & [R]ts)\ + AOT_hence \[[G]\<^sup>-\<^sup>v]s\ using "&E" by blast + AOT_thus \s =\<^sub>E v & \(s =\<^sub>E v)\ + by (metis \_minus_\E2 "\E"(4) "reductio-aa:1" s_eq_v "thm-neg=E") + qed + ultimately AOT_have \t =\<^sub>E u & s =\<^sub>E v\ + by (metis "\E"(2)) + AOT_thus \t =\<^sub>E u\ using "&E" by blast + qed + } + ultimately AOT_show \\!r ([F]r & [\?R\]rs)\ + using "\E"(2) "equi:1" "reductio-aa:2" by fast + qed +qed + + +AOT_theorem "approx-cont:1": \\F\G \(F \\<^sub>E G & \\F \\<^sub>E G)\ +proof - + let ?P = \\[\x E!x & \\<^bold>\E!x]\\ + AOT_have \\q\<^sub>0 & \\q\<^sub>0\ by (metis q\<^sub>0_prop) + AOT_hence 1: \\\x(E!x & \\<^bold>\E!x) & \\\x(E!x & \\<^bold>\E!x)\ + by (rule q\<^sub>0_def[THEN "=\<^sub>d\<^sub>fE"(2), rotated]) + (simp add: "log-prop-prop:2") + AOT_have \: \\\x [\?P\]x & \\\x [\?P\]x\ + apply (AOT_subst \[\?P\]x\ \E!x & \\<^bold>\E!x\ for: x) + apply (rule "beta-C-meta"[THEN "\E"]; "cqt:2[lambda]") + by (fact 1) + show ?thesis + proof (rule "\I"(1))+ + AOT_have \\[L]\<^sup>- \\<^sub>E [\?P\] & \\[L]\<^sup>- \\<^sub>E [\?P\]\ + proof (rule "&I"; rule "RM\"[THEN "\E"]; (rule "\I")?) + AOT_modally_strict { + AOT_assume A: \\\x [\?P\]x\ + AOT_show \[L]\<^sup>- \\<^sub>E [\?P\]\ + proof (safe intro!: "empty-approx:1"[unvarify F H, THEN "\E"] + "rel-neg-T:3" "&I") + AOT_show \[\?P\]\\ by "cqt:2[lambda]" + next + AOT_show \\\u [L\<^sup>-]u\ + proof (rule "raa-cor:2") + AOT_assume \\u [L\<^sup>-]u\ + then AOT_obtain u where \[L\<^sup>-]u\ + using "Ordinary.\E"[rotated] by blast + moreover AOT_have \\[L\<^sup>-]u\ + using "thm-noncont-e-e:2"[THEN "contingent-properties:2"[THEN "\\<^sub>d\<^sub>fE"], + THEN "&E"(2)] + by (metis "qml:2"[axiom_inst] "rule-ui:3" "\E") + ultimately AOT_show \p & \p\ for p + by (metis "raa-cor:3") + qed + next + AOT_show \\\v [\?P\]v\ + proof (rule "raa-cor:2") + AOT_assume \\v [\?P\]v\ + then AOT_obtain u where \[\?P\]u\ + using "Ordinary.\E"[rotated] by blast + AOT_hence \[\?P\]u\ + using "&E" by blast + AOT_hence \\x [\?P\]x\ + by (rule "\I") + AOT_thus \\x [\?P\]x & \\x [\?P\]x\ + using A "&I" by blast + qed + qed + } + next + AOT_show \\\\x [\?P\]x\ + using \ "&E" by blast + next + AOT_modally_strict { + AOT_assume A: \\x [\?P\]x\ + AOT_have B: \\[\?P\] \\<^sub>E [L]\<^sup>-\ + proof (safe intro!: "empty-approx:2"[unvarify F H, THEN "\E"] + "rel-neg-T:3" "&I") + AOT_show \[\?P\]\\ + by "cqt:2[lambda]" + next + AOT_obtain x where Px: \[\?P\]x\ + using A "\E" by blast + AOT_hence \E!x & \\<^bold>\E!x\ + by (rule "\\C"(1)) + AOT_hence 1: \\E!x\ + by (metis "T\" "&E"(1) "vdash-properties:10") + AOT_have \[\x \E!x]x\ + by (auto intro!: "\\C"(1) "cqt:2" 1) + AOT_hence \O!x\ + by (rule AOT_ordinary[THEN "=\<^sub>d\<^sub>fI"(2), rotated]) "cqt:2[lambda]" + AOT_hence \O!x & [\?P\]x\ + using Px "&I" by blast + AOT_thus \\u [\?P\]u\ + by (rule "\I") + next + AOT_show \\\u [L\<^sup>-]u\ + proof (rule "raa-cor:2") + AOT_assume \\u [L\<^sup>-]u\ + then AOT_obtain u where \[L\<^sup>-]u\ + using "Ordinary.\E"[rotated] by blast + moreover AOT_have \\[L\<^sup>-]u\ + using "thm-noncont-e-e:2"[THEN "contingent-properties:2"[THEN "\\<^sub>d\<^sub>fE"]] + by (metis "qml:2"[axiom_inst] "rule-ui:3" "\E" "&E"(2)) + ultimately AOT_show \p & \p\ for p + by (metis "raa-cor:3") + qed + qed + AOT_show \\[L]\<^sup>- \\<^sub>E [\?P\]\ + proof (rule "raa-cor:2") + AOT_assume \[L]\<^sup>- \\<^sub>E [\?P\]\ + AOT_hence \[\?P\] \\<^sub>E [L]\<^sup>-\ + apply (rule "eq-part:2"[unvarify F G, THEN "\E", rotated 2]) + apply "cqt:2[lambda]" + by (simp add: "rel-neg-T:3") + AOT_thus \[\?P\] \\<^sub>E [L]\<^sup>- & \[\?P\] \\<^sub>E [L]\<^sup>-\ + using B "&I" by blast + qed + } + next + AOT_show \\\x [\?P\]x\ + using \ "&E" by blast + qed + AOT_thus \\([L]\<^sup>- \\<^sub>E [\?P\] & \\[L]\<^sup>- \\<^sub>E [\?P\])\ + using "S5Basic:11" "\E"(2) by blast + next + AOT_show \[\x [E!]x & \\<^bold>\[E!]x]\\ + by "cqt:2" + next + AOT_show \[L]\<^sup>-\\ + by (simp add: "rel-neg-T:3") + qed +qed + + +AOT_theorem "approx-cont:2": + \\F\G \([\z \<^bold>\[F]z] \\<^sub>E G & \\[\z \<^bold>\[F]z] \\<^sub>E G)\ +proof - + let ?P = \\[\x E!x & \\<^bold>\E!x]\\ + AOT_have \\q\<^sub>0 & \\q\<^sub>0\ by (metis q\<^sub>0_prop) + AOT_hence 1: \\\x(E!x & \\<^bold>\E!x) & \\\x(E!x & \\<^bold>\E!x)\ + by (rule q\<^sub>0_def[THEN "=\<^sub>d\<^sub>fE"(2), rotated]) + (simp add: "log-prop-prop:2") + AOT_have \: \\\x [\?P\]x & \\\x [\?P\]x\ + apply (AOT_subst \[\?P\]x\ \E!x & \\<^bold>\E!x\ for: x) + apply (rule "beta-C-meta"[THEN "\E"]; "cqt:2") + by (fact 1) + show ?thesis + proof (rule "\I"(1))+ + AOT_have \\[\z \<^bold>\[L\<^sup>-]z] \\<^sub>E [\?P\] & \\[\z \<^bold>\[L\<^sup>-]z] \\<^sub>E [\?P\]\ + proof (rule "&I"; rule "RM\"[THEN "\E"]; (rule "\I")?) + AOT_modally_strict { + AOT_assume A: \\\x [\?P\]x\ + AOT_show \[\z \<^bold>\[L\<^sup>-]z] \\<^sub>E [\?P\]\ + proof (safe intro!: "empty-approx:1"[unvarify F H, THEN "\E"] + "rel-neg-T:3" "&I") + AOT_show \[\?P\]\\ by "cqt:2" + next + AOT_show \\\u [\z \<^bold>\[L\<^sup>-]z]u\ + proof (rule "raa-cor:2") + AOT_assume \\u [\z \<^bold>\[L\<^sup>-]z]u\ + then AOT_obtain u where \[\z \<^bold>\[L\<^sup>-]z]u\ + using "Ordinary.\E"[rotated] by blast + AOT_hence \\<^bold>\[L\<^sup>-]u\ + using "\\C"(1) "&E" by blast + moreover AOT_have \\\[L\<^sup>-]u\ + using "thm-noncont-e-e:2"[THEN "contingent-properties:2"[THEN "\\<^sub>d\<^sub>fE"]] + by (metis RN "qml:2"[axiom_inst] "rule-ui:3" "\E" "&E"(2)) + ultimately AOT_show \p & \p\ for p + by (metis "Act-Sub:3" "KBasic2:1" "\E"(1) "raa-cor:3" "\E") + qed + next + AOT_show \\\v [\?P\]v\ + proof (rule "raa-cor:2") + AOT_assume \\v [\?P\]v\ + then AOT_obtain u where \[\?P\]u\ + using "Ordinary.\E"[rotated] by blast + AOT_hence \[\?P\]u\ + using "&E" by blast + AOT_hence \\x [\?P\]x\ + by (rule "\I") + AOT_thus \\x [\?P\]x & \\x [\?P\]x\ + using A "&I" by blast + qed + next + AOT_show \[\z \<^bold>\[L\<^sup>-]z]\\ by "cqt:2" + qed + } + next + AOT_show \\\\x [\?P\]x\ using \ "&E" by blast + next + AOT_modally_strict { + AOT_assume A: \\x [\?P\]x\ + AOT_have B: \\[\?P\] \\<^sub>E [\z \<^bold>\[L\<^sup>-]z]\ + proof (safe intro!: "empty-approx:2"[unvarify F H, THEN "\E"] + "rel-neg-T:3" "&I") + AOT_show \[\?P\]\\ by "cqt:2" + next + AOT_obtain x where Px: \[\?P\]x\ + using A "\E" by blast + AOT_hence \E!x & \\<^bold>\E!x\ + by (rule "\\C"(1)) + AOT_hence \\E!x\ + by (metis "T\" "&E"(1) "\E") + AOT_hence \[\x \E!x]x\ + by (auto intro!: "\\C"(1) "cqt:2") + AOT_hence \O!x\ + by (rule AOT_ordinary[THEN "=\<^sub>d\<^sub>fI"(2), rotated]) "cqt:2" + AOT_hence \O!x & [\?P\]x\ + using Px "&I" by blast + AOT_thus \\u [\?P\]u\ + by (rule "\I") + next + AOT_show \\\u [\z \<^bold>\[L\<^sup>-]z]u\ + proof (rule "raa-cor:2") + AOT_assume \\u [\z \<^bold>\[L\<^sup>-]z]u\ + then AOT_obtain u where \[\z \<^bold>\[L\<^sup>-]z]u\ + using "Ordinary.\E"[rotated] by blast + AOT_hence \\<^bold>\[L\<^sup>-]u\ + using "\\C"(1) "&E" by blast + moreover AOT_have \\\[L\<^sup>-]u\ + using "thm-noncont-e-e:2"[THEN "contingent-properties:2"[THEN "\\<^sub>d\<^sub>fE"]] + by (metis RN "qml:2"[axiom_inst] "rule-ui:3" "\E" "&E"(2)) + ultimately AOT_show \p & \p\ for p + by (metis "Act-Sub:3" "KBasic2:1" "\E"(1) "raa-cor:3" "\E") + qed + next + AOT_show \[\z \<^bold>\[L\<^sup>-]z]\\ by "cqt:2" + qed + AOT_show \\[\z \<^bold>\[L\<^sup>-]z] \\<^sub>E [\?P\]\ + proof (rule "raa-cor:2") + AOT_assume \[\z \<^bold>\[L\<^sup>-]z] \\<^sub>E [\?P\]\ + AOT_hence \[\?P\] \\<^sub>E [\z \<^bold>\[L\<^sup>-]z]\ + by (rule "eq-part:2"[unvarify F G, THEN "\E", rotated 2]) + "cqt:2"+ + AOT_thus \[\?P\] \\<^sub>E [\z \<^bold>\[L\<^sup>-]z] & \[\?P\] \\<^sub>E [\z \<^bold>\[L\<^sup>-]z]\ + using B "&I" by blast + qed + } + next + AOT_show \\\x [\?P\]x\ + using \ "&E" by blast + qed + AOT_thus \\([\z \<^bold>\[L\<^sup>-]z] \\<^sub>E [\?P\] & \\[\z \<^bold>\[L\<^sup>-]z] \\<^sub>E [\?P\])\ + using "S5Basic:11" "\E"(2) by blast + next + AOT_show \[\x [E!]x & \\<^bold>\[E!]x]\\ by "cqt:2" + next + AOT_show \[L]\<^sup>-\\ + by (simp add: "rel-neg-T:3") + qed +qed + +notepad +begin + text\We already have defined being equivalent on the ordinary objects in the + Extended Relation Comprehension theory.\ + AOT_have \F \\<^sub>E G \\<^sub>d\<^sub>f F\ & G\ & \u ([F]u \ [G]u)\ for F G + using eqE by blast +end + +AOT_theorem "apE-eqE:1": \F \\<^sub>E G \ F \\<^sub>E G\ +proof(rule "\I") + AOT_assume 0: \F \\<^sub>E G\ + AOT_have \\R R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E G\ + proof (safe intro!: "\I"(1)[where \="\(=\<^sub>E)\"] "equi:2"[THEN "\\<^sub>d\<^sub>fI"] "&I" + "=E[denotes]" "cqt:2[const_var]"[axiom_inst] Ordinary.GEN + "\I" "equi:1"[THEN "\E"(2)]) + fix u + AOT_assume Fu: \[F]u\ + AOT_hence Gu: \[G]u\ + using "\\<^sub>d\<^sub>fE"[OF eqE, OF 0, THEN "&E"(2), + THEN "Ordinary.\E"[where \=u], THEN "\E"(1)] + Ordinary.\ Fu by blast + AOT_show \\v ([G]v & u =\<^sub>E v & \v' ([G]v' & u =\<^sub>E v' \ v' =\<^sub>E v))\ + by (safe intro!: "Ordinary.\I"[where \=u] "&I" GEN "\I" Ordinary.\ Gu + "ord=Eequiv:1"[THEN "\E", OF Ordinary.\] + "ord=Eequiv:2"[THEN "\E"] dest!: "&E"(2)) + next + fix v + AOT_assume Gv: \[G]v\ + AOT_hence Fv: \[F]v\ + using "\\<^sub>d\<^sub>fE"[OF eqE, OF 0, THEN "&E"(2), + THEN "Ordinary.\E"[where \=v], THEN "\E"(2)] + Ordinary.\ Gv by blast + AOT_show \\u ([F]u & u =\<^sub>E v & \v' ([F]v' & v' =\<^sub>E v \ v' =\<^sub>E u))\ + by (safe intro!: "Ordinary.\I"[where \=v] "&I" GEN "\I" Ordinary.\ Fv + "ord=Eequiv:1"[THEN "\E", OF Ordinary.\] + "ord=Eequiv:2"[THEN "\E"] dest!: "&E"(2)) + qed + AOT_thus \F \\<^sub>E G\ + by (rule "equi:3"[THEN "\\<^sub>d\<^sub>fI"]) +qed + +AOT_theorem "apE-eqE:2": \(F \\<^sub>E G & G \\<^sub>E H) \ F \\<^sub>E H\ +proof(rule "\I") + AOT_assume \F \\<^sub>E G & G \\<^sub>E H\ + AOT_hence \F \\<^sub>E G\ and \G \\<^sub>E H\ + using "apE-eqE:1"[THEN "\E"] "&E" by blast+ + AOT_thus \F \\<^sub>E H\ + by (metis Adjunction "eq-part:3" "vdash-properties:10") +qed + + +AOT_act_theorem "eq-part-act:1": \[\z \<^bold>\[F]z] \\<^sub>E F\ +proof (safe intro!: eqE[THEN "\\<^sub>d\<^sub>fI"] "&I" "cqt:2" Ordinary.GEN "\I") + fix u + AOT_have \[\z \<^bold>\[F]z]u \ \<^bold>\[F]u\ + by (rule "beta-C-meta"[THEN "\E"]) "cqt:2[lambda]" + also AOT_have \\ \ [F]u\ + using "act-conj-act:4" "logic-actual"[act_axiom_inst, THEN "\E"] by blast + finally AOT_show \[\z \<^bold>\[F]z]u \ [F]u\. +qed + +AOT_act_theorem "eq-part-act:2": \[\z \<^bold>\[F]z] \\<^sub>E F\ + by (safe intro!: "apE-eqE:1"[unvarify F, THEN "\E"] "eq-part-act:1") "cqt:2" + + +AOT_theorem "actuallyF:1": \\<^bold>\(F \\<^sub>E [\z \<^bold>\[F]z])\ +proof - + AOT_have 1: \\<^bold>\([F]x \ \<^bold>\[F]x)\ for x + by (meson "Act-Basic:5" "act-conj-act:4" "\E"(2) "Commutativity of \") + AOT_have \\<^bold>\([F]x \ [\z \<^bold>\[F]z]x)\ for x + apply (AOT_subst \[\z \<^bold>\[F]z]x\ \\<^bold>\[F]x\) + apply (rule "beta-C-meta"[THEN "\E"]) + apply "cqt:2[lambda]" + by (fact 1) + AOT_hence \O!x \ \<^bold>\([F]x \ [\z \<^bold>\[F]z]x)\ for x + by (metis "\I") + AOT_hence \\u \<^bold>\([F]u \ [\z \<^bold>\[F]z]u)\ + using "\I" by fast + AOT_hence 1: \\<^bold>\\u ([F]u \ [\z \<^bold>\[F]z]u)\ + by (metis "Ordinary.res-var-bound-reas[2]" "\E") + AOT_modally_strict { + AOT_have \[\z \<^bold>\[F]z]\\ by "cqt:2" + } note 2 = this + AOT_have \\<^bold>\(F \\<^sub>E [\z \<^bold>\[F]z])\ + apply (AOT_subst \F \\<^sub>E [\z \<^bold>\[F]z]\ \\u ([F]u \ [\z \<^bold>\[F]z]u)\) + using eqE[THEN "\Df", THEN "\S"(1), OF "&I", + OF "cqt:2[const_var]"[axiom_inst], OF 2] + by (auto simp: 1) + moreover AOT_have \\<^bold>\(F \\<^sub>E [\z \<^bold>\[F]z] \ F \\<^sub>E [\z \<^bold>\[F]z])\ + using "apE-eqE:1"[unvarify G, THEN "RA[2]", OF 2] by metis + ultimately AOT_show \\<^bold>\F \\<^sub>E [\z \<^bold>\[F]z]\ + by (metis "act-cond" "\E") +qed + +AOT_theorem "actuallyF:2": \Rigid([\z \<^bold>\[F]z])\ +proof(safe intro!: GEN "\I" "df-rigid-rel:1"[THEN "\\<^sub>d\<^sub>fI"] "&I") + AOT_show \[\z \<^bold>\[F]z]\\ by "cqt:2" +next + AOT_show \\\x ([\z \<^bold>\[F]z]x \ \[\z \<^bold>\[F]z]x)\ + proof(rule RN; rule GEN; rule "\I") + AOT_modally_strict { + fix x + AOT_assume \[\z \<^bold>\[F]z]x\ + AOT_hence \\<^bold>\[F]x\ + by (rule "\\C"(1)) + AOT_hence 1: \\\<^bold>\[F]x\ by (metis "Act-Basic:6" "\E"(1)) + AOT_show \\[\z \<^bold>\[F]z]x\ + apply (AOT_subst \[\z \<^bold>\[F]z]x\ \\<^bold>\[F]x\) + apply (rule "beta-C-meta"[THEN "\E"]) + apply "cqt:2[lambda]" + by (fact 1) + } + qed +qed + +AOT_theorem "approx-nec:1": \Rigid(F) \ F \\<^sub>E [\z \<^bold>\[F]z]\ +proof(rule "\I") + AOT_assume \Rigid([F])\ + AOT_hence A: \\\x ([F]x \ \[F]x)\ + using "df-rigid-rel:1"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2)] by blast + AOT_hence 0: \\x \([F]x \ \[F]x)\ + using CBF[THEN "\E"] by blast + AOT_hence 1: \\x ([F]x \ \[F]x)\ + using A "qml:2"[axiom_inst, THEN "\E"] by blast + AOT_have act_F_den: \[\z \<^bold>\[F]z]\\ + by "cqt:2" + AOT_show \F \\<^sub>E [\z \<^bold>\[F]z]\ + proof (safe intro!: "apE-eqE:1"[unvarify G, THEN "\E"] eqE[THEN "\\<^sub>d\<^sub>fI"] "&I" + "cqt:2" act_F_den Ordinary.GEN "\I" "\I") + fix u + AOT_assume \[F]u\ + AOT_hence \\[F]u\ + using 1[THEN "\E"(2), THEN "\E"] by blast + AOT_hence act_F_u: \\<^bold>\[F]u\ + by (metis "nec-imp-act" "\E") + AOT_show \[\z \<^bold>\[F]z]u\ + by (auto intro!: "\\C"(1) "cqt:2" act_F_u) + next + fix u + AOT_assume \[\z \<^bold>\[F]z]u\ + AOT_hence \\<^bold>\[F]u\ + by (rule "\\C"(1)) + AOT_thus \[F]u\ + using 0[THEN "\E"(2)] + by (metis "\E"(1) "sc-eq-fur:2" "\E") + qed +qed + + +AOT_theorem "approx-nec:2": + \F \\<^sub>E G \ \H ([\z \<^bold>\[H]z] \\<^sub>E F \ [\z \<^bold>\[H]z] \\<^sub>E G)\ +proof(rule "\I"; rule "\I") + AOT_assume 0: \F \\<^sub>E G\ + AOT_assume 0: \F \\<^sub>E G\ + AOT_hence \\H (H \\<^sub>E F \ H \\<^sub>E G)\ + using "eq-part:4"[THEN "\E"(1), OF 0] by blast + AOT_have \[\z \<^bold>\[H]z] \\<^sub>E F \ [\z \<^bold>\[H]z] \\<^sub>E G\ for H + by (rule "\E"(1)[OF "eq-part:4"[THEN "\E"(1), OF 0]]) "cqt:2" + AOT_thus \\H ([\z \<^bold>\[H]z] \\<^sub>E F \ [\z \<^bold>\[H]z] \\<^sub>E G)\ + by (rule GEN) +next + AOT_assume 0: \\H ([\z \<^bold>\[H]z] \\<^sub>E F \ [\z \<^bold>\[H]z] \\<^sub>E G)\ + AOT_obtain H where \Rigidifies(H,F)\ + using "rigid-der:3" "\E" by metis + AOT_hence H: \Rigid(H) & \x ([H]x \ [F]x)\ + using "df-rigid-rel:2"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_have H_rigid: \\\x ([H]x \ \[H]x)\ + using H[THEN "&E"(1), THEN "df-rigid-rel:1"[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(2)]. + AOT_hence \\x \([H]x \ \[H]x)\ + using "CBF" "vdash-properties:10" by blast + AOT_hence \\([H]x \ \[H]x)\ for x using "\E"(2) by blast + AOT_hence rigid: \[H]x \ \<^bold>\[H]x\ for x + by (metis "\E"(6) "oth-class-taut:3:a" "sc-eq-fur:2" "\E") + AOT_have \H \\<^sub>E F\ + proof (safe intro!: eqE[THEN "\\<^sub>d\<^sub>fI"] "&I" "cqt:2" Ordinary.GEN "\I") + AOT_show \[H]u \ [F]u\ for u using H[THEN "&E"(2)] "\E"(2) by fast + qed + AOT_hence \H \\<^sub>E F\ + by (rule "apE-eqE:2"[THEN "\E", OF "&I", rotated]) + (simp add: "eq-part:1") + AOT_hence F_approx_H: \F \\<^sub>E H\ + by (metis "eq-part:2" "\E") + moreover AOT_have H_eq_act_H: \H \\<^sub>E [\z \<^bold>\[H]z]\ + proof (safe intro!: eqE[THEN "\\<^sub>d\<^sub>fI"] "&I" "cqt:2" Ordinary.GEN "\I") + AOT_show \[H]u \ [\z \<^bold>\[H]z]u\ for u + apply (AOT_subst \[\z \<^bold>\[H]z]u\ \\<^bold>\[H]u\) + apply (rule "beta-C-meta"[THEN "\E"]) + apply "cqt:2[lambda]" + using rigid by blast + qed + AOT_have a: \F \\<^sub>E [\z \<^bold>\[H]z]\ + apply (rule "apE-eqE:2"[unvarify H, THEN "\E"]) + apply "cqt:2[lambda]" + using F_approx_H H_eq_act_H "&I" by blast + AOT_hence \[\z \<^bold>\[H]z] \\<^sub>E F\ + apply (rule "eq-part:2"[unvarify G, THEN "\E", rotated]) + by "cqt:2[lambda]" + AOT_hence b: \[\z \<^bold>\[H]z] \\<^sub>E G\ + by (rule 0[THEN "\E"(1), THEN "\E"(1), rotated]) "cqt:2" + AOT_show \F \\<^sub>E G\ + by (rule "eq-part:3"[unvarify G, THEN "\E", rotated, OF "&I", OF a, OF b]) + "cqt:2" +qed + +AOT_theorem "approx-nec:3": + \(Rigid(F) & Rigid(G)) \ \(F \\<^sub>E G \ \F \\<^sub>E G)\ +proof (rule "\I") + AOT_assume \Rigid(F) & Rigid(G)\ + AOT_hence \\\x([F]x \ \[F]x)\ and \\\x([G]x \ \[G]x)\ + using "df-rigid-rel:1"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2)] "&E" by blast+ + AOT_hence \\(\\x([F]x \ \[F]x) & \\x([G]x \ \[G]x))\ + using "KBasic:3" "4" "&I" "\E"(2) "vdash-properties:10" by meson + moreover AOT_have \\(\\x([F]x \ \[F]x) & \\x([G]x \ \[G]x)) \ + \(F \\<^sub>E G \ \F \\<^sub>E G)\ + proof(rule RM; rule "\I"; rule "\I") + AOT_modally_strict { + AOT_assume \\\x([F]x \ \[F]x) & \\x([G]x \ \[G]x)\ + AOT_hence \\\x([F]x \ \[F]x)\ and \\\x([G]x \ \[G]x)\ + using "&E" by blast+ + AOT_hence \\x\([F]x \ \[F]x)\ and \\x\([G]x \ \[G]x)\ + using CBF[THEN "\E"] by blast+ + AOT_hence F_nec: \\([F]x \ \[F]x)\ + and G_nec: \\([G]x \ \[G]x)\ for x + using "\E"(2) by blast+ + AOT_assume \F \\<^sub>E G\ + AOT_hence \\R R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E G\ + by (metis "\\<^sub>d\<^sub>fE" "equi:3") + then AOT_obtain R where \R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E G\ + using "\E"[rotated] by blast + AOT_hence C1: \\u ([F]u \ \!v ([G]v & [R]uv))\ + and C2: \\v ([G]v \ \!u ([F]u & [R]uv))\ + using "equi:2"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast+ + AOT_obtain R' where \Rigidifies(R', R)\ + using "rigid-der:3" "\E"[rotated] by blast + AOT_hence 1: \Rigid(R') & \x\<^sub>1...\x\<^sub>n ([R']x\<^sub>1...x\<^sub>n \ [R]x\<^sub>1...x\<^sub>n)\ + using "df-rigid-rel:2"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_hence \\\x\<^sub>1...\x\<^sub>n ([R']x\<^sub>1...x\<^sub>n \ \[R']x\<^sub>1...x\<^sub>n)\ + using "df-rigid-rel:1"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + AOT_hence \\x\<^sub>1...\x\<^sub>n (\[R']x\<^sub>1...x\<^sub>n \ \[R']x\<^sub>1...x\<^sub>n)\ + using "\E"(1) "rigid-rel-thms:1" by blast + AOT_hence D: \\x\<^sub>1\x\<^sub>2 (\[R']x\<^sub>1x\<^sub>2 \ \[R']x\<^sub>1x\<^sub>2)\ + using tuple_forall[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_have E: \\x\<^sub>1\x\<^sub>2 ([R']x\<^sub>1x\<^sub>2 \ [R]x\<^sub>1x\<^sub>2)\ + using tuple_forall[THEN "\\<^sub>d\<^sub>fE", OF 1[THEN "&E"(2)]] by blast + AOT_have \\u \([F]u \ \!v ([G]v & [R']uv))\ + and \\v \([G]v \ \!u ([F]u & [R']uv))\ + proof (safe intro!: Ordinary.GEN "\I") + fix u + AOT_show \\([F]u \ \!v ([G]v & [R']uv))\ + proof (rule "raa-cor:1") + AOT_assume \\\([F]u \ \!v ([G]v & [R']uv))\ + AOT_hence 1: \\\([F]u \ \!v ([G]v & [R']uv))\ + using "KBasic:11" "\E"(1) by blast + AOT_have \\([F]u & \\!v ([G]v & [R']uv))\ + apply (AOT_subst \[F]u & \\!v ([G]v & [R']uv)\ + \\([F]u \ \!v ([G]v & [R']uv))\) + apply (meson "\E"(6) "oth-class-taut:1:b" "oth-class-taut:3:a") + by (fact 1) + AOT_hence A: \\[F]u & \\\!v ([G]v & [R']uv)\ + using "KBasic2:3" "\E" by blast + AOT_hence \\[F]u\ + using F_nec "&E"(1) "\E"(1) "sc-eq-box-box:1" "\E" by blast + AOT_hence \[F]u\ + by (metis "qml:2"[axiom_inst] "\E") + AOT_hence \\!v ([G]v & [R]uv)\ + using C1[THEN "Ordinary.\E", THEN "\E"] by blast + AOT_hence \\v ([G]v & [R]uv & \v' ([G]v' & [R]uv' \ v' =\<^sub>E v))\ + using "equi:1"[THEN "\E"(1)] by auto + then AOT_obtain a where + a_prop: \O!a & ([G]a & [R]ua & \v' ([G]v' & [R]uv' \ v' =\<^sub>E a))\ + using "\E"[rotated] by blast + AOT_have \\v \([G]v & [R']uv & \v' ([G]v' & [R']uv' \ v' =\<^sub>E v))\ + proof(safe intro!: "\I"(2)[where \=a] "&I" a_prop[THEN "&E"(1)] + "KBasic:3"[THEN "\E"(2)]) + AOT_show \\[G]a\ + using a_prop[THEN "&E"(2), THEN "&E"(1), THEN "&E"(1)] + by (metis G_nec "qml:2"[axiom_inst] "\E") + next + AOT_show \\[R']ua\ + using D[THEN "\E"(2), THEN "\E"(2), THEN "\E"] + E[THEN "\E"(2), THEN "\E"(2), THEN "\E"(2), + OF a_prop[THEN "&E"(2), THEN "&E"(1), THEN "&E"(2)]] + by (metis "T\" "\E") + next + AOT_have \\v' \([G]v' & [R']uv' \ v' =\<^sub>E a)\ + proof (rule Ordinary.GEN; rule "raa-cor:1") + fix v' + AOT_assume \\\([G]v' & [R']uv' \ v' =\<^sub>E a)\ + AOT_hence \\\([G]v' & [R']uv' \ v' =\<^sub>E a)\ + by (metis "KBasic:11" "\E"(1)) + AOT_hence \\([G]v' & [R']uv' & \v' =\<^sub>E a)\ + by (AOT_subst \[G]v' & [R']uv' & \v' =\<^sub>E a\ + \\([G]v' & [R']uv' \ v' =\<^sub>E a)\) + (meson "\E"(6) "oth-class-taut:1:b" "oth-class-taut:3:a") + AOT_hence 1: \\[G]v'\ and 2: \\[R']uv'\ and 3: \\\v' =\<^sub>E a\ + using "KBasic2:3"[THEN "\E", THEN "&E"(1)] + "KBasic2:3"[THEN "\E", THEN "&E"(2)] by blast+ + AOT_have Gv': \[G]v'\ using G_nec 1 + by (meson "B\" "KBasic:13" "\E") + AOT_have \\[R']uv'\ + using 2 D[THEN "\E"(2), THEN "\E"(2), THEN "\E"] by blast + AOT_hence R'uv': \[R']uv'\ + by (metis "B\" "T\" "\E") + AOT_hence \[R]uv'\ + using E[THEN "\E"(2), THEN "\E"(2), THEN "\E"(1)] by blast + AOT_hence \v' =\<^sub>E a\ + using a_prop[THEN "&E"(2), THEN "&E"(2), THEN "Ordinary.\E", + THEN "\E", OF "&I", OF Gv'] by blast + AOT_hence \\(v' =\<^sub>E a)\ + by (metis "id-nec3:1" "\E"(4) "raa-cor:3") + moreover AOT_have \\\(v' =\<^sub>E a)\ + using 3 "KBasic:11" "\E"(2) by blast + ultimately AOT_show \\(v' =\<^sub>E a) & \\(v' =\<^sub>E a)\ + using "&I" by blast + qed + AOT_thus \\\v'([G]v' & [R']uv' \ v' =\<^sub>E a)\ + using "Ordinary.res-var-bound-reas[BF]" "\E" by fast + qed + AOT_hence \\\v ([G]v & [R']uv & \v' ([G]v' & [R']uv' \ v' =\<^sub>E v))\ + using "Ordinary.res-var-bound-reas[Buridan]" "\E" by fast + AOT_hence \\\!v ([G]v & [R']uv)\ + by (AOT_subst_thm "equi:1") + moreover AOT_have \\\\!v ([G]v & [R']uv)\ + using A[THEN "&E"(2)] "KBasic:11"[THEN "\E"(2)] by blast + ultimately AOT_show \\\!v ([G]v & [R']uv) & \\\!v ([G]v & [R']uv)\ + by (rule "&I") + qed + next + fix v + AOT_show \\([G]v \ \!u ([F]u & [R']uv))\ + proof (rule "raa-cor:1") + AOT_assume \\\([G]v \ \!u ([F]u & [R']uv))\ + AOT_hence 1: \\\([G]v \ \!u ([F]u & [R']uv))\ + using "KBasic:11" "\E"(1) by blast + AOT_hence \\([G]v & \\!u ([F]u & [R']uv))\ + by (AOT_subst \[G]v & \\!u ([F]u & [R']uv)\ + \\([G]v \ \!u ([F]u & [R']uv))\) + (meson "\E"(6) "oth-class-taut:1:b" "oth-class-taut:3:a") + AOT_hence A: \\[G]v & \\\!u ([F]u & [R']uv)\ + using "KBasic2:3" "\E" by blast + AOT_hence \\[G]v\ + using G_nec "&E"(1) "\E"(1) "sc-eq-box-box:1" "\E" by blast + AOT_hence \[G]v\ by (metis "qml:2"[axiom_inst] "\E") + AOT_hence \\!u ([F]u & [R]uv)\ + using C2[THEN "Ordinary.\E", THEN "\E"] by blast + AOT_hence \\u ([F]u & [R]uv & \u' ([F]u' & [R]u'v \ u' =\<^sub>E u))\ + using "equi:1"[THEN "\E"(1)] by auto + then AOT_obtain a where + a_prop: \O!a & ([F]a & [R]av & \u' ([F]u' & [R]u'v \ u' =\<^sub>E a))\ + using "\E"[rotated] by blast + AOT_have \\u \([F]u & [R']uv & \u' ([F]u' & [R']u'v \ u' =\<^sub>E u))\ + proof(safe intro!: "\I"(2)[where \=a] "&I" a_prop[THEN "&E"(1)] + "KBasic:3"[THEN "\E"(2)]) + AOT_show \\[F]a\ + using a_prop[THEN "&E"(2), THEN "&E"(1), THEN "&E"(1)] + by (metis F_nec "qml:2"[axiom_inst] "\E") + next + AOT_show \\[R']av\ + using D[THEN "\E"(2), THEN "\E"(2), THEN "\E"] + E[THEN "\E"(2), THEN "\E"(2), THEN "\E"(2), + OF a_prop[THEN "&E"(2), THEN "&E"(1), THEN "&E"(2)]] + by (metis "T\" "\E") + next + AOT_have \\u' \([F]u' & [R']u'v \ u' =\<^sub>E a)\ + proof (rule Ordinary.GEN; rule "raa-cor:1") + fix u' + AOT_assume \\\([F]u' & [R']u'v \ u' =\<^sub>E a)\ + AOT_hence \\\([F]u' & [R']u'v \ u' =\<^sub>E a)\ + by (metis "KBasic:11" "\E"(1)) + AOT_hence \\([F]u' & [R']u'v & \u' =\<^sub>E a)\ + by (AOT_subst \[F]u' & [R']u'v & \u' =\<^sub>E a\ + \\([F]u' & [R']u'v \ u' =\<^sub>E a)\) + (meson "\E"(6) "oth-class-taut:1:b" "oth-class-taut:3:a") + AOT_hence 1: \\[F]u'\ and 2: \\[R']u'v\ and 3: \\\u' =\<^sub>E a\ + using "KBasic2:3"[THEN "\E", THEN "&E"(1)] + "KBasic2:3"[THEN "\E", THEN "&E"(2)] by blast+ + AOT_have Fu': \[F]u'\ using F_nec 1 + by (meson "B\" "KBasic:13" "\E") + AOT_have \\[R']u'v\ + using 2 D[THEN "\E"(2), THEN "\E"(2), THEN "\E"] by blast + AOT_hence R'u'v: \[R']u'v\ + by (metis "B\" "T\" "\E") + AOT_hence \[R]u'v\ + using E[THEN "\E"(2), THEN "\E"(2), THEN "\E"(1)] by blast + AOT_hence \u' =\<^sub>E a\ + using a_prop[THEN "&E"(2), THEN "&E"(2), THEN "Ordinary.\E", + THEN "\E", OF "&I", OF Fu'] by blast + AOT_hence \\(u' =\<^sub>E a)\ + by (metis "id-nec3:1" "\E"(4) "raa-cor:3") + moreover AOT_have \\\(u' =\<^sub>E a)\ + using 3 "KBasic:11" "\E"(2) by blast + ultimately AOT_show \\(u' =\<^sub>E a) & \\(u' =\<^sub>E a)\ + using "&I" by blast + qed + AOT_thus \\\u'([F]u' & [R']u'v \ u' =\<^sub>E a)\ + using "Ordinary.res-var-bound-reas[BF]" "\E" by fast + qed + AOT_hence 1: \\\u ([F]u & [R']uv & \u' ([F]u' & [R']u'v \ u' =\<^sub>E u))\ + using "Ordinary.res-var-bound-reas[Buridan]" "\E" by fast + AOT_hence \\\!u ([F]u & [R']uv)\ + by (AOT_subst_thm "equi:1") + moreover AOT_have \\\\!u ([F]u & [R']uv)\ + using A[THEN "&E"(2)] "KBasic:11"[THEN "\E"(2)] by blast + ultimately AOT_show \\\!u ([F]u & [R']uv) & \\\!u ([F]u & [R']uv)\ + by (rule "&I") + qed + qed + AOT_hence \\\u ([F]u \ \!v ([G]v & [R']uv))\ + and \\\v ([G]v \ \!u ([F]u & [R']uv))\ + using "Ordinary.res-var-bound-reas[BF]"[THEN "\E"] by auto + moreover AOT_have \\[R']\\ and \\[F]\\ and \\[G]\\ + by (simp_all add: "ex:2:a") + ultimately AOT_have \\([R']\ & [F]\ & [G]\ & \u ([F]u \ \!v ([G]v & [R']uv)) & + \v ([G]v \ \!u ([F]u & [R']uv)))\ + using "KBasic:3" "&I" "\E"(2) by meson + AOT_hence \\R' |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E G\ + by (AOT_subst_def "equi:2") + AOT_hence \\R \R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E G\ + by (rule "\I"(2)) + AOT_hence \\\R R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E G\ + by (metis Buridan "\E") + AOT_thus \\F \\<^sub>E G\ + by (AOT_subst_def "equi:3") + } + qed + ultimately AOT_show \\(F \\<^sub>E G \ \F \\<^sub>E G)\ + using "\E" by blast +qed + + +AOT_define numbers :: \\ \ \ \ \\ (\Numbers'(_,_')\) + \Numbers(x,G) \\<^sub>d\<^sub>f A!x & G\ & \F(x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G)\ + +AOT_theorem "numbers[den]": + \\\ \ (Numbers(\, \) \ A!\ & \F(\[F] \ [\z \<^bold>\[F]z] \\<^sub>E \))\ + apply (safe intro!: numbers[THEN "\\<^sub>d\<^sub>fI"] "&I" "\I" "\I" "cqt:2" + dest!: numbers[THEN "\\<^sub>d\<^sub>fE"]) + using "&E" by blast+ + +AOT_theorem "num-tran:1": + \G \\<^sub>E H \ (Numbers(x, G) \ Numbers(x, H))\ +proof (safe intro!: "\I" "\I") + AOT_assume 0: \G \\<^sub>E H\ + AOT_assume \Numbers(x, G)\ + AOT_hence Ax: \A!x\ and \: \\F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G)\ + using numbers[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast+ + AOT_show \Numbers(x, H)\ + proof(safe intro!: numbers[THEN "\\<^sub>d\<^sub>fI"] "&I" Ax "cqt:2" GEN) + fix F + AOT_have \x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G\ + using \[THEN "\E"(2)]. + also AOT_have \\ \ [\z \<^bold>\[F]z] \\<^sub>E H\ + using 0 "approx-nec:2"[THEN "\E"(1), THEN "\E"(2)] by metis + finally AOT_show \x[F] \ [\z \<^bold>\[F]z] \\<^sub>E H\. + qed +next + AOT_assume \G \\<^sub>E H\ + AOT_hence 0: \H \\<^sub>E G\ + by (metis "eq-part:2" "\E") + AOT_assume \Numbers(x, H)\ + AOT_hence Ax: \A!x\ and \: \\F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E H)\ + using numbers[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast+ + AOT_show \Numbers(x, G)\ + proof(safe intro!: numbers[THEN "\\<^sub>d\<^sub>fI"] "&I" Ax "cqt:2" GEN) + fix F + AOT_have \x[F] \ [\z \<^bold>\[F]z] \\<^sub>E H\ + using \[THEN "\E"(2)]. + also AOT_have \\ \ [\z \<^bold>\[F]z] \\<^sub>E G\ + using 0 "approx-nec:2"[THEN "\E"(1), THEN "\E"(2)] by metis + finally AOT_show \x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G\. + qed +qed + +AOT_theorem "num-tran:2": + \(Numbers(x, G) & Numbers(x,H)) \ G \\<^sub>E H\ +proof (rule "\I"; frule "&E"(1); drule "&E"(2)) + AOT_assume \Numbers(x,G)\ + AOT_hence \\F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G)\ + using numbers[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + AOT_hence 1: \x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G\ for F + using "\E"(2) by blast + AOT_assume \Numbers(x,H)\ + AOT_hence \\F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E H)\ + using numbers[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + AOT_hence \x[F] \ [\z \<^bold>\[F]z] \\<^sub>E H\ for F + using "\E"(2) by blast + AOT_hence \[\z \<^bold>\[F]z] \\<^sub>E G \ [\z \<^bold>\[F]z] \\<^sub>E H\ for F + by (metis "1" "\E"(6)) + AOT_thus \G \\<^sub>E H\ + using "approx-nec:2"[THEN "\E"(2), OF GEN] by blast +qed + +AOT_theorem "num-tran:3": + \G \\<^sub>E H \ (Numbers(x, G) \ Numbers(x, H))\ + using "apE-eqE:1" "Hypothetical Syllogism" "num-tran:1" by blast + +AOT_theorem "pre-Hume": + \(Numbers(x,G) & Numbers(y,H)) \ (x = y \ G \\<^sub>E H)\ +proof(safe intro!: "\I" "\I"; frule "&E"(1); drule "&E"(2)) + AOT_assume \Numbers(x, G)\ + moreover AOT_assume \x = y\ + ultimately AOT_have \Numbers(y, G)\ by (rule "rule=E") + moreover AOT_assume \Numbers(y, H)\ + ultimately AOT_show \G \\<^sub>E H\ using "num-tran:2" "\E" "&I" by blast +next + AOT_assume \Numbers(x, G)\ + AOT_hence Ax: \A!x\ and xF: \\F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G)\ + using numbers[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast+ + AOT_assume \Numbers(y, H)\ + AOT_hence Ay: \A!y\ and yF: \\F (y[F] \ [\z \<^bold>\[F]z] \\<^sub>E H)\ + using numbers[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast+ + AOT_assume G_approx_H: \G \\<^sub>E H\ + AOT_show \x = y\ + proof(rule "ab-obey:1"[THEN "\E", THEN "\E", OF "&I", OF Ax, OF Ay]; rule GEN) + fix F + AOT_have \x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G\ + using xF[THEN "\E"(2)]. + also AOT_have \\ \ [\z \<^bold>\[F]z] \\<^sub>E H\ + using "approx-nec:2"[THEN "\E"(1), OF G_approx_H, THEN "\E"(2)]. + also AOT_have \\ \ y[F]\ + using yF[THEN "\E"(2), symmetric]. + finally AOT_show \x[F] \ y[F]\. + qed +qed + +AOT_theorem "two-num-not": + \\u\v(u \ v) \ \x\G\H(Numbers(x,G) & Numbers(x, H) & \G \\<^sub>E H)\ +proof (rule "\I") + AOT_have eqE_den: \[\x x =\<^sub>E y]\\ for y by "cqt:2" + AOT_assume \\u\v(u \ v)\ + then AOT_obtain c where Oc: \O!c\ and \\v (c \ v)\ + using "&E" "\E"[rotated] by blast + then AOT_obtain d where Od: \O!d\ and c_noteq_d: \c \ d\ + using "&E" "\E"[rotated] by blast + AOT_hence c_noteqE_d: \c \\<^sub>E d\ + using "=E-simple:2"[THEN "\E"] "=E-simple:2" "\E"(2) "modus-tollens:1" + "=-infix" "\\<^sub>d\<^sub>fE" "thm-neg=E" by fast + AOT_hence not_c_eqE_d: \\c =\<^sub>E d\ + using "\E"(1) "thm-neg=E" by blast + AOT_have \\x (A!x & \F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E [\x x =\<^sub>E c]))\ + by (simp add: "A-objects"[axiom_inst]) + then AOT_obtain a where a_prop: \A!a & \F (a[F] \ [\z \<^bold>\[F]z] \\<^sub>E [\x x =\<^sub>E c])\ + using "\E"[rotated] by blast + AOT_have \\x (A!x & \F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E [\x x =\<^sub>E d]))\ + by (simp add: "A-objects" "vdash-properties:1[2]") + then AOT_obtain b where b_prop: \A!b & \F (b[F] \ [\z \<^bold>\[F]z] \\<^sub>E [\x x =\<^sub>E d])\ + using "\E"[rotated] by blast + AOT_have num_a_eq_c: \Numbers(a, [\x x =\<^sub>E c])\ + by (safe intro!: numbers[THEN "\\<^sub>d\<^sub>fI"] "&I" a_prop[THEN "&E"(1)] + a_prop[THEN "&E"(2)]) "cqt:2" + moreover AOT_have num_b_eq_d: \Numbers(b, [\x x =\<^sub>E d])\ + by (safe intro!: numbers[THEN "\\<^sub>d\<^sub>fI"] "&I" b_prop[THEN "&E"(1)] + b_prop[THEN "&E"(2)]) "cqt:2" + moreover AOT_have \[\x x =\<^sub>E c] \\<^sub>E [\x x =\<^sub>E d]\ + proof (rule "equi:3"[THEN "\\<^sub>d\<^sub>fI"]) + let ?R = \\[\xy (x =\<^sub>E c & y =\<^sub>E d)]\\ + AOT_have Rcd: \[\?R\]cd\ + by (auto intro!: "\\C"(1) "cqt:2" "&I" prod_denotesI + "ord=Eequiv:1"[THEN "\E"] Od Oc) + AOT_show \\R R |: [\x x =\<^sub>E c] \<^sub>1\<^sub>-\<^sub>1\\<^sub>E [\x x =\<^sub>E d]\ + proof (safe intro!: "\I"(1)[where \=\?R\] "equi:2"[THEN "\\<^sub>d\<^sub>fI"] "&I" + eqE_den Ordinary.GEN "\I") + AOT_show \\?R\\\ by "cqt:2" + next + fix u + AOT_assume \[\x x =\<^sub>E c]u\ + AOT_hence \u =\<^sub>E c\ + by (metis "\\C"(1)) + AOT_hence u_is_c: \u = c\ + by (metis "=E-simple:2" "\E") + AOT_show \\!v ([\x x =\<^sub>E d]v & [\?R\]uv)\ + proof (safe intro!: "equi:1"[THEN "\E"(2)] "\I"(2)[where \=d] "&I" + Od Ordinary.GEN "\I") + AOT_show \[\x x =\<^sub>E d]d\ + by (auto intro!: "\\C"(1) "cqt:2" "ord=Eequiv:1"[THEN "\E", OF Od]) + next + AOT_show \[\?R\]ud\ + using u_is_c[symmetric] Rcd "rule=E" by fast + next + fix v + AOT_assume \[\x x =\<^sub>E d]v & [\?R\]uv\ + AOT_thus \v =\<^sub>E d\ + by (metis "\\C"(1) "&E"(1)) + qed + next + fix v + AOT_assume \[\x x =\<^sub>E d]v\ + AOT_hence \v =\<^sub>E d\ + by (metis "\\C"(1)) + AOT_hence v_is_d: \v = d\ + by (metis "=E-simple:2" "\E") + AOT_show \\!u ([\x x =\<^sub>E c]u & [\?R\]uv)\ + proof (safe intro!: "equi:1"[THEN "\E"(2)] "\I"(2)[where \=c] "&I" + Oc Ordinary.GEN "\I") + AOT_show \[\x x =\<^sub>E c]c\ + by (auto intro!: "\\C"(1) "cqt:2" "ord=Eequiv:1"[THEN "\E", OF Oc]) + next + AOT_show \[\?R\]cv\ + using v_is_d[symmetric] Rcd "rule=E" by fast + next + fix u + AOT_assume \[\x x =\<^sub>E c]u & [\?R\]uv\ + AOT_thus \u =\<^sub>E c\ + by (metis "\\C"(1) "&E"(1)) + qed + next + AOT_show \\?R\\\ + by "cqt:2" + qed + qed + ultimately AOT_have \a = b\ + using "pre-Hume"[unvarify G H, OF eqE_den, OF eqE_den, THEN "\E", + OF "&I", THEN "\E"(2)] by blast + AOT_hence num_a_eq_d: \Numbers(a, [\x x =\<^sub>E d])\ + using num_b_eq_d "rule=E" id_sym by fast + AOT_have not_equiv: \\[\x x =\<^sub>E c] \\<^sub>E [\x x =\<^sub>E d]\ + proof (rule "raa-cor:2") + AOT_assume \[\x x =\<^sub>E c] \\<^sub>E [\x x =\<^sub>E d]\ + AOT_hence \[\x x =\<^sub>E c]c \ [\x x =\<^sub>E d]c\ + using eqE[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2), THEN "\E"(2), THEN "\E"] Oc by blast + moreover AOT_have \[\x x =\<^sub>E c]c\ + by (auto intro!: "\\C"(1) "cqt:2" "ord=Eequiv:1"[THEN "\E", OF Oc]) + ultimately AOT_have \[\x x =\<^sub>E d]c\ + using "\E"(1) by blast + AOT_hence \c =\<^sub>E d\ + by (rule "\\C"(1)) + AOT_thus \c =\<^sub>E d & \c =\<^sub>E d\ + using not_c_eqE_d "&I" by blast + qed + AOT_show \\x \G \H (Numbers(x,G) & Numbers(x,H) & \G \\<^sub>E H)\ + apply (rule "\I"(2)[where \=a]) + apply (rule "\I"(1)[where \=\\[\x x =\<^sub>E c]\\]) + apply (rule "\I"(1)[where \=\\[\x x =\<^sub>E d]\\]) + by (safe intro!: eqE_den "&I" num_a_eq_c num_a_eq_d not_equiv) +qed + +AOT_theorem "num:1": \\x Numbers(x,G)\ + by (AOT_subst \Numbers(x,G)\ \[A!]x & \F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G)\ for: x) + (auto simp: "numbers[den]"[THEN "\E", OF "cqt:2[const_var]"[axiom_inst]] + "A-objects"[axiom_inst]) + +AOT_theorem "num:2": \\!x Numbers(x,G)\ + by (AOT_subst \Numbers(x,G)\ \[A!]x & \F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G)\ for: x) + (auto simp: "numbers[den]"[THEN "\E", OF "cqt:2[const_var]"[axiom_inst]] + "A-objects!") + +AOT_theorem "num-cont:1": + \\x\G(Numbers(x, G) & \\Numbers(x, G))\ +proof - + AOT_have \\F\G \([\z \<^bold>\[F]z] \\<^sub>E G & \\[\z \<^bold>\[F]z] \\<^sub>E G)\ + using "approx-cont:2". + then AOT_obtain F where \\G \([\z \<^bold>\[F]z] \\<^sub>E G & \\[\z \<^bold>\[F]z] \\<^sub>E G)\ + using "\E"[rotated] by blast + then AOT_obtain G where \\([\z \<^bold>\[F]z] \\<^sub>E G & \\[\z \<^bold>\[F]z] \\<^sub>E G)\ + using "\E"[rotated] by blast + AOT_hence \: \\[\z \<^bold>\[F]z] \\<^sub>E G\ and \: \\\[\z \<^bold>\[F]z] \\<^sub>E G\ + using "KBasic2:3"[THEN "\E"] "&E" "4\"[THEN "\E"] by blast+ + AOT_obtain a where \Numbers(a, G)\ + using "num:1" "\E"[rotated] by blast + moreover AOT_have \\\Numbers(a, G)\ + proof (rule "raa-cor:2") + AOT_assume \\Numbers(a, G)\ + AOT_hence \\([A!]a & G\ & \F (a[F] \ [\z \<^bold>\[F]z] \\<^sub>E G))\ + by (AOT_subst_def (reverse) numbers) + AOT_hence \\A!a\ and \\\F (a[F] \ [\z \<^bold>\[F]z] \\<^sub>E G)\ + using "KBasic:3"[THEN "\E"(1)] "&E" by blast+ + AOT_hence \\F \(a[F] \ [\z \<^bold>\[F]z] \\<^sub>E G)\ + using CBF[THEN "\E"] by blast + AOT_hence \\(a[F] \ [\z \<^bold>\[F]z] \\<^sub>E G)\ + using "\E"(2) by blast + AOT_hence A: \\(a[F] \ [\z \<^bold>\[F]z] \\<^sub>E G)\ + and B: \\([\z \<^bold>\[F]z] \\<^sub>E G \ a[F])\ + using "KBasic:4"[THEN "\E"(1)] "&E" by blast+ + AOT_have \\(\[\z \<^bold>\[F]z] \\<^sub>E G \ \a[F])\ + apply (AOT_subst \\[\z \<^bold>\[F]z] \\<^sub>E G \ \a[F]\ \a[F] \ [\z \<^bold>\[F]z] \\<^sub>E G\) + using "\I" "useful-tautologies:4" "useful-tautologies:5" apply presburger + by (fact A) + AOT_hence \\\a[F]\ + by (metis "KBasic:13" \ "\E") + AOT_hence \\a[F]\ + by (metis "KBasic:11" "en-eq:2[1]" "\E"(2) "\E"(4)) + AOT_hence \\\a[F]\ + by (metis "en-eq:3[1]" "\E"(4)) + moreover AOT_have \\a[F]\ + by (meson B \ "KBasic:13" "\E") + ultimately AOT_show \\a[F] & \\a[F]\ + using "&I" by blast + qed + + ultimately AOT_have \Numbers(a, G) & \\Numbers(a, G)\ + using "&I" by blast + AOT_hence \\G (Numbers(a, G) & \\Numbers(a, G))\ + by (rule "\I") + AOT_thus \\x\G (Numbers(x, G) & \\Numbers(x, G))\ + by (rule "\I") +qed + +AOT_theorem "num-cont:2": + \Rigid(G) \ \\x(Numbers(x,G) \ \Numbers(x,G))\ +proof(rule "\I") + AOT_assume \Rigid(G)\ + AOT_hence \\\z([G]z \ \[G]z)\ + using "df-rigid-rel:1"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2)] by blast + AOT_hence \\\\z([G]z \ \[G]z)\ by (metis "S5Basic:6" "\E"(1)) + moreover AOT_have \\\\z([G]z \ \[G]z) \ \\x(Numbers(x,G) \ \Numbers(x,G))\ + proof(rule RM; safe intro!: "\I" GEN) + AOT_modally_strict { + AOT_have act_den: \[\z \<^bold>\[F]z]\\ for F by "cqt:2[lambda]" + fix x + AOT_assume G_nec: \\\z([G]z \ \[G]z)\ + AOT_hence G_rigid: \Rigid(G)\ + using "df-rigid-rel:1"[THEN "\\<^sub>d\<^sub>fI", OF "&I"] "cqt:2" + by blast + AOT_assume \Numbers(x, G)\ + AOT_hence \[A!]x & G\ & \F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G)\ + using numbers[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_hence Ax: \[A!]x\ and \\F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G)\ + using "&E" by blast+ + AOT_hence \x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G\ for F + using "\E"(2) by blast + moreover AOT_have \\([\z \<^bold>\[F]z] \\<^sub>E G \ \[\z \<^bold>\[F]z] \\<^sub>E G)\ for F + using "approx-nec:3"[unvarify F, OF act_den, THEN "\E", OF "&I", + OF "actuallyF:2", OF G_rigid]. + moreover AOT_have \\(x[F] \ \x[F])\ for F + by (simp add: RN "pre-en-eq:1[1]") + ultimately AOT_have \\(x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G)\ for F + using "sc-eq-box-box:5" "\E" "qml:2"[axiom_inst] "&I" by meson + AOT_hence \\F \(x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G)\ + by (rule "\I") + AOT_hence 1: \\\F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G)\ + using BF[THEN "\E"] by fast + AOT_have \\G\\ + by (simp add: "ex:2:a") + moreover AOT_have \\[A!]x\ + using Ax "oa-facts:2" "\E" by blast + ultimately AOT_have \\(A!x & G\)\ + by (metis "KBasic:3" "&I" "\E"(2)) + AOT_hence \\(A!x & G\ & \F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G))\ + using 1 "KBasic:3" "&I" "\E"(2) by fast + AOT_thus \\Numbers(x, G)\ + by (AOT_subst_def numbers) + } + qed + ultimately AOT_show \\\x(Numbers(x,G) \ \Numbers(x,G))\ + using "\E" by blast +qed + +AOT_theorem "num-cont:3": + \\\x(Numbers(x, [\z \<^bold>\[G]z]) \ \Numbers(x, [\z \<^bold>\[G]z]))\ + by (rule "num-cont:2"[unvarify G, THEN "\E"]; + ("cqt:2[lambda]" | rule "actuallyF:2")) + +AOT_theorem "num-uniq": \\<^bold>\x Numbers(x, G)\\ + using "\E"(2) "A-Exists:2" "RA[2]" "num:2" by blast + +AOT_define num :: \\ \ \\<^sub>s\ (\#_\ [100] 100) + "num-def:1": \#G =\<^sub>d\<^sub>f \<^bold>\x Numbers(x, G)\ + +AOT_theorem "num-def:2": \#G\\ + using "num-def:1"[THEN "=\<^sub>d\<^sub>fI"(1)] "num-uniq" by simp + +AOT_theorem "num-can:1": + \#G = \<^bold>\x(A!x & \F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G))\ +proof - + AOT_have \\\x(Numbers(x,G) \ [A!]x & \F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G))\ + by (safe intro!: RN GEN "numbers[den]"[THEN "\E"] "cqt:2") + AOT_hence \\<^bold>\x Numbers(x, G) = \<^bold>\x([A!]x & \F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G))\ + using "num-uniq" "equiv-desc-eq:3"[THEN "\E", OF "&I"] by auto + thus ?thesis + by (rule "=\<^sub>d\<^sub>fI"(1)[OF "num-def:1", OF "num-uniq"]) +qed + +AOT_theorem "num-can:2": \#G = \<^bold>\x(A!x & \F (x[F] \ F \\<^sub>E G))\ +proof (rule id_trans[OF "num-can:1"]; rule "equiv-desc-eq:2"[THEN "\E"]; + safe intro!: "&I" "A-descriptions" GEN "Act-Basic:5"[THEN "\E"(2)] + "logic-actual-nec:3"[axiom_inst, THEN "\E"(2)]) + AOT_have act_den: \\<^bold>\\<^sub>\ [\z \<^bold>\[F]z]\\ for F + by "cqt:2" + AOT_have "eq-part:3[terms]": \\<^bold>\\<^sub>\ F \\<^sub>E G & F \\<^sub>E H \ G \\<^sub>E H\ for F G H + by (metis "&I" "eq-part:2" "eq-part:3" "\I" "&E" "\E") + fix x + { + fix F + AOT_have \\<^bold>\(F \\<^sub>E [\z \<^bold>\[F]z])\ + by (simp add: "actuallyF:1") + moreover AOT_have \\<^bold>\((F \\<^sub>E [\z \<^bold>\[F]z]) \ ([\z \<^bold>\[F]z] \\<^sub>E G \ F \\<^sub>E G))\ + by (auto intro!: "RA[2]" "\I" "\I" + simp: "eq-part:3"[unvarify G, OF act_den, THEN "\E", OF "&I"] + "eq-part:3[terms]"[unvarify G, OF act_den, THEN "\E", OF "&I"]) + ultimately AOT_have \\<^bold>\([\z \<^bold>\[F]z] \\<^sub>E G \ F \\<^sub>E G)\ + using "logic-actual-nec:2"[axiom_inst, THEN "\E"(1), THEN "\E"] by blast + + AOT_hence \\<^bold>\[\z \<^bold>\[F]z] \\<^sub>E G \ \<^bold>\F \\<^sub>E G\ + by (metis "Act-Basic:5" "\E"(1)) + AOT_hence 0: \(\<^bold>\x[F] \ \<^bold>\[\z \<^bold>\[F]z] \\<^sub>E G) \ (\<^bold>\x[F] \ \<^bold>\F \\<^sub>E G)\ + by (auto intro!: "\I" "\I" elim: "\E") + AOT_have \\<^bold>\(x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G) \ (\<^bold>\x[F] \ \<^bold>\[\z \<^bold>\[F]z] \\<^sub>E G)\ + by (simp add: "Act-Basic:5") + also AOT_have \\ \ (\<^bold>\x[F] \ \<^bold>\F \\<^sub>E G)\ using 0. + also AOT_have \\ \ \<^bold>\((x[F] \ F \\<^sub>E G))\ + by (meson "Act-Basic:5" "\E"(6) "oth-class-taut:3:a") + finally AOT_have 0: \\<^bold>\(x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G) \ \<^bold>\((x[F] \ F \\<^sub>E G))\. + } note 0 = this + AOT_have \\<^bold>\\F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G) \ \F \<^bold>\(x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G)\ + using "logic-actual-nec:3" "vdash-properties:1[2]" by blast + also AOT_have \\ \ \F \<^bold>\((x[F] \ F \\<^sub>E G))\ + apply (safe intro!: "\I" "\I" GEN) + using 0 "\E"(1) "\E"(2) "rule-ui:3" by blast+ + also AOT_have \\ \ \<^bold>\(\F (x[F] \ F \\<^sub>E G))\ + using "\E"(6) "logic-actual-nec:3"[axiom_inst] "oth-class-taut:3:a" by fast + finally AOT_have 0: \\<^bold>\\F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G) \ \<^bold>\(\F (x[F] \ F \\<^sub>E G))\. + AOT_have \\<^bold>\([A!]x & \F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G)) \ + (\<^bold>\A!x & \<^bold>\\F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G))\ + by (simp add: "Act-Basic:2") + also AOT_have \\ \ \<^bold>\[A!]x & \<^bold>\(\F (x[F] \ F \\<^sub>E G))\ + using 0 "oth-class-taut:4:f" "\E" by blast + also AOT_have \\ \ \<^bold>\(A!x & \F (x[F] \ F \\<^sub>E G))\ + using "Act-Basic:2" "\E"(6) "oth-class-taut:3:a" by blast + finally AOT_show \\<^bold>\([A!]x & \F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G)) \ + \<^bold>\([A!]x & \F (x[F] \ F \\<^sub>E G))\. +qed + +AOT_define NaturalCardinal :: \\ \ \\ (\NaturalCardinal'(_')\) + card: \NaturalCardinal(x) \\<^sub>d\<^sub>f \G(x = #G)\ + +AOT_theorem "natcard-nec": \NaturalCardinal(x) \ \NaturalCardinal(x)\ +proof(rule "\I") + AOT_assume \NaturalCardinal(x)\ + AOT_hence \\G(x = #G)\ using card[THEN "\\<^sub>d\<^sub>fE"] by blast + then AOT_obtain G where \x = #G\ using "\E"[rotated] by blast + AOT_hence \\x = #G\ by (metis "id-nec:2" "\E") + AOT_hence \\G \x = #G\ by (rule "\I") + AOT_hence \\\G x = #G\ by (metis Buridan "\E") + AOT_thus \\NaturalCardinal(x)\ + by (AOT_subst_def card) +qed + +AOT_act_theorem "hume:1": \Numbers(#G, G)\ + apply (rule "=\<^sub>d\<^sub>fI"(1)[OF "num-def:1"]) + apply (simp add: "num-uniq") + using "num-uniq" "vdash-properties:10" "y-in:3" by blast + +AOT_act_theorem "hume:2": \#F = #G \ F \\<^sub>E G\ + by (safe intro!: "pre-Hume"[unvarify x y, OF "num-def:2", + OF "num-def:2", THEN "\E"] "&I" "hume:1") + +AOT_act_theorem "hume:3": \#F = #G \ \R (R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>o\<^sub>n\<^sub>t\<^sub>oE G)\ + using "equi-rem-thm" + apply (AOT_subst (reverse) \R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>o\<^sub>n\<^sub>t\<^sub>oE G\ + \R |: F \<^sub>1\<^sub>-\<^sub>1\\<^sub>E G\ for: R :: \<\\\>\) + using "equi:3" "hume:2" "\E"(5) "\Df" by blast + +AOT_act_theorem "hume:4": \F \\<^sub>E G \ #F = #G\ + by (metis "apE-eqE:1" "deduction-theorem" "hume:2" "\E"(2) "\E") + +AOT_theorem "hume-strict:1": + \\x (Numbers(x, F) & Numbers(x, G)) \ F \\<^sub>E G\ +proof(safe intro!: "\I" "\I") + AOT_assume \\x (Numbers(x, F) & Numbers(x, G))\ + then AOT_obtain a where \Numbers(a, F) & Numbers(a, G)\ + using "\E"[rotated] by blast + AOT_thus \F \\<^sub>E G\ + using "num-tran:2" "\E" by blast +next + AOT_assume 0: \F \\<^sub>E G\ + moreover AOT_obtain b where num_b_F: \Numbers(b, F)\ + by (metis "instantiation" "num:1") + moreover AOT_have num_b_G: \Numbers(b, G)\ + using calculation "num-tran:1"[THEN "\E", THEN "\E"(1)] by blast + ultimately AOT_have \Numbers(b, F) & Numbers(b, G)\ + by (safe intro!: "&I") + AOT_thus \\x (Numbers(x, F) & Numbers(x, G))\ + by (rule "\I") +qed + +AOT_theorem "hume-strict:2": + \\x\y (Numbers(x, F) & + \z(Numbers(z,F) \ z = x) & + Numbers(y, G) & + \z (Numbers(z, G) \ z = y) & + x = y) \ + F \\<^sub>E G\ +proof(safe intro!: "\I" "\I") + AOT_assume \\x\y (Numbers(x, F) & \z(Numbers(z,F) \ z = x) & + Numbers(y, G) & \z (Numbers(z, G) \ z = y) & x = y)\ + then AOT_obtain x where + \\y (Numbers(x, F) & \z(Numbers(z,F) \ z = x) & Numbers(y, G) & + \z (Numbers(z, G) \ z = y) & x = y)\ + using "\E"[rotated] by blast + then AOT_obtain y where + \Numbers(x, F) & \z(Numbers(z,F) \ z = x) & Numbers(y, G) & + \z (Numbers(z, G) \ z = y) & x = y\ + using "\E"[rotated] by blast + AOT_hence \Numbers(x, F)\ and \Numbers(y,G)\ and \x = y\ + using "&E" by blast+ + AOT_hence \Numbers(y, F) & Numbers(y, G)\ + using "&I" "rule=E" by fast + AOT_hence \\y (Numbers(y, F) & Numbers(y, G))\ + by (rule "\I") + AOT_thus \F \\<^sub>E G\ + using "hume-strict:1"[THEN "\E"(1)] by blast +next + AOT_assume \F \\<^sub>E G\ + AOT_hence \\x (Numbers(x, F) & Numbers(x, G))\ + using "hume-strict:1"[THEN "\E"(2)] by blast + then AOT_obtain x where \Numbers(x, F) & Numbers(x, G)\ + using "\E"[rotated] by blast + moreover AOT_have \\z (Numbers(z, F) \ z = x)\ + and \\z (Numbers(z, G) \ z = x)\ + using calculation + by (auto intro!: GEN "\I" "pre-Hume"[THEN "\E", OF "&I", THEN "\E"(2), + rotated 2, OF "eq-part:1"] dest: "&E") + ultimately AOT_have \Numbers(x, F) & \z(Numbers(z,F) \ z = x) & + Numbers(x, G) & \z (Numbers(z, G) \ z = x) & x = x\ + by (auto intro!: "&I" "id-eq:1" dest: "&E") + AOT_thus \\x\y (Numbers(x, F) & \z(Numbers(z,F) \ z = x) & Numbers(y, G) & + \z (Numbers(z, G) \ z = y) & x = y)\ + by (auto intro!: "\I") +qed + +AOT_theorem unotEu: \\\y[\x O!x & x \\<^sub>E x]y\ +proof(rule "raa-cor:2") + AOT_assume \\y[\x O!x & x \\<^sub>E x]y\ + then AOT_obtain y where \[\x O!x & x \\<^sub>E x]y\ + using "\E"[rotated] by blast + AOT_hence 0: \O!y & y \\<^sub>E y\ + by (rule "\\C"(1)) + AOT_hence \\(y =\<^sub>E y)\ + using "&E"(2) "\E"(1) "thm-neg=E" by blast + moreover AOT_have \y =\<^sub>E y\ + by (metis 0[THEN "&E"(1)] "ord=Eequiv:1" "\E") + ultimately AOT_show \p & \p\ for p + by (metis "raa-cor:3") +qed + +AOT_define zero :: \\\<^sub>s\ (\0\) + "zero:1": \0 =\<^sub>d\<^sub>f #[\x O!x & x \\<^sub>E x]\ + +AOT_theorem "zero:2": \0\\ + by (rule "=\<^sub>d\<^sub>fI"(2)[OF "zero:1"]; rule "num-def:2"[unvarify G]; "cqt:2") + +AOT_theorem "zero-card": \NaturalCardinal(0)\ + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF "zero:1"]) + apply (rule "num-def:2"[unvarify G]; "cqt:2") + apply (rule card[THEN "\\<^sub>d\<^sub>fI"]) + apply (rule "\I"(1)[where \=\\[\x [O!]x & x \\<^sub>E x]\\]) + apply (rule "rule=I:1"; rule "num-def:2"[unvarify G]; "cqt:2") + by "cqt:2" + +AOT_theorem "eq-num:1": + \\<^bold>\Numbers(x, G) \ Numbers(x,[\z \<^bold>\[G]z])\ +proof - + AOT_have act_den: \\<^bold>\\<^sub>\ [\z \<^bold>\[F]z]\\ for F by "cqt:2" + AOT_have \\(\x(Numbers(x, G) & Numbers(x,[\z \<^bold>\[G]z])) \ G \\<^sub>E [\z \<^bold>\[G]z])\ + using "hume-strict:1"[unvarify G, OF act_den, THEN RN]. + AOT_hence \\<^bold>\(\x(Numbers(x, G) & Numbers(x,[\z \<^bold>\[G]z])) \ G \\<^sub>E [\z \<^bold>\[G]z])\ + using "nec-imp-act"[THEN "\E"] by fast + AOT_hence \\<^bold>\(\x(Numbers(x, G) & Numbers(x,[\z \<^bold>\[G]z])))\ + using "actuallyF:1" "Act-Basic:5" "\E"(1) "\E"(2) by fast + AOT_hence \\x \<^bold>\((Numbers(x, G) & Numbers(x,[\z \<^bold>\[G]z])))\ + by (metis "Act-Basic:10" "intro-elim:3:a") + then AOT_obtain a where \\<^bold>\(Numbers(a, G) & Numbers(a,[\z \<^bold>\[G]z]))\ + using "\E"[rotated] by blast + AOT_hence act_a_num_G: \\<^bold>\Numbers(a, G)\ + and act_a_num_actG: \\<^bold>\Numbers(a,[\z \<^bold>\[G]z])\ + using "Act-Basic:2" "&E" "\E"(1) by blast+ + AOT_hence num_a_act_g: \Numbers(a, [\z \<^bold>\[G]z])\ + using "num-cont:2"[unvarify G, OF act_den, THEN "\E", OF "actuallyF:2", + THEN CBF[THEN "\E"], THEN "\E"(2)] + by (metis "\E"(1) "sc-eq-fur:2" "vdash-properties:6") + AOT_have 0: \\<^bold>\\<^sub>\ Numbers(x, G) & Numbers(y, G) \ x = y\ for y + using "pre-Hume"[THEN "\E", THEN "\E"(2), rotated, OF "eq-part:1"] + "\I" by blast + show ?thesis + proof(safe intro!: "\I" "\I") + AOT_assume \\<^bold>\Numbers(x, G)\ + AOT_hence \\<^bold>\x = a\ + using 0[THEN "RA[2]", THEN "act-cond"[THEN "\E"], THEN "\E", + OF "Act-Basic:2"[THEN "\E"(2)], OF "&I"] + act_a_num_G by blast + AOT_hence \x = a\ by (metis "id-act:1" "\E"(2)) + AOT_hence \a = x\ using id_sym by auto + AOT_thus \Numbers(x, [\z \<^bold>\[G]z])\ + using "rule=E" num_a_act_g by fast + next + AOT_assume \Numbers(x, [\z \<^bold>\[G]z])\ + AOT_hence \a = x\ + using "pre-Hume"[unvarify G H, THEN "\E", OF act_den, OF act_den, OF "&I", + OF num_a_act_g, THEN "\E"(2)] + "eq-part:1"[unvarify F, OF act_den] by blast + AOT_thus \\<^bold>\Numbers(x, G)\ + using act_a_num_G "rule=E" by fast + qed +qed + +AOT_theorem "eq-num:2": \Numbers(x,[\z \<^bold>\[G]z]) \ x = #G\ +proof - + AOT_have 0: \\<^bold>\\<^sub>\ x = \<^bold>\x Numbers(x, G) \ \y (Numbers(y, [\z \<^bold>\[G]z]) \ y = x)\ for x + by (AOT_subst (reverse) \Numbers(x, [\z \<^bold>\[G]z])\ \\<^bold>\Numbers(x, G)\ for: x) + (auto simp: "eq-num:1" descriptions[axiom_inst]) + AOT_have \#G = \<^bold>\x Numbers(x, G) \ \y (Numbers(y, [\z \<^bold>\[G]z]) \ y = #G)\ + using 0[unvarify x, OF "num-def:2"]. + moreover AOT_have \#G = \<^bold>\x Numbers(x, G)\ + using "num-def:1" "num-uniq" "rule-id-df:1" by blast + ultimately AOT_have \\y (Numbers(y, [\z \<^bold>\[G]z]) \ y = #G)\ + using "\E" by blast + thus ?thesis using "\E"(2) by blast +qed + +AOT_theorem "eq-num:3": \Numbers(#G, [\y \<^bold>\[G]y])\ +proof - + AOT_have \#G = #G\ + by (simp add: "rule=I:1" "num-def:2") + thus ?thesis + using "eq-num:2"[unvarify x, OF "num-def:2", THEN "\E"(2)] by blast +qed + +AOT_theorem "eq-num:4": + \A!#G & \F (#G[F] \ [\z \<^bold>\[F]z] \\<^sub>E [\z \<^bold>\[G]z])\ + by (auto intro!: "&I" "eq-num:3"[THEN numbers[THEN "\\<^sub>d\<^sub>fE"], + THEN "&E"(1), THEN "&E"(1)] + "eq-num:3"[THEN numbers[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(2)]) + +AOT_theorem "eq-num:5": \#G[G]\ + by (auto intro!: "eq-num:4"[THEN "&E"(2), THEN "\E"(2), THEN "\E"(2)] + "eq-part:1"[unvarify F] simp: "cqt:2") + +AOT_theorem "eq-num:6": \Numbers(x, G) \ NaturalCardinal(x)\ +proof(rule "\I") + AOT_have act_den: \\<^bold>\\<^sub>\ [\z \<^bold>\[F]z]\\ for F + by "cqt:2" + AOT_obtain F where \Rigidifies(F, G)\ + by (metis "instantiation" "rigid-der:3") + AOT_hence \: \Rigid(F)\ and \\x([F]x \ [G]x)\ + using "df-rigid-rel:2"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2)] + "df-rigid-rel:2"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(1)] + by blast+ + AOT_hence \F \\<^sub>E G\ + by (auto intro!: eqE[THEN "\\<^sub>d\<^sub>fI"] "&I" "cqt:2" GEN "\I" elim: "\E"(2)) + moreover AOT_assume \Numbers(x, G)\ + ultimately AOT_have \Numbers(x, F)\ + using "num-tran:3"[THEN "\E", THEN "\E"(2)] by blast + moreover AOT_have \F \\<^sub>E [\z \<^bold>\[F]z]\ + using \ "approx-nec:1" "\E" by blast + ultimately AOT_have \Numbers(x, [\z \<^bold>\[F]z])\ + using "num-tran:1"[unvarify H, OF act_den, THEN "\E", THEN "\E"(1)] by blast + AOT_hence \x = #F\ + using "eq-num:2"[THEN "\E"(1)] by blast + AOT_hence \\F x = #F\ + by (rule "\I") + AOT_thus \NaturalCardinal(x)\ + using card[THEN "\\<^sub>d\<^sub>fI"] by blast +qed + +AOT_theorem "eq-df-num": \\G (x = #G) \ \G (Numbers(x,G))\ +proof(safe intro!: "\I" "\I") + AOT_assume \\G (x = #G)\ + then AOT_obtain P where \x = #P\ + using "\E"[rotated] by blast + AOT_hence \Numbers(x,[\z \<^bold>\[P]z])\ + using "eq-num:2"[THEN "\E"(2)] by blast + moreover AOT_have \[\z \<^bold>\[P]z]\\ by "cqt:2" + ultimately AOT_show \\G(Numbers(x,G))\ by (rule "\I") +next + AOT_assume \\G (Numbers(x,G))\ + then AOT_obtain Q where \Numbers(x,Q)\ + using "\E"[rotated] by blast + AOT_hence \NaturalCardinal(x)\ + using "eq-num:6"[THEN "\E"] by blast + AOT_thus \\G (x = #G)\ + using card[THEN "\\<^sub>d\<^sub>fE"] by blast +qed + +AOT_theorem "card-en": \NaturalCardinal(x) \ \F(x[F] \ x = #F)\ +proof(rule "\I"; rule GEN) + AOT_have act_den: \\<^bold>\\<^sub>\ [\z \<^bold>\[F]z]\\ for F by "cqt:2" + fix F + AOT_assume \NaturalCardinal(x)\ + AOT_hence \\F x = #F\ + using card[THEN "\\<^sub>d\<^sub>fE"] by blast + then AOT_obtain P where x_def: \x = #P\ + using "\E"[rotated] by blast + AOT_hence num_x_act_P: \Numbers(x,[\z \<^bold>\[P]z])\ + using "eq-num:2"[THEN "\E"(2)] by blast + AOT_have \#P[F] \ [\z \<^bold>\[F]z] \\<^sub>E [\z \<^bold>\[P]z]\ + using "eq-num:4"[THEN "&E"(2), THEN "\E"(2)] by blast + AOT_hence \x[F] \ [\z \<^bold>\[F]z] \\<^sub>E [\z \<^bold>\[P]z]\ + using x_def[symmetric] "rule=E" by fast + also AOT_have \\ \ Numbers(x, [\z \<^bold>\[F]z])\ + using "num-tran:1"[unvarify G H, OF act_den, OF act_den] + using "num-tran:2"[unvarify G H, OF act_den, OF act_den] + by (metis "&I" "deduction-theorem" "\I" "\E"(2) num_x_act_P) + also AOT_have \\ \ x = #F\ + using "eq-num:2" by blast + finally AOT_show \x[F] \ x = #F\. +qed + +AOT_theorem "0F:1": \\\u [F]u \ Numbers(0, F)\ +proof - + AOT_have unotEu_act_ord: \\\v[\x O!x & \<^bold>\x \\<^sub>E x]v\ + proof(rule "raa-cor:2") + AOT_assume \\v[\x O!x & \<^bold>\x \\<^sub>E x]v\ + then AOT_obtain y where \[\x O!x & \<^bold>\x \\<^sub>E x]y\ + using "\E"[rotated] "&E" by blast + AOT_hence 0: \O!y & \<^bold>\y \\<^sub>E y\ + by (rule "\\C"(1)) + AOT_have \\<^bold>\\(y =\<^sub>E y)\ + apply (AOT_subst \\(y =\<^sub>E y)\ \y \\<^sub>E y\) + apply (meson "\E"(2) "Commutativity of \" "thm-neg=E") + by (fact 0[THEN "&E"(2)]) + AOT_hence \\(y =\<^sub>E y)\ + by (metis "\\I" "Act-Sub:1" "id-act2:1" "\E"(4)) + moreover AOT_have \y =\<^sub>E y\ + by (metis 0[THEN "&E"(1)] "ord=Eequiv:1" "\E") + ultimately AOT_show \p & \p\ for p + by (metis "raa-cor:3") + qed + AOT_have \Numbers(0, [\y \<^bold>\[\x O!x & x \\<^sub>E x]y])\ + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF "zero:1"]) + apply (rule "num-def:2"[unvarify G]; "cqt:2") + apply (rule "eq-num:3"[unvarify G]) + by "cqt:2[lambda]" + AOT_hence numbers0: \Numbers(0, [\x [O!]x & \<^bold>\x \\<^sub>E x])\ + proof (rule "num-tran:3"[unvarify x G H, THEN "\E", THEN "\E"(1), rotated 4]) + AOT_show \[\y \<^bold>\[\x O!x & x \\<^sub>E x]y] \\<^sub>E [\x [O!]x & \<^bold>\x \\<^sub>E x]\ + proof (safe intro!: eqE[THEN "\\<^sub>d\<^sub>fI"] "&I" Ordinary.GEN "\I" "cqt:2") + fix u + AOT_have \[\y \<^bold>\[\x O!x & x \\<^sub>E x]y]u \ \<^bold>\[\x O!x & x \\<^sub>E x]u\ + by (rule "beta-C-meta"[THEN "\E"]; "cqt:2[lambda]") + also AOT_have \\ \ \<^bold>\(O!u & u \\<^sub>E u)\ + apply (AOT_subst \[\x O!x & x \\<^sub>E x]u\ \O!u & u \\<^sub>E u\) + apply (rule "beta-C-meta"[THEN "\E"]; "cqt:2[lambda]") + by (simp add: "oth-class-taut:3:a") + also AOT_have \\ \ (\<^bold>\O!u & \<^bold>\u \\<^sub>E u)\ + by (simp add: "Act-Basic:2") + also AOT_have \\ \ (O!u & \<^bold>\u \\<^sub>E u)\ + by (metis Ordinary.\ "&I" "&E"(2) "\I" "\I" "\E"(1) "oa-facts:7") + also AOT_have \\ \ [\x [O!]x & \<^bold>\x \\<^sub>E x]u\ + by (rule "beta-C-meta"[THEN "\E", symmetric]; "cqt:2[lambda]") + finally AOT_show \[\y \<^bold>\[\x O!x & x \\<^sub>E x]y]u \ [\x [O!]x & \<^bold>\x \\<^sub>E x]u\. + qed + qed(fact "zero:2" | "cqt:2")+ + show ?thesis + proof(safe intro!: "\I" "\I") + AOT_assume \\\u [F]u\ + moreover AOT_have \\\v [\x [O!]x & \<^bold>\x \\<^sub>E x]v\ + using unotEu_act_ord. + ultimately AOT_have 0: \F \\<^sub>E [\x [O!]x & \<^bold>\x \\<^sub>E x]\ + by (rule "empty-approx:1"[unvarify H, THEN "\E", rotated, OF "&I"]) "cqt:2" + AOT_thus \Numbers(0, F)\ + by (rule "num-tran:1"[unvarify x H, THEN "\E", + THEN "\E"(2), rotated, rotated]) + (fact "zero:2" numbers0 | "cqt:2[lambda]")+ + next + AOT_assume \Numbers(0, F)\ + AOT_hence 1: \F \\<^sub>E [\x [O!]x & \<^bold>\x \\<^sub>E x]\ + by (rule "num-tran:2"[unvarify x H, THEN "\E", rotated 2, OF "&I"]) + (fact numbers0 "zero:2" | "cqt:2[lambda]")+ + AOT_show \\\u [F]u\ + proof(rule "raa-cor:2") + AOT_have 0: \[\x [O!]x & \<^bold>\x \\<^sub>E x]\\ by "cqt:2[lambda]" + AOT_assume \\u [F]u\ + AOT_hence \\(F \\<^sub>E [\x [O!]x & \<^bold>\x \\<^sub>E x])\ + by (rule "empty-approx:2"[unvarify H, OF 0, THEN "\E", OF "&I"]) + (rule unotEu_act_ord) + AOT_thus \F \\<^sub>E [\x [O!]x & \<^bold>\x \\<^sub>E x] & \(F \\<^sub>E [\x [O!]x & \<^bold>\x \\<^sub>E x])\ + using 1 "&I" by blast + qed + qed +qed + +AOT_theorem "0F:2": \\\u \<^bold>\[F]u \ #F = 0\ +proof(rule "\I"; rule "\I") + AOT_assume 0: \\\u \<^bold>\[F]u\ + AOT_have \\\u [\z \<^bold>\[F]z]u\ + proof(rule "raa-cor:2") + AOT_assume \\u [\z \<^bold>\[F]z]u\ + then AOT_obtain u where \[\z \<^bold>\[F]z]u\ + using "Ordinary.\E"[rotated] by blast + AOT_hence \\<^bold>\[F]u\ + by (metis "betaC:1:a") + AOT_hence \\u \<^bold>\[F]u\ + by (rule "Ordinary.\I") + AOT_thus \\u \<^bold>\[F]u & \\u \<^bold>\[F]u\ + using 0 "&I" by blast + qed + AOT_hence \Numbers(0,[\z \<^bold>\[F]z])\ + by (safe intro!: "0F:1"[unvarify F, THEN "\E"(1)]) "cqt:2" + AOT_hence \0 = #F\ + by (rule "eq-num:2"[unvarify x, OF "zero:2", THEN "\E"(1)]) + AOT_thus \#F = 0\ using id_sym by blast +next + AOT_assume \#F = 0\ + AOT_hence \0 = #F\ using id_sym by blast + AOT_hence \Numbers(0,[\z \<^bold>\[F]z])\ + by (rule "eq-num:2"[unvarify x, OF "zero:2", THEN "\E"(2)]) + AOT_hence 0: \\\u [\z \<^bold>\[F]z]u\ + by (safe intro!: "0F:1"[unvarify F, THEN "\E"(2)]) "cqt:2" + AOT_show \\\u \<^bold>\[F]u\ + proof(rule "raa-cor:2") + AOT_assume \\u \<^bold>\[F]u\ + then AOT_obtain u where \\<^bold>\[F]u\ + using "Ordinary.\E"[rotated] by meson + AOT_hence \[\z \<^bold>\[F]z]u\ + by (auto intro!: "\\C" "cqt:2") + AOT_hence \\u [\z \<^bold>\[F]z]u\ + using "Ordinary.\I" by blast + AOT_thus \\u [\z \<^bold>\[F]z]u & \\u [\z \<^bold>\[F]z]u\ + using "&I" 0 by blast + qed +qed + +AOT_theorem "0F:3": \\\\u [F]u \ #F = 0\ +proof(rule "\I") + AOT_assume \\\\u [F]u\ + AOT_hence 0: \\\\u [F]u\ + using "KBasic2:1" "\E"(1) by blast + AOT_have \\\u [\z \<^bold>\[F]z]u\ + proof(rule "raa-cor:2") + AOT_assume \\u [\z \<^bold>\[F]z]u\ + then AOT_obtain u where \[\z \<^bold>\[F]z]u\ + using "Ordinary.\E"[rotated] by blast + AOT_hence \\<^bold>\[F]u\ + by (metis "betaC:1:a") + AOT_hence \\[F]u\ + by (metis "Act-Sub:3" "\E") + AOT_hence \\u \[F]u\ + by (rule "Ordinary.\I") + AOT_hence \\\u [F]u\ + using "Ordinary.res-var-bound-reas[CBF\]"[THEN "\E"] by blast + AOT_thus \\\u [F]u & \\\u [F]u\ + using 0 "&I" by blast + qed + AOT_hence \Numbers(0,[\z \<^bold>\[F]z])\ + by (safe intro!: "0F:1"[unvarify F, THEN "\E"(1)]) "cqt:2" + AOT_hence \0 = #F\ + by (rule "eq-num:2"[unvarify x, OF "zero:2", THEN "\E"(1)]) + AOT_thus \#F = 0\ using id_sym by blast +qed + +AOT_theorem "0F:4": \w \ \\u [F]u \ #[F]\<^sub>w = 0\ +proof (rule "rule-id-df:2:b"[OF "w-index", where \\<^sub>1\\<^sub>n="(_,_)", simplified]) + AOT_show \[\x\<^sub>1...x\<^sub>n w \ [F]x\<^sub>1...x\<^sub>n]\\ + by (simp add: "w-rel:3") +next + AOT_show \w \ \\u [F]u \ #[\x w \ [F]x] = 0\ + proof (rule "\I"; rule "\I") + AOT_assume \w \ \\u [F]u\ + AOT_hence 0: \\w \ \u [F]u\ + using "coherent:1"[unvarify p, OF "log-prop-prop:2", THEN "\E"(1)] by blast + AOT_have \\\u \<^bold>\[\x w \ [F]x]u\ + proof(rule "raa-cor:2") + AOT_assume \\u \<^bold>\[\x w \ [F]x]u\ + then AOT_obtain u where \\<^bold>\[\x w \ [F]x]u\ + using "Ordinary.\E"[rotated] by meson + AOT_hence \\<^bold>\w \ [F]u\ + by (AOT_subst (reverse) \w \ [F]u\ \[\x w \ [F]x]u\; + safe intro!: "beta-C-meta"[THEN "\E"] "w-rel:1"[THEN "\E"]) + "cqt:2" + AOT_hence 1: \w \ [F]u\ + using "rigid-truth-at:4"[unvarify p, OF "log-prop-prop:2", THEN "\E"(1)] + by blast + AOT_have \\([F]u \ \u [F]u)\ + using "Ordinary.\I" "\I" RN by simp + AOT_hence \w \ ([F]u \ \u [F]u)\ + using "fund:2"[unvarify p, OF "log-prop-prop:2", THEN "\E"(1)] + "PossibleWorld.\E" by fast + AOT_hence \w \ \u [F]u\ + using 1 "conj-dist-w:2"[unvarify p q, OF "log-prop-prop:2", + OF "log-prop-prop:2", THEN "\E"(1), + THEN "\E"] by blast + AOT_thus \w \ \u [F]u & \w \ \u [F]u\ + using 0 "&I" by blast + qed + AOT_thus \#[\x w \ [F]x] = 0\ + by (safe intro!: "0F:2"[unvarify F, THEN "\E"(1)] "w-rel:1"[THEN "\E"]) + "cqt:2" + next + AOT_assume \#[\x w \ [F]x] = 0\ + AOT_hence 0: \\\u \<^bold>\[\x w \ [F]x]u\ + by (safe intro!: "0F:2"[unvarify F, THEN "\E"(2)] "w-rel:1"[THEN "\E"]) + "cqt:2" + AOT_have \\w \ \u [F]u\ + proof (rule "raa-cor:2") + AOT_assume \w \ \u [F]u\ + AOT_hence \\x w \ (O!x & [F]x)\ + using "conj-dist-w:6"[THEN "\E"(1)] by fast + then AOT_obtain x where \w \ (O!x & [F]x)\ + using "\E"[rotated] by blast + AOT_hence \w \ O!x\ and Fx_in_w: \w \ [F]x\ + using "conj-dist-w:1"[unvarify p q] "\E"(1) "log-prop-prop:2" + "&E" by blast+ + AOT_hence \\O!x\ + using "fund:1"[unvarify p, OF "log-prop-prop:2", THEN "\E"(2)] + "PossibleWorld.\I" by simp + AOT_hence ord_x: \O!x\ + using "oa-facts:3"[THEN "\E"] by blast + AOT_have \\<^bold>\w \ [F]x\ + using "rigid-truth-at:4"[unvarify p, OF "log-prop-prop:2", THEN "\E"(2)] + Fx_in_w by blast + AOT_hence \\<^bold>\[\x w \ [F]x]x\ + by (AOT_subst \[\x w \ [F]x]x\ \w \ [F]x\; + safe intro!: "beta-C-meta"[THEN "\E"] "w-rel:1"[THEN "\E"]) "cqt:2" + AOT_hence \O!x & \<^bold>\[\x w \ [F]x]x\ + using ord_x "&I" by blast + AOT_hence \\x (O!x & \<^bold>\[\x w \ [F]x]x)\ + using "\I" by fast + AOT_thus \\u (\<^bold>\[\x w \ [F]x]u) & \\u \<^bold>\[\x w \ [F]x]u\ + using 0 "&I" by blast + qed + AOT_thus \w \ \\u[F]u\ + using "coherent:1"[unvarify p, OF "log-prop-prop:2", THEN "\E"(2)] by blast + qed +qed + +AOT_act_theorem "zero=:1": + \NaturalCardinal(x) \ \F (x[F] \ Numbers(x, F))\ +proof(safe intro!: "\I" GEN) + fix F + AOT_assume \NaturalCardinal(x)\ + AOT_hence \\F (x[F] \ x = #F)\ + by (metis "card-en" "\E") + AOT_hence 1: \x[F] \ x = #F\ + using "\E"(2) by blast + AOT_have 2: \x[F] \ x = \<^bold>\y(Numbers(y, F))\ + by (rule "num-def:1"[THEN "=\<^sub>d\<^sub>fE"(1)]) + (auto simp: 1 "num-uniq") + AOT_have \x = \<^bold>\y(Numbers(y, F)) \ Numbers(x, F)\ + using "y-in:1" by blast + moreover AOT_have \Numbers(x, F) \ x = \<^bold>\y(Numbers(y, F))\ + proof(rule "\I") + AOT_assume 1: \Numbers(x, F)\ + moreover AOT_obtain z where z_prop: \\y (Numbers(y, F) \ y = z)\ + using "num:2"[THEN "uniqueness:1"[THEN "\\<^sub>d\<^sub>fE"]] "\E"[rotated] "&E" by blast + ultimately AOT_have \x = z\ + using "\E"(2) "\E" by blast + AOT_hence \\y (Numbers(y, F) \ y = x)\ + using z_prop "rule=E" id_sym by fast + AOT_thus \x = \<^bold>\y(Numbers(y,F))\ + by (rule hintikka[THEN "\E"(2), OF "&I", rotated]) + (fact 1) + qed + ultimately AOT_have \x = \<^bold>\y(Numbers(y, F)) \ Numbers(x, F)\ + by (metis "\I") + AOT_thus \x[F] \ Numbers(x, F)\ + using 2 by (metis "\E"(5)) +qed + +AOT_act_theorem "zero=:2": \0[F] \ \\u[F]u\ +proof - + AOT_have \0[F] \ Numbers(0, F)\ + using "zero=:1"[unvarify x, OF "zero:2", THEN "\E", + OF "zero-card", THEN "\E"(2)]. + also AOT_have \\ \ \\u[F]u\ + using "0F:1"[symmetric]. + finally show ?thesis. +qed + +AOT_act_theorem "zero=:3": \\\u[F]u \ #F = 0\ +proof - + AOT_have \\\u[F]u \ 0[F]\ using "zero=:2"[symmetric]. + also AOT_have \\ \ 0 = #F\ + using "card-en"[unvarify x, OF "zero:2", THEN "\E", + OF "zero-card", THEN "\E"(2)]. + also AOT_have \\ \ #F = 0\ + by (simp add: "deduction-theorem" id_sym "\I") + finally show ?thesis. +qed + +AOT_define Hereditary :: \\ \ \ \ \\ (\Hereditary'(_,_')\) + "hered:1": + \Hereditary(F, R) \\<^sub>d\<^sub>f R\ & F\ & \x\y([R]xy \ ([F]x \ [F]y))\ + +AOT_theorem "hered:2": + \[\xy \F((\z([R]xz \ [F]z) & Hereditary(F,R)) \ [F]y)]\\ + by "cqt:2[lambda]" + +AOT_define StrongAncestral :: \\ \ \\ (\_\<^sup>*\) + "ances-df": + \R\<^sup>* =\<^sub>d\<^sub>f [\xy \F((\z([R]xz \ [F]z) & Hereditary(F,R)) \ [F]y)]\ + +AOT_theorem "ances": + \[R\<^sup>*]xy \ \F((\z([R]xz \ [F]z) & Hereditary(F,R)) \ [F]y)\ + apply (rule "=\<^sub>d\<^sub>fI"(1)[OF "ances-df"]) + apply "cqt:2[lambda]" + apply (rule "beta-C-meta"[THEN "\E", OF "hered:2", unvarify \\<^sub>1\\<^sub>n, + where \=\(_,_)\, simplified]) + by (simp add: "&I" "ex:1:a" prod_denotesI "rule-ui:3") + +AOT_theorem "anc-her:1": + \[R]xy \ [R\<^sup>*]xy\ +proof (safe intro!: "\I" ances[THEN "\E"(2)] GEN) + fix F + AOT_assume \\z ([R]xz \ [F]z) & Hereditary(F, R)\ + AOT_hence \[R]xy \ [F]y\ + using "\E"(2) "&E" by blast + moreover AOT_assume \[R]xy\ + ultimately AOT_show \[F]y\ + using "\E" by blast +qed + +AOT_theorem "anc-her:2": + \([R\<^sup>*]xy & \z([R]xz \ [F]z) & Hereditary(F,R)) \ [F]y\ +proof(rule "\I"; (frule "&E"(1); drule "&E"(2))+) + AOT_assume \[R\<^sup>*]xy\ + AOT_hence \(\z([R]xz \ [F]z) & Hereditary(F,R)) \ [F]y\ + using ances[THEN "\E"(1)] "\E"(2) by blast + moreover AOT_assume \\z([R]xz \ [F]z)\ + moreover AOT_assume \Hereditary(F,R)\ + ultimately AOT_show \[F]y\ + using "\E" "&I" by blast +qed + +AOT_theorem "anc-her:3": + \([F]x & [R\<^sup>*]xy & Hereditary(F, R)) \ [F]y\ +proof(rule "\I"; (frule "&E"(1); drule "&E"(2))+) + AOT_assume 1: \[F]x\ + AOT_assume 2: \Hereditary(F, R)\ + AOT_hence 3: \\x \y ([R]xy \ ([F]x \ [F]y))\ + using "hered:1"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + AOT_have \\z ([R]xz \ [F]z)\ + proof (rule GEN; rule "\I") + fix z + AOT_assume \[R]xz\ + moreover AOT_have \[R]xz \ ([F]x \ [F]z)\ + using 3 "\E"(2) by blast + ultimately AOT_show \[F]z\ + using 1 "\E" by blast + qed + moreover AOT_assume \[R\<^sup>*]xy\ + ultimately AOT_show \[F]y\ + by (auto intro!: 2 "anc-her:2"[THEN "\E"] "&I") +qed + +AOT_theorem "anc-her:4": \([R]xy & [R\<^sup>*]yz) \ [R\<^sup>*]xz\ +proof(rule "\I"; frule "&E"(1); drule "&E"(2)) + AOT_assume 0: \[R\<^sup>*]yz\ and 1: \[R]xy\ + AOT_show \[R\<^sup>*]xz\ + proof(safe intro!: ances[THEN "\E"(2)] GEN "&I" "\I"; + frule "&E"(1); drule "&E"(2)) + fix F + AOT_assume \\z ([R]xz \ [F]z)\ + AOT_hence 1: \[F]y\ + using 1 "\E"(2) "\E" by blast + AOT_assume 2: \Hereditary(F,R)\ + AOT_show \[F]z\ + by (rule "anc-her:3"[THEN "\E"]; auto intro!: "&I" 1 2 0) + qed +qed + +AOT_theorem "anc-her:5": \[R\<^sup>*]xy \ \z [R]zy\ +proof (rule "\I") + AOT_have 0: \[\y \x [R]xy]\\ by "cqt:2" + AOT_assume 1: \[R\<^sup>*]xy\ + AOT_have \[\y\x [R]xy]y\ + proof(rule "anc-her:2"[unvarify F, OF 0, THEN "\E"]; + safe intro!: "&I" GEN "\I" "hered:1"[THEN "\\<^sub>d\<^sub>fI"] "cqt:2" 0) + AOT_show \[R\<^sup>*]xy\ using 1. + next + fix z + AOT_assume \[R]xz\ + AOT_hence \\x [R]xz\ by (rule "\I") + AOT_thus \[\y\x [R]xy]z\ + by (auto intro!: "\\C"(1) "cqt:2") + next + fix x y + AOT_assume \[R]xy\ + AOT_hence \\x [R]xy\ by (rule "\I") + AOT_thus \[\y \x [R]xy]y\ + by (auto intro!: "\\C"(1) "cqt:2") + qed + AOT_thus \\z [R]zy\ + by (rule "\\C"(1)) +qed + +AOT_theorem "anc-her:6": \([R\<^sup>*]xy & [R\<^sup>*]yz) \ [R\<^sup>*]xz\ +proof (rule "\I"; frule "&E"(1); drule "&E"(2)) + AOT_assume \[R\<^sup>*]xy\ + AOT_hence \: \\z ([R]xz \ [F]z) & Hereditary(F,R) \ [F]y\ for F + using "\E"(2) ances[THEN "\E"(1)] by blast + AOT_assume \[R\<^sup>*]yz\ + AOT_hence \: \\z ([R]yz \ [F]z) & Hereditary(F,R) \ [F]z\ for F + using "\E"(2) ances[THEN "\E"(1)] by blast + AOT_show \[R\<^sup>*]xz\ + proof (rule ances[THEN "\E"(2)]; safe intro!: GEN "\I") + fix F + AOT_assume \: \\z ([R]xz \ [F]z) & Hereditary(F,R)\ + AOT_show \[F]z\ + proof (rule \[THEN "\E", OF "&I"]) + AOT_show \Hereditary(F,R)\ + using \[THEN "&E"(2)]. + next + AOT_show \\z ([R]yz \ [F]z)\ + proof(rule GEN; rule "\I") + fix z + AOT_assume \[R]yz\ + moreover AOT_have \[F]y\ + using \[THEN "\E", OF \]. + ultimately AOT_show \[F]z\ + using \[THEN "&E"(2), THEN "hered:1"[THEN "\\<^sub>d\<^sub>fE"], + THEN "&E"(2), THEN "\E"(2), THEN "\E"(2), + THEN "\E", THEN "\E"] + by blast + qed + qed + qed +qed + +AOT_define OneToOne :: \\ \ \\ (\1-1'(_')\) + "df-1-1:1": \1-1(R) \\<^sub>d\<^sub>f R\ & \x\y\z([R]xz & [R]yz \ x = y)\ + +AOT_define RigidOneToOne :: \\ \ \\ (\Rigid\<^sub>1\<^sub>-\<^sub>1'(_')\) + "df-1-1:2": \Rigid\<^sub>1\<^sub>-\<^sub>1(R) \\<^sub>d\<^sub>f 1-1(R) & Rigid(R)\ + +AOT_theorem "df-1-1:3": \Rigid\<^sub>1\<^sub>-\<^sub>1(R) \ \1-1(R)\ +proof(rule "\I") + AOT_assume \Rigid\<^sub>1\<^sub>-\<^sub>1(R)\ + AOT_hence \1-1(R)\ and RigidR: \Rigid(R)\ + using "df-1-1:2"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast+ + AOT_hence 1: \[R]xz & [R]yz \ x = y\ for x y z + using "df-1-1:1"[THEN "\\<^sub>d\<^sub>fE"] "&E"(2) "\E"(2) by blast + AOT_have 1: \[R]xz & [R]yz \ \x = y\ for x y z + by (AOT_subst (reverse) \\x = y\ \x = y\) + (auto simp: 1 "id-nec:2" "\I" "qml:2"[axiom_inst]) + AOT_have \\\x\<^sub>1...\x\<^sub>n ([R]x\<^sub>1...x\<^sub>n \ \[R]x\<^sub>1...x\<^sub>n)\ + using "df-rigid-rel:1"[THEN "\\<^sub>d\<^sub>fE", OF RigidR] "&E" by blast + AOT_hence \\x\<^sub>1...\x\<^sub>n \([R]x\<^sub>1...x\<^sub>n \ \[R]x\<^sub>1...x\<^sub>n)\ + using "CBF"[THEN "\E"] by fast + AOT_hence \\x\<^sub>1\x\<^sub>2 \([R]x\<^sub>1x\<^sub>2 \ \[R]x\<^sub>1x\<^sub>2)\ + using tuple_forall[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_hence \\([R]xy \ \[R]xy)\ for x y + using "\E"(2) by blast + AOT_hence \\(([R]xz \ \[R]xz) & ([R]yz \ \[R]yz))\ for x y z + by (metis "KBasic:3" "&I" "\E"(3) "raa-cor:3") + moreover AOT_have \\(([R]xz \ \[R]xz) & ([R]yz \ \[R]yz)) \ + \(([R]xz & [R]yz) \ \([R]xz & [R]yz))\ for x y z + by (rule RM) (metis "\I" "KBasic:3" "&I" "&E"(1) "&E"(2) "\E"(2) "\E") + ultimately AOT_have 2: \\(([R]xz & [R]yz) \ \([R]xz & [R]yz))\ for x y z + using "\E" by blast + AOT_hence 3: \\([R]xz & [R]yz \ x = y)\ for x y z + using "sc-eq-box-box:6"[THEN "\E", THEN "\E", OF 2, OF 1] by blast + AOT_hence 4: \\\x\y\z([R]xz & [R]yz \ x = y)\ + by (safe intro!: GEN BF[THEN "\E"] 3) + AOT_thus \\1-1(R)\ + by (AOT_subst_thm "df-1-1:1"[THEN "\Df", THEN "\S"(1), + OF "cqt:2[const_var]"[axiom_inst]]) +qed + +AOT_theorem "df-1-1:4": \\R(Rigid\<^sub>1\<^sub>-\<^sub>1(R) \ \Rigid\<^sub>1\<^sub>-\<^sub>1(R))\ +proof(rule GEN;rule "\I") +AOT_modally_strict { + fix R + AOT_assume 0: \Rigid\<^sub>1\<^sub>-\<^sub>1(R)\ + AOT_hence 1: \R\\ + by (meson "\\<^sub>d\<^sub>fE" "&E"(1) "df-1-1:1" "df-1-1:2") + AOT_hence 2: \\R\\ + using "exist-nec" "\E" by blast + AOT_have 4: \\1-1(R)\ + using "df-1-1:3"[unvarify R, OF 1, THEN "\E", OF 0]. + AOT_have \Rigid(R)\ + using 0 "\\<^sub>d\<^sub>fE"[OF "df-1-1:2"] "&E" by blast + AOT_hence \\\x\<^sub>1...\x\<^sub>n ([R]x\<^sub>1...x\<^sub>n \ \[R]x\<^sub>1...x\<^sub>n)\ + using "df-rigid-rel:1"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + AOT_hence \\\\x\<^sub>1...\x\<^sub>n ([R]x\<^sub>1...x\<^sub>n \ \[R]x\<^sub>1...x\<^sub>n)\ + by (metis "S5Basic:6" "\E"(1)) + AOT_hence \\Rigid(R)\ + apply (AOT_subst_def "df-rigid-rel:1") + using 2 "KBasic:3" "\S"(2) "\E"(2) by blast + AOT_thus \\Rigid\<^sub>1\<^sub>-\<^sub>1(R)\ + apply (AOT_subst_def "df-1-1:2") + using 4 "KBasic:3" "\S"(2) "\E"(2) by blast +} +qed + +AOT_define InDomainOf :: \\ \ \ \ \\ (\InDomainOf'(_,_')\) + "df-1-1:5": \InDomainOf(x, R) \\<^sub>d\<^sub>f \y [R]xy\ + +AOT_register_rigid_restricted_type + RigidOneToOneRelation: \Rigid\<^sub>1\<^sub>-\<^sub>1(\)\ +proof + AOT_modally_strict { + AOT_show \\\ Rigid\<^sub>1\<^sub>-\<^sub>1(\)\ + proof (rule "\I"(1)[where \=\\(=\<^sub>E)\\]) + AOT_show \Rigid\<^sub>1\<^sub>-\<^sub>1((=\<^sub>E))\ + proof (safe intro!: "df-1-1:2"[THEN "\\<^sub>d\<^sub>fI"] "&I" "df-1-1:1"[THEN "\\<^sub>d\<^sub>fI"] + GEN "\I" "df-rigid-rel:1"[THEN "\\<^sub>d\<^sub>fI"] "=E[denotes]") + fix x y z + AOT_assume \x =\<^sub>E z & y =\<^sub>E z\ + AOT_thus \x = y\ + by (metis "rule=E" "&E"(1) "Conjunction Simplification"(2) + "=E-simple:2" id_sym "\E") + next + AOT_have \\x\y \(x =\<^sub>E y \ \x =\<^sub>E y)\ + proof(rule GEN; rule GEN) + AOT_show \\(x =\<^sub>E y \ \x =\<^sub>E y)\ for x y + by (meson RN "deduction-theorem" "id-nec3:1" "\E"(1)) + qed + AOT_hence \\x\<^sub>1...\x\<^sub>n \([(=\<^sub>E)]x\<^sub>1...x\<^sub>n \ \[(=\<^sub>E)]x\<^sub>1...x\<^sub>n)\ + by (rule tuple_forall[THEN "\\<^sub>d\<^sub>fI"]) + AOT_thus \\\x\<^sub>1...\x\<^sub>n ([(=\<^sub>E)]x\<^sub>1...x\<^sub>n \ \[(=\<^sub>E)]x\<^sub>1...x\<^sub>n)\ + using BF[THEN "\E"] by fast + qed + qed(fact "=E[denotes]") + } +next + AOT_modally_strict { + AOT_show \Rigid\<^sub>1\<^sub>-\<^sub>1(\) \ \\\ for \ + proof(rule "\I") + AOT_assume \Rigid\<^sub>1\<^sub>-\<^sub>1(\)\ + AOT_hence \1-1(\)\ + using "df-1-1:2"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + AOT_thus \\\\ + using "df-1-1:1"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + qed + } +next + AOT_modally_strict { + AOT_show \\F(Rigid\<^sub>1\<^sub>-\<^sub>1(F) \ \Rigid\<^sub>1\<^sub>-\<^sub>1(F))\ + by (safe intro!: GEN "df-1-1:4"[THEN "\E"(2)]) + } +qed +AOT_register_variable_names + RigidOneToOneRelation: \ \ + +AOT_define IdentityRestrictedToDomain :: \\ \ \\ (\'(=\<^sub>_')\) + "id-d-R": \(=\<^sub>\) =\<^sub>d\<^sub>f [\xy \z ([\]xz & [\]yz)]\ + +syntax "_AOT_id_d_R_infix" :: \\ \ \ \ \ \ \\ ("(_ =\<^sub>_/ _)" [50, 51, 51] 50) +translations + "_AOT_id_d_R_infix \ \ \'" == + "CONST AOT_exe (CONST IdentityRestrictedToDomain \) (\,\')" + +AOT_theorem "id-R-thm:1": \x =\<^sub>\ y \ \z ([\]xz & [\]yz)\ +proof - + AOT_have 0: \[\xy \z ([\]xz & [\]yz)]\\ by "cqt:2" + show ?thesis + apply (rule "=\<^sub>d\<^sub>fI"(1)[OF "id-d-R"]) + apply (fact 0) + apply (rule "beta-C-meta"[THEN "\E", OF 0, unvarify \\<^sub>1\\<^sub>n, + where \=\(_,_)\, simplified]) + by (simp add: "&I" "ex:1:a" prod_denotesI "rule-ui:3") +qed + +AOT_theorem "id-R-thm:2": + \x =\<^sub>\ y \ (InDomainOf(x, \) & InDomainOf(y, \))\ +proof(rule "\I") + AOT_assume \x =\<^sub>\ y\ + AOT_hence \\z ([\]xz & [\]yz)\ + using "id-R-thm:1"[THEN "\E"(1)] by simp + then AOT_obtain z where z_prop: \[\]xz & [\]yz\ + using "\E"[rotated] by blast + AOT_show \InDomainOf(x, \) & InDomainOf(y, \)\ + proof (safe intro!: "&I" "df-1-1:5"[THEN "\\<^sub>d\<^sub>fI"]) + AOT_show \\y [\]xy\ + using z_prop[THEN "&E"(1)] "\I" by fast + next + AOT_show \\z [\]yz\ + using z_prop[THEN "&E"(2)] "\I" by fast + qed +qed + +AOT_theorem "id-R-thm:3": \x =\<^sub>\ y \ x = y\ +proof(rule "\I") + AOT_assume \x =\<^sub>\ y\ + AOT_hence \\z ([\]xz & [\]yz)\ + using "id-R-thm:1"[THEN "\E"(1)] by simp + then AOT_obtain z where z_prop: \[\]xz & [\]yz\ + using "\E"[rotated] by blast + AOT_thus \x = y\ + using "df-1-1:3"[THEN "\E", OF RigidOneToOneRelation.\, + THEN "qml:2"[axiom_inst, THEN "\E"], + THEN "\\<^sub>d\<^sub>fE"[OF "df-1-1:1"], THEN "&E"(2), + THEN "\E"(2), THEN "\E"(2), + THEN "\E"(2), THEN "\E"] + by blast +qed + +AOT_theorem "id-R-thm:4": + \(InDomainOf(x, \) \ InDomainOf(y, \)) \ (x =\<^sub>\ y \ x = y)\ +proof (rule "\I") + AOT_assume \InDomainOf(x, \) \ InDomainOf(y, \)\ + moreover { + AOT_assume \InDomainOf(x, \)\ + AOT_hence \\z [\]xz\ + by (metis "\\<^sub>d\<^sub>fE" "df-1-1:5") + then AOT_obtain z where z_prop: \[\]xz\ + using "\E"[rotated] by blast + AOT_have \x =\<^sub>\ y \ x = y\ + proof(safe intro!: "\I" "\I" "id-R-thm:3"[THEN "\E"]) + AOT_assume \x = y\ + AOT_hence \[\]yz\ + using z_prop "rule=E" by fast + AOT_hence \[\]xz & [\]yz\ + using z_prop "&I" by blast + AOT_hence \\z ([\]xz & [\]yz)\ + by (rule "\I") + AOT_thus \x =\<^sub>\ y\ + using "id-R-thm:1" "\E"(2) by blast + qed + } + moreover { + AOT_assume \InDomainOf(y, \)\ + AOT_hence \\z [\]yz\ + by (metis "\\<^sub>d\<^sub>fE" "df-1-1:5") + then AOT_obtain z where z_prop: \[\]yz\ + using "\E"[rotated] by blast + AOT_have \x =\<^sub>\ y \ x = y\ + proof(safe intro!: "\I" "\I" "id-R-thm:3"[THEN "\E"]) + AOT_assume \x = y\ + AOT_hence \[\]xz\ + using z_prop "rule=E" id_sym by fast + AOT_hence \[\]xz & [\]yz\ + using z_prop "&I" by blast + AOT_hence \\z ([\]xz & [\]yz)\ + by (rule "\I") + AOT_thus \x =\<^sub>\ y\ + using "id-R-thm:1" "\E"(2) by blast + qed + } + ultimately AOT_show \x =\<^sub>\ y \ x = y\ + by (metis "\E"(2) "raa-cor:1") +qed + +AOT_theorem "id-R-thm:5": \InDomainOf(x, \) \ x =\<^sub>\ x\ +proof (rule "\I") + AOT_assume \InDomainOf(x, \)\ + AOT_hence \\z [\]xz\ + by (metis "\\<^sub>d\<^sub>fE" "df-1-1:5") + then AOT_obtain z where z_prop: \[\]xz\ + using "\E"[rotated] by blast + AOT_hence \[\]xz & [\]xz\ + using "&I" by blast + AOT_hence \\z ([\]xz & [\]xz)\ + using "\I" by fast + AOT_thus \x =\<^sub>\ x\ + using "id-R-thm:1" "\E"(2) by blast +qed + +AOT_theorem "id-R-thm:6": \x =\<^sub>\ y \ y =\<^sub>\ x\ +proof(rule "\I") + AOT_assume 0: \x =\<^sub>\ y\ + AOT_hence 1: \InDomainOf(x,\) & InDomainOf(y,\)\ + using "id-R-thm:2"[THEN "\E"] by blast + AOT_hence \x =\<^sub>\ y \ x = y\ + using "id-R-thm:4"[THEN "\E", OF "\I"(1)] "&E" by blast + AOT_hence \x = y\ + using 0 by (metis "\E"(1)) + AOT_hence \y = x\ + using id_sym by blast + moreover AOT_have \y =\<^sub>\ x \ y = x\ + using "id-R-thm:4"[THEN "\E", OF "\I"(2)] 1 "&E" by blast + ultimately AOT_show \y =\<^sub>\ x\ + by (metis "\E"(2)) +qed + +AOT_theorem "id-R-thm:7": \x =\<^sub>\ y & y =\<^sub>\ z \ x =\<^sub>\ z\ +proof (rule "\I"; frule "&E"(1); drule "&E"(2)) + AOT_assume 0: \x =\<^sub>\ y\ + AOT_hence 1: \InDomainOf(x,\) & InDomainOf(y,\)\ + using "id-R-thm:2"[THEN "\E"] by blast + AOT_hence \x =\<^sub>\ y \ x = y\ + using "id-R-thm:4"[THEN "\E", OF "\I"(1)] "&E" by blast + AOT_hence x_eq_y: \x = y\ + using 0 by (metis "\E"(1)) + AOT_assume 2: \y =\<^sub>\ z\ + AOT_hence 3: \InDomainOf(y,\) & InDomainOf(z,\)\ + using "id-R-thm:2"[THEN "\E"] by blast + AOT_hence \y =\<^sub>\ z \ y = z\ + using "id-R-thm:4"[THEN "\E", OF "\I"(1)] "&E" by blast + AOT_hence \y = z\ + using 2 by (metis "\E"(1)) + AOT_hence x_eq_z: \x = z\ + using x_eq_y id_trans by blast + AOT_have \InDomainOf(x,\) & InDomainOf(z,\)\ + using 1 3 "&I" "&E" by meson + AOT_hence \x =\<^sub>\ z \ x = z\ + using "id-R-thm:4"[THEN "\E", OF "\I"(1)] "&E" by blast + AOT_thus \x =\<^sub>\ z\ + using x_eq_z "\E"(2) by blast +qed + +AOT_define WeakAncestral :: \\ \ \\ (\_\<^sup>+\) + "w-ances-df": \[\]\<^sup>+ =\<^sub>d\<^sub>f [\xy [\]\<^sup>*xy \ x =\<^sub>\ y]\ + +AOT_theorem "w-ances-df[den1]": \[\xy [\]\<^sup>*xy \ x =\<^sub>\ y]\\ + by "cqt:2" +AOT_theorem "w-ances-df[den2]": \[\]\<^sup>+\\ + using "w-ances-df[den1]" "=\<^sub>d\<^sub>fI"(1)[OF "w-ances-df"] by blast + +AOT_theorem "w-ances": \[\]\<^sup>+xy \ ([\]\<^sup>*xy \ x =\<^sub>\ y)\ +proof - + AOT_have 0: \[\xy [\\<^sup>*]xy \ x =\<^sub>\ y]\\ + by "cqt:2" + AOT_have 1: \\(AOT_term_of_var x,AOT_term_of_var y)\\\ + by (simp add: "&I" "ex:1:a" prod_denotesI "rule-ui:3") + have 2: \\[\\\<^sub>1...\\<^sub>n [\\<^sup>*]\\<^sub>1...\\<^sub>n \ [(=\<^sub>\)]\\<^sub>1...\\<^sub>n]xy\ = + \[\xy [\\<^sup>*]xy \ [(=\<^sub>\)]xy]xy\\ + by (simp add: cond_case_prod_eta) + show ?thesis + apply (rule "=\<^sub>d\<^sub>fI"(1)[OF "w-ances-df"]) + apply (fact "w-ances-df[den1]") + using "beta-C-meta"[THEN "\E", OF 0, unvarify \\<^sub>1\\<^sub>n, + where \=\(_,_)\, simplified, OF 1] 2 by simp +qed + +AOT_theorem "w-ances-her:1": \[\]xy \ [\]\<^sup>+xy\ +proof(rule "\I") + AOT_assume \[\]xy\ + AOT_hence \[\]\<^sup>*xy\ + using "anc-her:1"[THEN "\E"] by blast + AOT_thus \[\]\<^sup>+xy\ + using "w-ances"[THEN "\E"(2)] "\I" by blast +qed + +AOT_theorem "w-ances-her:2": + \[F]x & [\]\<^sup>+xy & Hereditary(F, \) \ [F]y\ +proof(rule "\I"; (frule "&E"(1); drule "&E"(2))+) + AOT_assume 0: \[F]x\ + AOT_assume 1: \Hereditary(F, \)\ + AOT_assume \[\]\<^sup>+xy\ + AOT_hence \[\]\<^sup>*xy \ x =\<^sub>\ y\ + using "w-ances"[THEN "\E"(1)] by simp + moreover { + AOT_assume \[\]\<^sup>*xy\ + AOT_hence \[F]y\ + using "anc-her:3"[THEN "\E", OF "&I", OF "&I"] 0 1 by blast + } + moreover { + AOT_assume \x =\<^sub>\ y\ + AOT_hence \x = y\ + using "id-R-thm:3"[THEN "\E"] by blast + AOT_hence \[F]y\ + using 0 "rule=E" by blast + } + ultimately AOT_show \[F]y\ + by (metis "\E"(3) "raa-cor:1") +qed + +AOT_theorem "w-ances-her:3": \([\]\<^sup>+xy & [\]yz) \ [\]\<^sup>*xz\ +proof(rule "\I"; frule "&E"(1); drule "&E"(2)) + AOT_assume \[\]\<^sup>+xy\ + moreover AOT_assume Ryz: \[\]yz\ + ultimately AOT_have \[\]\<^sup>*xy \ x =\<^sub>\ y\ + using "w-ances"[THEN "\E"(1)] by metis + moreover { + AOT_assume R_star_xy: \[\]\<^sup>*xy\ + AOT_have \[\]\<^sup>*xz\ + proof (safe intro!: ances[THEN "\E"(2)] "\I" GEN) + fix F + AOT_assume 0: \\z ([\]xz \ [F]z) & Hereditary(F,\)\ + AOT_hence \[F]y\ + using R_star_xy ances[THEN "\E"(1), OF R_star_xy, + THEN "\E"(2), THEN "\E"] by blast + AOT_thus \[F]z\ + using "hered:1"[THEN "\\<^sub>d\<^sub>fE", OF 0[THEN "&E"(2)], THEN "&E"(2)] + "\E"(2) "\E" Ryz by blast + qed + } + moreover { + AOT_assume \x =\<^sub>\ y\ + AOT_hence \x = y\ + using "id-R-thm:3"[THEN "\E"] by blast + AOT_hence \[\]xz\ + using Ryz "rule=E" id_sym by fast + AOT_hence \[\]\<^sup>*xz\ + by (metis "anc-her:1"[THEN "\E"]) + } + ultimately AOT_show \[\]\<^sup>*xz\ + by (metis "\E"(3) "raa-cor:1") +qed + +AOT_theorem "w-ances-her:4": \([\]\<^sup>*xy & [\]yz) \ [\]\<^sup>+xz\ +proof(rule "\I"; frule "&E"(1); drule "&E"(2)) + AOT_assume \[\]\<^sup>*xy\ + AOT_hence \[\]\<^sup>*xy \ x =\<^sub>\ y\ + using "\I" by blast + AOT_hence \[\]\<^sup>+xy\ + using "w-ances"[THEN "\E"(2)] by blast + moreover AOT_assume \[\]yz\ + ultimately AOT_have \[\]\<^sup>*xz\ + using "w-ances-her:3"[THEN "\E", OF "&I"] by simp + AOT_hence \[\]\<^sup>*xz \ x =\<^sub>\ z\ + using "\I" by blast + AOT_thus \[\]\<^sup>+xz\ + using "w-ances"[THEN "\E"(2)] by blast +qed + +AOT_theorem "w-ances-her:5": \([\]xy & [\]\<^sup>+yz) \ [\]\<^sup>*xz\ +proof(rule "\I"; frule "&E"(1); drule "&E"(2)) + AOT_assume 0: \[\]xy\ + AOT_assume \[\]\<^sup>+yz\ + AOT_hence \[\]\<^sup>*yz \ y =\<^sub>\ z\ + by (metis "\E"(1) "w-ances") + moreover { + AOT_assume \[\]\<^sup>*yz\ + AOT_hence \[\]\<^sup>*xz\ + using 0 by (metis "anc-her:4" Adjunction "\E") + } + moreover { + AOT_assume \y =\<^sub>\ z\ + AOT_hence \y = z\ + by (metis "id-R-thm:3" "\E") + AOT_hence \[\]xz\ + using 0 "rule=E" by fast + AOT_hence \[\]\<^sup>*xz\ + by (metis "anc-her:1" "\E") + } + ultimately AOT_show \[\]\<^sup>*xz\ by (metis "\E"(2) "reductio-aa:1") +qed + +AOT_theorem "w-ances-her:6": \([\]\<^sup>+xy & [\]\<^sup>+yz) \ [\]\<^sup>+xz\ +proof(rule "\I"; frule "&E"(1); drule "&E"(2)) + AOT_assume 0: \[\]\<^sup>+xy\ + AOT_hence 1: \[\]\<^sup>*xy \ x =\<^sub>\ y\ + by (metis "\E"(1) "w-ances") + AOT_assume 2: \[\]\<^sup>+yz\ + { + AOT_assume \x =\<^sub>\ y\ + AOT_hence \x = y\ + by (metis "id-R-thm:3" "\E") + AOT_hence \[\]\<^sup>+xz\ + using 2 "rule=E" id_sym by fast + } + moreover { + AOT_assume \\(x =\<^sub>\ y)\ + AOT_hence 3: \[\]\<^sup>*xy\ + using 1 by (metis "\E"(3)) + AOT_have \[\]\<^sup>*yz \ y =\<^sub>\ z\ + using 2 by (metis "\E"(1) "w-ances") + moreover { + AOT_assume \[\]\<^sup>*yz\ + AOT_hence \[\]\<^sup>*xz\ + using 3 by (metis "anc-her:6" Adjunction "\E") + AOT_hence \[\]\<^sup>+xz\ + by (metis "\I"(1) "\E"(2) "w-ances") + } + moreover { + AOT_assume \y =\<^sub>\ z\ + AOT_hence \y = z\ + by (metis "id-R-thm:3" "\E") + AOT_hence \[\]\<^sup>+xz\ + using 0 "rule=E" id_sym by fast + } + ultimately AOT_have \[\]\<^sup>+xz\ + by (metis "\E"(3) "reductio-aa:1") + } + ultimately AOT_show \[\]\<^sup>+xz\ + by (metis "reductio-aa:1") +qed + +AOT_theorem "w-ances-her:7": \[\]\<^sup>*xy \ \z([\]\<^sup>+xz & [\]zy)\ +proof(rule "\I") + AOT_assume 0: \[\]\<^sup>*xy\ + AOT_have 1: \\z ([\]xz \ [\]z) & Hereditary(\,\) \ [\]y\ if \\\\ for \ + using ances[THEN "\E"(1), THEN "\E"(1), OF 0] that by blast + AOT_have \[\y \z([\]\<^sup>+xz & [\]zy)]y\ + proof (rule 1[THEN "\E"]; "cqt:2[lambda]"?; + safe intro!: "&I" GEN "\I" "hered:1"[THEN "\\<^sub>d\<^sub>fI"] "cqt:2") + fix z + AOT_assume 0: \[\]xz\ + AOT_hence \\z [\]xz\ by (rule "\I") + AOT_hence \InDomainOf(x, \)\ by (metis "\\<^sub>d\<^sub>fI" "df-1-1:5") + AOT_hence \x =\<^sub>\ x\ by (metis "id-R-thm:5" "\E") + AOT_hence \[\]\<^sup>+xx\ by (metis "\I"(2) "\E"(2) "w-ances") + AOT_hence \[\]\<^sup>+xx & [\]xz\ using 0 "&I" by blast + AOT_hence \\y ([\]\<^sup>+xy & [\]yz)\ by (rule "\I") + AOT_thus \[\y \z ([\]\<^sup>+xz & [\]zy)]z\ + by (auto intro!: "\\C"(1) "cqt:2") + next + fix x' y + AOT_assume Rx'y: \[\]x'y\ + AOT_assume \[\y \z ([\]\<^sup>+xz & [\]zy)]x'\ + AOT_hence \\z ([\]\<^sup>+xz & [\]zx')\ + using "\\C"(1) by blast + then AOT_obtain c where c_prop: \[\]\<^sup>+xc & [\]cx'\ + using "\E"[rotated] by blast + AOT_hence \[\]\<^sup>*xx'\ + by (meson Rx'y "anc-her:1" "anc-her:6" Adjunction "\E" "w-ances-her:3") + AOT_hence \[\]\<^sup>*xx' \ x =\<^sub>\ x'\ by (rule "\I") + AOT_hence \[\]\<^sup>+xx'\ by (metis "\E"(2) "w-ances") + AOT_hence \[\]\<^sup>+xx' & [\]x'y\ using Rx'y by (metis "&I") + AOT_hence \\z ([\]\<^sup>+xz & [\]zy)\ by (rule "\I") + AOT_thus \[\y \z ([\]\<^sup>+xz & [\]zy)]y\ + by (auto intro!: "\\C"(1) "cqt:2") + qed + AOT_thus \\z([\]\<^sup>+xz & [\]zy)\ + using "\\C"(1) by fast +qed + +AOT_theorem "1-1-R:1": \([\]xy & [\]\<^sup>*zy) \ [\]\<^sup>+zx\ +proof(rule "\I"; frule "&E"(1); drule "&E"(2)) + AOT_assume \[\]\<^sup>*zy\ + AOT_hence \\x ([\]\<^sup>+zx & [\]xy)\ + using "w-ances-her:7"[THEN "\E"] by simp + then AOT_obtain a where a_prop: \[\]\<^sup>+za & [\]ay\ + using "\E"[rotated] by blast + moreover AOT_assume \[\]xy\ + ultimately AOT_have \x = a\ + using "df-1-1:2"[THEN "\\<^sub>d\<^sub>fE", OF RigidOneToOneRelation.\, THEN "&E"(1), + THEN "\\<^sub>d\<^sub>fE"[OF "df-1-1:1"], THEN "&E"(2), THEN "\E"(2), + THEN "\E"(2), THEN "\E"(2), THEN "\E", OF "&I"] + "&E" by blast + AOT_thus \[\]\<^sup>+zx\ + using a_prop[THEN "&E"(1)] "rule=E" id_sym by fast +qed + +AOT_theorem "1-1-R:2": \[\]xy \ (\[\]\<^sup>*xx \ \[\]\<^sup>*yy)\ +proof(rule "\I"; rule "useful-tautologies:5"[THEN "\E"]; rule "\I") + AOT_assume 0: \[\]xy\ + moreover AOT_assume \[\]\<^sup>*yy\ + ultimately AOT_have \[\]\<^sup>+yx\ + using "1-1-R:1"[THEN "\E", OF "&I"] by blast + AOT_thus \[\]\<^sup>*xx\ + using 0 by (metis "&I" "\E" "w-ances-her:5") +qed + +AOT_theorem "1-1-R:3": \\[\]\<^sup>*xx \ ([\]\<^sup>+xy \ \[\]\<^sup>*yy)\ +proof(safe intro!: "\I") + AOT_have 0: \[\z \[\]\<^sup>*zz]\\ by "cqt:2" + AOT_assume 1: \\[\]\<^sup>*xx\ + AOT_assume 2: \[\]\<^sup>+xy\ + AOT_have \[\z \[\]\<^sup>*zz]y\ + proof(rule "w-ances-her:2"[unvarify F, OF 0, THEN "\E"]; + safe intro!: "&I" "hered:1"[THEN "\\<^sub>d\<^sub>fI"] "cqt:2" GEN "\I") + AOT_show \[\z \[\]\<^sup>*zz]x\ + by (auto intro!: "\\C"(1) "cqt:2" simp: 1) + next + AOT_show \[\]\<^sup>+xy\ by (fact 2) + next + fix x y + AOT_assume \[\z \[\\<^sup>*]zz]x\ + AOT_hence \\[\]\<^sup>*xx\ by (rule "\\C"(1)) + moreover AOT_assume \[\]xy\ + ultimately AOT_have \\[\]\<^sup>*yy\ + using "1-1-R:2"[THEN "\E", THEN "\E"] by blast + AOT_thus \[\z \[\\<^sup>*]zz]y\ + by (auto intro!: "\\C"(1) "cqt:2") + qed + AOT_thus \\[\]\<^sup>*yy\ + using "\\C"(1) by blast +qed + +AOT_theorem "1-1-R:4": \[\]\<^sup>*xy \ InDomainOf(x,\)\ +proof(rule "\I"; rule "df-1-1:5"[THEN "\\<^sub>d\<^sub>fI"]) + AOT_assume 1: \[\]\<^sup>*xy\ + AOT_have \[\z [\\<^sup>*]xz \ \y [\]xy]y\ + proof (safe intro!: "anc-her:2"[unvarify F, THEN "\E"]; + safe intro!: "cqt:2" "&I" GEN "\I" "hered:1"[THEN "\\<^sub>d\<^sub>fI"]) + AOT_show \[\]\<^sup>*xy\ by (fact 1) + next + fix z + AOT_assume \[\]xz\ + AOT_thus \[\z [\\<^sup>*]xz \ \y [\]xy]z\ + by (safe intro!: "\\C"(1) "cqt:2") + (meson "\I" "existential:2[const_var]") + next + fix x' y + AOT_assume Rx'y: \[\]x'y\ + AOT_assume \[\z [\\<^sup>*]xz \ \y [\]xy]x'\ + AOT_hence 0: \[\\<^sup>*]xx' \ \y [\]xy\ by (rule "\\C"(1)) + AOT_have 1: \[\\<^sup>*]xy \ \y [\]xy\ + proof(rule "\I") + AOT_assume \[\]\<^sup>*xy\ + AOT_hence \[\]\<^sup>+xx'\ by (metis Rx'y "&I" "1-1-R:1" "\E") + AOT_hence \[\]\<^sup>*xx' \ x =\<^sub>\ x'\ by (metis "\E"(1) "w-ances") + moreover { + AOT_assume \[\]\<^sup>*xx'\ + AOT_hence \\y [\]xy\ using 0 by (metis "\E") + } + moreover { + AOT_assume \x =\<^sub>\ x'\ + AOT_hence \x = x'\ by (metis "id-R-thm:3" "\E") + AOT_hence \[\]xy\ using Rx'y "rule=E" id_sym by fast + AOT_hence \\y [\]xy\ by (rule "\I") + } + ultimately AOT_show \\y [\]xy\ + by (metis "\E"(3) "reductio-aa:1") + qed + AOT_show \[\z [\\<^sup>*]xz \ \y [\]xy]y\ + by (auto intro!: "\\C"(1) "cqt:2" 1) + qed + AOT_hence \[\\<^sup>*]xy \ \y [\]xy\ by (rule "\\C"(1)) + AOT_thus \\y [\]xy\ using 1 "\E" by blast +qed + +AOT_theorem "1-1-R:5": \[\]\<^sup>+xy \ InDomainOf(x,\)\ +proof (rule "\I") + AOT_assume \[\]\<^sup>+xy\ + AOT_hence \[\]\<^sup>*xy \ x =\<^sub>\ y\ + by (metis "\E"(1) "w-ances") + moreover { + AOT_assume \[\]\<^sup>*xy\ + AOT_hence \InDomainOf(x,\)\ + using "1-1-R:4" "\E" by blast + } + moreover { + AOT_assume \x =\<^sub>\ y\ + AOT_hence \InDomainOf(x,\)\ + by (metis "Conjunction Simplification"(1) "id-R-thm:2" "\E") + } + ultimately AOT_show \InDomainOf(x,\)\ + by (metis "\E"(3) "reductio-aa:1") +qed + +AOT_theorem "pre-ind": + \([F]z & \x\y(([\]\<^sup>+zx & [\]\<^sup>+zy) \ ([\]xy \ ([F]x \ [F]y)))) \ + \x ([\]\<^sup>+zx \ [F]x)\ +proof(safe intro!: "\I" GEN) + AOT_have den: \[\y [F]y & [\]\<^sup>+zy]\\ by "cqt:2" + fix x + AOT_assume \: \[F]z & \x\y(([\]\<^sup>+zx & [\]\<^sup>+zy) \ ([\]xy \ ([F]x \ [F]y)))\ + AOT_assume 0: \[\]\<^sup>+zx\ + + AOT_have \[\y [F]y & [\]\<^sup>+zy]x\ + proof (rule "w-ances-her:2"[unvarify F, OF den, THEN "\E"]; safe intro!: "&I") + AOT_show \[\y [F]y & [\]\<^sup>+zy]z\ + proof (safe intro!: "\\C"(1) "cqt:2" "&I") + AOT_show \[F]z\ using \ "&E" by blast + next + AOT_show \[\]\<^sup>+zz\ + by (rule "w-ances"[THEN "\E"(2), OF "\I"(2)]) + (meson "0" "id-R-thm:5" "1-1-R:5" "\E") + qed + next + AOT_show \[\]\<^sup>+zx\ by (fact 0) + next + AOT_show \Hereditary([\y [F]y & [\]\<^sup>+zy],\)\ + proof (safe intro!: "hered:1"[THEN "\\<^sub>d\<^sub>fI"] "&I" "cqt:2" GEN "\I") + fix x' y + AOT_assume 1: \[\]x'y\ + AOT_assume \[\y [F]y & [\]\<^sup>+zy]x'\ + AOT_hence 2: \[F]x' & [\]\<^sup>+zx'\ by (rule "\\C"(1)) + AOT_have \[\]\<^sup>*zy\ using 1 2[THEN "&E"(2)] + by (metis Adjunction "modus-tollens:1" "reductio-aa:1" "w-ances-her:3") + AOT_hence 3: \[\]\<^sup>+zy\ by (metis "\I"(1) "\E"(2) "w-ances") + AOT_show \[\y [F]y & [\]\<^sup>+zy]y\ + proof (safe intro!: "\\C"(1) "cqt:2" "&I" 3) + AOT_show \[F]y\ + proof (rule \[THEN "&E"(2), THEN "\E"(2), THEN "\E"(2), + THEN "\E", THEN "\E", THEN "\E"]) + AOT_show \[\]\<^sup>+zx' & [\]\<^sup>+zy\ + using 2 3 "&E" "&I" by blast + next + AOT_show \[\]x'y\ by (fact 1) + next + AOT_show \[F]x'\ using 2 "&E" by blast + qed + qed + qed + qed + AOT_thus \[F]x\ using "\\C"(1) "&E"(1) by fast +qed + +text\The following is not part of PLM, but a theorem of AOT. + It states that the predecessor relation coexists with numbering a property. + We will use this fact to derive the predecessor axiom, which asserts that the + predecessor relation denotes, from the fact that our models validate that + numbering a property denotes.\ +AOT_theorem pred_coex: + \[\xy \F\u ([F]u & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>u))]\ \ \F ([\x Numbers(x,F)]\)\ +proof(safe intro!: "\I" "\I" GEN) + fix F + let ?P = \\[\xy \F\u ([F]u & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>u))]\\ + AOT_assume \[\?P\]\\ + AOT_hence \\[\?P\]\\ + using "exist-nec" "\E" by blast + moreover AOT_have + \\[\?P\]\ \ \(\x\y(\F([F]x \ [F]y) \ (Numbers(x,F) \ Numbers(y,F))))\ + proof(rule RM; safe intro!: "\I" GEN) + AOT_modally_strict { + fix x y + AOT_assume pred_den: \[\?P\]\\ + AOT_hence pred_equiv: + \[\?P\]xy \ \F\u ([F]u & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>u))\ for x y + by (safe intro!: "beta-C-meta"[unvarify \\<^sub>1\\<^sub>n, where \=\(_,_)\, THEN "\E", + rotated, OF pred_den, simplified] + tuple_denotes[THEN "\\<^sub>d\<^sub>fI"] "&I" "cqt:2") + text\We show as a subproof that any natural cardinal that is not zero + has a predecessor.\ + AOT_have CardinalPredecessor: + \\y [\?P\]yx\ if card_x: \NaturalCardinal(x)\ and x_nonzero: \x \ 0\ for x + proof - + AOT_have \\G x = #G\ + using card[THEN "\\<^sub>d\<^sub>fE", OF card_x]. + AOT_hence \\G Numbers(x,G)\ + using "eq-df-num"[THEN "\E"(1)] by blast + then AOT_obtain G' where numxG': \Numbers(x,G')\ + using "\E"[rotated] by blast + AOT_obtain G where \Rigidifies(G,G')\ + using "rigid-der:3" "\E"[rotated] by blast + + AOT_hence H: \Rigid(G) & \x ([G]x \ [G']x)\ + using "df-rigid-rel:2"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_have H_rigid: \\\x ([G]x \ \[G]x)\ + using H[THEN "&E"(1), THEN "df-rigid-rel:1"[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(2)]. + AOT_hence \\x \([G]x \ \[G]x)\ + using "CBF" "\E" by blast + AOT_hence R: \\([G]x \ \[G]x)\ for x using "\E"(2) by blast + AOT_hence rigid: \[G]x \ \<^bold>\[G]x\ for x + by (metis "\E"(6) "oth-class-taut:3:a" "sc-eq-fur:2" "\E") + AOT_have \G \\<^sub>E G'\ + proof (safe intro!: eqE[THEN "\\<^sub>d\<^sub>fI"] "&I" "cqt:2" GEN "\I") + AOT_show \[G]x \ [G']x\ for x using H[THEN "&E"(2)] "\E"(2) by fast + qed + AOT_hence \G \\<^sub>E G'\ + by (rule "apE-eqE:2"[THEN "\E", OF "&I", rotated]) + (simp add: "eq-part:1") + AOT_hence numxG: \Numbers(x,G)\ + using "num-tran:1"[THEN "\E", THEN "\E"(2)] numxG' by blast + + { + AOT_assume \\\y(y \ x & [\?P\]yx)\ + AOT_hence \\y \(y \ x & [\?P\]yx)\ + using "cqt-further:4" "\E" by blast + AOT_hence \\(y \ x & [\?P\]yx)\ for y + using "\E"(2) by blast + AOT_hence 0: \\y \ x \ \[\?P\]yx\ for y + using "\\E" "intro-elim:3:c" "oth-class-taut:5:a" by blast + { + fix y + AOT_assume \[\?P\]yx\ + AOT_hence \\y \ x\ + using 0 "\\I" "con-dis-i-e:4:c" by blast + AOT_hence \y = x\ + using "=-infix" "\\<^sub>d\<^sub>fI" "raa-cor:4" by blast + } note Pxy_imp_eq = this + AOT_have \[\?P\]xx\ + proof(rule "raa-cor:1") + AOT_assume notPxx: \\[\?P\]xx\ + AOT_hence \\\F\u([F]u & Numbers(x,F) & Numbers(x,[F]\<^sup>-\<^sup>u))\ + using pred_equiv "intro-elim:3:c" by blast + AOT_hence \\F \\u([F]u & Numbers(x,F) & Numbers(x,[F]\<^sup>-\<^sup>u))\ + using "cqt-further:4"[THEN "\E"] by blast + AOT_hence \\\u([F]u & Numbers(x,F) & Numbers(x,[F]\<^sup>-\<^sup>u))\ for F + using "\E"(2) by blast + AOT_hence \\y \(O!y & ([F]y & Numbers(x,F) & Numbers(x,[F]\<^sup>-\<^sup>y)))\ for F + using "cqt-further:4"[THEN "\E"] by blast + AOT_hence 0: \\(O!u & ([F]u & Numbers(x,F) & Numbers(x,[F]\<^sup>-\<^sup>u)))\ for F u + using "\E"(2) by blast + AOT_have \\\\u [G]u\ + proof(rule "raa-cor:1") + AOT_assume \\\\\u [G]u\ + AOT_hence \\\u [G]u\ + using "\\<^sub>d\<^sub>fI" "conventions:5" by blast + AOT_hence \\u \[G]u\ + by (metis "Ordinary.res-var-bound-reas[BF\]"[THEN "\E"]) + then AOT_obtain u where posGu: \\[G]u\ + using "Ordinary.\E"[rotated] by meson + AOT_hence Gu: \[G]u\ + by (meson "B\" "K\" "\E" R) + AOT_have \\([G]u & Numbers(x,G) & Numbers(x,[G]\<^sup>-\<^sup>u))\ + using 0 Ordinary.\ + by (metis "con-dis-i-e:1" "raa-cor:1") + AOT_hence notnumx: \\Numbers(x,[G]\<^sup>-\<^sup>u)\ + using Gu numxG "con-dis-i-e:1" "raa-cor:5" by metis + AOT_obtain y where numy: \Numbers(y,[G]\<^sup>-\<^sup>u)\ + using "num:1"[unvarify G, OF "F-u[den]"] "\E"[rotated] by blast + AOT_hence \[G]u & Numbers(x,G) & Numbers(y,[G]\<^sup>-\<^sup>u)\ + using Gu numxG "&I" by blast + AOT_hence \\u ([G]u & Numbers(x,G) & Numbers(y,[G]\<^sup>-\<^sup>u))\ + by (rule "Ordinary.\I") + AOT_hence \\G\u ([G]u & Numbers(x,G) & Numbers(y,[G]\<^sup>-\<^sup>u))\ + by (rule "\I") + AOT_hence \[\?P\]yx\ + using pred_equiv[THEN "\E"(2)] by blast + AOT_hence \y = x\ using Pxy_imp_eq by blast + AOT_hence \Numbers(x,[G]\<^sup>-\<^sup>u)\ + using numy "rule=E" by fast + AOT_thus \p & \p\ for p using notnumx "reductio-aa:1" by blast + qed + AOT_hence \\\u [G]u\ + using "qml:2"[axiom_inst, THEN "\E"] by blast + AOT_hence num0G: \Numbers(0, G)\ + using "0F:1"[THEN "\E"(1)] by blast + AOT_hence \x = 0\ + using "pre-Hume"[unvarify x, THEN "\E", OF "zero:2", OF "&I", + THEN "\E"(2), OF num0G, OF numxG, OF "eq-part:1"] + id_sym by blast + moreover AOT_have \\x = 0\ + using x_nonzero + using "=-infix" "\\<^sub>d\<^sub>fE" by blast + ultimately AOT_show \p & \p\ for p using "reductio-aa:1" by blast + qed + } + AOT_hence \[\?P\]xx \ \y (y \ x & [\?P\]yx)\ + using "con-dis-i-e:3:a" "con-dis-i-e:3:b" "raa-cor:1" by blast + moreover { + AOT_assume \[\?P\]xx\ + AOT_hence \\y [\?P\]yx\ + by (rule "\I") + } + moreover { + AOT_assume \\y (y \ x & [\?P\]yx)\ + then AOT_obtain y where \y \ x & [\?P\]yx\ + using "\E"[rotated] by blast + AOT_hence \[\?P\]yx\ + using "&E" by blast + AOT_hence \\y [\?P\]yx\ + by (rule "\I") + } + ultimately AOT_show \\y [\?P\]yx\ + using "\E"(1) "\I" by blast + qed + + text\Given above lemma, we can show that if one of two indistinguishable objects + numbers a property, the other one numbers this property as well.\ + AOT_assume indist: \\F([F]x \ [F]y)\ + AOT_assume numxF: \Numbers(x,F)\ + AOT_hence 0: \NaturalCardinal(x)\ + by (metis "eq-num:6" "vdash-properties:10") + text\We show by case distinction that x equals y. + As first case we consider x to be non-zero.\ + { + AOT_assume \\(x = 0)\ + AOT_hence \x \ 0\ + by (metis "=-infix" "\\<^sub>d\<^sub>fI") + AOT_hence \\y [\?P\]yx\ + using CardinalPredecessor 0 by blast + then AOT_obtain z where Pxz: \[\?P\]zx\ + using "\E"[rotated] by blast + AOT_hence \[\y [\?P\]zy]x\ + by (safe intro!: "\\C" "cqt:2") + AOT_hence \[\y [\?P\]zy]y\ + by (safe intro!: indist[THEN "\E"(1), THEN "\E"(1)] "cqt:2") + AOT_hence Pyz: \[\?P\]zy\ + using "\\C"(1) by blast + AOT_hence \\F\u ([F]u & Numbers(y,F) & Numbers(z,[F]\<^sup>-\<^sup>u))\ + using Pyz pred_equiv[THEN "\E"(1)] by blast + then AOT_obtain F\<^sub>1 where \\u ([F\<^sub>1]u & Numbers(y,F\<^sub>1) & Numbers(z,[F\<^sub>1]\<^sup>-\<^sup>u))\ + using "\E"[rotated] by blast + then AOT_obtain u where u_prop: \[F\<^sub>1]u & Numbers(y,F\<^sub>1) & Numbers(z,[F\<^sub>1]\<^sup>-\<^sup>u)\ + using "Ordinary.\E"[rotated] by meson + AOT_have \\F\u ([F]u & Numbers(x,F) & Numbers(z,[F]\<^sup>-\<^sup>u))\ + using Pxz pred_equiv[THEN "\E"(1)] by blast + then AOT_obtain F\<^sub>2 where \\u ([F\<^sub>2]u & Numbers(x,F\<^sub>2) & Numbers(z,[F\<^sub>2]\<^sup>-\<^sup>u))\ + using "\E"[rotated] by blast + then AOT_obtain v where v_prop: \[F\<^sub>2]v & Numbers(x,F\<^sub>2) & Numbers(z,[F\<^sub>2]\<^sup>-\<^sup>v)\ + using "Ordinary.\E"[rotated] by meson + AOT_have \[F\<^sub>2]\<^sup>-\<^sup>v \\<^sub>E [F\<^sub>1]\<^sup>-\<^sup>u\ + using "hume-strict:1"[unvarify F G, THEN "\E"(1), OF "F-u[den]", + OF "F-u[den]", OF "\I"(2)[where \=z], OF "&I"] + v_prop u_prop "&E" by blast + AOT_hence \F\<^sub>2 \\<^sub>E F\<^sub>1\ + using "P'-eq"[THEN "\E", OF "&I", OF "&I"] + u_prop v_prop "&E" by meson + AOT_hence \x = y\ + using "pre-Hume"[THEN "\E", THEN "\E"(2), OF "&I"] + v_prop u_prop "&E" by blast + } + text\The second case handles x being equal to zero.\ + moreover { + fix u + AOT_assume x_is_zero: \x = 0\ + moreover AOT_have \Numbers(0,[\z z =\<^sub>E u]\<^sup>-\<^sup>u)\ + proof (safe intro!: "0F:1"[unvarify F, THEN "\E"(1)] "cqt:2" "raa-cor:2" + "F-u[den]"[unvarify F]) + AOT_assume \\v [[\z z =\<^sub>E u]\<^sup>-\<^sup>u]v\ + then AOT_obtain v where \[[\z z =\<^sub>E u]\<^sup>-\<^sup>u]v\ + using "Ordinary.\E"[rotated] by meson + AOT_hence \[\z z =\<^sub>E u]v & v \\<^sub>E u\ + by (auto intro: "F-u"[THEN "=\<^sub>d\<^sub>fE"(1), where \\<^sub>1\\<^sub>n="(_,_)", simplified] + intro!: "cqt:2" "F-u[equiv]"[unvarify F, THEN "\E"(1)] + "F-u[den]"[unvarify F]) + AOT_thus \p & \p\ for p + using "\\C" "thm-neg=E"[THEN "\E"(1)] "&E" "&I" + "raa-cor:3" by fast + qed + ultimately AOT_have 0: \Numbers(x,[\z z =\<^sub>E u]\<^sup>-\<^sup>u)\ + using "rule=E" id_sym by fast + AOT_have \\y Numbers(y,[\z z =\<^sub>E u])\ + by (safe intro!: "num:1"[unvarify G] "cqt:2") + then AOT_obtain z where \Numbers(z,[\z z =\<^sub>E u])\ + using "\E" by metis + moreover AOT_have \[\z z=\<^sub>E u]u\ + by (safe intro!: "\\C" "cqt:2" "ord=Eequiv:1"[THEN "\E"] Ordinary.\) + ultimately AOT_have + 1: \[\z z=\<^sub>E u]u & Numbers(z,[\z z=\<^sub>E u]) & Numbers(x,[\z z=\<^sub>E u]\<^sup>-\<^sup>u)\ + using 0 "&I" by auto + AOT_hence \\v([\z z=\<^sub>E u]v & Numbers(z,[\z z =\<^sub>E u]) & Numbers(x,[\z z=\<^sub>E u]\<^sup>-\<^sup>v))\ + by (rule "Ordinary.\I") + AOT_hence \\F\u([F]u & Numbers(z,[F]) & Numbers(x,[F]\<^sup>-\<^sup>u))\ + by (rule "\I"; "cqt:2") + AOT_hence Px1: \[\?P\]xz\ + using "beta-C-cor:2"[THEN "\E", OF pred_den, + THEN tuple_forall[THEN "\\<^sub>d\<^sub>fE"], THEN "\E"(2), + THEN "\E"(2), THEN "\E"(2)] by simp + AOT_hence \[\y [\?P\]yz]x\ + by (safe intro!: "\\C" "cqt:2") + AOT_hence \[\y [\?P\]yz]y\ + by (safe intro!: indist[THEN "\E"(1), THEN "\E"(1)] "cqt:2") + AOT_hence Py1: \[\?P\]yz\ + using "\\C" by blast + AOT_hence \\F\u([F]u & Numbers(z,[F]) & Numbers(y,[F]\<^sup>-\<^sup>u))\ + using "\\C" by fast + then AOT_obtain G where \\u([G]u & Numbers(z,[G]) & Numbers(y,[G]\<^sup>-\<^sup>u))\ + using "\E"[rotated] by blast + then AOT_obtain v where 2: \[G]v & Numbers(z,[G]) & Numbers(y,[G]\<^sup>-\<^sup>v)\ + using "Ordinary.\E"[rotated] by meson + with 1 2 AOT_have \[\z z =\<^sub>E u] \\<^sub>E G\ + by (auto intro!: "hume-strict:1"[unvarify F, THEN "\E"(1), rotated, + OF "\I"(2)[where \=z], OF "&I"] "cqt:2" + dest: "&E") + AOT_hence 3: \[\z z =\<^sub>E u]\<^sup>-\<^sup>u \\<^sub>E [G]\<^sup>-\<^sup>v\ + using 1 2 + by (safe_step intro!: "eqP'"[unvarify F, THEN "\E"]) + (auto dest: "&E" intro!: "cqt:2" "&I") + with 1 2 AOT_have \x = y\ + by (auto intro!: "pre-Hume"[unvarify G H, THEN "\E", + THEN "\E"(2), rotated 3, OF 3] + "F-u[den]"[unvarify F] "cqt:2" "&I" + dest: "&E") + } + ultimately AOT_have \x = y\ + using "\E"(1) "\I" "reductio-aa:1" by blast + text\Now since x numbers F, so does y.\ + AOT_hence \Numbers(y,F)\ + using numxF "rule=E" by fast + } note 0 = this + text\The only thing left is to generalize this result to a biconditional.\ + AOT_modally_strict { + fix x y + AOT_assume \[\?P\]\\ + moreover AOT_assume \\F([F]x \ [F]y)\ + moreover AOT_have \\F([F]y \ [F]x)\ + by (metis "cqt-basic:11" "intro-elim:3:a" calculation(2)) + ultimately AOT_show \Numbers(x,F) \ Numbers(y,F)\ + using 0 "\I" "\I" by auto + } + qed + ultimately AOT_show \[\x Numbers(x,F)]\\ + using "kirchner-thm:1"[THEN "\E"(2)] "\E" by fast +next + text\The converse can be shown by coexistence.\ + AOT_assume \\F [\x Numbers(x,F)]\\ + AOT_hence \[\x Numbers(x,F)]\\ for F + using "\E"(2) by blast + AOT_hence \\[\x Numbers(x,F)]\\ for F + using "exist-nec"[THEN "\E"] by blast + AOT_hence \\F \[\x Numbers(x,F)]\\ + by (rule GEN) + AOT_hence \\\F [\x Numbers(x,F)]\\ + using BF[THEN "\E"] by fast + moreover AOT_have + \\\F [\x Numbers(x,F)]\ \ + \\x \y (\F \u ([F]u & [\z Numbers(z,F)]y & [\z Numbers(z,[F]\<^sup>-\<^sup>u)]x) \ + \F \u ([F]u & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>u)))\ + proof(rule RM; safe intro!: "\I" GEN) + AOT_modally_strict { + fix x y + AOT_assume 0: \\F [\x Numbers(x,F)]\\ + AOT_show \\F \u ([F]u & [\z Numbers(z,F)]y & [\z Numbers(z,[F]\<^sup>-\<^sup>u)]x) \ + \F \u ([F]u & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>u))\ + proof(safe intro!: "\I" "\I") + AOT_assume \\F \u ([F]u & [\z Numbers(z,F)]y & [\z Numbers(z,[F]\<^sup>-\<^sup>u)]x)\ + then AOT_obtain F where + \\u ([F]u & [\z Numbers(z,F)]y & [\z Numbers(z,[F]\<^sup>-\<^sup>u)]x)\ + using "\E"[rotated] by blast + then AOT_obtain u where \[F]u & [\z Numbers(z,F)]y & [\z Numbers(z,[F]\<^sup>-\<^sup>u)]x\ + using "Ordinary.\E"[rotated] by meson + AOT_hence \[F]u & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>u)\ + by (auto intro!: "&I" dest: "&E" "\\C") + AOT_thus \\F \u ([F]u & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>u))\ + using "\I" "Ordinary.\I" by fast + next + AOT_assume \\F \u ([F]u & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>u))\ + then AOT_obtain F where \\u ([F]u & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>u))\ + using "\E"[rotated] by blast + then AOT_obtain u where \[F]u & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>u)\ + using "Ordinary.\E"[rotated] by meson + AOT_hence \[F]u & [\z Numbers(z,F)]y & [\z Numbers(z,[F]\<^sup>-\<^sup>u)]x\ + by (auto intro!: "&I" "\\C" 0[THEN "\E"(1)] "F-u[den]" + dest: "&E" intro: "cqt:2") + AOT_hence \\u([F]u & [\z Numbers(z,F)]y & [\z Numbers(z,[F]\<^sup>-\<^sup>u)]x)\ + by (rule "Ordinary.\I") + AOT_thus \\F\u([F]u & [\z Numbers(z,F)]y & [\z Numbers(z,[F]\<^sup>-\<^sup>u)]x)\ + by (rule "\I") + qed + } + qed + ultimately AOT_have + \\\x \y (\F \u ([F]u & [\z Numbers(z,F)]y & [\z Numbers(z,[F]\<^sup>-\<^sup>u)]x) \ + \F \u ([F]u & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>u)))\ + using "\E" by blast + AOT_thus \[\xy \F \u ([F]u & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>u))]\\ + by (rule "safe-ext[2]"[axiom_inst, THEN "\E", OF "&I", rotated]) "cqt:2" +qed + +text\The following is not part of PLM, but a consequence of extended relation + comprehension and can be used to @{emph \derive\} the predecessor axiom.\ +AOT_theorem numbers_prop_den: \[\x Numbers(x,G)]\\ +proof (rule "safe-ext"[axiom_inst, THEN "\E", OF "&I"]) + AOT_show \[\x A!x & [\x \F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G)]x]\\ + by "cqt:2" +next + AOT_have 0: \\<^bold>\\<^sub>\ [\x \F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G)]\\ + proof(safe intro!: Comprehension_3[THEN "\E"] "\I" RN GEN) + AOT_modally_strict { + fix F H + AOT_assume \\H \\<^sub>E F\ + AOT_hence \\\u ([H]u \ [F]u)\ + by (AOT_subst (reverse) \\u ([H]u \ [F]u)\ \H \\<^sub>E F\) + (safe intro!: "eqE"[THEN "\Df", THEN "\S"(1), OF "&I"] "cqt:2") + AOT_hence \\u \([H]u \ [F]u)\ + by (metis "Ordinary.res-var-bound-reas[CBF]" "\E") + AOT_hence \\([H]u \ [F]u)\ for u + using "Ordinary.\E" by fast + AOT_hence \\<^bold>\([H]u \ [F]u)\ for u + by (metis "nec-imp-act" "\E") + AOT_hence \\<^bold>\([F]u \ [H]u)\ for u + by (metis "Act-Basic:5" "Commutativity of \" "intro-elim:3:b") + AOT_hence \[\z \<^bold>\[F]z] \\<^sub>E [\z \<^bold>\[H]z]\ + by (safe intro!: "eqE"[THEN "\\<^sub>d\<^sub>fI"] "&I" "cqt:2" Ordinary.GEN; + AOT_subst \[\z \<^bold>\[F]z]u\ \\<^bold>\[F]u\ for: u F) + (auto intro!: "beta-C-meta"[THEN "\E"] "cqt:2" + "Act-Basic:5"[THEN "\E"(1)]) + AOT_hence \[\z \<^bold>\[F]z] \\<^sub>E [\z \<^bold>\[H]z]\ + by (safe intro!: "apE-eqE:1"[unvarify F G, THEN "\E"] "cqt:2") + AOT_thus \[\z \<^bold>\[F]z] \\<^sub>E G \ [\z \<^bold>\[H]z] \\<^sub>E G\ + using "\I" "eq-part:2[terms]" "eq-part:3[terms]" "\E" "\I" + by metis + } + qed + AOT_show \\\x (A!x & [\x \F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G)]x \ Numbers(x,G))\ + proof (safe intro!: RN GEN) + AOT_modally_strict { + fix x + AOT_show \A!x & [\x \F (x[F] \ [\z \<^bold>\[F]z] \\<^sub>E G)]x \ Numbers(x,G)\ + by (AOT_subst_def numbers; AOT_subst_thm "beta-C-meta"[THEN "\E", OF 0]) + (auto intro!: "beta-C-meta"[THEN "\E", OF 0] "\I" "\I" "&I" "cqt:2" + dest: "&E") + } + qed +qed + +text\The two theorems above allow us to derive + the predecessor axiom of PLM as theorem.\ + +AOT_theorem pred: \[\xy \F\u ([F]u & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>u))]\\ + using pred_coex numbers_prop_den["\I" G] "\E" by blast + +AOT_define Predecessor :: \\\ (\\\) + "pred-thm:1": + \\ =\<^sub>d\<^sub>f [\xy \F\u ([F]u & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>u))]\ + +AOT_theorem "pred-thm:2": \\\\ + using pred "pred-thm:1" "rule-id-df:2:b[zero]" by blast + +AOT_theorem "pred-thm:3": + \[\]xy \ \F\u ([F]u & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>u))\ + by (auto intro!: "beta-C-meta"[unvarify \\<^sub>1\\<^sub>n, where \=\(_,_)\, THEN "\E", + rotated, OF pred, simplified] + tuple_denotes[THEN "\\<^sub>d\<^sub>fI"] "&I" "cqt:2" pred + intro: "=\<^sub>d\<^sub>fI"(2)[OF "pred-thm:1"]) + +AOT_theorem "pred-1-1:1": \[\]xy \ \[\]xy\ +proof(rule "\I") + AOT_assume \[\]xy\ + AOT_hence \\F\u ([F]u & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>u))\ + using "\E"(1) "pred-thm:3" by fast + then AOT_obtain F where \\u ([F]u & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>u))\ + using "\E"[rotated] by blast + then AOT_obtain u where props: \[F]u & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>u)\ + using "Ordinary.\E"[rotated] by meson + AOT_obtain G where Ridigifies_G_F: \Rigidifies(G, F)\ + by (metis "instantiation" "rigid-der:3") + AOT_hence \: \\\x([G]x \ \[G]x)\ and \: \\x([G]x \ [F]x)\ + using "df-rigid-rel:2"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(1), + THEN "\\<^sub>d\<^sub>fE"[OF "df-rigid-rel:1"], THEN "&E"(2)] + "df-rigid-rel:2"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2)] by blast+ + + AOT_have rigid_num_nec: \Numbers(x,F) & Rigidifies(G,F) \ \Numbers(x,G)\ + for x G F + proof(rule "\I"; frule "&E"(1); drule "&E"(2)) + fix G F x + AOT_assume Numbers_xF: \Numbers(x,F)\ + AOT_assume \Rigidifies(G,F)\ + AOT_hence \: \Rigid(G)\ and \: \\x([G]x \ [F]x)\ + using "df-rigid-rel:2"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast+ + AOT_thus \\Numbers(x,G)\ + proof (safe intro!: + "num-cont:2"[THEN "\E", OF \, THEN "qml:2"[axiom_inst, THEN "\E"], + THEN "\E"(2), THEN "\E"] + "num-tran:3"[THEN "\E", THEN "\E"(1), rotated, OF Numbers_xF] + eqE[THEN "\\<^sub>d\<^sub>fI"] + "&I" "cqt:2[const_var]"[axiom_inst] Ordinary.GEN "\I") + AOT_show \[F]u \ [G]u\ for u + using \[THEN "\E"(2)] by (metis "\E"(6) "oth-class-taut:3:a") + qed + qed + AOT_have \\Numbers(y,G)\ + using rigid_num_nec[THEN "\E", OF "&I", OF props[THEN "&E"(1), THEN "&E"(2)], + OF Ridigifies_G_F]. + moreover { + AOT_have \Rigidifies([G]\<^sup>-\<^sup>u, [F]\<^sup>-\<^sup>u)\ + proof (safe intro!: "df-rigid-rel:1"[THEN "\\<^sub>d\<^sub>fI"] "df-rigid-rel:2"[THEN "\\<^sub>d\<^sub>fI"] + "&I" "F-u[den]" GEN "\I" "\I") + AOT_have \\\x([G]x \ \[G]x) \ \\x([[G]\<^sup>-\<^sup>u]x \ \[[G]\<^sup>-\<^sup>u]x)\ + proof (rule RM; safe intro!: "\I" GEN) + AOT_modally_strict { + fix x + AOT_assume 0: \\x([G]x \ \[G]x)\ + AOT_assume 1: \[[G]\<^sup>-\<^sup>u]x\ + AOT_have \[\x [G]x & x \\<^sub>E u]x\ + apply (rule "F-u"[THEN "=\<^sub>d\<^sub>fE"(1), where \\<^sub>1\\<^sub>n="(_,_)", simplified]) + apply "cqt:2[lambda]" + by (fact 1) + AOT_hence \[G]x & x \\<^sub>E u\ + by (rule "\\C"(1)) + AOT_hence 2: \\[G]x\ and 3: \\x \\<^sub>E u\ + using "&E" 0[THEN "\E"(2), THEN "\E"] "id-nec4:1" "\E"(1) by blast+ + AOT_show \\[[G]\<^sup>-\<^sup>u]x\ + apply (AOT_subst \[[G]\<^sup>-\<^sup>u]x\ \[G]x & x \\<^sub>E u\) + apply (rule "F-u"[THEN "=\<^sub>d\<^sub>fI"(1), where \\<^sub>1\\<^sub>n="(_,_)", simplified]) + apply "cqt:2[lambda]" + apply (rule "beta-C-meta"[THEN "\E"]) + apply "cqt:2[lambda]" + using 2 3 "KBasic:3" "\S"(2) "\E"(2) by blast + } + qed + AOT_thus \\\x([[G]\<^sup>-\<^sup>u]x \ \[[G]\<^sup>-\<^sup>u]x)\ using \ "\E" by blast + next + fix x + AOT_assume \[[G]\<^sup>-\<^sup>u]x\ + AOT_hence \[\x [G]x & x \\<^sub>E u]x\ + by (auto intro: "F-u"[THEN "=\<^sub>d\<^sub>fE"(1), where \\<^sub>1\\<^sub>n="(_,_)", simplified] + intro!: "cqt:2") + AOT_hence \[G]x & x \\<^sub>E u\ + by (rule "\\C"(1)) + AOT_hence \[F]x & x \\<^sub>E u\ + using \ "&I" "&E"(1) "&E"(2) "\E"(1) "rule-ui:3" by blast + AOT_hence \[\x [F]x & x \\<^sub>E u]x\ + by (auto intro!: "\\C"(1) "cqt:2") + AOT_thus \[[F]\<^sup>-\<^sup>u]x\ + by (auto intro: "F-u"[THEN "=\<^sub>d\<^sub>fI"(1), where \\<^sub>1\\<^sub>n="(_,_)", simplified] + intro!: "cqt:2") + next + fix x + AOT_assume \[[F]\<^sup>-\<^sup>u]x\ + AOT_hence \[\x [F]x & x \\<^sub>E u]x\ + by (auto intro: "F-u"[THEN "=\<^sub>d\<^sub>fE"(1), where \\<^sub>1\\<^sub>n="(_,_)", simplified] + intro!: "cqt:2") + AOT_hence \[F]x & x \\<^sub>E u\ + by (rule "\\C"(1)) + AOT_hence \[G]x & x \\<^sub>E u\ + using \ "&I" "&E"(1) "&E"(2) "\E"(2) "rule-ui:3" by blast + AOT_hence \[\x [G]x & x \\<^sub>E u]x\ + by (auto intro!: "\\C"(1) "cqt:2") + AOT_thus \[[G]\<^sup>-\<^sup>u]x\ + by (auto intro: "F-u"[THEN "=\<^sub>d\<^sub>fI"(1), where \\<^sub>1\\<^sub>n="(_,_)", simplified] + intro!: "cqt:2") + qed + AOT_hence \\Numbers(x,[G]\<^sup>-\<^sup>u)\ + using rigid_num_nec[unvarify F G, OF "F-u[den]", OF "F-u[den]", THEN "\E", + OF "&I", OF props[THEN "&E"(2)]] by blast + } + moreover AOT_have \\[G]u\ + using props[THEN "&E"(1), THEN "&E"(1), THEN \[THEN "\E"(2), THEN "\E"(2)]] + \[THEN "qml:2"[axiom_inst, THEN "\E"], THEN "\E"(2), THEN "\E"] + by blast + ultimately AOT_have \\([G]u & Numbers(y,G) & Numbers(x,[G]\<^sup>-\<^sup>u))\ + by (metis "KBasic:3" "&I" "\E"(2)) + AOT_hence \\u (\([G]u & Numbers(y,G) & Numbers(x,[G]\<^sup>-\<^sup>u)))\ + by (rule "Ordinary.\I") + AOT_hence \\\u ([G]u & Numbers(y,G) & Numbers(x,[G]\<^sup>-\<^sup>u))\ + using "Ordinary.res-var-bound-reas[Buridan]" "\E" by fast + AOT_hence \\F \\u ([F]u & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>u))\ + by (rule "\I") + AOT_hence 0: \\\F\u ([F]u & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>u))\ + using Buridan "vdash-properties:10" by fast + AOT_show \\[\]xy\ + by (AOT_subst \[\]xy\ \\F\u ([F]u & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>u))\; + simp add: "pred-thm:3" 0) +qed + +AOT_theorem "pred-1-1:2": \Rigid(\)\ + by (safe intro!: "df-rigid-rel:1"[THEN "\\<^sub>d\<^sub>fI"] "pred-thm:2" "&I" + RN tuple_forall[THEN "\\<^sub>d\<^sub>fI"]; + safe intro!: GEN "pred-1-1:1") + +AOT_theorem "pred-1-1:3": \1-1(\)\ +proof (safe intro!: "df-1-1:1"[THEN "\\<^sub>d\<^sub>fI"] "pred-thm:2" "&I" GEN "\I"; + frule "&E"(1); drule "&E"(2)) + fix x y z + AOT_assume \[\]xz\ + AOT_hence \\F\u ([F]u & Numbers(z,F) & Numbers(x,[F]\<^sup>-\<^sup>u))\ + using "pred-thm:3"[THEN "\E"(1)] by blast + then AOT_obtain F where \\u ([F]u & Numbers(z,F) & Numbers(x,[F]\<^sup>-\<^sup>u))\ + using "\E"[rotated] by blast + then AOT_obtain u where u_prop: \[F]u & Numbers(z,F) & Numbers(x,[F]\<^sup>-\<^sup>u)\ + using "Ordinary.\E"[rotated] by meson + AOT_assume \[\]yz\ + AOT_hence \\F\u ([F]u & Numbers(z,F) & Numbers(y,[F]\<^sup>-\<^sup>u))\ + using "pred-thm:3"[THEN "\E"(1)] by blast + then AOT_obtain G where \\u ([G]u & Numbers(z,G) & Numbers(y,[G]\<^sup>-\<^sup>u))\ + using "\E"[rotated] by blast + then AOT_obtain v where v_prop: \[G]v & Numbers(z,G) & Numbers(y,[G]\<^sup>-\<^sup>v)\ + using "Ordinary.\E"[rotated] by meson + AOT_show \x = y\ + proof (rule "pre-Hume"[unvarify G H, OF "F-u[den]", OF "F-u[den]", + THEN "\E", OF "&I", THEN "\E"(2)]) + AOT_show \Numbers(x, [F]\<^sup>-\<^sup>u)\ + using u_prop "&E" by blast + next + AOT_show \Numbers(y, [G]\<^sup>-\<^sup>v)\ + using v_prop "&E" by blast + next + AOT_have \F \\<^sub>E G\ + using u_prop[THEN "&E"(1), THEN "&E"(2)] + using v_prop[THEN "&E"(1), THEN "&E"(2)] + using "num-tran:2"[THEN "\E", OF "&I"] by blast + AOT_thus \[F]\<^sup>-\<^sup>u \\<^sub>E [G]\<^sup>-\<^sup>v\ + using u_prop[THEN "&E"(1), THEN "&E"(1)] + using v_prop[THEN "&E"(1), THEN "&E"(1)] + using eqP'[THEN "\E", OF "&I", OF "&I"] + by blast + qed +qed + +AOT_theorem "pred-1-1:4": \Rigid\<^sub>1\<^sub>-\<^sub>1(\)\ + by (meson "\\<^sub>d\<^sub>fI" "&I" "df-1-1:2" "pred-1-1:2" "pred-1-1:3") + +AOT_theorem "assume-anc:1": + \[\]\<^sup>* = [\xy \F((\z([\]xz \ [F]z) & Hereditary(F,\)) \ [F]y)]\ + apply (rule "=\<^sub>d\<^sub>fI"(1)[OF "ances-df"]) + apply "cqt:2[lambda]" + apply (rule "=I"(1)) + by "cqt:2[lambda]" + +AOT_theorem "assume-anc:2": \\\<^sup>*\\ + using "t=t-proper:1" "assume-anc:1" "vdash-properties:10" by blast + +AOT_theorem "assume-anc:3": + \[\\<^sup>*]xy \ \F((\z([\]xz \ [F]z) & \x'\y'([\]x'y' \ ([F]x' \ [F]y'))) \ [F]y)\ +proof - + AOT_have prod_den: \\<^bold>\\<^sub>\ \(AOT_term_of_var x\<^sub>1,AOT_term_of_var x\<^sub>2)\\\ + for x\<^sub>1 x\<^sub>2 :: \\ AOT_var\ + by (simp add: "&I" "ex:1:a" prod_denotesI "rule-ui:3") + AOT_have den: \[\xy \F((\z([\]xz \ [F]z) & Hereditary(F,\)) \ [F]y)]\\ + by "cqt:2[lambda]" + AOT_have 1: \[\\<^sup>*]xy \ \F((\z([\]xz \ [F]z) & Hereditary(F,\)) \ [F]y)\ + apply (rule "rule=E"[rotated, OF "assume-anc:1"[symmetric]]) + by (rule "beta-C-meta"[unvarify \\<^sub>1\\<^sub>n, OF prod_den, THEN "\E", + simplified, OF den, simplified]) + show ?thesis + apply (AOT_subst (reverse) \\x'\y' ([\]x'y' \ ([F]x' \ [F]y'))\ + \Hereditary(F,\)\ for: F :: \<\>\) + using "hered:1"[THEN "\Df", THEN "\S"(1), OF "&I", OF "pred-thm:2", + OF "cqt:2[const_var]"[axiom_inst]] apply blast + by (fact 1) +qed + +AOT_theorem "no-pred-0:1": \\\x [\]x 0\ +proof(rule "raa-cor:2") + AOT_assume \\x [\]x 0\ + then AOT_obtain a where \[\]a 0\ + using "\E"[rotated] by blast + AOT_hence \\F\u ([F]u & Numbers(0, F) & Numbers(a, [F]\<^sup>-\<^sup>u))\ + using "pred-thm:3"[unvarify y, OF "zero:2", THEN "\E"(1)] by blast + then AOT_obtain F where \\u ([F]u & Numbers(0, F) & Numbers(a, [F]\<^sup>-\<^sup>u))\ + using "\E"[rotated] by blast + then AOT_obtain u where \[F]u & Numbers(0, F) & Numbers(a, [F]\<^sup>-\<^sup>u)\ + using "Ordinary.\E"[rotated] by meson + AOT_hence \[F]u\ and num0_F: \Numbers(0, F)\ + using "&E" "&I" by blast+ + AOT_hence \\u [F]u\ + using "Ordinary.\I" by fast + moreover AOT_have \\\u [F]u\ + using num0_F "\E"(2) "0F:1" by blast + ultimately AOT_show \p & \p\ for p + by (metis "raa-cor:3") +qed + +AOT_theorem "no-pred-0:2": \\\x [\\<^sup>*]x 0\ +proof(rule "raa-cor:2") + AOT_assume \\x [\\<^sup>*]x 0\ + then AOT_obtain a where \[\\<^sup>*]a 0\ + using "\E"[rotated] by blast + AOT_hence \\z [\]z 0\ + using "anc-her:5"[unvarify R y, OF "zero:2", + OF "pred-thm:2", THEN "\E"] by auto + AOT_thus \\z [\]z 0 & \\z [\]z 0\ + by (metis "no-pred-0:1" "raa-cor:3") +qed + +AOT_theorem "no-pred-0:3": \\[\\<^sup>*]0 0\ + by (metis "existential:1" "no-pred-0:2" "reductio-aa:1" "zero:2") + +AOT_theorem "assume1:1": \(=\<^sub>\) = [\xy \z ([\]xz & [\]yz)]\ + apply (rule "=\<^sub>d\<^sub>fI"(1)[OF "id-d-R"]) + apply "cqt:2[lambda]" + apply (rule "=I"(1)) + by "cqt:2[lambda]" + +AOT_theorem "assume1:2": \x =\<^sub>\ y \ \z ([\]xz & [\]yz)\ +proof (rule "rule=E"[rotated, OF "assume1:1"[symmetric]]) + AOT_have prod_den: \\<^bold>\\<^sub>\ \(AOT_term_of_var x\<^sub>1,AOT_term_of_var x\<^sub>2)\\\ + for x\<^sub>1 x\<^sub>2 :: \\ AOT_var\ + by (simp add: "&I" "ex:1:a" prod_denotesI "rule-ui:3") + AOT_have 1: \[\xy \z ([\]xz & [\]yz)]\\ + by "cqt:2" + AOT_show \[\xy \z ([\]xz & [\]yz)]xy \ \z ([\]xz & [\]yz)\ + using "beta-C-meta"[THEN "\E", OF 1, unvarify \\<^sub>1\\<^sub>n, + OF prod_den, simplified] by blast +qed + +AOT_theorem "assume1:3": \[\]\<^sup>+ = [\xy [\]\<^sup>*xy \ x =\<^sub>\ y]\ + apply (rule "=\<^sub>d\<^sub>fI"(1)[OF "w-ances-df"]) + apply (simp add: "w-ances-df[den1]") + apply (rule "rule=E"[rotated, OF "assume1:1"[symmetric]]) + apply (rule "=\<^sub>d\<^sub>fI"(1)[OF "id-d-R"]) + apply "cqt:2[lambda]" + apply (rule "=I"(1)) + by "cqt:2[lambda]" + +AOT_theorem "assume1:4": \[\]\<^sup>+\\ + using "w-ances-df[den2]". + +AOT_theorem "assume1:5": \[\]\<^sup>+xy \ [\]\<^sup>*xy \ x =\<^sub>\ y\ +proof - + AOT_have 0: \[\xy [\]\<^sup>*xy \ x =\<^sub>\ y]\\ by "cqt:2" + AOT_have prod_den: \\<^bold>\\<^sub>\ \(AOT_term_of_var x\<^sub>1, AOT_term_of_var x\<^sub>2)\\\ + for x\<^sub>1 x\<^sub>2 :: \\ AOT_var\ + by (simp add: "&I" "ex:1:a" prod_denotesI "rule-ui:3") + show ?thesis + apply (rule "rule=E"[rotated, OF "assume1:3"[symmetric]]) + using "beta-C-meta"[THEN "\E", OF 0, unvarify \\<^sub>1\\<^sub>n, OF prod_den, simplified] + by (simp add: cond_case_prod_eta) +qed + +AOT_define NaturalNumber :: \\\ (\\\) + "nnumber:1": \\ =\<^sub>d\<^sub>f [\x [\]\<^sup>+0x]\ + +AOT_theorem "nnumber:2": \\\\ + by (rule "=\<^sub>d\<^sub>fI"(2)[OF "nnumber:1"]; "cqt:2[lambda]") + +AOT_theorem "nnumber:3": \[\]x \ [\]\<^sup>+0x\ + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF "nnumber:1"]) + apply "cqt:2[lambda]" + apply (rule "beta-C-meta"[THEN "\E"]) + by "cqt:2[lambda]" + +AOT_theorem "0-n": \[\]0\ +proof (safe intro!: "nnumber:3"[unvarify x, OF "zero:2", THEN "\E"(2)] + "assume1:5"[unvarify x y, OF "zero:2", OF "zero:2", THEN "\E"(2)] + "\I"(2) "assume1:2"[unvarify x y, OF "zero:2", OF "zero:2", THEN "\E"(2)]) + fix u + AOT_have den: \[\x O!x & x =\<^sub>E u]\\ by "cqt:2[lambda]" + AOT_obtain a where a_prop: \Numbers(a, [\x O!x & x =\<^sub>E u])\ + using "num:1"[unvarify G, OF den] "\E"[rotated] by blast + AOT_have \[\]0a\ + proof (safe intro!: "pred-thm:3"[unvarify x, OF "zero:2", THEN "\E"(2)] + "\I"(1)[where \=\\[\x O!x & x =\<^sub>E u]\\] + "Ordinary.\I"[where \=u] "&I" den + "0F:1"[unvarify F, OF "F-u[den]", unvarify F, + OF den, THEN "\E"(1)]) + AOT_show \[\x [O!]x & x =\<^sub>E u]u\ + by (auto intro!: "\\C"(1) "cqt:2" "&I" "ord=Eequiv:1"[THEN "\E"] + Ordinary.\) + next + AOT_show \Numbers(a,[\x [O!]x & x =\<^sub>E u])\ + using a_prop. + next + AOT_show \\\v [[\x [O!]x & x =\<^sub>E u]\<^sup>-\<^sup>u]v\ + proof(rule "raa-cor:2") + AOT_assume \\v [[\x [O!]x & x =\<^sub>E u]\<^sup>-\<^sup>u]v\ + then AOT_obtain v where \[[\x [O!]x & x =\<^sub>E u]\<^sup>-\<^sup>u]v\ + using "Ordinary.\E"[rotated] "&E" by blast + AOT_hence \[\z [\x [O!]x & x =\<^sub>E u]z & z \\<^sub>E u]v\ + apply (rule "F-u"[THEN "=\<^sub>d\<^sub>fE"(1), where \\<^sub>1\\<^sub>n="(_,_)", simplified, rotated]) + by "cqt:2[lambda]" + AOT_hence \[\x [O!]x & x =\<^sub>E u]v & v \\<^sub>E u\ + by (rule "\\C"(1)) + AOT_hence \v =\<^sub>E u\ and \v \\<^sub>E u\ + using "\\C"(1) "&E" by blast+ + AOT_hence \v =\<^sub>E u & \(v =\<^sub>E u)\ + by (metis "\E"(4) "reductio-aa:1" "thm-neg=E") + AOT_thus \p & \p\ for p + by (metis "raa-cor:1") + qed + qed + AOT_thus \\z ([\]0z & [\]0z)\ + by (safe intro!: "&I" "\I"(2)[where \=a]) +qed + +AOT_theorem "mod-col-num:1": \[\]x \ \[\]x\ +proof(rule "\I") + AOT_have nec0N: \[\x \[\]x]0\ + by (auto intro!: "\\C"(1) "cqt:2" simp: "zero:2" RN "0-n") + AOT_have 1: \[\x \[\]x]0 & + \x\y ([[\]\<^sup>+]0x & [[\]\<^sup>+]0y \ ([\]xy \ ([\x \[\]x]x \ [\x \[\]x]y))) \ + \x ([[\]\<^sup>+]0x \ [\x \[\]x]x)\ + by (auto intro!: "cqt:2" + intro: "pre-ind"[unconstrain \, unvarify \, OF "pred-thm:2", + THEN "\E", OF "pred-1-1:4", unvarify z, OF "zero:2", + unvarify F]) + AOT_have \\x ([[\]\<^sup>+]0x \ [\x \[\]x]x)\ + proof (rule 1[THEN "\E"]; safe intro!: "&I" GEN "\I" nec0N; + frule "&E"(1); drule "&E"(2)) + fix x y + AOT_assume \[\]xy\ + AOT_hence 0: \\[\]xy\ + by (metis "pred-1-1:1" "\E") + AOT_assume \[\x \[\]x]x\ + AOT_hence \\[\]x\ + by (rule "\\C"(1)) + AOT_hence \\([\]xy & [\]x)\ + by (metis "0" "KBasic:3" Adjunction "\E"(2) "\E") + moreover AOT_have \\([\]xy & [\]x) \ \[\]y\ + proof (rule RM; rule "\I"; frule "&E"(1); drule "&E"(2)) + AOT_modally_strict { + AOT_assume 0: \[\]xy\ + AOT_assume \[\]x\ + AOT_hence 1: \[[\]\<^sup>+]0x\ + by (metis "\E"(1) "nnumber:3") + AOT_show \[\]y\ + apply (rule "nnumber:3"[THEN "\E"(2)]) + apply (rule "assume1:5"[unvarify x, OF "zero:2", THEN "\E"(2)]) + apply (rule "\I"(1)) + apply (rule "w-ances-her:3"[unconstrain \, unvarify \, OF "pred-thm:2", + THEN "\E", OF "pred-1-1:4", unvarify x, + OF "zero:2", THEN "\E"]) + apply (rule "&I") + apply (fact 1) + by (fact 0) + } + qed + ultimately AOT_have \\[\]y\ + by (metis "\E") + AOT_thus \[\x \[\]x]y\ + by (auto intro!: "\\C"(1) "cqt:2") + qed + AOT_hence 0: \[[\]\<^sup>+]0x \ [\x \[\]x]x\ + using "\E"(2) by blast + AOT_assume \[\]x\ + AOT_hence \[[\]\<^sup>+]0x\ + by (metis "\E"(1) "nnumber:3") + AOT_hence \[\x \[\]x]x\ + using 0[THEN "\E"] by blast + AOT_thus \\[\]x\ + by (rule "\\C"(1)) +qed + +AOT_theorem "mod-col-num:2": \Rigid(\)\ + by (safe intro!: "df-rigid-rel:1"[THEN "\\<^sub>d\<^sub>fI"] "&I" RN GEN + "mod-col-num:1" "nnumber:2") + +AOT_register_rigid_restricted_type + Number: \[\]\\ +proof + AOT_modally_strict { + AOT_show \\x [\]x\ + by (rule "\I"(1)[where \=\\0\\]; simp add: "0-n" "zero:2") + } +next + AOT_modally_strict { + AOT_show \[\]\ \ \\\ for \ + by (simp add: "\I" "cqt:5:a[1]"[axiom_inst, THEN "\E", THEN "&E"(2)]) + } +next + AOT_modally_strict { + AOT_show \\x([\]x \ \[\]x)\ + by (simp add: GEN "mod-col-num:1") + } +qed +AOT_register_variable_names + Number: m n k i j + +AOT_theorem "0-pred": \\\n [\]n 0\ +proof (rule "raa-cor:2") + AOT_assume \\n [\]n 0\ + then AOT_obtain n where \[\]n 0\ + using "Number.\E"[rotated] by meson + AOT_hence \\x [\]x 0\ + using "&E" "\I" by fast + AOT_thus \\x [\]x 0 & \\x [\]x 0\ + using "no-pred-0:1" "&I" by auto +qed + +AOT_theorem "no-same-succ": + \\n\m\k([\]nk & [\]mk \ n = m)\ +proof(safe intro!: Number.GEN "\I") + fix n m k + AOT_assume \[\]nk & [\]mk\ + AOT_thus \n = m\ + by (safe intro!: "cqt:2[const_var]"[axiom_inst] "df-1-1:3"[ + unvarify R, OF "pred-thm:2", + THEN "\E", OF "pred-1-1:4", THEN "qml:2"[axiom_inst, THEN "\E"], + THEN "\\<^sub>d\<^sub>fE"[OF "df-1-1:1"], THEN "&E"(2), THEN "\E"(1), THEN "\E"(1), + THEN "\E"(1)[where \=\AOT_term_of_var (Number.Rep k)\], THEN "\E"]) +qed + +AOT_theorem induction: + \\F([F]0 & \n\m([\]nm \ ([F]n \ [F]m)) \ \n[F]n)\ +proof (safe intro!: GEN[where 'a=\<\>\] Number.GEN "&I" "\I"; + frule "&E"(1); drule "&E"(2)) + fix F n + AOT_assume F0: \[F]0\ + AOT_assume 0: \\n\m([\]nm \ ([F]n \ [F]m))\ + { + fix x y + AOT_assume \[[\]\<^sup>+]0x & [[\]\<^sup>+]0y\ + AOT_hence \[\]x\ and \[\]y\ + using "&E" "\E"(2) "nnumber:3" by blast+ + moreover AOT_assume \[\]xy\ + moreover AOT_assume \[F]x\ + ultimately AOT_have \[F]y\ + using 0[THEN "\E"(2), THEN "\E", THEN "\E"(2), THEN "\E", + THEN "\E", THEN "\E"] by blast + } note 1 = this + AOT_have 0: \[[\]\<^sup>+]0n\ + by (metis "\E"(1) "nnumber:3" Number.\) + AOT_show \[F]n\ + apply (rule "pre-ind"[unconstrain \, unvarify \, THEN "\E", OF "pred-thm:2", + OF "pred-1-1:4", unvarify z, OF "zero:2", THEN "\E", + THEN "\E"(2), THEN "\E"]; + safe intro!: 0 "&I" GEN "\I" F0) + using 1 by blast +qed + +AOT_theorem "suc-num:1": \[\]nx \ [\]x\ +proof(rule "\I") + AOT_have \[[\]\<^sup>+]0 n\ + by (meson Number.\ "\E"(1) "nnumber:3") + moreover AOT_assume \[\]nx\ + ultimately AOT_have \[[\]\<^sup>*]0 x\ + using "w-ances-her:3"[unconstrain \, unvarify \, OF "pred-thm:2", THEN "\E", + OF "pred-1-1:4", unvarify x, OF "zero:2", + THEN "\E", OF "&I"] + by blast + AOT_hence \[[\]\<^sup>+]0 x\ + using "assume1:5"[unvarify x, OF "zero:2", THEN "\E"(2), OF "\I"(1)] + by blast + AOT_thus \[\]x\ + by (metis "\E"(2) "nnumber:3") +qed + +AOT_theorem "suc-num:2": \[[\]\<^sup>*]nx \ [\]x\ +proof(rule "\I") + AOT_have \[[\]\<^sup>+]0 n\ + using Number.\ "\E"(1) "nnumber:3" by blast + AOT_assume \[[\]\<^sup>*]n x\ + AOT_hence \\F (\z ([\]nz \ [F]z) & \x'\y' ([\]x'y' \ ([F]x' \ [F]y')) \ [F]x)\ + using "assume-anc:3"[THEN "\E"(1)] by blast + AOT_hence \: \\z ([\]nz \ [\]z) & \x'\y' ([\]x'y' \ ([\]x' \ [\]y')) \ [\]x\ + using "\E"(1) "nnumber:2" by blast + AOT_show \[\]x\ + proof (safe intro!: \[THEN "\E"] GEN "\I" "&I") + AOT_show \[\]z\ if \[\]nz\ for z + using Number.\ "suc-num:1" that "\E" by blast + next + AOT_show \[\]y\ if \[\]xy\ and \[\]x\ for x y + using "suc-num:1"[unconstrain n, THEN "\E"] that "\E" by blast + qed +qed + +AOT_theorem "suc-num:3": \[\]\<^sup>+nx \ [\]x\ +proof (rule "\I") + AOT_assume \[\]\<^sup>+nx\ + AOT_hence \[\]\<^sup>*nx \ n =\<^sub>\ x\ + by (metis "assume1:5" "\E"(1)) + moreover { + AOT_assume \[\]\<^sup>*nx\ + AOT_hence \[\]x\ + by (metis "suc-num:2" "\E") + } + moreover { + AOT_assume \n =\<^sub>\ x\ + AOT_hence \n = x\ + using "id-R-thm:3"[unconstrain \, unvarify \, OF "pred-thm:2", + THEN "\E", OF "pred-1-1:4", THEN "\E"] by blast + AOT_hence \[\]x\ + by (metis "rule=E" Number.\) + } + ultimately AOT_show \[\]x\ + by (metis "\E"(3) "reductio-aa:1") +qed + +AOT_theorem "pred-num": \[\]xn \ [\]x\ +proof (rule "\I") + AOT_assume 0: \[\]xn\ + AOT_have \[[\]\<^sup>+]0 n\ + using Number.\ "\E"(1) "nnumber:3" by blast + AOT_hence \[[\]\<^sup>*]0 n \ 0 =\<^sub>\ n\ + using "assume1:5"[unvarify x, OF "zero:2"] by (metis "\E"(1)) + moreover { + AOT_assume \0 =\<^sub>\ n\ + AOT_hence \\z ([\]0z & [\]nz)\ + using "assume1:2"[unvarify x, OF "zero:2", THEN "\E"(1)] by blast + then AOT_obtain a where \[\]0a & [\]na\ using "\E"[rotated] by blast + AOT_hence \0 = n\ + using "pred-1-1:3"[THEN "df-1-1:1"[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(2), + THEN "\E"(1), OF "zero:2", THEN "\E"(2), + THEN "\E"(2), THEN "\E"] by blast + AOT_hence \[\]x 0\ + using 0 "rule=E" id_sym by fast + AOT_hence \\x [\]x 0\ + by (rule "\I") + AOT_hence \\x [\]x 0 & \\x [\]x 0\ + by (metis "no-pred-0:1" "raa-cor:3") + } + ultimately AOT_have \[[\]\<^sup>*]0n\ + by (metis "\E"(3) "raa-cor:1") + AOT_hence \\z ([[\]\<^sup>+]0z & [\]zn)\ + using "w-ances-her:7"[unconstrain \, unvarify \, OF "pred-thm:2", + THEN "\E", OF "pred-1-1:4", unvarify x, + OF "zero:2", THEN "\E"] by blast + then AOT_obtain b where b_prop: \[[\]\<^sup>+]0b & [\]bn\ + using "\E"[rotated] by blast + AOT_hence \[\]b\ + by (metis "&E"(1) "\E"(2) "nnumber:3") + moreover AOT_have \x = b\ + using "pred-1-1:3"[THEN "df-1-1:1"[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(2), + THEN "\E"(2), THEN "\E"(2), THEN "\E"(2), THEN "\E", + OF "&I", OF 0, OF b_prop[THEN "&E"(2)]]. + ultimately AOT_show \[\]x\ + using "rule=E" id_sym by fast +qed + +AOT_theorem "nat-card": \[\]x \ NaturalCardinal(x)\ +proof(rule "\I") + AOT_assume \[\]x\ + AOT_hence \[[\]\<^sup>+]0x\ + by (metis "\E"(1) "nnumber:3") + AOT_hence \[[\]\<^sup>*]0x \ 0 =\<^sub>\ x\ + using "assume1:5"[unvarify x, OF "zero:2", THEN "\E"(1)] by blast + moreover { + AOT_assume \[[\]\<^sup>*]0x\ + then AOT_obtain a where \[\]ax\ + using "anc-her:5"[unvarify R x, OF "zero:2", OF "pred-thm:2", THEN "\E"] + "\E"[rotated] by blast + AOT_hence \\F\u ([F]u & Numbers(x,F) & Numbers(a,[F]\<^sup>-\<^sup>u))\ + using "pred-thm:3"[THEN "\E"(1)] by blast + then AOT_obtain F where \\u ([F]u & Numbers(x,F) & Numbers(a,[F]\<^sup>-\<^sup>u))\ + using "\E"[rotated] by blast + then AOT_obtain u where \[F]u & Numbers(x,F) & Numbers(a,[F]\<^sup>-\<^sup>u)\ + using "Ordinary.\E"[rotated] by meson + AOT_hence \NaturalCardinal(x)\ + using "eq-num:6"[THEN "\E"] "&E" by blast + } + moreover { + AOT_assume \0 =\<^sub>\ x\ + AOT_hence \0 = x\ + using "id-R-thm:3"[unconstrain \, unvarify \, OF "pred-thm:2", + THEN "\E", OF "pred-1-1:4", unvarify x, + OF "zero:2", THEN "\E"] by blast + AOT_hence \NaturalCardinal(x)\ + by (metis "rule=E" "zero-card") + } + ultimately AOT_show \NaturalCardinal(x)\ + by (metis "\E"(2) "raa-cor:1") +qed + +AOT_theorem "pred-func:1": \[\]xy & [\]xz \ y = z\ +proof (rule "\I"; frule "&E"(1); drule "&E"(2)) + AOT_assume \[\]xy\ + AOT_hence \\F\u ([F]u & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>u))\ + using "pred-thm:3"[THEN "\E"(1)] by blast + then AOT_obtain F where \\u ([F]u & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>u))\ + using "\E"[rotated] by blast + then AOT_obtain a where + Oa: \O!a\ + and a_prop: \[F]a & Numbers(y,F) & Numbers(x,[F]\<^sup>-\<^sup>a)\ + using "\E"[rotated] "&E" by blast + AOT_assume \[\]xz\ + AOT_hence \\F\u ([F]u & Numbers(z,F) & Numbers(x,[F]\<^sup>-\<^sup>u))\ + using "pred-thm:3"[THEN "\E"(1)] by blast + then AOT_obtain G where \\u ([G]u & Numbers(z,G) & Numbers(x,[G]\<^sup>-\<^sup>u))\ + using "\E"[rotated] by blast + then AOT_obtain b where Ob: \O!b\ + and b_prop: \[G]b & Numbers(z,G) & Numbers(x,[G]\<^sup>-\<^sup>b)\ + using "\E"[rotated] "&E" by blast + AOT_have \[F]\<^sup>-\<^sup>a \\<^sub>E [G]\<^sup>-\<^sup>b\ + using "num-tran:2"[unvarify G H, OF "F-u[den]", OF "F-u[den]", + THEN "\E", OF "&I", OF a_prop[THEN "&E"(2)], + OF b_prop[THEN "&E"(2)]]. + AOT_hence \F \\<^sub>E G\ + using "P'-eq"[unconstrain u, THEN "\E", OF Oa, unconstrain v, THEN "\E", + OF Ob, THEN "\E", OF "&I", OF "&I"] + a_prop[THEN "&E"(1), THEN "&E"(1)] + b_prop[THEN "&E"(1), THEN "&E"(1)] by blast + AOT_thus \y = z\ + using "pre-Hume"[THEN "\E", THEN "\E"(2), OF "&I", + OF a_prop[THEN "&E"(1), THEN "&E"(2)], + OF b_prop[THEN "&E"(1), THEN "&E"(2)]] + by blast +qed + +AOT_theorem "pred-func:2": \[\]nm & [\]nk \ m = k\ + using "pred-func:1". + +AOT_theorem being_number_of_den: \[\x x = #G]\\ +proof (rule "safe-ext"[axiom_inst, THEN "\E"]; safe intro!: "&I" GEN RN) + AOT_show \[\x Numbers(x,[\z \<^bold>\[G]z])]\\ + by (rule numbers_prop_den[unvarify G]) "cqt:2[lambda]" +next + AOT_modally_strict { + AOT_show \Numbers(x,[\z \<^bold>\[G]z]) \ x = #G\ for x + using "eq-num:2". + } +qed + +axiomatization \_nat :: \\ \ nat\ where \_nat: \surj \_nat\ +text\Unfortunately, since the axiom requires the type @{typ \} + to have an infinite domain, @{command nitpick} can only find a potential model + and no genuine model. + However, since we could trivially choose @{typ \} as a copy of @{typ nat}, + we can still be assured that above axiom is consistent.\ +lemma \True\ nitpick[satisfy, user_axioms, card nat=1, expect = potential] .. + +AOT_axiom "modal-axiom": + \\x([\]x & x = #G) \ \\y([E!]y & \u (\<^bold>\[G]u \ u \\<^sub>E y))\ +proof(rule AOT_model_axiomI) AOT_modally_strict { + text\The actual extension on the ordinary objects of a property is the + set of ordinary urelements that exemplifies the property in the + designated actual world.\ + define act_\ext :: \<\> \ \ set\ where + \act_\ext \ \ \ . {x :: \ . [w\<^sub>0 \ [\]\\\ x\]}\ + text\Encoding a property with infinite actual extension on the ordinary objects + denotes a property by extended relation comprehension.\ + AOT_have enc_finite_act_\ext_den: + \\<^bold>\\<^sub>\ [\x \F(\\\\<^sub>\ w. finite (act_\ext F)\ & x[F])]\\ + proof(safe intro!: Comprehension_1[THEN "\E"] RN GEN "\I") + AOT_modally_strict { + fix F G + AOT_assume \\G \\<^sub>E F\ + AOT_hence \\<^bold>\G \\<^sub>E F\ + using "nec-imp-act"[THEN "\E"] by blast + AOT_hence \\<^bold>\(G\ & F\ & \u([G]u \ [F]u))\ + by (AOT_subst_def (reverse) eqE) + hence \[w\<^sub>0 \ [G]\\\ x\] = [w\<^sub>0 \ [F]\\\ x\]\ for x + by (auto dest!: "\E"(1) "\E" + simp: AOT_model_denotes_\_def AOT_sem_denotes AOT_sem_conj + AOT_model_\\_ordinary AOT_sem_act AOT_sem_equiv) + AOT_thus \\\\\<^sub>\ w. finite (act_\ext (AOT_term_of_var F))\ \ + \\\\<^sub>\ w. finite (act_\ext (AOT_term_of_var G))\\ + by (simp add: AOT_sem_not AOT_sem_equiv act_\ext_def + AOT_model_proposition_choice_simp) + } + qed + text\By coexistence, encoding only properties with finite actual extension + on the ordinary objects denotes.\ + AOT_have \[\x \F(x[F] \ \\\<^sub>\ w. finite (act_\ext F)\)]\\ + proof(rule "safe-ext"[axiom_inst, THEN "\E"]; safe intro!: "&I" RN GEN) + AOT_show \[\x \[\x \F(\\\\<^sub>\ w. finite (act_\ext F)\ & x[F])]x]\\ + by "cqt:2" + next + AOT_modally_strict { + fix x + AOT_show \\[\x \F (\\\\<^sub>\ w. finite (act_\ext F)\ & x[F])]x \ + \F(x[F] \ \\\<^sub>\ w. finite (act_\ext F)\)\ + by (AOT_subst \[\x \F (\\\\<^sub>\ w. finite (act_\ext F)\ & x[F])]x\ + \\F (\\\\<^sub>\ w. finite (act_\ext F)\ & x[F])\; + (rule "beta-C-meta"[THEN "\E"])?) + (auto simp: enc_finite_act_\ext_den AOT_sem_equiv AOT_sem_not + AOT_sem_forall AOT_sem_imp AOT_sem_conj AOT_sem_exists) + } + qed + text\We show by induction that any property encoded by a natural number + has a finite actual extension on the ordinary objects.\ + AOT_hence \[\x \F(x[F] \ \\\<^sub>\ w. finite (act_\ext F)\)]n\ for n + proof(rule induction[THEN "\E"(1), THEN "\E", THEN "Number.\E"]; + safe intro!: "&I" "Number.GEN" "\\C" "zero:2" "\I" "cqt:2" + dest!: "\\C") + AOT_show \\F(0[F] \ \\\<^sub>\ w. finite (act_\ext F)\)\ + proof(safe intro!: GEN "\I") + fix F + AOT_assume \0[F]\ + AOT_actually { + AOT_hence \\\u [F]u\ + using "zero=:2" "intro-elim:3:a" AOT_sem_enc_nec by blast + AOT_hence \\x \(O!x & [F]x)\ + using "cqt-further:4" "vdash-properties:10" by blast + hence \\([w\<^sub>0 \ [F]\\\ x\])\ for x + by (auto dest!: "\E"(1)[where \=\\\ x\] + simp: AOT_sem_not AOT_sem_conj AOT_model_\\_ordinary + "russell-axiom[exe,1].\_denotes_asm") + } + AOT_thus \\\\<^sub>\ w. finite (act_\ext (AOT_term_of_var F))\\ + by (auto simp: AOT_model_proposition_choice_simp act_\ext_def) + qed + next + fix n m + AOT_assume \[\]nm\ + AOT_hence \\F\u ([F]u & Numbers(m,F) & Numbers(n,[F]\<^sup>-\<^sup>u))\ + using "pred-thm:3"[THEN "\E"(1)] by blast + then AOT_obtain G where \\u ([G]u & Numbers(m,G) & Numbers(n,[G]\<^sup>-\<^sup>u))\ + using "\E"[rotated] by blast + then AOT_obtain u where 0: \[G]u & Numbers(m,G) & Numbers(n,[G]\<^sup>-\<^sup>u)\ + using "Ordinary.\E"[rotated] by meson + + AOT_assume n_prop: \\F(n[F] \ \\\<^sub>\ w. finite (act_\ext F)\)\ + AOT_show \\F(m[F] \ \\\<^sub>\ w. finite (act_\ext F)\)\ + proof(safe intro!: GEN "\I") + fix F + AOT_assume \m[F]\ + AOT_hence 1: \[\x \<^bold>\[F]x] \\<^sub>E G\ + using 0[THEN "&E"(1), THEN "&E"(2), THEN numbers[THEN "\\<^sub>d\<^sub>fE"], + THEN "&E"(2), THEN "\E"(2), THEN "\E"(1)] by auto + AOT_show \\\\<^sub>\ w. finite (act_\ext (AOT_term_of_var F))\\ + proof(rule "raa-cor:1") + AOT_assume \\\\\<^sub>\ w. finite (act_\ext (AOT_term_of_var F))\\ + hence inf: \infinite (act_\ext (AOT_term_of_var F))\ + by (auto simp: AOT_sem_not AOT_model_proposition_choice_simp) + then AOT_obtain v where act_F_v: \\<^bold>\[F]v\ + unfolding AOT_sem_act act_\ext_def + by (metis AOT_term_of_var_cases AOT_model_\\_ordinary + AOT_model_denotes_\_def Ordinary.Rep_cases \.disc(7) + mem_Collect_eq not_finite_existsD) + AOT_hence \[\x \<^bold>\[F]x]v\ + by (safe intro!: "\\C" "cqt:2") + AOT_hence \[\x \<^bold>\[F]x]\<^sup>-\<^sup>v \\<^sub>E [G]\<^sup>-\<^sup>u\ + by (safe intro!: eqP'[unvarify F, THEN "\E"] "&I" "cqt:2" 1 + 0[THEN "&E"(1), THEN "&E"(1)]) + moreover AOT_have \[\x \<^bold>\[F]x]\<^sup>-\<^sup>v \\<^sub>E [\x \<^bold>\[\y [F]y & y \\<^sub>E v]x]\ + proof(safe intro!: "apE-eqE:1"[unvarify F G, THEN "\E"] "cqt:2" + "F-u[den]"[unvarify F] eqE[THEN "\\<^sub>d\<^sub>fI"] "&I" + Ordinary.GEN) + fix u + AOT_have \[\x [\x \<^bold>\[F]x]x & x \\<^sub>E v]u \ [\x \<^bold>\[F]x]u & u \\<^sub>E v\ + by (safe intro!: "beta-C-meta"[THEN "\E"] "cqt:2") + also AOT_have \[\x \<^bold>\[F]x]u & u \\<^sub>E v \ \<^bold>\[F]u & u \\<^sub>E v\ + by (AOT_subst \[\x \<^bold>\[F]x]u\ \\<^bold>\[F]u\) + (safe intro!: "beta-C-meta"[THEN "\E"] "cqt:2" + "oth-class-taut:3:a") + also AOT_have \\<^bold>\[F]u & u \\<^sub>E v \ \<^bold>\([F]u & u \\<^sub>E v)\ + using "id-act2:2" AOT_sem_conj AOT_sem_equiv AOT_sem_act by auto + also AOT_have \\<^bold>\([F]u & u \\<^sub>E v) \ \<^bold>\[\y [F]y & y \\<^sub>E v]u\ + by (AOT_subst \[\y [F]y & y \\<^sub>E v]u\ \[F]u & u \\<^sub>E v\) + (safe intro!: "beta-C-meta"[THEN "\E"] "cqt:2" + "oth-class-taut:3:a") + also AOT_have \\<^bold>\[\y [F]y & y \\<^sub>E v]u \ [\x \<^bold>\[\y [F]y & y \\<^sub>E v]x]u\ + by (safe intro!: "beta-C-meta"[THEN "\E", symmetric] "cqt:2") + finally AOT_show \[[\x \<^bold>\[F]x]\<^sup>-\<^sup>v]u \ [\x \<^bold>\[\y [F]y & y \\<^sub>E v]x]u\ + by (auto intro!: "cqt:2" + intro: "rule-id-df:2:b"[OF "F-u", where \\<^sub>1\\<^sub>n=\(_,_)\, simplified]) + qed + ultimately AOT_have \[\x \<^bold>\[\y [F]y & y \\<^sub>E v]x] \\<^sub>E [G]\<^sup>-\<^sup>u\ + using "eq-part:2[terms]" "eq-part:3[terms]" "\E" by blast + AOT_hence \n[\y [F]y & y \\<^sub>E v]\ + by (safe intro!: 0[THEN "&E"(2), THEN numbers[THEN "\\<^sub>d\<^sub>fE"], + THEN "&E"(2), THEN "\E"(1), THEN "\E"(2)] "cqt:2") + hence finite: \finite (act_\ext \[\y [F]y & y \\<^sub>E v]\)\ + by (safe intro!: n_prop[THEN "\E"(1), THEN "\E", + simplified AOT_model_proposition_choice_simp] + "cqt:2") + obtain y where y_def: \\\ y = AOT_term_of_var (Ordinary.Rep v)\ + by (metis AOT_model_ordinary_\\ Ordinary.restricted_var_condition) + AOT_actually { + fix x + AOT_assume \[\y [F]y & y \\<^sub>E v]\\\ x\\ + AOT_hence \[F]\\\ x\\ + by (auto dest!: "\\C" "&E"(1)) + } + moreover AOT_actually { + AOT_have \[F]\\\ y\\ + unfolding y_def using act_F_v AOT_sem_act by blast + } + moreover AOT_actually { + fix x + assume noteq: \x \ y\ + AOT_assume \[F]\\\ x\\ + moreover AOT_have \\_x_den: \\\\ x\\\ + using AOT_sem_exe calculation by blast + moreover { + AOT_have \\(\\\ x\ =\<^sub>E v)\ + proof(rule "raa-cor:2") + AOT_assume \\\\ x\ =\<^sub>E v\ + AOT_hence \\\\ x\ = v\ + using "=E-simple:2"[unvarify x, THEN "\E", OF \\_x_den] + by blast + hence \\\ x = \\ y\ + unfolding y_def AOT_sem_eq + by meson + hence \x = y\ + by blast + AOT_thus \p & \p\ for p using noteq by blast + qed + AOT_hence \\\\ x\ \\<^sub>E v\ + by (safe intro!: "thm-neg=E"[unvarify x, THEN "\E"(2)] \\_x_den) + } + ultimately AOT_have \[\y [F]y & y \\<^sub>E v]\\\ x\\ + by (auto intro!: "\\C" "cqt:2" "&I") + } + ultimately have \(insert y (act_\ext \[\y [F]y & y \\<^sub>E v]\)) = + (act_\ext (AOT_term_of_var F))\ + unfolding act_\ext_def + by auto + hence \finite (act_\ext (AOT_term_of_var F))\ + using finite finite.insertI by metis + AOT_thus \p & \p\ for p + using inf by blast + qed + qed + qed + AOT_hence nat_enc_finite: \\F(n[F] \ \\\<^sub>\ w. finite (act_\ext F)\)\ for n + using "\\C"(1) by blast + + text\The main proof can now generate a witness, since we required + the domain of ordinary objects to be infinite.\ + AOT_show \\x ([\]x & x = #G) \ \\y (E!y & \u (\<^bold>\[G]u \ u \\<^sub>E y))\ + proof(safe intro!: "\I") + AOT_assume \\x ([\]x & x = #G)\ + then AOT_obtain n where \n = #G\ + using "Number.\E"[rotated] by meson + AOT_hence \Numbers(n,[\x \<^bold>\[G]x])\ + using "eq-num:3" "rule=E" id_sym by fast + AOT_hence \n[G]\ + by (auto intro!: numbers[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2), + THEN "\E"(2), THEN "\E"(2)] + "eq-part:1"[unvarify F] "cqt:2") + AOT_hence \\\\<^sub>\ w. finite (act_\ext (AOT_term_of_var G))\\ + using nat_enc_finite[THEN "\E"(2), THEN "\E"] by blast + hence finite: \finite (act_\ext (AOT_term_of_var G))\ + by (auto simp: AOT_model_proposition_choice_simp) + AOT_have \\u \\<^bold>\[G]u\ + proof(rule "raa-cor:1") + AOT_assume \\\u \\<^bold>\[G]u\ + AOT_hence \\x \(O!x & \\<^bold>\[G]x)\ + by (metis "cqt-further:4" "\E") + AOT_hence \\<^bold>\[G]x\ if \O!x\ for x + using "\E"(2) AOT_sem_conj AOT_sem_not that by blast + hence \[w\<^sub>0 \ [G]\\\ x\]\ for x + by (metis AOT_term_of_var_cases AOT_model_\\_ordinary + AOT_model_denotes_\_def AOT_sem_act \.disc(7)) + hence \(act_\ext (AOT_term_of_var G)) = UNIV\ + unfolding act_\ext_def by auto + moreover have \infinite (UNIV::\ set)\ + by (metis \_nat finite_imageI infinite_UNIV_char_0) + ultimately have \infinite (act_\ext (AOT_term_of_var G))\ + by simp + AOT_thus \p & \p\ for p using finite by blast + qed + then AOT_obtain x where x_prop: \O!x & \\<^bold>\[G]x\ + using "\E"[rotated] by blast + AOT_hence \\E!x\ + by (metis "betaC:1:a" "con-dis-i-e:2:a" AOT_sem_ordinary) + moreover AOT_have \\\u (\<^bold>\[G]u \ u \\<^sub>E x)\ + proof(safe intro!: RN GEN "\I") + AOT_modally_strict { + fix y + AOT_assume \O!y\ + AOT_assume 0: \\<^bold>\[G]y\ + AOT_show \y \\<^sub>E x\ + proof (safe intro!: "thm-neg=E"[THEN "\E"(2)] "raa-cor:2") + AOT_assume \y =\<^sub>E x\ + AOT_hence \y = x\ + by (metis "=E-simple:2" "vdash-properties:10") + hence \y = x\ + by (simp add: AOT_sem_eq AOT_term_of_var_inject) + AOT_hence \\\<^bold>\[G]y\ + using x_prop "&E" AOT_sem_not AOT_sem_act by metis + AOT_thus \\<^bold>\[G]y & \\<^bold>\[G]y\ + using 0 "&I" by blast + qed + } + qed + ultimately AOT_have \\(\u (\<^bold>\[G]u \ u \\<^sub>E x) & E!x)\ + using "KBasic:16"[THEN "\E", OF "&I"] by blast + AOT_hence \\(E!x & \u (\<^bold>\[G]u \ u \\<^sub>E x))\ + by (AOT_subst \E!x & \u (\<^bold>\[G]u \ u \\<^sub>E x)\ \\u (\<^bold>\[G]u \ u \\<^sub>E x) & E!x\) + (auto simp: "oth-class-taut:2:a") + AOT_hence \\y \(E!y & \u (\<^bold>\[G]u \ u \\<^sub>E y))\ + using "\I" by fast + AOT_thus \\\y (E!y & \u (\<^bold>\[G]u \ u \\<^sub>E y))\ + using "CBF\"[THEN "\E"] by fast + qed +} qed + +AOT_theorem "modal-lemma": + \\\u(\<^bold>\[G]u \ u \\<^sub>E v) \ \u(\<^bold>\[G]u \ u \\<^sub>E v)\ +proof(safe intro!: "\I" Ordinary.GEN) + AOT_modally_strict { + fix u + AOT_assume act_Gu: \\<^bold>\[G]u\ + AOT_have \\u (\<^bold>\[G]u \ u \\<^sub>E v) \ u \\<^sub>E v\ + proof(rule "\I") + AOT_assume \\u (\<^bold>\[G]u \ u \\<^sub>E v)\ + AOT_hence \\<^bold>\[G]u \ u \\<^sub>E v\ + using "Ordinary.\E" by fast + AOT_thus \u \\<^sub>E v\ + using act_Gu "\E" by blast + qed + } note 0 = this + AOT_have \: \\(\u (\<^bold>\[G]u \ u \\<^sub>E v) \ u \\<^sub>E v)\ if \\\<^bold>\[G]u\ for u + proof - + AOT_have \\\<^bold>\[G]u \ \(\u (\<^bold>\[G]u \ u \\<^sub>E v) \ u \\<^sub>E v)\ + apply (rule RM) using 0 "&E" "\I" by blast + thus ?thesis using that "\E" by blast + qed + fix u + AOT_assume 1: \\\u(\<^bold>\[G]u \ u \\<^sub>E v)\ + AOT_assume \\<^bold>\[G]u\ + AOT_hence \\\<^bold>\[G]u\ + by (metis "Act-Basic:6" "\E"(1)) + AOT_hence \\(\u (\<^bold>\[G]u \ u \\<^sub>E v) \ u \\<^sub>E v)\ + using Ordinary.\ \ by blast + AOT_hence \\u \\<^sub>E v\ + using 1 "K\"[THEN "\E", THEN "\E"] by blast + AOT_thus \u \\<^sub>E v\ + by (metis "id-nec4:2" "\E"(1)) +qed + +AOT_theorem "th-succ": \\n\!m [\]nm\ +proof(safe intro!: Number.GEN "\I" "uniqueness:1"[THEN "\\<^sub>d\<^sub>fI"]) + fix n + AOT_have \NaturalCardinal(n)\ + by (metis "nat-card" Number.\ "\E") + AOT_hence \\G(n = #G)\ + by (metis "\\<^sub>d\<^sub>fE" card) + then AOT_obtain G where n_num_G: \n = #G\ + using "\E"[rotated] by blast + AOT_hence \\n (n = #G)\ + by (rule "Number.\I") + AOT_hence \\\y ([E!]y & \u(\<^bold>\[G]u \ u \\<^sub>E y))\ + using "modal-axiom"[axiom_inst, THEN "\E"] by blast + AOT_hence \\y \([E!]y & \u(\<^bold>\[G]u \ u \\<^sub>E y))\ + using "BF\"[THEN "\E"] by auto + then AOT_obtain y where \\([E!]y & \u(\<^bold>\[G]u \ u \\<^sub>E y))\ + using "\E"[rotated] by blast + AOT_hence \\E!y\ and 2: \\\u(\<^bold>\[G]u \ u \\<^sub>E y)\ + using "KBasic2:3" "&E" "\E" by blast+ + AOT_hence Oy: \O!y\ + by (auto intro!: "\\C"(1) "cqt:2" intro: AOT_ordinary[THEN "=\<^sub>d\<^sub>fI"(2)]) + AOT_have 0: \\u(\<^bold>\[G]u \ u \\<^sub>E y)\ + using 2 "modal-lemma"[unconstrain v, THEN "\E", OF Oy, THEN "\E"] by simp + AOT_have 1: \[\x \<^bold>\[G]x \ x =\<^sub>E y]\\ + by "cqt:2" + AOT_obtain b where b_prop: \Numbers(b, [\x \<^bold>\[G]x \ x =\<^sub>E y])\ + using "num:1"[unvarify G, OF 1] "\E"[rotated] by blast + AOT_have Pnb: \[\]nb\ + proof(safe intro!: "pred-thm:3"[THEN "\E"(2)] + "\I"(1)[where \=\\[\x \<^bold>\[G]x \ x =\<^sub>E y]\\] + 1 "\I"(2)[where \=y] "&I" Oy b_prop) + AOT_show \[\x \<^bold>\[G]x \ x =\<^sub>E y]y\ + by (auto intro!: "\\C"(1) "cqt:2" "\I"(2) + "ord=Eequiv:1"[THEN "\E", OF Oy]) + next + AOT_have equinum: \[\x \<^bold>\[G]x \ x =\<^sub>E y]\<^sup>-\<^sup>y \\<^sub>E [\x \<^bold>\[G]x]\ + proof(rule "apE-eqE:1"[unvarify F G, THEN "\E"]; + ("cqt:2[lambda]" | rule "F-u[den]"[unvarify F]; "cqt:2[lambda]")?) + AOT_show \[\x \<^bold>\[G]x \ x =\<^sub>E y]\<^sup>-\<^sup>y \\<^sub>E [\x \<^bold>\[G]x]\ + proof (safe intro!: eqE[THEN "\\<^sub>d\<^sub>fI"] "&I" "F-u[den]"[unvarify F] + Ordinary.GEN "\I"; "cqt:2"?) + fix u + AOT_have \[[\x \<^bold>\[G]x \ [(=\<^sub>E)]xy]\<^sup>-\<^sup>y]u \ ([\x \<^bold>\[G]x \ x =\<^sub>E y]u) & u \\<^sub>E y\ + apply (rule "F-u"[THEN "=\<^sub>d\<^sub>fI"(1)[where \\<^sub>1\\<^sub>n=\(_,_)\], simplified]; "cqt:2"?) + by (rule "beta-C-cor:2"[THEN "\E", THEN "\E"(2)]; "cqt:2") + also AOT_have \\ \ (\<^bold>\[G]u \ u =\<^sub>E y) & u \\<^sub>E y\ + apply (AOT_subst \[\x \<^bold>\[G]x \ [(=\<^sub>E)]xy]u\ \\<^bold>\[G]u \ u =\<^sub>E y\) + apply (rule "beta-C-cor:2"[THEN "\E", THEN "\E"(2)]; "cqt:2") + using "oth-class-taut:3:a" by blast + also AOT_have \\ \ \<^bold>\[G]u\ + proof(safe intro!: "\I" "\I") + AOT_assume \(\<^bold>\[G]u \ u =\<^sub>E y) & u \\<^sub>E y\ + AOT_thus \\<^bold>\[G]u\ + by (metis "&E"(1) "&E"(2) "\E"(3) "\E"(1) "thm-neg=E") + next + AOT_assume \\<^bold>\[G]u\ + AOT_hence \u \\<^sub>E y\ and \\<^bold>\[G]u \ u =\<^sub>E y\ + using 0[THEN "\E"(2), THEN "\E", OF Ordinary.\, THEN "\E"] + "\I" by blast+ + AOT_thus \(\<^bold>\[G]u \ u =\<^sub>E y) & u \\<^sub>E y\ + using "&I" by simp + qed + also AOT_have \\ \ [\x \<^bold>\[G]x]u\ + by (rule "beta-C-cor:2"[THEN "\E", THEN "\E"(2), symmetric]; "cqt:2") + finally AOT_show \[[\x \<^bold>\[G]x \ [(=\<^sub>E)]xy]\<^sup>-\<^sup>y]u \ [\x \<^bold>\[G]x]u\. + qed + qed + AOT_have 2: \[\x \<^bold>\[G]x]\\ by "cqt:2[lambda]" + AOT_show \Numbers(n,[\x \<^bold>\[G]x \ x =\<^sub>E y]\<^sup>-\<^sup>y)\ + using "num-tran:1"[unvarify G H, OF 2, OF "F-u[den]"[unvarify F, OF 1], + THEN "\E", OF equinum, THEN "\E"(2), + OF "eq-num:2"[THEN "\E"(2), OF n_num_G]]. + qed + AOT_show \\\ ([\]\ & [\]n\ & \\ ([\]\ & [\]n\ \ \ = \))\ + proof(safe intro!: "\I"(2)[where \=b] "&I" Pnb "\I" GEN) + AOT_show \[\]b\ using "suc-num:1"[THEN "\E", OF Pnb]. + next + fix y + AOT_assume 0: \[\]y & [\]ny\ + AOT_show \y = b\ + apply (rule "pred-func:1"[THEN "\E"]) + using 0[THEN "&E"(2)] Pnb "&I" by blast + qed +qed + +(* Note the use of a bold '. *) +AOT_define Successor :: \\ \ \\<^sub>s\ (\_\<^bold>''\ [100] 100) + "def-suc": \n\<^bold>' =\<^sub>d\<^sub>f \<^bold>\m([\]nm)\ + +text\Note: not explicitly in PLM\ +AOT_theorem "def-suc[den1]": \\<^bold>\m([\]nm)\\ + using "A-Exists:2" "RA[2]" "\E"(2) "th-succ"[THEN "Number.\E"] by blast +text\Note: not explicitly in PLM\ +AOT_theorem "def-suc[den2]": shows \n\<^bold>'\\ + by (rule "def-suc"[THEN "=\<^sub>d\<^sub>fI"(1)]) + (auto simp: "def-suc[den1]") + +(* TODO: not in PLM *) +AOT_theorem suc_eq_desc: \n\<^bold>' = \<^bold>\m([\]nm)\ + by (rule "def-suc"[THEN "=\<^sub>d\<^sub>fI"(1)]) + (auto simp: "def-suc[den1]" "rule=I:1") + +AOT_theorem "suc-fact": \n = m \ n\<^bold>' = m\<^bold>'\ +proof (rule "\I") + AOT_assume 0: \n = m\ + AOT_show \n\<^bold>' = m\<^bold>'\ + apply (rule "rule=E"[rotated, OF 0]) + by (rule "=I"(1)[OF "def-suc[den2]"]) +qed + +AOT_theorem "ind-gnd": \m = 0 \ \n(m = n\<^bold>')\ +proof - + AOT_have \[[\]\<^sup>+]0m\ + using Number.\ "\E"(1) "nnumber:3" by blast + AOT_hence \[[\]\<^sup>*]0m \ 0 =\<^sub>\ m\ + using "assume1:5"[unvarify x, OF "zero:2", THEN "\E"(1)] by blast + moreover { + AOT_assume \[[\]\<^sup>*]0m\ + AOT_hence \\z ([[\]\<^sup>+]0z & [\]zm)\ + using "w-ances-her:7"[unconstrain \, unvarify \ x, OF "zero:2", + OF "pred-thm:2", THEN "\E", OF "pred-1-1:4", + THEN "\E"] + by blast + then AOT_obtain z where \: \[[\]\<^sup>+]0z\ and \: \[\]zm\ + using "&E" "\E"[rotated] by blast + AOT_have Nz: \[\]z\ + using \ "\E"(2) "nnumber:3" by blast + moreover AOT_have \m = z\<^bold>'\ + proof (rule "def-suc"[THEN "=\<^sub>d\<^sub>fI"(1)]; + safe intro!: "def-suc[den1]"[unconstrain n, THEN "\E", OF Nz] + "nec-hintikka-scheme"[THEN "\E"(2)] "&I" + GEN "\I" "Act-Basic:2"[THEN "\E"(2)]) + AOT_show \\<^bold>\[\]m\ using Number.\ + by (meson "mod-col-num:1" "nec-imp-act" "\E") + next + AOT_show \\<^bold>\[\]zm\ using \ + by (meson "nec-imp-act" "pred-1-1:1" "\E") + next + fix y + AOT_assume \\<^bold>\([\]y & [\]zy)\ + AOT_hence \\<^bold>\[\]y\ and \\<^bold>\[\]zy\ + using "Act-Basic:2" "&E" "\E"(1) by blast+ + AOT_hence 0: \[\]zy\ + by (metis RN "\E"(1) "pred-1-1:1" "sc-eq-fur:2" "\E") + AOT_thus \y = m\ + using "pred-func:1"[THEN "\E", OF "&I"] \ by metis + qed + ultimately AOT_have \[\]z & m = z\<^bold>'\ + by (rule "&I") + AOT_hence \\n m = n\<^bold>'\ + by (rule "\I") + hence ?thesis + by (rule "\I") + } + moreover { + AOT_assume \0 =\<^sub>\ m\ + AOT_hence \0 = m\ + using "id-R-thm:3"[unconstrain \, unvarify \ x, OF "zero:2", OF "pred-thm:2", + THEN "\E", OF "pred-1-1:4", THEN "\E"] + by auto + hence ?thesis using id_sym "\I" by blast + } + ultimately show ?thesis + by (metis "\E"(2) "raa-cor:1") +qed + +AOT_theorem "suc-thm": \[\]n n\<^bold>'\ +proof - + AOT_obtain x where m_is_n: \x = n\<^bold>'\ + using "free-thms:1"[THEN "\E"(1), OF "def-suc[den2]"] + using "\E" by metis + AOT_have \\<^bold>\([\]n\<^bold>' & [\]n n\<^bold>')\ + apply (rule "rule=E"[rotated, OF suc_eq_desc[symmetric]]) + apply (rule "actual-desc:4"[THEN "\E"]) + by (simp add: "def-suc[den1]") + AOT_hence \\<^bold>\[\]n\<^bold>'\ and \\<^bold>\[\]n n\<^bold>'\ + using "Act-Basic:2" "\E"(1) "&E" by blast+ + AOT_hence \\<^bold>\[\]nx\ + using m_is_n[symmetric] "rule=E" by fast+ + AOT_hence \[\]nx\ + by (metis RN "\E"(1) "pred-1-1:1" "sc-eq-fur:2" "\E") + thus ?thesis + using m_is_n "rule=E" by fast +qed + +AOT_define Numeral1 :: \\\<^sub>s\ ("1") + "numerals:1": \1 =\<^sub>d\<^sub>f 0\<^bold>'\ + +AOT_theorem "prec-facts:1": \[\]0 1\ + by (auto intro: "numerals:1"[THEN "rule-id-df:2:b[zero]", + OF "def-suc[den2]"[unconstrain n, unvarify \, + OF "zero:2", THEN "\E", OF "0-n"]] + "suc-thm"[unconstrain n, unvarify \, OF "zero:2", + THEN "\E", OF "0-n"]) + +(* TODO: more theorems *) + +(* Note: we forgo restricted variables for natural cardinals. *) +AOT_define Finite :: \\ \ \\ (\Finite'(_')\) + "inf-card:1": \Finite(x) \\<^sub>d\<^sub>f NaturalCardinal(x) & [\]x\ +AOT_define Infinite :: \\ \ \\ (\Infinite'(_')\) + "inf-card:2": \Infinite(x) \\<^sub>d\<^sub>f NaturalCardinal(x) & \Finite(x)\ + +AOT_theorem "inf-card-exist:1": \NaturalCardinal(#O!)\ + by (safe intro!: card[THEN "\\<^sub>d\<^sub>fI"] "\I"(1)[where \=\\O!\\] "=I" + "num-def:2"[unvarify G] "oa-exist:1") + +AOT_theorem "inf-card-exist:2": \Infinite(#O!)\ +proof (safe intro!: "inf-card:2"[THEN "\\<^sub>d\<^sub>fI"] "&I" "inf-card-exist:1") + AOT_show \\Finite(#O!)\ + proof(rule "raa-cor:2") + AOT_assume \Finite(#O!)\ + AOT_hence 0: \[\]#O!\ + using "inf-card:1"[THEN "\\<^sub>d\<^sub>fE"] "&E"(2) by blast + AOT_have \Numbers(#O!, [\z \<^bold>\O!z])\ + using "eq-num:3"[unvarify G, OF "oa-exist:1"]. + AOT_hence \#O! = #O!\ + using "eq-num:2"[unvarify x G, THEN "\E"(1), OF "oa-exist:1", + OF "num-def:2"[unvarify G], OF "oa-exist:1"] + by blast + AOT_hence \[\]#O! & #O! = #O!\ + using 0 "&I" by blast + AOT_hence \\x ([\]x & x = #O!)\ + using "num-def:2"[unvarify G, OF "oa-exist:1"] "\I"(1) by fast + AOT_hence \\\y ([E!]y & \u (\<^bold>\[O!]u \ u \\<^sub>E y))\ + using "modal-axiom"[axiom_inst, unvarify G, THEN "\E", OF "oa-exist:1"] by blast + AOT_hence \\y \([E!]y & \u (\<^bold>\[O!]u \ u \\<^sub>E y))\ + using "BF\"[THEN "\E"] by blast + then AOT_obtain b where \\([E!]b & \u (\<^bold>\[O!]u \ u \\<^sub>E b))\ + using "\E"[rotated] by blast + AOT_hence \\[E!]b\ and 2: \\\u (\<^bold>\[O!]u \ u \\<^sub>E b)\ + using "KBasic2:3"[THEN "\E"] "&E" by blast+ + AOT_hence \[\x \[E!]x]b\ + by (auto intro!: "\\C"(1) "cqt:2") + moreover AOT_have \O! = [\x \[E!]x]\ + by (rule "rule-id-df:1[zero]"[OF "oa:1"]) "cqt:2" + ultimately AOT_have b_ord: \O!b\ + using "rule=E" id_sym by fast + AOT_hence \\<^bold>\O!b\ + by (meson "\E"(1) "oa-facts:7") + moreover AOT_have 2: \\u (\<^bold>\[O!]u \ u \\<^sub>E b)\ + using "modal-lemma"[unvarify G, unconstrain v, OF "oa-exist:1", + THEN "\E", OF b_ord, THEN "\E", OF 2]. + ultimately AOT_have \b \\<^sub>E b\ + using "Ordinary.\E"[OF 2, unconstrain \, THEN "\E", + OF b_ord, THEN "\E"] by blast + AOT_hence \\(b =\<^sub>E b)\ + by (metis "\E"(1) "thm-neg=E") + moreover AOT_have \b =\<^sub>E b\ + using "ord=Eequiv:1"[THEN "\E", OF b_ord]. + ultimately AOT_show \p & \p\ for p + by (metis "raa-cor:3") + qed +qed + + + +(*<*) +end +(*>*) diff --git a/thys/AOT/AOT_PLM.thy b/thys/AOT/AOT_PLM.thy new file mode 100644 --- /dev/null +++ b/thys/AOT/AOT_PLM.thy @@ -0,0 +1,9080 @@ +(*<*) +theory AOT_PLM + imports AOT_Axioms +begin +(*>*) + +section\The Deductive System PLM\ +text\\label{PLM: 9}\ + +(* constrain sledgehammer to the abstraction layer *) +unbundle AOT_no_atp + +subsection\Primitive Rule of PLM: Modus Ponens\ +text\\label{PLM: 9.1}\ + +AOT_theorem "modus-ponens": + assumes \\\ and \\ \ \\ + shows \\\ + (* NOTE: semantics needed *) + using assms by (simp add: AOT_sem_imp) +lemmas MP = "modus-ponens" + +subsection\(Modally Strict) Proofs and Derivations\ +text\\label{PLM: 9.2}\ + +AOT_theorem "non-con-thm-thm": + assumes \\<^bold>\\<^sub>\ \\ + shows \\<^bold>\ \\ + using assms by simp + +AOT_theorem "vdash-properties:1[1]": + assumes \\ \ \\ + shows \\<^bold>\ \\ + (* NOTE: semantics needed *) + using assms unfolding AOT_model_act_axiom_def by blast + +text\Convenience attribute for instantiating modally-fragile axioms.\ +attribute_setup act_axiom_inst = + \Scan.succeed (Thm.rule_attribute [] + (K (fn thm => thm RS @{thm "vdash-properties:1[1]"})))\ + "Instantiate modally fragile axiom as modally fragile theorem." + +AOT_theorem "vdash-properties:1[2]": + assumes \\ \ \\<^sub>\\ + shows \\<^bold>\\<^sub>\ \\ + (* NOTE: semantics needed *) + using assms unfolding AOT_model_axiom_def by blast + +text\Convenience attribute for instantiating modally-strict axioms.\ +attribute_setup axiom_inst = + \Scan.succeed (Thm.rule_attribute [] + (K (fn thm => thm RS @{thm "vdash-properties:1[2]"})))\ + "Instantiate axiom as theorem." + +text\Convenience methods and theorem sets for applying "cqt:2".\ +method cqt_2_lambda_inst_prover = + (fast intro: AOT_instance_of_cqt_2_intro) +method "cqt:2[lambda]" = + (rule "cqt:2[lambda]"[axiom_inst]; cqt_2_lambda_inst_prover) +lemmas "cqt:2" = + "cqt:2[const_var]"[axiom_inst] "cqt:2[lambda]"[axiom_inst] + AOT_instance_of_cqt_2_intro +method "cqt:2" = (safe intro!: "cqt:2") + +AOT_theorem "vdash-properties:3": + assumes \\<^bold>\\<^sub>\ \\ + shows \\ \<^bold>\ \\ + using assms by blast + +AOT_theorem "vdash-properties:5": + assumes \\\<^sub>1 \<^bold>\ \\ and \\\<^sub>2 \<^bold>\ \ \ \\ + shows \\\<^sub>1, \\<^sub>2 \<^bold>\ \\ + using MP assms by blast + +AOT_theorem "vdash-properties:6": + assumes \\\ and \\ \ \\ + shows \\\ + using MP assms by blast + +AOT_theorem "vdash-properties:8": + assumes \\ \<^bold>\ \\ and \\ \<^bold>\ \\ + shows \\ \<^bold>\ \\ + using assms by argo + +AOT_theorem "vdash-properties:9": + assumes \\\ + shows \\ \ \\ + using MP "pl:1"[axiom_inst] assms by blast + +AOT_theorem "vdash-properties:10": + assumes \\ \ \\ and \\\ + shows \\\ + using MP assms by blast +lemmas "\E" = "vdash-properties:10" + +subsection\Two Fundamental Metarules: GEN and RN\ +text\\label{PLM: 9.3}\ + +AOT_theorem "rule-gen": + assumes \for arbitrary \: \{\}\ + shows \\\ \{\}\ + (* NOTE: semantics needed *) + using assms by (metis AOT_var_of_term_inverse AOT_sem_denotes AOT_sem_forall) +lemmas GEN = "rule-gen" + +AOT_theorem "RN[prem]": + assumes \\ \<^bold>\\<^sub>\ \\ + shows \\\ \<^bold>\\<^sub>\ \\\ + by (meson AOT_sem_box assms image_iff) (* NOTE: semantics needed *) +AOT_theorem RN: + assumes \\<^bold>\\<^sub>\ \\ + shows \\\\ + using "RN[prem]" assms by blast + +subsection\The Inferential Role of Definitions\ +text\\label{PLM: 9.4}\ + +AOT_axiom "df-rules-formulas[1]": + assumes \\ \\<^sub>d\<^sub>f \\ + shows \\ \ \\ + (* NOTE: semantics needed *) + using assms + by (auto simp: assms AOT_model_axiomI AOT_model_equiv_def AOT_sem_imp) +AOT_axiom "df-rules-formulas[2]": + assumes \\ \\<^sub>d\<^sub>f \\ + shows \\ \ \\ + (* NOTE: semantics needed *) + using assms + by (auto simp: AOT_model_axiomI AOT_model_equiv_def AOT_sem_imp) +(* NOTE: for convenience also state the above as regular theorems *) +AOT_theorem "df-rules-formulas[3]": + assumes \\ \\<^sub>d\<^sub>f \\ + shows \\ \ \\ + using "df-rules-formulas[1]"[axiom_inst, OF assms]. +AOT_theorem "df-rules-formulas[4]": + assumes \\ \\<^sub>d\<^sub>f \\ + shows \\ \ \\ + using "df-rules-formulas[2]"[axiom_inst, OF assms]. + + +AOT_axiom "df-rules-terms[1]": + assumes \\{\\<^sub>1...\\<^sub>n} =\<^sub>d\<^sub>f \{\\<^sub>1...\\<^sub>n}\ + shows \(\{\\<^sub>1...\\<^sub>n}\ \ \{\\<^sub>1...\\<^sub>n} = \{\\<^sub>1...\\<^sub>n}) & + (\\{\\<^sub>1...\\<^sub>n}\ \ \\{\\<^sub>1...\\<^sub>n}\)\ + (* NOTE: semantics needed *) + using assms + by (simp add: AOT_model_axiomI AOT_sem_conj AOT_sem_imp AOT_sem_eq + AOT_sem_not AOT_sem_denotes AOT_model_id_def) +AOT_axiom "df-rules-terms[2]": + assumes \\ =\<^sub>d\<^sub>f \\ + shows \(\\ \ \ = \) & (\\\ \ \\\)\ + by (metis "df-rules-terms[1]" case_unit_Unity assms) +(* NOTE: for convenience also state the above as regular theorems *) +AOT_theorem "df-rules-terms[3]": + assumes \\{\\<^sub>1...\\<^sub>n} =\<^sub>d\<^sub>f \{\\<^sub>1...\\<^sub>n}\ + shows \(\{\\<^sub>1...\\<^sub>n}\ \ \{\\<^sub>1...\\<^sub>n} = \{\\<^sub>1...\\<^sub>n}) & + (\\{\\<^sub>1...\\<^sub>n}\ \ \\{\\<^sub>1...\\<^sub>n}\)\ + using "df-rules-terms[1]"[axiom_inst, OF assms]. +AOT_theorem "df-rules-terms[4]": + assumes \\ =\<^sub>d\<^sub>f \\ + shows \(\\ \ \ = \) & (\\\ \ \\\)\ + using "df-rules-terms[2]"[axiom_inst, OF assms]. + +subsection\The Theory of Negations and Conditionals\ +text\\label{PLM: 9.5}\ + +AOT_theorem "if-p-then-p": \\ \ \\ + by (meson "pl:1"[axiom_inst] "pl:2"[axiom_inst] MP) + +AOT_theorem "deduction-theorem": + assumes \\ \<^bold>\ \\ + shows \\ \ \\ + (* NOTE: semantics needed *) + using assms by (simp add: AOT_sem_imp) +lemmas CP = "deduction-theorem" +lemmas "\I" = "deduction-theorem" + +AOT_theorem "ded-thm-cor:1": + assumes \\\<^sub>1 \<^bold>\ \ \ \\ and \\\<^sub>2 \<^bold>\ \ \ \\ + shows \\\<^sub>1, \\<^sub>2 \<^bold>\ \ \ \\ + using "\E" "\I" assms by blast +AOT_theorem "ded-thm-cor:2": + assumes \\\<^sub>1 \<^bold>\ \ \ (\ \ \)\ and \\\<^sub>2 \<^bold>\ \\ + shows \\\<^sub>1, \\<^sub>2 \<^bold>\ \ \ \\ + using "\E" "\I" assms by blast + +AOT_theorem "ded-thm-cor:3": + assumes \\ \ \\ and \\ \ \\ + shows \\ \ \\ + using "\E" "\I" assms by blast +declare "ded-thm-cor:3"[trans] +AOT_theorem "ded-thm-cor:4": + assumes \\ \ (\ \ \)\ and \\\ + shows \\ \ \\ + using "\E" "\I" assms by blast + +lemmas "Hypothetical Syllogism" = "ded-thm-cor:3" + +AOT_theorem "useful-tautologies:1": \\\\ \ \\ + by (metis "pl:3"[axiom_inst] "\I" "Hypothetical Syllogism") +AOT_theorem "useful-tautologies:2": \\ \ \\\\ + by (metis "pl:3"[axiom_inst] "\I" "ded-thm-cor:4") +AOT_theorem "useful-tautologies:3": \\\ \ (\ \ \)\ + by (meson "ded-thm-cor:4" "pl:3"[axiom_inst] "\I") +AOT_theorem "useful-tautologies:4": \(\\ \ \\) \ (\ \ \)\ + by (meson "pl:3"[axiom_inst] "Hypothetical Syllogism" "\I") +AOT_theorem "useful-tautologies:5": \(\ \ \) \ (\\ \ \\)\ + by (metis "useful-tautologies:4" "Hypothetical Syllogism" "\I") + +AOT_theorem "useful-tautologies:6": \(\ \ \\) \ (\ \ \\)\ + by (metis "\I" MP "useful-tautologies:4") + +AOT_theorem "useful-tautologies:7": \(\\ \ \) \ (\\ \ \)\ + by (metis "\I" MP "useful-tautologies:3" "useful-tautologies:5") + +AOT_theorem "useful-tautologies:8": \\ \ (\\ \ \(\ \ \))\ + by (metis "\I" MP "useful-tautologies:5") + +AOT_theorem "useful-tautologies:9": \(\ \ \) \ ((\\ \ \) \ \)\ + by (metis "\I" MP "useful-tautologies:6") + +AOT_theorem "useful-tautologies:10": \(\ \ \\) \ ((\ \ \) \ \\)\ + by (metis "\I" MP "pl:3"[axiom_inst]) + +AOT_theorem "dn-i-e:1": + assumes \\\ + shows \\\\\ + using MP "useful-tautologies:2" assms by blast +lemmas "\\I" = "dn-i-e:1" +AOT_theorem "dn-i-e:2": + assumes \\\\\ + shows \\\ + using MP "useful-tautologies:1" assms by blast +lemmas "\\E" = "dn-i-e:2" + +AOT_theorem "modus-tollens:1": + assumes \\ \ \\ and \\\\ + shows \\\\ + using MP "useful-tautologies:5" assms by blast +AOT_theorem "modus-tollens:2": + assumes \\ \ \\\ and \\\ + shows \\\\ + using "\\I" "modus-tollens:1" assms by blast +lemmas MT = "modus-tollens:1" "modus-tollens:2" + +AOT_theorem "contraposition:1[1]": + assumes \\ \ \\ + shows \\\ \ \\\ + using "\I" MT(1) assms by blast +AOT_theorem "contraposition:1[2]": + assumes \\\ \ \\\ + shows \\ \ \\ + using "\I" "\\E" MT(2) assms by blast + +AOT_theorem "contraposition:2": + assumes \\ \ \\\ + shows \\ \ \\\ + using "\I" MT(2) assms by blast + +AOT_theorem "reductio-aa:1": + assumes \\\ \<^bold>\ \\\ and \\\ \<^bold>\ \\ + shows \\\ + using "\I" "\\E" MT(2) assms by blast +AOT_theorem "reductio-aa:2": + assumes \\ \<^bold>\ \\\ and \\ \<^bold>\ \\ + shows \\\\ + using "reductio-aa:1" assms by blast +lemmas "RAA" = "reductio-aa:1" "reductio-aa:2" + +AOT_theorem "exc-mid": \\ \ \\\ + using "df-rules-formulas[4]" "if-p-then-p" MP + "conventions:2" by blast + +AOT_theorem "non-contradiction": \\(\ & \\)\ + using "df-rules-formulas[3]" MT(2) "useful-tautologies:2" + "conventions:1" by blast + +AOT_theorem "con-dis-taut:1": \(\ & \) \ \\ + by (meson "\I" "df-rules-formulas[3]" MP RAA(1) "conventions:1") +AOT_theorem "con-dis-taut:2": \(\ & \) \ \\ + by (metis "\I" "df-rules-formulas[3]" MT(2) RAA(2) + "\\E" "conventions:1") +lemmas "Conjunction Simplification" = "con-dis-taut:1" "con-dis-taut:2" + +AOT_theorem "con-dis-taut:3": \\ \ (\ \ \)\ + by (meson "contraposition:1[2]" "df-rules-formulas[4]" + MP "\I" "conventions:2") +AOT_theorem "con-dis-taut:4": \\ \ (\ \ \)\ + using "Hypothetical Syllogism" "df-rules-formulas[4]" + "pl:1"[axiom_inst] "conventions:2" by blast +lemmas "Disjunction Addition" = "con-dis-taut:3" "con-dis-taut:4" + +AOT_theorem "con-dis-taut:5": \\ \ (\ \ (\ & \))\ + by (metis "contraposition:2" "Hypothetical Syllogism" "\I" + "df-rules-formulas[4]" "conventions:1") +lemmas Adjunction = "con-dis-taut:5" + +AOT_theorem "con-dis-taut:6": \(\ & \) \ \\ + by (metis Adjunction "\I" "df-rules-formulas[4]" MP + "Conjunction Simplification"(1) "conventions:3") +lemmas "Idempotence of &" = "con-dis-taut:6" + +AOT_theorem "con-dis-taut:7": \(\ \ \) \ \\ +proof - + { + AOT_assume \\ \ \\ + AOT_hence \\\ \ \\ + using "conventions:2"[THEN "df-rules-formulas[3]"] MP by blast + AOT_hence \\\ using "if-p-then-p" RAA(1) MP by blast + } + moreover { + AOT_assume \\\ + AOT_hence \\ \ \\ using "Disjunction Addition"(1) MP by blast + } + ultimately AOT_show \(\ \ \) \ \\ + using "conventions:3"[THEN "df-rules-formulas[4]"] MP + by (metis Adjunction "\I") +qed +lemmas "Idempotence of \" = "con-dis-taut:7" + + +AOT_theorem "con-dis-i-e:1": + assumes \\\ and \\\ + shows \\ & \\ + using Adjunction MP assms by blast +lemmas "&I" = "con-dis-i-e:1" + +AOT_theorem "con-dis-i-e:2:a": + assumes \\ & \\ + shows \\\ + using "Conjunction Simplification"(1) MP assms by blast +AOT_theorem "con-dis-i-e:2:b": + assumes \\ & \\ + shows \\\ + using "Conjunction Simplification"(2) MP assms by blast +lemmas "&E" = "con-dis-i-e:2:a" "con-dis-i-e:2:b" + +AOT_theorem "con-dis-i-e:3:a": + assumes \\\ + shows \\ \ \\ + using "Disjunction Addition"(1) MP assms by blast +AOT_theorem "con-dis-i-e:3:b": + assumes \\\ + shows \\ \ \\ + using "Disjunction Addition"(2) MP assms by blast +AOT_theorem "con-dis-i-e:3:c": + assumes \\ \ \\ and \\ \ \\ and \\ \ \\ + shows \\ \ \\ + by (metis "con-dis-i-e:3:a" "Disjunction Addition"(2) + "df-rules-formulas[3]" MT(1) RAA(1) + "conventions:2" assms) +lemmas "\I" = "con-dis-i-e:3:a" "con-dis-i-e:3:b" "con-dis-i-e:3:c" + +AOT_theorem "con-dis-i-e:4:a": + assumes \\ \ \\ and \\ \ \\ and \\ \ \\ + shows \\\ + by (metis MP RAA(2) "df-rules-formulas[3]" "conventions:2" assms) +AOT_theorem "con-dis-i-e:4:b": + assumes \\ \ \\ and \\\\ + shows \\\ + using "con-dis-i-e:4:a" RAA(1) "\I" assms by blast +AOT_theorem "con-dis-i-e:4:c": + assumes \\ \ \\ and \\\\ + shows \\\ + using "con-dis-i-e:4:a" RAA(1) "\I" assms by blast +lemmas "\E" = "con-dis-i-e:4:a" "con-dis-i-e:4:b" "con-dis-i-e:4:c" + +AOT_theorem "raa-cor:1": + assumes \\\ \<^bold>\ \ & \\\ + shows \\\ + using "&E" "\E"(3) "\I"(2) RAA(2) assms by blast +AOT_theorem "raa-cor:2": + assumes \\ \<^bold>\ \ & \\\ + shows \\\\ + using "raa-cor:1" assms by blast +AOT_theorem "raa-cor:3": + assumes \\\ and \\\ \<^bold>\ \\\ + shows \\\ + using RAA assms by blast +AOT_theorem "raa-cor:4": + assumes \\\\ and \\\ \<^bold>\ \\ + shows \\\ + using RAA assms by blast +AOT_theorem "raa-cor:5": + assumes \\\ and \\ \<^bold>\ \\\ + shows \\\\ + using RAA assms by blast +AOT_theorem "raa-cor:6": + assumes \\\\ and \\ \<^bold>\ \\ + shows \\\\ + using RAA assms by blast + +AOT_theorem "oth-class-taut:1:a": \(\ \ \) \ \(\ & \\)\ + by (rule "conventions:3"[THEN "df-rules-formulas[4]", THEN "\E"]) + (metis "&E" "&I" "raa-cor:3" "\I" MP) +AOT_theorem "oth-class-taut:1:b": \\(\ \ \) \ (\ & \\)\ + by (rule "conventions:3"[THEN "df-rules-formulas[4]", THEN "\E"]) + (metis "&E" "&I" "raa-cor:3" "\I" MP) +AOT_theorem "oth-class-taut:1:c": \(\ \ \) \ (\\ \ \)\ + by (rule "conventions:3"[THEN "df-rules-formulas[4]", THEN "\E"]) + (metis "&I" "\I"(1, 2) "\E"(3) "\I" MP "raa-cor:1") + +AOT_theorem "oth-class-taut:2:a": \(\ & \) \ (\ & \)\ + by (rule "conventions:3"[THEN "df-rules-formulas[4]", THEN "\E"]) + (meson "&I" "&E" "\I") +lemmas "Commutativity of &" = "oth-class-taut:2:a" +AOT_theorem "oth-class-taut:2:b": \(\ & (\ & \)) \ ((\ & \) & \)\ + by (rule "conventions:3"[THEN "df-rules-formulas[4]", THEN "\E"]) + (metis "&I" "&E" "\I") +lemmas "Associativity of &" = "oth-class-taut:2:b" +AOT_theorem "oth-class-taut:2:c": \(\ \ \) \ (\ \ \)\ + by (rule "conventions:3"[THEN "df-rules-formulas[4]", THEN "\E"]) + (metis "&I" "\I"(1, 2) "\E"(1) "\I") +lemmas "Commutativity of \" = "oth-class-taut:2:c" +AOT_theorem "oth-class-taut:2:d": \(\ \ (\ \ \)) \ ((\ \ \) \ \)\ + by (rule "conventions:3"[THEN "df-rules-formulas[4]", THEN "\E"]) + (metis "&I" "\I"(1, 2) "\E"(1) "\I") +lemmas "Associativity of \" = "oth-class-taut:2:d" +AOT_theorem "oth-class-taut:2:e": \(\ \ \) \ (\ \ \)\ + by (rule "conventions:3"[THEN "df-rules-formulas[4]", THEN "\E"]; rule "&I"; + metis "&I" "df-rules-formulas[4]" "conventions:3" "&E" + "Hypothetical Syllogism" "\I" "df-rules-formulas[3]") +lemmas "Commutativity of \" = "oth-class-taut:2:e" +AOT_theorem "oth-class-taut:2:f": \(\ \ (\ \ \)) \ ((\ \ \) \ \)\ + using "conventions:3"[THEN "df-rules-formulas[4]"] + "conventions:3"[THEN "df-rules-formulas[3]"] + "\I" "\E" "&E" "&I" + by metis +lemmas "Associativity of \" = "oth-class-taut:2:f" + +AOT_theorem "oth-class-taut:3:a": \\ \ \\ + using "&I" "vdash-properties:6" "if-p-then-p" + "df-rules-formulas[4]" "conventions:3" by blast +AOT_theorem "oth-class-taut:3:b": \\ \ \\\\ + using "&I" "useful-tautologies:1" "useful-tautologies:2" "\E" + "df-rules-formulas[4]" "conventions:3" by blast +AOT_theorem "oth-class-taut:3:c": \\(\ \ \\)\ + by (metis "&E" "\E" RAA "df-rules-formulas[3]" "conventions:3") + +AOT_theorem "oth-class-taut:4:a": \(\ \ \) \ ((\ \ \) \ (\ \ \))\ + by (metis "\E" "\I") +AOT_theorem "oth-class-taut:4:b": \(\ \ \) \ (\\ \ \\)\ + using "conventions:3"[THEN "df-rules-formulas[4]"] + "conventions:3"[THEN "df-rules-formulas[3]"] + "\I" "\E" "&E" "&I" RAA by metis +AOT_theorem "oth-class-taut:4:c": \(\ \ \) \ ((\ \ \) \ (\ \ \))\ + using "conventions:3"[THEN "df-rules-formulas[4]"] + "conventions:3"[THEN "df-rules-formulas[3]"] + "\I" "\E" "&E" "&I" by metis +AOT_theorem "oth-class-taut:4:d": \(\ \ \) \ ((\ \ \) \ (\ \ \))\ + using "conventions:3"[THEN "df-rules-formulas[4]"] + "conventions:3"[THEN "df-rules-formulas[3]"] + "\I" "\E" "&E" "&I" by metis +AOT_theorem "oth-class-taut:4:e": \(\ \ \) \ ((\ & \) \ (\ & \))\ + using "conventions:3"[THEN "df-rules-formulas[4]"] + "conventions:3"[THEN "df-rules-formulas[3]"] + "\I" "\E" "&E" "&I" by metis +AOT_theorem "oth-class-taut:4:f": \(\ \ \) \ ((\ & \) \ (\ & \))\ + using "conventions:3"[THEN "df-rules-formulas[4]"] + "conventions:3"[THEN "df-rules-formulas[3]"] + "\I" "\E" "&E" "&I" by metis +AOT_theorem "oth-class-taut:4:g": \(\ \ \) \ ((\ & \) \ (\\ & \\))\ +proof(safe intro!: "conventions:3"[THEN "df-rules-formulas[4]", THEN "\E"] + "&I" "\I" + dest!: "conventions:3"[THEN "df-rules-formulas[3]", THEN "\E"]) + AOT_show \\ & \ \ (\\ & \\)\ if \(\ \ \) & (\ \ \)\ + using "&E" "\I" "\E" "&I" "raa-cor:1" "\I" "\E" that by metis +next + AOT_show \\\ if \\ & \ \ (\\ & \\)\ and \\\ + using that "\E" "&E" "raa-cor:3" by blast +next + AOT_show \\\ if \\ & \ \ (\\ & \\)\ and \\\ + using that "\E" "&E" "raa-cor:3" by blast +qed +AOT_theorem "oth-class-taut:4:h": \\(\ \ \) \ ((\ & \\) \ (\\ & \))\ +proof (safe intro!: "conventions:3"[THEN "df-rules-formulas[4]", THEN "\E"] + "&I" "\I") + AOT_show \\ & \\ \ (\\ & \)\ if \\(\ \ \)\ + by (metis that "&I" "\I"(1, 2) "\I" MT(1) "df-rules-formulas[4]" + "raa-cor:3" "conventions:3") +next + AOT_show \\(\ \ \)\ if \\ & \\ \ (\\ & \)\ + by (metis that "&E" "\E"(2) "\E" "df-rules-formulas[3]" + "raa-cor:3" "conventions:3") +qed +AOT_theorem "oth-class-taut:5:a": \(\ & \) \ \(\\ \ \\)\ + using "conventions:3"[THEN "df-rules-formulas[4]"] + "\I" "\E" "&E" "&I" "\I" "\E" RAA by metis +AOT_theorem "oth-class-taut:5:b": \(\ \ \) \ \(\\ & \\)\ + using "conventions:3"[THEN "df-rules-formulas[4]"] + "\I" "\E" "&E" "&I" "\I" "\E" RAA by metis +AOT_theorem "oth-class-taut:5:c": \\(\ & \) \ (\\ \ \\)\ + using "conventions:3"[THEN "df-rules-formulas[4]"] + "\I" "\E" "&E" "&I" "\I" "\E" RAA by metis +AOT_theorem "oth-class-taut:5:d": \\(\ \ \) \ (\\ & \\)\ + using "conventions:3"[THEN "df-rules-formulas[4]"] + "\I" "\E" "&E" "&I" "\I" "\E" RAA by metis + +lemmas DeMorgan = "oth-class-taut:5:c" "oth-class-taut:5:d" + +AOT_theorem "oth-class-taut:6:a": + \(\ & (\ \ \)) \ ((\ & \) \ (\ & \))\ + using "conventions:3"[THEN "df-rules-formulas[4]"] + "\I" "\E" "&E" "&I" "\I" "\E" RAA by metis +AOT_theorem "oth-class-taut:6:b": + \(\ \ (\ & \)) \ ((\ \ \) & (\ \ \))\ + using "conventions:3"[THEN "df-rules-formulas[4]"] + "\I" "\E" "&E" "&I" "\I" "\E" RAA by metis + +AOT_theorem "oth-class-taut:7:a": \((\ & \) \ \) \ (\ \ (\ \ \))\ + by (metis "&I" "\E" "\I") +lemmas Exportation = "oth-class-taut:7:a" +AOT_theorem "oth-class-taut:7:b": \(\ \ (\ \\)) \ ((\ & \) \ \)\ + by (metis "&E" "\E" "\I") +lemmas Importation = "oth-class-taut:7:b" + +AOT_theorem "oth-class-taut:8:a": + \(\ \ (\ \ \)) \ (\ \ (\ \ \))\ + using "conventions:3"[THEN "df-rules-formulas[4]"] "\I" "\E" "&E" "&I" + by metis +lemmas Permutation = "oth-class-taut:8:a" +AOT_theorem "oth-class-taut:8:b": + \(\ \ \) \ ((\ \ \) \ (\ \ (\ & \)))\ + by (metis "&I" "\E" "\I") +lemmas Composition = "oth-class-taut:8:b" +AOT_theorem "oth-class-taut:8:c": + \(\ \ \) \ ((\ \ \) \ ((\ \ \) \ \))\ + by (metis "\E"(2) "\E" "\I" RAA(1)) +AOT_theorem "oth-class-taut:8:d": + \((\ \ \) & (\ \ \)) \ ((\ & \) \ (\ & \))\ + by (metis "&E" "&I" "\E" "\I") +lemmas "Double Composition" = "oth-class-taut:8:d" +AOT_theorem "oth-class-taut:8:e": + \((\ & \) \ (\ & \)) \ (\ \ (\ \ \))\ + by (metis "conventions:3"[THEN "df-rules-formulas[4]"] + "conventions:3"[THEN "df-rules-formulas[3]"] + "\I" "\E" "&E" "&I") +AOT_theorem "oth-class-taut:8:f": + \((\ & \) \ (\ & \)) \ (\ \ (\ \ \))\ + by (metis "conventions:3"[THEN "df-rules-formulas[4]"] + "conventions:3"[THEN "df-rules-formulas[3]"] + "\I" "\E" "&E" "&I") +AOT_theorem "oth-class-taut:8:g": + \(\ \ \) \ ((\ \ \) \ (\ \ \))\ + by (metis "conventions:3"[THEN "df-rules-formulas[4]"] + "conventions:3"[THEN "df-rules-formulas[3]"] + "\I" "\E" "&E" "&I" "\I" "\E"(1)) +AOT_theorem "oth-class-taut:8:h": + \(\ \ \) \ ((\ \ \) \ (\ \ \))\ + by (metis "conventions:3"[THEN "df-rules-formulas[4]"] + "conventions:3"[THEN "df-rules-formulas[3]"] + "\I" "\E" "&E" "&I" "\I" "\E"(1)) +AOT_theorem "oth-class-taut:8:i": + \(\ \ (\ & \)) \ (\ \ (\ \ \))\ + by (metis "conventions:3"[THEN "df-rules-formulas[4]"] + "conventions:3"[THEN "df-rules-formulas[3]"] + "\I" "\E" "&E" "&I") + +AOT_theorem "intro-elim:1": + assumes \\ \ \\ and \\ \ \\ and \\ \ \\ + shows \\ \ \\ + by (metis assms "\I"(1, 2) "\E"(1) "\I" "\E" "&E"(1) + "conventions:3"[THEN "df-rules-formulas[3]"]) + +AOT_theorem "intro-elim:2": + assumes \\ \ \\ and \\ \ \\ + shows \\ \ \\ + by (meson "&I" "conventions:3" "df-rules-formulas[4]" MP assms) +lemmas "\I" = "intro-elim:2" + +AOT_theorem "intro-elim:3:a": + assumes \\ \ \\ and \\\ + shows \\\ + by (metis "\I"(1) "\I" "\E"(1) "intro-elim:1" assms) +AOT_theorem "intro-elim:3:b": + assumes \\ \ \\ and \\\ + shows \\\ + using "intro-elim:3:a" "Commutativity of \" assms by blast +AOT_theorem "intro-elim:3:c": + assumes \\ \ \\ and \\\\ + shows \\\\ + using "intro-elim:3:b" "raa-cor:3" assms by blast +AOT_theorem "intro-elim:3:d": + assumes \\ \ \\ and \\\\ + shows \\\\ + using "intro-elim:3:a" "raa-cor:3" assms by blast +AOT_theorem "intro-elim:3:e": + assumes \\ \ \\ and \\ \ \\ + shows \\ \ \\ + by (metis "\I" "\I" "intro-elim:3:a" "intro-elim:3:b" assms) +declare "intro-elim:3:e"[trans] +AOT_theorem "intro-elim:3:f": + assumes \\ \ \\ and \\ \ \\ + shows \\ \ \\ + by (metis "\I" "\I" "intro-elim:3:a" "intro-elim:3:b" assms) +lemmas "\E" = "intro-elim:3:a" "intro-elim:3:b" "intro-elim:3:c" + "intro-elim:3:d" "intro-elim:3:e" "intro-elim:3:f" + +declare "Commutativity of \"[THEN "\E"(1), sym] + +AOT_theorem "rule-eq-df:1": + assumes \\ \\<^sub>d\<^sub>f \\ + shows \\ \ \\ + by (simp add: "\I" "df-rules-formulas[3]" "df-rules-formulas[4]" assms) +lemmas "\Df" = "rule-eq-df:1" +AOT_theorem "rule-eq-df:2": + assumes \\ \\<^sub>d\<^sub>f \\ and \\\ + shows \\\ + using "\Df" "\E"(1) assms by blast +lemmas "\\<^sub>d\<^sub>fE" = "rule-eq-df:2" +AOT_theorem "rule-eq-df:3": + assumes \\ \\<^sub>d\<^sub>f \\ and \\\ + shows \\\ + using "\Df" "\E"(2) assms by blast +lemmas "\\<^sub>d\<^sub>fI" = "rule-eq-df:3" + +AOT_theorem "df-simplify:1": + assumes \\ \ (\ & \)\ and \\\ + shows \\ \ \\ + by (metis "&E"(2) "&I" "\E"(1, 2) "\I" "\I" assms) +(* Note: this is a slight variation from PLM *) +AOT_theorem "df-simplify:2": + assumes \\ \ (\ & \)\ and \\\ + shows \\ \ \\ + by (metis "&E"(1) "&I" "\E"(1, 2) "\I" "\I" assms) +lemmas "\S" = "df-simplify:1" "df-simplify:2" + +subsection\The Theory of Quantification\ +text\\label{PLM: 9.6}\ + +AOT_theorem "rule-ui:1": + assumes \\\ \{\}\ and \\\\ + shows \\{\}\ + using "\E" "cqt:1"[axiom_inst] assms by blast +AOT_theorem "rule-ui:2[const_var]": + assumes \\\ \{\}\ + shows \\{\}\ + by (simp add: "rule-ui:1" "cqt:2[const_var]"[axiom_inst] assms) +AOT_theorem "rule-ui:2[lambda]": + assumes \\F \{F}\ and \INSTANCE_OF_CQT_2(\)\ + shows \\{[\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]}\ + by (simp add: "rule-ui:1" "cqt:2[lambda]"[axiom_inst] assms) +AOT_theorem "rule-ui:3": + assumes \\\ \{\}\ + shows \\{\}\ + by (simp add: "rule-ui:2[const_var]" assms) +lemmas "\E" = "rule-ui:1" "rule-ui:2[const_var]" + "rule-ui:2[lambda]" "rule-ui:3" + +AOT_theorem "cqt-orig:1[const_var]": \\\ \{\} \ \{\}\ + by (simp add: "\E"(2) "\I") +AOT_theorem "cqt-orig:1[lambda]": + assumes \INSTANCE_OF_CQT_2(\)\ + shows \\F \{F} \ \{[\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]}\ + by (simp add: "\E"(3) "\I" assms) +AOT_theorem "cqt-orig:2": \\\ (\ \ \{\}) \ (\ \ \\ \{\})\ + by (metis "\I" GEN "vdash-properties:6" "\E"(4)) +AOT_theorem "cqt-orig:3": \\\ \{\} \ \{\}\ + using "cqt-orig:1[const_var]". + +AOT_theorem universal: + assumes \for arbitrary \: \{\}\ + shows \\\ \{\}\ + using GEN assms . +lemmas "\I" = universal + +(* Generalized mechanism for \I followed by \E *) +ML\ +fun get_instantiated_allI ctxt varname thm = let +val trm = Thm.concl_of thm +val trm = + case trm of (@{const Trueprop} $ (@{const AOT_model_valid_in} $ _ $ x)) => x + | _ => raise Term.TERM ("Expected simple theorem.", [trm]) +fun extractVars (Const (\<^const_name>\AOT_term_of_var\, _) $ Var v) = + (if fst (fst v) = fst varname then [Var v] else []) + | extractVars (t1 $ t2) = extractVars t1 @ extractVars t2 + | extractVars (Abs (_, _, t)) = extractVars t + | extractVars _ = [] +val vars = extractVars trm +val vars = fold Term.add_vars vars [] +val var = hd vars +val trmty = + case (snd var) of (Type (\<^type_name>\AOT_var\, [t])) => (t) + | _ => raise Term.TYPE ("Expected variable type.", [snd var], [Var var]) +val trm = Abs (Term.string_of_vname (fst var), trmty, Term.abstract_over ( + Const (\<^const_name>\AOT_term_of_var\, Type ("fun", [snd var, trmty])) + $ Var var, trm)) +val trm = Thm.cterm_of (Context.proof_of ctxt) trm +val ty = hd (Term.add_tvars (Thm.prop_of @{thm "\I"}) []) +val typ = Thm.ctyp_of (Context.proof_of ctxt) trmty +val allthm = Drule.instantiate_normalize (TVars.make [(ty, typ)], Vars.empty) @{thm "\I"} +val phi = hd (Term.add_vars (Thm.prop_of allthm) []) +val allthm = Drule.instantiate_normalize (TVars.empty, Vars.make [(phi,trm)]) allthm +in +allthm +end +\ + +attribute_setup "\I" = + \Scan.lift (Scan.repeat1 Args.var) >> (fn args => Thm.rule_attribute [] + (fn ctxt => fn thm => fold (fn arg => fn thm => + thm RS get_instantiated_allI ctxt arg thm) args thm))\ + "Quantify over a variable in a theorem using GEN." + +attribute_setup "unvarify" = + \Scan.lift (Scan.repeat1 Args.var) >> (fn args => Thm.rule_attribute [] + (fn ctxt => fn thm => + let + fun get_inst_allI arg thm = thm RS get_instantiated_allI ctxt arg thm + val thm = fold get_inst_allI args thm + val thm = fold (K (fn thm => thm RS @{thm "\E"(1)})) args thm + in + thm + end))\ + "Generalize a statement about variables to a statement about denoting terms." + +(* Note: rereplace-lem does not apply to the embedding *) + +AOT_theorem "cqt-basic:1": \\\\\ \{\,\} \ \\\\ \{\,\}\ + by (metis "\I" "\E"(2) "\I" "\I") + +AOT_theorem "cqt-basic:2": + \\\(\{\} \ \{\}) \ (\\(\{\} \ \{\}) & \\(\{\} \ \{\}))\ +proof (rule "\I"; rule "\I") + AOT_assume \\\(\{\} \ \{\})\ + AOT_hence \\{\} \ \{\}\ for \ using "\E"(2) by blast + AOT_hence \\{\} \ \{\}\ and \\{\} \ \{\}\ for \ + using "\E"(1,2) "\I" by blast+ + AOT_thus \\\(\{\} \ \{\}) & \\(\{\} \ \{\})\ + by (auto intro: "&I" "\I") +next + AOT_assume \\\(\{\} \ \{\}) & \\(\{\} \ \{\})\ + AOT_hence \\{\} \ \{\}\ and \\{\} \ \{\}\ for \ + using "\E"(2) "&E" by blast+ + AOT_hence \\{\} \ \{\}\ for \ + using "\I" by blast + AOT_thus \\\(\{\} \ \{\})\ by (auto intro: "\I") +qed + +AOT_theorem "cqt-basic:3": \\\(\{\} \ \{\}) \ (\\ \{\} \ \\ \{\})\ +proof(rule "\I") + AOT_assume \\\(\{\} \ \{\})\ + AOT_hence 1: \\{\} \ \{\}\ for \ using "\E"(2) by blast + { + AOT_assume \\\ \{\}\ + AOT_hence \\\ \{\}\ using 1 "\I" "\E"(4) "\E" by metis + } + moreover { + AOT_assume \\\ \{\}\ + AOT_hence \\\ \{\}\ using 1 "\I" "\E"(4) "\E" by metis + } + ultimately AOT_show \\\ \{\} \ \\ \{\}\ + using "\I" "\I" by auto +qed + +AOT_theorem "cqt-basic:4": \\\(\{\} & \{\}) \ (\\ \{\} & \\ \{\})\ +proof(rule "\I") + AOT_assume 0: \\\(\{\} & \{\})\ + AOT_have \\{\}\ and \\{\}\ for \ using "\E"(2) 0 "&E" by blast+ + AOT_thus \\\ \{\} & \\ \{\}\ + by (auto intro: "\I" "&I") +qed + +AOT_theorem "cqt-basic:5": \(\\\<^sub>1...\\\<^sub>n(\{\\<^sub>1...\\<^sub>n})) \ \{\\<^sub>1...\\<^sub>n}\ + using "cqt-orig:3" by blast + +AOT_theorem "cqt-basic:6": \\\\\ \{\} \ \\ \{\}\ + by (meson "\I" "\I" GEN "cqt-orig:1[const_var]") + +AOT_theorem "cqt-basic:7": \(\ \ \\ \{\}) \ \\(\ \ \{\})\ + by (metis "\I" "vdash-properties:6" "rule-ui:3" "\I" GEN) + +AOT_theorem "cqt-basic:8": \(\\ \{\} \ \\ \{\}) \ \\ (\{\} \ \{\})\ + by (simp add: "\I"(3) "\I" GEN "cqt-orig:1[const_var]") + +AOT_theorem "cqt-basic:9": + \(\\ (\{\} \ \{\}) & \\ (\{\} \ \{\})) \ \\(\{\} \ \{\})\ +proof - + { + AOT_assume \\\ (\{\} \ \{\})\ + moreover AOT_assume \\\ (\{\} \ \{\})\ + ultimately AOT_have \\{\} \ \{\}\ and \\{\} \ \{\}\ for \ + using "\E" by blast+ + AOT_hence \\{\} \ \{\}\ for \ by (metis "\E" "\I") + AOT_hence \\\(\{\} \ \{\})\ using "\I" by fast + } + thus ?thesis using "&I" "\I" "&E" by meson +qed + +AOT_theorem "cqt-basic:10": + \(\\(\{\} \ \{\}) & \\(\{\} \ \{\})) \ \\ (\{\} \ \{\})\ +proof(rule "\I"; rule "\I") + fix \ + AOT_assume \\\(\{\} \ \{\}) & \\(\{\} \ \{\})\ + AOT_hence \\{\} \ \{\}\ and \\{\} \ \{\}\ using "&E" "\E" by blast+ + AOT_thus \\{\} \ \{\}\ using "\I" "\E" by blast +qed + +AOT_theorem "cqt-basic:11": \\\(\{\} \ \{\}) \ \\ (\{\} \ \{\})\ +proof (rule "\I"; rule "\I") + AOT_assume 0: \\\(\{\} \ \{\})\ + { + fix \ + AOT_have \\{\} \ \{\}\ using 0 "\E" by blast + AOT_hence \\{\} \ \{\}\ using "\I" "\E" "\I" "\E" by metis + } + AOT_thus \\\(\{\} \ \{\})\ using "\I" by fast +next + AOT_assume 0: \\\(\{\} \ \{\})\ + { + fix \ + AOT_have \\{\} \ \{\}\ using 0 "\E" by blast + AOT_hence \\{\} \ \{\}\ using "\I" "\E" "\I" "\E" by metis + } + AOT_thus \\\(\{\} \ \{\})\ using "\I" by fast +qed + +AOT_theorem "cqt-basic:12": \\\ \{\} \ \\ (\{\} \ \{\})\ + by (simp add: "\E"(2) "\I" GEN) + +AOT_theorem "cqt-basic:13": \\\ \{\} \ \\ \{\}\ + using "\I" "\I" by blast + +AOT_theorem "cqt-basic:14": + \(\\\<^sub>1...\\\<^sub>n (\{\\<^sub>1...\\<^sub>n} \ \{\\<^sub>1...\\<^sub>n})) \ + ((\\\<^sub>1...\\\<^sub>n \{\\<^sub>1...\\<^sub>n}) \ (\\\<^sub>1...\\\<^sub>n \{\\<^sub>1...\\<^sub>n}))\ + using "cqt:3"[axiom_inst] by auto + +AOT_theorem "cqt-basic:15": + \(\\\<^sub>1...\\\<^sub>n (\ \ \{\\<^sub>1...\\<^sub>n})) \ (\ \ (\\\<^sub>1...\\\<^sub>n \{\\<^sub>1...\\<^sub>n}))\ + using "cqt-orig:2" by auto + +AOT_theorem "universal-cor": + assumes \for arbitrary \: \{\}\ + shows \\\ \{\}\ + using GEN assms . + +AOT_theorem "existential:1": + assumes \\{\}\ and \\\\ + shows \\\ \{\}\ +proof(rule "raa-cor:1") + AOT_assume \\\\ \{\}\ + AOT_hence \\\ \\{\}\ + using "\\<^sub>d\<^sub>fI" "conventions:4" RAA "&I" by blast + AOT_hence \\\{\}\ using assms(2) "\E"(1) "\E" by blast + AOT_thus \\{\} & \\{\}\ using assms(1) "&I" by blast +qed + +AOT_theorem "existential:2[const_var]": + assumes \\{\}\ + shows \\\ \{\}\ + using "existential:1" "cqt:2[const_var]"[axiom_inst] assms by blast + +AOT_theorem "existential:2[lambda]": + assumes \\{[\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]}\ and \INSTANCE_OF_CQT_2(\)\ + shows \\\ \{\}\ + using "existential:1" "cqt:2[lambda]"[axiom_inst] assms by blast +lemmas "\I" = "existential:1" "existential:2[const_var]" + "existential:2[lambda]" + +AOT_theorem "instantiation": + assumes \for arbitrary \: \{\} \<^bold>\ \\ and \\\ \{\}\ + shows \\\ + by (metis (no_types, lifting) "\\<^sub>d\<^sub>fE" GEN "raa-cor:3" "conventions:4" assms) +lemmas "\E" = "instantiation" + +AOT_theorem "cqt-further:1": \\\ \{\} \ \\ \{\}\ + using "\E"(4) "\I"(2) "\I" by metis + +AOT_theorem "cqt-further:2": \\\\ \{\} \ \\ \\{\}\ + using "\I" "\I"(2) "\I" RAA by metis + +AOT_theorem "cqt-further:3": \\\ \{\} \ \\\ \\{\}\ + using "\E"(4) "\E" "\I" RAA + by (metis "cqt-further:2" "\I" "modus-tollens:1") + +AOT_theorem "cqt-further:4": \\\\ \{\} \ \\ \\{\}\ + using "\I" "\I"(2)"\I" RAA by metis + +AOT_theorem "cqt-further:5": \\\ (\{\} & \{\}) \ (\\ \{\} & \\ \{\})\ + by (metis (no_types, lifting) "&E" "&I" "\E" "\I"(2) "\I") + +AOT_theorem "cqt-further:6": \\\ (\{\} \ \{\}) \ (\\ \{\} \ \\ \{\})\ + by (metis (mono_tags, lifting) "\E" "\I"(2) "\E"(3) "\I"(1, 2) "\I" RAA(2)) + +(* NOTE: vacuous in the embedding *) +AOT_theorem "cqt-further:7": \\\ \{\} \ \\ \{\}\ + by (simp add: "oth-class-taut:3:a") + +AOT_theorem "cqt-further:8": + \(\\ \{\} & \\ \{\}) \ \\ (\{\} \ \{\})\ + by (metis (mono_tags, lifting) "&E" "\I" "\E"(2) "\I" GEN) + +AOT_theorem "cqt-further:9": + \(\\\ \{\} & \\\ \{\}) \ \\ (\{\} \ \{\})\ + by (metis (mono_tags, lifting) "&E" "\I" "\I"(2) "\I" GEN "raa-cor:4") + +AOT_theorem "cqt-further:10": + \(\\ \{\} & \\\ \{\}) \ \\\ (\{\} \ \{\})\ +proof(rule "\I"; rule "raa-cor:2") + AOT_assume 0: \\\ \{\} & \\\ \{\}\ + then AOT_obtain \ where \\{\}\ using "\E" "&E"(1) by metis + moreover AOT_assume \\\ (\{\} \ \{\})\ + ultimately AOT_have \\{\}\ using "\E"(4) "\E"(1) by blast + AOT_hence \\\ \{\}\ using "\I" by blast + AOT_thus \\\ \{\} & \\\ \{\}\ using 0 "&E"(2) "&I" by blast +qed + +AOT_theorem "cqt-further:11": \\\\\ \{\,\} \ \\\\ \{\,\}\ + using "\I" "\I" "\I"(2) "\E" by metis + +subsection\Logical Existence, Identity, and Truth\ +text\\label{PLM: 9.7}\ + +AOT_theorem "log-prop-prop:1": \[\ \]\\ + using "cqt:2[lambda0]"[axiom_inst] by auto + +AOT_theorem "log-prop-prop:2": \\\\ + by (rule "\\<^sub>d\<^sub>fI"[OF "existence:3"]) "cqt:2[lambda]" + +AOT_theorem "exist-nec": \\\ \ \\\\ +proof - + AOT_have \\\ \\\\ + by (simp add: GEN RN "cqt:2[const_var]"[axiom_inst]) + AOT_thus \\\ \ \\\\ + using "cqt:1"[axiom_inst] "\E" by blast +qed + +(* TODO: replace this mechanism by a "proof by types" command *) +class AOT_Term_id = AOT_Term + + assumes "t=t-proper:1"[AOT]: \[v \ \ = \' \ \\]\ + and "t=t-proper:2"[AOT]: \[v \ \ = \' \ \'\]\ + +instance \ :: AOT_Term_id +proof + AOT_modally_strict { + AOT_show \\ = \' \ \\\ for \ \' + proof(rule "\I") + AOT_assume \\ = \'\ + AOT_hence \O!\ \ A!\\ + by (rule "\I"(3)[OF "\\<^sub>d\<^sub>fE"[OF "identity:1"]]) + (meson "\I" "\I"(1) "&E"(1))+ + AOT_thus \\\\ + by (rule "\E"(1)) + (metis "cqt:5:a"[axiom_inst] "\I" "\E" "&E"(2))+ + qed + } +next + AOT_modally_strict { + AOT_show \\ = \' \ \'\\ for \ \' + proof(rule "\I") + AOT_assume \\ = \'\ + AOT_hence \O!\' \ A!\'\ + by (rule "\I"(3)[OF "\\<^sub>d\<^sub>fE"[OF "identity:1"]]) + (meson "\I" "\I" "&E")+ + AOT_thus \\'\\ + by (rule "\E"(1)) + (metis "cqt:5:a"[axiom_inst] "\I" "\E" "&E"(2))+ + qed + } +qed + +instance rel :: (AOT_\s) AOT_Term_id +proof + AOT_modally_strict { + AOT_show \\ = \' \ \\\ for \ \' :: \<'a>\ + proof(rule "\I") + AOT_assume \\ = \'\ + AOT_thus \\\\ using "\\<^sub>d\<^sub>fE"[OF "identity:3"[of \ \']] "&E" by blast + qed + } +next + AOT_modally_strict { + AOT_show \\ = \' \ \'\\ for \ \' :: \<'a>\ + proof(rule "\I") + AOT_assume \\ = \'\ + AOT_thus \\'\\ using "\\<^sub>d\<^sub>fE"[OF "identity:3"[of \ \']] "&E" by blast + qed + } +qed + +instance \ :: AOT_Term_id +proof + AOT_modally_strict { + fix \ \ + AOT_show \\ = \ \ \\\ + proof(rule "\I") + AOT_assume \\ = \\ + AOT_thus \\\\ using "\\<^sub>d\<^sub>fE"[OF "identity:4"[of \ \]] "&E" by blast + qed + } +next + AOT_modally_strict { + fix \ \ + AOT_show \\ = \ \ \\\ + proof(rule "\I") + AOT_assume \\ = \\ + AOT_thus \\\\ using "\\<^sub>d\<^sub>fE"[OF "identity:4"[of \ \]] "&E" by blast + qed + } +qed + +instance prod :: (AOT_Term_id, AOT_Term_id) AOT_Term_id +proof + AOT_modally_strict { + fix \ \' :: \'a\'b\ + AOT_show \\ = \' \ \\\ + proof (induct \; induct \'; rule "\I") + fix \\<^sub>1 \\<^sub>1' :: 'a and \\<^sub>2 \\<^sub>2' :: 'b + AOT_assume \\(\\<^sub>1, \\<^sub>2)\ = \(\\<^sub>1', \\<^sub>2')\\ + AOT_hence \(\\<^sub>1 = \\<^sub>1') & (\\<^sub>2 = \\<^sub>2')\ by (metis "\\<^sub>d\<^sub>fE" tuple_identity_1) + AOT_hence \\\<^sub>1\\ and \\\<^sub>2\\ + using "t=t-proper:1" "&E" "vdash-properties:10" by blast+ + AOT_thus \\(\\<^sub>1, \\<^sub>2)\\\ by (metis "\\<^sub>d\<^sub>fI" "&I" tuple_denotes) + qed + } +next + AOT_modally_strict { + fix \ \' :: \'a\'b\ + AOT_show \\ = \' \ \'\\ + proof (induct \; induct \'; rule "\I") + fix \\<^sub>1 \\<^sub>1' :: 'a and \\<^sub>2 \\<^sub>2' :: 'b + AOT_assume \\(\\<^sub>1, \\<^sub>2)\ = \(\\<^sub>1', \\<^sub>2')\\ + AOT_hence \(\\<^sub>1 = \\<^sub>1') & (\\<^sub>2 = \\<^sub>2')\ by (metis "\\<^sub>d\<^sub>fE" tuple_identity_1) + AOT_hence \\\<^sub>1'\\ and \\\<^sub>2'\\ + using "t=t-proper:2" "&E" "vdash-properties:10" by blast+ + AOT_thus \\(\\<^sub>1', \\<^sub>2')\\\ by (metis "\\<^sub>d\<^sub>fI" "&I" tuple_denotes) + qed + } +qed + +(* This is the end of the "proof by types" and + makes the results available on new theorems *) +AOT_register_type_constraints + Term: \_::AOT_Term_id\ \_::AOT_Term_id\ +AOT_register_type_constraints + Individual: \\\ \_::{AOT_\s, AOT_Term_id}\ +AOT_register_type_constraints + Relation: \<_::{AOT_\s, AOT_Term_id}>\ + +AOT_theorem "id-rel-nec-equiv:1": + \\ = \' \ \\x\<^sub>1...\x\<^sub>n ([\]x\<^sub>1...x\<^sub>n \ [\']x\<^sub>1...x\<^sub>n)\ +proof(rule "\I") + AOT_assume assumption: \\ = \'\ + AOT_hence \\\\ and \\'\\ + using "t=t-proper:1" "t=t-proper:2" MP by blast+ + moreover AOT_have \\F\G (F = G \ ((\\x\<^sub>1...\x\<^sub>n ([F]x\<^sub>1...x\<^sub>n \ [F]x\<^sub>1...x\<^sub>n)) \ + \\x\<^sub>1...\x\<^sub>n ([F]x\<^sub>1...x\<^sub>n \ [G]x\<^sub>1...x\<^sub>n)))\ + apply (rule GEN)+ using "l-identity"[axiom_inst] by force + ultimately AOT_have \\ = \' \ ((\\x\<^sub>1...\x\<^sub>n ([\]x\<^sub>1...x\<^sub>n \ [\]x\<^sub>1...x\<^sub>n)) \ + \\x\<^sub>1...\x\<^sub>n ([\]x\<^sub>1...x\<^sub>n \ [\']x\<^sub>1...x\<^sub>n))\ + using "\E"(1) by blast + AOT_hence \(\\x\<^sub>1...\x\<^sub>n ([\]x\<^sub>1...x\<^sub>n \ [\]x\<^sub>1...x\<^sub>n)) \ + \\x\<^sub>1...\x\<^sub>n ([\]x\<^sub>1...x\<^sub>n \ [\']x\<^sub>1...x\<^sub>n)\ + using assumption "\E" by blast + moreover AOT_have \\\x\<^sub>1...\x\<^sub>n ([\]x\<^sub>1...x\<^sub>n \ [\]x\<^sub>1...x\<^sub>n)\ + by (simp add: RN "oth-class-taut:3:a" "universal-cor") + ultimately AOT_show \\\x\<^sub>1...\x\<^sub>n ([\]x\<^sub>1...x\<^sub>n \ [\']x\<^sub>1...x\<^sub>n)\ + using "\E" by blast +qed + +AOT_theorem "id-rel-nec-equiv:2": \\ = \ \ \(\ \ \)\ +proof(rule "\I") + AOT_assume assumption: \\ = \\ + AOT_hence \\\\ and \\\\ + using "t=t-proper:1" "t=t-proper:2" MP by blast+ + moreover AOT_have \\p\q (p = q \ ((\(p \ p) \ \(p \ q))))\ + apply (rule GEN)+ using "l-identity"[axiom_inst] by force + ultimately AOT_have \\ = \ \ (\(\ \ \) \ \(\ \ \))\ + using "\E"(1) by blast + AOT_hence \\(\ \ \) \ \(\ \ \)\ + using assumption "\E" by blast + moreover AOT_have \\(\ \ \)\ + by (simp add: RN "oth-class-taut:3:a" "universal-cor") + ultimately AOT_show \\(\ \ \)\ + using "\E" by blast +qed + +AOT_theorem "rule=E": + assumes \\{\}\ and \\ = \\ + shows \\{\}\ +proof - + AOT_have \\\\ and \\\\ + using assms(2) "t=t-proper:1" "t=t-proper:2" "\E" by blast+ + moreover AOT_have \\\\\(\ = \ \ (\{\} \ \{\}))\ + apply (rule GEN)+ using "l-identity"[axiom_inst] by blast + ultimately AOT_have \\ = \ \ (\{\} \ \{\})\ + using "\E"(1) by blast + AOT_thus \\{\}\ using assms "\E" by blast +qed + +AOT_theorem "propositions-lemma:1": \[\ \] = \\ +proof - + AOT_have \\\\ by (simp add: "log-prop-prop:2") + moreover AOT_have \\p [\ p] = p\ + using "lambda-predicates:3[zero]"[axiom_inst] "\I" by fast + ultimately AOT_show \[\ \] = \\ + using "\E" by blast +qed + +AOT_theorem "propositions-lemma:2": \[\ \] \ \\ +proof - + AOT_have \[\ \] \ [\ \]\ by (simp add: "oth-class-taut:3:a") + AOT_thus \[\ \] \ \\ using "propositions-lemma:1" "rule=E" by blast +qed + +text\propositions-lemma:3 through propositions-lemma:5 hold implicitly\ + +AOT_theorem "propositions-lemma:6": \(\ \ \) \ ([\ \] \ [\ \])\ + by (metis "\E"(1) "\E"(5) "Associativity of \" "propositions-lemma:2") + +text\dr-alphabetic-rules holds implicitly\ + +AOT_theorem "oa-exist:1": \O!\\ +proof - + AOT_have \[\x \[E!]x]\\ by "cqt:2[lambda]" + AOT_hence 1: \O! = [\x \[E!]x]\ + using "df-rules-terms[4]"[OF "oa:1", THEN "&E"(1)] "\E" by blast + AOT_show \O!\\ using "t=t-proper:1"[THEN "\E", OF 1] by simp +qed + +AOT_theorem "oa-exist:2": \A!\\ +proof - + AOT_have \[\x \\[E!]x]\\ by "cqt:2[lambda]" + AOT_hence 1: \A! = [\x \\[E!]x]\ + using "df-rules-terms[4]"[OF "oa:2", THEN "&E"(1)] "\E" by blast + AOT_show \A!\\ using "t=t-proper:1"[THEN "\E", OF 1] by simp +qed + +AOT_theorem "oa-exist:3": \O!x \ A!x\ +proof(rule "raa-cor:1") + AOT_assume \\(O!x \ A!x)\ + AOT_hence A: \\O!x\ and B: \\A!x\ + using "Disjunction Addition"(1) "modus-tollens:1" + "\I"(2) "raa-cor:5" by blast+ + AOT_have C: \O! = [\x \[E!]x]\ + by (rule "df-rules-terms[4]"[OF "oa:1", THEN "&E"(1), THEN "\E"]) "cqt:2" + AOT_have D: \A! = [\x \\[E!]x]\ + by (rule "df-rules-terms[4]"[OF "oa:2", THEN "&E"(1), THEN "\E"]) "cqt:2" + AOT_have E: \\[\x \[E!]x]x\ + using A C "rule=E" by fast + AOT_have F: \\[\x \\[E!]x]x\ + using B D "rule=E" by fast + AOT_have G: \[\x \[E!]x]x \ \[E!]x\ + by (rule "lambda-predicates:2"[axiom_inst, THEN "\E"]) "cqt:2" + AOT_have H: \[\x \\[E!]x]x \ \\[E!]x\ + by (rule "lambda-predicates:2"[axiom_inst, THEN "\E"]) "cqt:2" + AOT_show \\\[E!]x & \\\[E!]x\ using G E "\E" H F "\E" "&I" by metis +qed + +AOT_theorem "p-identity-thm2:1": \F = G \ \\x(x[F] \ x[G])\ +proof - + AOT_have \F = G \ F\ & G\ & \\x(x[F] \ x[G])\ + using "identity:2" "df-rules-formulas[3]" "df-rules-formulas[4]" + "\E" "&E" "\I" "\I" by blast + moreover AOT_have \F\\ and \G\\ + by (auto simp: "cqt:2[const_var]"[axiom_inst]) + ultimately AOT_show \F = G \ \\x(x[F] \ x[G])\ + using "\S"(1) "&I" by blast +qed + +AOT_theorem "p-identity-thm2:2[2]": + \F = G \ \y\<^sub>1([\x [F]xy\<^sub>1] = [\x [G]xy\<^sub>1] & [\x [F]y\<^sub>1x] = [\x [G]y\<^sub>1x])\ +proof - + AOT_have \F = G \ F\ & G\ & + \y\<^sub>1([\x [F]xy\<^sub>1] = [\x [G]xy\<^sub>1] & [\x [F]y\<^sub>1x] = [\x [G]y\<^sub>1x])\ + using "identity:3[2]" "df-rules-formulas[3]" "df-rules-formulas[4]" + "\E" "&E" "\I" "\I" by blast + moreover AOT_have \F\\ and \G\\ + by (auto simp: "cqt:2[const_var]"[axiom_inst]) + ultimately show ?thesis + using "\S"(1) "&I" by blast +qed + +AOT_theorem "p-identity-thm2:2[3]": + \F = G \ \y\<^sub>1\y\<^sub>2([\x [F]xy\<^sub>1y\<^sub>2] = [\x [G]xy\<^sub>1y\<^sub>2] & + [\x [F]y\<^sub>1xy\<^sub>2] = [\x [G]y\<^sub>1xy\<^sub>2] & + [\x [F]y\<^sub>1y\<^sub>2x] = [\x [G]y\<^sub>1y\<^sub>2x])\ +proof - + AOT_have \F = G \ F\ & G\ & \y\<^sub>1\y\<^sub>2([\x [F]xy\<^sub>1y\<^sub>2] = [\x [G]xy\<^sub>1y\<^sub>2] & + [\x [F]y\<^sub>1xy\<^sub>2] = [\x [G]y\<^sub>1xy\<^sub>2] & + [\x [F]y\<^sub>1y\<^sub>2x] = [\x [G]y\<^sub>1y\<^sub>2x])\ + using "identity:3[3]" "df-rules-formulas[3]" "df-rules-formulas[4]" + "\E" "&E" "\I" "\I" by blast + moreover AOT_have \F\\ and \G\\ + by (auto simp: "cqt:2[const_var]"[axiom_inst]) + ultimately show ?thesis + using "\S"(1) "&I" by blast +qed + +AOT_theorem "p-identity-thm2:2[4]": + \F = G \ \y\<^sub>1\y\<^sub>2\y\<^sub>3([\x [F]xy\<^sub>1y\<^sub>2y\<^sub>3] = [\x [G]xy\<^sub>1y\<^sub>2y\<^sub>3] & + [\x [F]y\<^sub>1xy\<^sub>2y\<^sub>3] = [\x [G]y\<^sub>1xy\<^sub>2y\<^sub>3] & + [\x [F]y\<^sub>1y\<^sub>2xy\<^sub>3] = [\x [G]y\<^sub>1y\<^sub>2xy\<^sub>3] & + [\x [F]y\<^sub>1y\<^sub>2y\<^sub>3x] = [\x [G]y\<^sub>1y\<^sub>2y\<^sub>3x])\ +proof - + AOT_have \F = G \ F\ & G\ & \y\<^sub>1\y\<^sub>2\y\<^sub>3([\x [F]xy\<^sub>1y\<^sub>2y\<^sub>3] = [\x [G]xy\<^sub>1y\<^sub>2y\<^sub>3] & + [\x [F]y\<^sub>1xy\<^sub>2y\<^sub>3] = [\x [G]y\<^sub>1xy\<^sub>2y\<^sub>3] & + [\x [F]y\<^sub>1y\<^sub>2xy\<^sub>3] = [\x [G]y\<^sub>1y\<^sub>2xy\<^sub>3] & + [\x [F]y\<^sub>1y\<^sub>2y\<^sub>3x] = [\x [G]y\<^sub>1y\<^sub>2y\<^sub>3x])\ + using "identity:3[4]" "df-rules-formulas[3]" "df-rules-formulas[4]" + "\E" "&E" "\I" "\I" by blast + moreover AOT_have \F\\ and \G\\ + by (auto simp: "cqt:2[const_var]"[axiom_inst]) + ultimately show ?thesis + using "\S"(1) "&I" by blast +qed + +AOT_theorem "p-identity-thm2:2": + \F = G \ \x\<^sub>1...\x\<^sub>n \AOT_sem_proj_id x\<^sub>1x\<^sub>n (\ \ . \[F]\\) (\ \ . \[G]\\)\\ +proof - + AOT_have \F = G \ F\ & G\ & + \x\<^sub>1...\x\<^sub>n \AOT_sem_proj_id x\<^sub>1x\<^sub>n (\ \ . \[F]\\) (\ \ . \[G]\\)\\ + using "identity:3" "df-rules-formulas[3]" "df-rules-formulas[4]" + "\E" "&E" "\I" "\I" by blast + moreover AOT_have \F\\ and \G\\ + by (auto simp: "cqt:2[const_var]"[axiom_inst]) + ultimately show ?thesis + using "\S"(1) "&I" by blast +qed + +AOT_theorem "p-identity-thm2:3": + \p = q \ [\x p] = [\x q]\ +proof - + AOT_have \p = q \ p\ & q\ & [\x p] = [\x q]\ + using "identity:4" "df-rules-formulas[3]" "df-rules-formulas[4]" + "\E" "&E" "\I" "\I" by blast + moreover AOT_have \p\\ and \q\\ + by (auto simp: "cqt:2[const_var]"[axiom_inst]) + ultimately show ?thesis + using "\S"(1) "&I" by blast +qed + +class AOT_Term_id_2 = AOT_Term_id + assumes "id-eq:1": \[v \ \ = \]\ + +instance \ :: AOT_Term_id_2 +proof + AOT_modally_strict { + fix x + { + AOT_assume \O!x\ + moreover AOT_have \\\F([F]x \ [F]x)\ + using RN GEN "oth-class-taut:3:a" by fast + ultimately AOT_have \O!x & O!x & \\F([F]x \ [F]x)\ using "&I" by simp + } + moreover { + AOT_assume \A!x\ + moreover AOT_have \\\F(x[F] \ x[F])\ + using RN GEN "oth-class-taut:3:a" by fast + ultimately AOT_have \A!x & A!x & \\F(x[F] \ x[F])\ using "&I" by simp + } + ultimately AOT_have \(O!x & O!x & \\F([F]x \ [F]x)) \ + (A!x & A!x & \\F(x[F] \ x[F]))\ + using "oa-exist:3" "\I"(1) "\I"(2) "\E"(3) "raa-cor:1" by blast + AOT_thus \x = x\ + using "identity:1"[THEN "df-rules-formulas[4]"] "\E" by blast + } +qed + +instance rel :: ("{AOT_\s,AOT_Term_id_2}") AOT_Term_id_2 +proof + AOT_modally_strict { + fix F :: "<'a> AOT_var" + AOT_have 0: \[\x\<^sub>1...x\<^sub>n [F]x\<^sub>1...x\<^sub>n] = F\ + by (simp add: "lambda-predicates:3"[axiom_inst]) + AOT_have \[\x\<^sub>1...x\<^sub>n [F]x\<^sub>1...x\<^sub>n]\\ + by "cqt:2[lambda]" + AOT_hence \[\x\<^sub>1...x\<^sub>n [F]x\<^sub>1...x\<^sub>n] = [\x\<^sub>1...x\<^sub>n [F]x\<^sub>1...x\<^sub>n]\ + using "lambda-predicates:1"[axiom_inst] "\E" by blast + AOT_show \F = F\ using "rule=E" 0 by force + } +qed + +instance \ :: AOT_Term_id_2 +proof + AOT_modally_strict { + fix p + AOT_have 0: \[\ p] = p\ + by (simp add: "lambda-predicates:3[zero]"[axiom_inst]) + AOT_have \[\ p]\\ + by (rule "cqt:2[lambda0]"[axiom_inst]) + AOT_hence \[\ p] = [\ p]\ + using "lambda-predicates:1[zero]"[axiom_inst] "\E" by blast + AOT_show \p = p\ using "rule=E" 0 by force + } +qed + +instance prod :: (AOT_Term_id_2, AOT_Term_id_2) AOT_Term_id_2 +proof + AOT_modally_strict { + fix \ :: \('a\'b) AOT_var\ + AOT_show \\ = \\ + proof (induct) + AOT_show \\ = \\ if \\\\ for \ :: \'a\'b\ + using that + proof (induct \) + fix \\<^sub>1 :: 'a and \\<^sub>2 :: 'b + AOT_assume \\(\\<^sub>1,\\<^sub>2)\\\ + AOT_hence \\\<^sub>1\\ and \\\<^sub>2\\ + using "\\<^sub>d\<^sub>fE" "&E" tuple_denotes by blast+ + AOT_hence \\\<^sub>1 = \\<^sub>1\ and \\\<^sub>2 = \\<^sub>2\ + using "id-eq:1"[unvarify \] by blast+ + AOT_thus \\(\\<^sub>1, \\<^sub>2)\ = \(\\<^sub>1, \\<^sub>2)\\ + by (metis "\\<^sub>d\<^sub>fI" "&I" tuple_identity_1) + qed + qed + } +qed + +AOT_register_type_constraints + Term: \_::AOT_Term_id_2\ \_::AOT_Term_id_2\ +AOT_register_type_constraints + Individual: \\\ \_::{AOT_\s, AOT_Term_id_2}\ +AOT_register_type_constraints + Relation: \<_::{AOT_\s, AOT_Term_id_2}>\ + +AOT_theorem "id-eq:2": \\ = \ \ \ = \\ + by (meson "rule=E" "deduction-theorem") + +AOT_theorem "id-eq:3": \\ = \ & \ = \ \ \ = \\ + using "rule=E" "\I" "&E" by blast + +AOT_theorem "id-eq:4": \\ = \ \ \\ (\ = \ \ \ = \)\ +proof (rule "\I"; rule "\I") + AOT_assume 0: \\ = \\ + AOT_hence 1: \\ = \\ using "id-eq:2" "\E" by blast + AOT_show \\\ (\ = \ \ \ = \)\ + by (rule GEN) (metis "\I" "\I" 0 "1" "rule=E") +next + AOT_assume \\\ (\ = \ \ \ = \)\ + AOT_hence \\ = \ \ \ = \\ using "\E"(2) by blast + AOT_hence \\ = \ \ \ = \\ using "\E"(1) "\I" by blast + AOT_hence \\ = \\ using "id-eq:1" "\E" by blast + AOT_thus \\ = \\ using "id-eq:2" "\E" by blast +qed + +AOT_theorem "rule=I:1": + assumes \\\\ + shows \\ = \\ +proof - + AOT_have \\\ (\ = \)\ + by (rule GEN) (metis "id-eq:1") + AOT_thus \\ = \\ using assms "\E" by blast +qed + +AOT_theorem "rule=I:2[const_var]": "\ = \" + using "id-eq:1". + +AOT_theorem "rule=I:2[lambda]": + assumes \INSTANCE_OF_CQT_2(\)\ + shows "[\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}] = [\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]" +proof - + AOT_have \\\ (\ = \)\ + by (rule GEN) (metis "id-eq:1") + moreover AOT_have \[\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\\ + using assms by (rule "cqt:2[lambda]"[axiom_inst]) + ultimately AOT_show \[\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}] = [\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\ + using assms "\E" by blast +qed + +lemmas "=I" = "rule=I:1" "rule=I:2[const_var]" "rule=I:2[lambda]" + +AOT_theorem "rule-id-df:1": + assumes \\{\\<^sub>1...\\<^sub>n} =\<^sub>d\<^sub>f \{\\<^sub>1...\\<^sub>n}\ and \\{\\<^sub>1...\\<^sub>n}\\ + shows \\{\\<^sub>1...\\<^sub>n} = \{\\<^sub>1...\\<^sub>n}\ +proof - + AOT_have \\{\\<^sub>1...\\<^sub>n}\ \ \{\\<^sub>1...\\<^sub>n} = \{\\<^sub>1...\\<^sub>n}\ + using "df-rules-terms[3]" assms(1) "&E" by blast + AOT_thus \\{\\<^sub>1...\\<^sub>n} = \{\\<^sub>1...\\<^sub>n}\ + using assms(2) "\E" by blast +qed + +AOT_theorem "rule-id-df:1[zero]": + assumes \\ =\<^sub>d\<^sub>f \\ and \\\\ + shows \\ = \\ +proof - + AOT_have \\\ \ \ = \\ + using "df-rules-terms[4]" assms(1) "&E" by blast + AOT_thus \\ = \\ + using assms(2) "\E" by blast +qed + +AOT_theorem "rule-id-df:2:a": + assumes \\{\\<^sub>1...\\<^sub>n} =\<^sub>d\<^sub>f \{\\<^sub>1...\\<^sub>n}\ and \\{\\<^sub>1...\\<^sub>n}\\ and \\{\{\\<^sub>1...\\<^sub>n}}\ + shows \\{\{\\<^sub>1...\\<^sub>n}}\ +proof - + AOT_have \\{\\<^sub>1...\\<^sub>n} = \{\\<^sub>1...\\<^sub>n}\ using "rule-id-df:1" assms(1,2) by blast + AOT_thus \\{\{\\<^sub>1...\\<^sub>n}}\ using assms(3) "rule=E" by blast +qed + +AOT_theorem "rule-id-df:2:a[2]": + assumes \\{\(\\<^sub>1,\\<^sub>2)\} =\<^sub>d\<^sub>f \{\(\\<^sub>1,\\<^sub>2)\}\ + and \\{\(\\<^sub>1,\\<^sub>2)\}\\ + and \\{\{\(\\<^sub>1,\\<^sub>2)\}}\ + shows \\{\{\(\\<^sub>1::'a::AOT_Term_id_2,\\<^sub>2::'b::AOT_Term_id_2)\}}\ +proof - + AOT_have \\{\(\\<^sub>1,\\<^sub>2)\} = \{\(\\<^sub>1,\\<^sub>2)\}\ + using "rule-id-df:1" assms(1,2) by auto + AOT_thus \\{\{\(\\<^sub>1,\\<^sub>2)\}}\ using assms(3) "rule=E" by blast +qed + +AOT_theorem "rule-id-df:2:a[zero]": + assumes \\ =\<^sub>d\<^sub>f \\ and \\\\ and \\{\}\ + shows \\{\}\ +proof - + AOT_have \\ = \\ using "rule-id-df:1[zero]" assms(1,2) by blast + AOT_thus \\{\}\ using assms(3) "rule=E" by blast +qed + +lemmas "=\<^sub>d\<^sub>fE" = "rule-id-df:2:a" "rule-id-df:2:a[zero]" + +AOT_theorem "rule-id-df:2:b": + assumes \\{\\<^sub>1...\\<^sub>n} =\<^sub>d\<^sub>f \{\\<^sub>1...\\<^sub>n}\ and \\{\\<^sub>1...\\<^sub>n}\\ and \\{\{\\<^sub>1...\\<^sub>n}}\ + shows \\{\{\\<^sub>1...\\<^sub>n}}\ +proof - + AOT_have \\{\\<^sub>1...\\<^sub>n} = \{\\<^sub>1...\\<^sub>n}\ + using "rule-id-df:1" assms(1,2) by blast + AOT_hence \\{\\<^sub>1...\\<^sub>n} = \{\\<^sub>1...\\<^sub>n}\ + using "rule=E" "=I"(1) "t=t-proper:1" "\E" by fast + AOT_thus \\{\{\\<^sub>1...\\<^sub>n}}\ using assms(3) "rule=E" by blast +qed + +AOT_theorem "rule-id-df:2:b[2]": + assumes \\{\(\\<^sub>1,\\<^sub>2)\} =\<^sub>d\<^sub>f \{\(\\<^sub>1,\\<^sub>2)\}\ + and \\{\(\\<^sub>1,\\<^sub>2)\}\\ + and \\{\{\(\\<^sub>1,\\<^sub>2)\}}\ + shows \\{\{\(\\<^sub>1::'a::AOT_Term_id_2,\\<^sub>2::'b::AOT_Term_id_2)\}}\ +proof - + AOT_have \\{\(\\<^sub>1,\\<^sub>2)\} = \{\(\\<^sub>1,\\<^sub>2)\}\ + using "=I"(1) "rule-id-df:2:a[2]" RAA(1) assms(1,2) "\I" by metis + AOT_hence \\{\(\\<^sub>1,\\<^sub>2)\} = \{\(\\<^sub>1,\\<^sub>2)\}\ + using "rule=E" "=I"(1) "t=t-proper:1" "\E" by fast + AOT_thus \\{\{\(\\<^sub>1,\\<^sub>2)\}}\ using assms(3) "rule=E" by blast +qed + +AOT_theorem "rule-id-df:2:b[zero]": + assumes \\ =\<^sub>d\<^sub>f \\ and \\\\ and \\{\}\ + shows \\{\}\ +proof - + AOT_have \\ = \\ using "rule-id-df:1[zero]" assms(1,2) by blast + AOT_hence \\ = \\ + using "rule=E" "=I"(1) "t=t-proper:1" "\E" by fast + AOT_thus \\{\}\ using assms(3) "rule=E" by blast +qed + +lemmas "=\<^sub>d\<^sub>fI" = "rule-id-df:2:b" "rule-id-df:2:b[zero]" + +AOT_theorem "free-thms:1": \\\ \ \\ (\ = \)\ + by (metis "\E" "rule=I:1" "t=t-proper:2" "\I" "\I"(1) "\I" "\E") + +AOT_theorem "free-thms:2": \\\ \{\} \ (\\ (\ = \) \ \{\})\ + by (metis "\E" "rule=E" "cqt:2[const_var]"[axiom_inst] "\I" "\E"(1)) + +AOT_theorem "free-thms:3[const_var]": \\\ (\ = \)\ + by (meson "\I"(2) "id-eq:1") + +AOT_theorem "free-thms:3[lambda]": + assumes \INSTANCE_OF_CQT_2(\)\ + shows \\\ (\ = [\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}])\ + by (meson "=I"(3) assms "cqt:2[lambda]"[axiom_inst] "existential:1") + +AOT_theorem "free-thms:4[rel]": + \([\]\\<^sub>1...\\<^sub>n \ \\<^sub>1...\\<^sub>n[\]) \ \\ (\ = \)\ + by (metis "rule=I:1" "&E"(1) "\E"(1) "cqt:5:a"[axiom_inst] + "cqt:5:b"[axiom_inst] "\I" "\I"(1)) + +AOT_theorem "free-thms:4[vars]": + \([\]\\<^sub>1...\\<^sub>n \ \\<^sub>1...\\<^sub>n[\]) \ \\\<^sub>1...\\\<^sub>n (\\<^sub>1...\\<^sub>n = \\<^sub>1...\\<^sub>n)\ + by (metis "rule=I:1" "&E"(2) "\E"(1) "cqt:5:a"[axiom_inst] + "cqt:5:b"[axiom_inst] "\I" "\I"(1)) + +AOT_theorem "free-thms:4[1,rel]": + \([\]\ \ \[\]) \ \\ (\ = \)\ + by (metis "rule=I:1" "&E"(1) "\E"(1) "cqt:5:a"[axiom_inst] + "cqt:5:b"[axiom_inst] "\I" "\I"(1)) +AOT_theorem "free-thms:4[1,1]": + \([\]\ \ \[\]) \ \\ (\ = \)\ + by (metis "rule=I:1" "&E"(2) "\E"(1) "cqt:5:a"[axiom_inst] + "cqt:5:b"[axiom_inst] "\I" "\I"(1)) + +AOT_theorem "free-thms:4[2,rel]": + \([\]\\<^sub>1\\<^sub>2 \ \\<^sub>1\\<^sub>2[\]) \ \\ (\ = \)\ + by (metis "rule=I:1" "&E"(1) "\E"(1) "cqt:5:a[2]"[axiom_inst] + "cqt:5:b[2]"[axiom_inst] "\I" "\I"(1)) +AOT_theorem "free-thms:4[2,1]": + \([\]\\<^sub>1\\<^sub>2 \ \\<^sub>1\\<^sub>2[\]) \ \\ (\ = \\<^sub>1)\ + by (metis "rule=I:1" "&E" "\E"(1) "cqt:5:a[2]"[axiom_inst] + "cqt:5:b[2]"[axiom_inst] "\I" "\I"(1)) +AOT_theorem "free-thms:4[2,2]": + \([\]\\<^sub>1\\<^sub>2 \ \\<^sub>1\\<^sub>2[\]) \ \\ (\ = \\<^sub>2)\ + by (metis "rule=I:1" "&E"(2) "\E"(1) "cqt:5:a[2]"[axiom_inst] + "cqt:5:b[2]"[axiom_inst] "\I" "\I"(1)) +AOT_theorem "free-thms:4[3,rel]": + \([\]\\<^sub>1\\<^sub>2\\<^sub>3 \ \\<^sub>1\\<^sub>2\\<^sub>3[\]) \ \\ (\ = \)\ + by (metis "rule=I:1" "&E"(1) "\E"(1) "cqt:5:a[3]"[axiom_inst] + "cqt:5:b[3]"[axiom_inst] "\I" "\I"(1)) +AOT_theorem "free-thms:4[3,1]": + \([\]\\<^sub>1\\<^sub>2\\<^sub>3 \ \\<^sub>1\\<^sub>2\\<^sub>3[\]) \ \\ (\ = \\<^sub>1)\ + by (metis "rule=I:1" "&E" "\E"(1) "cqt:5:a[3]"[axiom_inst] + "cqt:5:b[3]"[axiom_inst] "\I" "\I"(1)) +AOT_theorem "free-thms:4[3,2]": + \([\]\\<^sub>1\\<^sub>2\\<^sub>3 \ \\<^sub>1\\<^sub>2\\<^sub>3[\]) \ \\ (\ = \\<^sub>2)\ + by (metis "rule=I:1" "&E" "\E"(1) "cqt:5:a[3]"[axiom_inst] + "cqt:5:b[3]"[axiom_inst] "\I" "\I"(1)) +AOT_theorem "free-thms:4[3,3]": + \([\]\\<^sub>1\\<^sub>2\\<^sub>3 \ \\<^sub>1\\<^sub>2\\<^sub>3[\]) \ \\ (\ = \\<^sub>3)\ + by (metis "rule=I:1" "&E"(2) "\E"(1) "cqt:5:a[3]"[axiom_inst] + "cqt:5:b[3]"[axiom_inst] "\I" "\I"(1)) +AOT_theorem "free-thms:4[4,rel]": + \([\]\\<^sub>1\\<^sub>2\\<^sub>3\\<^sub>4 \ \\<^sub>1\\<^sub>2\\<^sub>3\\<^sub>4[\]) \ \\ (\ = \)\ + by (metis "rule=I:1" "&E"(1) "\E"(1) "cqt:5:a[4]"[axiom_inst] + "cqt:5:b[4]"[axiom_inst] "\I" "\I"(1)) +AOT_theorem "free-thms:4[4,1]": + \([\]\\<^sub>1\\<^sub>2\\<^sub>3\\<^sub>4 \ \\<^sub>1\\<^sub>2\\<^sub>3\\<^sub>4[\]) \ \\ (\ = \\<^sub>1)\ + by (metis "rule=I:1" "&E" "\E"(1) "cqt:5:a[4]"[axiom_inst] + "cqt:5:b[4]"[axiom_inst] "\I" "\I"(1)) +AOT_theorem "free-thms:4[4,2]": + \([\]\\<^sub>1\\<^sub>2\\<^sub>3\\<^sub>4 \ \\<^sub>1\\<^sub>2\\<^sub>3\\<^sub>4[\]) \ \\ (\ = \\<^sub>2)\ + by (metis "rule=I:1" "&E" "\E"(1) "cqt:5:a[4]"[axiom_inst] + "cqt:5:b[4]"[axiom_inst] "\I" "\I"(1)) +AOT_theorem "free-thms:4[4,3]": + \([\]\\<^sub>1\\<^sub>2\\<^sub>3\\<^sub>4 \ \\<^sub>1\\<^sub>2\\<^sub>3\\<^sub>4[\]) \ \\ (\ = \\<^sub>3)\ + by (metis "rule=I:1" "&E" "\E"(1) "cqt:5:a[4]"[axiom_inst] + "cqt:5:b[4]"[axiom_inst] "\I" "\I"(1)) +AOT_theorem "free-thms:4[4,4]": + \([\]\\<^sub>1\\<^sub>2\\<^sub>3\\<^sub>4 \ \\<^sub>1\\<^sub>2\\<^sub>3\\<^sub>4[\]) \ \\ (\ = \\<^sub>4)\ + by (metis "rule=I:1" "&E"(2) "\E"(1) "cqt:5:a[4]"[axiom_inst] + "cqt:5:b[4]"[axiom_inst] "\I" "\I"(1)) + +AOT_theorem "ex:1:a": \\\ \\\ + by (rule GEN) (fact "cqt:2[const_var]"[axiom_inst]) +AOT_theorem "ex:1:b": \\\\\(\ = \)\ + by (rule GEN) (fact "free-thms:3[const_var]") + +AOT_theorem "ex:2:a": \\\\\ + by (rule RN) (fact "cqt:2[const_var]"[axiom_inst]) +AOT_theorem "ex:2:b": \\\\(\ = \)\ + by (rule RN) (fact "free-thms:3[const_var]") + +AOT_theorem "ex:3:a": \\\\ \\\ + by (rule RN) (fact "ex:1:a") +AOT_theorem "ex:3:b": \\\\\\(\ = \)\ + by (rule RN) (fact "ex:1:b") + +AOT_theorem "ex:4:a": \\\ \\\\ + by (rule GEN; rule RN) (fact "cqt:2[const_var]"[axiom_inst]) +AOT_theorem "ex:4:b": \\\\\\(\ = \)\ + by (rule GEN; rule RN) (fact "free-thms:3[const_var]") + +AOT_theorem "ex:5:a": \\\\ \\\\ + by (rule RN) (simp add: "ex:4:a") +AOT_theorem "ex:5:b": \\\\\\\(\ = \)\ + by (rule RN) (simp add: "ex:4:b") + +AOT_theorem "all-self=:1": \\\\(\ = \)\ + by (rule RN; rule GEN) (fact "id-eq:1") +AOT_theorem "all-self=:2": \\\\(\ = \)\ + by (rule GEN; rule RN) (fact "id-eq:1") + +AOT_theorem "id-nec:1": \\ = \ \ \(\ = \)\ +proof(rule "\I") + AOT_assume \\ = \\ + moreover AOT_have \\(\ = \)\ + by (rule RN) (fact "id-eq:1") + ultimately AOT_show \\(\ = \)\ using "rule=E" by fast +qed + +AOT_theorem "id-nec:2": \\ = \ \ \(\ = \)\ +proof(rule "\I") + AOT_assume asm: \\ = \\ + moreover AOT_have \\\\ + using calculation "t=t-proper:1" "\E" by blast + moreover AOT_have \\(\ = \)\ + using calculation "all-self=:2" "\E"(1) by blast + ultimately AOT_show \\(\ = \)\ using "rule=E" by fast +qed + +AOT_theorem "term-out:1": \\{\} \ \\ (\ = \ & \{\})\ +proof (rule "\I"; rule "\I") + AOT_assume asm: \\{\}\ + AOT_show \\\ (\ = \ & \{\})\ + by (rule "\I"(2)[where \=\]; rule "&I") + (auto simp: "id-eq:1" asm) +next + AOT_assume 0: \\\ (\ = \ & \{\})\ + AOT_obtain \ where \\ = \ & \{\}\ + using "\E"[rotated, OF 0] by blast + AOT_thus \\{\}\ using "&E" "rule=E" by blast +qed + +AOT_theorem "term-out:2": \\\ \ (\{\} \ \\(\ = \ & \{\}))\ +proof(rule "\I") + AOT_assume \\\\ + moreover AOT_have \\\ (\{\} \ \\ (\ = \ & \{\}))\ + by (rule GEN) (fact "term-out:1") + ultimately AOT_show \\{\} \ \\(\ = \ & \{\})\ + using "\E" by blast +qed + +AOT_theorem "term-out:3": + \(\{\} & \\(\{\} \ \ = \)) \ \\(\{\} \ \ = \)\ + apply (rule "\I"; rule "\I") + apply (frule "&E"(1)) + apply (drule "&E"(2)) + apply (rule GEN; rule "\I"; rule "\I") + using "rule-ui:2[const_var]" "vdash-properties:5" + apply blast + apply (meson "rule=E" "id-eq:1") + apply (rule "&I") + using "id-eq:1" "\E"(2) "rule-ui:3" + apply blast + apply (rule GEN; rule "\I") + using "\E"(1) "rule-ui:2[const_var]" + by blast + +(* Note: generalized alphabetic variant of the last theorem. *) +AOT_theorem "term-out:4": + \(\{\} & \\(\{\} \ \ = \)) \ \\(\{\} \ \ = \)\ + using "term-out:3" . + +(* TODO: Provide a nicer mechanism for introducing custom binders. *) +AOT_define AOT_exists_unique :: \\ \ \ \ \\ "uniqueness:1": + \\AOT_exists_unique \\ \\<^sub>d\<^sub>f \\ (\{\} & \\ (\{\} \ \ = \))\ +syntax (input) "_AOT_exists_unique" :: \\ \ \ \ \\ ("\!_ _" [1,40]) +syntax (output) "_AOT_exists_unique" :: \\ \ \ \ \\ ("\!_'(_')" [1,40]) +AOT_syntax_print_translations + "_AOT_exists_unique \ \" <= "CONST AOT_exists_unique (_abs \ \)" +syntax + "_AOT_exists_unique_ellipse" :: \id_position \ id_position \ \ \ \\ + (\\!_...\!_ _\ [1,40]) +parse_ast_translation\ +[(\<^syntax_const>\_AOT_exists_unique_ellipse\, + fn ctx => fn [a,b,c] => Ast.mk_appl (Ast.Constant "AOT_exists_unique") + [parseEllipseList "_AOT_vars" ctx [a,b],c]), + (\<^syntax_const>\_AOT_exists_unique\, + AOT_restricted_binder + \<^const_name>\AOT_exists_unique\ + \<^const_syntax>\AOT_conj\)]\ +print_translation\AOT_syntax_print_translations [ + AOT_preserve_binder_abs_tr' + \<^const_syntax>\AOT_exists_unique\ + \<^syntax_const>\_AOT_exists_unique\ + (\<^syntax_const>\_AOT_exists_unique_ellipse\, true) + \<^const_name>\AOT_conj\, + AOT_binder_trans + @{theory} + @{binding "AOT_exists_unique_binder"} + \<^syntax_const>\_AOT_exists_unique\ +]\ + + +context AOT_meta_syntax +begin +notation AOT_exists_unique (binder "\<^bold>\\<^bold>!" 20) +end +context AOT_no_meta_syntax +begin +no_notation AOT_exists_unique (binder "\<^bold>\\<^bold>!" 20) +end + +AOT_theorem "uniqueness:2": \\!\ \{\} \ \\\\(\{\} \ \ = \)\ +proof(rule "\I"; rule "\I") + AOT_assume \\!\ \{\}\ + AOT_hence \\\ (\{\} & \\ (\{\} \ \ = \))\ + using "uniqueness:1" "\\<^sub>d\<^sub>fE" by blast + then AOT_obtain \ where \\{\} & \\ (\{\} \ \ = \)\ + using "instantiation"[rotated] by blast + AOT_hence \\\(\{\} \ \ = \)\ + using "term-out:3" "\E" by blast + AOT_thus \\\\\(\{\} \ \ = \)\ + using "\I" by fast +next + AOT_assume \\\\\(\{\} \ \ = \)\ + then AOT_obtain \ where \\\ (\{\} \ \ = \)\ + using "instantiation"[rotated] by blast + AOT_hence \\{\} & \\ (\{\} \ \ = \)\ + using "term-out:3" "\E" by blast + AOT_hence \\\ (\{\} & \\ (\{\} \ \ = \))\ + using "\I" by fast + AOT_thus \\!\ \{\}\ + using "uniqueness:1" "\\<^sub>d\<^sub>fI" by blast +qed + +AOT_theorem "uni-most": \\!\ \{\} \ \\\\((\{\} & \{\}) \ \ = \)\ +proof(rule "\I"; rule GEN; rule GEN; rule "\I") + fix \ \ + AOT_assume \\!\ \{\}\ + AOT_hence \\\\\(\{\} \ \ = \)\ + using "uniqueness:2" "\E" by blast + then AOT_obtain \ where \\\(\{\} \ \ = \)\ + using "instantiation"[rotated] by blast + moreover AOT_assume \\{\} & \{\}\ + ultimately AOT_have \\ = \\ and \\ = \\ + using "\E"(2) "&E" "\E"(1,2) by blast+ + AOT_thus \\ = \\ + by (metis "rule=E" "id-eq:2" "\E") +qed + +AOT_theorem "nec-exist-!": \\\(\{\} \ \\{\}) \ (\!\ \{\} \ \!\ \\{\})\ +proof (rule "\I"; rule "\I") + AOT_assume a: \\\(\{\} \ \\{\})\ + AOT_assume \\!\ \{\}\ + AOT_hence \\\ (\{\} & \\ (\{\} \ \ = \))\ + using "uniqueness:1" "\\<^sub>d\<^sub>fE" by blast + then AOT_obtain \ where \: \\{\} & \\ (\{\} \ \ = \)\ + using "instantiation"[rotated] by blast + AOT_have \\\{\}\ + using \ a "&E" "\E" "\E" by fast + moreover AOT_have \\\ (\\{\} \ \ = \)\ + apply (rule GEN; rule "\I") + using \[THEN "&E"(2), THEN "\E"(2), THEN "\E"] + "qml:2"[axiom_inst, THEN "\E"] by blast + ultimately AOT_have \(\\{\} & \\ (\\{\} \ \ = \))\ + using "&I" by blast + AOT_thus \\!\ \\{\}\ + using "uniqueness:1" "\\<^sub>d\<^sub>fI" "\I" by fast +qed + +subsection\The Theory of Actuality and Descriptions\ +text\\label{PLM: 9.8}\ + +AOT_theorem "act-cond": \\<^bold>\(\ \ \) \ (\<^bold>\\ \ \<^bold>\\)\ + using "\I" "\E"(1) "logic-actual-nec:2"[axiom_inst] by blast + +AOT_theorem "nec-imp-act": \\\ \ \<^bold>\\\ + by (metis "act-cond" "contraposition:1[2]" "\E"(4) + "qml:2"[THEN act_closure, axiom_inst] + "qml-act:2"[axiom_inst] RAA(1) "\E" "\I") + +AOT_theorem "act-conj-act:1": \\<^bold>\(\<^bold>\\ \ \)\ + using "\I" "\E"(2) "logic-actual-nec:2"[axiom_inst] + "logic-actual-nec:4"[axiom_inst] by blast + +AOT_theorem "act-conj-act:2": \\<^bold>\(\ \ \<^bold>\\)\ + by (metis "\I" "\E"(2, 4) "logic-actual-nec:2"[axiom_inst] + "logic-actual-nec:4"[axiom_inst] RAA(1)) + +AOT_theorem "act-conj-act:3": \(\<^bold>\\ & \<^bold>\\) \ \<^bold>\(\ & \)\ +proof - + AOT_have \\(\ \ (\ \ (\ & \)))\ + by (rule RN) (fact Adjunction) + AOT_hence \\<^bold>\(\ \ (\ \ (\ & \)))\ + using "nec-imp-act" "\E" by blast + AOT_hence \\<^bold>\\ \ \<^bold>\(\ \ (\ & \))\ + using "act-cond" "\E" by blast + moreover AOT_have \\<^bold>\(\ \ (\ & \)) \ (\<^bold>\\ \ \<^bold>\(\ & \))\ + by (fact "act-cond") + ultimately AOT_have \\<^bold>\\ \ (\<^bold>\\ \ \<^bold>\(\ & \))\ + using "\I" "\E" by metis + AOT_thus \(\<^bold>\\ & \<^bold>\\) \ \<^bold>\(\ & \)\ + by (metis Importation "\E") +qed + +AOT_theorem "act-conj-act:4": \\<^bold>\(\<^bold>\\ \ \)\ +proof - + AOT_have \(\<^bold>\(\<^bold>\\ \ \) & \<^bold>\(\ \ \<^bold>\\)) \ \<^bold>\((\<^bold>\\ \ \) & (\ \ \<^bold>\\))\ + by (fact "act-conj-act:3") + moreover AOT_have \\<^bold>\(\<^bold>\\ \ \) & \<^bold>\(\ \ \<^bold>\\)\ + using "&I" "act-conj-act:1" "act-conj-act:2" by simp + ultimately AOT_have \: \\<^bold>\((\<^bold>\\ \ \) & (\ \ \<^bold>\\))\ + using "\E" by blast + AOT_have \\<^bold>\(((\<^bold>\\ \ \) & (\ \ \<^bold>\\)) \ (\<^bold>\\ \ \))\ + using "conventions:3"[THEN "df-rules-formulas[2]", + THEN act_closure, axiom_inst] by blast + AOT_hence \\<^bold>\((\<^bold>\\ \ \) & (\ \ \<^bold>\\)) \ \<^bold>\(\<^bold>\\ \ \)\ + using "act-cond" "\E" by blast + AOT_thus \\<^bold>\(\<^bold>\\ \ \)\ using \ "\E" by blast +qed + +(* TODO: Consider introducing AOT_inductive. *) +inductive arbitrary_actualization for \ where + \arbitrary_actualization \ \\<^bold>\\\\ +| \arbitrary_actualization \ \\<^bold>\\\\ if \arbitrary_actualization \ \\ +declare arbitrary_actualization.cases[AOT] + arbitrary_actualization.induct[AOT] + arbitrary_actualization.simps[AOT] + arbitrary_actualization.intros[AOT] +syntax arbitrary_actualization :: \\' \ \' \ AOT_prop\ + ("ARBITRARY'_ACTUALIZATION'(_,_')") + +notepad +begin + AOT_modally_strict { + fix \ + AOT_have \ARBITRARY_ACTUALIZATION(\<^bold>\\ \ \, \<^bold>\(\<^bold>\\ \ \))\ + using AOT_PLM.arbitrary_actualization.intros by metis + AOT_have \ARBITRARY_ACTUALIZATION(\<^bold>\\ \ \, \<^bold>\\<^bold>\(\<^bold>\\ \ \))\ + using AOT_PLM.arbitrary_actualization.intros by metis + AOT_have \ARBITRARY_ACTUALIZATION(\<^bold>\\ \ \, \<^bold>\\<^bold>\\<^bold>\(\<^bold>\\ \ \))\ + using AOT_PLM.arbitrary_actualization.intros by metis + } +end + + +AOT_theorem "closure-act:1": + assumes \ARBITRARY_ACTUALIZATION(\<^bold>\\ \ \, \)\ + shows \\\ +using assms proof(induct) + case 1 + AOT_show \\<^bold>\(\<^bold>\\ \ \)\ + by (simp add: "act-conj-act:4") +next + case (2 \) + AOT_thus \\<^bold>\\\ + by (metis arbitrary_actualization.simps "\E"(1) + "logic-actual-nec:4"[axiom_inst]) +qed + +AOT_theorem "closure-act:2": \\\ \<^bold>\(\<^bold>\\{\} \ \{\})\ + by (simp add: "act-conj-act:4" "\I") + +AOT_theorem "closure-act:3": \\<^bold>\\\ \<^bold>\(\<^bold>\\{\} \ \{\})\ + by (metis (no_types, lifting) "act-conj-act:4" "\E"(1,2) "\I" + "logic-actual-nec:3"[axiom_inst] + "logic-actual-nec:4"[axiom_inst]) + +AOT_theorem "closure-act:4": \\<^bold>\\\\<^sub>1...\\\<^sub>n \<^bold>\(\<^bold>\\{\\<^sub>1...\\<^sub>n} \ \{\\<^sub>1...\\<^sub>n})\ + using "closure-act:3" . + +AOT_act_theorem "RA[1]": + assumes \\<^bold>\ \\ + shows \\<^bold>\ \<^bold>\\\ + \ \While this proof is rejected in PLM, + we merely state it as modally-fragile rule, + which addresses the concern in PLM.\ + using "\\E" assms "\E"(3) "logic-actual"[act_axiom_inst] + "logic-actual-nec:1"[axiom_inst] "modus-tollens:2" by blast +AOT_theorem "RA[2]": + assumes \\<^bold>\\<^sub>\ \\ + shows \\<^bold>\\<^sub>\ \<^bold>\\\ + \ \This rule is in fact a consequence of RN and + does not require an appeal to the semantics itself.\ + using RN assms "nec-imp-act" "vdash-properties:5" by blast +AOT_theorem "RA[3]": + assumes \\ \<^bold>\\<^sub>\ \\ + shows \\<^bold>\\ \<^bold>\\<^sub>\ \<^bold>\\\ + text\This rule is only derivable from the semantics, + but apparently no proof actually relies on it. + If this turns out to be required, it is valid to derive it from the + semantics just like RN, but we refrain from doing so, unless necessary.\ + (* using assms by (meson AOT_sem_act imageI) *) + oops \ \discard the rule\ + +AOT_act_theorem "ANeg:1": \\\<^bold>\\ \ \\\ + by (simp add: "RA[1]" "contraposition:1[1]" "deduction-theorem" + "\I" "logic-actual"[act_axiom_inst]) + +AOT_act_theorem "ANeg:2": \\\<^bold>\\\ \ \\ + using "ANeg:1" "\I" "\E"(5) "useful-tautologies:1" + "useful-tautologies:2" by blast + +AOT_theorem "Act-Basic:1": \\<^bold>\\ \ \<^bold>\\\\ + by (meson "\I"(1,2) "\E"(2) "logic-actual-nec:1"[axiom_inst] "raa-cor:1") + +AOT_theorem "Act-Basic:2": \\<^bold>\(\ & \) \ (\<^bold>\\ & \<^bold>\\)\ +proof (rule "\I"; rule "\I") + AOT_assume \\<^bold>\(\ & \)\ + moreover AOT_have \\<^bold>\((\ & \) \ \)\ + by (simp add: "RA[2]" "Conjunction Simplification"(1)) + moreover AOT_have \\<^bold>\((\ & \) \ \)\ + by (simp add: "RA[2]" "Conjunction Simplification"(2)) + ultimately AOT_show \\<^bold>\\ & \<^bold>\\\ + using "act-cond"[THEN "\E", THEN "\E"] "&I" by metis +next + AOT_assume \\<^bold>\\ & \<^bold>\\\ + AOT_thus \\<^bold>\(\ & \)\ + using "act-conj-act:3" "vdash-properties:6" by blast +qed + +AOT_theorem "Act-Basic:3": \\<^bold>\(\ \ \) \ (\<^bold>\(\ \ \) & \<^bold>\(\ \ \))\ +proof (rule "\I"; rule "\I") + AOT_assume \\<^bold>\(\ \ \)\ + moreover AOT_have \\<^bold>\((\ \ \) \ (\ \ \))\ + by (simp add: "RA[2]" "deduction-theorem" "\E"(1)) + moreover AOT_have \\<^bold>\((\ \ \) \ (\ \ \))\ + by (simp add: "RA[2]" "deduction-theorem" "\E"(2)) + ultimately AOT_show \\<^bold>\(\ \ \) & \<^bold>\(\ \ \)\ + using "act-cond"[THEN "\E", THEN "\E"] "&I" by metis +next + AOT_assume \\<^bold>\(\ \ \) & \<^bold>\(\ \ \)\ + AOT_hence \\<^bold>\((\ \ \) & (\ \ \))\ + by (metis "act-conj-act:3" "vdash-properties:10") + moreover AOT_have \\<^bold>\(((\ \ \) & (\ \ \)) \ (\ \ \))\ + by (simp add: "conventions:3" "RA[2]" "df-rules-formulas[2]" + "vdash-properties:1[2]") + ultimately AOT_show \\<^bold>\(\ \ \)\ + using "act-cond"[THEN "\E", THEN "\E"] by metis +qed + +AOT_theorem "Act-Basic:4": \(\<^bold>\(\ \ \) & \<^bold>\(\ \ \)) \ (\<^bold>\\ \ \<^bold>\\)\ +proof (rule "\I"; rule "\I") + AOT_assume 0: \\<^bold>\(\ \ \) & \<^bold>\(\ \ \)\ + AOT_show \\<^bold>\\ \ \<^bold>\\\ + using 0 "&E" "act-cond"[THEN "\E", THEN "\E"] "\I" "\I" by metis +next + AOT_assume \\<^bold>\\ \ \<^bold>\\\ + AOT_thus \\<^bold>\(\ \ \) & \<^bold>\(\ \ \)\ + by (metis "\I" "logic-actual-nec:2"[axiom_inst] "\E"(1,2) "&I") +qed + +AOT_theorem "Act-Basic:5": \\<^bold>\(\ \ \) \ (\<^bold>\\ \ \<^bold>\\)\ + using "Act-Basic:3" "Act-Basic:4" "\E"(5) by blast + +AOT_theorem "Act-Basic:6": \\<^bold>\\ \ \\<^bold>\\\ + by (simp add: "\I" "qml:2"[axiom_inst] "qml-act:1"[axiom_inst]) + +AOT_theorem "Act-Basic:7": \\<^bold>\\\ \ \\<^bold>\\\ + by (metis "Act-Basic:6" "\I" "\E" "\E"(1,2) "nec-imp-act" + "qml-act:2"[axiom_inst]) + +AOT_theorem "Act-Basic:8": \\\ \ \\<^bold>\\\ + using "Hypothetical Syllogism" "nec-imp-act" "qml-act:1"[axiom_inst] by blast + +AOT_theorem "Act-Basic:9": \\<^bold>\(\ \ \) \ (\<^bold>\\ \ \<^bold>\\)\ +proof (rule "\I"; rule "\I") + AOT_assume \\<^bold>\(\ \ \)\ + AOT_thus \\<^bold>\\ \ \<^bold>\\\ + proof (rule "raa-cor:3") + AOT_assume \\(\<^bold>\\ \ \<^bold>\\)\ + AOT_hence \\\<^bold>\\ & \\<^bold>\\\ + by (metis "\E"(1) "oth-class-taut:5:d") + AOT_hence \\<^bold>\\\ & \<^bold>\\\\ + using "logic-actual-nec:1"[axiom_inst, THEN "\E"(2)] "&E" "&I" by metis + AOT_hence \\<^bold>\(\\ & \\)\ + using "\E" "Act-Basic:2" by metis + moreover AOT_have \\<^bold>\((\\ & \\) \ \(\ \ \))\ + using "RA[2]" "\E"(6) "oth-class-taut:3:a" "oth-class-taut:5:d" by blast + moreover AOT_have \\<^bold>\(\\ & \\) \ \<^bold>\(\(\ \ \))\ + using calculation(2) by (metis "Act-Basic:5" "\E"(1)) + ultimately AOT_have \\<^bold>\(\(\ \ \))\ using "\E" by blast + AOT_thus \\\<^bold>\(\ \ \)\ + using "logic-actual-nec:1"[axiom_inst, THEN "\E"(1)] by auto + qed +next + AOT_assume \\<^bold>\\ \ \<^bold>\\\ + AOT_thus \\<^bold>\(\ \ \)\ + by (meson "RA[2]" "act-cond" "\I"(1) "\E"(1) "Disjunction Addition"(1,2)) +qed + +AOT_theorem "Act-Basic:10": \\<^bold>\\\ \{\} \ \\ \<^bold>\\{\}\ +proof - + AOT_have \: \\\<^bold>\\\ \\{\} \ \\\ \<^bold>\\\{\}\ + by (rule "oth-class-taut:4:b"[THEN "\E"(1)]) + (metis "logic-actual-nec:3"[axiom_inst]) + AOT_have \: \\\\ \<^bold>\\\{\} \ \\\ \\<^bold>\\{\}\ + by (rule "oth-class-taut:4:b"[THEN "\E"(1)]) + (rule "logic-actual-nec:1"[THEN universal_closure, + axiom_inst, THEN "cqt-basic:3"[THEN "\E"]]) + AOT_have \\<^bold>\(\\ \{\}) \ \<^bold>\(\\\ \\{\})\ + using "conventions:4"[THEN "df-rules-formulas[1]", + THEN act_closure, axiom_inst] + "conventions:4"[THEN "df-rules-formulas[2]", + THEN act_closure, axiom_inst] + "Act-Basic:4"[THEN "\E"(1)] "&I" "Act-Basic:5"[THEN "\E"(2)] by metis + also AOT_have \\ \ \\<^bold>\\\ \\{\}\ + by (simp add: "logic-actual-nec:1" "vdash-properties:1[2]") + also AOT_have \\ \ \\\ \<^bold>\ \\{\}\ using \ by blast + also AOT_have \\ \ \\\ \\<^bold>\ \{\}\ using \ by blast + also AOT_have \\ \ \\ \<^bold>\ \{\}\ + using "conventions:4"[THEN "\Df"] by (metis "\E"(6) "oth-class-taut:3:a") + finally AOT_show \\<^bold>\\\ \{\} \ \\ \<^bold>\\{\}\ . +qed + + +AOT_theorem "Act-Basic:11": + \\<^bold>\\\(\{\} \ \{\}) \ \\(\<^bold>\\{\} \ \<^bold>\\{\})\ +proof(rule "\I"; rule "\I") + AOT_assume \\<^bold>\\\(\{\} \ \{\})\ + AOT_hence \\\\<^bold>\(\{\} \ \{\})\ + using "logic-actual-nec:3"[axiom_inst, THEN "\E"(1)] by blast + AOT_hence \\<^bold>\(\{\} \ \{\})\ for \ using "\E" by blast + AOT_hence \\<^bold>\\{\} \ \<^bold>\\{\}\ for \ by (metis "Act-Basic:5" "\E"(1)) + AOT_thus \\\(\<^bold>\\{\} \ \<^bold>\\{\})\ by (rule "\I") +next + AOT_assume \\\(\<^bold>\\{\} \ \<^bold>\\{\})\ + AOT_hence \\<^bold>\\{\} \ \<^bold>\\{\}\ for \ using "\E" by blast + AOT_hence \\<^bold>\(\{\} \ \{\})\ for \ by (metis "Act-Basic:5" "\E"(2)) + AOT_hence \\\ \<^bold>\(\{\} \ \{\})\ by (rule "\I") + AOT_thus \\<^bold>\\\(\{\} \ \{\})\ + using "logic-actual-nec:3"[axiom_inst, THEN "\E"(2)] by fast +qed + +AOT_act_theorem "act-quant-uniq": + \\\(\<^bold>\\{\} \ \ = \) \ \\(\{\} \ \ = \)\ +proof(rule "\I"; rule "\I") + AOT_assume \\\(\<^bold>\\{\} \ \ = \)\ + AOT_hence \\<^bold>\\{\} \ \ = \\ for \ using "\E" by blast + AOT_hence \\{\} \ \ = \\ for \ + using "\I" "\I" "RA[1]" "\E"(1,2) "logic-actual"[act_axiom_inst] "\E" + by metis + AOT_thus \\\(\{\} \ \ = \)\ by (rule "\I") +next + AOT_assume \\\(\{\} \ \ = \)\ + AOT_hence \\{\} \ \ = \\ for \ using "\E" by blast + AOT_hence \\<^bold>\\{\} \ \ = \\ for \ + using "\I" "\I" "RA[1]" "\E"(1,2) "logic-actual"[act_axiom_inst] "\E" + by metis + AOT_thus \\\(\<^bold>\\{\} \ \ = \)\ by (rule "\I") +qed + +AOT_act_theorem "fund-cont-desc": \x = \<^bold>\x(\{x}) \ \z(\{z} \ z = x)\ + using descriptions[axiom_inst] "act-quant-uniq" "\E"(5) by fast + +AOT_act_theorem hintikka: \x = \<^bold>\x(\{x}) \ (\{x} & \z (\{z} \ z = x))\ + using "Commutativity of \"[THEN "\E"(1)] "term-out:3" + "fund-cont-desc" "\E"(5) by blast + + +locale russell_axiom = + fixes \ + assumes \_denotes_asm: "[v \ \{\}] \ [v \ \\]" +begin +AOT_act_theorem "russell-axiom": + \\{\<^bold>\x \{x}} \ \x(\{x} & \z(\{z} \ z = x) & \{x})\ +proof - + AOT_have b: \\x (x = \<^bold>\x \{x} \ (\{x} & \z(\{z} \ z = x)))\ + using hintikka "\I" by fast + show ?thesis + proof(rule "\I"; rule "\I") + AOT_assume c: \\{\<^bold>\x \{x}}\ + AOT_hence d: \\<^bold>\x \{x}\\ + using \_denotes_asm by blast + AOT_hence \\y (y = \<^bold>\x \{x})\ + by (metis "rule=I:1" "existential:1") + then AOT_obtain a where a_def: \a = \<^bold>\x \{x}\ + using "instantiation"[rotated] by blast + moreover AOT_have \a = \<^bold>\x \{x} \ (\{a} & \z(\{z} \ z = a))\ + using b "\E" by blast + ultimately AOT_have \\{a} & \z(\{z} \ z = a)\ + using "\E" by blast + moreover AOT_have \\{a}\ + proof - + AOT_have 1: \\x\y(x = y \ y = x)\ + by (simp add: "id-eq:2" "universal-cor") + AOT_have \a = \<^bold>\x \{x} \ \<^bold>\x \{x} = a\ + by (rule "\E"(1)[where \="\\<^bold>\x \{x}\"]; rule "\E"(2)[where \=a]) + (auto simp: 1 d "universal-cor") + AOT_thus \\{a}\ + using a_def c "rule=E" "\E" by blast + qed + ultimately AOT_have \\{a} & \z(\{z} \ z = a) & \{a}\ by (rule "&I") + AOT_thus \\x(\{x} & \z(\{z} \ z = x) & \{x})\ by (rule "\I") + next + AOT_assume \\x(\{x} & \z(\{z} \ z = x) & \{x})\ + then AOT_obtain b where g: \\{b} & \z(\{z} \ z = b) & \{b}\ + using "instantiation"[rotated] by blast + AOT_hence h: \b = \<^bold>\x \{x} \ (\{b} & \z(\{z} \ z = b))\ + using b "\E" by blast + AOT_have \\{b} & \z(\{z} \ z = b)\ and j: \\{b}\ + using g "&E" by blast+ + AOT_hence \b = \<^bold>\x \{x}\ using h "\E" by blast + AOT_thus \\{\<^bold>\x \{x}}\ using j "rule=E" by blast + qed +qed +end + +interpretation "russell-axiom[exe,1]": russell_axiom \\ \ . \[\]\\\ + by standard (metis "cqt:5:a[1]"[axiom_inst, THEN "\E"] "&E"(2)) +interpretation "russell-axiom[exe,2,1,1]": russell_axiom \\ \ . \[\]\\'\\ + by standard (metis "cqt:5:a[2]"[axiom_inst, THEN "\E"] "&E") +interpretation "russell-axiom[exe,2,1,2]": russell_axiom \\ \ . \[\]\'\\\ + by standard (metis "cqt:5:a[2]"[axiom_inst, THEN "\E"] "&E"(2)) +interpretation "russell-axiom[exe,2,2]": russell_axiom \\ \ . \[\]\\\\ + by standard (metis "cqt:5:a[2]"[axiom_inst, THEN "\E"] "&E"(2)) +interpretation "russell-axiom[exe,3,1,1]": russell_axiom \\ \ . \[\]\\'\''\\ + by standard (metis "cqt:5:a[3]"[axiom_inst, THEN "\E"] "&E") +interpretation "russell-axiom[exe,3,1,2]": russell_axiom \\ \ . \[\]\'\\''\\ + by standard (metis "cqt:5:a[3]"[axiom_inst, THEN "\E"] "&E") +interpretation "russell-axiom[exe,3,1,3]": russell_axiom \\ \ . \[\]\'\''\\\ + by standard (metis "cqt:5:a[3]"[axiom_inst, THEN "\E"] "&E"(2)) +interpretation "russell-axiom[exe,3,2,1]": russell_axiom \\ \ . \[\]\\\'\\ + by standard (metis "cqt:5:a[3]"[axiom_inst, THEN "\E"] "&E") +interpretation "russell-axiom[exe,3,2,2]": russell_axiom \\ \ . \[\]\\'\\\ + by standard (metis "cqt:5:a[3]"[axiom_inst, THEN "\E"] "&E"(2)) +interpretation "russell-axiom[exe,3,2,3]": russell_axiom \\ \ . \[\]\'\\\\ + by standard (metis "cqt:5:a[3]"[axiom_inst, THEN "\E"] "&E"(2)) +interpretation "russell-axiom[exe,3,3]": russell_axiom \\ \ . \[\]\\\\\ + by standard (metis "cqt:5:a[3]"[axiom_inst, THEN "\E"] "&E"(2)) + +interpretation "russell-axiom[enc,1]": russell_axiom \\ \ . \\[\]\\ + by standard (metis "cqt:5:b[1]"[axiom_inst, THEN "\E"] "&E"(2)) +interpretation "russell-axiom[enc,2,1]": russell_axiom \\ \ . \\\'[\]\\ + by standard (metis "cqt:5:b[2]"[axiom_inst, THEN "\E"] "&E") +interpretation "russell-axiom[enc,2,2]": russell_axiom \\ \ . \\'\[\]\\ + by standard (metis "cqt:5:b[2]"[axiom_inst, THEN "\E"] "&E"(2)) +interpretation "russell-axiom[enc,2,3]": russell_axiom \\ \ . \\\[\]\\ + by standard (metis "cqt:5:b[2]"[axiom_inst, THEN "\E"] "&E"(2)) +interpretation "russell-axiom[enc,3,1,1]": russell_axiom \\ \ . \\\'\''[\]\\ + by standard (metis "cqt:5:b[3]"[axiom_inst, THEN "\E"] "&E") +interpretation "russell-axiom[enc,3,1,2]": russell_axiom \\ \ . \\'\\''[\]\\ + by standard (metis "cqt:5:b[3]"[axiom_inst, THEN "\E"] "&E") +interpretation "russell-axiom[enc,3,1,3]": russell_axiom \\ \ . \\'\''\[\]\\ + by standard (metis "cqt:5:b[3]"[axiom_inst, THEN "\E"] "&E"(2)) +interpretation "russell-axiom[enc,3,2,1]": russell_axiom \\ \ . \\\\'[\]\\ + by standard (metis "cqt:5:b[3]"[axiom_inst, THEN "\E"] "&E") +interpretation "russell-axiom[enc,3,2,2]": russell_axiom \\ \ . \\\'\[\]\\ + by standard (metis "cqt:5:b[3]"[axiom_inst, THEN "\E"] "&E"(2)) +interpretation "russell-axiom[enc,3,2,3]": russell_axiom \\ \ . \\'\\[\]\\ + by standard (metis "cqt:5:b[3]"[axiom_inst, THEN "\E"] "&E"(2)) +interpretation "russell-axiom[enc,3,3]": russell_axiom \\ \ . \\\\[\]\\ + by standard (metis "cqt:5:b[3]"[axiom_inst, THEN "\E"] "&E"(2)) + +AOT_act_theorem "!-exists:1": \\<^bold>\x \{x}\ \ \!x \{x}\ +proof(rule "\I"; rule "\I") + AOT_assume \\<^bold>\x \{x}\\ + AOT_hence \\y (y = \<^bold>\x \{x})\ by (metis "rule=I:1" "existential:1") + then AOT_obtain a where \a = \<^bold>\x \{x}\ + using "instantiation"[rotated] by blast + AOT_hence \\{a} & \z (\{z} \ z = a)\ + using hintikka "\E" by blast + AOT_hence \\x (\{x} & \z (\{z} \ z = x))\ + by (rule "\I") + AOT_thus \\!x \{x}\ + using "uniqueness:1"[THEN "\\<^sub>d\<^sub>fI"] by blast +next + AOT_assume \\!x \{x}\ + AOT_hence \\x (\{x} & \z (\{z} \ z = x))\ + using "uniqueness:1"[THEN "\\<^sub>d\<^sub>fE"] by blast + then AOT_obtain b where \\{b} & \z (\{z} \ z = b)\ + using "instantiation"[rotated] by blast + AOT_hence \b = \<^bold>\x \{x}\ + using hintikka "\E" by blast + AOT_thus \\<^bold>\x \{x}\\ + by (metis "t=t-proper:2" "vdash-properties:6") +qed + +AOT_act_theorem "!-exists:2": \\y(y=\<^bold>\x \{x}) \ \!x \{x}\ + using "!-exists:1" "free-thms:1" "\E"(6) by blast + +AOT_act_theorem "y-in:1": \x = \<^bold>\x \{x} \ \{x}\ + using "&E"(1) "\I" hintikka "\E"(1) by blast + +(* Note: generalized alphabetic variant of the last theorem *) +AOT_act_theorem "y-in:2": \z = \<^bold>\x \{x} \ \{z}\ using "y-in:1". + +AOT_act_theorem "y-in:3": \\<^bold>\x \{x}\ \ \{\<^bold>\x \{x}}\ +proof(rule "\I") + AOT_assume \\<^bold>\x \{x}\\ + AOT_hence \\y (y = \<^bold>\x \{x})\ + by (metis "rule=I:1" "existential:1") + then AOT_obtain a where \a = \<^bold>\x \{x}\ + using "instantiation"[rotated] by blast + moreover AOT_have \\{a}\ + using calculation hintikka "\E"(1) "&E" by blast + ultimately AOT_show \\{\<^bold>\x \{x}}\ using "rule=E" by blast +qed + +AOT_act_theorem "y-in:4": \\y (y = \<^bold>\x \{x}) \ \{\<^bold>\x \{x}}\ + using "y-in:3"[THEN "\E"] "free-thms:1"[THEN "\E"(2)] "\I" by blast + + +AOT_theorem "act-quant-nec": + \\\ (\<^bold>\\{\} \ \ = \) \ \\(\<^bold>\\<^bold>\\{\} \ \ = \)\ +proof(rule "\I"; rule "\I") + AOT_assume \\\ (\<^bold>\\{\} \ \ = \)\ + AOT_hence \\<^bold>\\{\} \ \ = \\ for \ using "\E" by blast + AOT_hence \\<^bold>\\<^bold>\\{\} \ \ = \\ for \ + by (metis "Act-Basic:5" "act-conj-act:4" "\E"(1) "\E"(5)) + AOT_thus \\\(\<^bold>\\<^bold>\\{\} \ \ = \)\ + by (rule "\I") +next + AOT_assume \\\(\<^bold>\\<^bold>\\{\} \ \ = \)\ + AOT_hence \\<^bold>\\<^bold>\\{\} \ \ = \\ for \ using "\E" by blast + AOT_hence \\<^bold>\\{\} \ \ = \\ for \ + by (metis "Act-Basic:5" "act-conj-act:4" "\E"(1) "\E"(6)) + AOT_thus \\\ (\<^bold>\\{\} \ \ = \)\ + by (rule "\I") +qed + +AOT_theorem "equi-desc-descA:1": \x = \<^bold>\x \{x} \ x = \<^bold>\x(\<^bold>\\{x})\ +proof - + AOT_have \x = \<^bold>\x \{x} \ \z (\<^bold>\\{z} \ z = x)\ + using descriptions[axiom_inst] by blast + also AOT_have \... \ \z (\<^bold>\\<^bold>\\{z} \ z = x)\ + proof(rule "\I"; rule "\I"; rule "\I") + AOT_assume \\z (\<^bold>\\{z} \ z = x)\ + AOT_hence \\<^bold>\\{a} \ a = x\ for a + using "\E" by blast + AOT_thus \\<^bold>\\<^bold>\\{a} \ a = x\ for a + by (metis "Act-Basic:5" "act-conj-act:4" "\E"(1) "\E"(5)) + next + AOT_assume \\z (\<^bold>\\<^bold>\\{z} \ z = x)\ + AOT_hence \\<^bold>\\<^bold>\\{a} \ a = x\ for a + using "\E" by blast + AOT_thus \\<^bold>\\{a} \ a = x\ for a + by (metis "Act-Basic:5" "act-conj-act:4" "\E"(1) "\E"(6)) + qed + also AOT_have \... \ x = \<^bold>\x(\<^bold>\\{x})\ + using "Commutativity of \"[THEN "\E"(1)] descriptions[axiom_inst] by fast + finally show ?thesis . +qed + +AOT_theorem "equi-desc-descA:2": \\<^bold>\x \{x}\ \ \<^bold>\x \{x} = \<^bold>\x(\<^bold>\\{x})\ +proof(rule "\I") + AOT_assume \\<^bold>\x \{x}\\ + AOT_hence \\y (y = \<^bold>\x \{x})\ + by (metis "rule=I:1" "existential:1") + then AOT_obtain a where \a = \<^bold>\x \{x}\ + using "instantiation"[rotated] by blast + moreover AOT_have \a = \<^bold>\x(\<^bold>\\{x})\ + using calculation "equi-desc-descA:1"[THEN "\E"(1)] by blast + ultimately AOT_show \\<^bold>\x \{x} = \<^bold>\x(\<^bold>\\{x})\ + using "rule=E" by fast +qed + +AOT_theorem "nec-hintikka-scheme": + \x = \<^bold>\x \{x} \ \<^bold>\\{x} & \z(\<^bold>\\{z} \ z = x)\ +proof - + AOT_have \x = \<^bold>\x \{x} \ \z(\<^bold>\\{z} \ z = x)\ + using descriptions[axiom_inst] by blast + also AOT_have \\ \ (\<^bold>\\{x} & \z(\<^bold>\\{z} \ z = x))\ + using "Commutativity of \"[THEN "\E"(1)] "term-out:3" by fast + finally show ?thesis. +qed + +AOT_theorem "equiv-desc-eq:1": + \\<^bold>\\x(\{x} \ \{x}) \ \x (x = \<^bold>\x \{x} \ x = \<^bold>\x \{x})\ +proof(rule "\I"; rule "\I") + fix \ + AOT_assume \\<^bold>\\x(\{x} \ \{x})\ + AOT_hence \\<^bold>\(\{x} \ \{x})\ for x + using "logic-actual-nec:3"[axiom_inst, THEN "\E"(1)] "\E"(2) by blast + AOT_hence 0: \\<^bold>\\{x} \ \<^bold>\\{x}\ for x + by (metis "Act-Basic:5" "\E"(1)) + AOT_have \\ = \<^bold>\x \{x} \ \<^bold>\\{\} & \z(\<^bold>\\{z} \ z = \)\ + using "nec-hintikka-scheme" by blast + also AOT_have \... \ \<^bold>\\{\} & \z(\<^bold>\\{z} \ z = \)\ + proof (rule "\I"; rule "\I") + AOT_assume 1: \\<^bold>\\{\} & \z(\<^bold>\\{z} \ z = \)\ + AOT_hence \\<^bold>\\{z} \ z = \\ for z + using "&E" "\E" by blast + AOT_hence \\<^bold>\\{z} \ z = \\ for z + using 0 "\E" "\I" "\E" by metis + AOT_hence \\z(\<^bold>\\{z} \ z = \)\ + using "\I" by fast + moreover AOT_have \\<^bold>\\{\}\ + using "&E" 0[THEN "\E"(1)] 1 by blast + ultimately AOT_show \\<^bold>\\{\} & \z(\<^bold>\\{z} \ z = \)\ + using "&I" by blast + next + AOT_assume 1: \\<^bold>\\{\} & \z(\<^bold>\\{z} \ z = \)\ + AOT_hence \\<^bold>\\{z} \ z = \\ for z + using "&E" "\E" by blast + AOT_hence \\<^bold>\\{z} \ z = \\ for z + using 0 "\E" "\I" "\E" by metis + AOT_hence \\z(\<^bold>\\{z} \ z = \)\ + using "\I" by fast + moreover AOT_have \\<^bold>\\{\}\ + using "&E" 0[THEN "\E"(2)] 1 by blast + ultimately AOT_show \\<^bold>\\{\} & \z(\<^bold>\\{z} \ z = \)\ + using "&I" by blast + qed + also AOT_have \... \ \ = \<^bold>\x \{x}\ + using "Commutativity of \"[THEN "\E"(1)] "nec-hintikka-scheme" by blast + finally AOT_show \\ = \<^bold>\x \{x} \ \ = \<^bold>\x \{x}\ . +qed + +AOT_theorem "equiv-desc-eq:2": + \\<^bold>\x \{x}\ & \<^bold>\\x(\{x} \ \{x}) \ \<^bold>\x \{x} = \<^bold>\x \{x}\ +proof(rule "\I") + AOT_assume \\<^bold>\x \{x}\ & \<^bold>\\x(\{x} \ \{x})\ + AOT_hence 0: \\y (y = \<^bold>\x \{x})\ and + 1: \\x (x = \<^bold>\x \{x} \ x = \<^bold>\x \{x})\ + using "&E" "free-thms:1"[THEN "\E"(1)] "equiv-desc-eq:1" "\E" by blast+ + then AOT_obtain a where \a = \<^bold>\x \{x}\ + using "instantiation"[rotated] by blast + moreover AOT_have \a = \<^bold>\x \{x}\ + using calculation 1 "\E" "\E"(1) by fast + ultimately AOT_show \\<^bold>\x \{x} = \<^bold>\x \{x}\ + using "rule=E" by fast +qed + +AOT_theorem "equiv-desc-eq:3": + \\<^bold>\x \{x}\ & \\x(\{x} \ \{x}) \ \<^bold>\x \{x} = \<^bold>\x \{x}\ + using "\I" "equiv-desc-eq:2"[THEN "\E", OF "&I"] "&E" + "nec-imp-act"[THEN "\E"] by metis + +(* Note: this is a special case of "exist-nec" *) +AOT_theorem "equiv-desc-eq:4": \\<^bold>\x \{x}\ \ \\<^bold>\x \{x}\\ +proof(rule "\I") + AOT_assume \\<^bold>\x \{x}\\ + AOT_hence \\y (y = \<^bold>\x \{x})\ + by (metis "rule=I:1" "existential:1") + then AOT_obtain a where \a = \<^bold>\x \{x}\ + using "instantiation"[rotated] by blast + AOT_thus \\\<^bold>\x \{x}\\ + using "ex:2:a" "rule=E" by fast +qed + +AOT_theorem "equiv-desc-eq:5": \\<^bold>\x \{x}\ \ \y \(y = \<^bold>\x \{x})\ +proof(rule "\I") + AOT_assume \\<^bold>\x \{x}\\ + AOT_hence \\y (y = \<^bold>\x \{x})\ + by (metis "rule=I:1" "existential:1") + then AOT_obtain a where \a = \<^bold>\x \{x}\ + using "instantiation"[rotated] by blast + AOT_hence \\(a = \<^bold>\x \{x})\ + by (metis "id-nec:2" "vdash-properties:10") + AOT_thus \\y \(y = \<^bold>\x \{x})\ + by (rule "\I") +qed + +AOT_act_theorem "equiv-desc-eq2:1": + \\x (\{x} \ \{x}) \ \x (x = \<^bold>\x \{x} \ x = \<^bold>\x \{x})\ + using "\I" "logic-actual"[act_axiom_inst, THEN "\E"] + "equiv-desc-eq:1"[THEN "\E"] + "RA[1]" "deduction-theorem" by blast + +AOT_act_theorem "equiv-desc-eq2:2": + \\<^bold>\x \{x}\ & \x (\{x} \ \{x}) \ \<^bold>\x \{x} = \<^bold>\x \{x}\ + using "\I" "logic-actual"[act_axiom_inst, THEN "\E"] + "equiv-desc-eq:2"[THEN "\E", OF "&I"] + "RA[1]" "deduction-theorem" "&E" by metis + +context russell_axiom +begin +AOT_theorem "nec-russell-axiom": + \\{\<^bold>\x \{x}} \ \x(\<^bold>\\{x} & \z(\<^bold>\\{z} \ z = x) & \{x})\ +proof - + AOT_have b: \\x (x = \<^bold>\x \{x} \ (\<^bold>\\{x} & \z(\<^bold>\\{z} \ z = x)))\ + using "nec-hintikka-scheme" "\I" by fast + show ?thesis + proof(rule "\I"; rule "\I") + AOT_assume c: \\{\<^bold>\x \{x}}\ + AOT_hence d: \\<^bold>\x \{x}\\ + using \_denotes_asm by blast + AOT_hence \\y (y = \<^bold>\x \{x})\ + by (metis "rule=I:1" "existential:1") + then AOT_obtain a where a_def: \a = \<^bold>\x \{x}\ + using "instantiation"[rotated] by blast + moreover AOT_have \a = \<^bold>\x \{x} \ (\<^bold>\\{a} & \z(\<^bold>\\{z} \ z = a))\ + using b "\E" by blast + ultimately AOT_have \\<^bold>\\{a} & \z(\<^bold>\\{z} \ z = a)\ + using "\E" by blast + moreover AOT_have \\{a}\ + proof - + AOT_have 1: \\x\y(x = y \ y = x)\ + by (simp add: "id-eq:2" "universal-cor") + AOT_have \a = \<^bold>\x \{x} \ \<^bold>\x \{x} = a\ + by (rule "\E"(1)[where \="\\<^bold>\x \{x}\"]; rule "\E"(2)[where \=a]) + (auto simp: d "universal-cor" 1) + AOT_thus \\{a}\ + using a_def c "rule=E" "\E" by metis + qed + ultimately AOT_have \\<^bold>\\{a} & \z(\<^bold>\\{z} \ z = a) & \{a}\ + by (rule "&I") + AOT_thus \\x(\<^bold>\\{x} & \z(\<^bold>\\{z} \ z = x) & \{x})\ + by (rule "\I") + next + AOT_assume \\x(\<^bold>\\{x} & \z(\<^bold>\\{z} \ z = x) & \{x})\ + then AOT_obtain b where g: \\<^bold>\\{b} & \z(\<^bold>\\{z} \ z = b) & \{b}\ + using "instantiation"[rotated] by blast + AOT_hence h: \b = \<^bold>\x \{x} \ (\<^bold>\\{b} & \z(\<^bold>\\{z} \ z = b))\ + using b "\E" by blast + AOT_have \\<^bold>\\{b} & \z(\<^bold>\\{z} \ z = b)\ and j: \\{b}\ + using g "&E" by blast+ + AOT_hence \b = \<^bold>\x \{x}\ + using h "\E" by blast + AOT_thus \\{\<^bold>\x \{x}}\ + using j "rule=E" by blast + qed +qed +end + +AOT_theorem "actual-desc:1": \\<^bold>\x \{x}\ \ \!x \<^bold>\\{x}\ +proof (rule "\I"; rule "\I") + AOT_assume \\<^bold>\x \{x}\\ + AOT_hence \\y (y = \<^bold>\x \{x})\ + by (metis "rule=I:1" "existential:1") + then AOT_obtain a where \a = \<^bold>\x \{x}\ + using "instantiation"[rotated] by blast + moreover AOT_have \a = \<^bold>\x \{x} \ \z(\<^bold>\\{z} \ z = a)\ + using descriptions[axiom_inst] by blast + ultimately AOT_have \\z(\<^bold>\\{z} \ z = a)\ + using "\E" by blast + AOT_hence \\x\z(\<^bold>\\{z} \ z = x)\ by (rule "\I") + AOT_thus \\!x \<^bold>\\{x}\ + using "uniqueness:2"[THEN "\E"(2)] by fast +next + AOT_assume \\!x \<^bold>\\{x}\ + AOT_hence \\x\z(\<^bold>\\{z} \ z = x)\ + using "uniqueness:2"[THEN "\E"(1)] by fast + then AOT_obtain a where \\z(\<^bold>\\{z} \ z = a)\ + using "instantiation"[rotated] by blast + moreover AOT_have \a = \<^bold>\x \{x} \ \z(\<^bold>\\{z} \ z = a)\ + using descriptions[axiom_inst] by blast + ultimately AOT_have \a = \<^bold>\x \{x}\ + using "\E" by blast + AOT_thus \\<^bold>\x \{x}\\ + by (metis "t=t-proper:2" "vdash-properties:6") +qed + +AOT_theorem "actual-desc:2": \x = \<^bold>\x \{x} \ \<^bold>\\{x}\ + using "&E"(1) "contraposition:1[2]" "\E"(1) "nec-hintikka-scheme" + "reductio-aa:2" "vdash-properties:9" by blast + +(* Note: generalized alphabetic variant of the last theorem *) +AOT_theorem "actual-desc:3": \z = \<^bold>\x \{x} \ \<^bold>\\{z}\ + using "actual-desc:2". + +AOT_theorem "actual-desc:4": \\<^bold>\x \{x}\ \ \<^bold>\\{\<^bold>\x \{x}}\ +proof(rule "\I") + AOT_assume \\<^bold>\x \{x}\\ + AOT_hence \\y (y = \<^bold>\x \{x})\ by (metis "rule=I:1" "existential:1") + then AOT_obtain a where \a = \<^bold>\x \{x}\ using "instantiation"[rotated] by blast + AOT_thus \\<^bold>\\{\<^bold>\x \{x}}\ + using "actual-desc:2" "rule=E" "\E" by fast +qed + +AOT_theorem "actual-desc:5": \\<^bold>\x \{x} = \<^bold>\x \{x} \ \<^bold>\\x(\{x} \ \{x})\ +proof(rule "\I") + AOT_assume 0: \\<^bold>\x \{x} = \<^bold>\x \{x}\ + AOT_hence \_down: \\<^bold>\x \{x}\\ and \_down: \\<^bold>\x \{x}\\ + using "t=t-proper:1" "t=t-proper:2" "vdash-properties:6" by blast+ + AOT_hence \\y (y = \<^bold>\x \{x})\ and \\y (y = \<^bold>\x \{x})\ + by (metis "rule=I:1" "existential:1")+ + then AOT_obtain a and b where a_eq: \a = \<^bold>\x \{x}\ and b_eq: \b = \<^bold>\x \{x}\ + using "instantiation"[rotated] by metis + + AOT_have \\\\\ (\ = \ \ \ = \)\ + by (rule "\I"; rule "\I"; rule "id-eq:2") + AOT_hence \\\ (\<^bold>\x \{x} = \ \ \ = \<^bold>\x \{x})\ + using "\E" \_down by blast + AOT_hence \\<^bold>\x \{x} = \<^bold>\x \{x} \ \<^bold>\x \{x} = \<^bold>\x \{x}\ + using "\E" \_down by blast + AOT_hence 1: \\<^bold>\x \{x} = \<^bold>\x \{x}\ using 0 + "\E" by blast + + AOT_have \\<^bold>\\{x} \ \<^bold>\\{x}\ for x + proof(rule "\I"; rule "\I") + AOT_assume \\<^bold>\\{x}\ + moreover AOT_have \\<^bold>\\{x} \ x = a\ for x + using "nec-hintikka-scheme"[THEN "\E"(1), OF a_eq, THEN "&E"(2)] + "\E" by blast + ultimately AOT_have \x = a\ + using "\E" by blast + AOT_hence \x = \<^bold>\x \{x}\ + using a_eq "rule=E" by blast + AOT_hence \x = \<^bold>\x \{x}\ + using 0 "rule=E" by blast + AOT_thus \\<^bold>\\{x}\ + by (metis "actual-desc:3" "vdash-properties:6") + next + AOT_assume \\<^bold>\\{x}\ + moreover AOT_have \\<^bold>\\{x} \ x = b\ for x + using "nec-hintikka-scheme"[THEN "\E"(1), OF b_eq, THEN "&E"(2)] + "\E" by blast + ultimately AOT_have \x = b\ + using "\E" by blast + AOT_hence \x = \<^bold>\x \{x}\ + using b_eq "rule=E" by blast + AOT_hence \x = \<^bold>\x \{x}\ + using 1 "rule=E" by blast + AOT_thus \\<^bold>\\{x}\ + by (metis "actual-desc:3" "vdash-properties:6") + qed + AOT_hence \\<^bold>\(\{x} \ \{x})\ for x + by (metis "Act-Basic:5" "\E"(2)) + AOT_hence \\x \<^bold>\(\{x} \ \{x})\ + by (rule "\I") + AOT_thus \\<^bold>\\x (\{x} \ \{x})\ + using "logic-actual-nec:3"[axiom_inst, THEN "\E"(2)] by fast +qed + +AOT_theorem "!box-desc:1": \\!x \\{x} \ \y (y = \<^bold>\x \{x} \ \{y})\ +proof(rule "\I") + AOT_assume \\!x \\{x}\ + AOT_hence \: \\x (\\{x} & \z (\\{z} \ z = x))\ + using "uniqueness:1"[THEN "\\<^sub>d\<^sub>fE"] by blast + then AOT_obtain b where \: \\\{b} & \z (\\{z} \ z = b)\ + using "instantiation"[rotated] by blast + AOT_show \\y (y = \<^bold>\x \{x} \ \{y})\ + proof(rule GEN; rule "\I") + fix y + AOT_assume \y = \<^bold>\x \{x}\ + AOT_hence \\<^bold>\\{y} & \z (\<^bold>\\{z} \ z = y)\ + using "nec-hintikka-scheme"[THEN "\E"(1)] by blast + AOT_hence \\<^bold>\\{b} \ b = y\ + using "&E" "\E" by blast + moreover AOT_have \\<^bold>\\{b}\ + using \[THEN "&E"(1)] by (metis "nec-imp-act" "\E") + ultimately AOT_have \b = y\ + using "\E" by blast + moreover AOT_have \\{b}\ + using \[THEN "&E"(1)] by (metis "qml:2"[axiom_inst] "\E") + ultimately AOT_show \\{y}\ + using "rule=E" by blast + qed +qed + +AOT_theorem "!box-desc:2": + \\x (\{x} \ \\{x}) \ (\!x \{x} \ \y (y = \<^bold>\x \{x} \ \{y}))\ +proof(rule "\I"; rule "\I") + AOT_assume \\x (\{x} \ \\{x})\ + moreover AOT_assume \\!x \{x}\ + ultimately AOT_have \\!x \\{x}\ + using "nec-exist-!"[THEN "\E", THEN "\E"] by blast + AOT_thus \\y (y = \<^bold>\x \{x} \ \{y})\ + using "!box-desc:1" "\E" by blast +qed + +(* Note: vacuous in the embedding. *) +AOT_theorem "dr-alphabetic-thm": \\<^bold>\\ \{\}\ \ \<^bold>\\ \{\} = \<^bold>\\ \{\}\ + by (simp add: "rule=I:1" "\I") + +subsection\The Theory of Necessity\ +text\\label{PLM: 9.9}\ + +AOT_theorem "RM:1[prem]": + assumes \\ \<^bold>\\<^sub>\ \ \ \\ + shows \\\ \<^bold>\\<^sub>\ \\ \ \\\ +proof - + AOT_have \\\ \<^bold>\\<^sub>\ \(\ \ \)\ + using "RN[prem]" assms by blast + AOT_thus \\\ \<^bold>\\<^sub>\ \\ \ \\\ + by (metis "qml:1"[axiom_inst] "\E") +qed + +AOT_theorem "RM:1": + assumes \\<^bold>\\<^sub>\ \ \ \\ + shows \\<^bold>\\<^sub>\ \\ \ \\\ + using "RM:1[prem]" assms by blast + +lemmas RM = "RM:1" + +AOT_theorem "RM:2[prem]": + assumes \\ \<^bold>\\<^sub>\ \ \ \\ + shows \\\ \<^bold>\\<^sub>\ \\ \ \\\ +proof - + AOT_have \\ \<^bold>\\<^sub>\ \\ \ \\\ + using assms + by (simp add: "contraposition:1[1]") + AOT_hence \\\ \<^bold>\\<^sub>\ \\\ \ \\\\ + using "RM:1[prem]" by blast + AOT_thus \\\ \<^bold>\\<^sub>\ \\ \ \\\ + by (meson "\\<^sub>d\<^sub>fE" "\\<^sub>d\<^sub>fI" "conventions:5" "\I" "modus-tollens:1") +qed + +AOT_theorem "RM:2": + assumes \\<^bold>\\<^sub>\ \ \ \\ + shows \\<^bold>\\<^sub>\ \\ \ \\\ + using "RM:2[prem]" assms by blast + +lemmas "RM\" = "RM:2" + +AOT_theorem "RM:3[prem]": + assumes \\ \<^bold>\\<^sub>\ \ \ \\ + shows \\\ \<^bold>\\<^sub>\ \\ \ \\\ +proof - + AOT_have \\ \<^bold>\\<^sub>\ \ \ \\ and \\ \<^bold>\\<^sub>\ \ \ \\ + using assms "\E" "\I" by metis+ + AOT_hence \\\ \<^bold>\\<^sub>\ \\ \ \\\ and \\\ \<^bold>\\<^sub>\ \\ \ \\\ + using "RM:1[prem]" by metis+ + AOT_thus \\\ \<^bold>\\<^sub>\ \\ \ \\\ + by (simp add: "\I") +qed + +AOT_theorem "RM:3": + assumes \\<^bold>\\<^sub>\ \ \ \\ + shows \\<^bold>\\<^sub>\ \\ \ \\\ + using "RM:3[prem]" assms by blast + +lemmas RE = "RM:3" + +AOT_theorem "RM:4[prem]": + assumes \\ \<^bold>\\<^sub>\ \ \ \\ + shows \\\ \<^bold>\\<^sub>\ \\ \ \\\ +proof - + AOT_have \\ \<^bold>\\<^sub>\ \ \ \\ and \\ \<^bold>\\<^sub>\ \ \ \\ + using assms "\E" "\I" by metis+ + AOT_hence \\\ \<^bold>\\<^sub>\ \\ \ \\\ and \\\ \<^bold>\\<^sub>\ \\ \ \\\ + using "RM:2[prem]" by metis+ + AOT_thus \\\ \<^bold>\\<^sub>\ \\ \ \\\ + by (simp add: "\I") +qed + +AOT_theorem "RM:4": + assumes \\<^bold>\\<^sub>\ \ \ \\ + shows \\<^bold>\\<^sub>\ \\ \ \\\ + using "RM:4[prem]" assms by blast + +lemmas "RE\" = "RM:4" + +AOT_theorem "KBasic:1": \\\ \ \(\ \ \)\ + by (simp add: RM "pl:1"[axiom_inst]) + +AOT_theorem "KBasic:2": \\\\ \ \(\ \ \)\ + by (simp add: RM "useful-tautologies:3") + +AOT_theorem "KBasic:3": \\(\ & \) \ (\\ & \\)\ +proof (rule "\I"; rule "\I") + AOT_assume \\(\ & \)\ + AOT_thus \\\ & \\\ + by (meson RM "&I" "Conjunction Simplification"(1, 2) "\E") +next + AOT_have \\\ \ \(\ \ (\ & \))\ + by (simp add: "RM:1" Adjunction) + AOT_hence \\\ \ (\\ \ \(\ & \))\ + by (metis "Hypothetical Syllogism" "qml:1"[axiom_inst]) + moreover AOT_assume \\\ & \\\ + ultimately AOT_show \\(\ & \)\ + using "\E" "&E" by blast +qed + +AOT_theorem "KBasic:4": \\(\ \ \) \ (\(\ \ \) & \(\ \ \))\ +proof - + AOT_have \: \\((\ \ \) & (\ \ \)) \ (\(\ \ \) & \(\ \ \))\ + by (fact "KBasic:3") + AOT_modally_strict { + AOT_have \(\ \ \) \ ((\ \ \) & (\ \ \))\ + by (fact "conventions:3"[THEN "\Df"]) + } + AOT_hence \: \\(\ \ \) \ \((\ \ \) & (\ \ \))\ + by (rule RE) + with \ and \ AOT_show \\(\ \ \) \ (\(\ \ \) & \(\ \ \))\ + using "\E"(5) by blast +qed + +AOT_theorem "KBasic:5": \(\(\ \ \) & \(\ \ \)) \ (\\ \ \\)\ +proof - + AOT_have \\(\ \ \) \ (\\ \ \\)\ + by (fact "qml:1"[axiom_inst]) + moreover AOT_have \\(\ \ \) \ (\\ \ \\)\ + by (fact "qml:1"[axiom_inst]) + ultimately AOT_have \(\(\ \ \) & \(\ \ \)) \ ((\\ \ \\) & (\\ \ \\))\ + by (metis "&I" MP "Double Composition") + moreover AOT_have \((\\ \ \\) & (\\ \ \\)) \ (\\ \ \\)\ + using "conventions:3"[THEN "\\<^sub>d\<^sub>fI"] "\I" by blast + ultimately AOT_show \(\(\ \ \) & \(\ \ \)) \ (\\ \ \\)\ + by (metis "Hypothetical Syllogism") +qed + +AOT_theorem "KBasic:6": \\(\ \ \) \ (\\ \ \\)\ + using "KBasic:4" "KBasic:5" "deduction-theorem" "\E"(1) "\E" by blast +AOT_theorem "KBasic:7": \((\\ & \\) \ (\\\ & \\\)) \ \(\ \ \)\ +proof (rule "\I"; drule "\E"(1); (rule "\I")?) + AOT_assume \\\ & \\\ + AOT_hence \\\\ and \\\\ using "&E" by blast+ + AOT_hence \\(\ \ \)\ and \\(\ \ \)\ using "KBasic:1" "\E" by blast+ + AOT_hence \\(\ \ \) & \(\ \ \)\ using "&I" by blast + AOT_thus \\(\ \ \)\ by (metis "KBasic:4" "\E"(2)) +next + AOT_assume \\\\ & \\\\ + AOT_hence 0: \\(\\ & \\)\ using "KBasic:3"[THEN "\E"(2)] by blast + AOT_modally_strict { + AOT_have \(\\ & \\) \ (\ \ \)\ + by (metis "&E"(1) "&E"(2) "deduction-theorem" "\I" "reductio-aa:1") + } + AOT_hence \\(\\ & \\) \ \(\ \ \)\ + by (rule RM) + AOT_thus \\(\ \ \)\ using 0 "\E" by blast +qed(auto) + +AOT_theorem "KBasic:8": \\(\ & \) \ \(\ \ \)\ + by (meson "RM:1" "&E"(1) "&E"(2) "deduction-theorem" "\I") +AOT_theorem "KBasic:9": \\(\\ & \\) \ \(\ \ \)\ + by (metis "RM:1" "&E"(1) "&E"(2) "deduction-theorem" "\I" "raa-cor:4") +AOT_theorem "KBasic:10": \\\ \ \\\\\ + by (simp add: "RM:3" "oth-class-taut:3:b") +AOT_theorem "KBasic:11": \\\\ \ \\\\ +proof (rule "\I"; rule "\I") + AOT_show \\\\\ if \\\\\ + using that "\\<^sub>d\<^sub>fI" "conventions:5" "KBasic:10" "\E"(3) by blast +next + AOT_show \\\\\ if \\\\\ + using "\\<^sub>d\<^sub>fE" "conventions:5" "KBasic:10" "\E"(4) that by blast +qed +AOT_theorem "KBasic:12": \\\ \ \\\\\ +proof (rule "\I"; rule "\I") + AOT_show \\\\\\ if \\\\ + using "\\I" "KBasic:11" "\E"(3) that by blast +next + AOT_show \\\\ if \\\\\\ + using "KBasic:11" "\E"(1) "reductio-aa:1" that by blast +qed +AOT_theorem "KBasic:13": \\(\ \ \) \ (\\ \ \\)\ +proof - + AOT_have \\ \ \ \<^bold>\\<^sub>\ \ \ \\ by blast + AOT_hence \\(\ \ \) \<^bold>\\<^sub>\ \\ \ \\\ + using "RM:2[prem]" by blast + AOT_thus \\(\ \ \) \ (\\ \ \\)\ using "\I" by blast +qed +lemmas "K\" = "KBasic:13" +AOT_theorem "KBasic:14": \\\\ \ \\\\\\ + by (meson "RE\" "KBasic:11" "KBasic:12" "\E"(6) "oth-class-taut:3:a") +AOT_theorem "KBasic:15": \(\\ \ \\) \ \(\ \ \)\ +proof - + AOT_modally_strict { + AOT_have \\ \ (\ \ \)\ and \\ \ (\ \ \)\ + by (auto simp: "Disjunction Addition"(1) "Disjunction Addition"(2)) + } + AOT_hence \\\ \ \(\ \ \)\ and \\\ \ \(\ \ \)\ + using RM by blast+ + AOT_thus \(\\ \ \\) \ \(\ \ \)\ + by (metis "\E"(1) "deduction-theorem") +qed + +AOT_theorem "KBasic:16": \(\\ & \\) \ \(\ & \)\ + by (meson "KBasic:13" "RM:1" Adjunction "Hypothetical Syllogism" + Importation "\E") + +AOT_theorem "rule-sub-lem:1:a": + assumes \\<^bold>\\<^sub>\ \(\ \ \)\ + shows \\<^bold>\\<^sub>\ \\ \ \\\ + using "qml:2"[axiom_inst, THEN "\E", OF assms] + "\E"(1) "oth-class-taut:4:b" by blast + +AOT_theorem "rule-sub-lem:1:b": + assumes \\<^bold>\\<^sub>\ \(\ \ \)\ + shows \\<^bold>\\<^sub>\ (\ \ \) \ (\ \ \)\ + using "qml:2"[axiom_inst, THEN "\E", OF assms] + using "oth-class-taut:4:c" "vdash-properties:6" by blast + +AOT_theorem "rule-sub-lem:1:c": + assumes \\<^bold>\\<^sub>\ \(\ \ \)\ + shows \\<^bold>\\<^sub>\ (\ \ \) \ (\ \ \)\ + using "qml:2"[axiom_inst, THEN "\E", OF assms] + using "oth-class-taut:4:d" "vdash-properties:6" by blast + +AOT_theorem "rule-sub-lem:1:d": + assumes \for arbitrary \: \<^bold>\\<^sub>\ \(\{\} \ \{\})\ + shows \\<^bold>\\<^sub>\ \\ \{\} \ \\ \{\}\ +proof - + AOT_modally_strict { + AOT_have \\\ (\{\} \ \{\})\ + using "qml:2"[axiom_inst, THEN "\E", OF assms] "\I" by fast + AOT_hence 0: \\{\} \ \{\}\ for \ using "\E" by blast + AOT_show \\\ \{\} \ \\ \{\}\ + proof (rule "\I"; rule "\I") + AOT_assume \\\ \{\}\ + AOT_hence \\{\}\ for \ using "\E" by blast + AOT_hence \\{\}\ for \ using 0 "\E" by blast + AOT_thus \\\ \{\}\ by (rule "\I") + next + AOT_assume \\\ \{\}\ + AOT_hence \\{\}\ for \ using "\E" by blast + AOT_hence \\{\}\ for \ using 0 "\E" by blast + AOT_thus \\\ \{\}\ by (rule "\I") + qed + } +qed + +AOT_theorem "rule-sub-lem:1:e": + assumes \\<^bold>\\<^sub>\ \(\ \ \)\ + shows \\<^bold>\\<^sub>\ [\ \] \ [\ \]\ + using "qml:2"[axiom_inst, THEN "\E", OF assms] + using "\E"(1) "propositions-lemma:6" by blast + +AOT_theorem "rule-sub-lem:1:f": + assumes \\<^bold>\\<^sub>\ \(\ \ \)\ + shows \\<^bold>\\<^sub>\ \<^bold>\\ \ \<^bold>\\\ + using "qml:2"[axiom_inst, THEN "\E", OF assms, THEN "RA[2]"] + by (metis "Act-Basic:5" "\E"(1)) + +AOT_theorem "rule-sub-lem:1:g": + assumes \\<^bold>\\<^sub>\ \(\ \ \)\ + shows \\<^bold>\\<^sub>\ \\ \ \\\ + using "KBasic:6" assms "vdash-properties:6" by blast + +text\Note that instead of deriving @{text "rule-sub-lem:2"}, + @{text "rule-sub-lem:3"}, @{text "rule-sub-lem:4"}, + and @{text "rule-sub-nec"}, we construct substitution methods instead.\ + +class AOT_subst = + fixes AOT_subst :: "('a \ \) \ bool" + and AOT_subst_cond :: "'a \ 'a \ bool" + assumes AOT_subst: + "AOT_subst \ \ AOT_subst_cond \ \ \ [v \ \\ \\ \ \\ \\]" + +named_theorems AOT_substI + +instantiation \ :: AOT_subst +begin + +inductive AOT_subst_\ where + AOT_subst_\_id[AOT_substI]: + \AOT_subst_\ (\\. \)\ + | AOT_subst_\_const[AOT_substI]: + \AOT_subst_\ (\\. \)\ + | AOT_subst_\_not[AOT_substI]: + \AOT_subst_\ \ \ AOT_subst_\ (\ \. \\\{\}\)\ + | AOT_subst_\_imp[AOT_substI]: + \AOT_subst_\ \ \ AOT_subst_\ \ \ AOT_subst_\ (\ \. \\{\} \ \{\}\)\ + | AOT_subst_\_lambda0[AOT_substI]: + \AOT_subst_\ \ \ AOT_subst_\ (\ \. (AOT_lambda0 (\ \)))\ + | AOT_subst_\_act[AOT_substI]: + \AOT_subst_\ \ \ AOT_subst_\ (\ \. \\<^bold>\\{\}\)\ + | AOT_subst_\_box[AOT_substI]: + \AOT_subst_\ \ \ AOT_subst_\ (\ \. \\\{\}\)\ + | AOT_subst_\_by_def[AOT_substI]: + \(\ \ . AOT_model_equiv_def (\ \) (\ \)) \ + AOT_subst_\ \ \ AOT_subst_\ \\ + + +definition AOT_subst_cond_\ where + \AOT_subst_cond_\ \ \ \ \ . \ v . [v \ \ \ \]\ + +instance +proof + fix \ \ :: \ and \ :: \\ \ \\ + assume cond: \AOT_subst_cond \ \\ + assume \AOT_subst \\ + moreover AOT_have \\<^bold>\\<^sub>\ \ \ \\ + using cond unfolding AOT_subst_cond_\_def by blast + ultimately AOT_show \\<^bold>\\<^sub>\ \{\} \ \{\}\ + proof (induct arbitrary: \ \) + case AOT_subst_\_id + thus ?case + using "\E"(2) "oth-class-taut:4:b" "rule-sub-lem:1:a" by blast + next + case (AOT_subst_\_const \) + thus ?case + by (simp add: "oth-class-taut:3:a") + next + case (AOT_subst_\_not \) + thus ?case + by (simp add: RN "rule-sub-lem:1:a") + next + case (AOT_subst_\_imp \ \) + thus ?case + by (meson RN "\E"(5) "rule-sub-lem:1:b" "rule-sub-lem:1:c") + next + case (AOT_subst_\_lambda0 \) + thus ?case + by (simp add: RN "rule-sub-lem:1:e") + next + case (AOT_subst_\_act \) + thus ?case + by (simp add: RN "rule-sub-lem:1:f") + next + case (AOT_subst_\_box \) + thus ?case + by (simp add: RN "rule-sub-lem:1:g") + next + case (AOT_subst_\_by_def \ \) + AOT_modally_strict { + AOT_have \\{\} \ \{\}\ + using AOT_subst_\_by_def by simp + AOT_thus \\{\} \ \{\}\ + using "\Df"[OF AOT_subst_\_by_def(1), of _ \] + "\Df"[OF AOT_subst_\_by_def(1), of _ \] + by (metis "\E"(6) "oth-class-taut:3:a") + } + qed +qed +end + +instantiation "fun" :: (AOT_Term_id_2, AOT_subst) AOT_subst +begin + +definition AOT_subst_cond_fun :: \('a \ 'b) \ ('a \ 'b) \ bool\ where + \AOT_subst_cond_fun \ \ \ \ . \ \ . AOT_subst_cond (\ (AOT_term_of_var \)) + (\ (AOT_term_of_var \))\ + +inductive AOT_subst_fun :: \(('a \ 'b) \ \) \ bool\ where + AOT_subst_fun_const[AOT_substI]: + \AOT_subst_fun (\\. \)\ + | AOT_subst_fun_id[AOT_substI]: + \AOT_subst \ \ AOT_subst_fun (\\. \ (\ (AOT_term_of_var \)))\ + | AOT_subst_fun_all[AOT_substI]: + \AOT_subst \ \ (\ \ . AOT_subst_fun (\ (AOT_term_of_var \))) \ + AOT_subst_fun (\\ :: 'a \ 'b. \ \\\ \\ (\::'a) \\\)\ + | AOT_subst_fun_not[AOT_substI]: + \AOT_subst \ \ AOT_subst_fun (\\. \\\\ \\\)\ + | AOT_subst_fun_imp[AOT_substI]: + \AOT_subst \ \ AOT_subst \ \ AOT_subst_fun (\\. \\\ \\ \ \\ \\\)\ + | AOT_subst_fun_lambda0[AOT_substI]: + \AOT_subst \ \ AOT_subst_fun (\ \. (AOT_lambda0 (\ \)))\ + | AOT_subst_fun_act[AOT_substI]: + \AOT_subst \ \ AOT_subst_fun (\ \. \\<^bold>\\\ \\\)\ + | AOT_subst_fun_box[AOT_substI]: + \AOT_subst \ \ AOT_subst_fun (\ \. \\\\ \\\)\ + | AOT_subst_fun_def[AOT_substI]: + \(\ \ . AOT_model_equiv_def (\ \) (\ \)) \ + AOT_subst_fun \ \ AOT_subst_fun \\ + +instance proof + fix \ \ :: \'a \ 'b\ and \ :: \('a \ 'b) \ \\ + assume \AOT_subst \\ + moreover assume cond: \AOT_subst_cond \ \\ + ultimately AOT_show \\<^bold>\\<^sub>\ \\ \\ \ \\ \\\ + proof(induct) + case (AOT_subst_fun_const \) + then show ?case by (simp add: "oth-class-taut:3:a") + next + case (AOT_subst_fun_id \ x) + then show ?case by (simp add: AOT_subst AOT_subst_cond_fun_def) + next + next + case (AOT_subst_fun_all \ \) + AOT_have \\<^bold>\\<^sub>\ \(\{\, \\\} \ \{\, \\\})\ for \ + using AOT_subst_fun_all.hyps(3) AOT_subst_fun_all.prems RN by presburger + thus ?case using AOT_subst[OF AOT_subst_fun_all(1)] + by (simp add: RN "rule-sub-lem:1:d" + AOT_subst_cond_fun_def AOT_subst_cond_\_def) + next + case (AOT_subst_fun_not \) + then show ?case by (simp add: RN "rule-sub-lem:1:a") + next + case (AOT_subst_fun_imp \ \) + then show ?case + unfolding AOT_subst_cond_fun_def AOT_subst_cond_\_def + by (meson "\E"(5) "oth-class-taut:4:c" "oth-class-taut:4:d" "\E") + next + case (AOT_subst_fun_lambda0 \) + then show ?case by (simp add: RN "rule-sub-lem:1:e") + next + case (AOT_subst_fun_act \) + then show ?case by (simp add: RN "rule-sub-lem:1:f") + next + case (AOT_subst_fun_box \) + then show ?case by (simp add: RN "rule-sub-lem:1:g") + next + case (AOT_subst_fun_def \ \) + then show ?case + by (meson "df-rules-formulas[3]" "df-rules-formulas[4]" "\I" "\E"(5)) + qed +qed +end + +ML\ +fun prove_AOT_subst_tac ctxt = REPEAT (SUBGOAL (fn (trm,_) => let + fun findHeadConst (Const x) = SOME x + | findHeadConst (A $ _) = findHeadConst A + | findHeadConst _ = NONE + fun findDef (Const (\<^const_name>\AOT_model_equiv_def\, _) $ lhs $ _) + = findHeadConst lhs + | findDef (A $ B) = (case findDef A of SOME x => SOME x | _ => findDef B) + | findDef (Abs (_,_,c)) = findDef c + | findDef _ = NONE + val const_opt = (findDef trm) + val defs = case const_opt of SOME const => List.filter (fn thm => let + val concl = Thm.concl_of thm + val thmconst = (findDef concl) + in case thmconst of SOME (c,_) => fst const = c | _ => false end) + (AOT_Definitions.get ctxt) + | _ => [] + val tac = case defs of + [] => safe_step_tac (ctxt addSIs @{thms AOT_substI}) 1 + | _ => resolve_tac ctxt defs 1 + in tac end) 1) +fun getSubstThm ctxt reversed phi p q = let +val p_ty = Term.type_of p +val abs = HOLogic.mk_Trueprop (@{const AOT_subst(_)} $ phi) +val abs = Syntax.check_term ctxt abs +val substThm = Goal.prove ctxt [] [] abs + (fn {context=ctxt, prems=_} => prove_AOT_subst_tac ctxt) +val substThm = substThm RS @{thm AOT_subst} +in if reversed then let + val substThm = Drule.instantiate_normalize + (TVars.empty, Vars.make [((("\", 0), p_ty), Thm.cterm_of ctxt p), + ((("\", 0), p_ty), Thm.cterm_of ctxt q)]) substThm + val substThm = substThm RS @{thm "\E"(1)} + in substThm end +else + let + val substThm = Drule.instantiate_normalize + (TVars.empty, Vars.make [((("\", 0), p_ty), Thm.cterm_of ctxt p), + ((("\", 0), p_ty), Thm.cterm_of ctxt q)]) substThm + val substThm = substThm RS @{thm "\E"(2)} + in substThm end end +\ + +method_setup AOT_subst = \ +Scan.option (Scan.lift (Args.parens (Args.$$$ "reverse"))) -- +Scan.lift (Parse.embedded_inner_syntax -- Parse.embedded_inner_syntax) -- +Scan.option (Scan.lift (Args.$$$ "for" -- Args.colon) |-- +Scan.repeat1 (Scan.lift (Parse.embedded_inner_syntax) -- +Scan.option (Scan.lift (Args.$$$ "::" |-- Parse.embedded_inner_syntax)))) +>> (fn ((reversed,(raw_p,raw_q)),raw_bounds) => (fn ctxt => +(Method.SIMPLE_METHOD (Subgoal.FOCUS (fn {context = ctxt, params = _, + prems = prems, asms = asms, concl = concl, schematics = _} => +let +val thms = prems +val ctxt' = ctxt +val ctxt = Context_Position.set_visible false ctxt +val raw_bounds = case raw_bounds of SOME bounds => bounds | _ => [] + +val ctxt = (fold (fn (bound, ty) => fn ctxt => + let + val bound = AOT_read_term @{nonterminal \'} ctxt bound + val ty = Option.map (Syntax.read_typ ctxt) ty + val ctxt = case ty of SOME ty => let + val bound = Const ("_type_constraint_", Type ("fun", [ty,ty])) $ bound + val bound = Syntax.check_term ctxt bound + in Variable.declare_term bound ctxt end | _ => ctxt + in ctxt end)) raw_bounds ctxt + +val p = AOT_read_term @{nonterminal \'} ctxt raw_p +val p = Syntax.check_term ctxt p +val ctxt = Variable.declare_term p ctxt +val q = AOT_read_term @{nonterminal \'} ctxt raw_q +val q = Syntax.check_term ctxt q +val ctxt = Variable.declare_term q ctxt + +val bounds = (map (fn (bound, _) => + Syntax.check_term ctxt (AOT_read_term @{nonterminal \'} ctxt bound) +)) raw_bounds +val p = fold (fn bound => fn p => + Term.abs ("\", Term.type_of bound) (Term.abstract_over (bound,p))) + bounds p +val p = Syntax.check_term ctxt p +val p_ty = Term.type_of p + +val pat = @{const Trueprop} $ + (@{const AOT_model_valid_in} $ Var (("w",0), @{typ w}) $ + (Var (("\",0), Type (\<^type_name>\fun\, [p_ty, @{typ \}])) $ p)) +val univ = Unify.matchers (Context.Proof ctxt) [(pat, Thm.term_of concl)] +val univ = hd (Seq.list_of univ) (* TODO: consider all matches *) +val phi = the (Envir.lookup univ + (("\",0), Type (\<^type_name>\fun\, [p_ty, @{typ \}]))) + +val q = fold (fn bound => fn q => + Term.abs ("\", Term.type_of bound) (Term.abstract_over (bound,q))) bounds q +val q = Syntax.check_term ctxt q + +(* Reparse to report bounds as fixes. *) +val ctxt = Context_Position.restore_visible ctxt' ctxt +val ctxt' = ctxt +fun unsource str = fst (Input.source_content (Syntax.read_input str)) +val (_,ctxt') = Proof_Context.add_fixes (map (fn (str,_) => + (Binding.make (unsource str, Position.none), NONE, Mixfix.NoSyn)) raw_bounds) + ctxt' +val _ = (map (fn (x,_) => + Syntax.check_term ctxt (AOT_read_term @{nonterminal \'} ctxt' x))) + raw_bounds +val _ = AOT_read_term @{nonterminal \'} ctxt' raw_p +val _ = AOT_read_term @{nonterminal \'} ctxt' raw_q +val reversed = case reversed of SOME _ => true | _ => false +val simpThms = [@{thm AOT_subst_cond_\_def}, @{thm AOT_subst_cond_fun_def}] +in +resolve_tac ctxt [getSubstThm ctxt reversed phi p q] 1 +THEN simp_tac (ctxt addsimps simpThms) 1 +THEN (REPEAT (resolve_tac ctxt [@{thm allI}] 1)) +THEN (TRY (resolve_tac ctxt thms 1)) +end +) ctxt 1)))) +\ + +method_setup AOT_subst_def = \ +Scan.option (Scan.lift (Args.parens (Args.$$$ "reverse"))) -- +Attrib.thm +>> (fn (reversed,fact) => (fn ctxt => +(Method.SIMPLE_METHOD (Subgoal.FOCUS (fn {context = ctxt, params = _, + prems = prems, asms = asms, concl = concl, schematics = _} => +let +val c = Thm.concl_of fact +val (lhs, rhs) = case c of (\<^const>\Trueprop\ $ + (\<^const>\AOT_model_equiv_def\ $ lhs $ rhs)) => (lhs, rhs) + | _ => raise Fail "Definition expected." +val substCond = HOLogic.mk_Trueprop + (Const (\<^const_name>\AOT_subst_cond\, dummyT) $ lhs $ rhs) +val substCond = Syntax.check_term + (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) + substCond +val simpThms = [@{thm AOT_subst_cond_\_def}, + @{thm AOT_subst_cond_fun_def}, + fact RS @{thm "\Df"}] +val substCondThm = Goal.prove ctxt [] [] substCond + (fn {context=ctxt, prems=prems} => + (SUBGOAL (fn (trm,int) => + auto_tac (ctxt addsimps simpThms)) 1)) +val substThm = substCondThm RSN (2,@{thm AOT_subst}) +in +resolve_tac ctxt [substThm RS + (case reversed of NONE => @{thm "\E"(2)} | _ => @{thm "\E"(1)})] 1 +THEN prove_AOT_subst_tac ctxt +THEN (TRY (resolve_tac ctxt prems 1)) +end +) ctxt 1)))) +\ + +method_setup AOT_subst_thm = \ +Scan.option (Scan.lift (Args.parens (Args.$$$ "reverse"))) -- +Attrib.thm +>> (fn (reversed,fact) => (fn ctxt => +(Method.SIMPLE_METHOD (Subgoal.FOCUS (fn {context = ctxt, params = _, + prems = prems, asms = asms, concl = concl, schematics = _} => +let +val c = Thm.concl_of fact +val (lhs, rhs) = case c of + (\<^const>\Trueprop\ $ + (\<^const>\AOT_model_valid_in\ $ _ $ + (\<^const>\AOT_equiv\ $ lhs $ rhs))) => (lhs, rhs) + | _ => raise Fail "Equivalence expected." + +val substCond = HOLogic.mk_Trueprop + (Const (\<^const_name>\AOT_subst_cond\, dummyT) $ lhs $ rhs) +val substCond = Syntax.check_term + (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) + substCond +val simpThms = [@{thm AOT_subst_cond_\_def}, + @{thm AOT_subst_cond_fun_def}, + fact] +val substCondThm = Goal.prove ctxt [] [] substCond + (fn {context=ctxt, prems=prems} => + (SUBGOAL (fn (trm,int) => auto_tac (ctxt addsimps simpThms)) 1)) +val substThm = substCondThm RSN (2,@{thm AOT_subst}) +in +resolve_tac ctxt [substThm RS + (case reversed of NONE => @{thm "\E"(2)} | _ => @{thm "\E"(1)})] 1 +THEN prove_AOT_subst_tac ctxt +THEN (TRY (resolve_tac ctxt prems 1)) +end +) ctxt 1)))) +\ + +AOT_theorem "rule-sub-remark:1[1]": + assumes \\<^bold>\\<^sub>\ A!x \ \\E!x\ and \\A!x\ + shows \\\\E!x\ + by (AOT_subst (reverse) \\\E!x\ \A!x\) + (auto simp: assms) + +AOT_theorem "rule-sub-remark:1[2]": + assumes \\<^bold>\\<^sub>\ A!x \ \\E!x\ and \\\\E!x\ + shows \\A!x\ + by (AOT_subst \A!x\ \\\E!x\) + (auto simp: assms) + +AOT_theorem "rule-sub-remark:2[1]": + assumes \\<^bold>\\<^sub>\ [R]xy \ ([R]xy & ([Q]a \ \[Q]a))\ + and \p \ [R]xy\ + shows \p \ [R]xy & ([Q]a \ \[Q]a)\ + by (AOT_subst_thm (reverse) assms(1)) (simp add: assms(2)) + +AOT_theorem "rule-sub-remark:2[2]": + assumes \\<^bold>\\<^sub>\ [R]xy \ ([R]xy & ([Q]a \ \[Q]a))\ + and \p \ [R]xy & ([Q]a \ \[Q]a)\ + shows \p \ [R]xy\ + by (AOT_subst_thm assms(1)) (simp add: assms(2)) + +AOT_theorem "rule-sub-remark:3[1]": + assumes \for arbitrary x: \<^bold>\\<^sub>\ A!x \ \\E!x\ + and \\x A!x\ + shows \\x \\E!x\ + by (AOT_subst (reverse) \\\E!x\ \A!x\ for: x) + (auto simp: assms) + +AOT_theorem "rule-sub-remark:3[2]": + assumes \for arbitrary x: \<^bold>\\<^sub>\ A!x \ \\E!x\ + and \\x \\E!x\ + shows \\x A!x\ + by (AOT_subst \A!x\ \\\E!x\ for: x) + (auto simp: assms) + +AOT_theorem "rule-sub-remark:4[1]": + assumes \\<^bold>\\<^sub>\ \\[P]x \ [P]x\ and \\<^bold>\\\[P]x\ + shows \\<^bold>\[P]x\ + by (AOT_subst_thm (reverse) assms(1)) (simp add: assms(2)) + +AOT_theorem "rule-sub-remark:4[2]": + assumes \\<^bold>\\<^sub>\ \\[P]x \ [P]x\ and \\<^bold>\[P]x\ + shows \\<^bold>\\\[P]x\ + by (AOT_subst_thm assms(1)) (simp add: assms(2)) + +AOT_theorem "rule-sub-remark:5[1]": + assumes \\<^bold>\\<^sub>\ (\ \ \) \ (\\ \ \\)\ and \\(\ \ \)\ + shows \\(\\ \ \\)\ + by (AOT_subst_thm (reverse) assms(1)) (simp add: assms(2)) + +AOT_theorem "rule-sub-remark:5[2]": + assumes \\<^bold>\\<^sub>\ (\ \ \) \ (\\ \ \\)\ and \\(\\ \ \\)\ + shows \\(\ \ \)\ + by (AOT_subst_thm assms(1)) (simp add: assms(2)) + +AOT_theorem "rule-sub-remark:6[1]": + assumes \\<^bold>\\<^sub>\ \ \ \\ and \\(\ \ \)\ + shows \\(\ \ \)\ + by (AOT_subst_thm (reverse) assms(1)) (simp add: assms(2)) + +AOT_theorem "rule-sub-remark:6[2]": + assumes \\<^bold>\\<^sub>\ \ \ \\ and \\(\ \ \)\ + shows \\(\ \ \)\ + by (AOT_subst_thm assms(1)) (simp add: assms(2)) + +AOT_theorem "rule-sub-remark:7[1]": + assumes \\<^bold>\\<^sub>\ \ \ \\\\ and \\(\ \ \)\ + shows \\(\\\ \ \)\ + by (AOT_subst_thm (reverse) assms(1)) (simp add: assms(2)) + +AOT_theorem "rule-sub-remark:7[2]": + assumes \\<^bold>\\<^sub>\ \ \ \\\\ and \\(\\\ \ \)\ + shows \\(\ \ \)\ + by (AOT_subst_thm assms(1)) (simp add: assms(2)) + +AOT_theorem "KBasic2:1": \\\\ \ \\\\ + by (meson "conventions:5" "contraposition:2" + "Hypothetical Syllogism" "df-rules-formulas[3]" + "df-rules-formulas[4]" "\I" "useful-tautologies:1") + +AOT_theorem "KBasic2:2": \\(\ \ \) \ (\\ \ \\)\ +proof - + AOT_have \\(\ \ \) \ \\(\\ & \\)\ + by (simp add: "RE\" "oth-class-taut:5:b") + also AOT_have \\ \ \\(\\ & \\)\ + using "KBasic:11" "\E"(6) "oth-class-taut:3:a" by blast + also AOT_have \\ \ \(\\\ & \\\)\ + using "KBasic:3" "\E"(1) "oth-class-taut:4:b" by blast + also AOT_have \\ \ \(\\\ & \\\)\ + using "KBasic2:1" + by (AOT_subst \\\\\ \\\\\; AOT_subst \\\\\ \\\\\; + auto simp: "oth-class-taut:3:a") + also AOT_have \\ \ \\(\\ \ \\)\ + using "\E"(6) "oth-class-taut:3:b" "oth-class-taut:5:b" by blast + also AOT_have \\ \ \\ \ \\\ + by (simp add: "\I" "useful-tautologies:1" "useful-tautologies:2") + finally show ?thesis . +qed + +AOT_theorem "KBasic2:3": \\(\ & \) \ (\\ & \\)\ + by (metis "RM\" "&I" "Conjunction Simplification"(1,2) + "\I" "modus-tollens:1" "reductio-aa:1") + +AOT_theorem "KBasic2:4": \\(\ \ \) \ (\\ \ \\)\ +proof - + AOT_have \\(\ \ \) \ \(\\ \ \)\ + by (AOT_subst \\ \ \\ \\\ \ \\) + (auto simp: "oth-class-taut:1:c" "oth-class-taut:3:a") + also AOT_have \... \ \\\ \ \\\ + by (simp add: "KBasic2:2") + also AOT_have \... \ \\\ \ \\\ + by (AOT_subst \\\\\ \\\\\) + (auto simp: "KBasic:11" "oth-class-taut:3:a") + also AOT_have \... \ \\ \ \\\ + using "\E"(6) "oth-class-taut:1:c" "oth-class-taut:3:a" by blast + finally show ?thesis . +qed + +AOT_theorem "KBasic2:5": \\\\ \ \\\\\\ + using "conventions:5"[THEN "\Df"] + by (AOT_subst \\\\ \\\\\\; + AOT_subst \\\\\\\ \\\\\\\\\; + AOT_subst (reverse) \\\\\\\ \\\\\) + (auto simp: "oth-class-taut:3:b" "oth-class-taut:3:a") + + +AOT_theorem "KBasic2:6": \\(\ \ \) \ (\\ \ \\)\ +proof(rule "\I"; rule "raa-cor:1") + AOT_assume \\(\ \ \)\ + AOT_hence \\(\\ \ \)\ + using "conventions:2"[THEN "\Df"] + by (AOT_subst (reverse) \\\ \ \\ \\ \ \\) simp + AOT_hence 1: \\\\ \ \\\ + using "KBasic:13" "vdash-properties:10" by blast + AOT_assume \\(\\ \ \\)\ + AOT_hence \\\\\ and \\\\\ + using "&E" "\E"(1) "oth-class-taut:5:d" by blast+ + AOT_thus \\\ & \\\\ + using "&I"(1) 1[THEN "\E"] "KBasic:11" "\E"(4) "raa-cor:3" by blast +qed + +AOT_theorem "KBasic2:7": \(\(\ \ \) & \\\) \ \\\ +proof(rule "\I"; frule "&E"(1); drule "&E"(2)) + AOT_assume \\(\ \ \)\ + AOT_hence 1: \\\ \ \\\ + using "KBasic2:6" "\I"(2) "\E"(1) by blast + AOT_assume \\\\\ + AOT_hence \\\\\ using "KBasic:11" "\E"(2) by blast + AOT_thus \\\\ using 1 "\E"(2) by blast +qed + +AOT_theorem "T-S5-fund:1": \\ \ \\\ + by (meson "\\<^sub>d\<^sub>fI" "conventions:5" "contraposition:2" + "Hypothetical Syllogism" "\I" "qml:2"[axiom_inst]) +lemmas "T\" = "T-S5-fund:1" + +AOT_theorem "T-S5-fund:2": \\\\ \ \\\ +proof(rule "\I") + AOT_assume \\\\\ + AOT_hence \\\\\\\ + using "KBasic:14" "\E"(4) "raa-cor:3" by blast + moreover AOT_have \\\\ \ \\\\\ + by (fact "qml:3"[axiom_inst]) + ultimately AOT_have \\\\\\ + using "modus-tollens:1" by blast + AOT_thus \\\\ using "KBasic:12" "\E"(2) by blast +qed +lemmas "5\" = "T-S5-fund:2" + +AOT_theorem "Act-Sub:1": \\<^bold>\\ \ \\<^bold>\\\\ + by (AOT_subst \\<^bold>\\\\ \\\<^bold>\\\) + (auto simp: "logic-actual-nec:1"[axiom_inst] "oth-class-taut:3:b") + +AOT_theorem "Act-Sub:2": \\\ \ \<^bold>\\\\ + using "conventions:5"[THEN "\Df"] + by (AOT_subst \\\\ \\\\\\) + (metis "deduction-theorem" "\I" "\E"(1) "\E"(2) "\E"(3) + "logic-actual-nec:1"[axiom_inst] "qml-act:2"[axiom_inst]) + +AOT_theorem "Act-Sub:3": \\<^bold>\\ \ \\\ + using "conventions:5"[THEN "\Df"] + by (AOT_subst \\\\ \\\\\\) + (metis "Act-Sub:1" "\I" "\E"(4) "nec-imp-act" "reductio-aa:2" "\E") + +AOT_theorem "Act-Sub:4": \\<^bold>\\ \ \\<^bold>\\\ +proof (rule "\I"; rule "\I") + AOT_assume \\<^bold>\\\ + AOT_thus \\\<^bold>\\\ using "T\" "vdash-properties:10" by blast +next + AOT_assume \\\<^bold>\\\ + AOT_hence \\\\\<^bold>\\\ + using "\\<^sub>d\<^sub>fE" "conventions:5" by blast + AOT_hence \\\\<^bold>\\\\ + by (AOT_subst \\<^bold>\\\\ \\\<^bold>\\\) + (simp add: "logic-actual-nec:1"[axiom_inst]) + AOT_thus \\<^bold>\\\ + using "Act-Basic:1" "Act-Basic:6" "\E"(3) "\E"(4) + "reductio-aa:1" by blast +qed + +AOT_theorem "Act-Sub:5": \\\<^bold>\\ \ \<^bold>\\\\ + by (metis "Act-Sub:2" "Act-Sub:3" "Act-Sub:4" "\I" "\E"(1) "\E"(2) "\E") + +AOT_theorem "S5Basic:1": \\\ \ \\\\ + by (simp add: "\I" "qml:2"[axiom_inst] "qml:3"[axiom_inst]) + +AOT_theorem "S5Basic:2": \\\ \ \\\\ + by (simp add: "T\" "5\" "\I") + +AOT_theorem "S5Basic:3": \\ \ \\\\ + using "T\" "Hypothetical Syllogism" "qml:3"[axiom_inst] by blast +lemmas "B" = "S5Basic:3" + +AOT_theorem "S5Basic:4": \\\\ \ \\ + using "5\" "Hypothetical Syllogism" "qml:2"[axiom_inst] by blast +lemmas "B\" = "S5Basic:4" + +AOT_theorem "S5Basic:5": \\\ \ \\\\ + using "RM:1" "B" "5\" "Hypothetical Syllogism" by blast +lemmas "4" = "S5Basic:5" + +AOT_theorem "S5Basic:6": \\\ \ \\\\ + by (simp add: "4" "\I" "qml:2"[axiom_inst]) + +AOT_theorem "S5Basic:7": \\\\ \ \\\ + using "conventions:5"[THEN "\Df"] "oth-class-taut:3:b" + by (AOT_subst \\\\\ \\\\\\\; + AOT_subst \\\\ \\\\\\; + AOT_subst (reverse) \\\\\\\ \\\\\; + AOT_subst (reverse) \\\\\\ \\\\\) + (auto simp: "S5Basic:6" "if-p-then-p") + +lemmas "4\" = "S5Basic:7" + +AOT_theorem "S5Basic:8": \\\\ \ \\\ + by (simp add: "4\" "T\" "\I") + +AOT_theorem "S5Basic:9": \\(\ \ \\) \ (\\ \ \\)\ + apply (rule "\I"; rule "\I") + using "KBasic2:6" "5\" "\I"(3) "if-p-then-p" "vdash-properties:10" + apply blast + by (meson "KBasic:15" "4" "\I"(3) "\E"(1) "Disjunction Addition"(1) + "con-dis-taut:7" "intro-elim:1" "Commutativity of \") + +AOT_theorem "S5Basic:10": \\(\ \ \\) \ (\\ \ \\)\ +proof(rule "\I"; rule "\I") + AOT_assume \\(\ \ \\)\ + AOT_hence \\\ \ \\\\ + by (meson "KBasic2:6" "\I"(2) "\E"(1)) + AOT_thus \\\ \ \\\ + by (meson "B\" "4" "4\" "T\" "\I"(3)) +next + AOT_assume \\\ \ \\\ + AOT_hence \\\ \ \\\\ + by (meson "S5Basic:1" "B\" "S5Basic:6" "T\" "5\" "\I"(3) "intro-elim:1") + AOT_thus \\(\ \ \\)\ + by (meson "KBasic:15" "\I"(3) "\E"(1) "Disjunction Addition"(1,2)) +qed + +AOT_theorem "S5Basic:11": \\(\ & \\) \ (\\ & \\)\ +proof - + AOT_have \\(\ & \\) \ \\(\\ \ \\\)\ + by (AOT_subst \\ & \\\ \\(\\ \ \\\)\) + (auto simp: "oth-class-taut:5:a" "oth-class-taut:3:a") + also AOT_have \\ \ \\(\\ \ \\\)\ + by (AOT_subst \\\\\ \\\\\) + (auto simp: "KBasic2:1" "oth-class-taut:3:a") + also AOT_have \\ \ \\(\\ \ \\\)\ + using "KBasic:11" "\E"(6) "oth-class-taut:3:a" by blast + also AOT_have \\ \ \(\\\ \ \\\)\ + using "S5Basic:9" "\E"(1) "oth-class-taut:4:b" by blast + also AOT_have \\ \ \(\\\ \ \\\)\ + using "KBasic2:1" + by (AOT_subst \\\\\ \\\\\; AOT_subst \\\\\ \\\\\) + (auto simp: "oth-class-taut:3:a") + also AOT_have \\ \ \\ & \\\ + using "\E"(6) "oth-class-taut:3:a" "oth-class-taut:5:a" by blast + finally show ?thesis . +qed + +AOT_theorem "S5Basic:12": \\(\ & \\) \ (\\ & \\)\ +proof (rule "\I"; rule "\I") + AOT_assume \\(\ & \\)\ + AOT_hence \\\ & \\\\ + using "KBasic2:3" "vdash-properties:6" by blast + AOT_thus \\\ & \\\ + using "5\" "&I" "&E"(1) "&E"(2) "vdash-properties:6" by blast +next + AOT_assume \\\ & \\\ + moreover AOT_have \(\\\ & \\) \ \(\ & \\)\ + by (AOT_subst \\ & \\\ \\\ & \\) + (auto simp: "Commutativity of &" "KBasic:16") + ultimately AOT_show \\(\ & \\)\ + by (metis "4" "&I" "Conjunction Simplification"(1,2) "\E") +qed + +AOT_theorem "S5Basic:13": \\(\ \ \\) \ \(\\ \ \)\ +proof (rule "\I") + AOT_modally_strict { + AOT_have \\(\ \ \\) \ (\\ \ \)\ + by (meson "KBasic:13" "B\" "Hypothetical Syllogism" "\I") + } + AOT_hence \\\(\ \ \\) \ \(\\ \ \)\ + by (rule RM) + AOT_thus \\(\ \ \\) \ \(\\ \ \)\ + using "4" "Hypothetical Syllogism" by blast +next + AOT_modally_strict { + AOT_have \\(\\ \ \) \ (\ \ \\)\ + by (meson "B" "Hypothetical Syllogism" "\I" "qml:1"[axiom_inst]) + } + AOT_hence \\\(\\ \ \) \ \(\ \ \\)\ + by (rule RM) + AOT_thus \\(\\ \ \) \ \(\ \ \\)\ + using "4" "Hypothetical Syllogism" by blast +qed + +AOT_theorem "derived-S5-rules:1": + assumes \\ \<^bold>\\<^sub>\ \\ \ \\ + shows \\\ \<^bold>\\<^sub>\ \ \ \\\ +proof - + AOT_have \\\ \<^bold>\\<^sub>\ \\\ \ \\\ + using assms by (rule "RM:1[prem]") + AOT_thus \\\ \<^bold>\\<^sub>\ \ \ \\\ + using "B" "Hypothetical Syllogism" by blast +qed + +AOT_theorem "derived-S5-rules:2": + assumes \\ \<^bold>\\<^sub>\ \ \ \\\ + shows \\\ \<^bold>\\<^sub>\ \\ \ \\ +proof - + AOT_have \\\ \<^bold>\\<^sub>\ \\ \ \\\\ + using assms by (rule "RM:2[prem]") + AOT_thus \\\ \<^bold>\\<^sub>\ \\ \ \\ + using "B\" "Hypothetical Syllogism" by blast +qed + +AOT_theorem "BFs:1": \\\ \\{\} \ \\\ \{\}\ +proof - + AOT_modally_strict { + AOT_have \\\\ \\{\} \ \\\{\}\ for \ + using "cqt-orig:3" by (rule "RM\") + AOT_hence \\\\ \\{\} \ \\ \{\}\ + using "B\" "\I" "\E" "\I" by metis + } + thus ?thesis + using "derived-S5-rules:1" by blast +qed +lemmas "BF" = "BFs:1" + +AOT_theorem "BFs:2": \\\\ \{\} \ \\ \\{\}\ +proof - + AOT_have \\\\ \{\} \ \\{\}\ for \ + using RM "cqt-orig:3" by metis + thus ?thesis + using "cqt-orig:2"[THEN "\E"] "\I" by metis +qed +lemmas "CBF" = "BFs:2" + +AOT_theorem "BFs:3": \\\\ \{\} \ \\ \\{\}\ +proof(rule "\I") + AOT_modally_strict { + AOT_have \\\\ \\{\} \ \\ \\\{\}\ + using BF CBF "\I" by blast + } note \ = this + + AOT_assume \\\\ \{\}\ + AOT_hence \\\\(\\ \{\})\ + using "\\<^sub>d\<^sub>fE" "conventions:5" by blast + AOT_hence \\\\\ \\{\}\ + apply (AOT_subst \\\ \\{\}\ \\(\\ \{\})\) + using "\\<^sub>d\<^sub>fI" "conventions:3" "conventions:4" "&I" + "contraposition:2" "cqt-further:4" + "df-rules-formulas[3]" by blast + AOT_hence \\\\ \\\{\}\ + apply (AOT_subst (reverse) \\\ \\\{\}\ \\\\ \\{\}\) + using \ by blast + AOT_hence \\\\ \\\\\{\}\ + by (AOT_subst (reverse) \\\\\\{\}\ \\\\{\}\ for: \) + (simp add: "oth-class-taut:3:b") + AOT_hence \\\ \\\\{\}\ + by (rule "conventions:4"[THEN "\\<^sub>d\<^sub>fI"]) + AOT_thus \\\ \\{\}\ + using "conventions:5"[THEN "\Df"] + by (AOT_subst \\\{\}\ \\\\\{\}\ for: \) +qed +lemmas "BF\" = "BFs:3" + +AOT_theorem "BFs:4": \\\ \\{\} \ \\\ \{\}\ +proof(rule "\I") + AOT_assume \\\ \\{\}\ + AOT_hence \\\\ \\\{\}\ + using "conventions:4"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_hence \\\\ \\\{\}\ + using "KBasic2:1" + by (AOT_subst \\\\{\}\ \\\\{\}\ for: \) + moreover AOT_have \\\ \\\{\} \ \\\ \\{\}\ + using "\I" "BF" "CBF" by metis + ultimately AOT_have 1: \\\\\ \\{\}\ + using "\E"(3) by blast + AOT_show \\\\ \{\}\ + apply (rule "conventions:5"[THEN "\\<^sub>d\<^sub>fI"]) + apply (AOT_subst \\\ \{\}\ \\\\ \\{\}\) + apply (simp add: "conventions:4" "\Df") + apply (AOT_subst \\\\\ \\{\}\ \\\ \\{\}\) + by (auto simp: 1 "\I" "useful-tautologies:1" "useful-tautologies:2") +qed +lemmas "CBF\" = "BFs:4" + +AOT_theorem "sign-S5-thm:1": \\\ \\{\} \ \\\ \{\}\ +proof(rule "\I") + AOT_assume \\\ \\{\}\ + then AOT_obtain \ where \\\{\}\ using "\E" by metis + moreover AOT_have \\\\\ + by (simp add: "ex:1:a" "rule-ui:2[const_var]" RN) + moreover AOT_have \\\{\}, \\\ \<^bold>\\<^sub>\ \\\ \{\}\ for \ + proof - + AOT_have \\{\}, \\ \<^bold>\\<^sub>\ \\ \{\}\ using "existential:1" by blast + AOT_thus \\\{\}, \\\ \<^bold>\\<^sub>\ \\\ \{\}\ + using "RN[prem]"[where \="{\ \, \\\\}", simplified] by blast + qed + ultimately AOT_show \\\\ \{\}\ by blast +qed +lemmas Buridan = "sign-S5-thm:1" + +AOT_theorem "sign-S5-thm:2": \\\\ \{\} \ \\ \\{\}\ +proof - + AOT_have \\\ (\\\ \{\} \ \\{\})\ + by (simp add: "RM\" "cqt-orig:3" "\I") + AOT_thus \\\\ \{\} \ \\ \\{\}\ + using "\E"(4) "\I" "\E" "\I" by metis +qed +lemmas "Buridan\" = "sign-S5-thm:2" + +AOT_theorem "sign-S5-thm:3": + \\\\ (\{\} & \{\}) \ \(\\ \{\} & \\ \{\})\ + apply (rule "RM:2") + by (metis (no_types, lifting) "\E" "&I" "&E"(1) "&E"(2) "\I" "\I"(2)) + +AOT_theorem "sign-S5-thm:4": \\\\ (\{\} & \{\}) \ \\\ \{\}\ + apply (rule "RM:2") + by (meson "instantiation" "&E"(1) "\I" "\I"(2)) + +AOT_theorem "sign-S5-thm:5": + \(\\\ (\{\} \ \{\}) & \\\ (\{\} \ \{\})) \ \\\ (\{\} \ \{\})\ +proof - + { + fix \' \' \' + AOT_assume \\<^bold>\\<^sub>\ \' & \' \ \'\ + AOT_hence \\\' & \\' \ \\'\ + using "RN[prem]"[where \="{\', \'}"] apply simp + using "&E" "&I" "\E" "\I" by metis + } note R = this + show ?thesis by (rule R; fact AOT) +qed + +AOT_theorem "sign-S5-thm:6": + \(\\\ (\{\} \ \{\}) & \\\(\{\} \ \{\})) \ \\\(\{\} \ \{\})\ +proof - + { + fix \' \' \' + AOT_assume \\<^bold>\\<^sub>\ \' & \' \ \'\ + AOT_hence \\\' & \\' \ \\'\ + using "RN[prem]"[where \="{\', \'}"] apply simp + using "&E" "&I" "\E" "\I" by metis + } note R = this + show ?thesis by (rule R; fact AOT) +qed + +AOT_theorem "exist-nec2:1": \\\\ \ \\\ + using "B\" "RM\" "Hypothetical Syllogism" "exist-nec" by blast + +AOT_theorem "exists-nec2:2": \\\\ \ \\\\ + by (meson "Act-Sub:3" "Hypothetical Syllogism" "exist-nec" + "exist-nec2:1" "\I" "nec-imp-act") + +AOT_theorem "exists-nec2:3": \\\\ \ \\\\\ + using "KBasic2:1" "\I" "exist-nec2:1" "\E"(2) "modus-tollens:1" by blast + +AOT_theorem "exists-nec2:4": \\\\\ \ \\\\\ + by (metis "Act-Sub:3" "KBasic:12" "\I" "exist-nec" "exists-nec2:3" + "\I" "\E"(4) "nec-imp-act" "reductio-aa:1") + +AOT_theorem "id-nec2:1": \\\ = \ \ \ = \\ + using "B\" "RM\" "Hypothetical Syllogism" "id-nec:1" by blast + +AOT_theorem "id-nec2:2": \\ \ \ \ \\ \ \\ + apply (AOT_subst \\ \ \\ \\(\ = \)\) + using "=-infix"[THEN "\Df"] apply blast + using "KBasic2:1" "\I" "id-nec2:1" "\E"(2) "modus-tollens:1" by blast + +AOT_theorem "id-nec2:3": \\\ \ \ \ \ \ \\ + apply (AOT_subst \\ \ \\ \\(\ = \)\) + using "=-infix"[THEN "\Df"] apply blast + by (metis "KBasic:11" "\I" "id-nec:2" "\E"(3) "reductio-aa:2" "\E") + +AOT_theorem "id-nec2:4": \\\ = \ \ \\ = \\ + using "Hypothetical Syllogism" "id-nec2:1" "id-nec:1" by blast + +AOT_theorem "id-nec2:5": \\\ \ \ \ \\ \ \\ + using "id-nec2:3" "id-nec2:2" "\I" "\E" by metis + +AOT_theorem "sc-eq-box-box:1": \\(\ \ \\) \ (\\ \ \\)\ + apply (rule "\I"; rule "\I") + using "KBasic:13" "5\" "Hypothetical Syllogism" "\E" apply blast + by (metis "KBasic2:1" "KBasic:1" "KBasic:2" "S5Basic:13" "\E"(2) + "raa-cor:5" "\E") + +AOT_theorem "sc-eq-box-box:2": \(\(\ \ \\) \ (\\ \ \\)) \ (\\ \ \\)\ + by (metis "Act-Sub:3" "KBasic:13" "5\" "\E"(2) "\I" "\I" + "nec-imp-act" "raa-cor:2" "\E") + +AOT_theorem "sc-eq-box-box:3": \\(\ \ \\) \ (\\\ \ \\\)\ +proof (rule "\I"; rule "\I"; rule "\I") + AOT_assume \\(\ \ \\)\ + AOT_hence \\\ \ \\\ using "sc-eq-box-box:1" "\E" by blast + moreover AOT_assume \\\\\ + ultimately AOT_have \\\\\ + using "modus-tollens:1" by blast + AOT_thus \\\\\ + using "KBasic2:1" "\E"(2) by blast +next + AOT_assume \\(\ \ \\)\ + moreover AOT_assume \\\\\ + ultimately AOT_show \\\\\ + using "modus-tollens:1" "qml:2"[axiom_inst] "\E" by blast +qed + +AOT_theorem "sc-eq-box-box:4": + \(\(\ \ \\) & \(\ \ \\)) \ ((\\ \ \\) \ \(\ \ \))\ +proof(rule "\I"; rule "\I") + AOT_assume \: \\(\ \ \\) & \(\ \ \\)\ + AOT_assume \: \\\ \ \\\ + AOT_hence \(\\ & \\) \ (\\\ & \\\)\ + using "\E"(4) "oth-class-taut:4:g" "raa-cor:3" by blast + moreover { + AOT_assume \\\ & \\\ + AOT_hence \\(\ \ \)\ + using "KBasic:3" "KBasic:8" "\E"(2) "vdash-properties:10" by blast + } + moreover { + AOT_assume \\\\ & \\\\ + moreover AOT_have \\\\ \ \\\\ and \\\\ \ \\\\ + using \ "Conjunction Simplification"(1,2) + "sc-eq-box-box:3" "\E" by metis+ + ultimately AOT_have \\\\ & \\\\ + by (metis "&I" "Conjunction Simplification"(1,2) + "\E"(4) "modus-tollens:1" "raa-cor:3") + AOT_hence \\(\ \ \)\ + using "KBasic:3" "KBasic:9" "\E"(2) "\E" by blast + } + ultimately AOT_show \\(\ \ \)\ + using "\E"(2) "reductio-aa:1" by blast +qed + +AOT_theorem "sc-eq-box-box:5": + \(\(\ \ \\) & \(\ \ \\)) \ \((\ \ \) \ \(\ \ \))\ +proof (rule "\I") + AOT_assume \(\(\ \ \\) & \(\ \ \\))\ + AOT_hence \\(\(\ \ \\) & \(\ \ \\))\ + using 4[THEN "\E"] "&E" "&I" "KBasic:3" "\E"(2) by metis + moreover AOT_have \\(\(\ \ \\) & \(\ \ \\)) \ \((\ \ \) \ \(\ \ \))\ + proof (rule RM; rule "\I"; rule "\I") + AOT_modally_strict { + AOT_assume A: \(\(\ \ \\) & \(\ \ \\))\ + AOT_hence \\ \ \\\ and \\ \ \\\ + using "&E" "qml:2"[axiom_inst] "\E" by blast+ + moreover AOT_assume \\ \ \\ + ultimately AOT_have \\\ \ \\\ + using "\E" "qml:2"[axiom_inst] "\E" "\I" by meson + moreover AOT_have \(\\ \ \\) \ \(\ \ \)\ + using A "sc-eq-box-box:4" "\E" by blast + ultimately AOT_show \\(\ \ \)\ using "\E" by blast + } + qed + ultimately AOT_show \\((\ \ \) \ \(\ \ \))\ using "\E" by blast +qed + +AOT_theorem "sc-eq-box-box:6": \\(\ \ \\) \ ((\ \ \\) \ \(\ \ \))\ +proof (rule "\I"; rule "\I"; rule "raa-cor:1") + AOT_assume \\\(\ \ \)\ + AOT_hence \\\(\ \ \)\ + by (metis "KBasic:11" "\E"(1)) + AOT_hence \\(\ & \\)\ + by (AOT_subst \\ & \\\ \\(\ \ \)\) + (meson "Commutativity of \" "\E"(1) "oth-class-taut:1:b") + AOT_hence \\\\ and 2: \\\\\ + using "KBasic2:3"[THEN "\E"] "&E" by blast+ + moreover AOT_assume \\(\ \ \\)\ + ultimately AOT_have \\\\ + by (metis "\E"(1) "sc-eq-box-box:1" "\E") + AOT_hence \ + using "qml:2"[axiom_inst, THEN "\E"] by blast + moreover AOT_assume \\ \ \\\ + ultimately AOT_have \\\\ + using "\E" by blast + moreover AOT_have \\\\\ + using 2 "KBasic:12" "\\I" "intro-elim:3:d" by blast + ultimately AOT_show \\\ & \\\\ + using "&I" by blast +qed + +AOT_theorem "sc-eq-box-box:7": \\(\ \ \\) \ ((\ \ \<^bold>\\) \ \<^bold>\(\ \ \))\ +proof (rule "\I"; rule "\I"; rule "raa-cor:1") + AOT_assume \\\<^bold>\(\ \ \)\ + AOT_hence \\<^bold>\\(\ \ \)\ + by (metis "Act-Basic:1" "\E"(2)) + AOT_hence \\<^bold>\(\ & \\)\ + by (AOT_subst \\ & \\\ \\(\ \ \)\) + (meson "Commutativity of \" "\E"(1) "oth-class-taut:1:b") + AOT_hence \\<^bold>\\\ and 2: \\<^bold>\\\\ + using "Act-Basic:2"[THEN "\E"(1)] "&E" by blast+ + AOT_hence \\\\ + by (metis "Act-Sub:3" "\E") + moreover AOT_assume \\(\ \ \\)\ + ultimately AOT_have \\\\ + by (metis "\E"(1) "sc-eq-box-box:1" "\E") + AOT_hence \ + using "qml:2"[axiom_inst, THEN "\E"] by blast + moreover AOT_assume \\ \ \<^bold>\\\ + ultimately AOT_have \\<^bold>\\\ + using "\E" by blast + moreover AOT_have \\\<^bold>\\\ + using 2 by (meson "Act-Sub:1" "\E"(4) "raa-cor:3") + ultimately AOT_show \\<^bold>\\ & \\<^bold>\\\ + using "&I" by blast +qed + +AOT_theorem "sc-eq-fur:1": \\\<^bold>\\ \ \\<^bold>\\\ + using "Act-Basic:6" "Act-Sub:4" "\E"(6) by blast + +AOT_theorem "sc-eq-fur:2": \\(\ \ \\) \ (\<^bold>\\ \ \)\ + by (metis "B\" "Act-Sub:3" "KBasic:13" "T\" "Hypothetical Syllogism" + "\I" "\I" "nec-imp-act") + +AOT_theorem "sc-eq-fur:3": + \\\x (\{x} \ \\{x}) \ (\!x \{x} \ \<^bold>\x \{x}\)\ +proof (rule "\I"; rule "\I") + AOT_assume \\\x (\{x} \ \\{x})\ + AOT_hence A: \\x \(\{x} \ \\{x})\ + using CBF "\E" by blast + AOT_assume \\!x \{x}\ + then AOT_obtain a where a_def: \\{a} & \y (\{y} \ y = a)\ + using "\E"[rotated 1, OF "uniqueness:1"[THEN "\\<^sub>d\<^sub>fE"]] by blast + moreover AOT_have \\\{a}\ + using calculation A "\E"(2) "qml:2"[axiom_inst] "\E" "&E"(1) by blast + AOT_hence \\<^bold>\\{a}\ + using "nec-imp-act" "\E" by blast + moreover AOT_have \\y (\<^bold>\\{y} \ y = a)\ + proof (rule "\I"; rule "\I") + fix b + AOT_assume \\<^bold>\\{b}\ + AOT_hence \\\{b}\ + using "Act-Sub:3" "\E" by blast + moreover { + AOT_have \\(\{b} \ \\{b})\ + using A "\E"(2) by blast + AOT_hence \\\{b} \ \\{b}\ + using "KBasic:13" "5\" "Hypothetical Syllogism" "\E" by blast + } + ultimately AOT_have \\\{b}\ + using "\E" by blast + AOT_hence \\{b}\ + using "qml:2"[axiom_inst] "\E" by blast + AOT_thus \b = a\ + using a_def[THEN "&E"(2)] "\E"(2) "\E" by blast + qed + ultimately AOT_have \\<^bold>\\{a} & \y (\<^bold>\\{y} \ y = a)\ + using "&I" by blast + AOT_hence \\x (\<^bold>\\{x} & \y (\<^bold>\\{y} \ y = x))\ + using "\I" by fast + AOT_hence \\!x \<^bold>\\{x}\ + using "uniqueness:1"[THEN "\\<^sub>d\<^sub>fI"] by fast + AOT_thus \\<^bold>\x \{x}\\ + using "actual-desc:1"[THEN "\E"(2)] by blast +qed + +AOT_theorem "sc-eq-fur:4": +\\\x (\{x} \ \\{x}) \ (x = \<^bold>\x \{x} \ (\{x} & \z (\{z} \ z = x)))\ +proof (rule "\I") + AOT_assume \\\x (\{x} \ \\{x})\ + AOT_hence \\x \(\{x} \ \\{x})\ + using CBF "\E" by blast + AOT_hence A: \\<^bold>\\{\} \ \{\}\ for \ + using "sc-eq-fur:2" "\E" "\E" by fast + AOT_show \x = \<^bold>\x \{x} \ (\{x} & \z (\{z} \ z = x))\ + proof (rule "\I"; rule "\I") + AOT_assume \x = \<^bold>\x \{x}\ + AOT_hence B: \\<^bold>\\{x} & \z (\<^bold>\\{z} \ z = x)\ + using "nec-hintikka-scheme"[THEN "\E"(1)] by blast + AOT_show \\{x} & \z (\{z} \ z = x)\ + proof (rule "&I"; (rule "\I"; rule "\I")?) + AOT_show \\{x}\ + using A B[THEN "&E"(1)] "\E"(1) by blast + next + AOT_show \z = x\ if \\{z}\ for z + using that B[THEN "&E"(2)] "\E"(2) "\E" A[THEN "\E"(2)] by blast + qed + next + AOT_assume B: \\{x} & \z (\{z} \ z = x)\ + AOT_have \\<^bold>\\{x} & \z (\<^bold>\\{z} \ z = x)\ + proof(rule "&I"; (rule "\I"; rule "\I")?) + AOT_show \\<^bold>\\{x}\ + using B[THEN "&E"(1)] A[THEN "\E"(2)] by blast + next + AOT_show \b = x\ if \\<^bold>\\{b}\ for b + using A[THEN "\E"(1)] that + B[THEN "&E"(2), THEN "\E"(2), THEN "\E"] by blast + qed + AOT_thus \x = \<^bold>\x \{x}\ + using "nec-hintikka-scheme"[THEN "\E"(2)] by blast + qed +qed + +AOT_theorem "id-act:1": \\ = \ \ \<^bold>\\ = \\ + by (meson "Act-Sub:3" "Hypothetical Syllogism" + "id-nec2:1" "id-nec:2" "\I" "nec-imp-act") + +AOT_theorem "id-act:2": \\ \ \ \ \<^bold>\\ \ \\ +proof (AOT_subst \\ \ \\ \\(\ = \)\) + AOT_modally_strict { + AOT_show \\ \ \ \ \(\ = \)\ + by (simp add: "=-infix" "\Df") + } +next + AOT_show \\(\ = \) \ \<^bold>\\(\ = \)\ + proof (safe intro!: "\I" "\I") + AOT_assume \\\ = \\ + AOT_hence \\\<^bold>\\ = \\ using "id-act:1" "\E"(3) by blast + AOT_thus \\<^bold>\\\ = \\ + using "\\E" "Act-Sub:1" "\E"(3) by blast + next + AOT_assume \\<^bold>\\\ = \\ + AOT_hence \\\<^bold>\\ = \\ + using "\\I" "Act-Sub:1" "\E"(4) by blast + AOT_thus \\\ = \\ + using "id-act:1" "\E"(4) by blast + qed +qed + +AOT_theorem "A-Exists:1": \\<^bold>\\!\ \{\} \ \!\ \<^bold>\\{\}\ +proof - + AOT_have \\<^bold>\\!\ \{\} \ \<^bold>\\\\\ (\{\} \ \ = \)\ + by (AOT_subst \\!\ \{\}\ \\\\\ (\{\} \ \ = \)\) + (auto simp add: "oth-class-taut:3:a" "uniqueness:2") + also AOT_have \\ \ \\ \<^bold>\\\ (\{\} \ \ = \)\ + by (simp add: "Act-Basic:10") + also AOT_have \\ \ \\\\ \<^bold>\(\{\} \ \ = \)\ + by (AOT_subst \\<^bold>\\\ (\{\} \ \ = \)\ \\\ \<^bold>\(\{\} \ \ = \)\ for: \) + (auto simp: "logic-actual-nec:3"[axiom_inst] "oth-class-taut:3:a") + also AOT_have \\ \ \\\\ (\<^bold>\\{\} \ \<^bold>\\ = \)\ + by (AOT_subst (reverse) \\<^bold>\\{\} \ \<^bold>\\ = \\ + \\<^bold>\(\{\} \ \ = \)\ for: \ \ :: 'a) + (auto simp: "Act-Basic:5" "cqt-further:7") + also AOT_have \\ \ \\\\ (\<^bold>\\{\} \ \ = \)\ + by (AOT_subst (reverse) \\<^bold>\\ = \\ \\ = \\ for: \ \ :: 'a) + (auto simp: "id-act:1" "cqt-further:7") + also AOT_have \... \ \!\ \<^bold>\\{\}\ + using "uniqueness:2" "Commutativity of \"[THEN "\E"(1)] by fast + finally show ?thesis. +qed + +AOT_theorem "A-Exists:2": \\<^bold>\x \{x}\ \ \<^bold>\\!x \{x}\ + by (AOT_subst \\<^bold>\\!x \{x}\ \\!x \<^bold>\\{x}\) + (auto simp: "actual-desc:1" "A-Exists:1") + +AOT_theorem "id-act-desc:1": \\<^bold>\x (x = y)\\ +proof(rule "existence:1"[THEN "\\<^sub>d\<^sub>fI"]; rule "\I") + AOT_show \[\x E!x \ E!x]\<^bold>\x (x = y)\ + proof (rule "russell-axiom[exe,1].nec-russell-axiom"[THEN "\E"(2)]; + rule "\I"; (rule "&I")+) + AOT_show \\<^bold>\y = y\ by (simp add: "RA[2]" "id-eq:1") + next + AOT_show \\z (\<^bold>\z = y \ z = y)\ + apply (rule "\I") + using "id-act:1"[THEN "\E"(2)] "\I" by blast + next + AOT_show \[\x E!x \ E!x]y\ + proof (rule "lambda-predicates:2"[axiom_inst, THEN "\E", THEN "\E"(2)]) + AOT_show \[\x E!x \ E!x]\\ + by "cqt:2[lambda]" + next + AOT_show \E!y \ E!y\ + by (simp add: "if-p-then-p") + qed + qed +next + AOT_show \[\x E!x \ E!x]\\ + by "cqt:2[lambda]" +qed + +AOT_theorem "id-act-desc:2": \y = \<^bold>\x (x = y)\ + by (rule descriptions[axiom_inst, THEN "\E"(2)]; + rule "\I"; rule "id-act:1"[symmetric]) + +AOT_theorem "pre-en-eq:1[1]": \x\<^sub>1[F] \ \x\<^sub>1[F]\ + by (simp add: encoding "vdash-properties:1[2]") + +AOT_theorem "pre-en-eq:1[2]": \x\<^sub>1x\<^sub>2[F] \ \x\<^sub>1x\<^sub>2[F]\ +proof (rule "\I") + AOT_assume \x\<^sub>1x\<^sub>2[F]\ + AOT_hence \x\<^sub>1[\y [F]yx\<^sub>2]\ and \x\<^sub>2[\y [F]x\<^sub>1y]\ + using "nary-encoding[2]"[axiom_inst, THEN "\E"(1)] "&E" by blast+ + moreover AOT_have \[\y [F]yx\<^sub>2]\\ by "cqt:2" + moreover AOT_have \[\y [F]x\<^sub>1y]\\ by "cqt:2" + ultimately AOT_have \\x\<^sub>1[\y [F]yx\<^sub>2]\ and \\x\<^sub>2[\y [F]x\<^sub>1y]\ + using encoding[axiom_inst, unvarify F] "\E" "&I" by blast+ + note A = this + AOT_hence \\(x\<^sub>1[\y [F]yx\<^sub>2] & x\<^sub>2[\y [F]x\<^sub>1y])\ + using "KBasic:3"[THEN "\E"(2)] "&I" by blast + AOT_thus \\x\<^sub>1x\<^sub>2[F]\ + by (rule "nary-encoding[2]"[axiom_inst, THEN RN, + THEN "KBasic:6"[THEN "\E"], + THEN "\E"(2)]) +qed + +AOT_theorem "pre-en-eq:1[3]": \x\<^sub>1x\<^sub>2x\<^sub>3[F] \ \x\<^sub>1x\<^sub>2x\<^sub>3[F]\ +proof (rule "\I") + AOT_assume \x\<^sub>1x\<^sub>2x\<^sub>3[F]\ + AOT_hence \x\<^sub>1[\y [F]yx\<^sub>2x\<^sub>3]\ + and \x\<^sub>2[\y [F]x\<^sub>1yx\<^sub>3]\ + and \x\<^sub>3[\y [F]x\<^sub>1x\<^sub>2y]\ + using "nary-encoding[3]"[axiom_inst, THEN "\E"(1)] "&E" by blast+ + moreover AOT_have \[\y [F]yx\<^sub>2x\<^sub>3]\\ by "cqt:2" + moreover AOT_have \[\y [F]x\<^sub>1yx\<^sub>3]\\ by "cqt:2" + moreover AOT_have \[\y [F]x\<^sub>1x\<^sub>2y]\\ by "cqt:2" + ultimately AOT_have \\x\<^sub>1[\y [F]yx\<^sub>2x\<^sub>3]\ + and \\x\<^sub>2[\y [F]x\<^sub>1yx\<^sub>3]\ + and \\x\<^sub>3[\y [F]x\<^sub>1x\<^sub>2y]\ + using encoding[axiom_inst, unvarify F] "\E" by blast+ + note A = this + AOT_have B: \\(x\<^sub>1[\y [F]yx\<^sub>2x\<^sub>3] & x\<^sub>2[\y [F]x\<^sub>1yx\<^sub>3] & x\<^sub>3[\y [F]x\<^sub>1x\<^sub>2y])\ + by (rule "KBasic:3"[THEN "\E"(2)] "&I" A)+ + AOT_thus \\x\<^sub>1x\<^sub>2x\<^sub>3[F]\ + by (rule "nary-encoding[3]"[axiom_inst, THEN RN, + THEN "KBasic:6"[THEN "\E"], THEN "\E"(2)]) +qed + +AOT_theorem "pre-en-eq:1[4]": \x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F] \ \x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F]\ +proof (rule "\I") + AOT_assume \x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F]\ + AOT_hence \x\<^sub>1[\y [F]yx\<^sub>2x\<^sub>3x\<^sub>4]\ + and \x\<^sub>2[\y [F]x\<^sub>1yx\<^sub>3x\<^sub>4]\ + and \x\<^sub>3[\y [F]x\<^sub>1x\<^sub>2yx\<^sub>4]\ + and \x\<^sub>4[\y [F]x\<^sub>1x\<^sub>2x\<^sub>3y]\ + using "nary-encoding[4]"[axiom_inst, THEN "\E"(1)] "&E" by metis+ + moreover AOT_have \[\y [F]yx\<^sub>2x\<^sub>3x\<^sub>4]\\ by "cqt:2" + moreover AOT_have \[\y [F]x\<^sub>1yx\<^sub>3x\<^sub>4]\\ by "cqt:2" + moreover AOT_have \[\y [F]x\<^sub>1x\<^sub>2yx\<^sub>4]\\ by "cqt:2" + moreover AOT_have \[\y [F]x\<^sub>1x\<^sub>2x\<^sub>3y]\\ by "cqt:2" + ultimately AOT_have \\x\<^sub>1[\y [F]yx\<^sub>2x\<^sub>3x\<^sub>4]\ + and \\x\<^sub>2[\y [F]x\<^sub>1yx\<^sub>3x\<^sub>4]\ + and \\x\<^sub>3[\y [F]x\<^sub>1x\<^sub>2yx\<^sub>4]\ + and \\x\<^sub>4[\y [F]x\<^sub>1x\<^sub>2x\<^sub>3y]\ + using "\E" encoding[axiom_inst, unvarify F] by blast+ + note A = this + AOT_have B: \\(x\<^sub>1[\y [F]yx\<^sub>2x\<^sub>3x\<^sub>4] & + x\<^sub>2[\y [F]x\<^sub>1yx\<^sub>3x\<^sub>4] & + x\<^sub>3[\y [F]x\<^sub>1x\<^sub>2yx\<^sub>4] & + x\<^sub>4[\y [F]x\<^sub>1x\<^sub>2x\<^sub>3y])\ + by (rule "KBasic:3"[THEN "\E"(2)] "&I" A)+ + AOT_thus \\x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F]\ + by (rule "nary-encoding[4]"[axiom_inst, THEN RN, + THEN "KBasic:6"[THEN "\E"], THEN "\E"(2)]) +qed + +AOT_theorem "pre-en-eq:2[1]": \\x\<^sub>1[F] \ \\x\<^sub>1[F]\ +proof (rule "\I"; rule "raa-cor:1") + AOT_assume \\\\x\<^sub>1[F]\ + AOT_hence \\x\<^sub>1[F]\ + by (rule "conventions:5"[THEN "\\<^sub>d\<^sub>fI"]) + AOT_hence \x\<^sub>1[F]\ + by(rule "S5Basic:13"[THEN "\E"(1), OF "pre-en-eq:1[1]"[THEN RN], + THEN "qml:2"[axiom_inst, THEN "\E"], THEN "\E"]) + moreover AOT_assume \\x\<^sub>1[F]\ + ultimately AOT_show \x\<^sub>1[F] & \x\<^sub>1[F]\ by (rule "&I") +qed +AOT_theorem "pre-en-eq:2[2]": \\x\<^sub>1x\<^sub>2[F] \ \\x\<^sub>1x\<^sub>2[F]\ +proof (rule "\I"; rule "raa-cor:1") + AOT_assume \\\\x\<^sub>1x\<^sub>2[F]\ + AOT_hence \\x\<^sub>1x\<^sub>2[F]\ + by (rule "conventions:5"[THEN "\\<^sub>d\<^sub>fI"]) + AOT_hence \x\<^sub>1x\<^sub>2[F]\ + by(rule "S5Basic:13"[THEN "\E"(1), OF "pre-en-eq:1[2]"[THEN RN], + THEN "qml:2"[axiom_inst, THEN "\E"], THEN "\E"]) + moreover AOT_assume \\x\<^sub>1x\<^sub>2[F]\ + ultimately AOT_show \x\<^sub>1x\<^sub>2[F] & \x\<^sub>1x\<^sub>2[F]\ by (rule "&I") +qed + +AOT_theorem "pre-en-eq:2[3]": \\x\<^sub>1x\<^sub>2x\<^sub>3[F] \ \\x\<^sub>1x\<^sub>2x\<^sub>3[F]\ +proof (rule "\I"; rule "raa-cor:1") + AOT_assume \\\\x\<^sub>1x\<^sub>2x\<^sub>3[F]\ + AOT_hence \\x\<^sub>1x\<^sub>2x\<^sub>3[F]\ + by (rule "conventions:5"[THEN "\\<^sub>d\<^sub>fI"]) + AOT_hence \x\<^sub>1x\<^sub>2x\<^sub>3[F]\ + by(rule "S5Basic:13"[THEN "\E"(1), OF "pre-en-eq:1[3]"[THEN RN], + THEN "qml:2"[axiom_inst, THEN "\E"], THEN "\E"]) + moreover AOT_assume \\x\<^sub>1x\<^sub>2x\<^sub>3[F]\ + ultimately AOT_show \x\<^sub>1x\<^sub>2x\<^sub>3[F] & \x\<^sub>1x\<^sub>2x\<^sub>3[F]\ by (rule "&I") +qed + +AOT_theorem "pre-en-eq:2[4]": \\x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F] \ \\x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F]\ +proof (rule "\I"; rule "raa-cor:1") + AOT_assume \\\\x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F]\ + AOT_hence \\x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F]\ + by (rule "conventions:5"[THEN "\\<^sub>d\<^sub>fI"]) + AOT_hence \x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F]\ + by(rule "S5Basic:13"[THEN "\E"(1), OF "pre-en-eq:1[4]"[THEN RN], + THEN "qml:2"[axiom_inst, THEN "\E"], THEN "\E"]) + moreover AOT_assume \\x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F]\ + ultimately AOT_show \x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F] & \x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F]\ by (rule "&I") +qed + +AOT_theorem "en-eq:1[1]": \\x\<^sub>1[F] \ \x\<^sub>1[F]\ + using "pre-en-eq:1[1]"[THEN RN] "sc-eq-box-box:2" "\I" "\E" by metis +AOT_theorem "en-eq:1[2]": \\x\<^sub>1x\<^sub>2[F] \ \x\<^sub>1x\<^sub>2[F]\ + using "pre-en-eq:1[2]"[THEN RN] "sc-eq-box-box:2" "\I" "\E" by metis +AOT_theorem "en-eq:1[3]": \\x\<^sub>1x\<^sub>2x\<^sub>3[F] \ \x\<^sub>1x\<^sub>2x\<^sub>3[F]\ + using "pre-en-eq:1[3]"[THEN RN] "sc-eq-box-box:2" "\I" "\E" by fast +AOT_theorem "en-eq:1[4]": \\x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F] \ \x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F]\ + using "pre-en-eq:1[4]"[THEN RN] "sc-eq-box-box:2" "\I" "\E" by fast + +AOT_theorem "en-eq:2[1]": \x\<^sub>1[F] \ \x\<^sub>1[F]\ + by (simp add: "\I" "pre-en-eq:1[1]" "qml:2"[axiom_inst]) +AOT_theorem "en-eq:2[2]": \x\<^sub>1x\<^sub>2[F] \ \x\<^sub>1x\<^sub>2[F]\ + by (simp add: "\I" "pre-en-eq:1[2]" "qml:2"[axiom_inst]) +AOT_theorem "en-eq:2[3]": \x\<^sub>1x\<^sub>2x\<^sub>3[F] \ \x\<^sub>1x\<^sub>2x\<^sub>3[F]\ + by (simp add: "\I" "pre-en-eq:1[3]" "qml:2"[axiom_inst]) +AOT_theorem "en-eq:2[4]": \x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F] \ \x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F]\ + by (simp add: "\I" "pre-en-eq:1[4]" "qml:2"[axiom_inst]) + +AOT_theorem "en-eq:3[1]": \\x\<^sub>1[F] \ x\<^sub>1[F]\ + using "T\" "derived-S5-rules:2"[OF "pre-en-eq:1[1]"] "\I" by blast +AOT_theorem "en-eq:3[2]": \\x\<^sub>1x\<^sub>2[F] \ x\<^sub>1x\<^sub>2[F]\ + using "T\" "derived-S5-rules:2"[OF "pre-en-eq:1[2]"] "\I" by blast +AOT_theorem "en-eq:3[3]": \\x\<^sub>1x\<^sub>2x\<^sub>3[F] \ x\<^sub>1x\<^sub>2x\<^sub>3[F]\ + using "T\" "derived-S5-rules:2"[OF "pre-en-eq:1[3]"] "\I" by blast +AOT_theorem "en-eq:3[4]": \\x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F] \ x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F]\ + using "T\" "derived-S5-rules:2"[OF "pre-en-eq:1[4]"] "\I" by blast + +AOT_theorem "en-eq:4[1]": + \(x\<^sub>1[F] \ y\<^sub>1[G]) \ (\x\<^sub>1[F] \ \y\<^sub>1[G])\ + apply (rule "\I"; rule "\I"; rule "\I"; rule "\I") + using "qml:2"[axiom_inst, THEN "\E"] "\E"(1,2) "en-eq:2[1]" by blast+ +AOT_theorem "en-eq:4[2]": + \(x\<^sub>1x\<^sub>2[F] \ y\<^sub>1y\<^sub>2[G]) \ (\x\<^sub>1x\<^sub>2[F] \ \y\<^sub>1y\<^sub>2[G])\ + apply (rule "\I"; rule "\I"; rule "\I"; rule "\I") + using "qml:2"[axiom_inst, THEN "\E"] "\E"(1,2) "en-eq:2[2]" by blast+ +AOT_theorem "en-eq:4[3]": + \(x\<^sub>1x\<^sub>2x\<^sub>3[F] \ y\<^sub>1y\<^sub>2y\<^sub>3[G]) \ (\x\<^sub>1x\<^sub>2x\<^sub>3[F] \ \y\<^sub>1y\<^sub>2y\<^sub>3[G])\ + apply (rule "\I"; rule "\I"; rule "\I"; rule "\I") + using "qml:2"[axiom_inst, THEN "\E"] "\E"(1,2) "en-eq:2[3]" by blast+ +AOT_theorem "en-eq:4[4]": + \(x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F] \ y\<^sub>1y\<^sub>2y\<^sub>3y\<^sub>4[G]) \ (\x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F] \ \y\<^sub>1y\<^sub>2y\<^sub>3y\<^sub>4[G])\ + apply (rule "\I"; rule "\I"; rule "\I"; rule "\I") + using "qml:2"[axiom_inst, THEN "\E"] "\E"(1,2) "en-eq:2[4]" by blast+ + +AOT_theorem "en-eq:5[1]": + \\(x\<^sub>1[F] \ y\<^sub>1[G]) \ (\x\<^sub>1[F] \ \y\<^sub>1[G])\ + apply (rule "\I"; rule "\I") + using "en-eq:4[1]"[THEN "\E"(1)] "qml:2"[axiom_inst, THEN "\E"] + apply blast + using "sc-eq-box-box:4"[THEN "\E", THEN "\E"] + "&I"[OF "pre-en-eq:1[1]"[THEN RN], OF "pre-en-eq:1[1]"[THEN RN]] + by blast +AOT_theorem "en-eq:5[2]": + \\(x\<^sub>1x\<^sub>2[F] \ y\<^sub>1y\<^sub>2[G]) \ (\x\<^sub>1x\<^sub>2[F] \ \y\<^sub>1y\<^sub>2[G])\ + apply (rule "\I"; rule "\I") + using "en-eq:4[2]"[THEN "\E"(1)] "qml:2"[axiom_inst, THEN "\E"] + apply blast + using "sc-eq-box-box:4"[THEN "\E", THEN "\E"] + "&I"[OF "pre-en-eq:1[2]"[THEN RN], OF "pre-en-eq:1[2]"[THEN RN]] + by blast +AOT_theorem "en-eq:5[3]": + \\(x\<^sub>1x\<^sub>2x\<^sub>3[F] \ y\<^sub>1y\<^sub>2y\<^sub>3[G]) \ (\x\<^sub>1x\<^sub>2x\<^sub>3[F] \ \y\<^sub>1y\<^sub>2y\<^sub>3[G])\ + apply (rule "\I"; rule "\I") + using "en-eq:4[3]"[THEN "\E"(1)] "qml:2"[axiom_inst, THEN "\E"] + apply blast + using "sc-eq-box-box:4"[THEN "\E", THEN "\E"] + "&I"[OF "pre-en-eq:1[3]"[THEN RN], OF "pre-en-eq:1[3]"[THEN RN]] + by blast +AOT_theorem "en-eq:5[4]": + \\(x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F] \ y\<^sub>1y\<^sub>2y\<^sub>3y\<^sub>4[G]) \ (\x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F] \ \y\<^sub>1y\<^sub>2y\<^sub>3y\<^sub>4[G])\ + apply (rule "\I"; rule "\I") + using "en-eq:4[4]"[THEN "\E"(1)] "qml:2"[axiom_inst, THEN "\E"] + apply blast + using "sc-eq-box-box:4"[THEN "\E", THEN "\E"] + "&I"[OF "pre-en-eq:1[4]"[THEN RN], OF "pre-en-eq:1[4]"[THEN RN]] + by blast + +AOT_theorem "en-eq:6[1]": + \(x\<^sub>1[F] \ y\<^sub>1[G]) \ \(x\<^sub>1[F] \ y\<^sub>1[G])\ + using "en-eq:5[1]"[symmetric] "en-eq:4[1]" "\E"(5) by fast +AOT_theorem "en-eq:6[2]": + \(x\<^sub>1x\<^sub>2[F] \ y\<^sub>1y\<^sub>2[G]) \ \(x\<^sub>1x\<^sub>2[F] \ y\<^sub>1y\<^sub>2[G])\ + using "en-eq:5[2]"[symmetric] "en-eq:4[2]" "\E"(5) by fast +AOT_theorem "en-eq:6[3]": + \(x\<^sub>1x\<^sub>2x\<^sub>3[F] \ y\<^sub>1y\<^sub>2y\<^sub>3[G]) \ \(x\<^sub>1x\<^sub>2x\<^sub>3[F] \ y\<^sub>1y\<^sub>2y\<^sub>3[G])\ + using "en-eq:5[3]"[symmetric] "en-eq:4[3]" "\E"(5) by fast +AOT_theorem "en-eq:6[4]": + \(x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F] \ y\<^sub>1y\<^sub>2y\<^sub>3y\<^sub>4[G]) \ \(x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F] \ y\<^sub>1y\<^sub>2y\<^sub>3y\<^sub>4[G])\ + using "en-eq:5[4]"[symmetric] "en-eq:4[4]" "\E"(5) by fast + +AOT_theorem "en-eq:7[1]": \\x\<^sub>1[F] \ \\x\<^sub>1[F]\ + using "pre-en-eq:2[1]" "qml:2"[axiom_inst] "\I" by blast +AOT_theorem "en-eq:7[2]": \\x\<^sub>1x\<^sub>2[F] \ \\x\<^sub>1x\<^sub>2[F]\ + using "pre-en-eq:2[2]" "qml:2"[axiom_inst] "\I" by blast +AOT_theorem "en-eq:7[3]": \\x\<^sub>1x\<^sub>2x\<^sub>3[F] \ \\x\<^sub>1x\<^sub>2x\<^sub>3[F]\ + using "pre-en-eq:2[3]" "qml:2"[axiom_inst] "\I" by blast +AOT_theorem "en-eq:7[4]": \\x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F] \ \\x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F]\ + using "pre-en-eq:2[4]" "qml:2"[axiom_inst] "\I" by blast + +AOT_theorem "en-eq:8[1]": \\\x\<^sub>1[F] \ \x\<^sub>1[F]\ + using "en-eq:2[1]"[THEN "oth-class-taut:4:b"[THEN "\E"(1)]] + "KBasic:11" "\E"(5)[symmetric] by blast +AOT_theorem "en-eq:8[2]": \\\x\<^sub>1x\<^sub>2[F] \ \x\<^sub>1x\<^sub>2[F]\ + using "en-eq:2[2]"[THEN "oth-class-taut:4:b"[THEN "\E"(1)]] + "KBasic:11" "\E"(5)[symmetric] by blast +AOT_theorem "en-eq:8[3]": \\\x\<^sub>1x\<^sub>2x\<^sub>3[F] \ \x\<^sub>1x\<^sub>2x\<^sub>3[F]\ + using "en-eq:2[3]"[THEN "oth-class-taut:4:b"[THEN "\E"(1)]] + "KBasic:11" "\E"(5)[symmetric] by blast +AOT_theorem "en-eq:8[4]": \\\x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F] \ \x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F]\ + using "en-eq:2[4]"[THEN "oth-class-taut:4:b"[THEN "\E"(1)]] + "KBasic:11" "\E"(5)[symmetric] by blast + +AOT_theorem "en-eq:9[1]": \\\x\<^sub>1[F] \ \\x\<^sub>1[F]\ + using "en-eq:7[1]" "en-eq:8[1]" "\E"(5) by blast +AOT_theorem "en-eq:9[2]": \\\x\<^sub>1x\<^sub>2[F] \ \\x\<^sub>1x\<^sub>2[F]\ + using "en-eq:7[2]" "en-eq:8[2]" "\E"(5) by blast +AOT_theorem "en-eq:9[3]": \\\x\<^sub>1x\<^sub>2x\<^sub>3[F] \ \\x\<^sub>1x\<^sub>2x\<^sub>3[F]\ + using "en-eq:7[3]" "en-eq:8[3]" "\E"(5) by blast +AOT_theorem "en-eq:9[4]": \\\x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F] \ \\x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F]\ + using "en-eq:7[4]" "en-eq:8[4]" "\E"(5) by blast + +AOT_theorem "en-eq:10[1]": \\<^bold>\x\<^sub>1[F] \ x\<^sub>1[F]\ + by (metis "Act-Sub:3" "deduction-theorem" "\I" "\E"(1) + "nec-imp-act" "en-eq:3[1]" "pre-en-eq:1[1]") +AOT_theorem "en-eq:10[2]": \\<^bold>\x\<^sub>1x\<^sub>2[F] \ x\<^sub>1x\<^sub>2[F]\ + by (metis "Act-Sub:3" "deduction-theorem" "\I" "\E"(1) + "nec-imp-act" "en-eq:3[2]" "pre-en-eq:1[2]") +AOT_theorem "en-eq:10[3]": \\<^bold>\x\<^sub>1x\<^sub>2x\<^sub>3[F] \ x\<^sub>1x\<^sub>2x\<^sub>3[F]\ + by (metis "Act-Sub:3" "deduction-theorem" "\I" "\E"(1) + "nec-imp-act" "en-eq:3[3]" "pre-en-eq:1[3]") +AOT_theorem "en-eq:10[4]": \\<^bold>\x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F] \ x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[F]\ + by (metis "Act-Sub:3" "deduction-theorem" "\I" "\E"(1) + "nec-imp-act" "en-eq:3[4]" "pre-en-eq:1[4]") + +AOT_theorem "oa-facts:1": \O!x \ \O!x\ +proof(rule "\I") + AOT_modally_strict { + AOT_have \[\x \E!x]x \ \E!x\ + by (rule "lambda-predicates:2"[axiom_inst, THEN "\E"]) "cqt:2" + } note \ = this + AOT_assume \O!x\ + AOT_hence \[\x \E!x]x\ + by (rule "=\<^sub>d\<^sub>fE"(2)[OF AOT_ordinary, rotated 1]) "cqt:2" + AOT_hence \\E!x\ using \[THEN "\E"(1)] by blast + AOT_hence \\\E!x\ using "qml:3"[axiom_inst, THEN "\E"] by blast + AOT_hence \\[\x \E!x]x\ + by (AOT_subst \[\x \E!x]x\ \\E!x\) + (auto simp: \) + AOT_thus \\O!x\ + by (rule "=\<^sub>d\<^sub>fI"(2)[OF AOT_ordinary, rotated 1]) "cqt:2" +qed + +AOT_theorem "oa-facts:2": \A!x \ \A!x\ +proof(rule "\I") + AOT_modally_strict { + AOT_have \[\x \\E!x]x \ \\E!x\ + by (rule "lambda-predicates:2"[axiom_inst, THEN "\E"]) "cqt:2" + } note \ = this + AOT_assume \A!x\ + AOT_hence \[\x \\E!x]x\ + by (rule "=\<^sub>d\<^sub>fE"(2)[OF AOT_abstract, rotated 1]) "cqt:2" + AOT_hence \\\E!x\ using \[THEN "\E"(1)] by blast + AOT_hence \\\E!x\ using "KBasic2:1"[THEN "\E"(2)] by blast + AOT_hence \\\\E!x\ using "4"[THEN "\E"] by blast + AOT_hence \\\\E!x\ + using "KBasic2:1" + by (AOT_subst (reverse) \\\E!x\ \\\E!x\) blast + AOT_hence \\[\x \\E!x]x\ + by (AOT_subst \[\x \\E!x]x\ \\\E!x\) + (auto simp: \) + AOT_thus \\A!x\ + by (rule "=\<^sub>d\<^sub>fI"(2)[OF AOT_abstract, rotated 1]) "cqt:2[lambda]" +qed + +AOT_theorem "oa-facts:3": \\O!x \ O!x\ + using "oa-facts:1" "B\" "RM\" "Hypothetical Syllogism" by blast +AOT_theorem "oa-facts:4": \\A!x \ A!x\ + using "oa-facts:2" "B\" "RM\" "Hypothetical Syllogism" by blast + +AOT_theorem "oa-facts:5": \\O!x \ \O!x\ + by (meson "Act-Sub:3" "Hypothetical Syllogism" "\I" "nec-imp-act" + "oa-facts:1" "oa-facts:3") + +AOT_theorem "oa-facts:6": \\A!x \ \A!x\ + by (meson "Act-Sub:3" "Hypothetical Syllogism" "\I" "nec-imp-act" + "oa-facts:2" "oa-facts:4") + +AOT_theorem "oa-facts:7": \O!x \ \<^bold>\O!x\ + by (meson "Act-Sub:3" "Hypothetical Syllogism" "\I" "nec-imp-act" + "oa-facts:1" "oa-facts:3") + +AOT_theorem "oa-facts:8": \A!x \ \<^bold>\A!x\ + by (meson "Act-Sub:3" "Hypothetical Syllogism" "\I" "nec-imp-act" + "oa-facts:2" "oa-facts:4") + +subsection\The Theory of Relations\ +text\\label{PLM: 9.10}\ + +AOT_theorem "beta-C-meta": + \[\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n, \\<^sub>1...\\<^sub>n}]\ \ + ([\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n, \\<^sub>1...\\<^sub>n}]\\<^sub>1...\\<^sub>n \ \{\\<^sub>1...\\<^sub>n, \\<^sub>1...\\<^sub>n})\ + using "lambda-predicates:2"[axiom_inst] by blast + +AOT_theorem "beta-C-cor:1": + \(\\\<^sub>1...\\\<^sub>n([\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n, \\<^sub>1...\\<^sub>n}]\)) \ + \\\<^sub>1...\\\<^sub>n ([\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n, \\<^sub>1...\\<^sub>n}]\\<^sub>1...\\<^sub>n \ \{\\<^sub>1...\\<^sub>n, \\<^sub>1...\\<^sub>n})\ + apply (rule "cqt-basic:14"[where 'a='a, THEN "\E"]) + using "beta-C-meta" "\I" by fast + +AOT_theorem "beta-C-cor:2": + \[\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\ \ + \\\<^sub>1...\\\<^sub>n ([\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\\<^sub>1...\\<^sub>n \ \{\\<^sub>1...\\<^sub>n})\ + apply (rule "\I"; rule "\I") + using "beta-C-meta"[THEN "\E"] by fast + +(* TODO: add better syntax parsing for INSTANCE_OF_CQT_2 *) +theorem "beta-C-cor:3": + assumes \\\\<^sub>1\\<^sub>n. AOT_instance_of_cqt_2 (\ (AOT_term_of_var \\<^sub>1\\<^sub>n))\ + shows \[v \ \\\<^sub>1...\\\<^sub>n ([\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n, \\<^sub>1...\\<^sub>n}]\\<^sub>1...\\<^sub>n \ + \{\\<^sub>1...\\<^sub>n, \\<^sub>1...\\<^sub>n})]\ + using "cqt:2[lambda]"[axiom_inst, OF assms] + "beta-C-cor:1"[THEN "\E"] "\I" by fast + +AOT_theorem "betaC:1:a": \[\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\\<^sub>1...\\<^sub>n \<^bold>\\<^sub>\ \{\\<^sub>1...\\<^sub>n}\ +proof - + AOT_modally_strict { + AOT_assume \[\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\\<^sub>1...\\<^sub>n\ + moreover AOT_have \[\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\\ and \\\<^sub>1...\\<^sub>n\\ + using calculation "cqt:5:a"[axiom_inst, THEN "\E"] "&E" by blast+ + ultimately AOT_show \\{\\<^sub>1...\\<^sub>n}\ + using "beta-C-cor:2"[THEN "\E", THEN "\E"(1), THEN "\E"(1)] by blast + } +qed + +AOT_theorem "betaC:1:b": \\\{\\<^sub>1...\\<^sub>n} \<^bold>\\<^sub>\ \[\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\\<^sub>1...\\<^sub>n\ + using "betaC:1:a" "raa-cor:3" by blast + +lemmas "\\C" = "betaC:1:a" "betaC:1:b" + +AOT_theorem "betaC:2:a": + \[\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\, \\<^sub>1...\\<^sub>n\, \{\\<^sub>1...\\<^sub>n} \<^bold>\\<^sub>\ + [\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\\<^sub>1...\\<^sub>n\ +proof - + AOT_modally_strict { + AOT_assume 1: \[\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\\ + and 2: \\\<^sub>1...\\<^sub>n\\ + and 3: \\{\\<^sub>1...\\<^sub>n}\ + AOT_hence \[\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\\<^sub>1...\\<^sub>n\ + using "beta-C-cor:2"[THEN "\E", OF 1, THEN "\E"(1), THEN "\E"(2)] + by blast + } + AOT_thus \[\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\, \\<^sub>1...\\<^sub>n\, \{\\<^sub>1...\\<^sub>n} \<^bold>\\<^sub>\ + [\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\\<^sub>1...\\<^sub>n\ + by blast +qed + +AOT_theorem "betaC:2:b": + \[\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\, \\<^sub>1...\\<^sub>n\, \[\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\\<^sub>1...\\<^sub>n \<^bold>\\<^sub>\ + \\{\\<^sub>1...\\<^sub>n}\ + using "betaC:2:a" "raa-cor:3" by blast + +lemmas "\\C" = "betaC:2:a" "betaC:2:b" + +AOT_theorem "eta-conversion-lemma1:1": \\\ \ [\x\<^sub>1...x\<^sub>n [\]x\<^sub>1...x\<^sub>n] = \\ + using "lambda-predicates:3"[axiom_inst] "\I" "\E"(1) "\I" by fast + +(* Note: generalized alphabetic variant of the last theorem *) +AOT_theorem "eta-conversion-lemma1:2": \\\ \ [\\\<^sub>1...\\<^sub>n [\]\\<^sub>1...\\<^sub>n] = \\ + using "eta-conversion-lemma1:1". + +text\Note: not explicitly part of PLM.\ +AOT_theorem id_sym: + assumes \\ = \'\ + shows \\' = \\ + using "rule=E"[where \="\ \' . \\' = \\", rotated 1, OF assms] + "=I"(1)[OF "t=t-proper:1"[THEN "\E", OF assms]] by auto +declare id_sym[sym] + +text\Note: not explicitly part of PLM.\ +AOT_theorem id_trans: + assumes \\ = \'\ and \\' = \''\ + shows \\ = \''\ + using "rule=E" assms by blast +declare id_trans[trans] + +method "\C" for \ :: \<'a::{AOT_Term_id_2,AOT_\s}>\ = + (match conclusion in "[v \ \{\} = \'{\}]" for v \ \' \ \ + rule "rule=E"[rotated 1, OF "eta-conversion-lemma1:2" + [THEN "\E", of v "\[\]\", symmetric]]\) + +AOT_theorem "sub-des-lam:1": + \[\z\<^sub>1...z\<^sub>n \{z\<^sub>1...z\<^sub>n, \<^bold>\x \{x}}]\ & \<^bold>\x \{x} = \<^bold>\x \{x} \ + [\z\<^sub>1...z\<^sub>n \{z\<^sub>1...z\<^sub>n, \<^bold>\x \{x}}] = [\z\<^sub>1...z\<^sub>n \{z\<^sub>1...z\<^sub>n, \<^bold>\x \{x}}]\ +proof(rule "\I") + AOT_assume A: \[\z\<^sub>1...z\<^sub>n \{z\<^sub>1...z\<^sub>n, \<^bold>\x \{x}}]\ & \<^bold>\x \{x} = \<^bold>\x \{x}\ + AOT_show \[\z\<^sub>1...z\<^sub>n \{z\<^sub>1...z\<^sub>n, \<^bold>\x \{x}}] = [\z\<^sub>1...z\<^sub>n \{z\<^sub>1...z\<^sub>n, \<^bold>\x \{x}}]\ + using "rule=E"[where \="\ \ . \[\z\<^sub>1...z\<^sub>n \{z\<^sub>1...z\<^sub>n, \<^bold>\x \{x}}] = + [\z\<^sub>1...z\<^sub>n \{z\<^sub>1...z\<^sub>n, \}]\", + OF "=I"(1)[OF A[THEN "&E"(1)]], OF A[THEN "&E"(2)]] + by blast +qed + +AOT_theorem "sub-des-lam:2": + \\<^bold>\x \{x} = \<^bold>\x \{x} \ \{\<^bold>\x \{x}} = \{\<^bold>\x \{x}}\ for \ :: \\ \ \\ + using "rule=E"[where \="\ \ . \\{\<^bold>\x \{x}} = \{\}\", + OF "=I"(1)[OF "log-prop-prop:2"]] "\I" by blast + +AOT_theorem "prop-equiv": \F = G \ \x (x[F] \ x[G])\ +proof(rule "\I"; rule "\I") + AOT_assume \F = G\ + AOT_thus \\x (x[F] \ x[G])\ + by (rule "rule=E"[rotated]) (fact "oth-class-taut:3:a"[THEN GEN]) +next + AOT_assume \\x (x[F] \ x[G])\ + AOT_hence \x[F] \ x[G]\ for x + using "\E" by blast + AOT_hence \\(x[F] \ x[G])\ for x + using "en-eq:6[1]"[THEN "\E"(1)] by blast + AOT_hence \\x \(x[F] \ x[G])\ + by (rule GEN) + AOT_hence \\\x (x[F] \ x[G])\ + using BF[THEN "\E"] by fast + AOT_thus "F = G" + using "p-identity-thm2:1"[THEN "\E"(2)] by blast +qed + +AOT_theorem "relations:1": + assumes \INSTANCE_OF_CQT_2(\)\ + shows \\F \\x\<^sub>1...\x\<^sub>n ([F]x\<^sub>1...x\<^sub>n \ \{x\<^sub>1...x\<^sub>n})\ + apply (rule "\I"(1)[where \="\[\x\<^sub>1...x\<^sub>n \{x\<^sub>1...x\<^sub>n}]\"]) + using "cqt:2[lambda]"[OF assms, axiom_inst] + "beta-C-cor:2"[THEN "\E", THEN RN] by blast+ + +AOT_theorem "relations:2": + assumes \INSTANCE_OF_CQT_2(\)\ + shows \\F \\x ([F]x \ \{x})\ + using "relations:1" assms by blast + +AOT_theorem "block-paradox:1": \\[\x \G (x[G] & \[G]x)]\\ +proof(rule "raa-cor:2") + let ?K="\[\x \G (x[G] & \[G]x)]\" + AOT_assume A: \\?K\\\ + AOT_have \\x (A!x & \F (x[F] \ F = \?K\))\ + using "A-objects"[axiom_inst] by fast + then AOT_obtain a where \: \A!a & \F (a[F] \ F = \?K\)\ + using "\E"[rotated] by blast + AOT_show \p & \p\ for p + proof (rule "\E"(1)[OF "exc-mid"]; rule "\I") + AOT_assume B: \[\?K\]a\ + AOT_hence \\G (a[G] & \[G]a)\ + using "\\C" A by blast + then AOT_obtain P where \a[P] & \[P]a\ + using "\E"[rotated] by blast + moreover AOT_have \P = [\?K\]\ + using \[THEN "&E"(2), THEN "\E"(2), THEN "\E"(1)] + calculation[THEN "&E"(1)] by blast + ultimately AOT_have \\[\?K\]a\ + using "rule=E" "&E"(2) by fast + AOT_thus \p & \p\ + using B RAA by blast + next + AOT_assume B: \\[\?K\]a\ + AOT_hence \\\G (a[G] & \[G]a)\ + using "\\C" "cqt:2[const_var]"[of a, axiom_inst] A by blast + AOT_hence C: \\G \(a[G] & \[G]a)\ + using "cqt-further:4"[THEN "\E"] by blast + AOT_have \\G (a[G] \ [G]a)\ + by (AOT_subst \a[G] \ [G]a\ \\(a[G] & \[G]a)\ for: G) + (auto simp: "oth-class-taut:1:a" C) + AOT_hence \a[\?K\] \ [\?K\]a\ + using "\E" A by blast + moreover AOT_have \a[\?K\]\ + using \[THEN "&E"(2), THEN "\E"(1), OF A, THEN "\E"(2)] + using "=I"(1)[OF A] by blast + ultimately AOT_show \p & \p\ + using B "\E" RAA by blast + qed +qed + +AOT_theorem "block-paradox:2": \\\F \x([F]x \ \G(x[G] & \[G]x))\ +proof(rule RAA(2)) + AOT_assume \\F \x ([F]x \ \G (x[G] & \[G]x))\ + then AOT_obtain F where F_prop: \\x ([F]x \ \G (x[G] & \[G]x))\ + using "\E"[rotated] by blast + AOT_have \\x (A!x & \G (x[G] \ G = F))\ + using "A-objects"[axiom_inst] by fast + then AOT_obtain a where \: \A!a & \G (a[G] \ G = F)\ + using "\E"[rotated] by blast + AOT_show \\\F \x([F]x \ \G(x[G] & \[G]x))\ + proof (rule "\E"(1)[OF "exc-mid"]; rule "\I") + AOT_assume B: \[F]a\ + AOT_hence \\G (a[G] & \[G]a)\ + using F_prop[THEN "\E"(2), THEN "\E"(1)] by blast + then AOT_obtain P where \a[P] & \[P]a\ + using "\E"[rotated] by blast + moreover AOT_have \P = F\ + using \[THEN "&E"(2), THEN "\E"(2), THEN "\E"(1)] + calculation[THEN "&E"(1)] by blast + ultimately AOT_have \\[F]a\ + using "rule=E" "&E"(2) by fast + AOT_thus \\\F \x([F]x \ \G(x[G] & \[G]x))\ + using B RAA by blast + next + AOT_assume B: \\[F]a\ + AOT_hence \\\G (a[G] & \[G]a)\ + using "oth-class-taut:4:b"[THEN "\E"(1), + OF F_prop[THEN "\E"(2)[of _ _ a]], THEN "\E"(1)] + by simp + AOT_hence C: \\G \(a[G] & \[G]a)\ + using "cqt-further:4"[THEN "\E"] by blast + AOT_have \\G (a[G] \ [G]a)\ + by (AOT_subst \a[G] \ [G]a\ \\(a[G] & \[G]a)\ for: G) + (auto simp: "oth-class-taut:1:a" C) + AOT_hence \a[F] \ [F]a\ + using "\E" by blast + moreover AOT_have \a[F]\ + using \[THEN "&E"(2), THEN "\E"(2), of F, THEN "\E"(2)] + using "=I"(2) by blast + ultimately AOT_show \\\F \x([F]x \ \G(x[G] & \[G]x))\ + using B "\E" RAA by blast + qed +qed(simp) + +AOT_theorem "block-paradox:3": \\\y [\z z = y]\\ +proof(rule RAA(2)) + AOT_assume \: \\y [\z z = y]\\ + AOT_have \\x (A!x & \F (x[F] \ \y(F = [\z z = y] & \y[F])))\ + using "A-objects"[axiom_inst] by force + then AOT_obtain a where + a_prop: \A!a & \F (a[F] \ \y (F = [\z z = y] & \y[F]))\ + using "\E"[rotated] by blast + AOT_have \: \a[\z z = a] \ \y ([\z z = a] = [\z z = y] & \y[\z z = a])\ + using \[THEN "\E"(2)] a_prop[THEN "&E"(2), THEN "\E"(1)] by blast + AOT_show \\\y [\z z = y]\\ + proof (rule "\E"(1)[OF "exc-mid"]; rule "\I") + AOT_assume A: \a[\z z = a]\ + AOT_hence \\y ([\z z = a] = [\z z = y] & \y[\z z = a])\ + using \[THEN "\E"(1)] by blast + then AOT_obtain b where b_prop: \[\z z = a] = [\z z = b] & \b[\z z = a]\ + using "\E"[rotated] by blast + moreover AOT_have \a = a\ by (rule "=I") + moreover AOT_have \[\z z = a]\\ using \ "\E" by blast + moreover AOT_have \a\\ using "cqt:2[const_var]"[axiom_inst] . + ultimately AOT_have \[\z z = a]a\ using "\\C" by blast + AOT_hence \[\z z = b]a\ using "rule=E" b_prop[THEN "&E"(1)] by fast + AOT_hence \a = b\ using "\\C" by blast + AOT_hence \b[\z z = a]\ using A "rule=E" by fast + AOT_thus \\\y [\z z = y]\\ using b_prop[THEN "&E"(2)] RAA by blast + next + AOT_assume A: \\a[\z z = a]\ + AOT_hence \\\y ([\z z = a] = [\z z = y] & \y[\z z = a])\ + using \ "oth-class-taut:4:b"[THEN "\E"(1), THEN "\E"(1)] by blast + AOT_hence \\y \([\z z = a] = [\z z = y] & \y[\z z = a])\ + using "cqt-further:4"[THEN "\E"] by blast + AOT_hence \\([\z z = a] = [\z z = a] & \a[\z z = a])\ + using "\E" by blast + AOT_hence \[\z z = a] = [\z z = a] \ a[\z z = a]\ + by (metis "&I" "deduction-theorem" "raa-cor:4") + AOT_hence \a[\z z = a]\ using "=I"(1) \[THEN "\E"(2)] "\E" by blast + AOT_thus \\\y [\z z = y]\\ using A RAA by blast + qed +qed(simp) + +AOT_theorem "block-paradox:4": \\\y \F \x([F]x \ x = y)\ +proof(rule RAA(2)) + AOT_assume \: \\y \F \x([F]x \ x = y)\ + AOT_have \\x (A!x & \F (x[F] \ \z (\y([F]y \ y = z) & \z[F])))\ + using "A-objects"[axiom_inst] by force + then AOT_obtain a where + a_prop: \A!a & \F (a[F] \ \z (\y([F]y \ y = z) & \z[F]))\ + using "\E"[rotated] by blast + AOT_obtain F where F_prop: \\x ([F]x \ x = a)\ + using \[THEN "\E"(2)] "\E"[rotated] by blast + AOT_have \: \a[F] \ \z (\y ([F]y \ y = z) & \z[F])\ + using a_prop[THEN "&E"(2), THEN "\E"(2)] by blast + AOT_show \\\y \F \x([F]x \ x = y)\ + proof (rule "\E"(1)[OF "exc-mid"]; rule "\I") + AOT_assume A: \a[F]\ + AOT_hence \\z (\y ([F]y \ y = z) & \z[F])\ + using \[THEN "\E"(1)] by blast + then AOT_obtain b where b_prop: \\y ([F]y \ y = b) & \b[F]\ + using "\E"[rotated] by blast + moreover AOT_have \[F]a\ + using F_prop[THEN "\E"(2), THEN "\E"(2)] "=I"(2) by blast + ultimately AOT_have \a = b\ + using "\E"(2) "\E"(1) "&E" by fast + AOT_hence \a = b\ + using "\\C" by blast + AOT_hence \b[F]\ + using A "rule=E" by fast + AOT_thus \\\y \F \x([F]x \ x = y)\ + using b_prop[THEN "&E"(2)] RAA by blast + next + AOT_assume A: \\a[F]\ + AOT_hence \\\z (\y ([F]y \ y = z) & \z[F])\ + using \ "oth-class-taut:4:b"[THEN "\E"(1), THEN "\E"(1)] by blast + AOT_hence \\z \(\y ([F]y \ y = z) & \z[F])\ + using "cqt-further:4"[THEN "\E"] by blast + AOT_hence \\(\y ([F]y \ y = a) & \a[F])\ + using "\E" by blast + AOT_hence \\y ([F]y \ y = a) \ a[F]\ + by (metis "&I" "deduction-theorem" "raa-cor:4") + AOT_hence \a[F]\ using F_prop "\E" by blast + AOT_thus \\\y \F \x([F]x \ x = y)\ + using A RAA by blast + qed +qed(simp) + +AOT_theorem "block-paradox:5": \\\F\x\y([F]xy \ y = x)\ +proof(rule "raa-cor:2") + AOT_assume \\F\x\y([F]xy \ y = x)\ + then AOT_obtain F where F_prop: \\x\y([F]xy \ y = x)\ + using "\E"[rotated] by blast + { + fix x + AOT_have 1: \\y([F]xy \ y = x)\ + using F_prop "\E" by blast + AOT_have 2: \[\z [F]xz]\\ by "cqt:2" + moreover AOT_have \\y([\z [F]xz]y \ y = x)\ + proof(rule "\I") + fix y + AOT_have \[\z [F]xz]y \ [F]xy\ + using "beta-C-meta"[THEN "\E"] 2 by fast + also AOT_have \... \ y = x\ + using 1 "\E" by fast + finally AOT_show \[\z [F]xz]y \ y = x\. + qed + ultimately AOT_have \\F\y([F]y \ y = x)\ + using "\I" by fast + } + AOT_hence \\x\F\y([F]y \ y = x)\ + by (rule GEN) + AOT_thus \\x\F\y([F]y \ y = x) & \\x\F\y([F]y \ y = x)\ + using "&I" "block-paradox:4" by blast +qed + +AOT_act_theorem "block-paradox2:1": + \\x [G]x \ \[\x [G]\<^bold>\y (y = x & \H (x[H] & \[H]x))]\\ +proof(rule "\I"; rule "raa-cor:2") + AOT_assume antecedant: \\x [G]x\ + AOT_have Lemma: \\x ([G]\<^bold>\y(y = x & \H (x[H] & \[H]x)) \ \H (x[H] & \[H]x))\ + proof(rule GEN) + fix x + AOT_have A: \[G]\<^bold>\y (y = x & \H (x[H] & \[H]x)) \ + \!y (y = x & \H (x[H] & \[H]x))\ + proof(rule "\I"; rule "\I") + AOT_assume \[G]\<^bold>\y (y = x & \H (x[H] & \[H]x))\ + AOT_hence \\<^bold>\y (y = x & \H (x[H] & \[H]x))\\ + using "cqt:5:a"[axiom_inst, THEN "\E", THEN "&E"(2)] by blast + AOT_thus \\!y (y = x & \H (x[H] & \[H]x))\ + using "!-exists:1"[THEN "\E"(1)] by blast + next + AOT_assume A: \\!y (y = x & \H (x[H] & \[H]x))\ + AOT_obtain a where a_1: \a = x & \H (x[H] & \[H]x)\ + and a_2: \\z (z = x & \H (x[H] & \[H]x) \ z = a)\ + using "uniqueness:1"[THEN "\\<^sub>d\<^sub>fE", OF A] "&E" "\E"[rotated] by blast + AOT_have a_3: \[G]a\ + using antecedant "\E" by blast + AOT_show \[G]\<^bold>\y (y = x & \H (x[H] & \[H]x))\ + apply (rule "russell-axiom[exe,1].russell-axiom"[THEN "\E"(2)]) + apply (rule "\I"(2)) + using a_1 a_2 a_3 "&I" by blast + qed + also AOT_have B: \... \ \H (x[H] & \[H]x)\ + proof (rule "\I"; rule "\I") + AOT_assume A: \\!y (y = x & \H (x[H] & \[H]x))\ + AOT_obtain a where \a = x & \H (x[H] & \[H]x)\ + using "uniqueness:1"[THEN "\\<^sub>d\<^sub>fE", OF A] "&E" "\E"[rotated] by blast + AOT_thus \\H (x[H] & \[H]x)\ using "&E" by blast + next + AOT_assume \\H (x[H] & \[H]x)\ + AOT_hence \x = x & \H (x[H] & \[H]x)\ + using "id-eq:1" "&I" by blast + moreover AOT_have \\z (z = x & \H (x[H] & \[H]x) \ z = x)\ + by (simp add: "Conjunction Simplification"(1) "universal-cor") + ultimately AOT_show \\!y (y = x & \H (x[H] & \[H]x))\ + using "uniqueness:1"[THEN "\\<^sub>d\<^sub>fI"] "&I" "\I"(2) by fast + qed + finally AOT_show \([G]\<^bold>\y(y = x & \H (x[H] & \[H]x)) \ \H (x[H] & \[H]x))\ . + qed + + AOT_assume A: \[\x [G]\<^bold>\y (y = x & \H (x[H] & \[H]x))]\\ + AOT_have \: \\x ([\x [G]\<^bold>\y (y = x & \H (x[H] & \[H]x))]x \ + [G]\<^bold>\y(y = x & \H (x[H] & \[H]x)))\ + using "beta-C-meta"[THEN "\E", OF A] "\I" by fast + AOT_have \\x ([\x [G]\<^bold>\y (y = x & \H (x[H] & \[H]x))]x \ \H (x[H] & \[H]x))\ + using \ Lemma "cqt-basic:10"[THEN "\E"] "&I" by fast + AOT_hence \\F \x ([F]x \ \H (x[H] & \[H]x))\ + using "\I"(1) A by fast + AOT_thus \(\F \x ([F]x \ \H (x[H] & \[H]x))) & + (\\F \x ([F]x \ \H (x[H] & \[H]x)))\ + using "block-paradox:2" "&I" by blast +qed + +text\Note: Strengthens the above to a modally-strict theorem. + Not explicitly part of PLM.\ +AOT_theorem "block-paradox2:1[strict]": + \\x \<^bold>\[G]x \ \[\x [G]\<^bold>\y (y = x & \H (x[H] & \[H]x))]\\ +proof(rule "\I"; rule "raa-cor:2") + AOT_assume antecedant: \\x \<^bold>\[G]x\ + AOT_have Lemma: \\<^bold>\\x ([G]\<^bold>\y(y = x & \H (x[H] & \[H]x)) \ \H (x[H] & \[H]x))\ + proof(safe intro!: GEN "Act-Basic:5"[THEN "\E"(2)] + "logic-actual-nec:3"[axiom_inst, THEN "\E"(2)]) + fix x + AOT_have A: \\<^bold>\[G]\<^bold>\y (y = x & \H (x[H] & \[H]x)) \ + \!y \<^bold>\(y = x & \H (x[H] & \[H]x))\ + proof(rule "\I"; rule "\I") + AOT_assume \\<^bold>\[G]\<^bold>\y (y = x & \H (x[H] & \[H]x))\ + moreover AOT_have \\([G]\<^bold>\y (y = x & \H (x[H] & \[H]x)) \ + \\<^bold>\y (y = x & \H (x[H] & \[H]x))\)\ + proof(rule RN; rule "\I") + AOT_modally_strict { + AOT_assume \[G]\<^bold>\y (y = x & \H (x[H] & \[H]x))\ + AOT_hence \\<^bold>\y (y = x & \H (x[H] & \[H]x))\\ + using "cqt:5:a"[axiom_inst, THEN "\E", THEN "&E"(2)] by blast + AOT_thus \\\<^bold>\y (y = x & \H (x[H] & \[H]x))\\ + using "exist-nec"[THEN "\E"] by blast + } + qed + ultimately AOT_have \\<^bold>\\\<^bold>\y (y = x & \H (x[H] & \[H]x))\\ + using "act-cond"[THEN "\E", THEN "\E"] "nec-imp-act"[THEN "\E"] by blast + AOT_hence \\<^bold>\y (y = x & \H (x[H] & \[H]x))\\ + using "Act-Sub:3" "B\" "vdash-properties:10" by blast + AOT_thus \\!y \<^bold>\(y = x & \H (x[H] & \[H]x))\ + using "actual-desc:1"[THEN "\E"(1)] by blast + next + AOT_assume A: \\!y \<^bold>\(y = x & \H (x[H] & \[H]x))\ + AOT_obtain a where a_1: \\<^bold>\(a = x & \H (x[H] & \[H]x))\ + and a_2: \\z (\<^bold>\(z = x & \H (x[H] & \[H]x)) \ z = a)\ + using "uniqueness:1"[THEN "\\<^sub>d\<^sub>fE", OF A] "&E" "\E"[rotated] by blast + AOT_have a_3: \\<^bold>\[G]a\ + using antecedant "\E" by blast + moreover AOT_have \a = \<^bold>\y(y = x & \H (x[H] & \[H]x))\ + using "nec-hintikka-scheme"[THEN "\E"(2), OF "&I"] a_1 a_2 by auto + ultimately AOT_show \\<^bold>\[G]\<^bold>\y (y = x & \H (x[H] & \[H]x))\ + using "rule=E" by fast + qed + also AOT_have B: \... \ \<^bold>\\H (x[H] & \[H]x)\ + proof (rule "\I"; rule "\I") + AOT_assume A: \\!y \<^bold>\(y = x & \H (x[H] & \[H]x))\ + AOT_obtain a where \\<^bold>\(a = x & \H (x[H] & \[H]x))\ + using "uniqueness:1"[THEN "\\<^sub>d\<^sub>fE", OF A] "&E" "\E"[rotated] by blast + AOT_thus \\<^bold>\\H (x[H] & \[H]x)\ + using "Act-Basic:2"[THEN "\E"(1), THEN "&E"(2)] by blast + next + AOT_assume \\<^bold>\\H (x[H] & \[H]x)\ + AOT_hence \\<^bold>\x = x & \<^bold>\\H (x[H] & \[H]x)\ + using "id-eq:1" "&I" "RA[2]" by blast + AOT_hence \\<^bold>\(x = x & \H (x[H] & \[H]x))\ + using "act-conj-act:3" "Act-Basic:2" "\E" by blast + moreover AOT_have \\z (\<^bold>\(z = x & \H (x[H] & \[H]x)) \ z = x)\ + proof(safe intro!: GEN "\I") + fix z + AOT_assume \\<^bold>\(z = x & \H (x[H] & \[H]x))\ + AOT_hence \\<^bold>\(z = x)\ + using "Act-Basic:2"[THEN "\E"(1), THEN "&E"(1)] by blast + AOT_thus \z = x\ + by (metis "id-act:1" "intro-elim:3:b") + qed + ultimately AOT_show \\!y \<^bold>\(y = x & \H (x[H] & \[H]x))\ + using "uniqueness:1"[THEN "\\<^sub>d\<^sub>fI"] "&I" "\I"(2) by fast + qed + finally AOT_show \(\<^bold>\[G]\<^bold>\y(y = x & \H (x[H] & \[H]x)) \ \<^bold>\\H (x[H] & \[H]x))\. + qed + + AOT_assume A: \[\x [G]\<^bold>\y (y = x & \H (x[H] & \[H]x))]\\ + AOT_hence \\<^bold>\[\x [G]\<^bold>\y (y = x & \H (x[H] & \[H]x))]\\ + using "exist-nec" "\E" "nec-imp-act"[THEN "\E"] by blast + AOT_hence \\<^bold>\([\x [G]\<^bold>\y (y = x & \H (x[H] & \[H]x))]\ & + \x ([G]\<^bold>\y(y = x & \H (x[H] & \[H]x)) \ \H (x[H] & \[H]x)))\ + using Lemma "Act-Basic:2"[THEN "\E"(2)] "&I" by blast + moreover AOT_have \\<^bold>\([\x [G]\<^bold>\y (y = x & \H (x[H] & \[H]x))]\ & + \x ([G]\<^bold>\y(y = x & \H (x[H] & \[H]x)) \ \H (x[H] & \[H]x))) + \ \<^bold>\\p (p & \p)\ + proof (rule "logic-actual-nec:2"[axiom_inst, THEN "\E"(1)]; + rule "RA[2]"; rule "\I") + AOT_modally_strict { + AOT_assume 0: \[\x [G]\<^bold>\y (y = x & \H (x[H] & \[H]x))]\ & + \x ([G]\<^bold>\y(y = x & \H (x[H] & \[H]x)) \ \H (x[H] & \[H]x))\ + AOT_have \\F \x ([F]x \ \G (x[G] & \[G]x))\ + proof(rule "\I"(1)) + AOT_show \\x ([\x [G]\<^bold>\y (y = x & \H (x[H] & \[H]x))]x \ \H (x[H] & \[H]x))\ + proof(safe intro!: GEN "\I" "\I" "\\C" dest!: "\\C") + fix x + AOT_assume \[G]\<^bold>\y(y = x & \H (x[H] & \[H]x))\ + AOT_thus \\H (x[H] & \[H]x)\ + using 0 "&E" "\E"(2) "\E"(1) by blast + next + fix x + AOT_assume \\H (x[H] & \[H]x)\ + AOT_thus \[G]\<^bold>\y(y = x & \H (x[H] & \[H]x))\ + using 0 "&E" "\E"(2) "\E"(2) by blast + qed(auto intro!: 0[THEN "&E"(1)] "cqt:2") + next + AOT_show \[\x [G]\<^bold>\y (y = x & \H (x[H] & \[H]x))]\\ + using 0 "&E"(1) by blast + qed + AOT_thus \\p (p & \p)\ + using "block-paradox:2" "reductio-aa:1" by blast + } + qed + ultimately AOT_have \\<^bold>\\p (p & \p)\ + using "\E" by blast + AOT_hence \\p \<^bold>\(p & \p)\ + by (metis "Act-Basic:10" "intro-elim:3:a") + then AOT_obtain p where \\<^bold>\(p & \p)\ + using "\E"[rotated] by blast + moreover AOT_have \\\<^bold>\(p & \p)\ + using "non-contradiction"[THEN "RA[2]"] + by (meson "Act-Sub:1" "\\I" "intro-elim:3:d") + ultimately AOT_show \p & \p\ for p + by (metis "raa-cor:3") +qed + +AOT_act_theorem "block-paradox2:2": + \\G \[\x [G]\<^bold>\y (y = x & \H (x[H] & \[H]x))]\\ +proof(rule "\I"(1)) + AOT_have 0: \[\x \p (p \p)]\\ + by "cqt:2[lambda]" + moreover AOT_have \\x [\x \p (p \p)]x\ + apply (rule GEN) + apply (rule "beta-C-cor:2"[THEN "\E", OF 0, THEN "\E"(2), THEN "\E"(2)]) + using "if-p-then-p" GEN by fast + moreover AOT_have \\G (\x [G]x \ \[\x [G]\<^bold>\y (y = x & \H (x[H] & \[H]x))]\)\ + using "block-paradox2:1" "\I" by fast + ultimately AOT_show \\[\x [\x \p (p \p)]\<^bold>\y (y = x & \H (x[H] & \[H]x))]\\ + using "\E"(1) "\E" by blast +qed("cqt:2[lambda]") + +AOT_theorem propositions: \\p \(p \ \)\ +proof(rule "\I"(1)) + AOT_show \\(\ \ \)\ + by (simp add: RN "oth-class-taut:3:a") +next + AOT_show \\\\ + by (simp add: "log-prop-prop:2") +qed + +AOT_theorem "pos-not-equiv-ne:1": + \(\\\x\<^sub>1...\x\<^sub>n ([F]x\<^sub>1...x\<^sub>n \ [G]x\<^sub>1...x\<^sub>n)) \ F \ G\ +proof (rule "\I") + AOT_assume \\\\x\<^sub>1...\x\<^sub>n ([F]x\<^sub>1...x\<^sub>n \ [G]x\<^sub>1...x\<^sub>n)\ + AOT_hence \\\\x\<^sub>1...\x\<^sub>n ([F]x\<^sub>1...x\<^sub>n \ [G]x\<^sub>1...x\<^sub>n)\ + using "KBasic:11"[THEN "\E"(2)] by blast + AOT_hence \\(F = G)\ + using "id-rel-nec-equiv:1" "modus-tollens:1" by blast + AOT_thus \F \ G\ + using "=-infix"[THEN "\\<^sub>d\<^sub>fI"] by blast +qed + +AOT_theorem "pos-not-equiv-ne:2": \(\\(\{F} \ \{G})) \ F \ G\ +proof (rule "\I") + AOT_modally_strict { + AOT_have \\(\{F} \ \{G}) \ \(F = G)\ + proof (rule "\I"; rule "raa-cor:2") + AOT_assume 1: \F = G\ + AOT_hence \\{F} \ \{G}\ + using "l-identity"[axiom_inst, THEN "\E"] by blast + moreover { + AOT_have \G = F\ + using 1 id_sym by blast + AOT_hence \\{G} \ \{F}\ + using "l-identity"[axiom_inst, THEN "\E"] by blast + } + ultimately AOT_have \\{F} \ \{G}\ + using "\I" by blast + moreover AOT_assume \\(\{F} \ \{G})\ + ultimately AOT_show \(\{F} \ \{G}) & \(\{F} \ \{G})\ + using "&I" by blast + qed + } + AOT_hence \\\(\{F} \ \{G}) \ \\(F = G)\ + using "RM:2[prem]" by blast + moreover AOT_assume \\\(\{F} \ \{G})\ + ultimately AOT_have 0: \\\(F = G)\ using "\E" by blast + AOT_have \\(F \ G)\ + by (AOT_subst \F \ G\ \\(F = G)\) + (auto simp: "=-infix" "\Df" 0) + AOT_thus \F \ G\ + using "id-nec2:3"[THEN "\E"] by blast +qed + +AOT_theorem "pos-not-equiv-ne:2[zero]": \(\\(\{p} \ \{q})) \ p \ q\ +proof (rule "\I") + AOT_modally_strict { + AOT_have \\(\{p} \ \{q}) \ \(p = q)\ + proof (rule "\I"; rule "raa-cor:2") + AOT_assume 1: \p = q\ + AOT_hence \\{p} \ \{q}\ + using "l-identity"[axiom_inst, THEN "\E"] by blast + moreover { + AOT_have \q = p\ + using 1 id_sym by blast + AOT_hence \\{q} \ \{p}\ + using "l-identity"[axiom_inst, THEN "\E"] by blast + } + ultimately AOT_have \\{p} \ \{q}\ + using "\I" by blast + moreover AOT_assume \\(\{p} \ \{q})\ + ultimately AOT_show \(\{p} \ \{q}) & \(\{p} \ \{q})\ + using "&I" by blast + qed + } + AOT_hence \\\(\{p} \ \{q}) \ \\(p = q)\ + using "RM:2[prem]" by blast + moreover AOT_assume \\\(\{p} \ \{q})\ + ultimately AOT_have 0: \\\(p = q)\ using "\E" by blast + AOT_have \\(p \ q)\ + by (AOT_subst \p \ q\ \\(p = q)\) + (auto simp: 0 "=-infix" "\Df") + AOT_thus \p \ q\ + using "id-nec2:3"[THEN "\E"] by blast +qed + +AOT_theorem "pos-not-equiv-ne:3": + \(\\x\<^sub>1...\x\<^sub>n ([F]x\<^sub>1...x\<^sub>n \ [G]x\<^sub>1...x\<^sub>n)) \ F \ G\ + using "\I" "pos-not-equiv-ne:1"[THEN "\E"] "T\"[THEN "\E"] by blast + +AOT_theorem "pos-not-equiv-ne:4": \(\(\{F} \ \{G})) \ F \ G\ + using "\I" "pos-not-equiv-ne:2"[THEN "\E"] "T\"[THEN "\E"] by blast + +AOT_theorem "pos-not-equiv-ne:4[zero]": \(\(\{p} \ \{q})) \ p \ q\ + using "\I" "pos-not-equiv-ne:2[zero]"[THEN "\E"] + "T\"[THEN "\E"] by blast + +AOT_define relation_negation :: "\ \ \" ("_\<^sup>-") + "df-relation-negation": "[F]\<^sup>- =\<^sub>d\<^sub>f [\x\<^sub>1...x\<^sub>n \[F]x\<^sub>1...x\<^sub>n]" + +nonterminal \neg +syntax "" :: "\neg \ \" ("_") +syntax "" :: "\neg \ \" ("'(_')") + +AOT_define relation_negation_0 :: \\ \ \neg\ ("'(_')\<^sup>-") + "df-relation-negation[zero]": "(p)\<^sup>- =\<^sub>d\<^sub>f [\ \p]" + +AOT_theorem "rel-neg-T:1": \[\x\<^sub>1...x\<^sub>n \[\]x\<^sub>1...x\<^sub>n]\\ + by "cqt:2[lambda]" + +AOT_theorem "rel-neg-T:1[zero]": \[\ \\]\\ + using "cqt:2[lambda0]"[axiom_inst] by blast + +AOT_theorem "rel-neg-T:2": \[\]\<^sup>- = [\x\<^sub>1...x\<^sub>n \[\]x\<^sub>1...x\<^sub>n]\ + using "=I"(1)[OF "rel-neg-T:1"] + by (rule "=\<^sub>d\<^sub>fI"(1)[OF "df-relation-negation", OF "rel-neg-T:1"]) + +AOT_theorem "rel-neg-T:2[zero]": \(\)\<^sup>- = [\ \\]\ + using "=I"(1)[OF "rel-neg-T:1[zero]"] + by (rule "=\<^sub>d\<^sub>fI"(1)[OF "df-relation-negation[zero]", OF "rel-neg-T:1[zero]"]) + +AOT_theorem "rel-neg-T:3": \[\]\<^sup>-\\ + using "=\<^sub>d\<^sub>fI"(1)[OF "df-relation-negation", OF "rel-neg-T:1"] + "rel-neg-T:1" by blast + +AOT_theorem "rel-neg-T:3[zero]": \(\)\<^sup>-\\ + using "log-prop-prop:2" by blast + +AOT_theorem "thm-relation-negation:1": \[F]\<^sup>-x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n\ +proof - + AOT_have \[F]\<^sup>-x\<^sub>1...x\<^sub>n \ [\x\<^sub>1...x\<^sub>n \[F]x\<^sub>1...x\<^sub>n]x\<^sub>1...x\<^sub>n\ + using "rule=E"[rotated, OF "rel-neg-T:2"] + "rule=E"[rotated, OF "rel-neg-T:2"[THEN id_sym]] + "\I" "\I" by fast + also AOT_have \... \ \[F]x\<^sub>1...x\<^sub>n\ + using "beta-C-meta"[THEN "\E", OF "rel-neg-T:1"] by fast + finally show ?thesis. +qed + +AOT_theorem "thm-relation-negation:2": \\[F]\<^sup>-x\<^sub>1...x\<^sub>n \ [F]x\<^sub>1...x\<^sub>n\ + apply (AOT_subst \[F]x\<^sub>1...x\<^sub>n\ \\\[F]x\<^sub>1...x\<^sub>n\) + apply (simp add: "oth-class-taut:3:b") + apply (rule "oth-class-taut:4:b"[THEN "\E"(1)]) + using "thm-relation-negation:1". + +AOT_theorem "thm-relation-negation:3": \((p)\<^sup>-) \ \p\ +proof - + AOT_have \(p)\<^sup>- = [\ \p]\ using "rel-neg-T:2[zero]" by blast + AOT_hence \((p)\<^sup>-) \ [\ \p]\ + using "df-relation-negation[zero]" "log-prop-prop:2" + "oth-class-taut:3:a" "rule-id-df:2:a" by blast + also AOT_have \[\ \p] \ \p\ + by (simp add: "propositions-lemma:2") + finally show ?thesis. +qed + +AOT_theorem "thm-relation-negation:4": \(\((p)\<^sup>-)) \ p\ + using "thm-relation-negation:3"[THEN "\E"(1)] + "thm-relation-negation:3"[THEN "\E"(2)] + "\I" "\I" RAA by metis + +AOT_theorem "thm-relation-negation:5": \[F] \ [F]\<^sup>-\ +proof - + AOT_have \\([F] = [F]\<^sup>-)\ + proof (rule RAA(2)) + AOT_show \[F]x\<^sub>1...x\<^sub>n \ [F]x\<^sub>1...x\<^sub>n\ for x\<^sub>1x\<^sub>n + using "if-p-then-p". + next + AOT_assume \[F] = [F]\<^sup>-\ + AOT_hence \[F]\<^sup>- = [F]\ using id_sym by blast + AOT_hence \[F]x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n\ for x\<^sub>1x\<^sub>n + using "rule=E" "thm-relation-negation:1" by fast + AOT_thus \\([F]x\<^sub>1...x\<^sub>n \ [F]x\<^sub>1...x\<^sub>n)\ for x\<^sub>1x\<^sub>n + using "\E" RAA by metis + qed + thus ?thesis + using "\\<^sub>d\<^sub>fI" "=-infix" by blast +qed + +AOT_theorem "thm-relation-negation:6": \p \ (p)\<^sup>-\ +proof - + AOT_have \\(p = (p)\<^sup>-)\ + proof (rule RAA(2)) + AOT_show \p \ p\ + using "if-p-then-p". + next + AOT_assume \p = (p)\<^sup>-\ + AOT_hence \(p)\<^sup>- = p\ using id_sym by blast + AOT_hence \p \ \p\ + using "rule=E" "thm-relation-negation:3" by fast + AOT_thus \\(p \ p)\ + using "\E" RAA by metis + qed + thus ?thesis + using "\\<^sub>d\<^sub>fI" "=-infix" by blast +qed + +AOT_theorem "thm-relation-negation:7": \(p)\<^sup>- = (\p)\ + apply (rule "df-relation-negation[zero]"[THEN "=\<^sub>d\<^sub>fE"(1)]) + using "cqt:2[lambda0]"[axiom_inst] "rel-neg-T:2[zero]" + "propositions-lemma:1" id_trans by blast+ + +AOT_theorem "thm-relation-negation:8": \p = q \ (\p) = (\q)\ +proof(rule "\I") + AOT_assume \p = q\ + moreover AOT_have \(\p)\\ using "log-prop-prop:2". + moreover AOT_have \(\p) = (\p)\ using calculation(2) "=I" by blast + ultimately AOT_show \(\p) = (\q)\ + using "rule=E" by fast +qed + +AOT_theorem "thm-relation-negation:9": \p = q \ (p)\<^sup>- = (q)\<^sup>-\ +proof(rule "\I") + AOT_assume \p = q\ + AOT_hence \(\p) = (\q)\ using "thm-relation-negation:8" "\E" by blast + AOT_thus \(p)\<^sup>- = (q)\<^sup>-\ + using "thm-relation-negation:7" id_sym id_trans by metis +qed + +AOT_define Necessary :: \\ \ \\ ("Necessary'(_')") + "contingent-properties:1": + \Necessary([F]) \\<^sub>d\<^sub>f \\x\<^sub>1...\x\<^sub>n [F]x\<^sub>1...x\<^sub>n\ + +AOT_define Necessary0 :: \\ \ \\ ("Necessary0'(_')") + "contingent-properties:1[zero]": + \Necessary0(p) \\<^sub>d\<^sub>f \p\ + +AOT_define Impossible :: \\ \ \\ ("Impossible'(_')") + "contingent-properties:2": + \Impossible([F]) \\<^sub>d\<^sub>f F\ & \\x\<^sub>1...\x\<^sub>n \[F]x\<^sub>1...x\<^sub>n\ + +AOT_define Impossible0 :: \\ \ \\ ("Impossible0'(_')") + "contingent-properties:2[zero]": + \Impossible0(p) \\<^sub>d\<^sub>f \\p\ + +AOT_define NonContingent :: \\ \ \\ ("NonContingent'(_')") + "contingent-properties:3": + \NonContingent([F]) \\<^sub>d\<^sub>f Necessary([F]) \ Impossible([F])\ + +AOT_define NonContingent0 :: \\ \ \\ ("NonContingent0'(_')") + "contingent-properties:3[zero]": + \NonContingent0(p) \\<^sub>d\<^sub>f Necessary0(p) \ Impossible0(p)\ + +AOT_define Contingent :: \\ \ \\ ("Contingent'(_')") + "contingent-properties:4": + \Contingent([F]) \\<^sub>d\<^sub>f F\ & \(Necessary([F]) \ Impossible([F]))\ + +AOT_define Contingent0 :: \\ \ \\ ("Contingent0'(_')") + "contingent-properties:4[zero]": + \Contingent0(p) \\<^sub>d\<^sub>f \(Necessary0(p) \ Impossible0(p))\ + + +AOT_theorem "thm-cont-prop:1": \NonContingent([F]) \ NonContingent([F]\<^sup>-)\ +proof (rule "\I"; rule "\I") + AOT_assume \NonContingent([F])\ + AOT_hence \Necessary([F]) \ Impossible([F])\ + using "\\<^sub>d\<^sub>fE"[OF "contingent-properties:3"] by blast + moreover { + AOT_assume \Necessary([F])\ + AOT_hence \\(\x\<^sub>1...\x\<^sub>n [F]x\<^sub>1...x\<^sub>n)\ + using "\\<^sub>d\<^sub>fE"[OF "contingent-properties:1"] by blast + moreover AOT_modally_strict { + AOT_assume \\x\<^sub>1...\x\<^sub>n [F]x\<^sub>1...x\<^sub>n\ + AOT_hence \[F]x\<^sub>1...x\<^sub>n\ for x\<^sub>1x\<^sub>n using "\E" by blast + AOT_hence \\[F]\<^sup>-x\<^sub>1...x\<^sub>n\ for x\<^sub>1x\<^sub>n + by (meson "\E"(6) "oth-class-taut:3:a" + "thm-relation-negation:2" "\E"(1)) + AOT_hence \\x\<^sub>1...\x\<^sub>n \[F]\<^sup>-x\<^sub>1...x\<^sub>n\ using "\I" by fast + } + ultimately AOT_have \\(\x\<^sub>1...\x\<^sub>n \[F]\<^sup>-x\<^sub>1...x\<^sub>n)\ + using "RN[prem]"[where \="{\\x\<^sub>1...\x\<^sub>n [F]x\<^sub>1...x\<^sub>n\}", simplified] by blast + AOT_hence \Impossible([F]\<^sup>-)\ + using "\Df"[OF "contingent-properties:2", THEN "\S"(1), + OF "rel-neg-T:3", THEN "\E"(2)] + by blast + } + moreover { + AOT_assume \Impossible([F])\ + AOT_hence \\(\x\<^sub>1...\x\<^sub>n \[F]x\<^sub>1...x\<^sub>n)\ + using "\Df"[OF "contingent-properties:2", THEN "\S"(1), + OF "cqt:2[const_var]"[axiom_inst], THEN "\E"(1)] + by blast + moreover AOT_modally_strict { + AOT_assume \\x\<^sub>1...\x\<^sub>n \[F]x\<^sub>1...x\<^sub>n\ + AOT_hence \\[F]x\<^sub>1...x\<^sub>n\ for x\<^sub>1x\<^sub>n using "\E" by blast + AOT_hence \[F]\<^sup>-x\<^sub>1...x\<^sub>n\ for x\<^sub>1x\<^sub>n + by (meson "\E"(6) "oth-class-taut:3:a" + "thm-relation-negation:1" "\E"(1)) + AOT_hence \\x\<^sub>1...\x\<^sub>n [F]\<^sup>-x\<^sub>1...x\<^sub>n\ using "\I" by fast + } + ultimately AOT_have \\(\x\<^sub>1...\x\<^sub>n [F]\<^sup>-x\<^sub>1...x\<^sub>n)\ + using "RN[prem]"[where \="{\\x\<^sub>1...\x\<^sub>n \[F]x\<^sub>1...x\<^sub>n\}"] by blast + AOT_hence \Necessary([F]\<^sup>-)\ + using "\\<^sub>d\<^sub>fI"[OF "contingent-properties:1"] by blast + } + ultimately AOT_have \Necessary([F]\<^sup>-) \ Impossible([F]\<^sup>-)\ + using "\E"(1) "\I" "\I" by metis + AOT_thus \NonContingent([F]\<^sup>-)\ + using "\\<^sub>d\<^sub>fI"[OF "contingent-properties:3"] by blast +next + AOT_assume \NonContingent([F]\<^sup>-)\ + AOT_hence \Necessary([F]\<^sup>-) \ Impossible([F]\<^sup>-)\ + using "\\<^sub>d\<^sub>fE"[OF "contingent-properties:3"] by blast + moreover { + AOT_assume \Necessary([F]\<^sup>-)\ + AOT_hence \\(\x\<^sub>1...\x\<^sub>n [F]\<^sup>-x\<^sub>1...x\<^sub>n)\ + using "\\<^sub>d\<^sub>fE"[OF "contingent-properties:1"] by blast + moreover AOT_modally_strict { + AOT_assume \\x\<^sub>1...\x\<^sub>n [F]\<^sup>-x\<^sub>1...x\<^sub>n\ + AOT_hence \[F]\<^sup>-x\<^sub>1...x\<^sub>n\ for x\<^sub>1x\<^sub>n using "\E" by blast + AOT_hence \\[F]x\<^sub>1...x\<^sub>n\ for x\<^sub>1x\<^sub>n + by (meson "\E"(6) "oth-class-taut:3:a" + "thm-relation-negation:1" "\E"(2)) + AOT_hence \\x\<^sub>1...\x\<^sub>n \[F]x\<^sub>1...x\<^sub>n\ using "\I" by fast + } + ultimately AOT_have \\\x\<^sub>1...\x\<^sub>n \[F]x\<^sub>1...x\<^sub>n\ + using "RN[prem]"[where \="{\\x\<^sub>1...\x\<^sub>n [F]\<^sup>-x\<^sub>1...x\<^sub>n\}"] by blast + AOT_hence \Impossible([F])\ + using "\Df"[OF "contingent-properties:2", THEN "\S"(1), + OF "cqt:2[const_var]"[axiom_inst], THEN "\E"(2)] + by blast + } + moreover { + AOT_assume \Impossible([F]\<^sup>-)\ + AOT_hence \\(\x\<^sub>1...\x\<^sub>n \[F]\<^sup>-x\<^sub>1...x\<^sub>n)\ + using "\Df"[OF "contingent-properties:2", THEN "\S"(1), + OF "rel-neg-T:3", THEN "\E"(1)] + by blast + moreover AOT_modally_strict { + AOT_assume \\x\<^sub>1...\x\<^sub>n \[F]\<^sup>-x\<^sub>1...x\<^sub>n\ + AOT_hence \\[F]\<^sup>-x\<^sub>1...x\<^sub>n\ for x\<^sub>1x\<^sub>n using "\E" by blast + AOT_hence \[F]x\<^sub>1...x\<^sub>n\ for x\<^sub>1x\<^sub>n + using "thm-relation-negation:1"[THEN + "oth-class-taut:4:b"[THEN "\E"(1)], THEN "\E"(1)] + "useful-tautologies:1"[THEN "\E"] by blast + AOT_hence \\x\<^sub>1...\x\<^sub>n [F]x\<^sub>1...x\<^sub>n\ using "\I" by fast + } + ultimately AOT_have \\(\x\<^sub>1...\x\<^sub>n [F]x\<^sub>1...x\<^sub>n)\ + using "RN[prem]"[where \="{\\x\<^sub>1...\x\<^sub>n \[F]\<^sup>-x\<^sub>1...x\<^sub>n\}"] by blast + AOT_hence \Necessary([F])\ + using "\\<^sub>d\<^sub>fI"[OF "contingent-properties:1"] by blast + } + ultimately AOT_have \Necessary([F]) \ Impossible([F])\ + using "\E"(1) "\I" "\I" by metis + AOT_thus \NonContingent([F])\ + using "\\<^sub>d\<^sub>fI"[OF "contingent-properties:3"] by blast +qed + +AOT_theorem "thm-cont-prop:2": \Contingent([F]) \ \\x [F]x & \\x \[F]x\ +proof - + AOT_have \Contingent([F]) \ \(Necessary([F]) \ Impossible([F]))\ + using "contingent-properties:4"[THEN "\Df", THEN "\S"(1), + OF "cqt:2[const_var]"[axiom_inst]] + by blast + also AOT_have \... \ \Necessary([F]) & \Impossible([F])\ + using "oth-class-taut:5:d" by fastforce + also AOT_have \... \ \Impossible([F]) & \Necessary([F])\ + by (simp add: "Commutativity of &") + also AOT_have \... \ \\x [F]x & \Necessary([F])\ + proof (rule "oth-class-taut:4:e"[THEN "\E"]) + AOT_have \\Impossible([F]) \ \\\ \x [F]x\ + apply (rule "oth-class-taut:4:b"[THEN "\E"(1)]) + apply (AOT_subst \\x [F]x\ \\ \x \[F]x\) + apply (simp add: "conventions:4" "\Df") + apply (AOT_subst (reverse) \\\\x \[F]x\ \\x \[F]x\) + apply (simp add: "oth-class-taut:3:b") + using "contingent-properties:2"[THEN "\Df", THEN "\S"(1), + OF "cqt:2[const_var]"[axiom_inst]] + by blast + also AOT_have \... \ \\x [F]x\ + using "conventions:5"[THEN "\Df", symmetric] by blast + finally AOT_show \\Impossible([F]) \ \\x [F]x\ . + qed + also AOT_have \... \ \\x [F]x & \\x \[F]x\ + proof (rule "oth-class-taut:4:f"[THEN "\E"]) + AOT_have \\Necessary([F]) \ \\\\x \[F]x\ + apply (rule "oth-class-taut:4:b"[THEN "\E"(1)]) + apply (AOT_subst \\x \[F]x\ \\ \x \\[F]x\) + apply (simp add: "conventions:4" "\Df") + apply (AOT_subst (reverse) \\\[F]x\ \[F]x\ for: x) + apply (simp add: "oth-class-taut:3:b") + apply (AOT_subst (reverse) \\\\x [F]x\ \\x [F]x\) + by (auto simp: "oth-class-taut:3:b" "contingent-properties:1" "\Df") + also AOT_have \... \ \\x \[F]x\ + using "conventions:5"[THEN "\Df", symmetric] by blast + finally AOT_show \\Necessary([F]) \ \\x \[F]x\. + qed + finally show ?thesis. +qed + +AOT_theorem "thm-cont-prop:3": + \Contingent([F]) \ Contingent([F]\<^sup>-)\ for F::\<\> AOT_var\ +proof - + { + fix \ :: \<\>\ + AOT_assume \\\\ + moreover AOT_have \\F (Contingent([F]) \ \\x [F]x & \\x \[F]x)\ + using "thm-cont-prop:2" GEN by fast + ultimately AOT_have \Contingent([\]) \ \\x [\]x & \\x \[\]x\ + using "thm-cont-prop:2" "\E" by fast + } note 1 = this + AOT_have \Contingent([F]) \ \\x [F]x & \\x \[F]x\ + using "thm-cont-prop:2" by blast + also AOT_have \... \ \\x \[F]x & \\x [F]x\ + by (simp add: "Commutativity of &") + also AOT_have \... \ \\x [F]\<^sup>-x & \\x [F]x\ + by (AOT_subst \[F]\<^sup>-x\ \\[F]x\ for: x) + (auto simp: "thm-relation-negation:1" "oth-class-taut:3:a") + also AOT_have \... \ \\x [F]\<^sup>-x & \\x \[F]\<^sup>-x\ + by (AOT_subst (reverse) \[F]x\ \\[F]\<^sup>-x\ for: x) + (auto simp: "thm-relation-negation:2" "oth-class-taut:3:a") + also AOT_have \... \ Contingent([F]\<^sup>-)\ + using 1[OF "rel-neg-T:3", symmetric] by blast + finally show ?thesis. +qed + +AOT_define concrete_if_concrete :: \\\ ("L") + L_def: \L =\<^sub>d\<^sub>f [\x E!x \ E!x]\ + +AOT_theorem "thm-noncont-e-e:1": \Necessary(L)\ +proof - + AOT_modally_strict { + fix x + AOT_have \[\x E!x \ E!x]\\ by "cqt:2[lambda]" + moreover AOT_have \x\\ using "cqt:2[const_var]"[axiom_inst] by blast + moreover AOT_have \E!x \ E!x\ using "if-p-then-p" by blast + ultimately AOT_have \[\x E!x \ E!x]x\ + using "\\C" by blast + } + AOT_hence 0: \\\x [\x E!x \ E!x]x\ + using RN GEN by blast + show ?thesis + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF L_def]) + apply "cqt:2[lambda]" + by (rule "contingent-properties:1"[THEN "\\<^sub>d\<^sub>fI", OF 0]) +qed + +AOT_theorem "thm-noncont-e-e:2": \Impossible([L]\<^sup>-)\ +proof - + AOT_modally_strict { + fix x + + AOT_have 0: \\F (\[F]\<^sup>-x \ [F]x)\ + using "thm-relation-negation:2" GEN by fast + AOT_have \\[\x E!x \ E!x]\<^sup>-x \ [\x E!x \ E!x]x\ + by (rule 0[THEN "\E"(1)]) "cqt:2[lambda]" + moreover { + AOT_have \[\x E!x \ E!x]\\ by "cqt:2[lambda]" + moreover AOT_have \x\\ using "cqt:2[const_var]"[axiom_inst] by blast + moreover AOT_have \E!x \ E!x\ using "if-p-then-p" by blast + ultimately AOT_have \[\x E!x \ E!x]x\ + using "\\C" by blast + } + ultimately AOT_have \\[\x E!x \ E!x]\<^sup>-x\ + using "\E" by blast + } + AOT_hence 0: \\\x \[\x E!x \ E!x]\<^sup>-x\ + using RN GEN by fast + show ?thesis + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF L_def]) + apply "cqt:2[lambda]" + apply (rule "contingent-properties:2"[THEN "\\<^sub>d\<^sub>fI"]; rule "&I") + using "rel-neg-T:3" + apply blast + using 0 + by blast +qed + +AOT_theorem "thm-noncont-e-e:3": \NonContingent(L)\ + using "thm-noncont-e-e:1" + by (rule "contingent-properties:3"[THEN "\\<^sub>d\<^sub>fI", OF "\I"(1)]) + +AOT_theorem "thm-noncont-e-e:4": \NonContingent([L]\<^sup>-)\ +proof - + AOT_have 0: \\F (NonContingent([F]) \ NonContingent([F]\<^sup>-))\ + using "thm-cont-prop:1" "\I" by fast + moreover AOT_have 1: \L\\ + by (rule "=\<^sub>d\<^sub>fI"(2)[OF L_def]) "cqt:2[lambda]"+ + AOT_show \NonContingent([L]\<^sup>-)\ + using "\E"(1)[OF 0, OF 1, THEN "\E"(1), OF "thm-noncont-e-e:3"] by blast +qed + +AOT_theorem "thm-noncont-e-e:5": + \\F \G (F \ \G::<\>\ & NonContingent([F]) & NonContingent([G]))\ +proof (rule "\I")+ + { + AOT_have \\F [F] \ [F]\<^sup>-\ + using "thm-relation-negation:5" GEN by fast + moreover AOT_have \L\\ + by (rule "=\<^sub>d\<^sub>fI"(2)[OF L_def]) "cqt:2[lambda]"+ + ultimately AOT_have \L \ [L]\<^sup>-\ + using "\E" by blast + } + AOT_thus \L \ [L]\<^sup>- & NonContingent(L) & NonContingent([L]\<^sup>-)\ + using "thm-noncont-e-e:3" "thm-noncont-e-e:4" "&I" by metis +next + AOT_show \[L]\<^sup>-\\ + using "rel-neg-T:3" by blast +next + AOT_show \L\\ + by (rule "=\<^sub>d\<^sub>fI"(2)[OF L_def]) "cqt:2[lambda]"+ +qed + +AOT_theorem "lem-cont-e:1": \\\x ([F]x & \\[F]x) \ \\x (\[F]x & \[F]x)\ +proof - + AOT_have \\\x ([F]x & \\[F]x) \ \x \([F]x & \\[F]x)\ + using "BF\" "CBF\" "\I" by blast + also AOT_have \\ \ \x (\[F]x & \\[F]x)\ + by (AOT_subst \\([F]x & \\[F]x)\ \\[F]x & \\[F]x\ for: x) + (auto simp: "S5Basic:11" "cqt-further:7") + also AOT_have \\ \ \x (\\[F]x & \[F]x)\ + by (AOT_subst \\\[F]x & \[F]x\ \\[F]x & \\[F]x\ for: x) + (auto simp: "Commutativity of &" "cqt-further:7") + also AOT_have \\ \ \x \(\[F]x & \[F]x)\ + by (AOT_subst \\(\[F]x & \[F]x)\ \\\[F]x & \[F]x\ for: x) + (auto simp: "S5Basic:11" "oth-class-taut:3:a") + also AOT_have \\ \ \\x (\[F]x & \[F]x)\ + using "BF\" "CBF\" "\I" by fast + finally show ?thesis. +qed + +AOT_theorem "lem-cont-e:2": + \\\x ([F]x & \\[F]x) \ \\x ([F]\<^sup>-x & \\[F]\<^sup>-x)\ +proof - + AOT_have \\\x ([F]x & \\[F]x) \ \\x (\[F]x & \[F]x)\ + using "lem-cont-e:1". + also AOT_have \\ \ \\x ([F]\<^sup>-x & \\[F]\<^sup>-x)\ + apply (AOT_subst \\[F]\<^sup>-x\ \[F]x\ for: x) + apply (simp add: "thm-relation-negation:2") + apply (AOT_subst \[F]\<^sup>-x\ \\[F]x\ for: x) + apply (simp add: "thm-relation-negation:1") + by (simp add: "oth-class-taut:3:a") + finally show ?thesis. +qed + +AOT_theorem "thm-cont-e:1": \\\x (E!x & \\E!x)\ +proof (rule "CBF\"[THEN "\E"]) + AOT_have \\x \(E!x & \\<^bold>\E!x)\ + using "qml:4"[axiom_inst] "BF\"[THEN "\E"] by blast + then AOT_obtain a where \\(E!a & \\<^bold>\E!a)\ + using "\E"[rotated] by blast + AOT_hence \: \\E!a & \\\<^bold>\E!a\ + using "KBasic2:3"[THEN "\E"] by blast + AOT_have \: \\E!a & \\<^bold>\\E!a\ + by (AOT_subst \\<^bold>\\E!a\ \\\<^bold>\E!a\) + (auto simp: "logic-actual-nec:1"[axiom_inst] \) + AOT_have \: \\E!a & \<^bold>\\E!a\ + by (AOT_subst \\<^bold>\\E!a\ \\\<^bold>\\E!a\) + (auto simp add: "Act-Sub:4" \) + AOT_hence \\E!a & \\E!a\ + using "&E" "&I" "Act-Sub:3"[THEN "\E"] by blast + AOT_hence \\(E!a & \\E!a)\ + using "S5Basic:11"[THEN "\E"(2)] by simp + AOT_thus \\x \(E!x & \\E!x)\ + using "\I"(2) by fast +qed + +AOT_theorem "thm-cont-e:2": \\\x (\E!x & \E!x)\ +proof - + AOT_have \\F (\\x ([F]x & \\[F]x) \ \\x (\[F]x & \[F]x))\ + using "lem-cont-e:1" GEN by fast + AOT_hence \(\\x (E!x & \\E!x) \ \\x (\E!x & \E!x))\ + using "\E"(2) by blast + thus ?thesis using "thm-cont-e:1" "\E" by blast +qed + +AOT_theorem "thm-cont-e:3": \\\x E!x\ +proof (rule "CBF\"[THEN "\E"]) + AOT_obtain a where \\(E!a & \\E!a)\ + using "\E"[rotated, OF "thm-cont-e:1"[THEN "BF\"[THEN "\E"]]] by blast + AOT_hence \\E!a\ + using "KBasic2:3"[THEN "\E", THEN "&E"(1)] by blast + AOT_thus \\x \E!x\ using "\I" by fast +qed + +AOT_theorem "thm-cont-e:4": \\\x \E!x\ +proof (rule "CBF\"[THEN "\E"]) + AOT_obtain a where \\(E!a & \\E!a)\ + using "\E"[rotated, OF "thm-cont-e:1"[THEN "BF\"[THEN "\E"]]] by blast + AOT_hence \\\\E!a\ + using "KBasic2:3"[THEN "\E", THEN "&E"(2)] by blast + AOT_hence \\\E!a\ + using "4\"[THEN "\E"] by blast + AOT_thus \\x \\E!x\ using "\I" by fast +qed + +AOT_theorem "thm-cont-e:5": \Contingent([E!])\ +proof - + AOT_have \\F (Contingent([F]) \ \\x [F]x & \\x \[F]x)\ + using "thm-cont-prop:2" GEN by fast + AOT_hence \Contingent([E!]) \ \\x E!x & \\x \E!x\ + using "\E"(2) by blast + thus ?thesis + using "thm-cont-e:3" "thm-cont-e:4" "\E"(2) "&I" by blast +qed + +AOT_theorem "thm-cont-e:6": \Contingent([E!]\<^sup>-)\ +proof - + AOT_have \\F (Contingent([\F::<\>\]) \ Contingent([F]\<^sup>-))\ + using "thm-cont-prop:3" GEN by fast + AOT_hence \Contingent([E!]) \ Contingent([E!]\<^sup>-)\ + using "\E"(2) by fast + thus ?thesis using "thm-cont-e:5" "\E" by blast +qed + +AOT_theorem "thm-cont-e:7": + \\F\G (Contingent([\F::<\>\]) & Contingent([G]) & F \ G)\ +proof (rule "\I")+ + AOT_have \\F [\F::<\>\] \ [F]\<^sup>-\ + using "thm-relation-negation:5" GEN by fast + AOT_hence \[E!] \ [E!]\<^sup>-\ + using "\E" by fast + AOT_thus \Contingent([E!]) & Contingent([E!]\<^sup>-) & [E!] \ [E!]\<^sup>-\ + using "thm-cont-e:5" "thm-cont-e:6" "&I" by metis +next + AOT_show \E!\<^sup>-\\ + by (fact AOT) +qed("cqt:2") + +AOT_theorem "property-facts:1": + \NonContingent([F]) \ \\G (Contingent([G]) & G = F)\ +proof (rule "\I"; rule "raa-cor:2") + AOT_assume \NonContingent([F])\ + AOT_hence 1: \Necessary([F]) \ Impossible([F])\ + using "contingent-properties:3"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_assume \\G (Contingent([G]) & G = F)\ + then AOT_obtain G where \Contingent([G]) & G = F\ + using "\E"[rotated] by blast + AOT_hence \Contingent([F])\ using "rule=E" "&E" by blast + AOT_hence \\(Necessary([F]) \ Impossible([F]))\ + using "contingent-properties:4"[THEN "\Df", THEN "\S"(1), + OF "cqt:2[const_var]"[axiom_inst], THEN "\E"(1)] by blast + AOT_thus \(Necessary([F]) \ Impossible([F])) & + \(Necessary([F]) \ Impossible([F]))\ + using 1 "&I" by blast +qed + +AOT_theorem "property-facts:2": + \Contingent([F]) \ \\G (NonContingent([G]) & G = F)\ +proof (rule "\I"; rule "raa-cor:2") + AOT_assume \Contingent([F])\ + AOT_hence 1: \\(Necessary([F]) \ Impossible([F]))\ + using "contingent-properties:4"[THEN "\Df", THEN "\S"(1), + OF "cqt:2[const_var]"[axiom_inst], THEN "\E"(1)] by blast + AOT_assume \\G (NonContingent([G]) & G = F)\ + then AOT_obtain G where \NonContingent([G]) & G = F\ + using "\E"[rotated] by blast + AOT_hence \NonContingent([F])\ + using "rule=E" "&E" by blast + AOT_hence \Necessary([F]) \ Impossible([F])\ + using "contingent-properties:3"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_thus \(Necessary([F]) \ Impossible([F])) & + \(Necessary([F]) \ Impossible([F]))\ + using 1 "&I" by blast +qed + +AOT_theorem "property-facts:3": + \L \ [L]\<^sup>- & L \ E! & L \ E!\<^sup>- & [L]\<^sup>- \ [E!]\<^sup>- & E! \ [E!]\<^sup>-\ +proof - + AOT_have noneqI: \\ \ \'\ if \\{\}\ and \\\{\'}\ for \ and \ \' :: \<\>\ + apply (rule "=-infix"[THEN "\\<^sub>d\<^sub>fI"]; rule "raa-cor:2") + using "rule=E"[where \=\ and \=\ and \ = \'] that "&I" by blast + AOT_have contingent_denotes: \\\\ if \Contingent([\])\ for \ :: \<\>\ + using that "contingent-properties:4"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(1)] by blast + AOT_have not_noncontingent_if_contingent: + \\NonContingent([\])\ if \Contingent([\])\ for \ :: \<\>\ + proof(rule RAA(2)) + AOT_show \\(Necessary([\]) \ Impossible([\]))\ + using that "contingent-properties:4"[THEN "\Df", THEN "\S"(1), + OF contingent_denotes[OF that], THEN "\E"(1)] + by blast + next + AOT_assume \NonContingent([\])\ + AOT_thus \Necessary([\]) \ Impossible([\])\ + using "contingent-properties:3"[THEN "\\<^sub>d\<^sub>fE"] by blast + qed + + show ?thesis + proof (safe intro!: "&I") + AOT_show \L \ [L]\<^sup>-\ + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF L_def]) + apply "cqt:2[lambda]" + apply (rule "\E"(1)[where \="\ \ . \\ \ [\]\<^sup>-\"]) + apply (rule GEN) apply (fact AOT) + by "cqt:2[lambda]" + next + AOT_show \L \ E!\ + apply (rule noneqI) + using "thm-noncont-e-e:3" + not_noncontingent_if_contingent[OF "thm-cont-e:5"] + by auto + next + AOT_show \L \ E!\<^sup>-\ + apply (rule noneqI) + using "thm-noncont-e-e:3" apply fast + apply (rule not_noncontingent_if_contingent) + apply (rule "\E"(1)[ + where \="\ \ . \Contingent([\]) \ Contingent([\]\<^sup>-)\", + rotated, OF contingent_denotes, THEN "\E"(1), rotated]) + using "thm-cont-prop:3" GEN apply fast + using "thm-cont-e:5" by fast+ + next + AOT_show \[L]\<^sup>- \ E!\<^sup>-\ + apply (rule noneqI) + using "thm-noncont-e-e:4" apply fast + apply (rule not_noncontingent_if_contingent) + apply (rule "\E"(1)[ + where \="\ \ . \Contingent([\]) \ Contingent([\]\<^sup>-)\", + rotated, OF contingent_denotes, THEN "\E"(1), rotated]) + using "thm-cont-prop:3" GEN apply fast + using "thm-cont-e:5" by fast+ + next + AOT_show \E! \ E!\<^sup>-\ + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF L_def]) + apply "cqt:2[lambda]" + apply (rule "\E"(1)[where \="\ \ . \\ \ [\]\<^sup>-\"]) + apply (rule GEN) apply (fact AOT) + by "cqt:2" + qed +qed + +AOT_theorem "thm-cont-propos:1": + \NonContingent0(p) \ NonContingent0(((p)\<^sup>-))\ +proof(rule "\I"; rule "\I") + AOT_assume \NonContingent0(p)\ + AOT_hence \Necessary0(p) \ Impossible0(p)\ + using "contingent-properties:3[zero]"[THEN "\\<^sub>d\<^sub>fE"] by blast + moreover { + AOT_assume \Necessary0(p)\ + AOT_hence 1: \\p\ + using "contingent-properties:1[zero]"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_have \\\((p)\<^sup>-)\ + by (AOT_subst \\((p)\<^sup>-)\ \p\) + (auto simp add: 1 "thm-relation-negation:4") + AOT_hence \Impossible0(((p)\<^sup>-))\ + by (rule "contingent-properties:2[zero]"[THEN "\\<^sub>d\<^sub>fI"]) + } + moreover { + AOT_assume \Impossible0(p)\ + AOT_hence 1: \\\p\ + by (rule "contingent-properties:2[zero]"[THEN "\\<^sub>d\<^sub>fE"]) + AOT_have \\((p)\<^sup>-)\ + by (AOT_subst \((p)\<^sup>-)\ \\p\) + (auto simp: 1 "thm-relation-negation:3") + AOT_hence \Necessary0(((p)\<^sup>-))\ + by (rule "contingent-properties:1[zero]"[THEN "\\<^sub>d\<^sub>fI"]) + } + ultimately AOT_have \Necessary0(((p)\<^sup>-)) \ Impossible0(((p)\<^sup>-))\ + using "\E"(1) "\I" "\I" by metis + AOT_thus \NonContingent0(((p)\<^sup>-))\ + using "contingent-properties:3[zero]"[THEN "\\<^sub>d\<^sub>fI"] by blast +next + AOT_assume \NonContingent0(((p)\<^sup>-))\ + AOT_hence \Necessary0(((p)\<^sup>-)) \ Impossible0(((p)\<^sup>-))\ + using "contingent-properties:3[zero]"[THEN "\\<^sub>d\<^sub>fE"] by blast + moreover { + AOT_assume \Impossible0(((p)\<^sup>-))\ + AOT_hence 1: \\\((p)\<^sup>-)\ + by (rule "contingent-properties:2[zero]"[THEN "\\<^sub>d\<^sub>fE"]) + AOT_have \\p\ + by (AOT_subst (reverse) \p\ \\((p)\<^sup>-)\) + (auto simp: 1 "thm-relation-negation:4") + AOT_hence \Necessary0(p)\ + using "contingent-properties:1[zero]"[THEN "\\<^sub>d\<^sub>fI"] by blast + } + moreover { + AOT_assume \Necessary0(((p)\<^sup>-))\ + AOT_hence 1: \\((p)\<^sup>-)\ + by (rule "contingent-properties:1[zero]"[THEN "\\<^sub>d\<^sub>fE"]) + AOT_have \\\p\ + by (AOT_subst (reverse) \\p\ \((p)\<^sup>-)\) + (auto simp: 1 "thm-relation-negation:3") + AOT_hence \Impossible0(p)\ + by (rule "contingent-properties:2[zero]"[THEN "\\<^sub>d\<^sub>fI"]) + } + ultimately AOT_have \Necessary0(p) \ Impossible0(p)\ + using "\E"(1) "\I" "\I" by metis + AOT_thus \NonContingent0(p)\ + using "contingent-properties:3[zero]"[THEN "\\<^sub>d\<^sub>fI"] by blast +qed + +AOT_theorem "thm-cont-propos:2": \Contingent0(\) \ \\ & \\\\ +proof - + AOT_have \Contingent0(\) \ \(Necessary0(\) \ Impossible0(\))\ + using "contingent-properties:4[zero]"[THEN "\Df"] by simp + also AOT_have \\ \ \Necessary0(\) & \Impossible0(\)\ + by (fact AOT) + also AOT_have \\ \ \Impossible0(\) & \Necessary0(\)\ + by (fact AOT) + also AOT_have \\ \ \\ & \\\\ + apply (AOT_subst \\\\ \\\\\\) + apply (simp add: "conventions:5" "\Df") + apply (AOT_subst \Impossible0(\)\ \\\\\) + apply (simp add: "contingent-properties:2[zero]" "\Df") + apply (AOT_subst (reverse) \\\\\ \\\\\) + apply (simp add: "KBasic:11") + apply (AOT_subst \Necessary0(\)\ \\\\) + apply (simp add: "contingent-properties:1[zero]" "\Df") + by (simp add: "oth-class-taut:3:a") + finally show ?thesis. +qed + +AOT_theorem "thm-cont-propos:3": \Contingent0(p) \ Contingent0(((p)\<^sup>-))\ +proof - + AOT_have \Contingent0(p) \ \p & \\p\ using "thm-cont-propos:2". + also AOT_have \\ \ \\p & \p\ by (fact AOT) + also AOT_have \\ \ \((p)\<^sup>-) & \p\ + by (AOT_subst \((p)\<^sup>-)\ \\p\) + (auto simp: "thm-relation-negation:3" "oth-class-taut:3:a") + also AOT_have \\ \ \((p)\<^sup>-) & \\((p)\<^sup>-)\ + by (AOT_subst \\((p)\<^sup>-)\ \p\) + (auto simp: "thm-relation-negation:4" "oth-class-taut:3:a") + also AOT_have \\ \ Contingent0(((p)\<^sup>-))\ + using "thm-cont-propos:2"[symmetric] by blast + finally show ?thesis. +qed + +AOT_define noncontingent_prop :: \\\ ("p\<^sub>0") + p\<^sub>0_def: "(p\<^sub>0) =\<^sub>d\<^sub>f (\x (E!x \ E!x))" + +AOT_theorem "thm-noncont-propos:1": \Necessary0((p\<^sub>0))\ +proof(rule "contingent-properties:1[zero]"[THEN "\\<^sub>d\<^sub>fI"]) + AOT_show \\(p\<^sub>0)\ + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF p\<^sub>0_def]) + using "log-prop-prop:2" apply simp + using "if-p-then-p" RN GEN by fast +qed + +AOT_theorem "thm-noncont-propos:2": \Impossible0(((p\<^sub>0)\<^sup>-))\ +proof(rule "contingent-properties:2[zero]"[THEN "\\<^sub>d\<^sub>fI"]) + AOT_show \\\((p\<^sub>0)\<^sup>-)\ + apply (AOT_subst \((p\<^sub>0)\<^sup>-)\ \\p\<^sub>0\) + using "thm-relation-negation:3" GEN "\E"(1)[rotated, OF "log-prop-prop:2"] + apply fast + apply (AOT_subst (reverse) \\\p\<^sub>0\ \p\<^sub>0\) + apply (simp add: "oth-class-taut:3:b") + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF p\<^sub>0_def]) + using "log-prop-prop:2" apply simp + using "if-p-then-p" RN GEN by fast +qed + +AOT_theorem "thm-noncont-propos:3": \NonContingent0((p\<^sub>0))\ + apply(rule "contingent-properties:3[zero]"[THEN "\\<^sub>d\<^sub>fI"]) + using "thm-noncont-propos:1" "\I" by blast + +AOT_theorem "thm-noncont-propos:4": \NonContingent0(((p\<^sub>0)\<^sup>-))\ + apply(rule "contingent-properties:3[zero]"[THEN "\\<^sub>d\<^sub>fI"]) + using "thm-noncont-propos:2" "\I" by blast + +AOT_theorem "thm-noncont-propos:5": + \\p\q (NonContingent0((p)) & NonContingent0((q)) & p \ q)\ +proof(rule "\I")+ + AOT_have 0: \\ \ (\)\<^sup>-\ for \ + using "thm-relation-negation:6" "\I" + "\E"(1)[rotated, OF "log-prop-prop:2"] by fast + AOT_thus \NonContingent0((p\<^sub>0)) & NonContingent0(((p\<^sub>0)\<^sup>-)) & (p\<^sub>0) \ (p\<^sub>0)\<^sup>-\ + using "thm-noncont-propos:3" "thm-noncont-propos:4" "&I" by auto +qed(auto simp: "log-prop-prop:2") + +AOT_act_theorem "no-cnac": \\\x(E!x & \\<^bold>\E!x)\ +proof(rule "raa-cor:2") + AOT_assume \\x(E!x & \\<^bold>\E!x)\ + then AOT_obtain a where a: \E!a & \\<^bold>\E!a\ + using "\E"[rotated] by blast + AOT_hence \\<^bold>\\E!a\ + using "&E" "logic-actual-nec:1"[axiom_inst, THEN "\E"(2)] by blast + AOT_hence \\E!a\ + using "logic-actual"[act_axiom_inst, THEN "\E"] by blast + AOT_hence \E!a & \E!a\ + using a "&E" "&I" by blast + AOT_thus \p & \p\ for p using "raa-cor:1" by blast +qed + +AOT_theorem "pos-not-pna:1": \\\<^bold>\\x (E!x & \\<^bold>\E!x)\ +proof(rule "raa-cor:2") + AOT_assume \\<^bold>\\x (E!x & \\<^bold>\E!x)\ + AOT_hence \\x \<^bold>\(E!x & \\<^bold>\E!x)\ + using "Act-Basic:10"[THEN "\E"(1)] by blast + then AOT_obtain a where \\<^bold>\(E!a & \\<^bold>\E!a)\ + using "\E"[rotated] by blast + AOT_hence 1: \\<^bold>\E!a & \<^bold>\\\<^bold>\E!a\ + using "Act-Basic:2"[THEN "\E"(1)] by blast + AOT_hence \\\<^bold>\\<^bold>\E!a\ + using "&E"(2) "logic-actual-nec:1"[axiom_inst, THEN "\E"(1)] by blast + AOT_hence \\\<^bold>\E!a\ + using "logic-actual-nec:4"[axiom_inst, THEN "\E"(1)] RAA by blast + AOT_thus \p & \p\ for p using 1[THEN "&E"(1)] "&I" "raa-cor:1" by blast +qed + +AOT_theorem "pos-not-pna:2": \\\\x(E!x & \\<^bold>\E!x)\ +proof (rule RAA(1)) + AOT_show \\\<^bold>\\x (E!x & \\<^bold>\E!x)\ + using "pos-not-pna:1" by blast +next + AOT_assume \\\\\x (E!x & \\<^bold>\E!x)\ + AOT_hence \\\x (E!x & \\<^bold>\E!x)\ + using "KBasic:12"[THEN "\E"(2)] by blast + AOT_thus \\<^bold>\\x (E!x & \\<^bold>\E!x)\ + using "nec-imp-act"[THEN "\E"] by blast +qed + +AOT_theorem "pos-not-pna:3": \\x (\E!x & \\<^bold>\E!x)\ +proof - + AOT_obtain a where \\(E!a & \\<^bold>\E!a)\ + using "qml:4"[axiom_inst] "BF\"[THEN "\E"] "\E"[rotated] by blast + AOT_hence \: \\E!a\ and \: \\\\<^bold>\E!a\ + using "KBasic2:3"[THEN "\E"] "&E" by blast+ + AOT_have \\\\<^bold>\E!a\ + using \ "KBasic:11"[THEN "\E"(2)] by blast + AOT_hence \\\<^bold>\E!a\ + using "Act-Basic:6"[THEN "oth-class-taut:4:b"[THEN "\E"(1)], + THEN "\E"(2)] by blast + AOT_hence \\E!a & \\<^bold>\E!a\ using \ "&I" by blast + thus ?thesis using "\I" by fast +qed + +AOT_define contingent_prop :: \ ("q\<^sub>0") + q\<^sub>0_def: \(q\<^sub>0) =\<^sub>d\<^sub>f (\x (E!x & \\<^bold>\E!x))\ + +AOT_theorem q\<^sub>0_prop: \\q\<^sub>0 & \\q\<^sub>0\ + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF q\<^sub>0_def]) + apply (fact "log-prop-prop:2") + apply (rule "&I") + apply (fact "qml:4"[axiom_inst]) + by (fact "pos-not-pna:2") + +AOT_theorem "basic-prop:1": \Contingent0((q\<^sub>0))\ +proof(rule "contingent-properties:4[zero]"[THEN "\\<^sub>d\<^sub>fI"]) + AOT_have \\Necessary0((q\<^sub>0)) & \Impossible0((q\<^sub>0))\ + proof (rule "&I"; + rule "=\<^sub>d\<^sub>fI"(2)[OF q\<^sub>0_def]; + (rule "log-prop-prop:2" | rule "raa-cor:2")) + AOT_assume \Necessary0(\x (E!x & \\<^bold>\E!x))\ + AOT_hence \\\x (E!x & \\<^bold>\E!x)\ + using "contingent-properties:1[zero]"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_hence \\<^bold>\\x (E!x & \\<^bold>\E!x)\ + using "Act-Basic:8"[THEN "\E"] "qml:2"[axiom_inst, THEN "\E"] by blast + AOT_thus \\<^bold>\\x (E!x & \\<^bold>\E!x) & \\<^bold>\\x (E!x & \\<^bold>\E!x)\ + using "pos-not-pna:1" "&I" by blast + next + AOT_assume \Impossible0(\x (E!x & \\<^bold>\E!x))\ + AOT_hence \\\(\x (E!x & \\<^bold>\E!x))\ + using "contingent-properties:2[zero]"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_hence \\\(\x (E!x & \\<^bold>\E!x))\ + using "KBasic2:1"[THEN "\E"(1)] by blast + AOT_thus \\(\x (E!x & \\<^bold>\E!x)) & \\(\x (E!x & \\<^bold>\E!x))\ + using "qml:4"[axiom_inst] "&I" by blast + qed + AOT_thus \\(Necessary0((q\<^sub>0)) \ Impossible0((q\<^sub>0)))\ + using "oth-class-taut:5:d" "\E"(2) by blast +qed + +AOT_theorem "basic-prop:2": \\p Contingent0((p))\ + using "\I"(1)[rotated, OF "log-prop-prop:2"] "basic-prop:1" by blast + +AOT_theorem "basic-prop:3": \Contingent0(((q\<^sub>0)\<^sup>-))\ + apply (AOT_subst \((q\<^sub>0)\<^sup>-)\ \\q\<^sub>0\) + apply (insert "thm-relation-negation:3" "\I" + "\E"(1)[rotated, OF "log-prop-prop:2"]; fast) + apply (rule "contingent-properties:4[zero]"[THEN "\\<^sub>d\<^sub>fI"]) + apply (rule "oth-class-taut:5:d"[THEN "\E"(2)]) + apply (rule "&I") + apply (rule "contingent-properties:1[zero]"[THEN "df-rules-formulas[3]", + THEN "useful-tautologies:5"[THEN "\E"], THEN "\E"]) + apply (rule "conventions:5"[THEN "\\<^sub>d\<^sub>fE"]) + apply (rule "=\<^sub>d\<^sub>fE"(2)[OF q\<^sub>0_def]) + apply (rule "log-prop-prop:2") + apply (rule q\<^sub>0_prop[THEN "&E"(1)]) + apply (rule "contingent-properties:2[zero]"[THEN "df-rules-formulas[3]", + THEN "useful-tautologies:5"[THEN "\E"], THEN "\E"]) + apply (rule "conventions:5"[THEN "\\<^sub>d\<^sub>fE"]) + by (rule q\<^sub>0_prop[THEN "&E"(2)]) + +AOT_theorem "basic-prop:4": + \\p\q (p \ q & Contingent0(p) & Contingent0(q))\ +proof(rule "\I")+ + AOT_have 0: \\ \ (\)\<^sup>-\ for \ + using "thm-relation-negation:6" "\I" + "\E"(1)[rotated, OF "log-prop-prop:2"] by fast + AOT_show \(q\<^sub>0) \ (q\<^sub>0)\<^sup>- & Contingent0(q\<^sub>0) & Contingent0(((q\<^sub>0)\<^sup>-))\ + using "basic-prop:1" "basic-prop:3" "&I" 0 by presburger +qed(auto simp: "log-prop-prop:2") + +AOT_theorem "proposition-facts:1": + \NonContingent0(p) \ \\q (Contingent0(q) & q = p)\ +proof(rule "\I"; rule "raa-cor:2") + AOT_assume \NonContingent0(p)\ + AOT_hence 1: \Necessary0(p) \ Impossible0(p)\ + using "contingent-properties:3[zero]"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_assume \\q (Contingent0(q) & q = p)\ + then AOT_obtain q where \Contingent0(q) & q = p\ + using "\E"[rotated] by blast + AOT_hence \Contingent0(p)\ + using "rule=E" "&E" by fast + AOT_thus \(Necessary0(p) \ Impossible0(p)) & + \(Necessary0(p) \ Impossible0(p))\ + using "contingent-properties:4[zero]"[THEN "\\<^sub>d\<^sub>fE"] 1 "&I" by blast +qed + +AOT_theorem "proposition-facts:2": + \Contingent0(p) \ \\q (NonContingent0(q) & q = p)\ +proof(rule "\I"; rule "raa-cor:2") + AOT_assume \Contingent0(p)\ + AOT_hence 1: \\(Necessary0(p) \ Impossible0(p))\ + using "contingent-properties:4[zero]"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_assume \\q (NonContingent0(q) & q = p)\ + then AOT_obtain q where \NonContingent0(q) & q = p\ + using "\E"[rotated] by blast + AOT_hence \NonContingent0(p)\ + using "rule=E" "&E" by fast + AOT_thus \(Necessary0(p) \ Impossible0(p)) & + \(Necessary0(p) \ Impossible0(p))\ + using "contingent-properties:3[zero]"[THEN "\\<^sub>d\<^sub>fE"] 1 "&I" by blast +qed + +AOT_theorem "proposition-facts:3": + \(p\<^sub>0) \ (p\<^sub>0)\<^sup>- & (p\<^sub>0) \ (q\<^sub>0) & (p\<^sub>0) \ (q\<^sub>0)\<^sup>- & (p\<^sub>0)\<^sup>- \ (q\<^sub>0)\<^sup>- & (q\<^sub>0) \ (q\<^sub>0)\<^sup>-\ +proof - + { + fix \ \ \ + AOT_assume \\{\}\ + moreover AOT_assume \\\{\}\ + ultimately AOT_have \\(\{\} \ \{\})\ + using RAA "\E" by metis + moreover { + AOT_have \\p\q ((\(\{p} \ \{q})) \ p \ q)\ + by (rule "\I"; rule "\I"; rule "pos-not-equiv-ne:4[zero]") + AOT_hence \((\(\{\} \ \{\})) \ \ \ \)\ + using "\E" "log-prop-prop:2" by blast + } + ultimately AOT_have \\ \ \\ + using "\E" by blast + } note 0 = this + AOT_have contingent_neg: \Contingent0(\) \ Contingent0(((\)\<^sup>-))\ for \ + using "thm-cont-propos:3" "\I" + "\E"(1)[rotated, OF "log-prop-prop:2"] by fast + AOT_have not_noncontingent_if_contingent: + \\NonContingent0(\)\ if \Contingent0(\)\ for \ + apply (rule "contingent-properties:3[zero]"[THEN "\Df", + THEN "oth-class-taut:4:b"[THEN "\E"(1)], THEN "\E"(2)]) + using that "contingent-properties:4[zero]"[THEN "\\<^sub>d\<^sub>fE"] by blast + show ?thesis + apply (rule "&I")+ + using "thm-relation-negation:6" "\I" + "\E"(1)[rotated, OF "log-prop-prop:2"] + apply fast + apply (rule 0) + using "thm-noncont-propos:3" apply fast + apply (rule not_noncontingent_if_contingent) + apply (fact AOT) + apply (rule 0) + apply (rule "thm-noncont-propos:3") + apply (rule not_noncontingent_if_contingent) + apply (rule contingent_neg[THEN "\E"(1)]) + apply (fact AOT) + apply (rule 0) + apply (rule "thm-noncont-propos:4") + apply (rule not_noncontingent_if_contingent) + apply (rule contingent_neg[THEN "\E"(1)]) + apply (fact AOT) + using "thm-relation-negation:6" "\I" + "\E"(1)[rotated, OF "log-prop-prop:2"] by fast +qed + +AOT_define ContingentlyTrue :: \\ \ \\ ("ContingentlyTrue'(_')") + "cont-tf:1": \ContingentlyTrue(p) \\<^sub>d\<^sub>f p & \\p\ + +AOT_define ContingentlyFalse :: \\ \ \\ ("ContingentlyFalse'(_')") + "cont-tf:2": \ContingentlyFalse(p) \\<^sub>d\<^sub>f \p & \p\ + +AOT_theorem "cont-true-cont:1": + \ContingentlyTrue((p)) \ Contingent0((p))\ +proof(rule "\I") + AOT_assume \ContingentlyTrue((p))\ + AOT_hence 1: \p\ and 2: \\\p\ using "cont-tf:1"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast+ + AOT_have \\Necessary0((p))\ + apply (rule "contingent-properties:1[zero]"[THEN "\Df", + THEN "oth-class-taut:4:b"[THEN "\E"(1)], THEN "\E"(2)]) + using 2 "KBasic:11"[THEN "\E"(2)] by blast + moreover AOT_have \\Impossible0((p))\ + apply (rule "contingent-properties:2[zero]"[THEN "\Df", + THEN "oth-class-taut:4:b"[THEN "\E"(1)], THEN "\E"(2)]) + apply (rule "conventions:5"[THEN "\\<^sub>d\<^sub>fE"]) + using "T\"[THEN "\E", OF 1]. + ultimately AOT_have \\(Necessary0((p)) \ Impossible0((p)))\ + using DeMorgan(2)[THEN "\E"(2)] "&I" by blast + AOT_thus \Contingent0((p))\ + using "contingent-properties:4[zero]"[THEN "\\<^sub>d\<^sub>fI"] by blast +qed + +AOT_theorem "cont-true-cont:2": + \ContingentlyFalse((p)) \ Contingent0((p))\ +proof(rule "\I") + AOT_assume \ContingentlyFalse((p))\ + AOT_hence 1: \\p\ and 2: \\p\ using "cont-tf:2"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast+ + AOT_have \\Necessary0((p))\ + apply (rule "contingent-properties:1[zero]"[THEN "\Df", + THEN "oth-class-taut:4:b"[THEN "\E"(1)], THEN "\E"(2)]) + using "KBasic:11"[THEN "\E"(2)] "T\"[THEN "\E", OF 1] by blast + moreover AOT_have \\Impossible0((p))\ + apply (rule "contingent-properties:2[zero]"[THEN "\Df", + THEN "oth-class-taut:4:b"[THEN "\E"(1)], THEN "\E"(2)]) + apply (rule "conventions:5"[THEN "\\<^sub>d\<^sub>fE"]) + using 2. + ultimately AOT_have \\(Necessary0((p)) \ Impossible0((p)))\ + using DeMorgan(2)[THEN "\E"(2)] "&I" by blast + AOT_thus \Contingent0((p))\ + using "contingent-properties:4[zero]"[THEN "\\<^sub>d\<^sub>fI"] by blast +qed + +AOT_theorem "cont-true-cont:3": + \ContingentlyTrue((p)) \ ContingentlyFalse(((p)\<^sup>-))\ +proof(rule "\I"; rule "\I") + AOT_assume \ContingentlyTrue((p))\ + AOT_hence 0: \p & \\p\ using "cont-tf:1"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_have 1: \ContingentlyFalse(\p)\ + apply (rule "cont-tf:2"[THEN "\\<^sub>d\<^sub>fI"]) + apply (AOT_subst (reverse) \\\p\ p) + by (auto simp: "oth-class-taut:3:b" 0) + AOT_show \ContingentlyFalse(((p)\<^sup>-))\ + apply (AOT_subst \((p)\<^sup>-)\ \\p\) + by (auto simp: "thm-relation-negation:3" 1) +next + AOT_assume 1: \ContingentlyFalse(((p)\<^sup>-))\ + AOT_have \ContingentlyFalse(\p)\ + by (AOT_subst (reverse) \\p\ \((p)\<^sup>-)\) + (auto simp: "thm-relation-negation:3" 1) + AOT_hence \\\p & \\p\ using "cont-tf:2"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_hence \p & \\p\ + using "&I" "&E" "useful-tautologies:1"[THEN "\E"] by metis + AOT_thus \ContingentlyTrue((p))\ + using "cont-tf:1"[THEN "\\<^sub>d\<^sub>fI"] by blast +qed + +AOT_theorem "cont-true-cont:4": + \ContingentlyFalse((p)) \ ContingentlyTrue(((p)\<^sup>-))\ +proof(rule "\I"; rule "\I") + AOT_assume \ContingentlyFalse(p)\ + AOT_hence 0: \\p & \p\ + using "cont-tf:2"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_have \\p & \\\p\ + by (AOT_subst (reverse) \\\p\ p) + (auto simp: "oth-class-taut:3:b" 0) + AOT_hence 1: \ContingentlyTrue(\p)\ + by (rule "cont-tf:1"[THEN "\\<^sub>d\<^sub>fI"]) + AOT_show \ContingentlyTrue(((p)\<^sup>-))\ + by (AOT_subst \((p)\<^sup>-)\ \\p\) + (auto simp: "thm-relation-negation:3" 1) +next + AOT_assume 1: \ContingentlyTrue(((p)\<^sup>-))\ + AOT_have \ContingentlyTrue(\p)\ + by (AOT_subst (reverse) \\p\ \((p)\<^sup>-)\) + (auto simp add: "thm-relation-negation:3" 1) + AOT_hence 2: \\p & \\\p\ using "cont-tf:1"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_have \\p\ + by (AOT_subst p \\\p\) + (auto simp add: "oth-class-taut:3:b" 2[THEN "&E"(2)]) + AOT_hence \\p & \p\ using 2[THEN "&E"(1)] "&I" by blast + AOT_thus \ContingentlyFalse(p)\ + by (rule "cont-tf:2"[THEN "\\<^sub>d\<^sub>fI"]) +qed + +AOT_theorem "cont-true-cont:5": + \(ContingentlyTrue((p)) & Necessary0((q))) \ p \ q\ +proof (rule "\I"; frule "&E"(1); drule "&E"(2); rule "raa-cor:1") + AOT_assume \ContingentlyTrue((p))\ + AOT_hence \\\p\ + using "cont-tf:1"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + AOT_hence 0: \\\p\ using "KBasic:11"[THEN "\E"(2)] by blast + AOT_assume \Necessary0((q))\ + moreover AOT_assume \\(p \ q)\ + AOT_hence \p = q\ + using "=-infix"[THEN "\Df", + THEN "oth-class-taut:4:b"[THEN "\E"(1)], + THEN "\E"(1)] + "useful-tautologies:1"[THEN "\E"] by blast + ultimately AOT_have \Necessary0((p))\ using "rule=E" id_sym by blast + AOT_hence \\p\ + using "contingent-properties:1[zero]"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_thus \\p & \\p\ using 0 "&I" by blast +qed + +AOT_theorem "cont-true-cont:6": + \(ContingentlyFalse((p)) & Impossible0((q))) \ p \ q\ +proof (rule "\I"; frule "&E"(1); drule "&E"(2); rule "raa-cor:1") + AOT_assume \ContingentlyFalse((p))\ + AOT_hence \\p\ + using "cont-tf:2"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + AOT_hence 1: \\\\p\ + using "conventions:5"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_assume \Impossible0((q))\ + moreover AOT_assume \\(p \ q)\ + AOT_hence \p = q\ + using "=-infix"[THEN "\Df", + THEN "oth-class-taut:4:b"[THEN "\E"(1)], + THEN "\E"(1)] + "useful-tautologies:1"[THEN "\E"] by blast + ultimately AOT_have \Impossible0((p))\ using "rule=E" id_sym by blast + AOT_hence \\\p\ + using "contingent-properties:2[zero]"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_thus \\\p & \\\p\ using 1 "&I" by blast +qed + +AOT_act_theorem "q0cf:1": \ContingentlyFalse(q\<^sub>0)\ + apply (rule "cont-tf:2"[THEN "\\<^sub>d\<^sub>fI"]) + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF q\<^sub>0_def]) + apply (fact "log-prop-prop:2") + apply (rule "&I") + apply (fact "no-cnac") + by (fact "qml:4"[axiom_inst]) + +AOT_act_theorem "q0cf:2": \ContingentlyTrue(((q\<^sub>0)\<^sup>-))\ + apply (rule "cont-tf:1"[THEN "\\<^sub>d\<^sub>fI"]) + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF q\<^sub>0_def]) + apply (fact "log-prop-prop:2") + apply (rule "&I") + apply (rule "thm-relation-negation:3" + [unvarify p, OF "log-prop-prop:2", THEN "\E"(2)]) + apply (fact "no-cnac") + apply (rule "rule=E"[rotated, + OF "thm-relation-negation:7" + [unvarify p, OF "log-prop-prop:2", THEN id_sym]]) + apply (AOT_subst (reverse) \\\(\x (E!x & \\<^bold>\E!x))\ \\x (E!x & \\<^bold>\E!x)\) + by (auto simp: "oth-class-taut:3:b" "qml:4"[axiom_inst]) + +AOT_theorem "cont-tf-thm:1": \\p ContingentlyTrue((p))\ +proof(rule "\E"(1)[OF "exc-mid"]; rule "\I"; rule "\I") + AOT_assume \q\<^sub>0\ + AOT_hence \q\<^sub>0 & \\q\<^sub>0\ using q\<^sub>0_prop[THEN "&E"(2)] "&I" by blast + AOT_thus \ContingentlyTrue(q\<^sub>0)\ + by (rule "cont-tf:1"[THEN "\\<^sub>d\<^sub>fI"]) +next + AOT_assume \\q\<^sub>0\ + AOT_hence \\q\<^sub>0 & \q\<^sub>0\ using q\<^sub>0_prop[THEN "&E"(1)] "&I" by blast + AOT_hence \ContingentlyFalse(q\<^sub>0)\ + by (rule "cont-tf:2"[THEN "\\<^sub>d\<^sub>fI"]) + AOT_thus \ContingentlyTrue(((q\<^sub>0)\<^sup>-))\ + by (rule "cont-true-cont:4"[unvarify p, + OF "log-prop-prop:2", THEN "\E"(1)]) +qed(auto simp: "log-prop-prop:2") + + +AOT_theorem "cont-tf-thm:2": \\p ContingentlyFalse((p))\ +proof(rule "\E"(1)[OF "exc-mid"]; rule "\I"; rule "\I") + AOT_assume \q\<^sub>0\ + AOT_hence \q\<^sub>0 & \\q\<^sub>0\ using q\<^sub>0_prop[THEN "&E"(2)] "&I" by blast + AOT_hence \ContingentlyTrue(q\<^sub>0)\ + by (rule "cont-tf:1"[THEN "\\<^sub>d\<^sub>fI"]) + AOT_thus \ContingentlyFalse(((q\<^sub>0)\<^sup>-))\ + by (rule "cont-true-cont:3"[unvarify p, + OF "log-prop-prop:2", THEN "\E"(1)]) +next + AOT_assume \\q\<^sub>0\ + AOT_hence \\q\<^sub>0 & \q\<^sub>0\ using q\<^sub>0_prop[THEN "&E"(1)] "&I" by blast + AOT_thus \ContingentlyFalse(q\<^sub>0)\ + by (rule "cont-tf:2"[THEN "\\<^sub>d\<^sub>fI"]) +qed(auto simp: "log-prop-prop:2") + +AOT_theorem "property-facts1:1": \\F\x ([F]x & \\[F]x)\ +proof - + fix x + AOT_obtain p\<^sub>1 where \ContingentlyTrue((p\<^sub>1))\ + using "cont-tf-thm:1" "\E"[rotated] by blast + AOT_hence 1: \p\<^sub>1 & \\p\<^sub>1\ using "cont-tf:1"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_modally_strict { + AOT_have \for arbitrary p: \<^bold>\\<^sub>\ ([\z p]x \ p)\ + by (rule "beta-C-cor:3"[THEN "\E"(2)]) cqt_2_lambda_inst_prover + AOT_hence \for arbitrary p: \<^bold>\\<^sub>\ \ ([\z p]x \ p)\ + by (rule RN) + AOT_hence \\p \([\z p]x \ p)\ using GEN by fast + AOT_hence \\([\z p\<^sub>1]x \ p\<^sub>1)\ using "\E" by fast + } note 2 = this + AOT_hence \\([\z p\<^sub>1]x \ p\<^sub>1)\ using "\E" by blast + AOT_hence \[\z p\<^sub>1]x\ + using 1[THEN "&E"(1)] "qml:2"[axiom_inst, THEN "\E"] "\E"(2) by blast + moreover AOT_have \\\[\z p\<^sub>1]x\ + using 2[THEN "qml:2"[axiom_inst, THEN "\E"]] + apply (AOT_subst \[\z p\<^sub>1]x\ \p\<^sub>1\) + using 1[THEN "&E"(2)] by blast + ultimately AOT_have \[\z p\<^sub>1]x & \\[\z p\<^sub>1]x\ using "&I" by blast + AOT_hence \\x ([\z p\<^sub>1]x & \\[\z p\<^sub>1]x)\ using "\I"(2) by fast + moreover AOT_have \[\z p\<^sub>1]\\ by "cqt:2[lambda]" + ultimately AOT_show \\F\x ([F]x & \\[F]x)\ by (rule "\I"(1)) +qed + +AOT_theorem "property-facts1:2": \\F\x (\[F]x & \[F]x)\ +proof - + fix x + AOT_obtain p\<^sub>1 where \ContingentlyFalse((p\<^sub>1))\ + using "cont-tf-thm:2" "\E"[rotated] by blast + AOT_hence 1: \\p\<^sub>1 & \p\<^sub>1\ using "cont-tf:2"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_modally_strict { + AOT_have \for arbitrary p: \<^bold>\\<^sub>\ ([\z p]x \ p)\ + by (rule "beta-C-cor:3"[THEN "\E"(2)]) cqt_2_lambda_inst_prover + AOT_hence \for arbitrary p: \<^bold>\\<^sub>\ (\[\z p]x \ \p)\ + using "oth-class-taut:4:b" "\E" by blast + AOT_hence \for arbitrary p: \<^bold>\\<^sub>\ \(\[\z p]x \ \p)\ + by (rule RN) + AOT_hence \\p \(\[\z p]x \ \p)\ using GEN by fast + AOT_hence \\(\[\z p\<^sub>1]x \ \p\<^sub>1)\ using "\E" by fast + } note 2 = this + AOT_hence \\(\[\z p\<^sub>1]x \ \p\<^sub>1)\ using "\E" by blast + AOT_hence 3: \\[\z p\<^sub>1]x\ + using 1[THEN "&E"(1)] "qml:2"[axiom_inst, THEN "\E"] "\E"(2) by blast + AOT_modally_strict { + AOT_have \for arbitrary p: \<^bold>\\<^sub>\ ([\z p]x \ p)\ + by (rule "beta-C-cor:3"[THEN "\E"(2)]) cqt_2_lambda_inst_prover + AOT_hence \for arbitrary p: \<^bold>\\<^sub>\ \([\z p]x \ p)\ + by (rule RN) + AOT_hence \\p \([\z p]x \ p)\ using GEN by fast + AOT_hence \\([\z p\<^sub>1]x \ p\<^sub>1)\ using "\E" by fast + } note 4 = this + AOT_have \\[\z p\<^sub>1]x\ + using 4[THEN "qml:2"[axiom_inst, THEN "\E"]] + apply (AOT_subst \[\z p\<^sub>1]x\ \p\<^sub>1\) + using 1[THEN "&E"(2)] by blast + AOT_hence \\[\z p\<^sub>1]x & \[\z p\<^sub>1]x\ using 3 "&I" by blast + AOT_hence \\x (\[\z p\<^sub>1]x & \[\z p\<^sub>1]x)\ using "\I"(2) by fast + moreover AOT_have \[\z p\<^sub>1]\\ by "cqt:2[lambda]" + ultimately AOT_show \\F\x (\[F]x & \[F]x)\ by (rule "\I"(1)) +qed + +context +begin + +private AOT_lemma eqnotnec_123_Aux_\: \[L]x \ (E!x \ E!x)\ + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF L_def]) + apply "cqt:2[lambda]" + apply (rule "beta-C-meta"[THEN "\E"]) + by "cqt:2[lambda]" + +private AOT_lemma eqnotnec_123_Aux_\: \[\z \]x \ \\ + by (rule "beta-C-meta"[THEN "\E"]) "cqt:2[lambda]" + +private AOT_lemma eqnotnec_123_Aux_\: \\ \ \x([L]x \ [\z \]x)\ +proof(rule "\I"; rule "\I"; (rule "\I")?) + fix x + AOT_assume 1: \\\ + AOT_have \[L]x \ (E!x \ E!x)\ using eqnotnec_123_Aux_\. + also AOT_have \\ \ \\ + using "if-p-then-p" 1 "\I" "\I" by simp + also AOT_have \\ \ [\z \]x\ + using "Commutativity of \"[THEN "\E"(1)] eqnotnec_123_Aux_\ by blast + finally AOT_show \[L]x \ [\z \]x\. +next + fix x + AOT_assume \\x([L]x \ [\z \]x)\ + AOT_hence \[L]x \ [\z \]x\ using "\E" by blast + also AOT_have \\ \ \\ using eqnotnec_123_Aux_\. + finally AOT_have \\ \ [L]x\ + using "Commutativity of \"[THEN "\E"(1)] by blast + also AOT_have \\ \ E!x \ E!x\ using eqnotnec_123_Aux_\. + finally AOT_show \\\ using "\E" "if-p-then-p" by fast +qed +private lemmas eqnotnec_123_Aux_\ = + eqnotnec_123_Aux_\[THEN "oth-class-taut:4:b"[THEN "\E"(1)], + THEN "conventions:3"[THEN "\Df", THEN "\E"(1), THEN "&E"(1)], + THEN "RM\"] +private lemmas eqnotnec_123_Aux_\' = + eqnotnec_123_Aux_\[ + THEN "conventions:3"[THEN "\Df", THEN "\E"(1), THEN "&E"(1)], + THEN "RM\"] + +AOT_theorem "eqnotnec:1": \\F\G(\x([F]x \ [G]x) & \\\x([F]x \ [G]x))\ +proof- + AOT_obtain p\<^sub>1 where \ContingentlyTrue(p\<^sub>1)\ + using "cont-tf-thm:1" "\E"[rotated] by blast + AOT_hence \p\<^sub>1 & \\p\<^sub>1\ using "cont-tf:1"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_hence \\x ([L]x \ [\z p\<^sub>1]x) & \\\x([L]x \ [\z p\<^sub>1]x)\ + apply - apply (rule "&I") + using "&E" eqnotnec_123_Aux_\[THEN "\E"(1)] + eqnotnec_123_Aux_\ "\E" by fast+ + AOT_hence \\G (\x([L]x \ [G]x) & \\\x([L]x \ [G]x))\ + by (rule "\I") "cqt:2[lambda]" + AOT_thus \\F\G (\x([F]x \ [G]x) & \\\x([F]x \ [G]x))\ + apply (rule "\I") + by (rule "=\<^sub>d\<^sub>fI"(2)[OF L_def]) "cqt:2[lambda]"+ +qed + +AOT_theorem "eqnotnec:2": \\F\G(\\x([F]x \ [G]x) & \\x([F]x \ [G]x))\ +proof- + AOT_obtain p\<^sub>1 where \ContingentlyFalse(p\<^sub>1)\ + using "cont-tf-thm:2" "\E"[rotated] by blast + AOT_hence \\p\<^sub>1 & \p\<^sub>1\ using "cont-tf:2"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_hence \\\x ([L]x \ [\z p\<^sub>1]x) & \\x([L]x \ [\z p\<^sub>1]x)\ + apply - apply (rule "&I") + using eqnotnec_123_Aux_\[THEN "oth-class-taut:4:b"[THEN "\E"(1)], + THEN "\E"(1)] + "&E" eqnotnec_123_Aux_\' "\E" by fast+ + AOT_hence \\G (\\x([L]x \ [G]x) & \\x([L]x \ [G]x))\ + by (rule "\I") "cqt:2[lambda]" + AOT_thus \\F\G (\\x([F]x \ [G]x) & \\x([F]x \ [G]x))\ + apply (rule "\I") + by (rule "=\<^sub>d\<^sub>fI"(2)[OF L_def]) "cqt:2[lambda]"+ +qed + +AOT_theorem "eqnotnec:3": \\F\G(\<^bold>\\\x([F]x \ [G]x) & \\x([F]x \ [G]x))\ +proof- + AOT_have \\\<^bold>\q\<^sub>0\ + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF q\<^sub>0_def]) + apply (fact "log-prop-prop:2") + by (fact AOT) + AOT_hence \\<^bold>\\q\<^sub>0\ + using "logic-actual-nec:1"[axiom_inst, THEN "\E"(2)] by blast + AOT_hence \\<^bold>\\\x ([L]x \ [\z q\<^sub>0]x)\ + using eqnotnec_123_Aux_\[THEN "oth-class-taut:4:b"[THEN "\E"(1)], + THEN "conventions:3"[THEN "\Df", THEN "\E"(1), THEN "&E"(1)], + THEN "RA[2]", THEN "act-cond"[THEN "\E"], THEN "\E"] by blast + moreover AOT_have \\\x ([L]x \ [\z q\<^sub>0]x)\ + using eqnotnec_123_Aux_\'[THEN "\E"] q\<^sub>0_prop[THEN "&E"(1)] by blast + ultimately AOT_have \\<^bold>\\\x ([L]x \ [\z q\<^sub>0]x) & \\x ([L]x \ [\z q\<^sub>0]x)\ + using "&I" by blast + AOT_hence \\G (\<^bold>\\\x([L]x \ [G]x) & \\x([L]x \ [G]x))\ + by (rule "\I") "cqt:2[lambda]" + AOT_thus \\F\G (\<^bold>\\\x([F]x \ [G]x) & \\x([F]x \ [G]x))\ + apply (rule "\I") + by (rule "=\<^sub>d\<^sub>fI"(2)[OF L_def]) "cqt:2[lambda]"+ +qed + +end + +AOT_theorem "eqnotnec:4": \\F\G(\x([F]x \ [G]x) & \\\x([F]x \ [G]x))\ +proof(rule GEN) + fix F + AOT_have Aux_A: \\<^bold>\\<^sub>\ \ \ \x([F]x \ [\z [F]z & \]x)\ for \ + proof(rule "\I"; rule GEN) + AOT_modally_strict { + fix x + AOT_assume 0: \\\ + AOT_have \[\z [F]z & \]x \ [F]x & \\ + by (rule "beta-C-meta"[THEN "\E"]) "cqt:2[lambda]" + also AOT_have \... \ [F]x\ + apply (rule "\I"; rule "\I") + using "\E"(3)[rotated, OF "useful-tautologies:2"[THEN "\E"], OF 0] "&E" + apply blast + using 0 "&I" by blast + finally AOT_show \[F]x \ [\z [F]z & \]x\ + using "Commutativity of \"[THEN "\E"(1)] by blast + } + qed + + AOT_have Aux_B: \\<^bold>\\<^sub>\ \ \ \x([F]x \ [\z [F]z & \ \ \\]x)\ for \ + proof (rule "\I"; rule GEN) + AOT_modally_strict { + fix x + AOT_assume 0: \\\ + AOT_have \[\z ([F]z & \) \ \\]x \ (([F]x & \) \ \\)\ + by (rule "beta-C-meta"[THEN "\E"]) "cqt:2[lambda]" + also AOT_have \... \ [F]x\ + apply (rule "\I"; rule "\I") + using "\E"(3)[rotated, OF "useful-tautologies:2"[THEN "\E"], OF 0] + "&E" + apply blast + apply (rule "\I"(1)) using 0 "&I" by blast + finally AOT_show \[F]x \ [\z ([F]z & \) \ \\]x\ + using "Commutativity of \"[THEN "\E"(1)] by blast + } + qed + + AOT_have Aux_C: + \\<^bold>\\<^sub>\ \\\ \ \\\z([\z [F]z & \]z \ [\z [F]z & \ \ \\]z)\ for \ + proof(rule "RM\"; rule "\I"; rule "raa-cor:2") + AOT_modally_strict { + AOT_assume 0: \\\\ + AOT_assume \\z ([\z [F]z & \]z \ [\z [F]z & \ \ \\]z)\ + AOT_hence \[\z [F]z & \]z \ [\z [F]z & \ \ \\]z\ for z + using "\E" by blast + moreover AOT_have \[\z [F]z & \]z \ [F]z & \\ for z + by (rule "beta-C-meta"[THEN "\E"]) "cqt:2[lambda]" + moreover AOT_have \[\z ([F]z & \) \ \\]z \ (([F]z & \) \ \\)\ for z + by (rule "beta-C-meta"[THEN "\E"]) "cqt:2[lambda]" + ultimately AOT_have \[F]z & \ \ (([F]z & \) \ \\)\ for z + using "Commutativity of \"[THEN "\E"(1)] "\E"(5) by meson + moreover AOT_have \(([F]z & \) \ \\)\ for z using 0 "\I" by blast + ultimately AOT_have \\\ using "\E" "&E" by metis + AOT_thus \\ & \\\ using 0 "&I" by blast + } + qed + + AOT_have Aux_D: \\\z ([F]z \ [\z [F]z & \]z) \ + (\\\x ([\z [F]z & \]x \ [\z [F]z & \ \ \\]x) \ + \\\x ([F]x \ [\z [F]z & \ \ \\]x))\ for \ + proof (rule "\I") + AOT_assume A: \\\z([F]z \ [\z [F]z & \]z)\ + AOT_show \\\\x ([\z [F]z & \]x \ [\z [F]z & \ \ \\]x) \ + \\\x ([F]x \ [\z [F]z & \ \ \\]x)\ + proof(rule "\I"; rule "KBasic:13"[THEN "\E"]; + rule "RN[prem]"[where \="{\\z([F]z \ [\z [F]z & \]z)\}", simplified]; + (rule "useful-tautologies:5"[THEN "\E"]; rule "\I")?) + AOT_modally_strict { + AOT_assume \\z ([F]z \ [\z [F]z & \]z)\ + AOT_hence 1: \[F]z \ [\z [F]z & \]z\ for z + using "\E" by blast + AOT_assume \\x ([F]x \ [\z [F]z & \ \ \\]x)\ + AOT_hence 2: \[F]z \ [\z [F]z & \ \ \\]z\ for z + using "\E" by blast + AOT_have \[\z [F]z & \]z \ [\z [F]z & \ \ \\]z\ for z + using "\E" 1 2 by meson + AOT_thus \\x ([\z [F]z & \]x \ [\z [F]z & \ \ \\]x)\ + by (rule GEN) + } + next + AOT_modally_strict { + AOT_assume \\z ([F]z \ [\z [F]z & \]z)\ + AOT_hence 1: \[F]z \ [\z [F]z & \]z\ for z + using "\E" by blast + AOT_assume \\x ([\z [F]z & \]x \ [\z [F]z & \ \ \\]x)\ + AOT_hence 2: \[\z [F]z & \]z \ [\z [F]z & \ \ \\]z\ for z + using "\E" by blast + AOT_have \[F]z \ [\z [F]z & \ \ \\]z\ for z + using 1 2 "\E" by meson + AOT_thus \ \x ([F]x \ [\z [F]z & \ \ \\]x)\ + by (rule GEN) + } + qed(auto simp: A) + qed + + AOT_obtain p\<^sub>1 where p\<^sub>1_prop: \p\<^sub>1 & \\p\<^sub>1\ + using "cont-tf-thm:1" "\E"[rotated] + "cont-tf:1"[THEN "\\<^sub>d\<^sub>fE"] by blast + { + AOT_assume 1: \\\x([F]x \ [\z [F]z & p\<^sub>1]x)\ + AOT_have 2: \\x([F]x \ [\z [F]z & p\<^sub>1 \ \p\<^sub>1]x)\ + using Aux_B[THEN "\E", OF p\<^sub>1_prop[THEN "&E"(1)]]. + AOT_have \\\\x([\z [F]z & p\<^sub>1]x \ [\z [F]z & p\<^sub>1 \ \p\<^sub>1]x)\ + using Aux_C[THEN "\E", OF p\<^sub>1_prop[THEN "&E"(2)]]. + AOT_hence 3: \\\\x([F]x \ [\z [F]z & p\<^sub>1 \ \p\<^sub>1]x)\ + using Aux_D[THEN "\E", OF 1, THEN "\E"(1)] by blast + AOT_hence \\x([F]x \ [\z [F]z & p\<^sub>1 \ \p\<^sub>1]x) & + \\\x([F]x \ [\z [F]z & p\<^sub>1 \ \p\<^sub>1]x)\ + using 2 "&I" by blast + AOT_hence \\G (\x ([F]x \ [G]x) & \\\x([F]x \ [G]x))\ + by (rule "\I"(1)) "cqt:2[lambda]" + } + moreover { + AOT_assume 2: \\\\x([F]x \ [\z [F]z & p\<^sub>1]x)\ + AOT_hence \\\\x([F]x \ [\z [F]z & p\<^sub>1]x)\ + using "KBasic:11"[THEN "\E"(1)] by blast + AOT_hence \\x ([F]x \ [\z [F]z & p\<^sub>1]x) & \\\x([F]x \ [\z [F]z & p\<^sub>1]x)\ + using Aux_A[THEN "\E", OF p\<^sub>1_prop[THEN "&E"(1)]] "&I" by blast + AOT_hence \\G (\x ([F]x \ [G]x) & \\\x([F]x \ [G]x))\ + by (rule "\I"(1)) "cqt:2[lambda]" + } + ultimately AOT_show \\G (\x ([F]x \ [G]x) & \\\x([F]x \ [G]x))\ + using "\E"(1)[OF "exc-mid"] "\I" by blast +qed + +AOT_theorem "eqnotnec:5": \\F\G(\\x([F]x \ [G]x) & \\x([F]x \ [G]x))\ +proof(rule GEN) + fix F + AOT_have Aux_A: \\<^bold>\\<^sub>\ \\ \ \\x([F]x \ [\z [F]z & \]x)\ for \ + proof(rule "RM\"; rule "\I"; rule GEN) + AOT_modally_strict { + fix x + AOT_assume 0: \\\ + AOT_have \[\z [F]z & \]x \ [F]x & \\ + by (rule "beta-C-meta"[THEN "\E"]) "cqt:2[lambda]" + also AOT_have \... \ [F]x\ + apply (rule "\I"; rule "\I") + using "\E"(3)[rotated, OF "useful-tautologies:2"[THEN "\E"], OF 0] "&E" + apply blast + using 0 "&I" by blast + finally AOT_show \[F]x \ [\z [F]z & \]x\ + using "Commutativity of \"[THEN "\E"(1)] by blast + } + qed + + AOT_have Aux_B: \\<^bold>\\<^sub>\ \\ \ \\x([F]x \ [\z [F]z & \ \ \\]x)\ for \ + proof (rule "RM\"; rule "\I"; rule GEN) + AOT_modally_strict { + fix x + AOT_assume 0: \\\ + AOT_have \[\z ([F]z & \) \ \\]x \ (([F]x & \) \ \\)\ + by (rule "beta-C-meta"[THEN "\E"]) "cqt:2[lambda]" + also AOT_have \... \ [F]x\ + apply (rule "\I"; rule "\I") + using "\E"(3)[rotated, OF "useful-tautologies:2"[THEN "\E"], OF 0] "&E" + apply blast + apply (rule "\I"(1)) using 0 "&I" by blast + finally AOT_show \[F]x \ [\z ([F]z & \) \ \\]x\ + using "Commutativity of \"[THEN "\E"(1)] by blast + } + qed + + AOT_have Aux_C: \\<^bold>\\<^sub>\ \\ \ \\z([\z [F]z & \]z \ [\z [F]z & \ \ \\]z)\ for \ + proof(rule "\I"; rule "raa-cor:2") + AOT_modally_strict { + AOT_assume 0: \\\\ + AOT_assume \\z ([\z [F]z & \]z \ [\z [F]z & \ \ \\]z)\ + AOT_hence \[\z [F]z & \]z \ [\z [F]z & \ \ \\]z\ for z + using "\E" by blast + moreover AOT_have \[\z [F]z & \]z \ [F]z & \\ for z + by (rule "beta-C-meta"[THEN "\E"]) "cqt:2[lambda]" + moreover AOT_have \[\z ([F]z & \) \ \\]z \ (([F]z & \) \ \\)\ for z + by (rule "beta-C-meta"[THEN "\E"]) "cqt:2[lambda]" + ultimately AOT_have \[F]z & \ \ (([F]z & \) \ \\)\ for z + using "Commutativity of \"[THEN "\E"(1)] "\E"(5) by meson + moreover AOT_have \(([F]z & \) \ \\)\ for z + using 0 "\I" by blast + ultimately AOT_have \\\ using "\E" "&E" by metis + AOT_thus \\ & \\\ using 0 "&I" by blast + } + qed + + AOT_have Aux_D: \\z ([F]z \ [\z [F]z & \]z) \ + (\\x ([\z [F]z & \]x \ [\z [F]z & \ \ \\]x) \ + \\x ([F]x \ [\z [F]z & \ \ \\]x))\ for \ + proof (rule "\I"; rule "\I"; + (rule "useful-tautologies:5"[THEN "\E"]; rule "\I")?) + AOT_modally_strict { + AOT_assume \\z ([F]z \ [\z [F]z & \]z)\ + AOT_hence 1: \[F]z \ [\z [F]z & \]z\ for z + using "\E" by blast + AOT_assume \\x ([F]x \ [\z [F]z & \ \ \\]x)\ + AOT_hence 2: \[F]z \ [\z [F]z & \ \ \\]z\ for z + using "\E" by blast + AOT_have \[\z [F]z & \]z \ [\z [F]z & \ \ \\]z\ for z + using "\E" 1 2 by meson + AOT_thus \\x ([\z [F]z & \]x \ [\z [F]z & \ \ \\]x)\ + by (rule GEN) + } + next + AOT_modally_strict { + AOT_assume \\z ([F]z \ [\z [F]z & \]z)\ + AOT_hence 1: \[F]z \ [\z [F]z & \]z\ for z + using "\E" by blast + AOT_assume \\x ([\z [F]z & \]x \ [\z [F]z & \ \ \\]x)\ + AOT_hence 2: \[\z [F]z & \]z \ [\z [F]z & \ \ \\]z\ for z + using "\E" by blast + AOT_have \[F]z \ [\z [F]z & \ \ \\]z\ for z + using 1 2 "\E" by meson + AOT_thus \ \x ([F]x \ [\z [F]z & \ \ \\]x)\ + by (rule GEN) + } + qed + + AOT_obtain p\<^sub>1 where p\<^sub>1_prop: \\p\<^sub>1 & \p\<^sub>1\ + using "cont-tf-thm:2" "\E"[rotated] "cont-tf:2"[THEN "\\<^sub>d\<^sub>fE"] by blast + { + AOT_assume 1: \\x([F]x \ [\z [F]z & p\<^sub>1]x)\ + AOT_have 2: \\\x([F]x \ [\z [F]z & p\<^sub>1 \ \p\<^sub>1]x)\ + using Aux_B[THEN "\E", OF p\<^sub>1_prop[THEN "&E"(2)]]. + AOT_have \\\x([\z [F]z & p\<^sub>1]x \ [\z [F]z & p\<^sub>1 \ \p\<^sub>1]x)\ + using Aux_C[THEN "\E", OF p\<^sub>1_prop[THEN "&E"(1)]]. + AOT_hence 3: \\\x([F]x \ [\z [F]z & p\<^sub>1 \ \p\<^sub>1]x)\ + using Aux_D[THEN "\E", OF 1, THEN "\E"(1)] by blast + AOT_hence \\\x([F]x \ [\z [F]z & p\<^sub>1 \ \p\<^sub>1]x) & + \\x([F]x \ [\z [F]z & p\<^sub>1 \ \p\<^sub>1]x)\ + using 2 "&I" by blast + AOT_hence \\G (\\x ([F]x \ [G]x) & \\x([F]x \ [G]x))\ + by (rule "\I"(1)) "cqt:2[lambda]" + } + moreover { + AOT_assume 2: \\\x([F]x \ [\z [F]z & p\<^sub>1]x)\ + AOT_hence \\\x([F]x \ [\z [F]z & p\<^sub>1]x)\ + using "KBasic:11"[THEN "\E"(1)] by blast + AOT_hence \\\x ([F]x \ [\z [F]z & p\<^sub>1]x) & + \\x([F]x \ [\z [F]z & p\<^sub>1]x)\ + using Aux_A[THEN "\E", OF p\<^sub>1_prop[THEN "&E"(2)]] "&I" by blast + AOT_hence \\G (\\x ([F]x \ [G]x) & \\x([F]x \ [G]x))\ + by (rule "\I"(1)) "cqt:2[lambda]" + } + ultimately AOT_show \\G (\\x ([F]x \ [G]x) & \\x([F]x \ [G]x))\ + using "\E"(1)[OF "exc-mid"] "\I" by blast +qed + +AOT_theorem "eqnotnec:6": \\F\G(\<^bold>\\\x([F]x \ [G]x) & \\x([F]x \ [G]x))\ +proof(rule GEN) + fix F + AOT_have Aux_A: \\<^bold>\\<^sub>\ \\ \ \\x([F]x \ [\z [F]z & \]x)\ for \ + proof(rule "RM\"; rule "\I"; rule GEN) + AOT_modally_strict { + fix x + AOT_assume 0: \\\ + AOT_have \[\z [F]z & \]x \ [F]x & \\ + by (rule "beta-C-meta"[THEN "\E"]) "cqt:2[lambda]" + also AOT_have \... \ [F]x\ + apply (rule "\I"; rule "\I") + using "\E"(3)[rotated, OF "useful-tautologies:2"[THEN "\E"], OF 0] + "&E" + apply blast + using 0 "&I" by blast + finally AOT_show \[F]x \ [\z [F]z & \]x\ + using "Commutativity of \"[THEN "\E"(1)] by blast + } + qed + + AOT_have Aux_B: \\<^bold>\\<^sub>\ \\ \ \\x([F]x \ [\z [F]z & \ \ \\]x)\ for \ + proof (rule "RM\"; rule "\I"; rule GEN) + AOT_modally_strict { + fix x + AOT_assume 0: \\\ + AOT_have \[\z ([F]z & \) \ \\]x \ (([F]x & \) \ \\)\ + by (rule "beta-C-meta"[THEN "\E"]) "cqt:2[lambda]" + also AOT_have \... \ [F]x\ + apply (rule "\I"; rule "\I") + using "\E"(3)[rotated, OF "useful-tautologies:2"[THEN "\E"], OF 0] "&E" + apply blast + apply (rule "\I"(1)) using 0 "&I" by blast + finally AOT_show \[F]x \ [\z ([F]z & \) \ \\]x\ + using "Commutativity of \"[THEN "\E"(1)] by blast + } + qed + + AOT_have Aux_C: + \\<^bold>\\<^sub>\ \<^bold>\\\ \ \<^bold>\\\z([\z [F]z & \]z \ [\z [F]z & \ \ \\]z)\ for \ + proof(rule "act-cond"[THEN "\E"]; rule "RA[2]"; rule "\I"; rule "raa-cor:2") + AOT_modally_strict { + AOT_assume 0: \\\\ + AOT_assume \\z ([\z [F]z & \]z \ [\z [F]z & \ \ \\]z)\ + AOT_hence \[\z [F]z & \]z \ [\z [F]z & \ \ \\]z\ for z + using "\E" by blast + moreover AOT_have \[\z [F]z & \]z \ [F]z & \\ for z + by (rule "beta-C-meta"[THEN "\E"]) "cqt:2[lambda]" + moreover AOT_have \[\z ([F]z & \) \ \\]z \ (([F]z & \) \ \\)\ for z + by (rule "beta-C-meta"[THEN "\E"]) "cqt:2[lambda]" + ultimately AOT_have \[F]z & \ \ (([F]z & \) \ \\)\ for z + using "Commutativity of \"[THEN "\E"(1)] "\E"(5) by meson + moreover AOT_have \(([F]z & \) \ \\)\ for z + using 0 "\I" by blast + ultimately AOT_have \\\ using "\E" "&E" by metis + AOT_thus \\ & \\\ using 0 "&I" by blast + } + qed + + AOT_have \\(\z ([F]z \ [\z [F]z & \]z) \ + (\\x ([\z [F]z & \]x \ [\z [F]z & \ \ \\]x) \ + \\x ([F]x \ [\z [F]z & \ \ \\]x)))\ for \ + proof (rule RN; rule "\I") + AOT_modally_strict { + AOT_assume \\z ([F]z \ [\z [F]z & \]z)\ + AOT_thus \\\x ([\z [F]z & \]x \ [\z [F]z & \ \ \\]x) \ + \\x ([F]x \ [\z [F]z & \ \ \\]x)\ + apply - + proof(rule "\I"; (rule "useful-tautologies:5"[THEN "\E"]; rule "\I")?) + AOT_assume \\z ([F]z \ [\z [F]z & \]z)\ + AOT_hence 1: \[F]z \ [\z [F]z & \]z\ for z + using "\E" by blast + AOT_assume \\x ([F]x \ [\z [F]z & \ \ \\]x)\ + AOT_hence 2: \[F]z \ [\z [F]z & \ \ \\]z\ for z + using "\E" by blast + AOT_have \[\z [F]z & \]z \ [\z [F]z & \ \ \\]z\ for z + using "\E" 1 2 by meson + AOT_thus \\x ([\z [F]z & \]x \ [\z [F]z & \ \ \\]x)\ + by (rule GEN) + next + AOT_assume \\z ([F]z \ [\z [F]z & \]z)\ + AOT_hence 1: \[F]z \ [\z [F]z & \]z\ for z + using "\E" by blast + AOT_assume \\x ([\z [F]z & \]x \ [\z [F]z & \ \ \\]x)\ + AOT_hence 2: \[\z [F]z & \]z \ [\z [F]z & \ \ \\]z\ for z + using "\E" by blast + AOT_have \[F]z \ [\z [F]z & \ \ \\]z\ for z + using 1 2 "\E" by meson + AOT_thus \ \x ([F]x \ [\z [F]z & \ \ \\]x)\ + by (rule GEN) + qed + } + qed + AOT_hence \\<^bold>\(\z ([F]z \ [\z [F]z & \]z) \ + (\\x ([\z [F]z & \]x \ [\z [F]z & \ \ \\]x) \ + \\x ([F]x \ [\z [F]z & \ \ \\]x)))\ for \ + using "nec-imp-act"[THEN "\E"] by blast + AOT_hence \\<^bold>\\z ([F]z \ [\z [F]z & \]z) \ + \<^bold>\(\\x ([\z [F]z & \]x \ [\z [F]z & \ \ \\]x) \ + \\x ([F]x \ [\z [F]z & \ \ \\]x))\ for \ + using "act-cond"[THEN "\E"] by blast + AOT_hence Aux_D: \\<^bold>\\z ([F]z \ [\z [F]z & \]z) \ + (\<^bold>\\\x ([\z [F]z & \]x \ [\z [F]z & \ \ \\]x) \ + \<^bold>\\\x ([F]x \ [\z [F]z & \ \ \\]x))\ for \ + by (auto intro!: "\I" "Act-Basic:5"[THEN "\E"(1)] dest!: "\E") + + AOT_have \\\<^bold>\q\<^sub>0\ + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF q\<^sub>0_def]) + apply (fact "log-prop-prop:2") + by (fact AOT) + AOT_hence q\<^sub>0_prop_1: \\<^bold>\\q\<^sub>0\ + using "logic-actual-nec:1"[axiom_inst, THEN "\E"(2)] by blast + { + AOT_assume 1: \\<^bold>\\x([F]x \ [\z [F]z & q\<^sub>0]x)\ + AOT_have 2: \\\x([F]x \ [\z [F]z & q\<^sub>0 \ \q\<^sub>0]x)\ + using Aux_B[THEN "\E", OF q\<^sub>0_prop[THEN "&E"(1)]]. + AOT_have \\<^bold>\\\x([\z [F]z & q\<^sub>0]x \ [\z [F]z & q\<^sub>0 \ \q\<^sub>0]x)\ + using Aux_C[THEN "\E", OF q\<^sub>0_prop_1]. + AOT_hence 3: \\<^bold>\\\x([F]x \ [\z [F]z & q\<^sub>0 \ \q\<^sub>0]x)\ + using Aux_D[THEN "\E", OF 1, THEN "\E"(1)] by blast + AOT_hence \\<^bold>\\\x([F]x \ [\z [F]z & q\<^sub>0 \ \q\<^sub>0]x) & + \\x([F]x \ [\z [F]z & q\<^sub>0 \ \q\<^sub>0]x)\ + using 2 "&I" by blast + AOT_hence \\G (\<^bold>\\\x ([F]x \ [G]x) & \\x([F]x \ [G]x))\ + by (rule "\I"(1)) "cqt:2[lambda]" + } + moreover { + AOT_assume 2: \\\<^bold>\\x([F]x \ [\z [F]z & q\<^sub>0]x)\ + AOT_hence \\<^bold>\\\x([F]x \ [\z [F]z & q\<^sub>0]x)\ + using "logic-actual-nec:1"[axiom_inst, THEN "\E"(2)] by blast + AOT_hence \\<^bold>\\\x ([F]x \ [\z [F]z & q\<^sub>0]x) & \\x([F]x \ [\z [F]z & q\<^sub>0]x)\ + using Aux_A[THEN "\E", OF q\<^sub>0_prop[THEN "&E"(1)]] "&I" by blast + AOT_hence \\G (\<^bold>\\\x ([F]x \ [G]x) & \\x([F]x \ [G]x))\ + by (rule "\I"(1)) "cqt:2[lambda]" + } + ultimately AOT_show \\G (\<^bold>\\\x ([F]x \ [G]x) & \\x([F]x \ [G]x))\ + using "\E"(1)[OF "exc-mid"] "\I" by blast +qed + +AOT_theorem "oa-contingent:1": \O! \ A!\ +proof(rule "\\<^sub>d\<^sub>fI"[OF "=-infix"]; rule "raa-cor:2") + fix x + AOT_assume 1: \O! = A!\ + AOT_hence \[\x \E!x] = A!\ + by (rule "=\<^sub>d\<^sub>fE"(2)[OF AOT_ordinary, rotated]) "cqt:2[lambda]" + AOT_hence \[\x \E!x] = [\x \\E!x]\ + by (rule "=\<^sub>d\<^sub>fE"(2)[OF AOT_abstract, rotated]) "cqt:2[lambda]" + moreover AOT_have \[\x \E!x]x \ \E!x\ + by (rule "beta-C-meta"[THEN "\E"]) "cqt:2[lambda]" + ultimately AOT_have \[\x \\E!x]x \ \E!x\ + using "rule=E" by fast + moreover AOT_have \[\x \\E!x]x \ \\E!x\ + by (rule "beta-C-meta"[THEN "\E"]) "cqt:2[lambda]" + ultimately AOT_have \\E!x \ \\E!x\ + using "\E"(6) "Commutativity of \"[THEN "\E"(1)] by blast + AOT_thus "(\E!x \ \\E!x) & \(\E!x \ \\E!x)" + using "oth-class-taut:3:c" "&I" by blast +qed + +AOT_theorem "oa-contingent:2": \O!x \ \A!x\ +proof - + AOT_have \O!x \ [\x \E!x]x\ + apply (rule "\I"; rule "\I") + apply (rule "=\<^sub>d\<^sub>fE"(2)[OF AOT_ordinary]) + apply "cqt:2[lambda]" + apply argo + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF AOT_ordinary]) + apply "cqt:2[lambda]" + by argo + also AOT_have \\ \ \E!x\ + by (rule "beta-C-meta"[THEN "\E"]) "cqt:2[lambda]" + also AOT_have \\ \ \\\E!x\ + using "oth-class-taut:3:b". + also AOT_have \\ \ \[\x \\E!x]x\ + by (rule "beta-C-meta"[THEN "\E", + THEN "oth-class-taut:4:b"[THEN "\E"(1)], symmetric]) + "cqt:2" + also AOT_have \\ \ \A!x\ + apply (rule "\I"; rule "\I") + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF AOT_abstract]) + apply "cqt:2[lambda]" + apply argo + apply (rule "=\<^sub>d\<^sub>fE"(2)[OF AOT_abstract]) + apply "cqt:2[lambda]" + by argo + finally show ?thesis. +qed + +AOT_theorem "oa-contingent:3": \A!x \ \O!x\ + by (AOT_subst \A!x\ \\\A!x\) + (auto simp add: "oth-class-taut:3:b" "oa-contingent:2"[THEN + "oth-class-taut:4:b"[THEN "\E"(1)], symmetric]) + +AOT_theorem "oa-contingent:4": \Contingent(O!)\ +proof (rule "thm-cont-prop:2"[unvarify F, OF "oa-exist:1", THEN "\E"(2)]; + rule "&I") + AOT_have \\\x E!x\ using "thm-cont-e:3" . + AOT_hence \\x \E!x\ using "BF\"[THEN "\E"] by blast + then AOT_obtain a where \\E!a\ using "\E"[rotated] by blast + AOT_hence \[\x \E!x]a\ + by (rule "beta-C-meta"[THEN "\E", THEN "\E"(2), rotated]) "cqt:2" + AOT_hence \O!a\ + by (rule "=\<^sub>d\<^sub>fI"(2)[OF AOT_ordinary, rotated]) "cqt:2" + AOT_hence \\x O!x\ using "\I" by blast + AOT_thus \\\x O!x\ using "T\"[THEN "\E"] by blast +next + AOT_obtain a where \A!a\ + using "A-objects"[axiom_inst] "\E"[rotated] "&E" by blast + AOT_hence \\O!a\ using "oa-contingent:3"[THEN "\E"(1)] by blast + AOT_hence \\x \O!x\ using "\I" by fast + AOT_thus \\\x \O!x\ using "T\"[THEN "\E"] by blast +qed + +AOT_theorem "oa-contingent:5": \Contingent(A!)\ +proof (rule "thm-cont-prop:2"[unvarify F, OF "oa-exist:2", THEN "\E"(2)]; + rule "&I") + AOT_obtain a where \A!a\ + using "A-objects"[axiom_inst] "\E"[rotated] "&E" by blast + AOT_hence \\x A!x\ using "\I" by fast + AOT_thus \\\x A!x\ using "T\"[THEN "\E"] by blast +next + AOT_have \\\x E!x\ using "thm-cont-e:3" . + AOT_hence \\x \E!x\ using "BF\"[THEN "\E"] by blast + then AOT_obtain a where \\E!a\ using "\E"[rotated] by blast + AOT_hence \[\x \E!x]a\ + by (rule "beta-C-meta"[THEN "\E", THEN "\E"(2), rotated]) "cqt:2[lambda]" + AOT_hence \O!a\ + by (rule "=\<^sub>d\<^sub>fI"(2)[OF AOT_ordinary, rotated]) "cqt:2[lambda]" + AOT_hence \\A!a\ using "oa-contingent:2"[THEN "\E"(1)] by blast + AOT_hence \\x \A!x\ using "\I" by fast + AOT_thus \\\x \A!x\ using "T\"[THEN "\E"] by blast +qed + +AOT_theorem "oa-contingent:7": \O!\<^sup>-x \ \A!\<^sup>-x\ +proof - + AOT_have \O!x \ \A!x\ + using "oa-contingent:2" by blast + also AOT_have \\ \ A!\<^sup>-x\ + using "thm-relation-negation:1"[symmetric, unvarify F, OF "oa-exist:2"]. + finally AOT_have 1: \O!x \ A!\<^sup>-x\. + + AOT_have \A!x \ \O!x\ + using "oa-contingent:3" by blast + also AOT_have \\ \ O!\<^sup>-x\ + using "thm-relation-negation:1"[symmetric, unvarify F, OF "oa-exist:1"]. + finally AOT_have 2: \A!x \ O!\<^sup>-x\. + + AOT_show \O!\<^sup>-x \ \A!\<^sup>-x\ + using 1[THEN "oth-class-taut:4:b"[THEN "\E"(1)]] + "oa-contingent:3"[of _ x] 2[symmetric] + "\E"(5) by blast +qed + +AOT_theorem "oa-contingent:6": \O!\<^sup>- \ A!\<^sup>-\ +proof (rule "=-infix"[THEN "\\<^sub>d\<^sub>fI"]; rule "raa-cor:2") + AOT_assume 1: \O!\<^sup>- = A!\<^sup>-\ + fix x + AOT_have \A!\<^sup>-x \ O!\<^sup>-x\ + apply (rule "rule=E"[rotated, OF 1]) + by (fact "oth-class-taut:3:a") + AOT_hence \A!\<^sup>-x \ \A!\<^sup>-x\ + using "oa-contingent:7" "\E" by fast + AOT_thus \(A!\<^sup>-x \ \A!\<^sup>-x) & \(A!\<^sup>-x \ \A!\<^sup>-x)\ + using "oth-class-taut:3:c" "&I" by blast +qed + +AOT_theorem "oa-contingent:8": \Contingent(O!\<^sup>-)\ + using "thm-cont-prop:3"[unvarify F, OF "oa-exist:1", THEN "\E"(1), + OF "oa-contingent:4"]. + +AOT_theorem "oa-contingent:9": \Contingent(A!\<^sup>-)\ + using "thm-cont-prop:3"[unvarify F, OF "oa-exist:2", THEN "\E"(1), + OF "oa-contingent:5"]. + +AOT_define WeaklyContingent :: \\ \ \\ (\WeaklyContingent'(_')\) + "df-cont-nec": + \WeaklyContingent([F]) \\<^sub>d\<^sub>f Contingent([F]) & \x (\[F]x \ \[F]x)\ + +AOT_theorem "cont-nec-fact1:1": + \WeaklyContingent([F]) \ WeaklyContingent([F]\<^sup>-)\ +proof - + AOT_have \WeaklyContingent([F]) \ Contingent([F]) & \x (\[F]x \ \[F]x)\ + using "df-cont-nec"[THEN "\Df"] by blast + also AOT_have \... \ Contingent([F]\<^sup>-) & \x (\[F]x \ \[F]x)\ + apply (rule "oth-class-taut:8:f"[THEN "\E"(2)]; rule "\I") + using "thm-cont-prop:3". + also AOT_have \\ \ Contingent([F]\<^sup>-) & \x (\[F]\<^sup>-x \ \[F]\<^sup>-x)\ + proof (rule "oth-class-taut:8:e"[THEN "\E"(2)]; + rule "\I"; rule "\I"; rule "\I"; rule GEN; rule "\I") + fix x + AOT_assume 0: \\x (\[F]x \ \[F]x)\ + AOT_assume 1: \\[F]\<^sup>-x\ + AOT_have \\\[F]x\ + by (AOT_subst (reverse) \\[F]x\ \[F]\<^sup>-x\) + (auto simp add: "thm-relation-negation:1" 1) + AOT_hence 2: \\\[F]x\ + using "KBasic:11"[THEN "\E"(2)] by blast + AOT_show \\[F]\<^sup>-x\ + proof (rule "raa-cor:1") + AOT_assume 3: \\\[F]\<^sup>-x\ + AOT_have \\\\[F]x\ + by (AOT_subst (reverse) \\[F]x\ \[F]\<^sup>-x\) + (auto simp add: "thm-relation-negation:1" 3) + AOT_hence \\[F]x\ + using "conventions:5"[THEN "\\<^sub>d\<^sub>fI"] by simp + AOT_hence \\[F]x\ using 0 "\E" "\E" by fast + AOT_thus \\[F]x & \\[F]x\ using "&I" 2 by blast + qed + next + fix x + AOT_assume 0: \\x (\[F]\<^sup>-x \ \[F]\<^sup>-x)\ + AOT_assume 1: \\[F]x\ + AOT_have \\\[F]\<^sup>-x\ + by (AOT_subst \\[F]\<^sup>-x\ \[F]x\) + (auto simp: "thm-relation-negation:2" 1) + AOT_hence 2: \\\[F]\<^sup>-x\ + using "KBasic:11"[THEN "\E"(2)] by blast + AOT_show \\[F]x\ + proof (rule "raa-cor:1") + AOT_assume 3: \\\[F]x\ + AOT_have \\\\[F]\<^sup>-x\ + by (AOT_subst \\[F]\<^sup>-x\ \[F]x\) + (auto simp add: "thm-relation-negation:2" 3) + AOT_hence \\[F]\<^sup>-x\ + using "conventions:5"[THEN "\\<^sub>d\<^sub>fI"] by simp + AOT_hence \\[F]\<^sup>-x\ using 0 "\E" "\E" by fast + AOT_thus \\[F]\<^sup>-x & \\[F]\<^sup>-x\ using "&I" 2 by blast + qed + qed + also AOT_have \\ \ WeaklyContingent([F]\<^sup>-)\ + using "df-cont-nec"[THEN "\Df", symmetric] by blast + finally show ?thesis. +qed + +AOT_theorem "cont-nec-fact1:2": + \(WeaklyContingent([F]) & \WeaklyContingent([G])) \ F \ G\ +proof (rule "\I"; rule "=-infix"[THEN "\\<^sub>d\<^sub>fI"]; rule "raa-cor:2") + AOT_assume 1: \WeaklyContingent([F]) & \WeaklyContingent([G])\ + AOT_hence \WeaklyContingent([F])\ using "&E" by blast + moreover AOT_assume \F = G\ + ultimately AOT_have \WeaklyContingent([G])\ + using "rule=E" by blast + AOT_thus \WeaklyContingent([G]) & \WeaklyContingent([G])\ + using 1 "&I" "&E" by blast +qed + +AOT_theorem "cont-nec-fact2:1": \WeaklyContingent(O!)\ +proof (rule "df-cont-nec"[THEN "\\<^sub>d\<^sub>fI"]; rule "&I") + AOT_show \Contingent(O!)\ + using "oa-contingent:4". +next + AOT_show \\x (\[O!]x \ \[O!]x)\ + apply (rule GEN; rule "\I") + using "oa-facts:5"[THEN "\E"(1)] by blast +qed + + +AOT_theorem "cont-nec-fact2:2": \WeaklyContingent(A!)\ +proof (rule "df-cont-nec"[THEN "\\<^sub>d\<^sub>fI"]; rule "&I") + AOT_show \Contingent(A!)\ + using "oa-contingent:5". +next + AOT_show \\x (\[A!]x \ \[A!]x)\ + apply (rule GEN; rule "\I") + using "oa-facts:6"[THEN "\E"(1)] by blast +qed + +AOT_theorem "cont-nec-fact2:3": \\WeaklyContingent(E!)\ +proof (rule "df-cont-nec"[THEN "\Df", + THEN "oth-class-taut:4:b"[THEN "\E"(1)], + THEN "\E"(2)]; + rule DeMorgan(1)[THEN "\E"(2)]; rule "\I"(2); rule "raa-cor:2") + AOT_have \\\x (E!x & \\<^bold>\E!x)\ using "qml:4"[axiom_inst]. + AOT_hence \\x \(E!x & \\<^bold>\E!x)\ using "BF\"[THEN "\E"] by blast + then AOT_obtain a where \\(E!a & \\<^bold>\E!a)\ using "\E"[rotated] by blast + AOT_hence 1: \\E!a & \\\<^bold>\E!a\ using "KBasic2:3"[THEN "\E"] by simp + moreover AOT_assume \\x (\[E!]x \ \[E!]x)\ + ultimately AOT_have \\E!a\ using "&E" "\E" "\E" by fast + AOT_hence \\<^bold>\E!a\ using "nec-imp-act"[THEN "\E"] by blast + AOT_hence \\\<^bold>\E!a\ using "qml-act:1"[axiom_inst, THEN "\E"] by blast + moreover AOT_have \\\\<^bold>\E!a\ + using "KBasic:11"[THEN "\E"(2)] 1[THEN "&E"(2)] by meson + ultimately AOT_have \\\<^bold>\E!a & \\\<^bold>\E!a\ using "&I" by blast + AOT_thus \p & \p\ for p using "raa-cor:1" by blast +qed + +AOT_theorem "cont-nec-fact2:4": \\WeaklyContingent(L)\ + apply (rule "df-cont-nec"[THEN "\Df", + THEN "oth-class-taut:4:b"[THEN "\E"(1)], + THEN "\E"(2)]; + rule DeMorgan(1)[THEN "\E"(2)]; rule "\I"(1)) + apply (rule "contingent-properties:4" + [THEN "\Df", + THEN "oth-class-taut:4:b"[THEN "\E"(1)], + THEN "\E"(2)]) + apply (rule DeMorgan(1)[THEN "\E"(2)]; + rule "\I"(2); + rule "useful-tautologies:2"[THEN "\E"]) + using "thm-noncont-e-e:3"[THEN "contingent-properties:3"[THEN "\\<^sub>d\<^sub>fE"]]. + +AOT_theorem "cont-nec-fact2:5": \O! \ E! & O! \ E!\<^sup>- & O! \ L & O! \ L\<^sup>-\ +proof - + AOT_have 1: \L\\ + by (rule "=\<^sub>d\<^sub>fI"(2)[OF L_def]) "cqt:2[lambda]"+ + { + fix \ and \ \' :: \<\>\ + AOT_have A: \\(\{\'} \ \{\})\ if \\{\}\ and \\\{\'}\ + proof (rule "raa-cor:2") + AOT_assume \\{\'} \ \{\}\ + AOT_hence \\{\'}\ using that(1) "\E" by blast + AOT_thus \\{\'} & \\{\'}\ using that(2) "&I" by blast + qed + AOT_have \\' \ \\ if \\\\ and \\'\\ and \\{\}\ and \\\{\'}\ + using "pos-not-equiv-ne:4"[unvarify F G, THEN "\E", + OF that(1,2), OF A[OF that(3, 4)]]. + } note 0 = this + show ?thesis + apply(safe intro!: "&I"; rule 0) + apply "cqt:2" + using "oa-exist:1" apply blast + using "cont-nec-fact2:3" apply fast + apply (rule "useful-tautologies:2"[THEN "\E"]) + using "cont-nec-fact2:1" apply fast + using "rel-neg-T:3" apply fast + using "oa-exist:1" apply blast + using "cont-nec-fact1:1"[THEN "oth-class-taut:4:b"[THEN "\E"(1)], + THEN "\E"(1), rotated, OF "cont-nec-fact2:3"] apply fast + apply (rule "useful-tautologies:2"[THEN "\E"]) + using "cont-nec-fact2:1" apply blast + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF L_def]; "cqt:2[lambda]") + using "oa-exist:1" apply fast + using "cont-nec-fact2:4" apply fast + apply (rule "useful-tautologies:2"[THEN "\E"]) + using "cont-nec-fact2:1" apply fast + using "rel-neg-T:3" apply fast + using "oa-exist:1" apply fast + apply (rule "cont-nec-fact1:1"[unvarify F, + THEN "oth-class-taut:4:b"[THEN "\E"(1)], + THEN "\E"(1), rotated, OF "cont-nec-fact2:4"]) + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF L_def]; "cqt:2[lambda]") + apply (rule "useful-tautologies:2"[THEN "\E"]) + using "cont-nec-fact2:1" by blast +qed + +AOT_theorem "cont-nec-fact2:6": \A! \ E! & A! \ E!\<^sup>- & A! \ L & A! \ L\<^sup>-\ +proof - + AOT_have 1: \L\\ + by (rule "=\<^sub>d\<^sub>fI"(2)[OF L_def]) "cqt:2[lambda]"+ + { + fix \ and \ \' :: \<\>\ + AOT_have A: \\(\{\'} \ \{\})\ if \\{\}\ and \\\{\'}\ + proof (rule "raa-cor:2") + AOT_assume \\{\'} \ \{\}\ + AOT_hence \\{\'}\ using that(1) "\E" by blast + AOT_thus \\{\'} & \\{\'}\ using that(2) "&I" by blast + qed + AOT_have \\' \ \\ if \\\\ and \\'\\ and \\{\}\ and \\\{\'}\ + using "pos-not-equiv-ne:4"[unvarify F G, THEN "\E", + OF that(1,2), OF A[OF that(3, 4)]]. + } note 0 = this + show ?thesis + apply(safe intro!: "&I"; rule 0) + apply "cqt:2" + using "oa-exist:2" apply blast + using "cont-nec-fact2:3" apply fast + apply (rule "useful-tautologies:2"[THEN "\E"]) + using "cont-nec-fact2:2" apply fast + using "rel-neg-T:3" apply fast + using "oa-exist:2" apply blast + using "cont-nec-fact1:1"[THEN "oth-class-taut:4:b"[THEN "\E"(1)], + THEN "\E"(1), rotated, OF "cont-nec-fact2:3"] apply fast + apply (rule "useful-tautologies:2"[THEN "\E"]) + using "cont-nec-fact2:2" apply blast + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF L_def]; "cqt:2[lambda]") + using "oa-exist:2" apply fast + using "cont-nec-fact2:4" apply fast + apply (rule "useful-tautologies:2"[THEN "\E"]) + using "cont-nec-fact2:2" apply fast + using "rel-neg-T:3" apply fast + using "oa-exist:2" apply fast + apply (rule "cont-nec-fact1:1"[unvarify F, + THEN "oth-class-taut:4:b"[THEN "\E"(1)], + THEN "\E"(1), rotated, OF "cont-nec-fact2:4"]) + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF L_def]; "cqt:2[lambda]") + apply (rule "useful-tautologies:2"[THEN "\E"]) + using "cont-nec-fact2:2" by blast +qed + +AOT_define necessary_or_contingently_false :: \\ \ \\ ("\<^bold>\_" [49] 54) + \\<^bold>\p \\<^sub>d\<^sub>f \p \ (\\<^bold>\p & \p)\ + +AOT_theorem sixteen: + shows \\F\<^sub>1\F\<^sub>2\F\<^sub>3\F\<^sub>4\F\<^sub>5\F\<^sub>6\F\<^sub>7\F\<^sub>8\F\<^sub>9\F\<^sub>1\<^sub>0\F\<^sub>1\<^sub>1\F\<^sub>1\<^sub>2\F\<^sub>1\<^sub>3\F\<^sub>1\<^sub>4\F\<^sub>1\<^sub>5\F\<^sub>1\<^sub>6 ( + \F\<^sub>1::<\>\ \ F\<^sub>2 & F\<^sub>1 \ F\<^sub>3 & F\<^sub>1 \ F\<^sub>4 & F\<^sub>1 \ F\<^sub>5 & F\<^sub>1 \ F\<^sub>6 & F\<^sub>1 \ F\<^sub>7 & + F\<^sub>1 \ F\<^sub>8 & F\<^sub>1 \ F\<^sub>9 & F\<^sub>1 \ F\<^sub>1\<^sub>0 & F\<^sub>1 \ F\<^sub>1\<^sub>1 & F\<^sub>1 \ F\<^sub>1\<^sub>2 & F\<^sub>1 \ F\<^sub>1\<^sub>3 & + F\<^sub>1 \ F\<^sub>1\<^sub>4 & F\<^sub>1 \ F\<^sub>1\<^sub>5 & F\<^sub>1 \ F\<^sub>1\<^sub>6 & + F\<^sub>2 \ F\<^sub>3 & F\<^sub>2 \ F\<^sub>4 & F\<^sub>2 \ F\<^sub>5 & F\<^sub>2 \ F\<^sub>6 & F\<^sub>2 \ F\<^sub>7 & F\<^sub>2 \ F\<^sub>8 & + F\<^sub>2 \ F\<^sub>9 & F\<^sub>2 \ F\<^sub>1\<^sub>0 & F\<^sub>2 \ F\<^sub>1\<^sub>1 & F\<^sub>2 \ F\<^sub>1\<^sub>2 & F\<^sub>2 \ F\<^sub>1\<^sub>3 & F\<^sub>2 \ F\<^sub>1\<^sub>4 & + F\<^sub>2 \ F\<^sub>1\<^sub>5 & F\<^sub>2 \ F\<^sub>1\<^sub>6 & + F\<^sub>3 \ F\<^sub>4 & F\<^sub>3 \ F\<^sub>5 & F\<^sub>3 \ F\<^sub>6 & F\<^sub>3 \ F\<^sub>7 & F\<^sub>3 \ F\<^sub>8 & F\<^sub>3 \ F\<^sub>9 & F\<^sub>3 \ F\<^sub>1\<^sub>0 & + F\<^sub>3 \ F\<^sub>1\<^sub>1 & F\<^sub>3 \ F\<^sub>1\<^sub>2 & F\<^sub>3 \ F\<^sub>1\<^sub>3 & F\<^sub>3 \ F\<^sub>1\<^sub>4 & F\<^sub>3 \ F\<^sub>1\<^sub>5 & F\<^sub>3 \ F\<^sub>1\<^sub>6 & + F\<^sub>4 \ F\<^sub>5 & F\<^sub>4 \ F\<^sub>6 & F\<^sub>4 \ F\<^sub>7 & F\<^sub>4 \ F\<^sub>8 & F\<^sub>4 \ F\<^sub>9 & F\<^sub>4 \ F\<^sub>1\<^sub>0 & F\<^sub>4 \ F\<^sub>1\<^sub>1 & + F\<^sub>4 \ F\<^sub>1\<^sub>2 & F\<^sub>4 \ F\<^sub>1\<^sub>3 & F\<^sub>4 \ F\<^sub>1\<^sub>4 & F\<^sub>4 \ F\<^sub>1\<^sub>5 & F\<^sub>4 \ F\<^sub>1\<^sub>6 & + F\<^sub>5 \ F\<^sub>6 & F\<^sub>5 \ F\<^sub>7 & F\<^sub>5 \ F\<^sub>8 & F\<^sub>5 \ F\<^sub>9 & F\<^sub>5 \ F\<^sub>1\<^sub>0 & F\<^sub>5 \ F\<^sub>1\<^sub>1 & F\<^sub>5 \ F\<^sub>1\<^sub>2 & + F\<^sub>5 \ F\<^sub>1\<^sub>3 & F\<^sub>5 \ F\<^sub>1\<^sub>4 & F\<^sub>5 \ F\<^sub>1\<^sub>5 & F\<^sub>5 \ F\<^sub>1\<^sub>6 & + F\<^sub>6 \ F\<^sub>7 & F\<^sub>6 \ F\<^sub>8 & F\<^sub>6 \ F\<^sub>9 & F\<^sub>6 \ F\<^sub>1\<^sub>0 & F\<^sub>6 \ F\<^sub>1\<^sub>1 & F\<^sub>6 \ F\<^sub>1\<^sub>2 & F\<^sub>6 \ F\<^sub>1\<^sub>3 & + F\<^sub>6 \ F\<^sub>1\<^sub>4 & F\<^sub>6 \ F\<^sub>1\<^sub>5 & F\<^sub>6 \ F\<^sub>1\<^sub>6 & + F\<^sub>7 \ F\<^sub>8 & F\<^sub>7 \ F\<^sub>9 & F\<^sub>7 \ F\<^sub>1\<^sub>0 & F\<^sub>7 \ F\<^sub>1\<^sub>1 & F\<^sub>7 \ F\<^sub>1\<^sub>2 & F\<^sub>7 \ F\<^sub>1\<^sub>3 & F\<^sub>7 \ F\<^sub>1\<^sub>4 & + F\<^sub>7 \ F\<^sub>1\<^sub>5 & F\<^sub>7 \ F\<^sub>1\<^sub>6 & + F\<^sub>8 \ F\<^sub>9 & F\<^sub>8 \ F\<^sub>1\<^sub>0 & F\<^sub>8 \ F\<^sub>1\<^sub>1 & F\<^sub>8 \ F\<^sub>1\<^sub>2 & F\<^sub>8 \ F\<^sub>1\<^sub>3 & F\<^sub>8 \ F\<^sub>1\<^sub>4 & F\<^sub>8 \ F\<^sub>1\<^sub>5 & + F\<^sub>8 \ F\<^sub>1\<^sub>6 & + F\<^sub>9 \ F\<^sub>1\<^sub>0 & F\<^sub>9 \ F\<^sub>1\<^sub>1 & F\<^sub>9 \ F\<^sub>1\<^sub>2 & F\<^sub>9 \ F\<^sub>1\<^sub>3 & F\<^sub>9 \ F\<^sub>1\<^sub>4 & F\<^sub>9 \ F\<^sub>1\<^sub>5 & F\<^sub>9 \ F\<^sub>1\<^sub>6 & + F\<^sub>1\<^sub>0 \ F\<^sub>1\<^sub>1 & F\<^sub>1\<^sub>0 \ F\<^sub>1\<^sub>2 & F\<^sub>1\<^sub>0 \ F\<^sub>1\<^sub>3 & F\<^sub>1\<^sub>0 \ F\<^sub>1\<^sub>4 & F\<^sub>1\<^sub>0 \ F\<^sub>1\<^sub>5 & F\<^sub>1\<^sub>0 \ F\<^sub>1\<^sub>6 & + F\<^sub>1\<^sub>1 \ F\<^sub>1\<^sub>2 & F\<^sub>1\<^sub>1 \ F\<^sub>1\<^sub>3 & F\<^sub>1\<^sub>1 \ F\<^sub>1\<^sub>4 & F\<^sub>1\<^sub>1 \ F\<^sub>1\<^sub>5 & F\<^sub>1\<^sub>1 \ F\<^sub>1\<^sub>6 & + F\<^sub>1\<^sub>2 \ F\<^sub>1\<^sub>3 & F\<^sub>1\<^sub>2 \ F\<^sub>1\<^sub>4 & F\<^sub>1\<^sub>2 \ F\<^sub>1\<^sub>5 & F\<^sub>1\<^sub>2 \ F\<^sub>1\<^sub>6 & + F\<^sub>1\<^sub>3 \ F\<^sub>1\<^sub>4 & F\<^sub>1\<^sub>3 \ F\<^sub>1\<^sub>5 & F\<^sub>1\<^sub>3 \ F\<^sub>1\<^sub>6 & + F\<^sub>1\<^sub>4 \ F\<^sub>1\<^sub>5 & F\<^sub>1\<^sub>4 \ F\<^sub>1\<^sub>6 & + F\<^sub>1\<^sub>5 \ F\<^sub>1\<^sub>6)\ +proof - + AOT_have Delta_pos: \\<^bold>\\ \ \\\ for \ + proof(rule "\I") + AOT_assume \\<^bold>\\\ + AOT_hence \\\ \ (\\<^bold>\\ & \\)\ + using "\\<^sub>d\<^sub>fE"[OF necessary_or_contingently_false] by blast + moreover { + AOT_assume \\\\ + AOT_hence \\\\ + by (metis "B\" "T\" "vdash-properties:10") + } + moreover { + AOT_assume \\\<^bold>\\ & \\\ + AOT_hence \\\\ + using "&E" by blast + } + ultimately AOT_show \\\\ + by (metis "\E"(2) "raa-cor:1") + qed + + AOT_have act_and_not_nec_not_delta: \\\<^bold>\\\ if \\<^bold>\\\ and \\\\\ for \ + using "\\<^sub>d\<^sub>fE" "&E"(1) "\E"(2) necessary_or_contingently_false + "raa-cor:3" that(1,2) by blast + AOT_have act_and_pos_not_not_delta: \\\<^bold>\\\ if \\<^bold>\\\ and \\\\\ for \ + using "KBasic:11" act_and_not_nec_not_delta "\E"(2) that(1,2) by blast + AOT_have impossible_delta: \\\<^bold>\\\ if \\\\\ for \ + using Delta_pos "modus-tollens:1" that by blast + AOT_have not_act_and_pos_delta: \\<^bold>\\\ if \\\<^bold>\\\ and \\\\ for \ + by (meson "\\<^sub>d\<^sub>fI" "&I" "\I"(2) necessary_or_contingently_false that(1,2)) + AOT_have nec_delta: \\<^bold>\\\ if \\\\ for \ + using "\\<^sub>d\<^sub>fI" "\I"(1) necessary_or_contingently_false that by blast + + AOT_obtain a where a_prop: \A!a\ + using "A-objects"[axiom_inst] "\E"[rotated] "&E" by blast + AOT_obtain b where b_prop: \\[E!]b & \\<^bold>\[E!]b\ + using "pos-not-pna:3" using "\E"[rotated] by blast + + AOT_have b_ord: \[O!]b\ + proof(rule "=\<^sub>d\<^sub>fI"(2)[OF AOT_ordinary]) + AOT_show \[\x \[E!]x]\\ by "cqt:2[lambda]" + next + AOT_show \[\x \[E!]x]b\ + proof (rule "\\C"(1); ("cqt:2[lambda]")?) + AOT_show \b\\ by (rule "cqt:2[const_var]"[axiom_inst]) + AOT_show \\[E!]b\ by (fact b_prop[THEN "&E"(1)]) + qed + qed + + AOT_have nec_not_L_neg: \\\[L\<^sup>-]x\ for x + using "thm-noncont-e-e:2" "contingent-properties:2"[THEN "\\<^sub>d\<^sub>fE"] "&E" + CBF[THEN "\E"] "\E" by blast + AOT_have nec_L: \\[L]x\ for x + using "thm-noncont-e-e:1" "contingent-properties:1"[THEN "\\<^sub>d\<^sub>fE"] + CBF[THEN "\E"] "\E" by blast + + AOT_have act_ord_b: \\<^bold>\[O!]b\ + using b_ord "\E"(1) "oa-facts:7" by blast + AOT_have delta_ord_b: \\<^bold>\[O!]b\ + by (meson "\\<^sub>d\<^sub>fI" b_ord "\I"(1) necessary_or_contingently_false + "oa-facts:1" "\E") + AOT_have not_act_ord_a: \\\<^bold>\[O!]a\ + by (meson a_prop "\E"(1) "\E"(3) "oa-contingent:3" "oa-facts:7") + AOT_have not_delta_ord_a: \\\<^bold>\[O!]a\ + by (metis Delta_pos "\E"(4) not_act_ord_a "oa-facts:3" "oa-facts:7" + "reductio-aa:1" "\E") + + AOT_have not_act_abs_b: \\\<^bold>\[A!]b\ + by (meson b_ord "\E"(1) "\E"(3) "oa-contingent:2" "oa-facts:8") + AOT_have not_delta_abs_b: \\\<^bold>\[A!]b\ + proof(rule "raa-cor:2") + AOT_assume \\<^bold>\[A!]b\ + AOT_hence \\[A!]b\ + by (metis Delta_pos "vdash-properties:10") + AOT_thus \[A!]b & \[A!]b\ + by (metis b_ord "&I" "\E"(1) "oa-contingent:2" + "oa-facts:4" "\E") + qed + AOT_have act_abs_a: \\<^bold>\[A!]a\ + using a_prop "\E"(1) "oa-facts:8" by blast + AOT_have delta_abs_a: \\<^bold>\[A!]a\ + by (metis "\\<^sub>d\<^sub>fI" a_prop "oa-facts:2" "\E" "\I"(1) + necessary_or_contingently_false) + + AOT_have not_act_concrete_b: \\\<^bold>\[E!]b\ + using b_prop "&E"(2) by blast + AOT_have delta_concrete_b: \\<^bold>\[E!]b\ + proof (rule "\\<^sub>d\<^sub>fI"[OF necessary_or_contingently_false]; + rule "\I"(2); rule "&I") + AOT_show \\\<^bold>\[E!]b\ using b_prop "&E"(2) by blast + next + AOT_show \\[E!]b\ using b_prop "&E"(1) by blast + qed + AOT_have not_act_concrete_a: \\\<^bold>\[E!]a\ + proof (rule "raa-cor:2") + AOT_assume \\<^bold>\[E!]a\ + AOT_hence 1: \\[E!]a\ by (metis "Act-Sub:3" "\E") + AOT_have \[A!]a\ by (simp add: a_prop) + AOT_hence \[\x \\[E!]x]a\ + by (rule "=\<^sub>d\<^sub>fE"(2)[OF AOT_abstract, rotated]) "cqt:2" + AOT_hence \\\[E!]a\ using "\\C"(1) by blast + AOT_thus \\[E!]a & \\[E!]a\ using 1 "&I" by blast + qed + AOT_have not_delta_concrete_a: \\\<^bold>\[E!]a\ + proof (rule "raa-cor:2") + AOT_assume \\<^bold>\[E!]a\ + AOT_hence 1: \\[E!]a\ by (metis Delta_pos "vdash-properties:10") + AOT_have \[A!]a\ by (simp add: a_prop) + AOT_hence \[\x \\[E!]x]a\ + by (rule "=\<^sub>d\<^sub>fE"(2)[OF AOT_abstract, rotated]) "cqt:2[lambda]" + AOT_hence \\\[E!]a\ using "\\C"(1) by blast + AOT_thus \\[E!]a & \\[E!]a\ using 1 "&I" by blast + qed + + AOT_have not_act_q_zero: \\\<^bold>\q\<^sub>0\ + by (meson "log-prop-prop:2" "pos-not-pna:1" + q\<^sub>0_def "reductio-aa:1" "rule-id-df:2:a[zero]") + AOT_have delta_q_zero: \\<^bold>\q\<^sub>0\ + proof(rule "\\<^sub>d\<^sub>fI"[OF necessary_or_contingently_false]; + rule "\I"(2); rule "&I") + AOT_show \\\<^bold>\q\<^sub>0\ using not_act_q_zero. + AOT_show \\q\<^sub>0\ by (meson "&E"(1) q\<^sub>0_prop) + qed + AOT_have act_not_q_zero: \\<^bold>\\q\<^sub>0\ + using "Act-Basic:1" "\E"(2) not_act_q_zero by blast + AOT_have not_delta_not_q_zero: \\\<^bold>\\q\<^sub>0\ + using "\\<^sub>d\<^sub>fE" "conventions:5" "Act-Basic:1" act_and_not_nec_not_delta + "&E"(1) "\E"(2) not_act_q_zero q\<^sub>0_prop by blast + + AOT_have \[L\<^sup>-]\\ by (simp add: "rel-neg-T:3") + moreover AOT_have \\\<^bold>\[L\<^sup>-]b & \\<^bold>\[L\<^sup>-]b & \\<^bold>\[L\<^sup>-]a & \\<^bold>\[L\<^sup>-]a\ + proof (safe intro!: "&I") + AOT_show \\\<^bold>\[L\<^sup>-]b\ + by (meson "\E"(1) "logic-actual-nec:1"[axiom_inst] "nec-imp-act" + nec_not_L_neg "\E") + AOT_show \\\<^bold>\[L\<^sup>-]b\ + by (meson Delta_pos "KBasic2:1" "\E"(1) + "modus-tollens:1" nec_not_L_neg) + AOT_show \\\<^bold>\[L\<^sup>-]a\ + by (meson "\E"(1) "logic-actual-nec:1"[axiom_inst] + "nec-imp-act" nec_not_L_neg "\E") + AOT_show \\\<^bold>\[L\<^sup>-]a\ + using Delta_pos "KBasic2:1" "\E"(1) "modus-tollens:1" + nec_not_L_neg by blast + qed + ultimately AOT_obtain F\<^sub>0 where \\\<^bold>\[F\<^sub>0]b & \\<^bold>\[F\<^sub>0]b & \\<^bold>\[F\<^sub>0]a & \\<^bold>\[F\<^sub>0]a\ + using "\I"(1)[rotated, THEN "\E"[rotated]] by fastforce + AOT_hence \\\<^bold>\[F\<^sub>0]b\ and \\\<^bold>\[F\<^sub>0]b\ and \\\<^bold>\[F\<^sub>0]a\ and \\\<^bold>\[F\<^sub>0]a\ + using "&E" by blast+ + note props = this + + let ?\ = "\[\y [A!]y & q\<^sub>0]\" + AOT_modally_strict { + AOT_have \[\?\\]\\ by "cqt:2[lambda]" + } note 1 = this + moreover AOT_have \\\<^bold>\[\?\\]b & \\<^bold>\[\?\\]b & \\<^bold>\[\?\\]a & \<^bold>\[\?\\]a\ + proof (safe intro!: "&I"; AOT_subst \[\y A!y & q\<^sub>0]x\ \A!x & q\<^sub>0\ for: x) + AOT_show \\\<^bold>\([A!]b & q\<^sub>0)\ + using "Act-Basic:2" "&E"(1) "\E"(1) not_act_abs_b "raa-cor:3" by blast + next AOT_show \\\<^bold>\([A!]b & q\<^sub>0)\ + by (metis Delta_pos "KBasic2:3" "&E"(1) "\E"(4) not_act_abs_b + "oa-facts:4" "oa-facts:8" "raa-cor:3" "\E") + next AOT_show \\\<^bold>\([A!]a & q\<^sub>0)\ + using "Act-Basic:2" "&E"(2) "\E"(1) not_act_q_zero + "raa-cor:3" by blast + next AOT_show \\<^bold>\([A!]a & q\<^sub>0)\ + proof (rule not_act_and_pos_delta) + AOT_show \\\<^bold>\([A!]a & q\<^sub>0)\ + using "Act-Basic:2" "&E"(2) "\E"(4) not_act_q_zero + "raa-cor:3" by blast + next AOT_show \\([A!]a & q\<^sub>0)\ + by (metis "&I" "\E" Delta_pos "KBasic:16" "&E"(1) delta_abs_a + "\E"(1) "oa-facts:6" q\<^sub>0_prop) + qed + qed(auto simp: "beta-C-meta"[THEN "\E", OF 1]) + ultimately AOT_obtain F\<^sub>1 where \\\<^bold>\[F\<^sub>1]b & \\<^bold>\[F\<^sub>1]b & \\<^bold>\[F\<^sub>1]a & \<^bold>\[F\<^sub>1]a\ + using "\I"(1)[rotated, THEN "\E"[rotated]] by fastforce + AOT_hence \\\<^bold>\[F\<^sub>1]b\ and \\\<^bold>\[F\<^sub>1]b\ and \\\<^bold>\[F\<^sub>1]a\ and \\<^bold>\[F\<^sub>1]a\ + using "&E" by blast+ + note props = props this + + let ?\ = "\[\y [A!]y & \q\<^sub>0]\" + AOT_modally_strict { + AOT_have \[\?\\]\\ by "cqt:2[lambda]" + } note 1 = this + moreover AOT_have \\\<^bold>\[\?\\]b & \\<^bold>\[\?\\]b & \<^bold>\[\?\\]a & \\<^bold>\[\?\\]a\ + proof (safe intro!: "&I"; AOT_subst \[\y A!y & \q\<^sub>0]x\ \A!x & \q\<^sub>0\ for: x) + AOT_show \\\<^bold>\([A!]b & \q\<^sub>0)\ + using "Act-Basic:2" "&E"(1) "\E"(1) not_act_abs_b "raa-cor:3" by blast + next AOT_show \\\<^bold>\([A!]b & \q\<^sub>0)\ + by (meson "RM\" Delta_pos "Conjunction Simplification"(1) "\E"(4) + "modus-tollens:1" not_act_abs_b "oa-facts:4" "oa-facts:8") + next AOT_show \\<^bold>\([A!]a & \q\<^sub>0)\ + by (metis "Act-Basic:1" "Act-Basic:2" act_abs_a "&I" "\E"(2) + "\E"(3) not_act_q_zero "raa-cor:3") + next AOT_show \\\<^bold>\([A!]a & \q\<^sub>0)\ + proof (rule act_and_not_nec_not_delta) + AOT_show \\<^bold>\([A!]a & \q\<^sub>0)\ + by (metis "Act-Basic:1" "Act-Basic:2" act_abs_a "&I" "\E"(2) + "\E"(3) not_act_q_zero "raa-cor:3") + next + AOT_show \\\([A!]a & \q\<^sub>0)\ + by (metis "KBasic2:1" "KBasic:3" "&E"(1) "&E"(2) "\E"(4) + q\<^sub>0_prop "raa-cor:3") + qed + qed(auto simp: "beta-C-meta"[THEN "\E", OF 1]) + ultimately AOT_obtain F\<^sub>2 where \\\<^bold>\[F\<^sub>2]b & \\<^bold>\[F\<^sub>2]b & \<^bold>\[F\<^sub>2]a & \\<^bold>\[F\<^sub>2]a\ + using "\I"(1)[rotated, THEN "\E"[rotated]] by fastforce + AOT_hence \\\<^bold>\[F\<^sub>2]b\ and \\\<^bold>\[F\<^sub>2]b\ and \\<^bold>\[F\<^sub>2]a\ and \\\<^bold>\[F\<^sub>2]a\ + using "&E" by blast+ + note props = props this + + AOT_have abstract_prop: \\\<^bold>\[A!]b & \\<^bold>\[A!]b & \<^bold>\[A!]a & \<^bold>\[A!]a\ + using act_abs_a "&I" delta_abs_a not_act_abs_b not_delta_abs_b + by presburger + then AOT_obtain F\<^sub>3 where \\\<^bold>\[F\<^sub>3]b & \\<^bold>\[F\<^sub>3]b & \<^bold>\[F\<^sub>3]a & \<^bold>\[F\<^sub>3]a\ + using "\I"(1)[rotated, THEN "\E"[rotated]] "oa-exist:2" by fastforce + AOT_hence \\\<^bold>\[F\<^sub>3]b\ and \\\<^bold>\[F\<^sub>3]b\ and \\<^bold>\[F\<^sub>3]a\ and \\<^bold>\[F\<^sub>3]a\ + using "&E" by blast+ + note props = props this + + AOT_have \\\<^bold>\[E!]b & \<^bold>\[E!]b & \\<^bold>\[E!]a & \\<^bold>\[E!]a\ + by (meson "&I" delta_concrete_b not_act_concrete_a + not_act_concrete_b not_delta_concrete_a) + then AOT_obtain F\<^sub>4 where \\\<^bold>\[F\<^sub>4]b & \<^bold>\[F\<^sub>4]b & \\<^bold>\[F\<^sub>4]a & \\<^bold>\[F\<^sub>4]a\ + using "\I"(1)[rotated, THEN "\E"[rotated]] + by fastforce + AOT_hence \\\<^bold>\[F\<^sub>4]b\ and \\<^bold>\[F\<^sub>4]b\ and \\\<^bold>\[F\<^sub>4]a\ and \\\<^bold>\[F\<^sub>4]a\ + using "&E" by blast+ + note props = props this + + AOT_modally_strict { + AOT_have \[\y q\<^sub>0]\\ by "cqt:2[lambda]" + } note 1 = this + moreover AOT_have \\\<^bold>\[\y q\<^sub>0]b & \<^bold>\[\y q\<^sub>0]b & \\<^bold>\[\y q\<^sub>0]a & \<^bold>\[\y q\<^sub>0]a\ + by (safe intro!: "&I"; AOT_subst \[\y q\<^sub>0]b\ \q\<^sub>0\ for: b) + (auto simp: not_act_q_zero delta_q_zero "beta-C-meta"[THEN "\E", OF 1]) + ultimately AOT_obtain F\<^sub>5 where \\\<^bold>\[F\<^sub>5]b & \<^bold>\[F\<^sub>5]b & \\<^bold>\[F\<^sub>5]a & \<^bold>\[F\<^sub>5]a\ + using "\I"(1)[rotated, THEN "\E"[rotated]] + by fastforce + AOT_hence \\\<^bold>\[F\<^sub>5]b\ and \\<^bold>\[F\<^sub>5]b\ and \\\<^bold>\[F\<^sub>5]a\ and \\<^bold>\[F\<^sub>5]a\ + using "&E" by blast+ + note props = props this + + let ?\ = "\[\y [E!]y \ ([A!]y & \q\<^sub>0)]\" + AOT_modally_strict { + AOT_have \[\?\\]\\ by "cqt:2[lambda]" + } note 1 = this + moreover AOT_have \\\<^bold>\[\?\\]b & \<^bold>\[\?\\]b & \<^bold>\[\?\\]a & \\<^bold>\[\?\\]a\ + proof(safe intro!: "&I"; + AOT_subst \[\y E!y \ (A!y & \q\<^sub>0)]x\ \E!x \ (A!x & \q\<^sub>0)\ for: x) + AOT_have \\<^bold>\\([A!]b & \q\<^sub>0)\ + by (metis "Act-Basic:1" "Act-Basic:2" abstract_prop "&E"(1) "\E"(2) + "\E"(1) "raa-cor:3") + moreover AOT_have \\\<^bold>\[E!]b\ + using b_prop "&E"(2) by blast + ultimately AOT_have 2: \\<^bold>\(\[E!]b & \([A!]b & \q\<^sub>0))\ + by (metis "Act-Basic:2" "Act-Sub:1" "&I" "\E"(3) "raa-cor:1") + AOT_have \\<^bold>\\([E!]b \ ([A!]b & \q\<^sub>0))\ + by (AOT_subst \\([E!]b \ ([A!]b & \q\<^sub>0))\ \\[E!]b & \([A!]b & \q\<^sub>0)\) + (auto simp: "oth-class-taut:5:d" 2) + AOT_thus \\\<^bold>\([E!]b \ ([A!]b & \q\<^sub>0))\ + by (metis "\\I" "Act-Sub:1" "\E"(4)) + next + AOT_show \\<^bold>\([E!]b \ ([A!]b & \q\<^sub>0))\ + proof (rule not_act_and_pos_delta) + AOT_show \\\<^bold>\([E!]b \ ([A!]b & \q\<^sub>0))\ + by (metis "Act-Basic:2" "Act-Basic:9" "\E"(2) "raa-cor:3" + "Conjunction Simplification"(1) "\E"(4) + "modus-tollens:1" not_act_abs_b not_act_concrete_b) + next + AOT_show \\([E!]b \ ([A!]b & \q\<^sub>0))\ + using "KBasic2:2" b_prop "&E"(1) "\I"(1) "\E"(3) "raa-cor:3" by blast + qed + next AOT_show \\<^bold>\([E!]a \ ([A!]a & \q\<^sub>0))\ + by (metis "Act-Basic:1" "Act-Basic:2" "Act-Basic:9" act_abs_a "&I" + "\I"(2) "\E"(2) "\E"(3) not_act_q_zero "raa-cor:1") + next AOT_show \\\<^bold>\([E!]a \ ([A!]a & \q\<^sub>0))\ + proof (rule act_and_not_nec_not_delta) + AOT_show \\<^bold>\([E!]a \ ([A!]a & \q\<^sub>0))\ + by (metis "Act-Basic:1" "Act-Basic:2" "Act-Basic:9" act_abs_a "&I" + "\I"(2) "\E"(2) "\E"(3) not_act_q_zero "raa-cor:1") + next + AOT_have \\\[E!]a\ + by (metis "\\<^sub>d\<^sub>fI" "conventions:5" "&I" "\I"(2) + necessary_or_contingently_false + not_act_concrete_a not_delta_concrete_a "raa-cor:3") + moreover AOT_have \\\([A!]a & \q\<^sub>0)\ + by (metis "KBasic2:1" "KBasic:11" "KBasic:3" + "&E"(1,2) "\E"(1) q\<^sub>0_prop "raa-cor:3") + ultimately AOT_have \\(\[E!]a & \([A!]a & \q\<^sub>0))\ + by (metis "KBasic:16" "&I" "vdash-properties:10") + AOT_hence \\\([E!]a \ ([A!]a & \q\<^sub>0))\ + by (metis "RE\" "\E"(2) "oth-class-taut:5:d") + AOT_thus \\\([E!]a \ ([A!]a & \q\<^sub>0))\ + by (metis "KBasic:12" "\E"(1) "raa-cor:3") + qed + qed(auto simp: "beta-C-meta"[THEN "\E", OF 1]) + ultimately AOT_obtain F\<^sub>6 where \\\<^bold>\[F\<^sub>6]b & \<^bold>\[F\<^sub>6]b & \<^bold>\[F\<^sub>6]a & \\<^bold>\[F\<^sub>6]a\ + using "\I"(1)[rotated, THEN "\E"[rotated]] by fastforce + AOT_hence \\\<^bold>\[F\<^sub>6]b\ and \\<^bold>\[F\<^sub>6]b\ and \\<^bold>\[F\<^sub>6]a\ and \\\<^bold>\[F\<^sub>6]a\ + using "&E" by blast+ + note props = props this + + let ?\ = "\[\y [A!]y \ [E!]y]\" + AOT_modally_strict { + AOT_have \[\?\\]\\ by "cqt:2[lambda]" + } note 1 = this + moreover AOT_have \\\<^bold>\[\?\\]b & \<^bold>\[\?\\]b & \<^bold>\[\?\\]a & \<^bold>\[\?\\]a\ + proof(safe intro!: "&I"; AOT_subst \[\y A!y \ E!y]x\ \A!x \ E!x\ for: x) + AOT_show \\\<^bold>\([A!]b \ [E!]b)\ + using "Act-Basic:9" "\E"(2) "\E"(4) not_act_abs_b + not_act_concrete_b "raa-cor:3" by blast + next AOT_show \\<^bold>\([A!]b \ [E!]b)\ + proof (rule not_act_and_pos_delta) + AOT_show \\\<^bold>\([A!]b \ [E!]b)\ + using "Act-Basic:9" "\E"(2) "\E"(4) not_act_abs_b + not_act_concrete_b "raa-cor:3" by blast + next AOT_show \\([A!]b \ [E!]b)\ + using "KBasic2:2" b_prop "&E"(1) "\I"(2) "\E"(2) by blast + qed + next AOT_show \\<^bold>\([A!]a \ [E!]a)\ + by (meson "Act-Basic:9" act_abs_a "\I"(1) "\E"(2)) + next AOT_show \\<^bold>\([A!]a \ [E!]a) \ + proof (rule nec_delta) + AOT_show \\([A!]a \ [E!]a)\ + by (metis "KBasic:15" act_abs_a act_and_not_nec_not_delta + "Disjunction Addition"(1) delta_abs_a "raa-cor:3" "\E") + qed + qed(auto simp: "beta-C-meta"[THEN "\E", OF 1]) + ultimately AOT_obtain F\<^sub>7 where \\\<^bold>\[F\<^sub>7]b & \<^bold>\[F\<^sub>7]b & \<^bold>\[F\<^sub>7]a & \<^bold>\[F\<^sub>7]a\ + using "\I"(1)[rotated, THEN "\E"[rotated]] by fastforce + AOT_hence \\\<^bold>\[F\<^sub>7]b\ and \\<^bold>\[F\<^sub>7]b\ and \\<^bold>\[F\<^sub>7]a\ and \\<^bold>\[F\<^sub>7]a\ + using "&E" by blast+ + note props = props this + + let ?\ = "\[\y [O!]y & \[E!]y]\" + AOT_modally_strict { + AOT_have \[\?\\]\\ by "cqt:2[lambda]" + } note 1 = this + moreover AOT_have \\<^bold>\[\?\\]b & \\<^bold>\[\?\\]b & \\<^bold>\[\?\\]a & \\<^bold>\[\?\\]a\ + proof(safe intro!: "&I"; AOT_subst \[\y O!y & \E!y]x\ \O!x & \E!x\ for: x) + AOT_show \\<^bold>\([O!]b & \[E!]b)\ + by (metis "Act-Basic:1" "Act-Basic:2" act_ord_b "&I" "\E"(2) + "\E"(3) not_act_concrete_b "raa-cor:3") + next AOT_show \\\<^bold>\([O!]b & \[E!]b)\ + by (metis (no_types, opaque_lifting) "conventions:5" "Act-Sub:1" "RM:1" + act_and_not_nec_not_delta "act-conj-act:3" + act_ord_b b_prop "&I" "&E"(1) "Conjunction Simplification"(2) + "df-rules-formulas[3]" + "\E"(3) "raa-cor:1" "\E") + next AOT_show \\\<^bold>\([O!]a & \[E!]a)\ + using "Act-Basic:2" "&E"(1) "\E"(1) not_act_ord_a "raa-cor:3" by blast + next AOT_have \\\([O!]a & \[E!]a)\ + by (metis "KBasic2:3" "&E"(1) "\E"(4) not_act_ord_a "oa-facts:3" + "oa-facts:7" "raa-cor:3" "vdash-properties:10") + AOT_thus \\\<^bold>\([O!]a & \[E!]a)\ + by (rule impossible_delta) + qed(auto simp: "beta-C-meta"[THEN "\E", OF 1]) + ultimately AOT_obtain F\<^sub>8 where \\<^bold>\[F\<^sub>8]b & \\<^bold>\[F\<^sub>8]b & \\<^bold>\[F\<^sub>8]a & \\<^bold>\[F\<^sub>8]a\ + using "\I"(1)[rotated, THEN "\E"[rotated]] by fastforce + AOT_hence \\<^bold>\[F\<^sub>8]b\ and \\\<^bold>\[F\<^sub>8]b\ and \\\<^bold>\[F\<^sub>8]a\ and \\\<^bold>\[F\<^sub>8]a\ + using "&E" by blast+ + note props = props this + + let ?\ = "\[\y \[E!]y & ([O!]y \ q\<^sub>0)]\" + AOT_modally_strict { + AOT_have \[\?\\]\\ by "cqt:2[lambda]" + } note 1 = this + moreover AOT_have \\<^bold>\[\?\\]b & \\<^bold>\[\?\\]b & \\<^bold>\[\?\\]a & \<^bold>\[\?\\]a\ + proof(safe intro!: "&I"; + AOT_subst \[\y \E!y & (O!y \ q\<^sub>0)]x\ \\E!x & (O!x \ q\<^sub>0)\ for: x) + AOT_show \\<^bold>\(\[E!]b & ([O!]b \ q\<^sub>0))\ + by (metis "Act-Basic:1" "Act-Basic:2" "Act-Basic:9" act_ord_b "&I" + "\I"(1) "\E"(2) "\E"(3) not_act_concrete_b "raa-cor:1") + next AOT_show \\\<^bold>\(\[E!]b & ([O!]b \ q\<^sub>0))\ + proof (rule act_and_pos_not_not_delta) + AOT_show \\<^bold>\(\[E!]b & ([O!]b \ q\<^sub>0))\ + by (metis "Act-Basic:1" "Act-Basic:2" "Act-Basic:9" act_ord_b "&I" + "\I"(1) "\E"(2) "\E"(3) not_act_concrete_b "raa-cor:1") + next + AOT_show \\\(\[E!]b & ([O!]b \ q\<^sub>0))\ + proof (AOT_subst \\(\[E!]b & ([O!]b \ q\<^sub>0))\ \[E!]b \ \([O!]b \ q\<^sub>0)\) + AOT_modally_strict { + AOT_show \\(\[E!]b & ([O!]b \ q\<^sub>0)) \ [E!]b \ \([O!]b \ q\<^sub>0)\ + by (metis "&I" "&E"(1,2) "\I"(1,2) "\E"(2) + "\I" "\I" "reductio-aa:1") + } + next + AOT_show \\([E!]b \ \([O!]b \ q\<^sub>0))\ + using "KBasic2:2" b_prop "&E"(1) "\I"(1) "\E"(3) + "raa-cor:3" by blast + qed + qed + next + AOT_show \\\<^bold>\(\[E!]a & ([O!]a \ q\<^sub>0))\ + using "Act-Basic:2" "Act-Basic:9" "&E"(2) "\E"(3) "\E"(1) + not_act_ord_a not_act_q_zero "reductio-aa:2" by blast + next + AOT_show \\<^bold>\(\[E!]a & ([O!]a \ q\<^sub>0))\ + proof (rule not_act_and_pos_delta) + AOT_show \\\<^bold>\(\[E!]a & ([O!]a \ q\<^sub>0))\ + by (metis "Act-Basic:2" "Act-Basic:9" "&E"(2) "\E"(3) "\E"(1) + not_act_ord_a not_act_q_zero "reductio-aa:2") + next + AOT_have \\\[E!]a\ + using "KBasic2:1" "\E"(2) not_act_and_pos_delta not_act_concrete_a + not_delta_concrete_a "raa-cor:5" by blast + moreover AOT_have \\([O!]a \ q\<^sub>0)\ + by (metis "KBasic2:2" "&E"(1) "\I"(2) "\E"(3) q\<^sub>0_prop "raa-cor:3") + ultimately AOT_show \\(\[E!]a & ([O!]a \ q\<^sub>0))\ + by (metis "KBasic:16" "&I" "vdash-properties:10") + qed + qed(auto simp: "beta-C-meta"[THEN "\E", OF 1]) + ultimately AOT_obtain F\<^sub>9 where \\<^bold>\[F\<^sub>9]b & \\<^bold>\[F\<^sub>9]b & \\<^bold>\[F\<^sub>9]a & \<^bold>\[F\<^sub>9]a\ + using "\I"(1)[rotated, THEN "\E"[rotated]] by fastforce + AOT_hence \\<^bold>\[F\<^sub>9]b\ and \\\<^bold>\[F\<^sub>9]b\ and \\\<^bold>\[F\<^sub>9]a\ and \\<^bold>\[F\<^sub>9]a\ + using "&E" by blast+ + note props = props this + + AOT_modally_strict { + AOT_have \[\y \q\<^sub>0]\\ by "cqt:2[lambda]" + } note 1 = this + moreover AOT_have \\<^bold>\[\y \q\<^sub>0]b & \\<^bold>\[\y \q\<^sub>0]b & \<^bold>\[\y \q\<^sub>0]a & \\<^bold>\[\y \q\<^sub>0]a\ + by (safe intro!: "&I"; AOT_subst \[\y \q\<^sub>0]x\ \\q\<^sub>0\ for: x) + (auto simp: act_not_q_zero not_delta_not_q_zero + "beta-C-meta"[THEN "\E", OF 1]) + ultimately AOT_obtain F\<^sub>1\<^sub>0 where \\<^bold>\[F\<^sub>1\<^sub>0]b & \\<^bold>\[F\<^sub>1\<^sub>0]b & \<^bold>\[F\<^sub>1\<^sub>0]a & \\<^bold>\[F\<^sub>1\<^sub>0]a\ + using "\I"(1)[rotated, THEN "\E"[rotated]] by fastforce + AOT_hence \\<^bold>\[F\<^sub>1\<^sub>0]b\ and \\\<^bold>\[F\<^sub>1\<^sub>0]b\ and \\<^bold>\[F\<^sub>1\<^sub>0]a\ and \\\<^bold>\[F\<^sub>1\<^sub>0]a\ + using "&E" by blast+ + note props = props this + + AOT_modally_strict { + AOT_have \[\y \[E!]y]\\ by "cqt:2[lambda]" + } note 1 = this + moreover AOT_have \\<^bold>\[\y \[E!]y]b & \\<^bold>\[\y \[E!]y]b & + \<^bold>\[\y \[E!]y]a & \<^bold>\[\y \[E!]y]a\ + proof (safe intro!: "&I"; AOT_subst \[\y \[E!]y]x\ \\[E!]x\ for: x) + AOT_show \\<^bold>\\[E!]b\ + using "Act-Basic:1" "\E"(2) not_act_concrete_b by blast + next AOT_show \\\<^bold>\\[E!]b\ + using "\\<^sub>d\<^sub>fE" "conventions:5" "Act-Basic:1" act_and_not_nec_not_delta + b_prop "&E"(1) "\E"(2) not_act_concrete_b by blast + next AOT_show \\<^bold>\\[E!]a\ + using "Act-Basic:1" "\E"(2) not_act_concrete_a by blast + next AOT_show \\<^bold>\\[E!]a\ + using "KBasic2:1" "\E"(2) nec_delta not_act_and_pos_delta + not_act_concrete_a not_delta_concrete_a "reductio-aa:1" + by blast + qed(auto simp: "beta-C-meta"[THEN "\E", OF 1]) + ultimately AOT_obtain F\<^sub>1\<^sub>1 where \\<^bold>\[F\<^sub>1\<^sub>1]b & \\<^bold>\[F\<^sub>1\<^sub>1]b & \<^bold>\[F\<^sub>1\<^sub>1]a & \<^bold>\[F\<^sub>1\<^sub>1]a\ + using "\I"(1)[rotated, THEN "\E"[rotated]] by fastforce + AOT_hence \\<^bold>\[F\<^sub>1\<^sub>1]b\ and \\\<^bold>\[F\<^sub>1\<^sub>1]b\ and \\<^bold>\[F\<^sub>1\<^sub>1]a\ and \\<^bold>\[F\<^sub>1\<^sub>1]a\ + using "&E" by blast+ + note props = props this + + AOT_have \\<^bold>\[O!]b & \<^bold>\[O!]b & \\<^bold>\[O!]a & \\<^bold>\[O!]a\ + by (simp add: act_ord_b "&I" delta_ord_b not_act_ord_a not_delta_ord_a) + then AOT_obtain F\<^sub>1\<^sub>2 where \\<^bold>\[F\<^sub>1\<^sub>2]b & \<^bold>\[F\<^sub>1\<^sub>2]b & \\<^bold>\[F\<^sub>1\<^sub>2]a & \\<^bold>\[F\<^sub>1\<^sub>2]a\ + using "oa-exist:1" "\I"(1)[rotated, THEN "\E"[rotated]] by fastforce + AOT_hence \\<^bold>\[F\<^sub>1\<^sub>2]b\ and \\<^bold>\[F\<^sub>1\<^sub>2]b\ and \\\<^bold>\[F\<^sub>1\<^sub>2]a\ and \\\<^bold>\[F\<^sub>1\<^sub>2]a\ + using "&E" by blast+ + note props = props this + + let ?\ = "\[\y [O!]y \ q\<^sub>0]\" + AOT_modally_strict { + AOT_have \[\?\\]\\ by "cqt:2[lambda]" + } note 1 = this + moreover AOT_have \\<^bold>\[\?\\]b & \<^bold>\[\?\\]b & \\<^bold>\[\?\\]a & \<^bold>\[\?\\]a\ + proof (safe intro!: "&I"; AOT_subst \[\y O!y \ q\<^sub>0]x\ \O!x \ q\<^sub>0\ for: x) + AOT_show \\<^bold>\([O!]b \ q\<^sub>0)\ + by (meson "Act-Basic:9" act_ord_b "\I"(1) "\E"(2)) + next AOT_show \\<^bold>\([O!]b \ q\<^sub>0)\ + by (meson "KBasic:15" b_ord "\I"(1) nec_delta "oa-facts:1" "\E") + next AOT_show \\\<^bold>\([O!]a \ q\<^sub>0)\ + using "Act-Basic:9" "\E"(2) "\E"(4) not_act_ord_a + not_act_q_zero "raa-cor:3" by blast + next AOT_show \\<^bold>\([O!]a \ q\<^sub>0)\ + proof (rule not_act_and_pos_delta) + AOT_show \\\<^bold>\([O!]a \ q\<^sub>0)\ + using "Act-Basic:9" "\E"(2) "\E"(4) not_act_ord_a + not_act_q_zero "raa-cor:3" by blast + next AOT_show \\([O!]a \ q\<^sub>0)\ + using "KBasic2:2" "&E"(1) "\I"(2) "\E"(2) q\<^sub>0_prop by blast + qed + qed(auto simp: "beta-C-meta"[THEN "\E", OF 1]) + ultimately AOT_obtain F\<^sub>1\<^sub>3 where \\<^bold>\[F\<^sub>1\<^sub>3]b & \<^bold>\[F\<^sub>1\<^sub>3]b & \\<^bold>\[F\<^sub>1\<^sub>3]a & \<^bold>\[F\<^sub>1\<^sub>3]a\ + using "\I"(1)[rotated, THEN "\E"[rotated]] by fastforce + AOT_hence \\<^bold>\[F\<^sub>1\<^sub>3]b\ and \\<^bold>\[F\<^sub>1\<^sub>3]b\ and \\\<^bold>\[F\<^sub>1\<^sub>3]a\ and \\<^bold>\[F\<^sub>1\<^sub>3]a\ + using "&E" by blast+ + note props = props this + + let ?\ = "\[\y [O!]y \ \q\<^sub>0]\" + AOT_modally_strict { + AOT_have \[\?\\]\\ by "cqt:2[lambda]" + } note 1 = this + moreover AOT_have \\<^bold>\[\?\\]b & \<^bold>\[\?\\]b & \<^bold>\[\?\\]a & \\<^bold>\[\?\\]a\ + proof (safe intro!: "&I"; AOT_subst \[\y O!y \ \q\<^sub>0]x\ \O!x \ \q\<^sub>0\ for: x) + AOT_show \\<^bold>\([O!]b \ \q\<^sub>0)\ + by (meson "Act-Basic:9" act_not_q_zero "\I"(2) "\E"(2)) + next AOT_show \\<^bold>\([O!]b \ \q\<^sub>0)\ + by (meson "KBasic:15" b_ord "\I"(1) nec_delta "oa-facts:1" "\E") + next AOT_show \\<^bold>\([O!]a \ \q\<^sub>0)\ + by (meson "Act-Basic:9" act_not_q_zero "\I"(2) "\E"(2)) + next AOT_show \\\<^bold>\([O!]a \ \q\<^sub>0)\ + proof(rule act_and_pos_not_not_delta) + AOT_show \\<^bold>\([O!]a \ \q\<^sub>0)\ + by (meson "Act-Basic:9" act_not_q_zero "\I"(2) "\E"(2)) + next + AOT_have \\\[O!]a\ + using "KBasic2:1" "\E"(2) not_act_and_pos_delta + not_act_ord_a not_delta_ord_a "raa-cor:6" by blast + moreover AOT_have \\q\<^sub>0\ + by (meson "&E"(1) q\<^sub>0_prop) + ultimately AOT_have 2: \\(\[O!]a & q\<^sub>0)\ + by (metis "KBasic:16" "&I" "vdash-properties:10") + AOT_show \\\([O!]a \ \q\<^sub>0)\ + proof (AOT_subst (reverse) \\([O!]a \ \q\<^sub>0)\ \\[O!]a & q\<^sub>0\) + AOT_modally_strict { + AOT_show \\[O!]a & q\<^sub>0 \ \([O!]a \ \q\<^sub>0)\ + by (metis "&I" "&E"(1) "&E"(2) "\I"(1) "\I"(2) + "\E"(3) "deduction-theorem" "\I" "raa-cor:3") + } + next + AOT_show \\(\[O!]a & q\<^sub>0)\ + using "2" by blast + qed + qed + qed(auto simp: "beta-C-meta"[THEN "\E", OF 1]) + ultimately AOT_obtain F\<^sub>1\<^sub>4 where \\<^bold>\[F\<^sub>1\<^sub>4]b & \<^bold>\[F\<^sub>1\<^sub>4]b & \<^bold>\[F\<^sub>1\<^sub>4]a & \\<^bold>\[F\<^sub>1\<^sub>4]a\ + using "\I"(1)[rotated, THEN "\E"[rotated]] by fastforce + AOT_hence \\<^bold>\[F\<^sub>1\<^sub>4]b\ and \\<^bold>\[F\<^sub>1\<^sub>4]b\ and \\<^bold>\[F\<^sub>1\<^sub>4]a\ and \\\<^bold>\[F\<^sub>1\<^sub>4]a\ + using "&E" by blast+ + note props = props this + + AOT_have \[L]\\ + by (rule "=\<^sub>d\<^sub>fI"(2)[OF L_def]) "cqt:2[lambda]"+ + moreover AOT_have \\<^bold>\[L]b & \<^bold>\[L]b & \<^bold>\[L]a & \<^bold>\[L]a\ + proof (safe intro!: "&I") + AOT_show \\<^bold>\[L]b\ + by (meson nec_L "nec-imp-act" "vdash-properties:10") + next AOT_show \\<^bold>\[L]b\ using nec_L nec_delta by blast + next AOT_show \\<^bold>\[L]a\ by (meson nec_L "nec-imp-act" "\E") + next AOT_show \\<^bold>\[L]a\ using nec_L nec_delta by blast + qed + ultimately AOT_obtain F\<^sub>1\<^sub>5 where \\<^bold>\[F\<^sub>1\<^sub>5]b & \<^bold>\[F\<^sub>1\<^sub>5]b & \<^bold>\[F\<^sub>1\<^sub>5]a & \<^bold>\[F\<^sub>1\<^sub>5]a\ + using "\I"(1)[rotated, THEN "\E"[rotated]] by fastforce + AOT_hence \\<^bold>\[F\<^sub>1\<^sub>5]b\ and \\<^bold>\[F\<^sub>1\<^sub>5]b\ and \\<^bold>\[F\<^sub>1\<^sub>5]a\ and \\<^bold>\[F\<^sub>1\<^sub>5]a\ + using "&E" by blast+ + note props = props this + + show ?thesis + by (rule "\I"(2)[where \=F\<^sub>0]; rule "\I"(2)[where \=F\<^sub>1]; + rule "\I"(2)[where \=F\<^sub>2]; rule "\I"(2)[where \=F\<^sub>3]; + rule "\I"(2)[where \=F\<^sub>4]; rule "\I"(2)[where \=F\<^sub>5]; + rule "\I"(2)[where \=F\<^sub>6]; rule "\I"(2)[where \=F\<^sub>7]; + rule "\I"(2)[where \=F\<^sub>8]; rule "\I"(2)[where \=F\<^sub>9]; + rule "\I"(2)[where \=F\<^sub>1\<^sub>0]; rule "\I"(2)[where \=F\<^sub>1\<^sub>1]; + rule "\I"(2)[where \=F\<^sub>1\<^sub>2]; rule "\I"(2)[where \=F\<^sub>1\<^sub>3]; + rule "\I"(2)[where \=F\<^sub>1\<^sub>4]; rule "\I"(2)[where \=F\<^sub>1\<^sub>5]; + safe intro!: "&I") + (match conclusion in "[?v \ [F] \ [G]]" for F G \ \ + match props in A: "[?v \ \\{F}]" for \ \ \ + match (\) in "\a . ?p" \ \fail\ \ "\a . a" \ \fail\ \ _ \ \ + match props in B: "[?v \ \{G}]" \ \ + fact "pos-not-equiv-ne:4"[where F=F and G=G and \=\, THEN "\E", + OF "oth-class-taut:4:h"[THEN "\E"(2)], + OF "Disjunction Addition"(2)[THEN "\E"], + OF "&I", OF A, OF B]\\\\)+ +qed + +subsection\The Theory of Objects\ +text\\label{PLM: 9.11}\ + +AOT_theorem "o-objects-exist:1": \\\x O!x\ +proof(rule RN) + AOT_modally_strict { + AOT_obtain a where \\(E!a & \\<^bold>\[E!]a)\ + using "\E"[rotated, OF "qml:4"[axiom_inst, THEN "BF\"[THEN "\E"]]] + by blast + AOT_hence 1: \\E!a\ by (metis "KBasic2:3" "&E"(1) "\E") + AOT_have \[\x \[E!]x]a\ + proof (rule "\\C"(1); "cqt:2[lambda]"?) + AOT_show \a\\ using "cqt:2[const_var]"[axiom_inst] by blast + next + AOT_show \\E!a\ by (fact 1) + qed + AOT_hence \O!a\ by (rule "=\<^sub>d\<^sub>fI"(2)[OF AOT_ordinary, rotated]) "cqt:2" + AOT_thus \\x [O!]x\ by (rule "\I") + } +qed + +AOT_theorem "o-objects-exist:2": \\\x A!x\ +proof (rule RN) + AOT_modally_strict { + AOT_obtain a where \[A!]a\ + using "A-objects"[axiom_inst] "\E"[rotated] "&E" by blast + AOT_thus \\x A!x\ using "\I" by blast + } +qed + +AOT_theorem "o-objects-exist:3": \\\\x O!x\ + by (rule RN) + (metis (no_types, opaque_lifting) "\E" "cqt-orig:1[const_var]" + "\E"(4) "modus-tollens:1" "o-objects-exist:2" "oa-contingent:2" + "qml:2"[axiom_inst] "reductio-aa:2") + +AOT_theorem "o-objects-exist:4": \\\\x A!x\ + by (rule RN) + (metis (mono_tags, opaque_lifting) "\E" "cqt-orig:1[const_var]" + "\E"(1) "modus-tollens:1" "o-objects-exist:1" "oa-contingent:2" + "qml:2"[axiom_inst] "\E") + +AOT_theorem "o-objects-exist:5": \\\\x E!x\ +proof (rule RN; rule "raa-cor:2") + AOT_modally_strict { + AOT_assume \\x E!x\ + moreover AOT_obtain a where abs: \A!a\ + using "o-objects-exist:2"[THEN "qml:2"[axiom_inst, THEN "\E"]] + "\E"[rotated] by blast + ultimately AOT_have \E!a\ using "\E" by blast + AOT_hence 1: \\E!a\ by (metis "T\" "\E") + AOT_have \[\y \E!y]a\ + proof (rule "\\C"(1); "cqt:2[lambda]"?) + AOT_show \a\\ using "cqt:2[const_var]"[axiom_inst]. + next + AOT_show \\E!a\ by (fact 1) + qed + AOT_hence \O!a\ + by (rule "=\<^sub>d\<^sub>fI"(2)[OF AOT_ordinary, rotated]) "cqt:2[lambda]" + AOT_hence \\A!a\ by (metis "\E"(1) "oa-contingent:2") + AOT_thus \p & \p\ for p using abs by (metis "raa-cor:3") + } +qed + +AOT_theorem partition: \\\x (O!x & A!x)\ +proof(rule "raa-cor:2") + AOT_assume \\x (O!x & A!x)\ + then AOT_obtain a where \O!a & A!a\ + using "\E"[rotated] by blast + AOT_thus \p & \p\ for p + by (metis "&E"(1) "Conjunction Simplification"(2) "\E"(1) + "modus-tollens:1" "oa-contingent:2" "raa-cor:3") +qed + +AOT_define eq_E :: \\\ ("'(=\<^sub>E')") + "=E": \(=\<^sub>E) =\<^sub>d\<^sub>f [\xy O!x & O!y & \\F ([F]x \ [F]y)]\ + +syntax "_AOT_eq_E_infix" :: \\ \ \ \ \\ (infixl "=\<^sub>E" 50) +translations + "_AOT_eq_E_infix \ \'" == "CONST AOT_exe (CONST eq_E) (CONST Pair \ \')" +print_translation\ +AOT_syntax_print_translations +[(\<^const_syntax>\AOT_exe\, fn ctxt => fn [ + Const (\<^const_name>\eq_E\, _), + Const (\<^const_syntax>\Pair\, _) $ lhs $ rhs +] => Const (\<^syntax_const>\_AOT_eq_E_infix\, dummyT) $ lhs $ rhs)]\ + +text\Note: Not explicitly mentioned as theorem in PLM.\ +AOT_theorem "=E[denotes]": \[(=\<^sub>E)]\\ + by (rule "=\<^sub>d\<^sub>fI"(2)[OF "=E"]) "cqt:2[lambda]"+ + +AOT_theorem "=E-simple:1": \x =\<^sub>E y \ (O!x & O!y & \\F ([F]x \ [F]y))\ +proof - + AOT_have 1: \[\xy [O!]x & [O!]y & \\F ([F]x \ [F]y)]\\ by "cqt:2" + show ?thesis + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF "=E"]; "cqt:2[lambda]"?) + using "beta-C-meta"[THEN "\E", OF 1, unvarify \\<^sub>1\\<^sub>n, of "(_,_)", + OF tuple_denotes[THEN "\\<^sub>d\<^sub>fI"], OF "&I", + OF "cqt:2[const_var]"[axiom_inst], + OF "cqt:2[const_var]"[axiom_inst]] + by fast +qed + +AOT_theorem "=E-simple:2": \x =\<^sub>E y \ x = y\ +proof (rule "\I") + AOT_assume \x =\<^sub>E y\ + AOT_hence \O!x & O!y & \\F ([F]x \ [F]y)\ + using "=E-simple:1"[THEN "\E"(1)] by blast + AOT_thus \x = y\ + using "\\<^sub>d\<^sub>fI"[OF "identity:1"] "\I" by blast +qed + +AOT_theorem "id-nec3:1": \x =\<^sub>E y \ \(x =\<^sub>E y)\ +proof (rule "\I"; rule "\I") + AOT_assume \x =\<^sub>E y\ + AOT_hence \O!x & O!y & \\F ([F]x \ [F]y)\ + using "=E-simple:1" "\E" by blast + AOT_hence \\O!x & \O!y & \\\F ([F]x \ [F]y)\ + by (metis "S5Basic:6" "&I" "&E"(1) "&E"(2) "\E"(4) + "oa-facts:1" "raa-cor:3" "vdash-properties:10") + AOT_hence \\(O!x & O!y & \\F ([F]x \ [F]y))\ + by (metis "&E"(1) "&E"(2) "\E"(2) "KBasic:3" "&I") + AOT_thus \\(x =\<^sub>E y)\ + using "=E-simple:1" + by (AOT_subst \x =\<^sub>E y\ \O!x & O!y & \\F ([F]x \ [F]y)\) auto +next + AOT_assume \\(x =\<^sub>E y)\ + AOT_thus \x =\<^sub>E y\ using "qml:2"[axiom_inst, THEN "\E"] by blast +qed + +AOT_theorem "id-nec3:2": \\(x =\<^sub>E y) \ x =\<^sub>E y\ + by (meson "RE\" "S5Basic:2" "id-nec3:1" "\E"(1,5) "Commutativity of \") + +AOT_theorem "id-nec3:3": \\(x =\<^sub>E y) \ \(x =\<^sub>E y)\ + by (meson "id-nec3:1" "id-nec3:2" "\E"(5)) + +syntax "_AOT_non_eq_E" :: \\\ ("'(\\<^sub>E')") +translations + (\) "(\\<^sub>E)" == (\) "(=\<^sub>E)\<^sup>-" +syntax "_AOT_non_eq_E_infix" :: \\ \ \ \ \\ (infixl "\\<^sub>E" 50) +translations + "_AOT_non_eq_E_infix \ \'" == + "CONST AOT_exe (CONST relation_negation (CONST eq_E)) (CONST Pair \ \')" +print_translation\ +AOT_syntax_print_translations +[(\<^const_syntax>\AOT_exe\, fn ctxt => fn [ + Const (\<^const_syntax>\relation_negation\, _) $ Const (\<^const_name>\eq_E\, _), + Const (\<^const_syntax>\Pair\, _) $ lhs $ rhs +] => Const (\<^syntax_const>\_AOT_non_eq_E_infix\, dummyT) $ lhs $ rhs)]\ +AOT_theorem "thm-neg=E": \x \\<^sub>E y \ \(x =\<^sub>E y)\ +proof - + AOT_have \: \[\x\<^sub>1...x\<^sub>2 \(=\<^sub>E)x\<^sub>1...x\<^sub>2]\\ by "cqt:2" + AOT_have \x \\<^sub>E y \ [\x\<^sub>1...x\<^sub>2 \(=\<^sub>E)x\<^sub>1...x\<^sub>2]xy\ + by (rule "=\<^sub>d\<^sub>fI"(1)[OF "df-relation-negation", OF \]) + (meson "oth-class-taut:3:a") + also AOT_have \\ \ \(=\<^sub>E)xy\ + by (safe intro!: "beta-C-meta"[THEN "\E", unvarify \\<^sub>1\\<^sub>n] "cqt:2" + tuple_denotes[THEN "\\<^sub>d\<^sub>fI"] "&I") + finally show ?thesis. +qed + +AOT_theorem "id-nec4:1": \x \\<^sub>E y \ \(x \\<^sub>E y)\ +proof - + AOT_have \x \\<^sub>E y \ \(x =\<^sub>E y)\ using "thm-neg=E". + also AOT_have \\ \ \\(x =\<^sub>E y)\ + by (meson "id-nec3:2" "\E"(1) "Commutativity of \" "oth-class-taut:4:b") + also AOT_have \\ \ \\(x =\<^sub>E y)\ + by (meson "KBasic2:1" "\E"(2) "Commutativity of \") + also AOT_have \\ \ \(x \\<^sub>E y)\ + by (AOT_subst (reverse) \\(x =\<^sub>E y)\ \x \\<^sub>E y\) + (auto simp: "thm-neg=E" "oth-class-taut:3:a") + finally show ?thesis. +qed + +AOT_theorem "id-nec4:2": \\(x \\<^sub>E y) \ (x \\<^sub>E y)\ + by (meson "RE\" "S5Basic:2" "id-nec4:1" "\E"(2,5) "Commutativity of \") + +AOT_theorem "id-nec4:3": \\(x \\<^sub>E y) \ \(x \\<^sub>E y)\ + by (meson "id-nec4:1" "id-nec4:2" "\E"(5)) + +AOT_theorem "id-act2:1": \x =\<^sub>E y \ \<^bold>\x =\<^sub>E y\ + by (meson "Act-Basic:5" "Act-Sub:2" "RA[2]" "id-nec3:2" "\E"(1,6)) +AOT_theorem "id-act2:2": \x \\<^sub>E y \ \<^bold>\x \\<^sub>E y\ + by (meson "Act-Basic:5" "Act-Sub:2" "RA[2]" "id-nec4:2" "\E"(1,6)) + +AOT_theorem "ord=Eequiv:1": \O!x \ x =\<^sub>E x\ +proof (rule "\I") + AOT_assume 1: \O!x\ + AOT_show \x =\<^sub>E x\ + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF "=E"]) apply "cqt:2[lambda]" + apply (rule "\\C"(1)) + apply "cqt:2[lambda]" + apply (simp add: "&I" "cqt:2[const_var]"[axiom_inst] prod_denotesI) + by (simp add: "1" RN "&I" "oth-class-taut:3:a" "universal-cor") +qed + +AOT_theorem "ord=Eequiv:2": \x =\<^sub>E y \ y =\<^sub>E x\ +proof(rule CP) + AOT_assume 1: \x =\<^sub>E y\ + AOT_hence 2: \x = y\ by (metis "=E-simple:2" "vdash-properties:10") + AOT_have \O!x\ using 1 by (meson "&E"(1) "=E-simple:1" "\E"(1)) + AOT_hence \x =\<^sub>E x\ using "ord=Eequiv:1" "\E" by blast + AOT_thus \y =\<^sub>E x\ using "rule=E"[rotated, OF 2] by fast +qed + +AOT_theorem "ord=Eequiv:3": \(x =\<^sub>E y & y =\<^sub>E z) \ x =\<^sub>E z\ +proof (rule CP) + AOT_assume 1: \x =\<^sub>E y & y =\<^sub>E z\ + AOT_hence \x = y & y = z\ + by (metis "&I" "&E"(1) "&E"(2) "=E-simple:2" "vdash-properties:6") + AOT_hence \x = z\ by (metis "id-eq:3" "vdash-properties:6") + moreover AOT_have \x =\<^sub>E x\ + using 1[THEN "&E"(1)] "&E"(1) "=E-simple:1" "\E"(1) + "ord=Eequiv:1" "\E" by blast + ultimately AOT_show \x =\<^sub>E z\ + using "rule=E" by fast +qed + +AOT_theorem "ord-=E=:1": \(O!x \ O!y) \ \(x = y \ x =\<^sub>E y)\ +proof(rule CP) + AOT_assume \O!x \ O!y\ + moreover { + AOT_assume \O!x\ + AOT_hence \\O!x\ by (metis "oa-facts:1" "vdash-properties:10") + moreover { + AOT_modally_strict { + AOT_have \O!x \ (x = y \ x =\<^sub>E y)\ + proof (rule "\I"; rule "\I"; rule "\I") + AOT_assume \O!x\ + AOT_hence \x =\<^sub>E x\ by (metis "ord=Eequiv:1" "\E") + moreover AOT_assume \x = y\ + ultimately AOT_show \x =\<^sub>E y\ using "rule=E" by fast + next + AOT_assume \x =\<^sub>E y\ + AOT_thus \x = y\ by (metis "=E-simple:2" "\E") + qed + } + AOT_hence \\O!x \ \(x = y \ x =\<^sub>E y)\ by (metis "RM:1") + } + ultimately AOT_have \\(x = y \ x =\<^sub>E y)\ using "\E" by blast + } + moreover { + AOT_assume \O!y\ + AOT_hence \\O!y\ by (metis "oa-facts:1" "vdash-properties:10") + moreover { + AOT_modally_strict { + AOT_have \O!y \ (x = y \ x =\<^sub>E y)\ + proof (rule "\I"; rule "\I"; rule "\I") + AOT_assume \O!y\ + AOT_hence \y =\<^sub>E y\ by (metis "ord=Eequiv:1" "\E") + moreover AOT_assume \x = y\ + ultimately AOT_show \x =\<^sub>E y\ using "rule=E" id_sym by fast + next + AOT_assume \x =\<^sub>E y\ + AOT_thus \x = y\ by (metis "=E-simple:2" "\E") + qed + } + AOT_hence \\O!y \ \(x = y \ x =\<^sub>E y)\ by (metis "RM:1") + } + ultimately AOT_have \\(x = y \ x =\<^sub>E y)\ using "\E" by blast + } + ultimately AOT_show \\(x = y \ x =\<^sub>E y)\ by (metis "\E"(3) "raa-cor:1") +qed + +AOT_theorem "ord-=E=:2": \O!y \ [\x x = y]\\ +proof (rule "\I"; rule "safe-ext"[axiom_inst, THEN "\E"]; rule "&I") + AOT_show \[\x x =\<^sub>E y]\\ by "cqt:2[lambda]" +next + AOT_assume \O!y\ + AOT_hence 1: \\(x = y \ x =\<^sub>E y)\ for x + using "ord-=E=:1" "\E" "\I" by blast + AOT_have \\(x =\<^sub>E y \ x = y)\ for x + by (AOT_subst \x =\<^sub>E y \ x = y\ \x = y \ x =\<^sub>E y\) + (auto simp add: "Commutativity of \" 1) + AOT_hence \\x \(x =\<^sub>E y \ x = y)\ by (rule GEN) + AOT_thus \\\x (x =\<^sub>E y \ x = y)\ by (rule BF[THEN "\E"]) +qed + + +AOT_theorem "ord-=E=:3": \[\xy O!x & O!y & x = y]\\ +proof (rule "safe-ext[2]"[axiom_inst, THEN "\E"]; rule "&I") + AOT_show \[\xy O!x & O!y & x =\<^sub>E y]\\ by "cqt:2[lambda]" +next + AOT_show \\\x\y ([O!]x & [O!]y & x =\<^sub>E y \ [O!]x & [O!]y & x = y)\ + proof (rule RN; rule GEN; rule GEN; rule "\I"; rule "\I") + AOT_modally_strict { + AOT_show \[O!]x & [O!]y & x = y\ if \[O!]x & [O!]y & x =\<^sub>E y\ for x y + by (metis "&I" "&E"(1) "Conjunction Simplification"(2) "=E-simple:2" + "modus-tollens:1" "raa-cor:1" that) + } + next + AOT_modally_strict { + AOT_show \[O!]x & [O!]y & x =\<^sub>E y\ if \[O!]x & [O!]y & x = y\ for x y + apply(safe intro!: "&I") + apply (metis that[THEN "&E"(1), THEN "&E"(1)]) + apply (metis that[THEN "&E"(1), THEN "&E"(2)]) + using "rule=E"[rotated, OF that[THEN "&E"(2)]] + "ord=Eequiv:1"[THEN "\E", OF that[THEN "&E"(1), THEN "&E"(1)]] + by fast + } + qed +qed + +AOT_theorem "ind-nec": \\F ([F]x \ [F]y) \ \\F ([F]x \ [F]y)\ +proof(rule "\I") + AOT_assume \\F ([F]x \ [F]y)\ + moreover AOT_have \[\x \\F ([F]x \ [F]y)]\\ by "cqt:2[lambda]" + ultimately AOT_have \[\x \\F ([F]x \ [F]y)]x \ [\x \\F ([F]x \ [F]y)]y\ + using "\E" by blast + moreover AOT_have \[\x \\F ([F]x \ [F]y)]y\ + apply (rule "\\C"(1)) + apply "cqt:2[lambda]" + apply (fact "cqt:2[const_var]"[axiom_inst]) + by (simp add: RN GEN "oth-class-taut:3:a") + ultimately AOT_have \[\x \\F ([F]x \ [F]y)]x\ using "\E" by blast + AOT_thus \\\F ([F]x \ [F]y)\ + using "\\C"(1) by blast +qed + +AOT_theorem "ord=E:1": \(O!x & O!y) \ (\F ([F]x \ [F]y) \ x =\<^sub>E y)\ +proof (rule "\I"; rule "\I") + AOT_assume \\F ([F]x \ [F]y)\ + AOT_hence \\\F ([F]x \ [F]y)\ + using "ind-nec"[THEN "\E"] by blast + moreover AOT_assume \O!x & O!y\ + ultimately AOT_have \O!x & O!y & \\F ([F]x \ [F]y)\ + using "&I" by blast + AOT_thus \x =\<^sub>E y\ using "=E-simple:1"[THEN "\E"(2)] by blast +qed + +AOT_theorem "ord=E:2": \(O!x & O!y) \ (\F ([F]x \ [F]y) \ x = y)\ +proof (rule "\I"; rule "\I") + AOT_assume \O!x & O!y\ + moreover AOT_assume \\F ([F]x \ [F]y)\ + ultimately AOT_have \x =\<^sub>E y\ + using "ord=E:1" "\E" by blast + AOT_thus \x = y\ using "=E-simple:2"[THEN "\E"] by blast +qed + +AOT_theorem "ord=E2:1": + \(O!x & O!y) \ (x \ y \ [\z z =\<^sub>E x] \ [\z z =\<^sub>E y])\ +proof (rule "\I"; rule "\I"; rule "\I"; + rule "\\<^sub>d\<^sub>fI"[OF "=-infix"]; rule "raa-cor:2") + AOT_assume 0: \O!x & O!y\ + AOT_assume \x \ y\ + AOT_hence 1: \\(x = y)\ using "\\<^sub>d\<^sub>fE"[OF "=-infix"] by blast + AOT_assume \[\z z =\<^sub>E x] = [\z z =\<^sub>E y]\ + moreover AOT_have \[\z z =\<^sub>E x]x\ + apply (rule "\\C"(1)) + apply "cqt:2[lambda]" + apply (fact "cqt:2[const_var]"[axiom_inst]) + using "ord=Eequiv:1"[THEN "\E", OF 0[THEN "&E"(1)]]. + ultimately AOT_have \[\z z =\<^sub>E y]x\ using "rule=E" by fast + AOT_hence \x =\<^sub>E y\ using "\\C"(1) by blast + AOT_hence \x = y\ by (metis "=E-simple:2" "vdash-properties:6") + AOT_thus \x = y & \(x = y)\ using 1 "&I" by blast +next + AOT_assume \[\z z =\<^sub>E x] \ [\z z =\<^sub>E y]\ + AOT_hence 0: \\([\z z =\<^sub>E x] = [\z z =\<^sub>E y])\ + using "\\<^sub>d\<^sub>fE"[OF "=-infix"] by blast + AOT_have \[\z z =\<^sub>E x]\\ by "cqt:2[lambda]" + AOT_hence \[\z z =\<^sub>E x] = [\z z =\<^sub>E x]\ + by (metis "rule=I:1") + moreover AOT_assume \x = y\ + ultimately AOT_have \[\z z =\<^sub>E x] = [\z z =\<^sub>E y]\ + using "rule=E" by fast + AOT_thus \[\z z =\<^sub>E x] = [\z z =\<^sub>E y] & \([\z z =\<^sub>E x] = [\z z =\<^sub>E y])\ + using 0 "&I" by blast +qed + +AOT_theorem "ord=E2:2": + \(O!x & O!y) \ (x \ y \ [\z z = x] \ [\z z = y])\ +proof (rule "\I"; rule "\I"; rule "\I"; + rule "\\<^sub>d\<^sub>fI"[OF "=-infix"]; rule "raa-cor:2") + AOT_assume 0: \O!x & O!y\ + AOT_assume \x \ y\ + AOT_hence 1: \\(x = y)\ using "\\<^sub>d\<^sub>fE"[OF "=-infix"] by blast + AOT_assume \[\z z = x] = [\z z = y]\ + moreover AOT_have \[\z z = x]x\ + apply (rule "\\C"(1)) + apply (fact "ord-=E=:2"[THEN "\E", OF 0[THEN "&E"(1)]]) + apply (fact "cqt:2[const_var]"[axiom_inst]) + by (simp add: "id-eq:1") + ultimately AOT_have \[\z z = y]x\ using "rule=E" by fast + AOT_hence \x = y\ using "\\C"(1) by blast + AOT_thus \x = y & \(x = y)\ using 1 "&I" by blast +next + AOT_assume 0: \O!x & O!y\ + AOT_assume \[\z z = x] \ [\z z = y]\ + AOT_hence 1: \\([\z z = x] = [\z z = y])\ + using "\\<^sub>d\<^sub>fE"[OF "=-infix"] by blast + AOT_have \[\z z = x]\\ + by (fact "ord-=E=:2"[THEN "\E", OF 0[THEN "&E"(1)]]) + AOT_hence \[\z z = x] = [\z z = x]\ + by (metis "rule=I:1") + moreover AOT_assume \x = y\ + ultimately AOT_have \[\z z = x] = [\z z = y]\ + using "rule=E" by fast + AOT_thus \[\z z = x] = [\z z = y] & \([\z z = x] = [\z z = y])\ + using 1 "&I" by blast +qed + +AOT_theorem ordnecfail: \O!x \ \\\F x[F]\ + by (meson "RM:1" "\I" nocoder[axiom_inst] "oa-facts:1" "\E") + +AOT_theorem "ab-obey:1": \(A!x & A!y) \ (\F (x[F] \ y[F]) \ x = y)\ +proof (rule "\I"; rule "\I") + AOT_assume 1: \A!x & A!y\ + AOT_assume \\F (x[F] \ y[F])\ + AOT_hence \x[F] \ y[F]\ for F using "\E" by blast + AOT_hence \\(x[F] \ y[F])\ for F by (metis "en-eq:6[1]" "\E"(1)) + AOT_hence \\F \(x[F] \ y[F])\ by (rule GEN) + AOT_hence \\\F (x[F] \ y[F])\ by (rule BF[THEN "\E"]) + AOT_thus \x = y\ + using "\\<^sub>d\<^sub>fI"[OF "identity:1", OF "\I"(2)] 1 "&I" by blast +qed + +AOT_theorem "ab-obey:2": + \(\F (x[F] & \y[F]) \ \F (y[F] & \x[F])) \ x \ y\ +proof (rule "\I"; rule "\\<^sub>d\<^sub>fI"[OF "=-infix"]; rule "raa-cor:2") + AOT_assume 1: \x = y\ + AOT_assume \\F (x[F] & \y[F]) \ \F (y[F] & \x[F])\ + moreover { + AOT_assume \\F (x[F] & \y[F])\ + then AOT_obtain F where \x[F] & \y[F]\ + using "\E"[rotated] by blast + moreover AOT_have \y[F]\ + using calculation[THEN "&E"(1)] 1 "rule=E" by fast + ultimately AOT_have \p & \p\ for p + by (metis "Conjunction Simplification"(2) "modus-tollens:2" "raa-cor:3") + } + moreover { + AOT_assume \\F (y[F] & \x[F])\ + then AOT_obtain F where \y[F] & \x[F]\ + using "\E"[rotated] by blast + moreover AOT_have \\y[F]\ + using calculation[THEN "&E"(2)] 1 "rule=E" by fast + ultimately AOT_have \p & \p\ for p + by (metis "Conjunction Simplification"(1) "modus-tollens:1" "raa-cor:3") + } + ultimately AOT_show \p & \p\ for p + by (metis "\E"(3) "raa-cor:1") +qed + +AOT_theorem "encoders-are-abstract": \\F x[F] \ A!x\ + by (meson "deduction-theorem" "\E"(2) "modus-tollens:2" nocoder + "oa-contingent:3" "vdash-properties:1[2]") + +AOT_theorem "denote=:1": \\H\x x[H]\ + by (rule GEN; rule "existence:2[1]"[THEN "\\<^sub>d\<^sub>fE"]; "cqt:2") + +AOT_theorem "denote=:2": \\G\x\<^sub>1...\x\<^sub>n x\<^sub>1...x\<^sub>n[H]\ + by (rule GEN; rule "existence:2"[THEN "\\<^sub>d\<^sub>fE"]; "cqt:2") + +AOT_theorem "denote=:2[2]": \\G\x\<^sub>1\x\<^sub>2 x\<^sub>1x\<^sub>2[H]\ + by (rule GEN; rule "existence:2[2]"[THEN "\\<^sub>d\<^sub>fE"]; "cqt:2") + +AOT_theorem "denote=:2[3]": \\G\x\<^sub>1\x\<^sub>2\x\<^sub>3 x\<^sub>1x\<^sub>2x\<^sub>3[H]\ + by (rule GEN; rule "existence:2[3]"[THEN "\\<^sub>d\<^sub>fE"]; "cqt:2") + +AOT_theorem "denote=:2[4]": \\G\x\<^sub>1\x\<^sub>2\x\<^sub>3\x\<^sub>4 x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[H]\ + by (rule GEN; rule "existence:2[4]"[THEN "\\<^sub>d\<^sub>fE"]; "cqt:2") + +AOT_theorem "denote=:3": \\x x[\] \ \H (H = \)\ + using "existence:2[1]" "free-thms:1" "\E"(2,5) + "Commutativity of \" "\Df" by blast + +AOT_theorem "denote=:4": \(\x\<^sub>1...\x\<^sub>n x\<^sub>1...x\<^sub>n[\]) \ \H (H = \)\ + using "existence:2" "free-thms:1" "\E"(6) "\Df" by blast + +AOT_theorem "denote=:4[2]": \(\x\<^sub>1\x\<^sub>2 x\<^sub>1x\<^sub>2[\]) \ \H (H = \)\ + using "existence:2[2]" "free-thms:1" "\E"(6) "\Df" by blast + +AOT_theorem "denote=:4[3]": \(\x\<^sub>1\x\<^sub>2\x\<^sub>3 x\<^sub>1x\<^sub>2x\<^sub>3[\]) \ \H (H = \)\ + using "existence:2[3]" "free-thms:1" "\E"(6) "\Df" by blast + +AOT_theorem "denote=:4[4]": \(\x\<^sub>1\x\<^sub>2\x\<^sub>3\x\<^sub>4 x\<^sub>1x\<^sub>2x\<^sub>3x\<^sub>4[\]) \ \H (H = \)\ + using "existence:2[4]" "free-thms:1" "\E"(6) "\Df" by blast + +AOT_theorem "A-objects!": \\!x (A!x & \F (x[F] \ \{F}))\ +proof (rule "uniqueness:1"[THEN "\\<^sub>d\<^sub>fI"]) + AOT_obtain a where a_prop: \A!a & \F (a[F] \ \{F})\ + using "A-objects"[axiom_inst] "\E"[rotated] by blast + AOT_have \(A!\ & \F (\[F] \ \{F})) \ \ = a\ for \ + proof (rule "\I") + AOT_assume \_prop: \[A!]\ & \F (\[F] \ \{F})\ + AOT_hence \\[F] \ \{F}\ for F + using "\E" "&E" by blast + AOT_hence \\[F] \ a[F]\ for F + using a_prop[THEN "&E"(2)] "\E" "\E"(2,5) + "Commutativity of \" by fast + AOT_hence \\F (\[F] \ a[F])\ by (rule GEN) + AOT_thus \\ = a\ + using "ab-obey:1"[THEN "\E", + OF "&I"[OF \_prop[THEN "&E"(1)], OF a_prop[THEN "&E"(1)]], + THEN "\E"] by blast + qed + AOT_hence \\\ ((A!\ & \F (\[F] \ \{F})) \ \ = a)\ by (rule GEN) + AOT_thus \\\ ([A!]\ & \F (\[F] \ \{F}) & + \\ ([A!]\ & \F (\[F] \ \{F}) \ \ = \))\ + using "\I" using a_prop "&I" by fast +qed + +AOT_theorem "obj-oth:1": \\!x (A!x & \F (x[F] \ [F]y))\ + using "A-objects!" by fast + +AOT_theorem "obj-oth:2": \\!x (A!x & \F (x[F] \ [F]y & [F]z))\ + using "A-objects!" by fast + +AOT_theorem "obj-oth:3": \\!x (A!x & \F (x[F] \ [F]y \ [F]z))\ + using "A-objects!" by fast + +AOT_theorem "obj-oth:4": \\!x (A!x & \F (x[F] \ \[F]y))\ + using "A-objects!" by fast + +AOT_theorem "obj-oth:5": \\!x (A!x & \F (x[F] \ F = G))\ + using "A-objects!" by fast + +AOT_theorem "obj-oth:6": \\!x (A!x & \F (x[F] \ \\y([G]y \ [F]y)))\ + using "A-objects!" by fast + +AOT_theorem "A-descriptions": \\<^bold>\x (A!x & \F (x[F] \ \{F}))\\ + by (rule "A-Exists:2"[THEN "\E"(2)]; rule "RA[2]"; rule "A-objects!") + +AOT_act_theorem "thm-can-terms2": + \y = \<^bold>\x(A!x & \F (x[F] \ \{F})) \ (A!y & \F (y[F] \ \{F}))\ + using "y-in:2" by blast + +AOT_theorem "can-ab2": \y = \<^bold>\x(A!x & \F (x[F] \ \{F})) \ A!y\ +proof(rule "\I") + AOT_assume \y = \<^bold>\x(A!x & \F (x[F] \ \{F}))\ + AOT_hence \\<^bold>\(A!y & \F (y[F] \ \{F}))\ + using "actual-desc:2"[THEN "\E"] by blast + AOT_hence \\<^bold>\A!y\ by (metis "Act-Basic:2" "&E"(1) "\E"(1)) + AOT_thus \A!y\ by (metis "\E"(2) "oa-facts:8") +qed + +AOT_act_theorem "desc-encode:1": \\<^bold>\x(A!x & \F (x[F] \ \{F}))[F] \ \{F}\ +proof - + AOT_have \\<^bold>\x(A!x & \F (x[F] \ \{F}))\\ + by (simp add: "A-descriptions") + AOT_hence \A!\<^bold>\x(A!x & \F (x[F] \ \{F})) & + \F(\<^bold>\x(A!x & \F (x[F] \ \{F}))[F] \ \{F})\ + using "y-in:3"[THEN "\E"] by blast + AOT_thus \\<^bold>\x(A!x & \F (x[F] \ \{F}))[F] \ \{F}\ + using "&E" "\E" by blast +qed + +AOT_act_theorem "desc-encode:2": \\<^bold>\x(A!x & \F (x[F] \ \{F}))[G] \ \{G}\ + using "desc-encode:1". + +AOT_theorem "desc-nec-encode:1": + \\<^bold>\x (A!x & \F (x[F] \ \{F}))[F] \ \<^bold>\\{F}\ +proof - + AOT_have 0: \\<^bold>\x(A!x & \F (x[F] \ \{F}))\\ + by (simp add: "A-descriptions") + AOT_hence \\<^bold>\(A!\<^bold>\x(A!x & \F (x[F] \ \{F})) & + \F(\<^bold>\x(A!x & \F (x[F] \ \{F}))[F] \ \{F}))\ + using "actual-desc:4"[THEN "\E"] by blast + AOT_hence \\<^bold>\\F (\<^bold>\x(A!x & \F (x[F] \ \{F}))[F] \ \{F})\ + using "Act-Basic:2" "&E"(2) "\E"(1) by blast + AOT_hence \\F \<^bold>\(\<^bold>\x(A!x & \F (x[F] \ \{F}))[F] \ \{F})\ + using "\E"(1) "logic-actual-nec:3" "vdash-properties:1[2]" by blast + AOT_hence \\<^bold>\(\<^bold>\x(A!x & \F (x[F] \ \{F}))[F] \ \{F})\ + using "\E" by blast + AOT_hence \\<^bold>\\<^bold>\x(A!x & \F (x[F] \ \{F}))[F] \ \<^bold>\\{F}\ + using "Act-Basic:5" "\E"(1) by blast + AOT_thus \\<^bold>\x(A!x & \F (x[F] \ \{F}))[F] \ \<^bold>\\{F}\ + using "en-eq:10[1]"[unvarify x\<^sub>1, OF 0] "\E"(6) by blast +qed + +AOT_theorem "desc-nec-encode:2": + \\<^bold>\x (A!x & \F (x[F] \ \{F}))[G] \ \<^bold>\\{G}\ + using "desc-nec-encode:1". + +AOT_theorem "Box-desc-encode:1": \\\{G} \ \<^bold>\x(A!x & \F (x[F] \ \{G}))[G]\ + by (rule "\I"; rule "desc-nec-encode:2"[THEN "\E"(2)]) + (meson "nec-imp-act" "vdash-properties:10") + +AOT_theorem "Box-desc-encode:2": +\\\{G} \ \(\<^bold>\x(A!x & \F (x[F] \ \{G}))[G] \ \{G})\ +proof(rule CP) + AOT_assume \\\{G}\ + AOT_hence \\\\{G}\ by (metis "S5Basic:6" "\E"(1)) + moreover AOT_have \\\\{G} \ \(\<^bold>\x(A!x & \F (x[F] \ \{G}))[G] \ \{G})\ + proof (rule RM; rule "\I") + AOT_modally_strict { + AOT_assume 1: \\\{G}\ + AOT_hence \\<^bold>\x(A!x & \F (x[F] \ \{G}))[G]\ + using "Box-desc-encode:1" "\E" by blast + moreover AOT_have \\{G}\ + using 1 by (meson "qml:2"[axiom_inst] "\E") + ultimately AOT_show \\<^bold>\x(A!x & \F (x[F] \ \{G}))[G] \ \{G}\ + using "\I" "\I" by simp + } + qed + ultimately AOT_show \\(\<^bold>\x(A!x & \F (x[F] \ \{G}))[G] \ \{G})\ + using "\E" by blast +qed + +definition rigid_condition where + \rigid_condition \ \ \v . [v \ \\ (\{\} \ \\{\})]\ +syntax rigid_condition :: \id_position \ AOT_prop\ ("RIGID'_CONDITION'(_')") + +AOT_theorem "strict-can:1[E]": + assumes \RIGID_CONDITION(\)\ + shows \\\ (\{\} \ \\{\})\ + using assms[unfolded rigid_condition_def] by auto + +AOT_theorem "strict-can:1[I]": + assumes \\<^bold>\\<^sub>\ \\ (\{\} \ \\{\})\ + shows \RIGID_CONDITION(\)\ + using assms rigid_condition_def by auto + +AOT_theorem "box-phi-a:1": + assumes \RIGID_CONDITION(\)\ + shows \(A!x & \F (x[F] \ \{F})) \ \(A!x & \F (x[F] \ \{F}))\ +proof (rule "\I") + AOT_assume a: \A!x & \F (x[F] \ \{F})\ + AOT_hence b: \\A!x\ + by (metis "Conjunction Simplification"(1) "oa-facts:2" "\E") + AOT_have \x[F] \ \{F}\ for F + using a[THEN "&E"(2)] "\E" by blast + moreover AOT_have \\(x[F] \ \x[F])\ for F + by (meson "pre-en-eq:1[1]" RN) + moreover AOT_have \\(\{F} \ \\{F})\ for F + using RN "strict-can:1[E]"[OF assms] "\E" by blast + ultimately AOT_have \\(x[F] \ \{F})\ for F + using "sc-eq-box-box:5" "qml:2"[axiom_inst, THEN "\E"] "\E" "&I" by metis + AOT_hence \\F \(x[F] \ \{F})\ by (rule GEN) + AOT_hence \\\F (x[F] \ \{F})\ by (rule BF[THEN "\E"]) + AOT_thus \\([A!]x & \F (x[F] \ \{F}))\ + using b "KBasic:3" "\S"(1) "\E"(2) by blast +qed + +AOT_theorem "box-phi-a:2": + assumes \RIGID_CONDITION(\)\ + shows \y = \<^bold>\x(A!x & \F (x[F] \ \{F})) \ (A!y & \F (y[F] \ \{F}))\ +proof(rule "\I") + AOT_assume \y = \<^bold>\x(A!x & \F (x[F] \ \{F}))\ + AOT_hence \\<^bold>\(A!y & \F (y[F] \ \{F}))\ + using "actual-desc:2"[THEN "\E"] by fast + AOT_hence abs: \\<^bold>\A!y\ and \\<^bold>\\F (y[F] \ \{F})\ + using "Act-Basic:2" "&E" "\E"(1) by blast+ + AOT_hence \\F \<^bold>\(y[F] \ \{F})\ + by (metis "\E"(1) "logic-actual-nec:3" "vdash-properties:1[2]") + AOT_hence \\<^bold>\(y[F] \ \{F})\ for F + using "\E" by blast + AOT_hence \\<^bold>\y[F] \ \<^bold>\\{F}\ for F + by (metis "Act-Basic:5" "\E"(1)) + AOT_hence \y[F] \ \{F}\ for F + using "sc-eq-fur:2"[THEN "\E", + OF "strict-can:1[E]"[OF assms, + THEN "\E"(2)[where \=F], THEN RN]] + by (metis "en-eq:10[1]" "\E"(6)) + AOT_hence \\F (y[F] \ \{F})\ by (rule GEN) + AOT_thus \[A!]y & \F (y[F] \ \{F})\ + using abs "&I" "\E"(2) "oa-facts:8" by blast +qed + +AOT_theorem "box-phi-a:3": + assumes \RIGID_CONDITION(\)\ + shows \\<^bold>\x(A!x & \F (x[F] \ \{F}))[F] \ \{F}\ + using "desc-nec-encode:2" + "sc-eq-fur:2"[THEN "\E", + OF "strict-can:1[E]"[OF assms, + THEN "\E"(2)[where \=F], THEN RN]] + "\E"(5) by blast + +AOT_define Null :: \\ \ \\ ("Null'(_')") + "df-null-uni:1": \Null(x) \\<^sub>d\<^sub>f A!x & \\F x[F]\ + +AOT_define Universal :: \\ \ \\ ("Universal'(_')") + "df-null-uni:2": \Universal(x) \\<^sub>d\<^sub>f A!x & \F x[F]\ + +AOT_theorem "null-uni-uniq:1": \\!x Null(x)\ +proof (rule "uniqueness:1"[THEN "\\<^sub>d\<^sub>fI"]) + AOT_obtain a where a_prop: \A!a & \F (a[F] \ \(F = F))\ + using "A-objects"[axiom_inst] "\E"[rotated] by fast + AOT_have a_null: \\a[F]\ for F + proof (rule "raa-cor:2") + AOT_assume \a[F]\ + AOT_hence \\(F = F)\ using a_prop[THEN "&E"(2)] "\E" "\E" by blast + AOT_hence \F = F & \(F = F)\ by (metis "id-eq:1" "raa-cor:3") + AOT_thus \p & \p\ for p by (metis "raa-cor:1") + qed + AOT_have \Null(a) & \\ (Null(\) \ \ = a)\ + proof (rule "&I") + AOT_have \\\F a[F]\ + using a_null by (metis "instantiation" "reductio-aa:1") + AOT_thus \Null(a)\ + using "df-null-uni:1"[THEN "\\<^sub>d\<^sub>fI"] a_prop[THEN "&E"(1)] "&I" by metis + next + AOT_show \\\ (Null(\) \ \ = a)\ + proof (rule GEN; rule "\I") + fix \ + AOT_assume a: \Null(\)\ + AOT_hence \\\F \[F]\ + using "df-null-uni:1"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + AOT_hence \_null: \\\[F]\ for F + by (metis "existential:2[const_var]" "reductio-aa:1") + AOT_have \\F (\[F] \ a[F])\ + apply (rule GEN; rule "\I"; rule CP) + using "raa-cor:3" \_null a_null by blast+ + moreover AOT_have \A!\\ + using a "df-null-uni:1"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + ultimately AOT_show \\ = a\ + using a_prop[THEN "&E"(1)] "ab-obey:1"[THEN "\E", THEN "\E"] + "&I" by blast + qed + qed + AOT_thus \\\ (Null(\) & \\ (Null(\) \ \ = \))\ + using "\I"(2) by fast +qed + +AOT_theorem "null-uni-uniq:2": \\!x Universal(x)\ +proof (rule "uniqueness:1"[THEN "\\<^sub>d\<^sub>fI"]) + AOT_obtain a where a_prop: \A!a & \F (a[F] \ F = F)\ + using "A-objects"[axiom_inst] "\E"[rotated] by fast + AOT_hence aF: \a[F]\ for F using "&E" "\E" "\E" "id-eq:1" by fast + AOT_hence \Universal(a)\ + using "df-null-uni:2"[THEN "\\<^sub>d\<^sub>fI"] "&I" a_prop[THEN "&E"(1)] GEN by blast + moreover AOT_have \\\ (Universal(\) \ \ = a)\ + proof (rule GEN; rule "\I") + fix \ + AOT_assume \Universal(\)\ + AOT_hence abs_\: \A!\\ and \\[F]\ for F + using "df-null-uni:2"[THEN "\\<^sub>d\<^sub>fE"] "&E" "\E" by blast+ + AOT_hence \\[F] \ a[F]\ for F + using aF by (metis "deduction-theorem" "\I") + AOT_hence \\F (\[F] \ a[F])\ by (rule GEN) + AOT_thus \\ = a\ + using a_prop[THEN "&E"(1)] "ab-obey:1"[THEN "\E", THEN "\E"] + "&I" abs_\ by blast + qed + ultimately AOT_show \\\ (Universal(\) & \\ (Universal(\) \ \ = \))\ + using "&I" "\I" by fast +qed + +AOT_theorem "null-uni-uniq:3": \\<^bold>\x Null(x)\\ + using "A-Exists:2" "RA[2]" "\E"(2) "null-uni-uniq:1" by blast + +AOT_theorem "null-uni-uniq:4": \\<^bold>\x Universal(x)\\ + using "A-Exists:2" "RA[2]" "\E"(2) "null-uni-uniq:2" by blast + +AOT_define Null_object :: \\\<^sub>s\ (\a\<^sub>\\) + "df-null-uni-terms:1": \a\<^sub>\ =\<^sub>d\<^sub>f \<^bold>\x Null(x)\ + +AOT_define Universal_object :: \\\<^sub>s\ (\a\<^sub>V\) + "df-null-uni-terms:2": \a\<^sub>V =\<^sub>d\<^sub>f \<^bold>\x Universal(x)\ + +AOT_theorem "null-uni-facts:1": \Null(x) \ \Null(x)\ +proof (rule "\I") + AOT_assume \Null(x)\ + AOT_hence x_abs: \A!x\ and x_null: \\\F x[F]\ + using "df-null-uni:1"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast+ + AOT_have \\x[F]\ for F using x_null + using "existential:2[const_var]" "reductio-aa:1" + by metis + AOT_hence \\\x[F]\ for F by (metis "en-eq:7[1]" "\E"(1)) + AOT_hence \\F \\x[F]\ by (rule GEN) + AOT_hence \\\F \x[F]\ by (rule BF[THEN "\E"]) + moreover AOT_have \\\F \x[F] \ \\\F x[F]\ + apply (rule RM) + by (metis (full_types) "instantiation" "cqt:2[const_var]"[axiom_inst] + "\I" "reductio-aa:1" "rule-ui:1") + ultimately AOT_have \\\\F x[F]\ + by (metis "\E") + moreover AOT_have \\A!x\ using x_abs + using "oa-facts:2" "vdash-properties:10" by blast + ultimately AOT_have r: \\(A!x & \\F x[F])\ + by (metis "KBasic:3" "&I" "\E"(3) "raa-cor:3") + AOT_show \\Null(x)\ + by (AOT_subst \Null(x)\ \A!x & \\F x[F]\) + (auto simp: "df-null-uni:1" "\Df" r) +qed + +AOT_theorem "null-uni-facts:2": \Universal(x) \ \Universal(x)\ +proof (rule "\I") + AOT_assume \Universal(x)\ + AOT_hence x_abs: \A!x\ and x_univ: \\F x[F]\ + using "df-null-uni:2"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast+ + AOT_have \x[F]\ for F using x_univ "\E" by blast + AOT_hence \\x[F]\ for F by (metis "en-eq:2[1]" "\E"(1)) + AOT_hence \\F \x[F]\ by (rule GEN) + AOT_hence \\\F x[F]\ by (rule BF[THEN "\E"]) + moreover AOT_have \\A!x\ using x_abs + using "oa-facts:2" "vdash-properties:10" by blast + ultimately AOT_have r: \\(A!x & \F x[F])\ + by (metis "KBasic:3" "&I" "\E"(3) "raa-cor:3") + AOT_show \\Universal(x)\ + by (AOT_subst \Universal(x)\ \A!x & \F x[F]\) + (auto simp add: "df-null-uni:2" "\Df" r) +qed + +AOT_theorem "null-uni-facts:3": \Null(a\<^sub>\)\ + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF "df-null-uni-terms:1"]) + apply (simp add: "null-uni-uniq:3") + using "actual-desc:4"[THEN "\E", OF "null-uni-uniq:3"] + "sc-eq-fur:2"[THEN "\E", + OF "null-uni-facts:1"[unvarify x, THEN RN, OF "null-uni-uniq:3"], + THEN "\E"(1)] + by blast + +AOT_theorem "null-uni-facts:4": \Universal(a\<^sub>V)\ + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF "df-null-uni-terms:2"]) + apply (simp add: "null-uni-uniq:4") + using "actual-desc:4"[THEN "\E", OF "null-uni-uniq:4"] + "sc-eq-fur:2"[THEN "\E", + OF "null-uni-facts:2"[unvarify x, THEN RN, OF "null-uni-uniq:4"], + THEN "\E"(1)] + by blast + +AOT_theorem "null-uni-facts:5": \a\<^sub>\ \ a\<^sub>V\ +proof (rule "=\<^sub>d\<^sub>fI"(2)[OF "df-null-uni-terms:1", OF "null-uni-uniq:3"]; + rule "=\<^sub>d\<^sub>fI"(2)[OF "df-null-uni-terms:2", OF "null-uni-uniq:4"]; + rule "\\<^sub>d\<^sub>fI"[OF "=-infix"]; + rule "raa-cor:2") + AOT_obtain x where nullx: \Null(x)\ + by (metis "instantiation" "df-null-uni-terms:1" "existential:1" + "null-uni-facts:3" "null-uni-uniq:3" "rule-id-df:2:b[zero]") + AOT_hence act_null: \\<^bold>\Null(x)\ + by (metis "nec-imp-act" "null-uni-facts:1" "\E") + AOT_assume \\<^bold>\x Null(x) = \<^bold>\x Universal(x)\ + AOT_hence \\<^bold>\\x(Null(x) \ Universal(x))\ + using "actual-desc:5"[THEN "\E"] by blast + AOT_hence \\x \<^bold>\(Null(x) \ Universal(x))\ + by (metis "\E"(1) "logic-actual-nec:3" "vdash-properties:1[2]") + AOT_hence \\<^bold>\Null(x) \ \<^bold>\Universal(x)\ + using "Act-Basic:5" "\E"(1) "rule-ui:3" by blast + AOT_hence \\<^bold>\Universal(x)\ using act_null "\E" by blast + AOT_hence \Universal(x)\ + by (metis RN "\E"(1) "null-uni-facts:2" "sc-eq-fur:2" "\E") + AOT_hence \\F x[F]\ using "\\<^sub>d\<^sub>fE"[OF "df-null-uni:2"] "&E" by metis + moreover AOT_have \\\F x[F]\ + using nullx "\\<^sub>d\<^sub>fE"[OF "df-null-uni:1"] "&E" by metis + ultimately AOT_show \p & \p\ for p + by (metis "cqt-further:1" "raa-cor:3" "\E") +qed + +AOT_theorem "null-uni-facts:6": \a\<^sub>\ = \<^bold>\x(A!x & \F (x[F] \ F \ F))\ +proof (rule "ab-obey:1"[unvarify x y, THEN "\E", THEN "\E"]) + AOT_show \\<^bold>\x([A!]x & \F (x[F] \ F \ F))\\ + by (simp add: "A-descriptions") +next + AOT_show \a\<^sub>\\\ + by (rule "=\<^sub>d\<^sub>fI"(2)[OF "df-null-uni-terms:1", OF "null-uni-uniq:3"]) + (simp add: "null-uni-uniq:3") +next + AOT_have \\<^bold>\x([A!]x & \F (x[F] \ F \ F))\\ + by (simp add: "A-descriptions") + AOT_hence 1: \\<^bold>\x([A!]x & \F (x[F] \ F \ F)) = \<^bold>\x([A!]x & \F (x[F] \ F \ F))\ + using "rule=I:1" by blast + AOT_show \[A!]a\<^sub>\ & [A!]\<^bold>\x([A!]x & \F (x[F] \ F \ F))\ + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF "df-null-uni-terms:1", OF "null-uni-uniq:3"]; + rule "&I") + apply (meson "\\<^sub>d\<^sub>fE" "Conjunction Simplification"(1) + "df-null-uni:1" "df-null-uni-terms:1" "null-uni-facts:3" + "null-uni-uniq:3" "rule-id-df:2:a[zero]" "\E") + using "can-ab2"[unvarify y, OF "A-descriptions", THEN "\E", OF 1]. +next + AOT_show \\F (a\<^sub>\[F] \ \<^bold>\x([A!]x & \F (x[F] \ F \ F))[F])\ + proof (rule GEN) + fix F + AOT_have \\a\<^sub>\[F]\ + by (rule "=\<^sub>d\<^sub>fI"(2)[OF "df-null-uni-terms:1", OF "null-uni-uniq:3"]) + (metis (no_types, lifting) "\\<^sub>d\<^sub>fE" "&E"(2) "\I"(2) "\E"(3) "\I"(2) + "df-null-uni:1" "df-null-uni-terms:1" "null-uni-facts:3" + "raa-cor:2" "rule-id-df:2:a[zero]" + "russell-axiom[enc,1].\_denotes_asm") + moreover AOT_have \\\<^bold>\x([A!]x & \F (x[F] \ F \ F))[F]\ + proof(rule "raa-cor:2") + AOT_assume 0: \\<^bold>\x([A!]x & \F (x[F] \ F \ F))[F]\ + AOT_hence \\<^bold>\(F \ F)\ + using "desc-nec-encode:2"[THEN "\E"(1), OF 0] by blast + moreover AOT_have \\\<^bold>\(F \ F)\ + using "\\<^sub>d\<^sub>fE" "id-act:2" "id-eq:1" "\E"(2) + "=-infix" "raa-cor:3" by blast + ultimately AOT_show \\<^bold>\(F \ F) & \\<^bold>\(F \ F)\ by (rule "&I") + qed + ultimately AOT_show \a\<^sub>\[F] \ \<^bold>\x([A!]x & \F (x[F] \ F \ F))[F]\ + using "deduction-theorem" "\I" "raa-cor:4" by blast + qed +qed + +AOT_theorem "null-uni-facts:7": \a\<^sub>V = \<^bold>\x(A!x & \F (x[F] \ F = F))\ +proof (rule "ab-obey:1"[unvarify x y, THEN "\E", THEN "\E"]) + AOT_show \\<^bold>\x([A!]x & \F (x[F] \ F = F))\\ + by (simp add: "A-descriptions") +next + AOT_show \a\<^sub>V\\ + by (rule "=\<^sub>d\<^sub>fI"(2)[OF "df-null-uni-terms:2", OF "null-uni-uniq:4"]) + (simp add: "null-uni-uniq:4") +next + AOT_have \\<^bold>\x([A!]x & \F (x[F] \ F = F))\\ + by (simp add: "A-descriptions") + AOT_hence 1: \\<^bold>\x([A!]x & \F (x[F] \ F = F)) = \<^bold>\x([A!]x & \F (x[F] \ F = F))\ + using "rule=I:1" by blast + AOT_show \[A!]a\<^sub>V & [A!]\<^bold>\x([A!]x & \F (x[F] \ F = F))\ + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF "df-null-uni-terms:2", OF "null-uni-uniq:4"]; + rule "&I") + apply (meson "\\<^sub>d\<^sub>fE" "Conjunction Simplification"(1) "df-null-uni:2" + "df-null-uni-terms:2" "null-uni-facts:4" "null-uni-uniq:4" + "rule-id-df:2:a[zero]" "\E") + using "can-ab2"[unvarify y, OF "A-descriptions", THEN "\E", OF 1]. +next + AOT_show \\F (a\<^sub>V[F] \ \<^bold>\x([A!]x & \F (x[F] \ F = F))[F])\ + proof (rule GEN) + fix F + AOT_have \a\<^sub>V[F]\ + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF "df-null-uni-terms:2", OF "null-uni-uniq:4"]) + using "\\<^sub>d\<^sub>fE" "&E"(2) "df-null-uni:2" "df-null-uni-terms:2" + "null-uni-facts:4" "null-uni-uniq:4" "rule-id-df:2:a[zero]" + "rule-ui:3" by blast + moreover AOT_have \\<^bold>\x([A!]x & \F (x[F] \ F = F))[F]\ + using "RA[2]" "desc-nec-encode:2" "id-eq:1" "\E"(2) by fastforce + ultimately AOT_show \a\<^sub>V[F] \ \<^bold>\x([A!]x & \F (x[F] \ F = F))[F]\ + using "deduction-theorem" "\I" by simp + qed +qed + +AOT_theorem "aclassical:1": +\\R\x\y(A!x & A!y & x \ y & [\z [R]zx] = [\z [R]zy])\ +proof(rule GEN) + fix R + AOT_obtain a where a_prop: + \A!a & \F (a[F] \ \y(A!y & F = [\z [R]zy] & \y[F]))\ + using "A-objects"[axiom_inst] "\E"[rotated] by fast + AOT_have a_enc: \a[\z [R]za]\ + proof (rule "raa-cor:1") + AOT_assume 0: \\a[\z [R]za]\ + AOT_hence \\\y(A!y & [\z [R]za] = [\z [R]zy] & \y[\z [R]za])\ + by (rule a_prop[THEN "&E"(2), THEN "\E"(1)[where \="\[\z [R]za]\"], + THEN "oth-class-taut:4:b"[THEN "\E"(1)], + THEN "\E"(1), rotated]) + "cqt:2[lambda]" + AOT_hence \\y \(A!y & [\z [R]za] = [\z [R]zy] & \y[\z [R]za])\ + using "cqt-further:4" "vdash-properties:10" by blast + AOT_hence \\(A!a & [\z [R]za] = [\z [R]za] & \a[\z [R]za])\ + using "\E" by blast + AOT_hence \(A!a & [\z [R]za] = [\z [R]za]) \ a[\z [R]za]\ + by (metis "&I" "deduction-theorem" "raa-cor:3") + moreover AOT_have \[\z [R]za] = [\z [R]za]\ + by (rule "=I") "cqt:2[lambda]" + ultimately AOT_have \a[\z [R]za]\ + using a_prop[THEN "&E"(1)] "\E" "&I" by blast + AOT_thus \a[\z [R]za] & \a[\z [R]za]\ + using 0 "&I" by blast + qed + AOT_hence \\y(A!y & [\z [R]za] = [\z [R]zy] & \y[\z [R]za])\ + by (rule a_prop[THEN "&E"(2), THEN "\E"(1), THEN "\E"(1), rotated]) + "cqt:2" + then AOT_obtain b where b_prop: + \A!b & [\z [R]za] = [\z [R]zb] & \b[\z [R]za]\ + using "\E"[rotated] by blast + AOT_have \a \ b\ + apply (rule "\\<^sub>d\<^sub>fI"[OF "=-infix"]) + using a_enc b_prop[THEN "&E"(2)] + using "\\I" "rule=E" id_sym "\E"(4) "oth-class-taut:3:a" + "raa-cor:3" "reductio-aa:1" by fast + AOT_hence \A!a & A!b & a \ b & [\z [R]za] = [\z [R]zb]\ + using b_prop "&E" a_prop "&I" by meson + AOT_hence \\y (A!a & A!y & a \ y & [\z [R]za] = [\z [R]zy])\ by (rule "\I") + AOT_thus \\x\y (A!x & A!y & x \ y & [\z [R]zx] = [\z [R]zy])\ by (rule "\I") +qed + +AOT_theorem "aclassical:2": + \\R\x\y(A!x & A!y & x \ y & [\z [R]xz] = [\z [R]yz])\ +proof(rule GEN) + fix R + AOT_obtain a where a_prop: + \A!a & \F (a[F] \ \y(A!y & F = [\z [R]yz] & \y[F]))\ + using "A-objects"[axiom_inst] "\E"[rotated] by fast + AOT_have a_enc: \a[\z [R]az]\ + proof (rule "raa-cor:1") + AOT_assume 0: \\a[\z [R]az]\ + AOT_hence \\\y(A!y & [\z [R]az] = [\z [R]yz] & \y[\z [R]az])\ + by (rule a_prop[THEN "&E"(2), THEN "\E"(1)[where \="\[\z [R]az]\"], + THEN "oth-class-taut:4:b"[THEN "\E"(1)], + THEN "\E"(1), rotated]) + "cqt:2[lambda]" + AOT_hence \\y \(A!y & [\z [R]az] = [\z [R]yz] & \y[\z [R]az])\ + using "cqt-further:4" "vdash-properties:10" by blast + AOT_hence \\(A!a & [\z [R]az] = [\z [R]az] & \a[\z [R]az])\ + using "\E" by blast + AOT_hence \(A!a & [\z [R]az] = [\z [R]az]) \ a[\z [R]az]\ + by (metis "&I" "deduction-theorem" "raa-cor:3") + moreover AOT_have \[\z [R]az] = [\z [R]az]\ + by (rule "=I") "cqt:2[lambda]" + ultimately AOT_have \a[\z [R]az]\ + using a_prop[THEN "&E"(1)] "\E" "&I" by blast + AOT_thus \a[\z [R]az] & \a[\z [R]az]\ + using 0 "&I" by blast + qed + AOT_hence \\y(A!y & [\z [R]az] = [\z [R]yz] & \y[\z [R]az])\ + by (rule a_prop[THEN "&E"(2), THEN "\E"(1), THEN "\E"(1), rotated]) + "cqt:2" + then AOT_obtain b where b_prop: + \A!b & [\z [R]az] = [\z [R]bz] & \b[\z [R]az]\ + using "\E"[rotated] by blast + AOT_have \a \ b\ + apply (rule "\\<^sub>d\<^sub>fI"[OF "=-infix"]) + using a_enc b_prop[THEN "&E"(2)] + using "\\I" "rule=E" id_sym "\E"(4) "oth-class-taut:3:a" + "raa-cor:3" "reductio-aa:1" by fast + AOT_hence \A!a & A!b & a \ b & [\z [R]az] = [\z [R]bz]\ + using b_prop "&E" a_prop "&I" by meson + AOT_hence \\y (A!a & A!y & a \ y & [\z [R]az] = [\z [R]yz])\ by (rule "\I") + AOT_thus \\x\y (A!x & A!y & x \ y & [\z [R]xz] = [\z [R]yz])\ by (rule "\I") +qed + +AOT_theorem "aclassical:3": + \\F\x\y(A!x & A!y & x \ y & [\ [F]x] = [\ [F]y])\ +proof(rule GEN) + fix R + AOT_obtain a where a_prop: + \A!a & \F (a[F] \ \y(A!y & F = [\z [R]y] & \y[F]))\ + using "A-objects"[axiom_inst] "\E"[rotated] by fast + AOT_have den: \[\z [R]a]\\ by "cqt:2[lambda]" + AOT_have a_enc: \a[\z [R]a]\ + proof (rule "raa-cor:1") + AOT_assume 0: \\a[\z [R]a]\ + AOT_hence \\\y(A!y & [\z [R]a] = [\z [R]y] & \y[\z [R]a])\ + by (safe intro!: a_prop[THEN "&E"(2), THEN "\E"(1)[where \=\\[\z [R]a]\\], + THEN "oth-class-taut:4:b"[THEN "\E"(1)], + THEN "\E"(1), rotated] "cqt:2") + AOT_hence \\y \(A!y & [\z [R]a] = [\z [R]y] & \y[\z [R]a])\ + using "cqt-further:4" "\E" by blast + AOT_hence \\(A!a & [\z [R]a] = [\z [R]a] & \a[\z [R]a])\ using "\E" by blast + AOT_hence \(A!a & [\z [R]a] = [\z [R]a]) \ a[\z [R]a]\ + by (metis "&I" "deduction-theorem" "raa-cor:3") + AOT_hence \a[\z [R]a]\ + using a_prop[THEN "&E"(1)] "\E" "&I" + by (metis "rule=I:1" den) + AOT_thus \a[\z [R]a] & \a[\z [R]a]\ by (metis "0" "raa-cor:3") + qed + AOT_hence \\y(A!y & [\z [R]a] = [\z [R]y] & \y[\z [R]a])\ + by (rule a_prop[THEN "&E"(2), THEN "\E"(1), OF den, THEN "\E"(1), rotated]) + then AOT_obtain b where b_prop: \A!b & [\z [R]a] = [\z [R]b] & \b[\z [R]a]\ + using "\E"[rotated] by blast + AOT_have 1: \a \ b\ + apply (rule "\\<^sub>d\<^sub>fI"[OF "=-infix"]) + using a_enc b_prop[THEN "&E"(2)] + using "\\I" "rule=E" id_sym "\E"(4) "oth-class-taut:3:a" + "raa-cor:3" "reductio-aa:1" by fast + AOT_have a: \[\ [R]a] = ([R]a)\ + apply (rule "lambda-predicates:3[zero]"[axiom_inst, unvarify p]) + by (meson "log-prop-prop:2") + AOT_have b: \[\ [R]b] = ([R]b)\ + apply (rule "lambda-predicates:3[zero]"[axiom_inst, unvarify p]) + by (meson "log-prop-prop:2") + AOT_have \[\ [R]a] = [\ [R]b]\ + apply (rule "rule=E"[rotated, OF a[THEN id_sym]]) + apply (rule "rule=E"[rotated, OF b[THEN id_sym]]) + apply (rule "identity:4"[THEN "\\<^sub>d\<^sub>fI", OF "&I", rotated]) + using b_prop "&E" apply blast + apply (safe intro!: "&I") + by (simp add: "log-prop-prop:2")+ + AOT_hence \A!a & A!b & a \ b & [\ [R]a] = [\ [R]b]\ + using 1 a_prop[THEN "&E"(1)] b_prop[THEN "&E"(1), THEN "&E"(1)] + "&I" by auto + AOT_hence \\y (A!a & A!y & a \ y & [\ [R]a] = [\ [R]y])\ by (rule "\I") + AOT_thus \\x\y (A!x & A!y & x \ y & [\ [R]x] = [\ [R]y])\ by (rule "\I") +qed + +AOT_theorem aclassical2: \\x\y (A!x & A!y & x \ y & \F ([F]x \ [F]y))\ +proof - + AOT_have \\x \y ([A!]x & [A!]y & x \ y & + [\z [\xy \F ([F]x \ [F]y)]zx] = + [\z [\xy \F ([F]x \ [F]y)]zy])\ + by (rule "aclassical:1"[THEN "\E"(1)[where \="\[\xy \F ([F]x \ [F]y)]\"]]) + "cqt:2" + then AOT_obtain x where \\y ([A!]x & [A!]y & x \ y & + [\z [\xy \F ([F]x \ [F]y)]zx] = + [\z [\xy \F ([F]x \ [F]y)]zy])\ + using "\E"[rotated] by blast + then AOT_obtain y where 0: \([A!]x & [A!]y & x \ y & + [\z [\xy \F ([F]x \ [F]y)]zx] = + [\z [\xy \F ([F]x \ [F]y)]zy])\ + using "\E"[rotated] by blast + AOT_have \[\z [\xy \F ([F]x \ [F]y)]zx]x\ + by (auto intro!: "\\C"(1) "cqt:2" + simp: "&I" "ex:1:a" prod_denotesI "rule-ui:3" + "oth-class-taut:3:a" "universal-cor") + AOT_hence \[\z [\xy \F ([F]x \ [F]y)]zy]x\ + by (rule "rule=E"[rotated, OF 0[THEN "&E"(2)]]) + AOT_hence \[\xy \F ([F]x \ [F]y)]xy\ + by (rule "\\C"(1)) + AOT_hence \\F ([F]x \ [F]y)\ + using "\\C"(1) old.prod.case by fast + AOT_hence \[A!]x & [A!]y & x \ y & \F ([F]x \ [F]y)\ + using 0 "&E" "&I" by blast + AOT_hence \\y ([A!]x & [A!]y & x \ y & \F ([F]x \ [F]y))\ by (rule "\I") + AOT_thus \\x\y ([A!]x & [A!]y & x \ y & \F ([F]x \ [F]y))\ by (rule "\I"(2)) +qed + +AOT_theorem "kirchner-thm:1": + \[\x \{x}]\ \ \\x\y(\F([F]x \ [F]y) \ (\{x} \ \{y}))\ +proof(rule "\I"; rule "\I") + AOT_assume \[\x \{x}]\\ + AOT_hence \\[\x \{x}]\\ by (metis "exist-nec" "vdash-properties:10") + moreover AOT_have \\[\x \{x}]\ \ \\x\y(\F([F]x \ [F]y) \ (\{x} \ \{y}))\ + proof (rule "RM:1"; rule "\I"; rule GEN; rule GEN; rule "\I") + AOT_modally_strict { + fix x y + AOT_assume 0: \[\x \{x}]\\ + moreover AOT_assume \\F([F]x \ [F]y)\ + ultimately AOT_have \[\x \{x}]x \ [\x \{x}]y\ + using "\E" by blast + AOT_thus \(\{x} \ \{y})\ + using "beta-C-meta"[THEN "\E", OF 0] "\E"(6) by meson + } + qed + ultimately AOT_show \\\x\y(\F([F]x \ [F]y) \ (\{x} \ \{y}))\ + using "\E" by blast +next + AOT_have \\\x\y(\F([F]x \ [F]y) \ (\{x} \ \{y})) \ + \\y(\x(\F([F]x \ [F]y) & \{x}) \ \{y})\ + proof(rule "RM:1"; rule "\I"; rule GEN) + AOT_modally_strict { + AOT_assume \\x\y(\F([F]x \ [F]y) \ (\{x} \ \{y}))\ + AOT_hence indisc: \\{x} \ \{y}\ if \\F([F]x \ [F]y)\ for x y + using "\E"(2) "\E" that by blast + AOT_show \(\x(\F([F]x \ [F]y) & \{x}) \ \{y})\ for y + proof (rule "raa-cor:1") + AOT_assume \\(\x(\F([F]x \ [F]y) & \{x}) \ \{y})\ + AOT_hence \(\x(\F([F]x \ [F]y) & \{x}) & \\{y}) \ + (\(\x(\F([F]x \ [F]y) & \{x})) & \{y})\ + using "\E"(1) "oth-class-taut:4:h" by blast + moreover { + AOT_assume 0: \\x(\F([F]x \ [F]y) & \{x}) & \\{y}\ + AOT_obtain a where \\F([F]a \ [F]y) & \{a}\ + using "\E"[rotated, OF 0[THEN "&E"(1)]] by blast + AOT_hence \\{y}\ + using indisc[THEN "\E"(1)] "&E" by blast + AOT_hence \p & \p\ for p + using 0[THEN "&E"(2)] "&I" "raa-cor:3" by blast + } + moreover { + AOT_assume 0: \(\(\x(\F([F]x \ [F]y) & \{x})) & \{y})\ + AOT_hence \\x \(\F([F]x \ [F]y) & \{x})\ + using "&E"(1) "cqt-further:4" "\E" by blast + AOT_hence \\(\F([F]y \ [F]y) & \{y})\ + using "\E" by blast + AOT_hence \\\F([F]y \ [F]y) \ \\{y}\ + using "\E"(1) "oth-class-taut:5:c" by blast + moreover AOT_have \\F([F]y \ [F]y)\ + by (simp add: "oth-class-taut:3:a" "universal-cor") + ultimately AOT_have \\\{y}\ by (metis "\\I" "\E"(2)) + AOT_hence \p & \p\ for p + using 0[THEN "&E"(2)] "&I" "raa-cor:3" by blast + } + ultimately AOT_show \p & \p\ for p + using "\E"(3) "raa-cor:1" by blast + qed + } + qed + moreover AOT_assume \\\x\y(\F([F]x \ [F]y) \ (\{x} \ \{y}))\ + ultimately AOT_have \\\y(\x(\F([F]x \ [F]y) & \{x}) \ \{y})\ + using "\E" by blast + AOT_thus \[\x \{x}]\\ + by (rule "safe-ext"[axiom_inst, THEN "\E", OF "&I", rotated]) "cqt:2" +qed + +AOT_theorem "kirchner-thm:2": + \[\x\<^sub>1...x\<^sub>n \{x\<^sub>1...x\<^sub>n}]\ \ \\x\<^sub>1...\x\<^sub>n\y\<^sub>1...\y\<^sub>n + (\F([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) \ (\{x\<^sub>1...x\<^sub>n} \ \{y\<^sub>1...y\<^sub>n}))\ +proof(rule "\I"; rule "\I") + AOT_assume \[\x\<^sub>1...x\<^sub>n \{x\<^sub>1...x\<^sub>n}]\\ + AOT_hence \\[\x\<^sub>1...x\<^sub>n \{x\<^sub>1...x\<^sub>n}]\\ by (metis "exist-nec" "\E") + moreover AOT_have \\[\x\<^sub>1...x\<^sub>n \{x\<^sub>1...x\<^sub>n}]\ \ \\x\<^sub>1...\x\<^sub>n\y\<^sub>1...\y\<^sub>n + (\F([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) \ (\{x\<^sub>1...x\<^sub>n} \ \{y\<^sub>1...y\<^sub>n}))\ + proof (rule "RM:1"; rule "\I"; rule GEN; rule GEN; rule "\I") + AOT_modally_strict { + fix x\<^sub>1x\<^sub>n y\<^sub>1y\<^sub>n :: \'a AOT_var\ + AOT_assume 0: \[\x\<^sub>1...x\<^sub>n \{x\<^sub>1...x\<^sub>n}]\\ + moreover AOT_assume \\F([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n)\ + ultimately AOT_have \[\x\<^sub>1...x\<^sub>n \{x\<^sub>1...x\<^sub>n}]x\<^sub>1...x\<^sub>n \ + [\x\<^sub>1...x\<^sub>n \{x\<^sub>1...x\<^sub>n}]y\<^sub>1...y\<^sub>n\ + using "\E" by blast + AOT_thus \(\{x\<^sub>1...x\<^sub>n} \ \{y\<^sub>1...y\<^sub>n})\ + using "beta-C-meta"[THEN "\E", OF 0] "\E"(6) by meson + } + qed + ultimately AOT_show \\\x\<^sub>1...\x\<^sub>n\y\<^sub>1...\y\<^sub>n( + \F([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) \ (\{x\<^sub>1...x\<^sub>n} \ \{y\<^sub>1...y\<^sub>n}) + )\ + using "\E" by blast +next + AOT_have \ + \(\x\<^sub>1...\x\<^sub>n\y\<^sub>1...\y\<^sub>n + (\F([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) \ (\{x\<^sub>1...x\<^sub>n} \ \{y\<^sub>1...y\<^sub>n}))) + \ \\y\<^sub>1...\y\<^sub>n + ((\x\<^sub>1...\x\<^sub>n(\F([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) & \{x\<^sub>1...x\<^sub>n})) \ + \{y\<^sub>1...y\<^sub>n})\ + proof(rule "RM:1"; rule "\I"; rule GEN) + AOT_modally_strict { + AOT_assume \\x\<^sub>1...\x\<^sub>n\y\<^sub>1...\y\<^sub>n + (\F([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) \ (\{x\<^sub>1...x\<^sub>n} \ \{y\<^sub>1...y\<^sub>n}))\ + AOT_hence indisc: \\{x\<^sub>1...x\<^sub>n} \ \{y\<^sub>1...y\<^sub>n}\ + if \\F([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n)\ for x\<^sub>1x\<^sub>n y\<^sub>1y\<^sub>n + using "\E"(2) "\E" that by blast + AOT_show \(\x\<^sub>1...\x\<^sub>n(\F([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) & \{x\<^sub>1...x\<^sub>n})) \ + \{y\<^sub>1...y\<^sub>n}\ for y\<^sub>1y\<^sub>n + proof (rule "raa-cor:1") + AOT_assume \\((\x\<^sub>1...\x\<^sub>n(\F([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) & \{x\<^sub>1...x\<^sub>n})) \ + \{y\<^sub>1...y\<^sub>n})\ + AOT_hence \((\x\<^sub>1...\x\<^sub>n(\F([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) + & \{x\<^sub>1...x\<^sub>n})) + & \\{y\<^sub>1...y\<^sub>n}) \ + (\(\x\<^sub>1...\x\<^sub>n(\F([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) & \{x\<^sub>1...x\<^sub>n})) + & \{y\<^sub>1...y\<^sub>n})\ + using "\E"(1) "oth-class-taut:4:h" by blast + moreover { + AOT_assume 0: \(\x\<^sub>1...\x\<^sub>n(\F([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) & \{x\<^sub>1...x\<^sub>n})) + & \\{y\<^sub>1...y\<^sub>n}\ + AOT_obtain a\<^sub>1a\<^sub>n where \\F([F]a\<^sub>1...a\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) & \{a\<^sub>1...a\<^sub>n}\ + using "\E"[rotated, OF 0[THEN "&E"(1)]] by blast + AOT_hence \\{y\<^sub>1...y\<^sub>n}\ + using indisc[THEN "\E"(1)] "&E" by blast + AOT_hence \p & \p\ for p + using 0[THEN "&E"(2)] "&I" "raa-cor:3" by blast + } + moreover { + AOT_assume 0: \\(\x\<^sub>1...\x\<^sub>n(\F([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) & \{x\<^sub>1...x\<^sub>n})) + & \{y\<^sub>1...y\<^sub>n}\ + AOT_hence \\x\<^sub>1...\x\<^sub>n \(\F([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) & \{x\<^sub>1...x\<^sub>n})\ + using "&E"(1) "cqt-further:4" "\E" by blast + AOT_hence \\(\F([F]y\<^sub>1...y\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) & \{y\<^sub>1...y\<^sub>n})\ + using "\E" by blast + AOT_hence \\\F([F]y\<^sub>1...y\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) \ \\{y\<^sub>1...y\<^sub>n}\ + using "\E"(1) "oth-class-taut:5:c" by blast + moreover AOT_have \\F([F]y\<^sub>1...y\<^sub>n \ [F]y\<^sub>1...y\<^sub>n)\ + by (simp add: "oth-class-taut:3:a" "universal-cor") + ultimately AOT_have \\\{y\<^sub>1...y\<^sub>n}\ + by (metis "\\I" "\E"(2)) + AOT_hence \p & \p\ for p + using 0[THEN "&E"(2)] "&I" "raa-cor:3" by blast + } + ultimately AOT_show \p & \p\ for p + using "\E"(3) "raa-cor:1" by blast + qed + } + qed + moreover AOT_assume \\\x\<^sub>1...\x\<^sub>n\y\<^sub>1...\y\<^sub>n + (\F([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) \ (\{x\<^sub>1...x\<^sub>n} \ \{y\<^sub>1...y\<^sub>n}))\ + ultimately AOT_have \\\y\<^sub>1...\y\<^sub>n + ((\x\<^sub>1...\x\<^sub>n(\F([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) & \{x\<^sub>1...x\<^sub>n})) \ + \{y\<^sub>1...y\<^sub>n})\ + using "\E" by blast + AOT_thus \[\x\<^sub>1...x\<^sub>n \{x\<^sub>1...x\<^sub>n}]\\ + by (rule "safe-ext"[axiom_inst, THEN "\E", OF "&I", rotated]) "cqt:2" +qed + +AOT_theorem "kirchner-thm-cor:1": + \[\x \{x}]\ \ \x\y(\F([F]x \ [F]y) \ \(\{x} \ \{y}))\ +proof(rule "\I"; rule GEN; rule GEN; rule "\I") + fix x y + AOT_assume \[\x \{x}]\\ + AOT_hence \\\x\y (\F ([F]x \ [F]y) \ (\{x} \ \{y}))\ + by (rule "kirchner-thm:1"[THEN "\E"(1)]) + AOT_hence \\x\\y (\F ([F]x \ [F]y) \ (\{x} \ \{y}))\ + using CBF[THEN "\E"] by blast + AOT_hence \\\y (\F ([F]x \ [F]y) \ (\{x} \ \{y}))\ + using "\E" by blast + AOT_hence \\y \(\F ([F]x \ [F]y) \ (\{x} \ \{y}))\ + using CBF[THEN "\E"] by blast + AOT_hence \\(\F ([F]x \ [F]y) \ (\{x} \ \{y}))\ + using "\E" by blast + AOT_hence \\\F ([F]x \ [F]y) \ \(\{x} \ \{y})\ + using "qml:1"[axiom_inst] "vdash-properties:6" by blast + moreover AOT_assume \\F([F]x \ [F]y)\ + ultimately AOT_show \\(\{x} \ \{y})\ using "\E" "ind-nec" by blast +qed + +AOT_theorem "kirchner-thm-cor:2": + \[\x\<^sub>1...x\<^sub>n \{x\<^sub>1...x\<^sub>n}]\ \ \x\<^sub>1...\x\<^sub>n\y\<^sub>1...\y\<^sub>n + (\F([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) \ \(\{x\<^sub>1...x\<^sub>n} \ \{y\<^sub>1...y\<^sub>n}))\ +proof(rule "\I"; rule GEN; rule GEN; rule "\I") + fix x\<^sub>1x\<^sub>n y\<^sub>1y\<^sub>n + AOT_assume \[\x\<^sub>1...x\<^sub>n \{x\<^sub>1...x\<^sub>n}]\\ + AOT_hence 0: \\\x\<^sub>1...\x\<^sub>n\y\<^sub>1...\y\<^sub>n + (\F ([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) \ (\{x\<^sub>1...x\<^sub>n} \ \{y\<^sub>1...y\<^sub>n}))\ + by (rule "kirchner-thm:2"[THEN "\E"(1)]) + AOT_have \\x\<^sub>1...\x\<^sub>n\y\<^sub>1...\y\<^sub>n + \(\F ([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) \ (\{x\<^sub>1...x\<^sub>n} \ \{y\<^sub>1...y\<^sub>n}))\ + proof(rule GEN; rule GEN) + fix x\<^sub>1x\<^sub>n y\<^sub>1y\<^sub>n + AOT_show \\(\F ([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) \ (\{x\<^sub>1...x\<^sub>n} \ \{y\<^sub>1...y\<^sub>n}))\ + apply (rule "RM:1"[THEN "\E", rotated, OF 0]; rule "\I") + using "\E" by blast + qed + AOT_hence \\y\<^sub>1...\y\<^sub>n \(\F ([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) \ + (\{x\<^sub>1...x\<^sub>n} \ \{y\<^sub>1...y\<^sub>n}))\ + using "\E" by blast + AOT_hence \\(\F ([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) \ (\{x\<^sub>1...x\<^sub>n} \ \{y\<^sub>1...y\<^sub>n}))\ + using "\E" by blast + AOT_hence \\(\F ([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) \ (\{x\<^sub>1...x\<^sub>n} \ \{y\<^sub>1...y\<^sub>n}))\ + using "\E" by blast + AOT_hence 0: \\\F ([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) \ \(\{x\<^sub>1...x\<^sub>n} \ \{y\<^sub>1...y\<^sub>n})\ + using "qml:1"[axiom_inst] "vdash-properties:6" by blast + moreover AOT_assume \\F([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n)\ + moreover AOT_have \[\x\<^sub>1...x\<^sub>n \\F ([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n)]\\ by "cqt:2" + ultimately AOT_have \[\x\<^sub>1...x\<^sub>n \\F ([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n)]x\<^sub>1...x\<^sub>n \ + [\x\<^sub>1...x\<^sub>n \\F ([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n)]y\<^sub>1...y\<^sub>n\ + using "\E" by blast + moreover AOT_have \[\x\<^sub>1...x\<^sub>n \\F ([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n)]y\<^sub>1...y\<^sub>n\ + apply (rule "\\C"(1)) + apply "cqt:2[lambda]" + apply (fact "cqt:2[const_var]"[axiom_inst]) + by (simp add: RN GEN "oth-class-taut:3:a") + ultimately AOT_have \[\x\<^sub>1...x\<^sub>n \\F ([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n)]x\<^sub>1...x\<^sub>n\ + using "\E"(2) by blast + AOT_hence \\\F ([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n)\ + using "\\C"(1) by blast + AOT_thus \\(\{x\<^sub>1...x\<^sub>n} \ \{y\<^sub>1...y\<^sub>n})\ using "\E" 0 by blast +qed + +subsection\Propositional Properties\ +text\\label{PLM: 9.12}\ + +AOT_define propositional :: \\ \ \\ (\Propositional'(_')\) + "prop-prop1": \Propositional([F]) \\<^sub>d\<^sub>f \p(F = [\y p])\ + +AOT_theorem "prop-prop2:1": \\p [\y p]\\ + by (rule GEN) "cqt:2[lambda]" + +AOT_theorem "prop-prop2:2": \[\\ \]\\ + by "cqt:2[lambda]" + +AOT_theorem "prop-prop2:3": \F = [\y p] \ \\x([F]x \ p)\ +proof (rule "\I") + AOT_assume 0: \F = [\y p]\ + AOT_show \\\x([F]x \ p)\ + by (rule "rule=E"[rotated, OF 0[symmetric]]; + rule RN; rule GEN; rule "beta-C-meta"[THEN "\E"]) + "cqt:2[lambda]" +qed + +AOT_theorem "prop-prop2:4": \Propositional([F]) \ \Propositional([F])\ +proof(rule "\I") + AOT_assume \Propositional([F])\ + AOT_hence \\p(F = [\y p])\ + using "\\<^sub>d\<^sub>fE"[OF "prop-prop1"] by blast + then AOT_obtain p where \F = [\y p]\ + using "\E"[rotated] by blast + AOT_hence \\(F = [\y p])\ + using "id-nec:2" "modus-tollens:1" "raa-cor:3" by blast + AOT_hence \\p \(F = [\y p])\ + using "\I" by fast + AOT_hence 0: \\\p (F = [\y p])\ + by (metis Buridan "vdash-properties:10") + AOT_thus \\Propositional([F])\ + using "prop-prop1"[THEN "\Df"] + by (AOT_subst \Propositional([F])\ \\p (F = [\y p])\) auto +qed + +AOT_define indicriminate :: \\ \ \\ ("Indiscriminate'(_')") + "prop-indis": \Indiscriminate([F]) \\<^sub>d\<^sub>f F\ & \(\x [F]x \ \x [F]x)\ + +AOT_theorem "prop-in-thm": \Propositional([\]) \ Indiscriminate([\])\ +proof(rule "\I") + AOT_assume \Propositional([\])\ + AOT_hence \\p \ = [\y p]\ using "\\<^sub>d\<^sub>fE"[OF "prop-prop1"] by blast + then AOT_obtain p where \_def: \\ = [\y p]\ using "\E"[rotated] by blast + AOT_show \Indiscriminate([\])\ + proof (rule "\\<^sub>d\<^sub>fI"[OF "prop-indis"]; rule "&I") + AOT_show \\\\ + using \_def by (meson "t=t-proper:1" "vdash-properties:6") + next + AOT_show \\(\x [\]x \ \x [\]x)\ + proof (rule "rule=E"[rotated, OF \_def[symmetric]]; + rule RN; rule "\I"; rule GEN) + AOT_modally_strict { + AOT_assume \\x [\y p]x\ + then AOT_obtain a where \[\y p]a\ using "\E"[rotated] by blast + AOT_hence 0: \p\ by (metis "\\C"(1)) + AOT_show \[\y p]x\ for x + apply (rule "\\C"(1)) + apply "cqt:2[lambda]" + apply (fact "cqt:2[const_var]"[axiom_inst]) + by (fact 0) + } + qed + qed +qed + +AOT_theorem "prop-in-f:1": \Necessary([F]) \ Indiscriminate([F])\ +proof (rule "\I") + AOT_assume \Necessary([F])\ + AOT_hence 0: \\\x\<^sub>1...\x\<^sub>n [F]x\<^sub>1...x\<^sub>n\ + using "\\<^sub>d\<^sub>fE"[OF "contingent-properties:1"] by blast + AOT_show \Indiscriminate([F])\ + by (rule "\\<^sub>d\<^sub>fI"[OF "prop-indis"]) + (metis "0" "KBasic:1" "&I" "ex:1:a" "rule-ui:2[const_var]" "\E") +qed + +AOT_theorem "prop-in-f:2": \Impossible([F]) \ Indiscriminate([F])\ +proof (rule "\I") + AOT_modally_strict { + AOT_have \\x \[F]x \ (\x [F]x \ \x [F]x)\ + by (metis "\E" "cqt-orig:3" "Hypothetical Syllogism" "\I" "raa-cor:3") + } + AOT_hence 0: \\\x \[F]x \ \(\x [F]x \ \x [F]x)\ + by (rule "RM:1") + AOT_assume \Impossible([F])\ + AOT_hence \\\x \[F]x\ + using "\\<^sub>d\<^sub>fE"[OF "contingent-properties:2"] "&E" by blast + AOT_hence 1: \\(\x [F]x \ \x [F]x)\ + using 0 "\E" by blast + AOT_show \Indiscriminate([F])\ + by (rule "\\<^sub>d\<^sub>fI"[OF "prop-indis"]; rule "&I") + (simp add: "ex:1:a" "rule-ui:2[const_var]" 1)+ +qed + +AOT_theorem "prop-in-f:3:a": \\Indiscriminate([E!])\ +proof(rule "raa-cor:2") + AOT_assume \Indiscriminate([E!])\ + AOT_hence 0: \\(\x [E!]x \ \x [E!]x)\ + using "\\<^sub>d\<^sub>fE"[OF "prop-indis"] "&E" by blast + AOT_hence \\\x [E!]x \ \\x [E!]x\ + using "KBasic:13" "vdash-properties:10" by blast + moreover AOT_have \\\x [E!]x\ + by (simp add: "thm-cont-e:3") + ultimately AOT_have \\\x [E!]x\ + by (metis "vdash-properties:6") + AOT_thus \p & \p\ for p + by (metis "\\<^sub>d\<^sub>fE" "conventions:5" "o-objects-exist:5" "reductio-aa:1") +qed + +AOT_theorem "prop-in-f:3:b": \\Indiscriminate([E!]\<^sup>-)\ +proof (rule "rule=E"[rotated, OF "rel-neg-T:2"[symmetric]]; + rule "raa-cor:2") + AOT_assume \Indiscriminate([\x \[E!]x])\ + AOT_hence 0: \\(\x [\x \[E!]x]x \ \x [\x \[E!]x]x)\ + using "\\<^sub>d\<^sub>fE"[OF "prop-indis"] "&E" by blast + AOT_hence \\\x [\x \[E!]x]x \ \\x [\x \[E!]x]x\ + using "\E" "qml:1" "vdash-properties:1[2]" by blast + moreover AOT_have \\\x [\x \[E!]x]x\ + apply (AOT_subst \[\x \E!x]x\ \\E!x\ for: x) + apply (rule "beta-C-meta"[THEN "\E"]) + apply "cqt:2" + by (metis (full_types) "B\" RN "T\" "cqt-further:2" + "o-objects-exist:5" "\E") + ultimately AOT_have 1: \\\x [\x \[E!]x]x\ + by (metis "vdash-properties:6") + AOT_hence \\\x \[E!]x\ + by (AOT_subst (reverse) \\[E!]x\ \[\x \[E!]x]x\ for: x) + (auto intro!: "cqt:2" "beta-C-meta"[THEN "\E"]) + AOT_hence \\x \\[E!]x\ by (metis "CBF" "vdash-properties:10") + moreover AOT_obtain a where abs_a: \O!a\ + using "\E" "o-objects-exist:1" "qml:2"[axiom_inst] "\E" by blast + ultimately AOT_have \\\[E!]a\ using "\E" by blast + AOT_hence 2: \\\[E!]a\ by (metis "\\<^sub>d\<^sub>fE" "conventions:5" "reductio-aa:1") + AOT_have \A!a\ + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF AOT_abstract]) + apply "cqt:2[lambda]" + apply (rule "\\C"(1)) + apply "cqt:2[lambda]" + using "cqt:2[const_var]"[axiom_inst] apply blast + by (fact 2) + AOT_thus \p & \p\ for p using abs_a + by (metis "\E"(1) "oa-contingent:2" "reductio-aa:1") +qed + +AOT_theorem "prop-in-f:3:c": \\Indiscriminate(O!)\ +proof(rule "raa-cor:2") + AOT_assume \Indiscriminate(O!)\ + AOT_hence 0: \\(\x O!x \ \x O!x)\ + using "\\<^sub>d\<^sub>fE"[OF "prop-indis"] "&E" by blast + AOT_hence \\\x O!x \ \\x O!x\ + using "qml:1"[axiom_inst] "vdash-properties:6" by blast + moreover AOT_have \\\x O!x\ + using "o-objects-exist:1" by blast + ultimately AOT_have \\\x O!x\ + by (metis "vdash-properties:6") + AOT_thus \p & \p\ for p + by (metis "o-objects-exist:3" "qml:2"[axiom_inst] "raa-cor:3" "\E") +qed + +AOT_theorem "prop-in-f:3:d": \\Indiscriminate(A!)\ +proof(rule "raa-cor:2") + AOT_assume \Indiscriminate(A!)\ + AOT_hence 0: \\(\x A!x \ \x A!x)\ + using "\\<^sub>d\<^sub>fE"[OF "prop-indis"] "&E" by blast + AOT_hence \\\x A!x \ \\x A!x\ + using "qml:1"[axiom_inst] "vdash-properties:6" by blast + moreover AOT_have \\\x A!x\ + using "o-objects-exist:2" by blast + ultimately AOT_have \\\x A!x\ + by (metis "vdash-properties:6") + AOT_thus \p & \p\ for p + by (metis "o-objects-exist:4" "qml:2"[axiom_inst] "raa-cor:3" "\E") +qed + +AOT_theorem "prop-in-f:4:a": \\Propositional(E!)\ + using "modus-tollens:1" "prop-in-f:3:a" "prop-in-thm" by blast + +AOT_theorem "prop-in-f:4:b": \\Propositional(E!\<^sup>-)\ + using "modus-tollens:1" "prop-in-f:3:b" "prop-in-thm" by blast + +AOT_theorem "prop-in-f:4:c": \\Propositional(O!)\ + using "modus-tollens:1" "prop-in-f:3:c" "prop-in-thm" by blast + +AOT_theorem "prop-in-f:4:d": \\Propositional(A!)\ + using "modus-tollens:1" "prop-in-f:3:d" "prop-in-thm" by blast + +AOT_theorem "prop-prop-nec:1": \\\p (F = [\y p]) \ \p(F = [\y p])\ +proof(rule "\I") + AOT_assume \\\p (F = [\y p])\ + AOT_hence \\p \(F = [\y p])\ + by (metis "BF\" "\E") + then AOT_obtain p where \\(F = [\y p])\ + using "\E"[rotated] by blast + AOT_hence \F = [\y p]\ + by (metis "derived-S5-rules:2" emptyE "id-nec:2" "\E") + AOT_thus \\p(F = [\y p])\ by (rule "\I") +qed + +AOT_theorem "prop-prop-nec:2": \\p (F \ [\y p]) \ \\p(F \ [\y p])\ +proof(rule "\I") + AOT_assume \\p (F \ [\y p])\ + AOT_hence \(F \ [\y p])\ for p + using "\E" by blast + AOT_hence \\(F \ [\y p])\ for p + by (rule "id-nec2:2"[unvarify \, THEN "\E", rotated]) "cqt:2" + AOT_hence \\p \(F \ [\y p])\ by (rule GEN) + AOT_thus \\\p (F \ [\y p])\ using BF[THEN "\E"] by fast +qed + +AOT_theorem "prop-prop-nec:3": \\p (F = [\y p]) \ \\p(F = [\y p])\ +proof(rule "\I") + AOT_assume \\p (F = [\y p])\ + then AOT_obtain p where \(F = [\y p])\ using "\E"[rotated] by blast + AOT_hence \\(F = [\y p])\ by (metis "id-nec:2" "\E") + AOT_hence \\p\(F = [\y p])\ by (rule "\I") + AOT_thus \\\p(F = [\y p])\ by (metis Buridan "\E") +qed + +AOT_theorem "prop-prop-nec:4": \\\p (F \ [\y p]) \ \p(F \ [\y p])\ +proof(rule "\I") + AOT_assume \\\p (F \ [\y p])\ + AOT_hence \\p \(F \ [\y p])\ by (metis "Buridan\" "\E") + AOT_hence \\(F \ [\y p])\ for p + using "\E" by blast + AOT_hence \F \ [\y p]\ for p + by (rule "id-nec2:3"[unvarify \, THEN "\E", rotated]) "cqt:2" + AOT_thus \\p (F \ [\y p])\ by (rule GEN) +qed + +AOT_theorem "enc-prop-nec:1": + \\\F (x[F] \ \p(F = [\y p])) \ \F(x[F] \ \p (F = [\y p]))\ +proof(rule "\I"; rule GEN; rule "\I") + fix F + AOT_assume \\\F (x[F] \ \p(F = [\y p]))\ + AOT_hence \\F \(x[F] \ \p(F = [\y p]))\ + using "Buridan\" "vdash-properties:10" by blast + AOT_hence 0: \\(x[F] \ \p(F = [\y p]))\ using "\E" by blast + AOT_assume \x[F]\ + AOT_hence \\x[F]\ by (metis "en-eq:2[1]" "\E"(1)) + AOT_hence \\\p(F = [\y p])\ + using 0 by (metis "KBasic2:4" "\E"(1) "vdash-properties:10") + AOT_thus \\p(F = [\y p])\ + using "prop-prop-nec:1"[THEN "\E"] by blast +qed + +AOT_theorem "enc-prop-nec:2": + \\F (x[F] \ \p(F = [\y p])) \ \\F(x[F] \ \p (F = [\y p]))\ + using "derived-S5-rules:1"[where \="{}", simplified, OF "enc-prop-nec:1"] + by blast + +(*<*) +end +(*>*) \ No newline at end of file diff --git a/thys/AOT/AOT_PossibleWorlds.thy b/thys/AOT/AOT_PossibleWorlds.thy new file mode 100644 --- /dev/null +++ b/thys/AOT/AOT_PossibleWorlds.thy @@ -0,0 +1,3195 @@ +(*<*) +theory AOT_PossibleWorlds + imports AOT_PLM AOT_BasicLogicalObjects AOT_RestrictedVariables +begin +(*>*) + +section\Possible Worlds\ + +AOT_define Situation :: \\ \ \\ (\Situation'(_')\) + situations: \Situation(x) \\<^sub>d\<^sub>f A!x & \F (x[F] \ Propositional([F]))\ + +AOT_theorem "T-sit": \TruthValue(x) \ Situation(x)\ +proof(rule "\I") + AOT_assume \TruthValue(x)\ + AOT_hence \\p TruthValueOf(x,p)\ + using "T-value"[THEN "\\<^sub>d\<^sub>fE"] by blast + then AOT_obtain p where \TruthValueOf(x,p)\ using "\E"[rotated] by blast + AOT_hence \: \A!x & \F (x[F] \ \q((q \ p) & F = [\y q]))\ + using "tv-p"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_show \Situation(x)\ + proof(rule situations[THEN "\\<^sub>d\<^sub>fI"]; safe intro!: "&I" GEN "\I" \[THEN "&E"(1)]) + fix F + AOT_assume \x[F]\ + AOT_hence \\q((q \ p) & F = [\y q])\ + using \[THEN "&E"(2), THEN "\E"(2)[where \=F], THEN "\E"(1)] by argo + then AOT_obtain q where \(q \ p) & F = [\y q]\ using "\E"[rotated] by blast + AOT_hence \\p F = [\y p]\ using "&E"(2) "\I"(2) by metis + AOT_thus \Propositional([F])\ + by (metis "\\<^sub>d\<^sub>fI" "prop-prop1") + qed +qed + +AOT_theorem "possit-sit:1": \Situation(x) \ \Situation(x)\ +proof(rule "\I"; rule "\I") + AOT_assume \Situation(x)\ + AOT_hence 0: \A!x & \F (x[F] \ Propositional([F]))\ + using situations[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_have 1: \\(A!x & \F (x[F] \ Propositional([F])))\ + proof(rule "KBasic:3"[THEN "\E"(2)]; rule "&I") + AOT_show \\A!x\ using 0[THEN "&E"(1)] by (metis "oa-facts:2"[THEN "\E"]) + next + AOT_have \\F (x[F] \ Propositional([F])) \ \\F (x[F] \ Propositional([F]))\ + by (AOT_subst \Propositional([F])\ \\p (F = [\y p])\ for: F :: \<\>\) + (auto simp: "prop-prop1" "\Df" "enc-prop-nec:2") + AOT_thus \\\F (x[F] \ Propositional([F]))\ + using 0[THEN "&E"(2)] "\E" by blast + qed + AOT_show \\Situation(x)\ + by (AOT_subst \Situation(x)\ \A!x & \F (x[F] \ Propositional([F]))\) + (auto simp: 1 "\Df" situations) +next + AOT_show \Situation(x)\ if \\Situation(x)\ + using "qml:2"[axiom_inst, THEN "\E", OF that]. +qed + +AOT_theorem "possit-sit:2": \\Situation(x) \ Situation(x)\ + using "possit-sit:1" + by (metis "RE\" "S5Basic:2" "\E"(1) "\E"(5) "Commutativity of \") + +AOT_theorem "possit-sit:3": \\Situation(x) \ \Situation(x)\ + using "possit-sit:1" "possit-sit:2" by (meson "\E"(5)) + +AOT_theorem "possit-sit:4": \\<^bold>\Situation(x) \ Situation(x)\ + by (meson "Act-Basic:5" "Act-Sub:2" "RA[2]" "\E"(1) "\E"(6) "possit-sit:2") + +AOT_theorem "possit-sit:5": \Situation(\p)\ +proof (safe intro!: situations[THEN "\\<^sub>d\<^sub>fI"] "&I" GEN "\I" "prop-prop1"[THEN "\\<^sub>d\<^sub>fI"]) + AOT_have \\F \p[F]\ + using "tv-id:2"[THEN "prop-enc"[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(2)] + "existential:1" "prop-prop2:2" by blast + AOT_thus \A!\p\ + by (safe intro!: "encoders-are-abstract"[unvarify x, THEN "\E"] + "t=t-proper:2"[THEN "\E", OF "ext-p-tv:3"]) +next + fix F + AOT_assume \\p[F]\ + AOT_hence \\<^bold>\x(A!x & \F (x[F] \ \q ((q \ p) & F = [\y q])))[F]\ + using "tv-id:1" "rule=E" by fast + AOT_hence \\<^bold>\\q ((q \ p) & F = [\y q])\ + using "\E"(1) "desc-nec-encode:1" by fast + AOT_hence \\q \<^bold>\((q \ p) & F = [\y q])\ + by (metis "Act-Basic:10" "\E"(1)) + then AOT_obtain q where \\<^bold>\((q \ p) & F = [\y q])\ using "\E"[rotated] by blast + AOT_hence \\<^bold>\F = [\y q]\ by (metis "Act-Basic:2" "con-dis-i-e:2:b" "intro-elim:3:a") + AOT_hence \F = [\y q]\ + using "id-act:1"[unvarify \, THEN "\E"(2)] by (metis "prop-prop2:2") + AOT_thus \\p F = [\y p]\ + using "\I" by fast +qed + +AOT_theorem "possit-sit:6": \Situation(\)\ +proof - + AOT_have true_def: \\<^bold>\\<^sub>\ \ = \<^bold>\x (A!x & \F (x[F] \ \p(p & F = [\y p])))\ + by (simp add: "A-descriptions" "rule-id-df:1[zero]" "the-true:1") + AOT_hence true_den: \\<^bold>\\<^sub>\ \\\ + using "t=t-proper:1" "vdash-properties:6" by blast + AOT_have \\<^bold>\TruthValue(\)\ + using "actual-desc:2"[unvarify x, OF true_den, THEN "\E", OF true_def] + using "TV-lem2:1"[unvarify x, OF true_den, THEN "RA[2]", + THEN "act-cond"[THEN "\E"], THEN "\E"] + by blast + AOT_hence \\<^bold>\Situation(\)\ + using "T-sit"[unvarify x, OF true_den, THEN "RA[2]", + THEN "act-cond"[THEN "\E"], THEN "\E"] by blast + AOT_thus \Situation(\)\ + using "possit-sit:4"[unvarify x, OF true_den, THEN "\E"(1)] by blast +qed + +AOT_theorem "possit-sit:7": \Situation(\)\ +proof - + AOT_have true_def: \\<^bold>\\<^sub>\ \ = \<^bold>\x (A!x & \F (x[F] \ \p(\p & F = [\y p])))\ + by (simp add: "A-descriptions" "rule-id-df:1[zero]" "the-true:2") + AOT_hence true_den: \\<^bold>\\<^sub>\ \\\ + using "t=t-proper:1" "vdash-properties:6" by blast + AOT_have \\<^bold>\TruthValue(\)\ + using "actual-desc:2"[unvarify x, OF true_den, THEN "\E", OF true_def] + using "TV-lem2:2"[unvarify x, OF true_den, THEN "RA[2]", + THEN "act-cond"[THEN "\E"], THEN "\E"] + by blast + AOT_hence \\<^bold>\Situation(\)\ + using "T-sit"[unvarify x, OF true_den, THEN "RA[2]", + THEN "act-cond"[THEN "\E"], THEN "\E"] by blast + AOT_thus \Situation(\)\ + using "possit-sit:4"[unvarify x, OF true_den, THEN "\E"(1)] by blast +qed + +AOT_register_rigid_restricted_type + Situation: \Situation(\)\ +proof + AOT_modally_strict { + fix p + AOT_obtain x where \TruthValueOf(x,p)\ + by (metis "instantiation" "p-has-!tv:1") + AOT_hence \\p TruthValueOf(x,p)\ by (rule "\I") + AOT_hence \TruthValue(x)\ by (metis "\\<^sub>d\<^sub>fI" "T-value") + AOT_hence \Situation(x)\ using "T-sit"[THEN "\E"] by blast + AOT_thus \\x Situation(x)\ by (rule "\I") + } +next + AOT_modally_strict { + AOT_show \Situation(\) \ \\\ for \ + proof (rule "\I") + AOT_assume \Situation(\)\ + AOT_hence \A!\\ by (metis "\\<^sub>d\<^sub>fE" "&E"(1) situations) + AOT_thus \\\\ by (metis "russell-axiom[exe,1].\_denotes_asm") + qed + } +next + AOT_modally_strict { + AOT_show \\\(Situation(\) \ \Situation(\))\ + using "possit-sit:1"[THEN "conventions:3"[THEN "\\<^sub>d\<^sub>fE"], + THEN "&E"(1)] GEN by fast + } +qed + +AOT_register_variable_names + Situation: s + +AOT_define TruthInSituation :: \\ \ \ \ \\ ("(_ \/ _)" [100, 40] 100) + "true-in-s": \s \ p \\<^sub>d\<^sub>f s\<^bold>\p\ + +notepad +begin + (* Verify precedence. *) + fix x p q + have \\x \ p \ q\ = \(x \ p) \ q\\ + by simp + have \\x \ p & q\ = \(x \ p) & q\\ + by simp + have \\x \ \p\ = \x \ (\p)\\ + by simp + have \\x \ \p\ = \x \ (\p)\\ + by simp + have \\x \ \<^bold>\p\ = \x \ (\<^bold>\p)\\ + by simp + have \\\x \ p\ = \\(x \ p)\\ + by simp + have \\\x \ p\ = \\(x \ p)\\ + by simp +end + + +AOT_theorem lem1: \Situation(x) \ (x \ p \ x[\y p])\ +proof (rule "\I"; rule "\I"; rule "\I") + AOT_assume \Situation(x)\ + AOT_assume \x \ p\ + AOT_hence \x\<^bold>\p\ + using "true-in-s"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + AOT_thus \x[\y p]\ using "prop-enc"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast +next + AOT_assume 1: \Situation(x)\ + AOT_assume \x[\y p]\ + AOT_hence \x\<^bold>\p\ + using "prop-enc"[THEN "\\<^sub>d\<^sub>fI", OF "&I", OF "cqt:2"(1)] by blast + AOT_thus \x \ p\ + using "true-in-s"[THEN "\\<^sub>d\<^sub>fI"] 1 "&I" by blast +qed + +AOT_theorem "lem2:1": \s \ p \ \s \ p\ +proof - + AOT_have sit: \Situation(s)\ + by (simp add: Situation.\) + AOT_have \s \ p \ s[\y p]\ + using lem1[THEN "\E", OF sit] by blast + also AOT_have \\ \ \s[\y p]\ + by (rule "en-eq:2[1]"[unvarify F]) "cqt:2[lambda]" + also AOT_have \\ \ \s \ p\ + using lem1[THEN RM, THEN "\E", OF "possit-sit:1"[THEN "\E"(1), OF sit]] + by (metis "KBasic:6" "\E"(2) "Commutativity of \" "\E") + finally show ?thesis. +qed + +AOT_theorem "lem2:2": \\s \ p \ s \ p\ +proof - + AOT_have \\(s \ p \ \s \ p)\ + using "possit-sit:1"[THEN "\E"(1), OF Situation.\] + "lem2:1"[THEN "conventions:3"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(1)]] + RM[OF "\I", THEN "\E"] by blast + thus ?thesis by (metis "B\" "S5Basic:13" "T\" "\I" "\E"(1) "\E") +qed + +AOT_theorem "lem2:3": \\s \ p \ \s \ p\ + using "lem2:1" "lem2:2" by (metis "\E"(5)) + +AOT_theorem "lem2:4": \\<^bold>\(s \ p) \ s \ p\ +proof - + AOT_have \\(s \ p \ \s \ p)\ + using "possit-sit:1"[THEN "\E"(1), OF Situation.\] + "lem2:1"[THEN "conventions:3"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(1)]] + RM[OF "\I", THEN "\E"] by blast + thus ?thesis + using "sc-eq-fur:2"[THEN "\E"] by blast +qed + +AOT_theorem "lem2:5": \\s \ p \ \\s \ p\ + by (metis "KBasic2:1" "contraposition:1[2]" "\I" "\I" "\E"(3) "\E"(4) "lem2:2") + +AOT_theorem "sit-identity": \s = s' \ \p(s \ p \ s' \ p)\ +proof(rule "\I"; rule "\I") + AOT_assume \s = s'\ + moreover AOT_have \\p(s \ p \ s \ p)\ + by (simp add: "oth-class-taut:3:a" "universal-cor") + ultimately AOT_show \\p(s \ p \ s' \ p)\ + using "rule=E" by fast +next + AOT_assume a: \\p (s \ p \ s' \ p)\ + AOT_show \s = s'\ + proof(safe intro!: "ab-obey:1"[THEN "\E", THEN "\E"] "&I" GEN "\I" "\I") + AOT_show \A!s\ using Situation.\ "\\<^sub>d\<^sub>fE" "&E"(1) situations by blast + next + AOT_show \A!s'\ using Situation.\ "\\<^sub>d\<^sub>fE" "&E"(1) situations by blast + next + fix F + AOT_assume 0: \s[F]\ + AOT_hence \\p (F = [\y p])\ + using Situation.\[THEN situations[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(2), + THEN "\E"(2)[where \=F], THEN "\E"] + "prop-prop1"[THEN "\\<^sub>d\<^sub>fE"] by blast + then AOT_obtain p where F_def: \F = [\y p]\ + using "\E" by metis + AOT_hence \s[\y p]\ + using 0 "rule=E" by blast + AOT_hence \s \ p\ + using lem1[THEN "\E", OF Situation.\, THEN "\E"(2)] by blast + AOT_hence \s' \ p\ + using a[THEN "\E"(2)[where \=p], THEN "\E"(1)] by blast + AOT_hence \s'[\y p]\ + using lem1[THEN "\E", OF Situation.\, THEN "\E"(1)] by blast + AOT_thus \s'[F]\ + using F_def[symmetric] "rule=E" by blast + next + fix F + AOT_assume 0: \s'[F]\ + AOT_hence \\p (F = [\y p])\ + using Situation.\[THEN situations[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(2), + THEN "\E"(2)[where \=F], THEN "\E"] + "prop-prop1"[THEN "\\<^sub>d\<^sub>fE"] by blast + then AOT_obtain p where F_def: \F = [\y p]\ + using "\E" by metis + AOT_hence \s'[\y p]\ + using 0 "rule=E" by blast + AOT_hence \s' \ p\ + using lem1[THEN "\E", OF Situation.\, THEN "\E"(2)] by blast + AOT_hence \s \ p\ + using a[THEN "\E"(2)[where \=p], THEN "\E"(2)] by blast + AOT_hence \s[\y p]\ + using lem1[THEN "\E", OF Situation.\, THEN "\E"(1)] by blast + AOT_thus \s[F]\ + using F_def[symmetric] "rule=E" by blast + qed +qed + +AOT_define PartOfSituation :: \\ \ \ \ \\ (infixl \\\ 80) + "sit-part-whole": \s \ s' \\<^sub>d\<^sub>f \p (s \ p \ s' \ p)\ + +AOT_theorem "part:1": \s \ s\ + by (rule "sit-part-whole"[THEN "\\<^sub>d\<^sub>fI"]) + (safe intro!: "&I" Situation.\ GEN "\I") + +AOT_theorem "part:2": \s \ s' & s \ s' \ \(s' \ s)\ +proof(rule "\I"; frule "&E"(1); drule "&E"(2); rule "raa-cor:2") + AOT_assume 0: \s \ s'\ + AOT_hence a: \s \ p \ s' \ p\ for p + using "\E"(2) "sit-part-whole"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + AOT_assume \s' \ s\ + AOT_hence b: \s' \ p \ s \ p\ for p + using "\E"(2) "sit-part-whole"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + AOT_have \\p (s \ p \ s' \ p)\ + using a b by (simp add: "\I" "universal-cor") + AOT_hence 1: \s = s'\ + using "sit-identity"[THEN "\E"(2)] by metis + AOT_assume \s \ s'\ + AOT_hence \\(s = s')\ + by (metis "\\<^sub>d\<^sub>fE" "=-infix") + AOT_thus \s = s' & \(s = s')\ + using 1 "&I" by blast +qed + +AOT_theorem "part:3": \s \ s' & s' \ s'' \ s \ s''\ +proof(rule "\I"; frule "&E"(1); drule "&E"(2); + safe intro!: "&I" GEN "\I" "sit-part-whole"[THEN "\\<^sub>d\<^sub>fI"] Situation.\) + fix p + AOT_assume \s \ p\ + moreover AOT_assume \s \ s'\ + ultimately AOT_have \s' \ p\ + using "sit-part-whole"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2), + THEN "\E"(2)[where \=p], THEN "\E"] by blast + moreover AOT_assume \s' \ s''\ + ultimately AOT_show \s'' \ p\ + using "sit-part-whole"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2), + THEN "\E"(2)[where \=p], THEN "\E"] by blast +qed + +AOT_theorem "sit-identity2:1": \s = s' \ s \ s' & s' \ s\ +proof (safe intro!: "\I" "&I" "\I") + AOT_show \s \ s'\ if \s = s'\ + using "rule=E" "part:1" that by blast +next + AOT_show \s' \ s\ if \s = s'\ + using "rule=E" "part:1" that[symmetric] by blast +next + AOT_assume \s \ s' & s' \ s\ + AOT_thus \s = s'\ using "part:2"[THEN "\E", OF "&I"] + by (metis "\\<^sub>d\<^sub>fI" "&E"(1) "&E"(2) "=-infix" "raa-cor:3") +qed + +AOT_theorem "sit-identity2:2": \s = s' \ \s'' (s'' \ s \ s'' \ s')\ +proof(safe intro!: "\I" "\I" Situation.GEN "sit-identity"[THEN "\E"(2)] + GEN[where 'a=\]) + AOT_show \s'' \ s'\ if \s'' \ s\ and \s = s'\ for s'' + using "rule=E" that by blast +next + AOT_show \s'' \ s\ if \s'' \ s'\ and \s = s'\ for s'' + using "rule=E" id_sym that by blast +next + AOT_show \s' \ p\ if \s \ p\ and \\s'' (s'' \ s \ s'' \ s')\ for p + using "sit-part-whole"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2), + OF that(2)[THEN "Situation.\E", THEN "\E"(1), OF "part:1"], + THEN "\E"(2), THEN "\E", OF that(1)]. +next + AOT_show \s \ p\ if \s' \ p\ and \\s'' (s'' \ s \ s'' \ s')\ for p + using "sit-part-whole"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2), + OF that(2)[THEN "Situation.\E", THEN "\E"(2), OF "part:1"], + THEN "\E"(2), THEN "\E", OF that(1)]. +qed + +AOT_define Persistent :: \\ \ \\ (\Persistent'(_')\) + persistent: \Persistent(p) \\<^sub>d\<^sub>f \s (s \ p \ \s' (s \ s' \ s' \ p))\ + +AOT_theorem "pers-prop": \\p Persistent(p)\ + by (safe intro!: GEN[where 'a=\] Situation.GEN persistent[THEN "\\<^sub>d\<^sub>fI"] "\I") + (simp add: "sit-part-whole"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2), THEN "\E"(2), THEN "\E"]) + +AOT_define NullSituation :: \\ \ \\ (\NullSituation'(_')\) + "df-null-trivial:1": \NullSituation(s) \\<^sub>d\<^sub>f \\p s \ p\ + +AOT_define TrivialSituation :: \\ \ \\ (\TrivialSituation'(_')\) + "df-null-trivial:2": \TrivialSituation(s) \\<^sub>d\<^sub>f \p s \ p\ + +AOT_theorem "thm-null-trivial:1": \\!x NullSituation(x)\ +proof (AOT_subst \NullSituation(x)\ \A!x & \F (x[F] \ F \ F)\ for: x) + AOT_modally_strict { + AOT_show \NullSituation(x) \ A!x & \F (x[F] \ F \ F)\ for x + proof (safe intro!: "\I" "\I" "df-null-trivial:1"[THEN "\\<^sub>d\<^sub>fI"] + dest!: "df-null-trivial:1"[THEN "\\<^sub>d\<^sub>fE"]) + AOT_assume 0: \Situation(x) & \\p x \ p\ + AOT_have 1: \A!x\ + using 0[THEN "&E"(1), THEN situations[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(1)]. + AOT_have 2: \x[F] \ \p F = [\y p]\ for F + using 0[THEN "&E"(1), THEN situations[THEN "\\<^sub>d\<^sub>fE"], + THEN "&E"(2), THEN "\E"(2)] + by (metis "\\<^sub>d\<^sub>fE" "\I" "prop-prop1" "\E") + AOT_show \A!x & \F (x[F] \ F \ F)\ + proof (safe intro!: "&I" 1 GEN "\I" "\I") + fix F + AOT_assume \x[F]\ + moreover AOT_obtain p where \F = [\y p]\ + using calculation 2[THEN "\E"] "\E"[rotated] by blast + ultimately AOT_have \x[\y p]\ + by (metis "rule=E") + AOT_hence \x \ p\ + using lem1[THEN "\E", OF 0[THEN "&E"(1)], THEN "\E"(2)] by blast + AOT_hence \\p (x \ p)\ + by (rule "\I") + AOT_thus \F \ F\ + using 0[THEN "&E"(2)] "raa-cor:1" "&I" by blast + next + fix F :: \<\> AOT_var\ + AOT_assume \F \ F\ + AOT_hence \\(F = F)\ by (metis "\\<^sub>d\<^sub>fE" "=-infix") + moreover AOT_have \F = F\ + by (simp add: "id-eq:1") + ultimately AOT_show \x[F]\ using "&I" "raa-cor:1" by blast + qed + next + AOT_assume 0: \A!x & \F (x[F] \ F \ F)\ + AOT_hence \x[F] \ F \ F\ for F + using "\E" "&E" by blast + AOT_hence 1: \\x[F]\ for F + using "\\<^sub>d\<^sub>fE" "id-eq:1" "=-infix" "reductio-aa:1" "\E"(1) by blast + AOT_show \Situation(x) & \\p x \ p\ + proof (safe intro!: "&I" situations[THEN "\\<^sub>d\<^sub>fI"] 0[THEN "&E"(1)] GEN "\I") + AOT_show \Propositional([F])\ if \x[F]\ for F + using that 1 "&I" "raa-cor:1" by fast + next + AOT_show \\\p x \ p\ + proof(rule "raa-cor:2") + AOT_assume \\p x \ p\ + then AOT_obtain p where \x \ p\ using "\E"[rotated] by blast + AOT_hence \x[\y p]\ + using "\\<^sub>d\<^sub>fE" "&E"(1) "\E"(1) lem1 "modus-tollens:1" + "raa-cor:3" "true-in-s" by fast + moreover AOT_have \\x[\y p]\ + by (rule 1[unvarify F]) "cqt:2[lambda]" + ultimately AOT_show \p & \p\ for p using "&I" "raa-cor:1" by blast + qed + qed + qed + } +next + AOT_show \\!x ([A!]x & \F (x[F] \ F \ F))\ + by (simp add: "A-objects!") +qed + + +AOT_theorem "thm-null-trivial:2": \\!x TrivialSituation(x)\ +proof (AOT_subst \TrivialSituation(x)\ \A!x & \F (x[F] \ \p F = [\y p])\ for: x) + AOT_modally_strict { + AOT_show \TrivialSituation(x) \ A!x & \F (x[F] \ \p F = [\y p])\ for x + proof (safe intro!: "\I" "\I" "df-null-trivial:2"[THEN "\\<^sub>d\<^sub>fI"] + dest!: "df-null-trivial:2"[THEN "\\<^sub>d\<^sub>fE"]) + AOT_assume 0: \Situation(x) & \p x \ p\ + AOT_have 1: \A!x\ + using 0[THEN "&E"(1), THEN situations[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(1)]. + AOT_have 2: \x[F] \ \p F = [\y p]\ for F + using 0[THEN "&E"(1), THEN situations[THEN "\\<^sub>d\<^sub>fE"], + THEN "&E"(2), THEN "\E"(2)] + by (metis "\\<^sub>d\<^sub>fE" "deduction-theorem" "prop-prop1" "\E") + AOT_show \A!x & \F (x[F] \ \p F = [\y p])\ + proof (safe intro!: "&I" 1 GEN "\I" "\I" 2) + fix F + AOT_assume \\p F = [\y p]\ + then AOT_obtain p where \F = [\y p]\ + using "\E"[rotated] by blast + moreover AOT_have \x \ p\ + using 0[THEN "&E"(2)] "\E" by blast + ultimately AOT_show \x[F]\ + by (metis 0 "rule=E" "&E"(1) id_sym "\E"(2) lem1 + "Commutativity of \" "\E") + qed + next + AOT_assume 0: \A!x & \F (x[F] \ \p F = [\y p])\ + AOT_hence 1: \x[F] \ \p F = [\y p]\ for F + using "\E" "&E" by blast + AOT_have 2: \Situation(x)\ + proof (safe intro!: "&I" situations[THEN "\\<^sub>d\<^sub>fI"] 0[THEN "&E"(1)] GEN "\I") + AOT_show \Propositional([F])\ if \x[F]\ for F + using 1[THEN "\E"(1), OF that] + by (metis "\\<^sub>d\<^sub>fI" "prop-prop1") + qed + AOT_show \Situation(x) & \p (x \ p)\ + proof (safe intro!: "&I" 2 0[THEN "&E"(1)] GEN "\I") + AOT_have \x[\y p] \ \q [\y p] = [\y q]\ for p + by (rule 1[unvarify F, where \="\[\y p]\"]) "cqt:2[lambda]" + moreover AOT_have \\q [\y p] = [\y q]\ for p + by (rule "\I"(2)[where \=p]) + (simp add: "rule=I:1" "prop-prop2:2") + ultimately AOT_have \x[\y p]\ for p by (metis "\E"(2)) + AOT_thus \x \ p\ for p + by (metis "2" "\E"(2) lem1 "\E") + qed + qed + } +next + AOT_show \\!x ([A!]x & \F (x[F] \ \p F = [\y p]))\ + by (simp add: "A-objects!") +qed + +AOT_theorem "thm-null-trivial:3": \\<^bold>\x NullSituation(x)\\ + by (meson "A-Exists:2" "RA[2]" "\E"(2) "thm-null-trivial:1") + +AOT_theorem "thm-null-trivial:4": \\<^bold>\x TrivialSituation(x)\\ + using "A-Exists:2" "RA[2]" "\E"(2) "thm-null-trivial:2" by blast + +AOT_define TheNullSituation :: \\\<^sub>s\ (\\<^bold>s\<^sub>\\) + "df-the-null-sit:1": \\<^bold>s\<^sub>\ =\<^sub>d\<^sub>f \<^bold>\x NullSituation(x)\ + +AOT_define TheTrivialSituation :: \\\<^sub>s\ (\\<^bold>s\<^sub>V\) + "df-the-null-sit:2": \\<^bold>s\<^sub>V =\<^sub>d\<^sub>f \<^bold>\x TrivialSituation(x)\ + +AOT_theorem "null-triv-sc:1": \NullSituation(x) \ \NullSituation(x)\ +proof(safe intro!: "\I" dest!: "df-null-trivial:1"[THEN "\\<^sub>d\<^sub>fE"]; + frule "&E"(1); drule "&E"(2)) + AOT_assume 1: \\\p (x \ p)\ + AOT_assume 0: \Situation(x)\ + AOT_hence \\Situation(x)\ by (metis "\E"(1) "possit-sit:1") + moreover AOT_have \\\\p (x \ p)\ + proof(rule "raa-cor:1") + AOT_assume \\\\\p (x \ p)\ + AOT_hence \\\p (x \ p)\ + by (metis "\\<^sub>d\<^sub>fI" "conventions:5") + AOT_hence \\p \(x \ p)\ by (metis "BF\" "\E") + then AOT_obtain p where \\(x \ p)\ using "\E"[rotated] by blast + AOT_hence \x \ p\ + by (metis "\E"(1) "lem2:2"[unconstrain s, THEN "\E", OF 0]) + AOT_hence \\p x \ p\ using "\I" by fast + AOT_thus \\p x \ p & \\p x \ p\ using 1 "&I" by blast + qed + ultimately AOT_have 2: \\(Situation(x) & \\p x \ p)\ + by (metis "KBasic:3" "&I" "\E"(2)) + AOT_show \\NullSituation(x)\ + by (AOT_subst \NullSituation(x)\ \Situation(x) & \\p x \ p\) + (auto simp: "df-null-trivial:1" "\Df" 2) +qed + + +AOT_theorem "null-triv-sc:2": \TrivialSituation(x) \ \TrivialSituation(x)\ +proof(safe intro!: "\I" dest!: "df-null-trivial:2"[THEN "\\<^sub>d\<^sub>fE"]; + frule "&E"(1); drule "&E"(2)) + AOT_assume 0: \Situation(x)\ + AOT_hence 1: \\Situation(x)\ by (metis "\E"(1) "possit-sit:1") + AOT_assume \\p x \ p\ + AOT_hence \x \ p\ for p + using "\E" by blast + AOT_hence \\x \ p\ for p + using 0 "\E"(1) "lem2:1"[unconstrain s, THEN "\E"] by blast + AOT_hence \\p \x \ p\ + by (rule GEN) + AOT_hence \\\p x \ p\ + by (rule BF[THEN "\E"]) + AOT_hence 2: \\(Situation(x) & \p x \ p)\ + using 1 by (metis "KBasic:3" "&I" "\E"(2)) + AOT_show \\TrivialSituation(x)\ + by (AOT_subst \TrivialSituation(x)\ \Situation(x) & \p x \ p\) + (auto simp: "df-null-trivial:2" "\Df" 2) +qed + +AOT_theorem "null-triv-sc:3": \NullSituation(\<^bold>s\<^sub>\)\ + by (safe intro!: "df-the-null-sit:1"[THEN "=\<^sub>d\<^sub>fI"(2)] "thm-null-trivial:3" + "rule=I:1"[OF "thm-null-trivial:3"] + "!box-desc:2"[THEN "\E", THEN "\E", rotated, OF "thm-null-trivial:1", + OF "\I", OF "null-triv-sc:1", THEN "\E"(1), THEN "\E"]) + +AOT_theorem "null-triv-sc:4": \TrivialSituation(\<^bold>s\<^sub>V)\ + by (safe intro!: "df-the-null-sit:2"[THEN "=\<^sub>d\<^sub>fI"(2)] "thm-null-trivial:4" + "rule=I:1"[OF "thm-null-trivial:4"] + "!box-desc:2"[THEN "\E", THEN "\E", rotated, OF "thm-null-trivial:2", + OF "\I", OF "null-triv-sc:2", THEN "\E"(1), THEN "\E"]) + +AOT_theorem "null-triv-facts:1": \NullSituation(x) \ Null(x)\ +proof (safe intro!: "\I" "\I" "df-null-uni:1"[THEN "\\<^sub>d\<^sub>fI"] + "df-null-trivial:1"[THEN "\\<^sub>d\<^sub>fI"] + dest!: "df-null-uni:1"[THEN "\\<^sub>d\<^sub>fE"] "df-null-trivial:1"[THEN "\\<^sub>d\<^sub>fE"]) + AOT_assume 0: \Situation(x) & \\p x \ p\ + AOT_have 1: \x[F] \ \p F = [\y p]\ for F + using 0[THEN "&E"(1), THEN situations[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(2), THEN "\E"(2)] + by (metis "\\<^sub>d\<^sub>fE" "deduction-theorem" "prop-prop1" "\E") + AOT_show \A!x & \\F x[F]\ + proof (safe intro!: "&I" 0[THEN "&E"(1), THEN situations[THEN "\\<^sub>d\<^sub>fE"], + THEN "&E"(1)]; + rule "raa-cor:2") + AOT_assume \\F x[F]\ + then AOT_obtain F where F_prop: \x[F]\ + using "\E"[rotated] by blast + AOT_hence \\p F = [\y p]\ + using 1[THEN "\E"] by blast + then AOT_obtain p where \F = [\y p]\ + using "\E"[rotated] by blast + AOT_hence \x[\y p]\ + by (metis "rule=E" F_prop) + AOT_hence \x \ p\ + using lem1[THEN "\E", OF 0[THEN "&E"(1)], THEN "\E"(2)] by blast + AOT_hence \\p x \ p\ + by (rule "\I") + AOT_thus \\p x \ p & \\p x \ p\ + using 0[THEN "&E"(2)] "&I" by blast + qed +next + AOT_assume 0: \A!x & \\F x[F]\ + AOT_have \Situation(x)\ + apply (rule situations[THEN "\\<^sub>d\<^sub>fI", OF "&I", OF 0[THEN "&E"(1)]]; rule GEN) + using 0[THEN "&E"(2)] by (metis "\I" "existential:2[const_var]" "raa-cor:3") + moreover AOT_have \\\p x \ p\ + proof (rule "raa-cor:2") + AOT_assume \\p x \ p\ + then AOT_obtain p where \x \ p\ by (metis "instantiation") + AOT_hence \x[\y p]\ by (metis "\\<^sub>d\<^sub>fE" "&E"(2) "prop-enc" "true-in-s") + AOT_hence \\F x[F]\ by (rule "\I") "cqt:2[lambda]" + AOT_thus \\F x[F] & \\F x[F]\ using 0[THEN "&E"(2)] "&I" by blast + qed + ultimately AOT_show \Situation(x) & \\p x \ p\ using "&I" by blast +qed + +AOT_theorem "null-triv-facts:2": \\<^bold>s\<^sub>\ = a\<^sub>\\ + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF "df-the-null-sit:1"]) + apply (fact "thm-null-trivial:3") + apply (rule "=\<^sub>d\<^sub>fI"(2)[OF "df-null-uni-terms:1"]) + apply (fact "null-uni-uniq:3") + apply (rule "equiv-desc-eq:3"[THEN "\E"]) + apply (rule "&I") + apply (fact "thm-null-trivial:3") + by (rule RN; rule GEN; rule "null-triv-facts:1") + +AOT_theorem "null-triv-facts:3": \\<^bold>s\<^sub>V \ a\<^sub>V\ +proof(rule "=-infix"[THEN "\\<^sub>d\<^sub>fI"]) + AOT_have \Universal(a\<^sub>V)\ + by (simp add: "null-uni-facts:4") + AOT_hence 0: \a\<^sub>V[A!]\ + using "df-null-uni:2"[THEN "\\<^sub>d\<^sub>fE"] "&E" "\E"(1) + by (metis "cqt:5:a" "vdash-properties:10" "vdash-properties:1[2]") + moreover AOT_have 1: \\\<^bold>s\<^sub>V[A!]\ + proof(rule "raa-cor:2") + AOT_have \Situation(\<^bold>s\<^sub>V)\ + using "\\<^sub>d\<^sub>fE" "&E"(1) "df-null-trivial:2" "null-triv-sc:4" by blast + AOT_hence \\F (\<^bold>s\<^sub>V[F] \ Propositional([F]))\ + by (metis "\\<^sub>d\<^sub>fE" "&E"(2) situations) + moreover AOT_assume \\<^bold>s\<^sub>V[A!]\ + ultimately AOT_have \Propositional(A!)\ + using "\E"(1)[rotated, OF "oa-exist:2"] "\E" by blast + AOT_thus \Propositional(A!) & \Propositional(A!)\ + using "prop-in-f:4:d" "&I" by blast + qed + AOT_show \\(\<^bold>s\<^sub>V = a\<^sub>V)\ + proof (rule "raa-cor:2") + AOT_assume \\<^bold>s\<^sub>V = a\<^sub>V\ + AOT_hence \\<^bold>s\<^sub>V[A!]\ using 0 "rule=E" id_sym by fast + AOT_thus \\<^bold>s\<^sub>V[A!] & \\<^bold>s\<^sub>V[A!]\ using 1 "&I" by blast + qed +qed + +definition ConditionOnPropositionalProperties :: \(<\> \ \) \ bool\ where + "cond-prop": \ConditionOnPropositionalProperties \ \ \ . \ v . + [v \ \F (\{F} \ Propositional([F]))]\ + +syntax ConditionOnPropositionalProperties :: \id_position \ AOT_prop\ + ("CONDITION'_ON'_PROPOSITIONAL'_PROPERTIES'(_')") + +AOT_theorem "cond-prop[E]": + assumes \CONDITION_ON_PROPOSITIONAL_PROPERTIES(\)\ + shows \\F (\{F} \ Propositional([F]))\ + using assms[unfolded "cond-prop"] by auto + +AOT_theorem "cond-prop[I]": + assumes \\<^bold>\\<^sub>\ \F (\{F} \ Propositional([F]))\ + shows \CONDITION_ON_PROPOSITIONAL_PROPERTIES(\)\ + using assms "cond-prop" by metis + +AOT_theorem "pre-comp-sit": + assumes \CONDITION_ON_PROPOSITIONAL_PROPERTIES(\)\ + shows \(Situation(x) & \F (x[F] \ \{F})) \ (A!x & \F (x[F] \ \{F}))\ +proof(rule "\I"; rule "\I") + AOT_assume \Situation(x) & \F (x[F] \ \{F})\ + AOT_thus \A!x & \F (x[F] \ \{F})\ + using "&E" situations[THEN "\\<^sub>d\<^sub>fE"] "&I" by blast +next + AOT_assume 0: \A!x & \F (x[F] \ \{F})\ + AOT_show \Situation(x) & \F (x[F] \ \{F})\ + proof (safe intro!: situations[THEN "\\<^sub>d\<^sub>fI"] "&I") + AOT_show \A!x\ using 0[THEN "&E"(1)]. + next + AOT_show \\F (x[F] \ Propositional([F]))\ + proof(rule GEN; rule "\I") + fix F + AOT_assume \x[F]\ + AOT_hence \\{F}\ + using 0[THEN "&E"(2)] "\E" "\E" by blast + AOT_thus \Propositional([F])\ + using "cond-prop[E]"[OF assms] "\E" "\E" by blast + qed + next + AOT_show \\F (x[F] \ \{F})\ using 0 "&E" by blast + qed +qed + +AOT_theorem "comp-sit:1": + assumes \CONDITION_ON_PROPOSITIONAL_PROPERTIES(\)\ + shows \\s \F(s[F] \ \{F})\ + by (AOT_subst \Situation(x) & \F(x[F] \ \{F})\ \A!x & \F (x[F] \ \{F})\ for: x) + (auto simp: "pre-comp-sit"[OF assms] "A-objects"[where \=\, axiom_inst]) + +AOT_theorem "comp-sit:2": + assumes \CONDITION_ON_PROPOSITIONAL_PROPERTIES(\)\ + shows \\!s \F(s[F] \ \{F})\ + by (AOT_subst \Situation(x) & \F(x[F] \ \{F})\ \A!x & \F (x[F] \ \{F})\ for: x) + (auto simp: assms "pre-comp-sit" "pre-comp-sit"[OF assms] "A-objects!") + +AOT_theorem "can-sit-desc:1": + assumes \CONDITION_ON_PROPOSITIONAL_PROPERTIES(\)\ + shows \\<^bold>\s(\F (s[F] \ \{F}))\\ + using "comp-sit:2"[OF assms] "A-Exists:2" "RA[2]" "\E"(2) by blast + +AOT_theorem "can-sit-desc:2": + assumes \CONDITION_ON_PROPOSITIONAL_PROPERTIES(\)\ + shows \\<^bold>\s(\F (s[F] \ \{F})) = \<^bold>\x(A!x & \F (x[F] \ \{F}))\ + by (auto intro!: "equiv-desc-eq:2"[THEN "\E", OF "&I", + OF "can-sit-desc:1"[OF assms]] + "RA[2]" GEN "pre-comp-sit"[OF assms]) + +AOT_theorem "strict-sit": + assumes \RIGID_CONDITION(\)\ + and \CONDITION_ON_PROPOSITIONAL_PROPERTIES(\)\ + shows \y = \<^bold>\s(\F (s[F] \ \{F})) \ \F (y[F] \ \{F})\ + using "rule=E"[rotated, OF "can-sit-desc:2"[OF assms(2), symmetric]] + "box-phi-a:2"[OF assms(1)] "\E" "\I" "&E" by fast + +(* TODO: exercise (479) sit-lit *) + +AOT_define actual :: \\ \ \\ (\Actual'(_')\) + \Actual(s) \\<^sub>d\<^sub>f \p (s \ p \ p)\ + +AOT_theorem "act-and-not-pos": \\s (Actual(s) & \\Actual(s))\ +proof - + AOT_obtain q\<^sub>1 where q\<^sub>1_prop: \q\<^sub>1 & \\q\<^sub>1\ + by (metis "\\<^sub>d\<^sub>fE" "instantiation" "cont-tf:1" "cont-tf-thm:1") + AOT_have \\s (\F (s[F] \ F = [\y q\<^sub>1]))\ + proof (safe intro!: "comp-sit:1" "cond-prop[I]" GEN "\I") + AOT_modally_strict { + AOT_show \Propositional([F])\ if \F = [\y q\<^sub>1]\ for F + using "\\<^sub>d\<^sub>fI" "existential:2[const_var]" "prop-prop1" that by fastforce + } + qed + then AOT_obtain s\<^sub>1 where s_prop: \\F (s\<^sub>1[F] \ F = [\y q\<^sub>1])\ + using "Situation.\E"[rotated] by meson + AOT_have \Actual(s\<^sub>1)\ + proof(safe intro!: actual[THEN "\\<^sub>d\<^sub>fI"] "&I" GEN "\I" s_prop Situation.\) + fix p + AOT_assume \s\<^sub>1 \ p\ + AOT_hence \s\<^sub>1[\y p]\ + by (metis "\\<^sub>d\<^sub>fE" "&E"(2) "prop-enc" "true-in-s") + AOT_hence \[\y p] = [\y q\<^sub>1]\ + by (rule s_prop[THEN "\E"(1), THEN "\E"(1), rotated]) "cqt:2[lambda]" + AOT_hence \p = q\<^sub>1\ by (metis "\E"(2) "p-identity-thm2:3") + AOT_thus \p\ using q\<^sub>1_prop[THEN "&E"(1)] "rule=E" id_sym by fast + qed + moreover AOT_have \\\Actual(s\<^sub>1)\ + proof(rule "raa-cor:1"; drule "KBasic:12"[THEN "\E"(2)]) + AOT_assume \\Actual(s\<^sub>1)\ + AOT_hence \\(Situation(s\<^sub>1) & \p (s\<^sub>1 \ p \ p))\ + using actual[THEN "\Df", THEN "conventions:3"[THEN "\\<^sub>d\<^sub>fE"], + THEN "&E"(1), THEN RM, THEN "\E"] by blast + AOT_hence \\\p (s\<^sub>1 \ p \ p)\ + by (metis "RM:1" "Conjunction Simplification"(2) "\E") + AOT_hence \\p \(s\<^sub>1 \ p \ p)\ + by (metis "CBF" "vdash-properties:10") + AOT_hence \\(s\<^sub>1 \ q\<^sub>1 \ q\<^sub>1)\ + using "\E" by blast + AOT_hence \\s\<^sub>1 \ q\<^sub>1 \ \q\<^sub>1\ + by (metis "\E" "qml:1" "vdash-properties:1[2]") + moreover AOT_have \s\<^sub>1 \ q\<^sub>1\ + using s_prop[THEN "\E"(1), THEN "\E"(2), + THEN lem1[THEN "\E", OF Situation.\, THEN "\E"(2)]] + "rule=I:1" "prop-prop2:2" by blast + ultimately AOT_have \\q\<^sub>1\ + using "\\<^sub>d\<^sub>fE" "&E"(1) "\E"(1) "lem2:1" "true-in-s" "\E" by fast + AOT_thus \\\q\<^sub>1 & \\\q\<^sub>1\ + using "KBasic:12"[THEN "\E"(1)] q\<^sub>1_prop[THEN "&E"(2)] "&I" by blast + qed + ultimately AOT_have \(Actual(s\<^sub>1) & \\Actual(s\<^sub>1))\ + using s_prop "&I" by blast + thus ?thesis + by (rule "Situation.\I") +qed + +AOT_theorem "actual-s:1": \\s Actual(s)\ +proof - + AOT_obtain s where \(Actual(s) & \\Actual(s))\ + using "act-and-not-pos" "Situation.\E"[rotated] by meson + AOT_hence \Actual(s)\ using "&E" "&I" by metis + thus ?thesis by (rule "Situation.\I") +qed + +AOT_theorem "actual-s:2": \\s \Actual(s)\ +proof(rule "\I"(1)[where \=\\\<^bold>s\<^sub>V\\]; (rule "&I")?) + AOT_show \Situation(\<^bold>s\<^sub>V)\ + using "\\<^sub>d\<^sub>fE" "&E"(1) "df-null-trivial:2" "null-triv-sc:4" by blast +next + AOT_show \\Actual(\<^bold>s\<^sub>V)\ + proof(rule "raa-cor:2") + AOT_assume 0: \Actual(\<^bold>s\<^sub>V)\ + AOT_obtain p\<^sub>1 where notp\<^sub>1: \\p\<^sub>1\ + by (metis "\E" "\I"(1) "log-prop-prop:2" "non-contradiction") + AOT_have \\<^bold>s\<^sub>V \ p\<^sub>1\ + using "null-triv-sc:4"[THEN "\\<^sub>d\<^sub>fE"[OF "df-null-trivial:2"], THEN "&E"(2)] + "\E" by blast + AOT_hence \p\<^sub>1\ + using 0[THEN actual[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(2), THEN "\E"(2), THEN "\E"] + by blast + AOT_thus \p & \p\ for p using notp\<^sub>1 by (metis "raa-cor:3") + qed +next + AOT_show \\<^bold>s\<^sub>V\\ + using "df-the-null-sit:2" "rule-id-df:2:b[zero]" "thm-null-trivial:4" by blast +qed + +AOT_theorem "actual-s:3": \\p\s(Actual(s) \ \s \ p)\ +proof - + AOT_obtain p\<^sub>1 where notp\<^sub>1: \\p\<^sub>1\ + by (metis "\E" "\I"(1) "log-prop-prop:2" "non-contradiction") + AOT_have \\s (Actual(s) \ \(s \ p\<^sub>1))\ + proof (rule Situation.GEN; rule "\I"; rule "raa-cor:2") + fix s + AOT_assume \Actual(s)\ + moreover AOT_assume \s \ p\<^sub>1\ + ultimately AOT_have \p\<^sub>1\ + using actual[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2), THEN "\E"(2), THEN "\E"] by blast + AOT_thus \p\<^sub>1 & \p\<^sub>1\ + using notp\<^sub>1 "&I" by simp + qed + thus ?thesis by (rule "\I") +qed + +AOT_theorem comp: + \\s (s' \ s & s'' \ s & \s''' (s' \ s''' & s'' \ s''' \ s \ s'''))\ +proof - + have cond_prop: \ConditionOnPropositionalProperties (\ \ . \s'[\] \ s''[\]\)\ + proof(safe intro!: "cond-prop[I]" GEN "oth-class-taut:8:c"[THEN "\E", THEN "\E"]; + rule "\I") + AOT_modally_strict { + fix F + AOT_have \Situation(s')\ + by (simp add: Situation.restricted_var_condition) + AOT_hence \s'[F] \ Propositional([F])\ + using "situations"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2), THEN "\E"(2)] by blast + moreover AOT_assume \s'[F]\ + ultimately AOT_show \Propositional([F])\ + using "\E" by blast + } + next + AOT_modally_strict { + fix F + AOT_have \Situation(s'')\ + by (simp add: Situation.restricted_var_condition) + AOT_hence \s''[F] \ Propositional([F])\ + using "situations"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2), THEN "\E"(2)] by blast + moreover AOT_assume \s''[F]\ + ultimately AOT_show \Propositional([F])\ + using "\E" by blast + } + qed + AOT_obtain s\<^sub>3 where \: \\F (s\<^sub>3[F] \ s'[F] \ s''[F])\ + using "comp-sit:1"[OF cond_prop] "Situation.\E"[rotated] by meson + AOT_have \s' \ s\<^sub>3 & s'' \ s\<^sub>3 & \s''' (s' \ s''' & s'' \ s''' \ s\<^sub>3 \ s''')\ + proof(safe intro!: "&I" "\\<^sub>d\<^sub>fI"[OF "true-in-s"] "\\<^sub>d\<^sub>fI"[OF "prop-enc"] + "Situation.GEN" "GEN"[where 'a=\] "\I" + "sit-part-whole"[THEN "\\<^sub>d\<^sub>fI"] + Situation.\ "cqt:2[const_var]"[axiom_inst]) + fix p + AOT_assume \s' \ p\ + AOT_hence \s'[\x p]\ + by (metis "&E"(2) "prop-enc" "\\<^sub>d\<^sub>fE" "true-in-s") + AOT_thus \s\<^sub>3[\x p]\ + using \[THEN "\E"(1),OF "prop-prop2:2", THEN "\E"(2), OF "\I"(1)] by blast + next + fix p + AOT_assume \s'' \ p\ + AOT_hence \s''[\x p]\ + by (metis "&E"(2) "prop-enc" "\\<^sub>d\<^sub>fE" "true-in-s") + AOT_thus \s\<^sub>3[\x p]\ + using \[THEN "\E"(1),OF "prop-prop2:2", THEN "\E"(2), OF "\I"(2)] by blast + next + fix s p + AOT_assume 0: \s' \ s & s'' \ s\ + AOT_assume \s\<^sub>3 \ p\ + AOT_hence \s\<^sub>3[\x p]\ + by (metis "&E"(2) "prop-enc" "\\<^sub>d\<^sub>fE" "true-in-s") + AOT_hence \s'[\x p] \ s''[\x p]\ + using \[THEN "\E"(1),OF "prop-prop2:2", THEN "\E"(1)] by blast + moreover { + AOT_assume \s'[\x p]\ + AOT_hence \s' \ p\ + by (safe intro!: "prop-enc"[THEN "\\<^sub>d\<^sub>fI"] "true-in-s"[THEN "\\<^sub>d\<^sub>fI"] "&I" + Situation.\ "cqt:2[const_var]"[axiom_inst]) + moreover AOT_have \s' \ p \ s \ p\ + using "sit-part-whole"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2)] 0[THEN "&E"(1)] + "\E"(2) by blast + ultimately AOT_have \s \ p\ + using "\E" by blast + AOT_hence \s[\x p]\ + using "true-in-s"[THEN "\\<^sub>d\<^sub>fE"] "prop-enc"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + } + moreover { + AOT_assume \s''[\x p]\ + AOT_hence \s'' \ p\ + by (safe intro!: "prop-enc"[THEN "\\<^sub>d\<^sub>fI"] "true-in-s"[THEN "\\<^sub>d\<^sub>fI"] "&I" + Situation.\ "cqt:2[const_var]"[axiom_inst]) + moreover AOT_have \s'' \ p \ s \ p\ + using "sit-part-whole"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2)] 0[THEN "&E"(2)] + "\E"(2) by blast + ultimately AOT_have \s \ p\ + using "\E" by blast + AOT_hence \s[\x p]\ + using "true-in-s"[THEN "\\<^sub>d\<^sub>fE"] "prop-enc"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + } + ultimately AOT_show \s[\x p]\ + by (metis "\E"(1) "\I") + qed + thus ?thesis + using "Situation.\I" by fast +qed + +AOT_theorem "act-sit:1": \Actual(s) \ (s \ p \ [\y p]s)\ +proof (safe intro!: "\I") + AOT_assume \Actual(s)\ + AOT_hence p if \s \ p\ + using actual[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2), THEN "\E"(2), THEN "\E"] that by blast + moreover AOT_assume \s \ p\ + ultimately AOT_have p by blast + AOT_thus \[\y p]s\ + by (safe intro!: "\\C"(1) "cqt:2") +qed + +AOT_theorem "act-sit:2": + \(Actual(s') & Actual(s'')) \ \x (Actual(x) & s' \ x & s'' \ x)\ +proof(rule "\I"; frule "&E"(1); drule "&E"(2)) + AOT_assume act_s': \Actual(s')\ + AOT_assume act_s'': \Actual(s'')\ + have "cond-prop": \ConditionOnPropositionalProperties + (\ \ . \\p (\ = [\y p] & (s' \ p \ s'' \ p))\)\ + proof (safe intro!: "cond-prop[I]" "\I" "\I" "prop-prop1"[THEN "\\<^sub>d\<^sub>fI"]) + AOT_modally_strict { + fix \ + AOT_assume \\p (\ = [\y p] & (s' \ p \ s'' \ p))\ + then AOT_obtain p where \\ = [\y p]\ using "\E"[rotated] "&E" by blast + AOT_thus \\p \ = [\y p]\ by (rule "\I") + } + qed + have rigid: \rigid_condition (\ \ . \\p (\ = [\y p] & (s' \ p \ s'' \ p))\)\ + proof(safe intro!: "strict-can:1[I]" "\I" GEN) + AOT_modally_strict { + fix F + AOT_assume \\p (F = [\y p] & (s' \ p \ s'' \ p))\ + then AOT_obtain p\<^sub>1 where p\<^sub>1_prop: \F = [\y p\<^sub>1] & (s' \ p\<^sub>1 \ s'' \ p\<^sub>1)\ + using "\E"[rotated] by blast + AOT_hence \\(F = [\y p\<^sub>1])\ + using "&E"(1) "id-nec:2" "vdash-properties:10" by blast + moreover AOT_have \\(s' \ p\<^sub>1 \ s'' \ p\<^sub>1)\ + proof(rule "\E"; (rule "\I"; rule "KBasic:15"[THEN "\E"])?) + AOT_show \s' \ p\<^sub>1 \ s'' \ p\<^sub>1\ using p\<^sub>1_prop "&E" by blast + next + AOT_show \\s' \ p\<^sub>1 \ \s'' \ p\<^sub>1\ if \s' \ p\<^sub>1\ + apply (rule "\I"(1)) + using "\\<^sub>d\<^sub>fE" "&E"(1) "\E"(1) "lem2:1" that "true-in-s" by blast + next + AOT_show \\s' \ p\<^sub>1 \ \s'' \ p\<^sub>1\ if \s'' \ p\<^sub>1\ + apply (rule "\I"(2)) + using "\\<^sub>d\<^sub>fE" "&E"(1) "\E"(1) "lem2:1" that "true-in-s" by blast + qed + ultimately AOT_have \\(F = [\y p\<^sub>1] & (s' \ p\<^sub>1 \ s'' \ p\<^sub>1))\ + by (metis "KBasic:3" "&I" "\E"(2)) + AOT_hence \\p \(F = [\y p] & (s' \ p \ s'' \ p))\ by (rule "\I") + AOT_thus \\\p (F = [\y p] & (s' \ p \ s'' \ p))\ + using Buridan[THEN "\E"] by fast + } + qed + + AOT_have desc_den: \\<^bold>\s(\F (s[F] \ \p (F = [\y p] & (s' \ p \ s'' \ p))))\\ + by (rule "can-sit-desc:1"[OF "cond-prop"]) + AOT_obtain x\<^sub>0 + where x\<^sub>0_prop1: \x\<^sub>0 = \<^bold>\s(\F (s[F] \ \p (F = [\y p] & (s' \ p \ s'' \ p))))\ + by (metis (no_types, lifting) "\E" "rule=I:1" desc_den "\I"(1) id_sym) + AOT_hence x\<^sub>0_sit: \Situation(x\<^sub>0)\ + using "actual-desc:3"[THEN "\E"] "Act-Basic:2" "&E"(1) "\E"(1) + "possit-sit:4" by blast + + AOT_have 1: \\F (x\<^sub>0[F] \ \p (F = [\y p] & (s' \ p \ s'' \ p)))\ + using "strict-sit"[OF rigid, OF "cond-prop", THEN "\E", OF x\<^sub>0_prop1]. + AOT_have 2: \(x\<^sub>0 \ p) \ (s' \ p \ s'' \ p)\ for p + proof (rule "\I"; rule "\I") + AOT_assume \x\<^sub>0 \ p\ + AOT_hence \x\<^sub>0[\y p]\ using lem1[THEN "\E", OF x\<^sub>0_sit, THEN "\E"(1)] by blast + then AOT_obtain q where \[\y p] = [\y q] & (s' \ q \ s'' \ q)\ + using 1[THEN "\E"(1)[where \="\[\y p]\"], OF "prop-prop2:2", THEN "\E"(1)] + "\E"[rotated] by blast + AOT_thus \s' \ p \ s'' \ p\ + by (metis "rule=E" "&E"(1) "&E"(2) "\I"(1) "\I"(2) + "\E"(1) "deduction-theorem" id_sym "\E"(2) "p-identity-thm2:3") + next + AOT_assume \s' \ p \ s'' \ p\ + AOT_hence \[\y p] = [\y p] & (s' \ p \ s'' \ p)\ + by (metis "rule=I:1" "&I" "prop-prop2:2") + AOT_hence \\q ([\y p] = [\y q] & (s' \ q \ s'' \ q))\ + by (rule "\I") + AOT_hence \x\<^sub>0[\y p]\ + using 1[THEN "\E"(1), OF "prop-prop2:2", THEN "\E"(2)] by blast + AOT_thus \x\<^sub>0 \ p\ + by (metis "\\<^sub>d\<^sub>fI" "&I" "ex:1:a" "prop-enc" "rule-ui:2[const_var]" + x\<^sub>0_sit "true-in-s") + qed + + AOT_have \Actual(x\<^sub>0) & s' \ x\<^sub>0 & s'' \ x\<^sub>0\ + proof(safe intro!: "\I" "&I" "\I"(1) actual[THEN "\\<^sub>d\<^sub>fI"] x\<^sub>0_sit GEN + "sit-part-whole"[THEN "\\<^sub>d\<^sub>fI"]) + fix p + AOT_assume \x\<^sub>0 \ p\ + AOT_hence \s' \ p \ s'' \ p\ + using 2 "\E"(1) by metis + AOT_thus \p\ + using act_s' act_s'' + actual[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2), THEN "\E"(2), THEN "\E"] + by (metis "\E"(3) "reductio-aa:1") + next + AOT_show \x\<^sub>0 \ p\ if \s' \ p\ for p + using 2[THEN "\E"(2), OF "\I"(1), OF that]. + next + AOT_show \x\<^sub>0 \ p\ if \s'' \ p\ for p + using 2[THEN "\E"(2), OF "\I"(2), OF that]. + next + AOT_show \Situation(s')\ + using act_s'[THEN actual[THEN "\\<^sub>d\<^sub>fE"]] "&E" by blast + next + AOT_show \Situation(s'')\ + using act_s''[THEN actual[THEN "\\<^sub>d\<^sub>fE"]] "&E" by blast + qed + AOT_thus \\x (Actual(x) & s' \ x & s'' \ x)\ + by (rule "\I") +qed + +AOT_define Consistent :: \\ \ \\ (\Consistent'(_')\) + cons: \Consistent(s) \\<^sub>d\<^sub>f \\p (s \ p & s \ \p)\ + +AOT_theorem "sit-cons": \Actual(s) \ Consistent(s)\ +proof(safe intro!: "\I" cons[THEN "\\<^sub>d\<^sub>fI"] "&I" Situation.\ + dest!: actual[THEN "\\<^sub>d\<^sub>fE"]; frule "&E"(1); drule "&E"(2)) + AOT_assume 0: \\p (s \ p \ p)\ + AOT_show \\\p (s \ p & s \ \p)\ + proof (rule "raa-cor:2") + AOT_assume \\p (s \ p & s \ \p)\ + then AOT_obtain p where \s \ p & s \ \p\ + using "\E"[rotated] by blast + AOT_hence \p & \p\ + using 0[THEN "\E"(1)[where \=\\\p\\, THEN "\E"], OF "log-prop-prop:2"] + 0[THEN "\E"(2)[where \=p], THEN "\E"] "&E" "&I" by blast + AOT_thus \p & \p\ for p by (metis "raa-cor:1") + qed +qed + +AOT_theorem "cons-rigid:1": \\Consistent(s) \ \\Consistent(s)\ +proof (rule "\I"; rule "\I") + AOT_assume \\Consistent(s)\ + AOT_hence \\p (s \ p & s \ \p)\ + using cons[THEN "\\<^sub>d\<^sub>fI", OF "&I", OF Situation.\] + by (metis "raa-cor:3") + then AOT_obtain p where p_prop: \s \ p & s \ \p\ + using "\E"[rotated] by blast + AOT_hence \\s \ p\ + using "&E"(1) "\E"(1) "lem2:1" by blast + moreover AOT_have \\s \ \p\ + using p_prop "T\" "&E" "\E"(1) + "modus-tollens:1" "raa-cor:3" "lem2:3"[unvarify p] + "log-prop-prop:2" by metis + ultimately AOT_have \\(s \ p & s \ \p)\ + by (metis "KBasic:3" "&I" "\E"(2)) + AOT_hence \\p \(s \ p & s \ \p)\ + by (rule "\I") + AOT_hence \\\p(s \ p & s \ \p)\ + by (metis Buridan "vdash-properties:10") + AOT_thus \\\Consistent(s)\ + apply (rule "qml:1"[axiom_inst, THEN "\E", THEN "\E", rotated]) + apply (rule RN) + using "\\<^sub>d\<^sub>fE" "&E"(2) cons "deduction-theorem" "raa-cor:3" by blast +next + AOT_assume \\\Consistent(s)\ + AOT_thus \\Consistent(s)\ using "qml:2"[axiom_inst, THEN "\E"] by auto +qed + +AOT_theorem "cons-rigid:2": \\Consistent(x) \ Consistent(x)\ +proof(rule "\I"; rule "\I") + AOT_assume 0: \\Consistent(x)\ + AOT_have \\(Situation(x) & \\p (x \ p & x \ \p))\ + apply (AOT_subst \Situation(x) & \\p (x \ p & x \ \p)\ \Consistent(x)\) + using cons "\E"(2) "Commutativity of \" "\Df" apply blast + by (simp add: 0) + AOT_hence \\Situation(x)\ and 1: \\\\p (x \ p & x \ \p)\ + using "RM\" "Conjunction Simplification"(1) "Conjunction Simplification"(2) + "modus-tollens:1" "raa-cor:3" by blast+ + AOT_hence 2: \Situation(x)\ by (metis "\E"(1) "possit-sit:2") + AOT_have 3: \\\\p (x \ p & x \ \p)\ + using 2 using 1 "KBasic:11" "\E"(2) by blast + AOT_show \Consistent(x)\ + proof (rule "raa-cor:1") + AOT_assume \\Consistent(x)\ + AOT_hence \\p (x \ p & x \ \p)\ + using 0 "\\<^sub>d\<^sub>fE" "conventions:5" 2 "cons-rigid:1"[unconstrain s, THEN "\E"] + "modus-tollens:1" "raa-cor:3" "\E"(4) by meson + then AOT_obtain p where \x \ p\ and 4: \x \ \p\ + using "\E"[rotated] "&E" by blast + AOT_hence \\x \ p\ + by (metis "2" "\E"(1) "lem2:1"[unconstrain s, THEN "\E"]) + moreover AOT_have \\x \ \p\ + using 4 "lem2:1"[unconstrain s, unvarify p, THEN "\E"] + by (metis 2 "\E"(1) "log-prop-prop:2") + ultimately AOT_have \\(x \ p & x \ \p)\ + by (metis "KBasic:3" "&I" "\E"(3) "raa-cor:3") + AOT_hence \\p \(x \ p & x \ \p)\ + by (metis "existential:1" "log-prop-prop:2") + AOT_hence \\\p (x \ p & x \ \p)\ + by (metis Buridan "vdash-properties:10") + AOT_thus \p & \p\ for p + using 3 "&I" by (metis "raa-cor:3") + qed +next + AOT_show \\Consistent(x)\ if \Consistent(x)\ + using "T\" that "vdash-properties:10" by blast +qed + +AOT_define possible :: \\ \ \\ (\Possible'(_')\) + pos: \Possible(s) \\<^sub>d\<^sub>f \Actual(s)\ + +AOT_theorem "sit-pos:1": \Actual(s) \ Possible(s)\ + apply(rule "\I"; rule pos[THEN "\\<^sub>d\<^sub>fI"]; rule "&I") + apply (meson "\\<^sub>d\<^sub>fE" actual "&E"(1)) + using "T\" "vdash-properties:10" by blast + +AOT_theorem "sit-pos:2": \\p ((s \ p) & \\p) \ \Possible(s)\ +proof(rule "\I") + AOT_assume \\p ((s \ p) & \\p)\ + then AOT_obtain p where a: \(s \ p) & \\p\ + using "\E"[rotated] by blast + AOT_hence \\(s \ p)\ + using "&E" by (metis "T\" "\E"(1) "lem2:3" "vdash-properties:10") + moreover AOT_have \\\p\ + using a[THEN "&E"(2)] by (metis "KBasic2:1" "\E"(2)) + ultimately AOT_have \\(s \ p & \p)\ + by (metis "KBasic:3" "&I" "\E"(3) "raa-cor:3") + AOT_hence \\p \(s \ p & \p)\ + by (rule "\I") + AOT_hence 1: \\\q (s \ q & \q)\ + by (metis Buridan "vdash-properties:10") + AOT_have \\\\q (s \ q \ q)\ + apply (AOT_subst \s \ q \ q\ \\(s \ q & \q)\ for: q) + apply (simp add: "oth-class-taut:1:a") + apply (AOT_subst \\\q \(s \ q & \q)\ \\q (s \ q & \q)\) + by (auto simp: "conventions:4" "df-rules-formulas[3]" "df-rules-formulas[4]" "\I" 1) + AOT_hence 0: \\\\q (s \ q \ q)\ + by (metis "\\<^sub>d\<^sub>fE" "conventions:5" "raa-cor:3") + AOT_show \\Possible(s)\ + apply (AOT_subst \Possible(s)\ \Situation(s) & \Actual(s)\) + apply (simp add: pos "\Df") + apply (AOT_subst \Actual(s)\ \Situation(s) & \q (s \ q \ q)\) + using actual "\Df" apply presburger + by (metis "0" "KBasic2:3" "&E"(2) "raa-cor:3" "vdash-properties:10") +qed + +AOT_theorem "pos-cons-sit:1": \Possible(s) \ Consistent(s)\ + by (auto simp: "sit-cons"[THEN "RM\", THEN "\E", + THEN "cons-rigid:2"[THEN "\E"(1)]] + intro!: "\I" dest!: pos[THEN "\\<^sub>d\<^sub>fE"] "&E"(2)) + +AOT_theorem "pos-cons-sit:2": \\s (Consistent(s) & \Possible(s))\ +proof - + AOT_obtain q\<^sub>1 where \q\<^sub>1 & \\q\<^sub>1\ + using "\\<^sub>d\<^sub>fE" "instantiation" "cont-tf:1" "cont-tf-thm:1" by blast + have "cond-prop": \ConditionOnPropositionalProperties + (\ \ . \\ = [\y q\<^sub>1 & \q\<^sub>1]\)\ + by (auto intro!: "cond-prop[I]" GEN "\I" "prop-prop1"[THEN "\\<^sub>d\<^sub>fI"] + "\I"(1)[where \=\\q\<^sub>1 & \q\<^sub>1\\, rotated, OF "log-prop-prop:2"]) + have rigid: \rigid_condition (\ \ . \\ = [\y q\<^sub>1 & \q\<^sub>1]\)\ + by (auto intro!: "strict-can:1[I]" GEN "\I" simp: "id-nec:2"[THEN "\E"]) + + AOT_obtain x where x_prop: \x = \<^bold>\s (\F (s[F] \ F = [\y q\<^sub>1 & \q\<^sub>1]))\ + using "ex:1:b"[THEN "\E"(1), OF "can-sit-desc:1", OF "cond-prop"] + "\E"[rotated] by blast + AOT_hence 0: \\<^bold>\(Situation(x) & \F (x[F] \ F = [\y q\<^sub>1 & \q\<^sub>1]))\ + using "\E" "actual-desc:2" by blast + AOT_hence \\<^bold>\(Situation(x))\ by (metis "Act-Basic:2" "&E"(1) "\E"(1)) + AOT_hence s_sit: \Situation(x)\ by (metis "\E"(1) "possit-sit:4") + AOT_have s_enc_prop: \\F (x[F] \ F = [\y q\<^sub>1 & \q\<^sub>1])\ + using "strict-sit"[OF rigid, OF "cond-prop", THEN "\E", OF x_prop]. + AOT_hence \x[\y q\<^sub>1 & \q\<^sub>1]\ + using "\E"(1)[rotated, OF "prop-prop2:2"] + "rule=I:1"[OF "prop-prop2:2"] "\E" by blast + AOT_hence \x \ (q\<^sub>1 & \q\<^sub>1)\ + using lem1[THEN "\E", OF s_sit, unvarify p, THEN "\E"(2), OF "log-prop-prop:2"] + by blast + AOT_hence \\(x \ (q\<^sub>1 & \q\<^sub>1))\ + using "lem2:1"[unconstrain s, THEN "\E", OF s_sit, unvarify p, + OF "log-prop-prop:2", THEN "\E"(1)] by blast + moreover AOT_have \\(x \ (q\<^sub>1 & \q\<^sub>1) \ \Actual(x))\ + proof(rule RN; rule "\I"; rule "raa-cor:2") + AOT_modally_strict { + AOT_assume \Actual(x)\ + AOT_hence \\p (x \ p \ p)\ + using actual[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2)] by blast + moreover AOT_assume \x \ (q\<^sub>1 & \q\<^sub>1)\ + ultimately AOT_show \q\<^sub>1 & \q\<^sub>1\ + using "\E"(1)[rotated, OF "log-prop-prop:2"] "\E" by metis + } + qed + ultimately AOT_have nec_not_actual_s: \\\Actual(x)\ + using "qml:1"[axiom_inst, THEN "\E", THEN "\E"] by blast + AOT_have 1: \\\p (x \ p & x \ \p)\ + proof (rule "raa-cor:2") + AOT_assume \\p (x \ p & x \ \p)\ + then AOT_obtain p where \x \ p & x \ \p\ + using "\E"[rotated] by blast + AOT_hence \x[\y p] & x[\y \p]\ + using lem1[unvarify p, THEN "\E", OF "log-prop-prop:2", + OF s_sit, THEN "\E"(1)] "&I" "&E" by metis + AOT_hence \[\y p] = [\y q\<^sub>1 & \q\<^sub>1]\ and \[\y \p] = [\y q\<^sub>1 & \q\<^sub>1]\ + by (auto intro!: "prop-prop2:2" s_enc_prop[THEN "\E"(1), THEN "\E"(1)] + elim: "&E") + AOT_hence i: \[\y p] = [\y \p]\ by (metis "rule=E" id_sym) + { + AOT_assume 0: \p\ + AOT_have \[\y p]x\ for x + by (auto intro!: "\\C"(1) "cqt:2" 0) + AOT_hence \[\y \p]x\ for x using i "rule=E" by fast + AOT_hence \\p\ + using "\\C"(1) by auto + } + moreover { + AOT_assume 0: \\p\ + AOT_have \[\y \p]x\ for x + by (auto intro!: "\\C"(1) "cqt:2" 0) + AOT_hence \[\y p]x\ for x using i[symmetric] "rule=E" by fast + AOT_hence \p\ + using "\\C"(1) by auto + } + ultimately AOT_show \p & \p\ for p by (metis "raa-cor:1" "raa-cor:3") + qed + AOT_have 2: \\Possible(x)\ + proof(rule "raa-cor:2") + AOT_assume \Possible(x)\ + AOT_hence \\Actual(x)\ + by (metis "\\<^sub>d\<^sub>fE" "&E"(2) pos) + moreover AOT_have \\\Actual(x)\ using nec_not_actual_s + using "\\<^sub>d\<^sub>fE" "conventions:5" "reductio-aa:2" by blast + ultimately AOT_show \\Actual(x) & \\Actual(x)\ by (rule "&I") + qed + show ?thesis + by(rule "\I"(2)[where \=x]; safe intro!: "&I" 2 s_sit cons[THEN "\\<^sub>d\<^sub>fI"] 1) +qed + +AOT_theorem "sit-classical:1": \\p (s \ p \ p) \ \q(s \ \q \ \s \ q)\ +proof(rule "\I"; rule GEN) + fix q + AOT_assume \\p (s \ p \ p)\ + AOT_hence \s \ q \ q\ and \s \ \q \ \q\ + using "\E"(1)[rotated, OF "log-prop-prop:2"] by blast+ + AOT_thus \s \ \q \ \s \ q\ + by (metis "deduction-theorem" "\I" "\E"(1) "\E"(2) "\E"(4)) +qed + +AOT_theorem "sit-classical:2": + \\p (s \ p \ p) \ \q\r((s \ (q \ r)) \ (s \ q \ s \ r))\ +proof(rule "\I"; rule GEN; rule GEN) + fix q r + AOT_assume \\p (s \ p \ p)\ + AOT_hence \: \s \ q \ q\ and \: \s \ r \ r\ and \: \(s \ (q \ r)) \ (q \ r)\ + using "\E"(1)[rotated, OF "log-prop-prop:2"] by blast+ + AOT_show \(s \ (q \ r)) \ (s \ q \ s \ r)\ + proof (safe intro!: "\I" "\I") + AOT_assume \s \ (q \ r)\ + moreover AOT_assume \s \ q\ + ultimately AOT_show \s \ r\ + using \ \ \ by (metis "\E"(1) "\E"(2) "vdash-properties:10") + next + AOT_assume \s \ q \ s \ r\ + AOT_thus \s \ (q \ r)\ + using \ \ \ by (metis "deduction-theorem" "\E"(1) "\E"(2) "\E") + qed +qed + +AOT_theorem "sit-classical:3": + \\p (s \ p \ p) \ ((s \ \\ \{\}) \ \\ s \ \{\})\ +proof (rule "\I") + AOT_assume \\p (s \ p \ p)\ + AOT_hence \: \s \ \{\} \ \{\}\ and \: \s \ \\ \{\} \ \\ \{\}\ for \ + using "\E"(1)[rotated, OF "log-prop-prop:2"] by blast+ + AOT_show \s \ \\ \{\} \ \\ s \ \{\}\ + proof (safe intro!: "\I" "\I" GEN) + fix \ + AOT_assume \s \ \\ \{\}\ + AOT_hence \\{\}\ using \ "\E"(2) "\E"(1) by blast + AOT_thus \s \ \{\}\ using \ "\E"(2) by blast + next + AOT_assume \\\ s \ \{\}\ + AOT_hence \s \ \{\}\ for \ using "\E"(2) by blast + AOT_hence \\{\}\ for \ using \ "\E"(1) by blast + AOT_hence \\\ \{\}\ by (rule GEN) + AOT_thus \s \ \\ \{\}\ using \ "\E"(2) by blast + qed +qed + +AOT_theorem "sit-classical:4": \\p (s \ p \ p) \ \q (s \ \q \ \s \ q)\ +proof(rule "\I"; rule GEN; rule "\I") + fix q + AOT_assume \\p (s \ p \ p)\ + AOT_hence \: \s \ q \ q\ and \: \s \ \q \ \q\ + using "\E"(1)[rotated, OF "log-prop-prop:2"] by blast+ + AOT_assume \s \ \q\ + AOT_hence \\q\ using \ "\E"(1) by blast + AOT_hence \q\ using "qml:2"[axiom_inst, THEN "\E"] by blast + AOT_hence \s \ q\ using \ "\E"(2) by blast + AOT_thus \\s \ q\ using "\\<^sub>d\<^sub>fE" "&E"(1) "\E"(1) "lem2:1" "true-in-s" by blast +qed + +AOT_theorem "sit-classical:5": + \\p (s \ p \ p) \ \q(\(s \ q) & \(s \ \ q))\ +proof (rule "\I") + AOT_obtain r where A: \r\ and \\\r\ + by (metis "&E"(1) "&E"(2) "\\<^sub>d\<^sub>fE" "instantiation" "cont-tf:1" "cont-tf-thm:1") + AOT_hence B: \\\r\ + using "KBasic:11" "\E"(2) by blast + moreover AOT_assume asm: \\ p (s \ p \ p)\ + AOT_hence \s \ r\ + using "\E"(2) A "\E"(2) by blast + AOT_hence 1: \\s \ r\ + using "\\<^sub>d\<^sub>fE" "&E"(1) "\E"(1) "lem2:1" "true-in-s" by blast + AOT_have \s \ \\r\ + using asm[THEN "\E"(1)[rotated, OF "log-prop-prop:2"], THEN "\E"(2)] B by blast + AOT_hence \\s \ \r\ + using "sit-classical:1"[THEN "\E", OF asm, + THEN "\E"(1)[rotated, OF "log-prop-prop:2"], THEN "\E"(1)] by blast + AOT_hence \\s \ r & \s \ \r\ + using 1 "&I" by blast + AOT_thus \\r (\s \ r & \s \ \r)\ + by (rule "\I") +qed + +AOT_theorem "sit-classical:6": + \\s \p (s \ p \ p)\ +proof - + have "cond-prop": \ConditionOnPropositionalProperties + (\ \ . \\q (q & \ = [\y q])\)\ + proof (safe intro!: "cond-prop[I]" GEN "\I") + fix F + AOT_modally_strict { + AOT_assume \\q (q & F = [\y q])\ + then AOT_obtain q where \q & F = [\y q]\ + using "\E"[rotated] by blast + AOT_hence \F = [\y q]\ + using "&E" by blast + AOT_hence \\q F = [\y q]\ + by (rule "\I") + AOT_thus \Propositional([F])\ + by (metis "\\<^sub>d\<^sub>fI" "prop-prop1") + } + qed + AOT_have \\s \F (s[F] \ \q (q & F = [\y q]))\ + using "comp-sit:1"[OF "cond-prop"]. + then AOT_obtain s\<^sub>0 where s\<^sub>0_prop: \\F (s\<^sub>0[F] \ \q (q & F = [\y q]))\ + using "Situation.\E"[rotated] by meson + AOT_have \\p (s\<^sub>0 \ p \ p)\ + proof(safe intro!: GEN "\I" "\I") + fix p + AOT_assume \s\<^sub>0 \ p\ + AOT_hence \s\<^sub>0[\y p]\ + using lem1[THEN "\E", OF Situation.\, THEN "\E"(1)] by blast + AOT_hence \\q (q & [\y p] = [\y q])\ + using s\<^sub>0_prop[THEN "\E"(1)[rotated, OF "prop-prop2:2"], THEN "\E"(1)] by blast + then AOT_obtain q\<^sub>1 where q\<^sub>1_prop: \q\<^sub>1 & [\y p] = [\y q\<^sub>1]\ + using "\E"[rotated] by blast + AOT_hence \p = q\<^sub>1\ + by (metis "&E"(2) "\E"(2) "p-identity-thm2:3") + AOT_thus \p\ + using q\<^sub>1_prop[THEN "&E"(1)] "rule=E" id_sym by fast + next + fix p + AOT_assume \p\ + moreover AOT_have \[\y p] = [\y p]\ + by (simp add: "rule=I:1"[OF "prop-prop2:2"]) + ultimately AOT_have \p & [\y p] = [\y p]\ + using "&I" by blast + AOT_hence \\q (q & [\y p] = [\y q])\ + by (rule "\I") + AOT_hence \s\<^sub>0[\y p]\ + using s\<^sub>0_prop[THEN "\E"(1)[rotated, OF "prop-prop2:2"], THEN "\E"(2)] by blast + AOT_thus \s\<^sub>0 \ p\ + using lem1[THEN "\E", OF Situation.\, THEN "\E"(2)] by blast + qed + AOT_hence \\p (s\<^sub>0 \ p \ p)\ + using "&I" by blast + AOT_thus \\s \p (s \ p \ p)\ + by (rule "Situation.\I") +qed + +AOT_define PossibleWorld :: \\ \ \\ (\PossibleWorld'(_')\) + "world:1": \PossibleWorld(x) \\<^sub>d\<^sub>f Situation(x) & \\p(x \ p \ p)\ + +AOT_theorem "world:2": \\x PossibleWorld(x)\ +proof - + AOT_obtain s where s_prop: \\p (s \ p \ p)\ + using "sit-classical:6" "Situation.\E"[rotated] by meson + AOT_have \\p (s \ p \ p)\ + proof(safe intro!: GEN "\I" "\I") + fix p + AOT_assume \s \ p\ + AOT_thus \p\ + using s_prop[THEN "\E"(2), THEN "\E"(1)] by blast + next + fix p + AOT_assume \p\ + AOT_thus \s \ p\ + using s_prop[THEN "\E"(2), THEN "\E"(2)] by blast + qed + AOT_hence \\\p (s \ p \ p)\ + by (metis "T\"[THEN "\E"]) + AOT_hence \\\p (s \ p \ p)\ + using s_prop "&I" by blast + AOT_hence \PossibleWorld(s)\ + using "world:1"[THEN "\\<^sub>d\<^sub>fI"] Situation.\ "&I" by blast + AOT_thus \\x PossibleWorld(x)\ + by (rule "\I") +qed + +AOT_theorem "world:3": \PossibleWorld(\) \ \\\ +proof (rule "\I") + AOT_assume \PossibleWorld(\)\ + AOT_hence \Situation(\)\ + using "world:1"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + AOT_hence \A!\\ + by (metis "\\<^sub>d\<^sub>fE" "&E"(1) situations) + AOT_thus \\\\ + by (metis "russell-axiom[exe,1].\_denotes_asm") +qed + +AOT_theorem "rigid-pw:1": \PossibleWorld(x) \ \PossibleWorld(x)\ +proof(safe intro!: "\I" "\I") + AOT_assume \PossibleWorld(x)\ + AOT_hence \Situation(x) & \\p(x \ p \ p)\ + using "world:1"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_hence \\Situation(x) & \\\p(x \ p \ p)\ + by (metis "S5Basic:1" "&I" "&E"(1) "&E"(2) "\E"(1) "possit-sit:1") + AOT_hence 0: \\(Situation(x) & \\p(x \ p \ p))\ + by (metis "KBasic:3" "\E"(2)) + AOT_show \\PossibleWorld(x)\ + by (AOT_subst \PossibleWorld(x)\ \Situation(x) & \\p(x \ p \ p)\) + (auto simp: "\Df" "world:1" 0) +next + AOT_show \PossibleWorld(x)\ if \\PossibleWorld(x)\ + using that "qml:2"[axiom_inst, THEN "\E"] by blast +qed + +AOT_theorem "rigid-pw:2": \\PossibleWorld(x) \ PossibleWorld(x)\ + using "rigid-pw:1" + by (meson "RE\" "S5Basic:2" "\E"(2) "\E"(6) "Commutativity of \") + +AOT_theorem "rigid-pw:3": \\PossibleWorld(x) \ \PossibleWorld(x)\ + using "rigid-pw:1" "rigid-pw:2" by (meson "\E"(5)) + +AOT_theorem "rigid-pw:4": \\<^bold>\PossibleWorld(x) \ PossibleWorld(x)\ + by (metis "Act-Sub:3" "\I" "\I" "\E"(6) "nec-imp-act" "rigid-pw:1" "rigid-pw:2") + +AOT_register_rigid_restricted_type + PossibleWorld: \PossibleWorld(\)\ +proof + AOT_modally_strict { + AOT_show \\x PossibleWorld(x)\ using "world:2". + } +next + AOT_modally_strict { + AOT_show \PossibleWorld(\) \ \\\ for \ using "world:3". + } +next + AOT_modally_strict { + AOT_show \\\(PossibleWorld(\) \ \PossibleWorld(\))\ + by (meson GEN "\I" "\E"(1) "rigid-pw:1") + } +qed +AOT_register_variable_names + PossibleWorld: w + +AOT_theorem "world-pos": \Possible(w)\ +proof (safe intro!: "\\<^sub>d\<^sub>fE"[OF "world:1", OF PossibleWorld.\, THEN "&E"(1)] + pos[THEN "\\<^sub>d\<^sub>fI"] "&I" ) + AOT_have \\\p (w \ p \ p)\ + using "world:1"[THEN "\\<^sub>d\<^sub>fE", OF PossibleWorld.\, THEN "&E"(2)]. + AOT_hence \\\p (w \ p \ p)\ + proof (rule "RM\"[THEN "\E", rotated]; safe intro!: "\I" GEN) + AOT_modally_strict { + fix p + AOT_assume \\p (w \ p \ p)\ + AOT_hence \w \ p \ p\ using "\E"(2) by blast + moreover AOT_assume \w \ p\ + ultimately AOT_show p using "\E"(1) by blast + } + qed + AOT_hence 0: \\(Situation(w) & \p (w \ p \ p))\ + using "world:1"[THEN "\\<^sub>d\<^sub>fE", OF PossibleWorld.\, THEN "&E"(1), + THEN "possit-sit:1"[THEN "\E"(1)]] + by (metis "KBasic:16" "&I" "vdash-properties:10") + AOT_show \\Actual(w)\ + by (AOT_subst \Actual(w)\ \Situation(w) & \p (w \ p \ p)\) + (auto simp: actual "\Df" 0) +qed + +AOT_theorem "world-cons:1": \Consistent(w)\ + using "world-pos" + using "pos-cons-sit:1"[unconstrain s, THEN "\E", THEN "\E"] + by (meson "\\<^sub>d\<^sub>fE" "&E"(1) pos) + +AOT_theorem "world-cons:2": \\TrivialSituation(w)\ +proof(rule "raa-cor:2") + AOT_assume \TrivialSituation(w)\ + AOT_hence \Situation(w) & \p w \ p\ + using "df-null-trivial:2"[THEN "\\<^sub>d\<^sub>fE"] by blast + AOT_hence 0: \\w \ (\p (p & \p))\ + using "&E" + by (metis "Buridan\" "T\" "&E"(2) "\E"(1) "lem2:3"[unconstrain s, THEN "\E"] + "log-prop-prop:2" "rule-ui:1" "universal-cor" "\E") + AOT_have \\\p (w \ p \ p)\ + using PossibleWorld.\ "world:1"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2)] by metis + AOT_hence \\p \(w \ p \ p)\ + using "Buridan\"[THEN "\E"] by blast + AOT_hence \\(w \ (\p (p & \p)) \ (\p (p & \p)))\ + by (metis "log-prop-prop:2" "rule-ui:1") + AOT_hence \\(w \ (\p (p & \p)) \ (\p (p & \p)))\ + using "RM\"[THEN "\E"] "\I" "\E"(1) by meson + AOT_hence \\(\p (p & \p))\ using 0 + by (metis "KBasic2:4" "\E"(1) "\E") + moreover AOT_have \\\(\p (p & \p))\ + by (metis "instantiation" "KBasic2:1" RN "\E"(1) "raa-cor:2") + ultimately AOT_show \\(\p (p & \p)) & \\(\p (p & \p))\ + using "&I" by blast +qed + +AOT_theorem "rigid-truth-at:1": \w \ p \ \w \ p\ + using "lem2:1"[unconstrain s, THEN "\E", + OF PossibleWorld.\[THEN "world:1"[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(1)]]. + +AOT_theorem "rigid-truth-at:2": \\w \ p \ w \ p\ + using "lem2:2"[unconstrain s, THEN "\E", + OF PossibleWorld.\[THEN "world:1"[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(1)]]. + +AOT_theorem "rigid-truth-at:3": \\w \ p \ \w \ p\ + using "lem2:3"[unconstrain s, THEN "\E", + OF PossibleWorld.\[THEN "world:1"[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(1)]]. + +AOT_theorem "rigid-truth-at:4": \\<^bold>\w \ p \ w \ p\ + using "lem2:4"[unconstrain s, THEN "\E", + OF PossibleWorld.\[THEN "world:1"[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(1)]]. + +AOT_theorem "rigid-truth-at:5": \\w \ p \ \\w \ p\ + using "lem2:5"[unconstrain s, THEN "\E", + OF PossibleWorld.\[THEN "world:1"[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(1)]]. + +AOT_define Maximal :: \\ \ \\ (\Maximal'(_')\) + max: \Maximal(s) \\<^sub>d\<^sub>f \p (s \ p \ s \ \p)\ + +AOT_theorem "world-max": \Maximal(w)\ +proof(safe intro!: PossibleWorld.\[THEN "\\<^sub>d\<^sub>fE"[OF "world:1"], THEN "&E"(1)] + GEN "\\<^sub>d\<^sub>fI"[OF max] "&I" ) + fix q + AOT_have \\(w \ q \ w \ \q)\ + proof(rule "RM\"[THEN "\E"]; (rule "\I")?) + AOT_modally_strict { + AOT_assume \\p (w \ p \ p)\ + AOT_hence \w \ q \ q\ and \w \ \q \ \q\ + using "\E"(1)[rotated, OF "log-prop-prop:2"] by blast+ + AOT_thus \w \ q \ w \ \q\ + by (metis "\I"(1) "\I"(2) "\E"(3) "reductio-aa:1") + } + next + AOT_show \\\p (w \ p \ p)\ + using PossibleWorld.\[THEN "\\<^sub>d\<^sub>fE"[OF "world:1"], THEN "&E"(2)]. + qed + AOT_hence \\w \ q \ \w \ \q\ + using "KBasic2:2"[THEN "\E"(1)] by blast + AOT_thus \w \ q \ w \ \q\ + using "lem2:2"[unconstrain s, THEN "\E", unvarify p, + OF PossibleWorld.\[THEN "\\<^sub>d\<^sub>fE"[OF "world:1"], THEN "&E"(1)], + THEN "\E"(1), OF "log-prop-prop:2"] + by (metis "\I"(1) "\I"(2) "\E"(3) "raa-cor:2") +qed + +AOT_theorem "world=maxpos:1": \Maximal(x) \ \Maximal(x)\ +proof (AOT_subst \Maximal(x)\ \Situation(x) & \p (x \ p \ x \ \p)\; + safe intro!: max "\Df" "\I"; frule "&E"(1); drule "&E"(2)) + AOT_assume sit_x: \Situation(x)\ + AOT_hence nec_sit_x: \\Situation(x)\ + by (metis "\E"(1) "possit-sit:1") + AOT_assume \\p (x \ p \ x \ \p)\ + AOT_hence \x \ p \ x \ \p\ for p + using "\E"(1)[rotated, OF "log-prop-prop:2"] by blast + AOT_hence \\x \ p \ \x \ \p\ for p + using "lem2:1"[unconstrain s, THEN "\E", OF sit_x, unvarify p, + OF "log-prop-prop:2", THEN "\E"(1)] + by (metis "\I"(1) "\I"(2) "\E"(2) "raa-cor:1") + AOT_hence \\(x \ p \ x \ \p)\ for p + by (metis "KBasic:15" "\E") + AOT_hence \\p \(x \ p \ x \ \p)\ + by (rule GEN) + AOT_hence \\\p (x \ p \ x \ \p)\ + by (rule BF[THEN "\E"]) + AOT_thus \\(Situation(x) & \p (x \ p \ x \ \p))\ + using nec_sit_x by (metis "KBasic:3" "&I" "\E"(2)) +qed + +AOT_theorem "world=maxpos:2": \PossibleWorld(x) \ Maximal(x) & Possible(x)\ +proof(safe intro!: "\I" "\I" "&I" "world-pos"[unconstrain w, THEN "\E"] + "world-max"[unconstrain w, THEN "\E"]; + frule "&E"(2); drule "&E"(1)) + AOT_assume pos_x: \Possible(x)\ + AOT_have \\(Situation(x) & \p(x \ p \ p))\ + apply (AOT_subst (reverse) \Situation(x) & \p(x \ p \ p)\ \Actual(x)\) + using actual "\Df" apply presburger + using "\\<^sub>d\<^sub>fE" "&E"(2) pos pos_x by blast + AOT_hence 0: \\\p(x \ p \ p)\ + by (metis "KBasic2:3" "&E"(2) "vdash-properties:6") + AOT_assume max_x: \Maximal(x)\ + AOT_hence sit_x: \Situation(x)\ by (metis "\\<^sub>d\<^sub>fE" max_x "&E"(1) max) + AOT_have \\Maximal(x)\ using "world=maxpos:1"[THEN "\E", OF max_x] by simp + moreover AOT_have \\Maximal(x) \ \(\p(x \ p \ p) \ \p (x \ p \ p))\ + proof(safe intro!: "\I" RM GEN) + AOT_modally_strict { + fix p + AOT_assume 0: \Maximal(x)\ + AOT_assume 1: \\p (x \ p \ p)\ + AOT_show \x \ p \ p\ + proof(safe intro!: "\I" "\I" 1[THEN "\E"(2), THEN "\E"]; rule "raa-cor:1") + AOT_assume \\x \ p\ + AOT_hence \x \ \p\ + using 0[THEN "\\<^sub>d\<^sub>fE"[OF max], THEN "&E"(2), THEN "\E"(2)] + 1 by (metis "\E"(2)) + AOT_hence \\p\ + using 1[THEN "\E"(1), OF "log-prop-prop:2", THEN "\E"] by blast + moreover AOT_assume p + ultimately AOT_show \p & \p\ using "&I" by blast + qed + } + qed + ultimately AOT_have \\(\p(x \ p \ p) \ \p (x \ p \ p))\ + using "\E" by blast + AOT_hence \\\p(x \ p \ p) \ \\p(x \ p \ p)\ + by (metis "KBasic:13"[THEN "\E"]) + AOT_hence \\\p(x \ p \ p)\ + using 0 "\E" by blast + AOT_thus \PossibleWorld(x)\ + using "\\<^sub>d\<^sub>fI"[OF "world:1", OF "&I", OF sit_x] by blast +qed + +AOT_define NecImpl :: \\ \ \ \ \\ (infixl \\\ 26) + "nec-impl-p:1": \p \ q \\<^sub>d\<^sub>f \(p \ q)\ +AOT_define NecEquiv :: \\ \ \ \ \\ (infixl \\\ 21) + "nec-impl-p:2": \p \ q \\<^sub>d\<^sub>f (p \ q) & (q \ p)\ + +AOT_theorem "nec-equiv-nec-im": \p \ q \ \(p \ q)\ +proof(safe intro!: "\I" "\I") + AOT_assume \p \ q\ + AOT_hence \(p \ q)\ and \(q \ p)\ + using "nec-impl-p:2"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast+ + AOT_hence \\(p \ q)\ and \\(q \ p)\ + using "nec-impl-p:1"[THEN "\\<^sub>d\<^sub>fE"] by blast+ + AOT_thus \\(p \ q)\ by (metis "KBasic:4" "&I" "\E"(2)) +next + AOT_assume \\(p \ q)\ + AOT_hence \\(p \ q)\ and \\(q \ p)\ + using "KBasic:4" "&E" "\E"(1) by blast+ + AOT_hence \(p \ q)\ and \(q \ p)\ + using "nec-impl-p:1"[THEN "\\<^sub>d\<^sub>fI"] by blast+ + AOT_thus \p \ q\ + using "nec-impl-p:2"[THEN "\\<^sub>d\<^sub>fI"] "&I" by blast +qed + +(* TODO: PLM: discuss these; still not in PLM *) +AOT_theorem world_closed_lem_1_a: + \(s \ (\ & \)) \ (\p (s \ p \ p) \ (s \ \ & s \ \))\ +proof(safe intro!: "\I") + AOT_assume \\ p (s \ p \ p)\ + AOT_hence \s \ (\ & \) \ (\ & \)\ and \s \ \ \ \\ and \s \ \ \ \\ + using "\E"(1)[rotated, OF "log-prop-prop:2"] by blast+ + moreover AOT_assume \s \ (\ & \)\ + ultimately AOT_show \s \ \ & s \ \\ + by (metis "&I" "&E"(1) "&E"(2) "\E"(1) "\E"(2)) +qed + +AOT_theorem world_closed_lem_1_b: + \(s \ \ & (\ \ q)) \ (\p (s \ p \ p) \ s \ q)\ +proof(safe intro!: "\I") + AOT_assume \\ p (s \ p \ p)\ + AOT_hence \s \ \ \ \\ for \ + using "\E"(1)[rotated, OF "log-prop-prop:2"] by blast + moreover AOT_assume \s \ \ & (\ \ q)\ + ultimately AOT_show \s \ q\ + by (metis "&E"(1) "&E"(2) "\E"(1) "\E"(2) "\E") +qed + +AOT_theorem world_closed_lem_1_c: + \(s \ \ & s \ (\ \ \)) \ (\p (s \ p \ p) \ s \ \)\ +proof(safe intro!: "\I") + AOT_assume \\ p (s \ p \ p)\ + AOT_hence \s \ \ \ \\ for \ + using "\E"(1)[rotated, OF "log-prop-prop:2"] by blast + moreover AOT_assume \s \ \ & s \ (\ \ \)\ + ultimately AOT_show \s \ \\ + by (metis "&E"(1) "&E"(2) "\E"(1) "\E"(2) "\E") +qed + +AOT_theorem "world-closed-lem:1[0]": + \q \ (\p (s \ p \ p) \ s \ q)\ + by (meson "\I" "\E"(2) "log-prop-prop:2" "rule-ui:1") + +AOT_theorem "world-closed-lem:1[1]": + \s \ p\<^sub>1 & (p\<^sub>1 \ q) \ (\p (s \ p \ p) \ s \ q)\ + using world_closed_lem_1_b. + +AOT_theorem "world-closed-lem:1[2]": + \s \ p\<^sub>1 & s \ p\<^sub>2 & ((p\<^sub>1 & p\<^sub>2) \ q) \ (\p (s \ p \ p) \ s \ q)\ + using world_closed_lem_1_b world_closed_lem_1_a + by (metis (full_types) "&I" "&E" "\I" "\E") + +AOT_theorem "world-closed-lem:1[3]": + \s \ p\<^sub>1 & s \ p\<^sub>2 & s \ p\<^sub>3 & ((p\<^sub>1 & p\<^sub>2 & p\<^sub>3) \ q) \ (\p (s \ p \ p) \ s \ q)\ + using world_closed_lem_1_b world_closed_lem_1_a + by (metis (full_types) "&I" "&E" "\I" "\E") + +AOT_theorem "world-closed-lem:1[4]": + \s \ p\<^sub>1 & s \ p\<^sub>2 & s \ p\<^sub>3 & s \ p\<^sub>4 & ((p\<^sub>1 & p\<^sub>2 & p\<^sub>3 & p\<^sub>4) \ q) \ + (\p (s \ p \ p) \ s \ q)\ + using world_closed_lem_1_b world_closed_lem_1_a + by (metis (full_types) "&I" "&E" "\I" "\E") + +AOT_theorem "coherent:1": \w \ \p \ \w \ p\ +proof(safe intro!: "\I" "\I") + AOT_assume 1: \w \ \p\ + AOT_show \\w \ p\ + proof(rule "raa-cor:2") + AOT_assume \w \ p\ + AOT_hence \w \ p & w \ \p\ using 1 "&I" by blast + AOT_hence \\q (w \ q & w \ \q)\ by (rule "\I") + moreover AOT_have \\\q (w \ q & w \ \q)\ + using "world-cons:1"[THEN "\\<^sub>d\<^sub>fE"[OF cons], THEN "&E"(2)]. + ultimately AOT_show \\q (w \ q & w \ \q) & \\q (w \ q & w \ \q)\ + using "&I" by blast + qed +next + AOT_assume \\w \ p\ + AOT_thus \w \ \p\ + using "world-max"[THEN "\\<^sub>d\<^sub>fE"[OF max], THEN "&E"(2)] + by (metis "\E"(2) "log-prop-prop:2" "rule-ui:1") +qed + +AOT_theorem "coherent:2": \w \ p \ \w \ \p\ + by (metis "coherent:1" "deduction-theorem" "\I" "\E"(1) "\E"(2) "raa-cor:3") + +AOT_theorem "act-world:1": \\w \p (w \ p \ p)\ +proof - + AOT_obtain s where s_prop: \\p (s \ p \ p)\ + using "sit-classical:6" "Situation.\E"[rotated] by meson + AOT_hence \\\p (s \ p \ p)\ + by (metis "T\" "vdash-properties:10") + AOT_hence \PossibleWorld(s)\ + using "world:1"[THEN "\\<^sub>d\<^sub>fI"] Situation.\ "&I" by blast + AOT_hence \PossibleWorld(s) & \p (s \ p \ p)\ + using "&I" s_prop by blast + thus ?thesis by (rule "\I") +qed + +AOT_theorem "act-world:2": \\!w Actual(w)\ +proof - + AOT_obtain w where w_prop: \\p (w \ p \ p)\ + using "act-world:1" "PossibleWorld.\E"[rotated] by meson + AOT_have sit_s: \Situation(w)\ + using PossibleWorld.\ "world:1"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(1)] by blast + show ?thesis + proof (safe intro!: "uniqueness:1"[THEN "\\<^sub>d\<^sub>fI"] "\I"(2) "&I" GEN "\I" + PossibleWorld.\ actual[THEN "\\<^sub>d\<^sub>fI"] sit_s + "sit-identity"[unconstrain s, unconstrain s', THEN "\E", + THEN "\E", THEN "\E"(2)] "\I" + w_prop[THEN "\E"(2), THEN "\E"(1)]) + AOT_show \PossibleWorld(w)\ using PossibleWorld.\. + next + AOT_show \Situation(w)\ + by (simp add: sit_s) + next + fix y p + AOT_assume w_asm: \PossibleWorld(y) & Actual(y)\ + AOT_assume \w \ p\ + AOT_hence p: \p\ + using w_prop[THEN "\E"(2), THEN "\E"(1)] by blast + AOT_show \y \ p\ + proof(rule "raa-cor:1") + AOT_assume \\y \ p\ + AOT_hence \y \ \p\ + by (metis "coherent:1"[unconstrain w, THEN "\E"] "&E"(1) "\E"(2) w_asm) + AOT_hence \\p\ + using w_asm[THEN "&E"(2), THEN actual[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(2), + THEN "\E"(1), rotated, OF "log-prop-prop:2"] + "\E" by blast + AOT_thus \p & \p\ using p "&I" by blast + qed + next + AOT_show \w \ p\ if \y \ p\ and \PossibleWorld(y) & Actual(y)\ for p y + using that(2)[THEN "&E"(2), THEN actual[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(2), + THEN "\E"(2), THEN "\E", OF that(1)] + w_prop[THEN "\E"(2), THEN "\E"(2)] by blast + next + AOT_show \Situation(y)\ if \PossibleWorld(y) & Actual(y)\ for y + using that[THEN "&E"(1)] "world:1"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(1)] by blast + next + AOT_show \Situation(w)\ + using sit_s by blast + qed(simp) +qed + +AOT_theorem "pre-walpha": \\<^bold>\w Actual(w)\\ + using "A-Exists:2" "RA[2]" "act-world:2" "\E"(2) by blast + +AOT_define TheActualWorld :: \\\<^sub>s\ (\\<^bold>w\<^sub>\\) + "w-alpha": \\<^bold>w\<^sub>\ =\<^sub>d\<^sub>f \<^bold>\w Actual(w)\ + +(* TODO: not in PLM *) +AOT_theorem true_in_truth_act_true: \\ \ p \ \<^bold>\p\ +proof(safe intro!: "\I" "\I") + AOT_have true_def: \\<^bold>\\<^sub>\ \ = \<^bold>\x (A!x & \F (x[F] \ \p(p & F = [\y p])))\ + by (simp add: "A-descriptions" "rule-id-df:1[zero]" "the-true:1") + AOT_hence true_den: \\<^bold>\\<^sub>\ \\\ + using "t=t-proper:1" "vdash-properties:6" by blast + { + AOT_assume \\ \ p\ + AOT_hence \\[\y p]\ + by (metis "\\<^sub>d\<^sub>fE" "con-dis-i-e:2:b" "prop-enc" "true-in-s") + AOT_hence \\<^bold>\x(A!x & \F (x[F] \ \q (q & F = [\y q])))[\y p]\ + using "rule=E" true_def true_den by fast + AOT_hence \\<^bold>\\q (q & [\y p] = [\y q])\ + using "\E"(1) "desc-nec-encode:1"[unvarify F] "prop-prop2:2" by fast + AOT_hence \\q \<^bold>\(q & [\y p] = [\y q])\ + by (metis "Act-Basic:10" "\E"(1)) + then AOT_obtain q where \\<^bold>\(q & [\y p] = [\y q])\ + using "\E"[rotated] by blast + AOT_hence actq: \\<^bold>\q\ and \\<^bold>\[\y p] = [\y q]\ + using "Act-Basic:2" "intro-elim:3:a" "&E" by blast+ + AOT_hence \[\y p] = [\y q]\ + using "id-act:1"[unvarify \ \, THEN "\E"(2)] "prop-prop2:2" by blast + AOT_hence \p = q\ + by (metis "intro-elim:3:b" "p-identity-thm2:3") + AOT_thus \\<^bold>\p\ + using actq "rule=E" id_sym by blast + } + { + AOT_assume \\<^bold>\p\ + AOT_hence \\<^bold>\(p & [\y p] = [\y p])\ + by (auto intro!: "Act-Basic:2"[THEN "\E"(2)] "&I" + intro: "RA[2]" "=I"(1)[OF "prop-prop2:2"]) + AOT_hence \\q \<^bold>\(q & [\y p] = [\y q])\ + using "\I" by fast + AOT_hence \\<^bold>\\q (q & [\y p] = [\y q])\ + by (metis "Act-Basic:10" "\E"(2)) + AOT_hence \\<^bold>\x(A!x & \F (x[F] \ \q (q & F = [\y q])))[\y p]\ + using "\E"(2) "desc-nec-encode:1"[unvarify F] "prop-prop2:2" by fast + AOT_hence \\[\y p]\ + using "rule=E" true_def true_den id_sym by fast + AOT_thus \\ \ p\ + by (safe intro!: "true-in-s"[THEN "\\<^sub>d\<^sub>fI"] "&I" "possit-sit:6" + "prop-enc"[THEN "\\<^sub>d\<^sub>fI"] true_den) + } +qed + +AOT_theorem "T-world": \\ = \<^bold>w\<^sub>\\ +proof - + AOT_have true_den: \\<^bold>\\<^sub>\ \\\ + using "Situation.res-var:3" "possit-sit:6" "\E" by blast + AOT_have \\<^bold>\\p (\ \ p \ p)\ + proof (safe intro!: "logic-actual-nec:3"[axiom_inst, THEN "\E"(2)] GEN + "logic-actual-nec:2"[axiom_inst, THEN "\E"(2)] "\I") + fix p + AOT_assume \\<^bold>\\ \ p\ + AOT_hence \\ \ p\ + using "lem2:4"[unconstrain s, unvarify \, OF true_den, + THEN "\E", OF "possit-sit:6"] "\E"(1) by blast + AOT_thus \\<^bold>\p\ using true_in_truth_act_true "\E"(1) by blast + qed + moreover AOT_have \\<^bold>\(Situation(\) & \p (\ \ p \ p)) \ \<^bold>\Actual(\)\ for \ + using actual[THEN "\Df", THEN "conventions:3"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2)], + THEN "RA[2]", THEN "act-cond"[THEN "\E"]]. + ultimately AOT_have act_act_true: \\<^bold>\Actual(\)\ + using "possit-sit:4"[unvarify x, OF true_den, THEN "\E"(2), OF "possit-sit:6"] + "Act-Basic:2"[THEN "\E"(2), OF "&I"] "\E" by blast + AOT_hence \\Actual(\)\ by (metis "Act-Sub:3" "vdash-properties:10") + AOT_hence \Possible(\)\ + by (safe intro!: pos[THEN "\\<^sub>d\<^sub>fI"] "&I" "possit-sit:6") + moreover AOT_have \Maximal(\)\ + proof (safe intro!: max[THEN "\\<^sub>d\<^sub>fI"] "&I" "possit-sit:6" GEN) + fix p + AOT_have \\<^bold>\p \ \<^bold>\\p\ + by (simp add: "Act-Basic:1") + moreover AOT_have \\ \ p\ if \\<^bold>\p\ + using that true_in_truth_act_true[THEN "\E"(2)] by blast + moreover AOT_have \\ \ \p\ if \\<^bold>\\p\ + using that true_in_truth_act_true[unvarify p, THEN "\E"(2)] + "log-prop-prop:2" by blast + ultimately AOT_show \\ \ p \ \ \ \p\ + using "\I"(3) "\I" by blast + qed + ultimately AOT_have \PossibleWorld(\)\ + by (safe intro!: "world=maxpos:2"[unvarify x, OF true_den, THEN "\E"(2)] "&I") + AOT_hence \\<^bold>\PossibleWorld(\)\ + using "rigid-pw:4"[unvarify x, OF true_den, THEN "\E"(2)] by blast + AOT_hence 1: \\<^bold>\(PossibleWorld(\) & Actual(\))\ + using act_act_true "Act-Basic:2" "df-simplify:2" "intro-elim:3:b" by blast + AOT_have \\<^bold>w\<^sub>\ = \<^bold>\w(Actual(w))\ + using "rule-id-df:1[zero]"[OF "w-alpha", OF "pre-walpha"] by simp + moreover AOT_have w_act_den: \\<^bold>w\<^sub>\\\ + using calculation "t=t-proper:1" "\E" by blast + ultimately AOT_have \\z (\<^bold>\(PossibleWorld(z) & Actual(z)) \ z = \<^bold>w\<^sub>\)\ + using "nec-hintikka-scheme"[unvarify x] "\E"(1) "&E" by blast + AOT_thus \\ = \<^bold>w\<^sub>\\ + using "\E"(1)[rotated, OF true_den] 1 "\E" by blast +qed + +AOT_act_theorem "truth-at-alpha:1": \p \ \<^bold>w\<^sub>\ = \<^bold>\x (ExtensionOf(x, p))\ + by (metis "rule=E" "T-world" "deduction-theorem" "ext-p-tv:3" id_sym "\I" + "\E"(1) "\E"(2) "q-True:1") + +AOT_act_theorem "truth-at-alpha:2": \p \ \<^bold>w\<^sub>\ \ p\ +proof - + AOT_have \PossibleWorld(\<^bold>w\<^sub>\)\ + using "&E"(1) "pre-walpha" "rule-id-df:2:b[zero]" "\E" + "w-alpha" "y-in:3" by blast + AOT_hence sit_w_alpha: \Situation(\<^bold>w\<^sub>\)\ + by (metis "\\<^sub>d\<^sub>fE" "&E"(1) "world:1") + AOT_have w_alpha_den: \\<^bold>w\<^sub>\\\ + using "pre-walpha" "rule-id-df:2:b[zero]" "w-alpha" by blast + AOT_have \p \ \\<^bold>\p\ + using "q-True:3" by force + moreover AOT_have \\ = \<^bold>w\<^sub>\\ + using "T-world" by auto + ultimately AOT_have \p \ \<^bold>w\<^sub>\\<^bold>\p\ + using "rule=E" by fast + moreover AOT_have \\<^bold>w\<^sub>\ \<^bold>\ p \ \<^bold>w\<^sub>\ \ p\ + using lem1[unvarify x, OF w_alpha_den, THEN "\E", OF sit_w_alpha] + using "\S"(1) "\E"(1) "Commutativity of \" "\Df" sit_w_alpha "true-in-s" by blast + ultimately AOT_show \p \ \<^bold>w\<^sub>\ \ p\ + by (metis "\E"(5)) +qed + +AOT_theorem "alpha-world:1": \PossibleWorld(\<^bold>w\<^sub>\)\ +proof - + AOT_have 0: \\<^bold>w\<^sub>\ = \<^bold>\w Actual(w)\ + using "pre-walpha" "rule-id-df:1[zero]" "w-alpha" by blast + AOT_hence walpha_den: \\<^bold>w\<^sub>\\\ + by (metis "t=t-proper:1" "vdash-properties:6") + AOT_have \\<^bold>\(PossibleWorld(\<^bold>w\<^sub>\) & Actual(\<^bold>w\<^sub>\))\ + by (rule "actual-desc:2"[unvarify x, OF walpha_den, THEN "\E"]) (fact 0) + AOT_hence \\<^bold>\PossibleWorld(\<^bold>w\<^sub>\)\ + by (metis "Act-Basic:2" "&E"(1) "\E"(1)) + AOT_thus \PossibleWorld(\<^bold>w\<^sub>\)\ + using "rigid-pw:4"[unvarify x, OF walpha_den, THEN "\E"(1)] + by blast +qed + +AOT_theorem "alpha-world:2": \Maximal(\<^bold>w\<^sub>\)\ +proof - + AOT_have \\<^bold>w\<^sub>\\\ + using "pre-walpha" "rule-id-df:2:b[zero]" "w-alpha" by blast + then AOT_obtain x where x_def: \x = \<^bold>w\<^sub>\\ + by (metis "instantiation" "rule=I:1" "existential:1" id_sym) + AOT_hence \PossibleWorld(x)\ using "alpha-world:1" "rule=E" id_sym by fast + AOT_hence \Maximal(x)\ by (metis "world-max"[unconstrain w, THEN "\E"]) + AOT_thus \Maximal(\<^bold>w\<^sub>\)\ using x_def "rule=E" by blast +qed + +AOT_theorem "t-at-alpha-strict": \\<^bold>w\<^sub>\ \ p \ \<^bold>\p\ +proof - + AOT_have 0: \\<^bold>w\<^sub>\ = \<^bold>\w Actual(w)\ + using "pre-walpha" "rule-id-df:1[zero]" "w-alpha" by blast + AOT_hence walpha_den: \\<^bold>w\<^sub>\\\ + by (metis "t=t-proper:1" "vdash-properties:6") + AOT_have 1: \\<^bold>\(PossibleWorld(\<^bold>w\<^sub>\) & Actual(\<^bold>w\<^sub>\))\ + by (rule "actual-desc:2"[unvarify x, OF walpha_den, THEN "\E"]) (fact 0) + AOT_have walpha_sit: \Situation(\<^bold>w\<^sub>\)\ + by (meson "\\<^sub>d\<^sub>fE" "alpha-world:2" "&E"(1) max) + { + fix p + AOT_have 2: \Situation(x) \ (\<^bold>\x \ p \ x \ p)\ for x + using "lem2:4"[unconstrain s] by blast + AOT_assume \\<^bold>w\<^sub>\ \ p\ + AOT_hence \: \\<^bold>\\<^bold>w\<^sub>\ \ p\ + using 2[unvarify x, OF walpha_den, THEN "\E", OF walpha_sit, THEN "\E"(2)] + by argo + AOT_have 3: \\<^bold>\Actual(\<^bold>w\<^sub>\)\ + using "1" "Act-Basic:2" "&E"(2) "\E"(1) by blast + AOT_have \\<^bold>\(Situation(\<^bold>w\<^sub>\) & \q(\<^bold>w\<^sub>\ \ q \ q))\ + apply (AOT_subst (reverse) \Situation(\<^bold>w\<^sub>\) & \q(\<^bold>w\<^sub>\ \ q \ q)\ \Actual(\<^bold>w\<^sub>\)\) + using actual "\Df" apply blast + by (fact 3) + AOT_hence \\<^bold>\\q(\<^bold>w\<^sub>\ \ q \ q)\ by (metis "Act-Basic:2" "&E"(2) "\E"(1)) + AOT_hence \\q \<^bold>\(\<^bold>w\<^sub>\ \ q \ q)\ + using "logic-actual-nec:3"[axiom_inst, THEN "\E"(1)] by blast + AOT_hence \\<^bold>\(\<^bold>w\<^sub>\ \ p \ p)\ using "\E"(2) by blast + AOT_hence \\<^bold>\(\<^bold>w\<^sub>\ \ p) \ \<^bold>\p\ by (metis "act-cond" "vdash-properties:10") + AOT_hence \\<^bold>\p\ using \ "\E" by blast + } + AOT_hence 2: \\<^bold>w\<^sub>\ \ p \ \<^bold>\p\ for p by (rule "\I") + AOT_have walpha_sit: \Situation(\<^bold>w\<^sub>\)\ + using "\\<^sub>d\<^sub>fE" "alpha-world:2" "&E"(1) max by blast + show ?thesis + proof(safe intro!: "\I" "\I" 2) + AOT_assume actp: \\<^bold>\p\ + AOT_show \\<^bold>w\<^sub>\ \ p\ + proof(rule "raa-cor:1") + AOT_assume \\\<^bold>w\<^sub>\ \ p\ + AOT_hence \\<^bold>w\<^sub>\ \ \p\ + using "alpha-world:2"[THEN max[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(2), + THEN "\E"(1), OF "log-prop-prop:2"] + by (metis "\E"(2)) + AOT_hence \\<^bold>\\p\ + using 2[unvarify p, OF "log-prop-prop:2", THEN "\E"] by blast + AOT_hence \\\<^bold>\p\ by (metis "\\I" "Act-Sub:1" "\E"(4)) + AOT_thus \\<^bold>\p & \\<^bold>\p\ using actp "&I" by blast + qed + qed +qed + +AOT_act_theorem "not-act": \w \ \<^bold>w\<^sub>\ \ \Actual(w)\ +proof (rule "\I"; rule "raa-cor:2") + AOT_assume \w \ \<^bold>w\<^sub>\\ + AOT_hence 0: \\(w = \<^bold>w\<^sub>\)\ by (metis "\\<^sub>d\<^sub>fE" "=-infix") + AOT_have walpha_den: \\<^bold>w\<^sub>\\\ + using "pre-walpha" "rule-id-df:2:b[zero]" "w-alpha" by blast + AOT_have walpha_sit: \Situation(\<^bold>w\<^sub>\)\ + using "\\<^sub>d\<^sub>fE" "alpha-world:2" "&E"(1) max by blast + AOT_assume act_w: \Actual(w)\ + AOT_hence w_sit: \Situation(w)\ by (metis "\\<^sub>d\<^sub>fE" actual "&E"(1)) + AOT_have sid: \Situation(x') \ (w = x' \ \p (w \ p \ x' \ p))\ for x' + using "sit-identity"[unconstrain s', unconstrain s, THEN "\E", OF w_sit] + by blast + AOT_have \w = \<^bold>w\<^sub>\\ + proof(safe intro!: GEN sid[unvarify x', OF walpha_den, THEN "\E", + OF walpha_sit, THEN "\E"(2)] "\I" "\I") + fix p + AOT_assume \w \ p\ + AOT_hence \p\ + using actual[THEN "\\<^sub>d\<^sub>fE", OF act_w, THEN "&E"(2), THEN "\E"(2), THEN "\E"] + by blast + AOT_hence \\<^bold>\p\ + by (metis "RA[1]") + AOT_thus \\<^bold>w\<^sub>\ \ p\ + using "t-at-alpha-strict"[THEN "\E"(2)] by blast + next + fix p + AOT_assume \\<^bold>w\<^sub>\ \ p\ + AOT_hence \\<^bold>\p\ + using "t-at-alpha-strict"[THEN "\E"(1)] by blast + AOT_hence p: \p\ + using "logic-actual"[act_axiom_inst, THEN "\E"] by blast + AOT_show \w \ p\ + proof(rule "raa-cor:1") + AOT_assume \\w \ p\ + AOT_hence \w \ \p\ + by (metis "coherent:1" "\E"(2)) + AOT_hence \\p\ + using actual[THEN "\\<^sub>d\<^sub>fE", OF act_w, THEN "&E"(2), THEN "\E"(1), + OF "log-prop-prop:2", THEN "\E"] by blast + AOT_thus \p & \p\ using p "&I" by blast + qed + qed + AOT_thus \w = \<^bold>w\<^sub>\ & \(w = \<^bold>w\<^sub>\)\ using 0 "&I" by blast +qed + +AOT_act_theorem "w-alpha-part": \Actual(s) \ s \ \<^bold>w\<^sub>\\ +proof(safe intro!: "\I" "\I" "sit-part-whole"[THEN "\\<^sub>d\<^sub>fI"] "&I" GEN + dest!: "sit-part-whole"[THEN "\\<^sub>d\<^sub>fE"]) + AOT_show \Situation(s)\ if \Actual(s)\ + using "\\<^sub>d\<^sub>fE" actual "&E"(1) that by blast +next + AOT_show \Situation(\<^bold>w\<^sub>\)\ + using "\\<^sub>d\<^sub>fE" "alpha-world:2" "&E"(1) max by blast +next + fix p + AOT_assume \Actual(s)\ + moreover AOT_assume \s \ p\ + ultimately AOT_have \p\ + using actual[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2), THEN "\E"(2), THEN "\E"] by blast + AOT_thus \\<^bold>w\<^sub>\ \ p\ + by (metis "\E"(1) "truth-at-alpha:2") +next + AOT_assume 0: \Situation(s) & Situation(\<^bold>w\<^sub>\) & \p (s \ p \ \<^bold>w\<^sub>\ \ p)\ + AOT_hence \s \ p \ \<^bold>w\<^sub>\ \ p\ for p + using "&E" "\E"(2) by blast + AOT_hence \s \ p \ p\ for p + by (metis "\I" "\E"(2) "truth-at-alpha:2" "\E") + AOT_hence \\p (s \ p \ p)\ by (rule GEN) + AOT_thus \Actual(s)\ + using actual[THEN "\\<^sub>d\<^sub>fI", OF "&I", OF 0[THEN "&E"(1), THEN "&E"(1)]] by blast +qed + +AOT_act_theorem "act-world2:1": \\<^bold>w\<^sub>\ \ p \ [\y p]\<^bold>w\<^sub>\\ + apply (AOT_subst \[\y p]\<^bold>w\<^sub>\\ p) + apply (rule "beta-C-meta"[THEN "\E", OF "prop-prop2:2", unvarify \\<^sub>1\\<^sub>n]) + using "pre-walpha" "rule-id-df:2:b[zero]" "w-alpha" apply blast + using "\E"(2) "Commutativity of \" "truth-at-alpha:2" by blast + +AOT_act_theorem "act-world2:2": \p \ \<^bold>w\<^sub>\ \ [\y p]\<^bold>w\<^sub>\\ +proof - + AOT_have \p \ [\y p]\<^bold>w\<^sub>\\ + apply (rule "beta-C-meta"[THEN "\E", OF "prop-prop2:2", + unvarify \\<^sub>1\\<^sub>n, symmetric]) + using "pre-walpha" "rule-id-df:2:b[zero]" "w-alpha" by blast + also AOT_have \\ \ \<^bold>w\<^sub>\ \ [\y p]\<^bold>w\<^sub>\\ + by (meson "log-prop-prop:2" "rule-ui:1" "truth-at-alpha:2" "universal-cor") + finally show ?thesis. +qed + +AOT_theorem "fund-lem:1": \\p \ \\w (w \ p)\ +proof (rule "RM\"; rule "\I"; rule "raa-cor:1") + AOT_modally_strict { + AOT_obtain w where w_prop: \\q (w \ q \ q)\ + using "act-world:1" "PossibleWorld.\E"[rotated] by meson + AOT_assume p: \p\ + AOT_assume 0: \\\w (w \ p)\ + AOT_have \\w \(w \ p)\ + apply (AOT_subst \PossibleWorld(x) \ \x \ p\ + \\(PossibleWorld(x) & x \ p)\ for: x) + apply (metis "&I" "&E"(1) "&E"(2) "\I" "\I" "modus-tollens:2") + using "0" "cqt-further:4" "vdash-properties:10" by blast + AOT_hence \\(w \ p)\ + using PossibleWorld.\ "rule-ui:3" "\E" by blast + AOT_hence \\p\ + using w_prop[THEN "\E"(2), THEN "\E"(2)] + by (metis "raa-cor:3") + AOT_thus \p & \p\ + using p "&I" by blast + } +qed + +AOT_theorem "fund-lem:2": \\\w (w \ p) \ \w (w \ p)\ +proof (rule "\I") + AOT_assume \\\w (w \ p)\ + AOT_hence \\w \(w \ p)\ + using "PossibleWorld.res-var-bound-reas[BF\]"[THEN "\E"] by auto + then AOT_obtain w where \\(w \ p)\ + using "PossibleWorld.\E"[rotated] by meson + moreover AOT_have \Situation(w)\ + by (metis "\\<^sub>d\<^sub>fE" "&E"(1) pos "world-pos") + ultimately AOT_have \w \ p\ + using "lem2:2"[unconstrain s, THEN "\E"] "\E" by blast + AOT_thus \\w w \ p\ + by (rule "PossibleWorld.\I") +qed + +AOT_theorem "fund-lem:3": \p \ \s(\q (s \ q \ q) \ s \ p)\ +proof(safe intro!: "\I" Situation.GEN) + fix s + AOT_assume \p\ + moreover AOT_assume \\q (s \ q \ q)\ + ultimately AOT_show \s \ p\ + using "\E"(2) "\E"(2) by blast +qed + +AOT_theorem "fund-lem:4": \\p \ \\s(\q (s \ q \ q) \ s \ p)\ + using "fund-lem:3" by (rule RM) + +AOT_theorem "fund-lem:5": \\\s \{s} \ \s \\{s}\ +proof(safe intro!: "\I" Situation.GEN) + fix s + AOT_assume \\\s \{s}\ + AOT_hence \\s \\{s}\ + using "Situation.res-var-bound-reas[CBF]"[THEN "\E"] by blast + AOT_thus \\\{s}\ + using "Situation.\E" by fast +qed + +text\Note: not explicit in PLM.\ +AOT_theorem "fund-lem:5[world]": \\\w \{w} \ \w \\{w}\ +proof(safe intro!: "\I" PossibleWorld.GEN) + fix w + AOT_assume \\\w \{w}\ + AOT_hence \\w \\{w}\ + using "PossibleWorld.res-var-bound-reas[CBF]"[THEN "\E"] by blast + AOT_thus \\\{w}\ + using "PossibleWorld.\E" by fast +qed + +AOT_theorem "fund-lem:6": \\w w \ p \ \\w w \ p\ +proof(rule "\I") + AOT_assume \\w (w \ p)\ + AOT_hence 1: \PossibleWorld(w) \ (w \ p)\ for w + using "\E"(2) by blast + AOT_show \\\w w \ p\ + proof(rule "raa-cor:1") + AOT_assume \\\\w w \ p\ + AOT_hence \\\\w w \ p\ + by (metis "KBasic:11" "\E"(1)) + AOT_hence \\\x (\(PossibleWorld(x) \ x \ p))\ + apply (rule "RM\"[THEN "\E", rotated]) + by (simp add: "cqt-further:2") + AOT_hence \\x \(\(PossibleWorld(x) \ x \ p))\ + by (metis "BF\" "vdash-properties:10") + then AOT_obtain x where x_prop: \\\(PossibleWorld(x) \ x \ p)\ + using "\E"[rotated] by blast + AOT_have \\(PossibleWorld(x) & \x \ p)\ + apply (AOT_subst \PossibleWorld(x) & \x \ p\ + \\(PossibleWorld(x) \ x \ p)\) + apply (meson "\E"(6) "oth-class-taut:1:b" "oth-class-taut:3:a") + by(fact x_prop) + AOT_hence 2: \\PossibleWorld(x) & \\x \ p\ + by (metis "KBasic2:3" "vdash-properties:10") + AOT_hence \PossibleWorld(x)\ + using "&E"(1) "\E"(1) "rigid-pw:2" by blast + AOT_hence \\(x \ p)\ + using 2[THEN "&E"(2)] 1[unconstrain w, THEN "\E"] "\E" + "rigid-truth-at:1"[unconstrain w, THEN "\E"] + by (metis "\E"(1)) + moreover AOT_have \\\(x \ p)\ + using 2[THEN "&E"(2)] by (metis "\\I" "KBasic:12" "\E"(4)) + ultimately AOT_show \p & \p\ for p + by (metis "raa-cor:3") + qed +qed + +AOT_theorem "fund-lem:7": \\\w(w \ p) \ \p\ +proof(rule RM; rule "\I") + AOT_modally_strict { + AOT_obtain w where w_prop: \\p (w \ p \ p)\ + using "act-world:1" "PossibleWorld.\E"[rotated] by meson + AOT_assume \\w (w \ p)\ + AOT_hence \w \ p\ + using "PossibleWorld.\E" by fast + AOT_thus \p\ + using w_prop[THEN "\E"(2), THEN "\E"(1)] by blast + } +qed + +AOT_theorem "fund:1": \\p \ \w w \ p\ +proof (rule "\I"; rule "\I") + AOT_assume \\p\ + AOT_thus \\w w \ p\ + by (metis "fund-lem:1" "fund-lem:2" "\E") +next + AOT_assume \\w w \ p\ + then AOT_obtain w where w_prop: \w \ p\ + using "PossibleWorld.\E"[rotated] by meson + AOT_hence \\\p (w \ p \ p)\ + using "world:1"[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2)] PossibleWorld.\ "&E" by blast + AOT_hence \\p \(w \ p \ p)\ + by (metis "Buridan\" "\E") + AOT_hence 1: \\(w \ p \ p)\ + by (metis "log-prop-prop:2" "rule-ui:1") + AOT_have \\((w \ p \ p) & (p \ w \ p))\ + apply (AOT_subst \(w \ p \ p) & (p \ w \ p)\ \w \ p \ p\) + apply (meson "conventions:3" "\E"(6) "oth-class-taut:3:a" "\Df") + by (fact 1) + AOT_hence \\(w \ p \ p)\ + by (metis "RM\" "Conjunction Simplification"(1) "\E") + moreover AOT_have \\(w \ p)\ + using w_prop by (metis "\E"(1) "rigid-truth-at:1") + ultimately AOT_show \\p\ + by (metis "KBasic2:4" "\E"(1) "\E") +qed + +AOT_theorem "fund:2": \\p \ \w (w \ p)\ +proof - + AOT_have 0: \\w (w \ \p \ \w \ p)\ + apply (rule PossibleWorld.GEN) + using "coherent:1" by blast + AOT_have \\\p \ \w (w \ \p)\ + using "fund:1"[unvarify p, OF "log-prop-prop:2"] by blast + also AOT_have \\ \ \w \(w \ p)\ + proof(safe intro!: "\I" "\I") + AOT_assume \\w w \ \p\ + then AOT_obtain w where w_prop: \w \ \p\ + using "PossibleWorld.\E"[rotated] by meson + AOT_hence \\w \ p\ + using 0[THEN "PossibleWorld.\E", THEN "\E"(1)] "&E" by blast + AOT_thus \\w \w \ p\ + by (rule "PossibleWorld.\I") + next + AOT_assume \\w \w \ p\ + then AOT_obtain w where w_prop: \\w \ p\ + using "PossibleWorld.\E"[rotated] by meson + AOT_hence \w \ \p\ + using 0[THEN "\E"(2), THEN "\E", THEN "\E"(1)] "&E" + by (metis "coherent:1" "\E"(2)) + AOT_thus \\w w \ \p\ + by (rule "PossibleWorld.\I") + qed + finally AOT_have \\\\p \ \\w \w \ p\ + by (meson "\E"(1) "oth-class-taut:4:b") + AOT_hence \\p \ \\w \w \ p\ + by (metis "KBasic:12" "\E"(5)) + also AOT_have \\ \ \w w \ p\ + proof(safe intro!: "\I" "\I") + AOT_assume \\\w \w \ p\ + AOT_hence 0: \\x (\(PossibleWorld(x) & \x \ p))\ + by (metis "cqt-further:4" "\E") + AOT_show \\w w \ p\ + apply (AOT_subst \PossibleWorld(x) \ x \ p\ + \\(PossibleWorld(x) & \x \ p)\ for: x) + using "oth-class-taut:1:a" apply presburger + by (fact 0) + next + AOT_assume 0: \\w w \ p\ + AOT_have \\x (\(PossibleWorld(x) & \x \ p))\ + by (AOT_subst (reverse) \\(PossibleWorld(x) & \x \ p)\ + \PossibleWorld(x) \ x \ p\ for: x) + (auto simp: "oth-class-taut:1:a" 0) + AOT_thus \\\w \w \ p\ + by (metis "\E" "raa-cor:3" "rule-ui:3") + qed + finally AOT_show \\p \ \w w \ p\. +qed + +AOT_theorem "fund:3": \\\p \ \\w w \ p\ + by (metis (full_types) "contraposition:1[1]" "\I" "fund:1" "\I" "\E"(1,2)) + +AOT_theorem "fund:4": \\\p \ \w \w \p\ + apply (AOT_subst \\w \w \ p\ \\ \w w \ p\) + apply (AOT_subst \PossibleWorld(x) \ x \ p\ + \\(PossibleWorld(x) & \x \ p)\ for: x) + by (auto simp add: "oth-class-taut:1:a" "conventions:4" "\Df" RN + "fund:2" "rule-sub-lem:1:a") + +AOT_theorem "nec-dia-w:1": \\p \ \w w \ \p\ +proof - + AOT_have \\p \ \\p\ + using "S5Basic:2" by blast + also AOT_have \... \ \w w \ \p\ + using "fund:1"[unvarify p, OF "log-prop-prop:2"] by blast + finally show ?thesis. +qed + +AOT_theorem "nec-dia-w:2": \\p \ \w w \ \p\ +proof - + AOT_have \\p \ \\p\ + using 4 "qml:2"[axiom_inst] "\I" by blast + also AOT_have \... \ \w w \ \p\ + using "fund:2"[unvarify p, OF "log-prop-prop:2"] by blast + finally show ?thesis. +qed + +AOT_theorem "nec-dia-w:3": \\p \ \w w \ \p\ +proof - + AOT_have \\p \ \\p\ + by (simp add: "4\" "T\" "\I") + also AOT_have \... \ \w w \ \p\ + using "fund:1"[unvarify p, OF "log-prop-prop:2"] by blast + finally show ?thesis. +qed + +AOT_theorem "nec-dia-w:4": \\p \ \w w \ \p\ +proof - + AOT_have \\p \ \\p\ + by (simp add: "S5Basic:1") + also AOT_have \... \ \w w \ \p\ + using "fund:2"[unvarify p, OF "log-prop-prop:2"] by blast + finally show ?thesis. +qed + +AOT_theorem "conj-dist-w:1": \w \ (p & q) \ ((w \ p) & (w \ q))\ +proof(safe intro!: "\I" "\I") + AOT_assume \w \ (p & q)\ + AOT_hence 0: \\w \ (p & q)\ + using "rigid-truth-at:1"[unvarify p, THEN "\E"(1), OF "log-prop-prop:2"] + by blast + AOT_modally_strict { + AOT_have \\p (w \ p \ p) \ ((w \ (\ & \)) \ (w \ \ & w \ \))\ for w \ \ + proof(safe intro!: "\I") + AOT_assume \\ p (w \ p \ p)\ + AOT_hence \w \ (\ & \) \ (\ & \)\ and \w \ \ \ \\ and \w \ \ \ \\ + using "\E"(1)[rotated, OF "log-prop-prop:2"] by blast+ + moreover AOT_assume \w \ (\ & \)\ + ultimately AOT_show \w \ \ & w \ \\ + by (metis "&I" "&E"(1) "&E"(2) "\E"(1) "\E"(2)) + qed + } + AOT_hence \\\p (w \ p \ p) \ \(w \ (\ & \) \ w \ \ & w \ \)\ for w \ \ + by (rule "RM\") + moreover AOT_have pos: \\\p (w \ p \ p)\ + using "world:1"[THEN "\\<^sub>d\<^sub>fE", OF PossibleWorld.\] "&E" by blast + ultimately AOT_have \\(w \ (p & q) \ w \ p & w \ q)\ using "\E" by blast + AOT_hence \\(w \ p) & \(w \ q)\ + by (metis 0 "KBasic2:3" "KBasic2:4" "\E"(1) "vdash-properties:10") + AOT_thus \w \ p & w \ q\ + using "rigid-truth-at:2"[unvarify p, THEN "\E"(1), OF "log-prop-prop:2"] + "&E" "&I" by meson +next + AOT_assume \w \ p & w \ q\ + AOT_hence \\w \ p & \w \ q\ + using "rigid-truth-at:1"[unvarify p, THEN "\E"(1), OF "log-prop-prop:2"] + "&E" "&I" by blast + AOT_hence 0: \\(w \ p & w \ q)\ + by (metis "KBasic:3" "\E"(2)) + AOT_modally_strict { + AOT_have \\p (w \ p \ p) \ ((w \ \ & w \ \) \ (w \ (\ & \)))\ for w \ \ + proof(safe intro!: "\I") + AOT_assume \\ p (w \ p \ p)\ + AOT_hence \w \ (\ & \) \ (\ & \)\ and \w \ \ \ \\ and \w \ \ \ \\ + using "\E"(1)[rotated, OF "log-prop-prop:2"] by blast+ + moreover AOT_assume \w \ \ & w \ \\ + ultimately AOT_show \w \ (\ & \)\ + by (metis "&I" "&E"(1) "&E"(2) "\E"(1) "\E"(2)) + qed + } + AOT_hence \\\p (w \ p \ p) \ \((w \ \ & w \ \) \ w \ (\ & \))\ for w \ \ + by (rule "RM\") + moreover AOT_have pos: \\\p (w \ p \ p)\ + using "world:1"[THEN "\\<^sub>d\<^sub>fE", OF PossibleWorld.\] "&E" by blast + ultimately AOT_have \\((w \ p & w \ q) \ w \ (p & q))\ + using "\E" by blast + AOT_hence \\(w \ (p & q))\ + by (metis 0 "KBasic2:4" "\E"(1) "vdash-properties:10") + AOT_thus \w \ (p & q)\ + using "rigid-truth-at:2"[unvarify p, THEN "\E"(1), OF "log-prop-prop:2"] + by blast +qed + +AOT_theorem "conj-dist-w:2": \w \ (p \ q) \ ((w \ p) \ (w \ q))\ +proof(safe intro!: "\I" "\I") + AOT_assume \w \ (p \ q)\ + AOT_hence 0: \\w \ (p \ q)\ + using "rigid-truth-at:1"[unvarify p, THEN "\E"(1), OF "log-prop-prop:2"] + by blast + AOT_assume \w \ p\ + AOT_hence 1: \\w \ p\ + by (metis "T\" "\E"(1) "rigid-truth-at:3" "\E") + AOT_modally_strict { + AOT_have \\p (w \ p \ p) \ ((w \ (\ \ \)) \ (w \ \ \ w \ \))\ for w \ \ + proof(safe intro!: "\I") + AOT_assume \\ p (w \ p \ p)\ + AOT_hence \w \ (\ \ \) \ (\ \ \)\ and \w \ \ \ \\ and \w \ \ \ \\ + using "\E"(1)[rotated, OF "log-prop-prop:2"] by blast+ + moreover AOT_assume \w \ (\ \ \)\ + moreover AOT_assume \w \ \\ + ultimately AOT_show \w \ \\ + by (metis "\E"(1) "\E"(2) "\E") + qed + } + AOT_hence \\\p (w \ p \ p) \ \(w \ (\ \ \) \ (w \ \ \ w \ \))\ for w \ \ + by (rule "RM\") + moreover AOT_have pos: \\\p (w \ p \ p)\ + using "world:1"[THEN "\\<^sub>d\<^sub>fE", OF PossibleWorld.\] "&E" by blast + ultimately AOT_have \\(w \ (p \ q) \ (w \ p \ w \ q))\ + using "\E" by blast + AOT_hence \\(w \ p \ w \ q)\ + by (metis 0 "KBasic2:4" "\E"(1) "\E") + AOT_hence \\w \ q\ + by (metis 1 "KBasic2:4" "\E"(1) "\E") + AOT_thus \w \ q\ + using "rigid-truth-at:2"[unvarify p, THEN "\E"(1), OF "log-prop-prop:2"] + "&E" "&I" by meson +next + AOT_assume \w \ p \ w \ q\ + AOT_hence \\(w \ p) \ w \ q\ + by (metis "\I"(1) "\I"(2) "reductio-aa:1" "\E") + AOT_hence \w \ \p \ w \ q\ + by (metis "coherent:1" "\I"(1) "\I"(2) "\E"(2) "\E"(2) "reductio-aa:1") + AOT_hence 0: \\(w \ \p \ w \ q)\ + using "rigid-truth-at:1"[unvarify p, THEN "\E"(1), OF "log-prop-prop:2"] + by (metis "KBasic:15" "\I"(1) "\I"(2) "\E"(2) "reductio-aa:1" "\E") + AOT_modally_strict { + AOT_have \\p (w \ p \ p) \ ((w \ \\ \ w \ \) \ (w \ (\ \ \)))\ for w \ \ + proof(safe intro!: "\I") + AOT_assume \\ p (w \ p \ p)\ + moreover AOT_assume \w \ \\ \ w \ \\ + ultimately AOT_show \w \ (\ \ \)\ + by (metis "\E"(2) "\I" "\E"(1) "\E"(2) "log-prop-prop:2" + "reductio-aa:1" "rule-ui:1") + qed + } + AOT_hence \\\p (w \ p \ p) \ \((w \ \\ \ w \ \) \ w \ (\ \ \))\ for w \ \ + by (rule "RM\") + moreover AOT_have pos: \\\p (w \ p \ p)\ + using "world:1"[THEN "\\<^sub>d\<^sub>fE", OF PossibleWorld.\] "&E" by blast + ultimately AOT_have \\((w \ \p \ w \ q) \ w \ (p \ q))\ + using "\E" by blast + AOT_hence \\(w \ (p \ q))\ + by (metis 0 "KBasic2:4" "\E"(1) "\E") + AOT_thus \w \ (p \ q)\ + using "rigid-truth-at:2"[unvarify p, THEN "\E"(1), OF "log-prop-prop:2"] + by blast +qed + +AOT_theorem "conj-dist-w:3": \w \ (p \ q) \ ((w \ p) \ (w \ q))\ +proof(safe intro!: "\I" "\I") + AOT_assume \w \ (p \ q)\ + AOT_hence 0: \\w \ (p \ q)\ + using "rigid-truth-at:1"[unvarify p, THEN "\E"(1), OF "log-prop-prop:2"] + by blast + AOT_modally_strict { + AOT_have \\p (w \ p \ p) \ ((w \ (\ \ \)) \ (w \ \ \ w \ \))\ for w \ \ + proof(safe intro!: "\I") + AOT_assume \\ p (w \ p \ p)\ + AOT_hence \w \ (\ \ \) \ (\ \ \)\ and \w \ \ \ \\ and \w \ \ \ \\ + using "\E"(1)[rotated, OF "log-prop-prop:2"] by blast+ + moreover AOT_assume \w \ (\ \ \)\ + ultimately AOT_show \w \ \ \ w \ \\ + by (metis "\I"(1) "\I"(2) "\E"(3) "\E"(1) "\E"(2) "reductio-aa:1") + qed + } + AOT_hence \\\p (w \ p \ p) \ \(w \ (\ \ \) \ (w \ \ \ w \ \))\ for w \ \ + by (rule "RM\") + moreover AOT_have pos: \\\p (w \ p \ p)\ + using "world:1"[THEN "\\<^sub>d\<^sub>fE", OF PossibleWorld.\] "&E" by blast + ultimately AOT_have \\(w \ (p \ q) \ (w \ p \ w \ q))\ using "\E" by blast + AOT_hence \\(w \ p \ w \ q)\ + by (metis 0 "KBasic2:4" "\E"(1) "vdash-properties:10") + AOT_hence \\w \ p \ \w \ q\ + using "KBasic2:2"[THEN "\E"(1)] by blast + AOT_thus \w \ p \ w \ q\ + using "rigid-truth-at:2"[unvarify p, THEN "\E"(1), OF "log-prop-prop:2"] + by (metis "\I"(1) "\I"(2) "\E"(2) "reductio-aa:1") +next + AOT_assume \w \ p \ w \ q\ + AOT_hence 0: \\(w \ p \ w \ q)\ + using "rigid-truth-at:1"[unvarify p, THEN "\E"(1), OF "log-prop-prop:2"] + by (metis "KBasic:15" "\I"(1) "\I"(2) "\E"(2) "reductio-aa:1" "\E") + AOT_modally_strict { + AOT_have \\p (w \ p \ p) \ ((w \ \ \ w \ \) \ (w \ (\ \ \)))\ for w \ \ + proof(safe intro!: "\I") + AOT_assume \\ p (w \ p \ p)\ + moreover AOT_assume \w \ \ \ w \ \\ + ultimately AOT_show \w \ (\ \ \)\ + by (metis "\I"(1) "\I"(2) "\E"(2) "\E"(1) "\E"(2) + "log-prop-prop:2" "reductio-aa:1" "rule-ui:1") + qed + } + AOT_hence \\\p (w \ p \ p) \ \((w \ \ \ w \ \) \ w \ (\ \ \))\ for w \ \ + by (rule "RM\") + moreover AOT_have pos: \\\p (w \ p \ p)\ + using "world:1"[THEN "\\<^sub>d\<^sub>fE", OF PossibleWorld.\] "&E" by blast + ultimately AOT_have \\((w \ p \ w \ q) \ w \ (p \ q))\ + using "\E" by blast + AOT_hence \\(w \ (p \ q))\ + by (metis 0 "KBasic2:4" "\E"(1) "\E") + AOT_thus \w \ (p \ q)\ + using "rigid-truth-at:2"[unvarify p, THEN "\E"(1), OF "log-prop-prop:2"] + by blast +qed + +AOT_theorem "conj-dist-w:4": \w \ (p \ q) \ ((w \ p) \ (w \ q))\ +proof(rule "\I"; rule "\I") + AOT_assume \w \ (p \ q)\ + AOT_hence 0: \\w \ (p \ q)\ + using "rigid-truth-at:1"[unvarify p, THEN "\E"(1), OF "log-prop-prop:2"] + by blast + AOT_modally_strict { + AOT_have \\p (w \ p \ p) \ ((w \ (\ \ \)) \ (w \ \ \ w \ \))\ for w \ \ + proof(safe intro!: "\I") + AOT_assume \\ p (w \ p \ p)\ + AOT_hence \w \ (\ \ \) \ (\ \ \)\ and \w \ \ \ \\ and \w \ \ \ \\ + using "\E"(1)[rotated, OF "log-prop-prop:2"] by blast+ + moreover AOT_assume \w \ (\ \ \)\ + ultimately AOT_show \w \ \ \ w \ \\ + by (metis "\E"(2) "\E"(5) "Commutativity of \") + qed + } + AOT_hence \\\p (w \ p \ p) \ \(w \ (\ \ \) \ (w \ \ \ w \ \))\ for w \ \ + by (rule "RM\") + moreover AOT_have pos: \\\p (w \ p \ p)\ + using "world:1"[THEN "\\<^sub>d\<^sub>fE", OF PossibleWorld.\] "&E" by blast + ultimately AOT_have \\(w \ (p \ q) \ (w \ p \ w \ q))\ + using "\E" by blast + AOT_hence 1: \\(w \ p \ w \ q)\ + by (metis 0 "KBasic2:4" "\E"(1) "vdash-properties:10") + AOT_have \\((w \ p \ w \ q) & (w \ q \ w \ p))\ + apply (AOT_subst \(w \ p \ w \ q) & (w \ q \ w \ p)\ \w \ p \ w \ q\) + apply (meson "\\<^sub>d\<^sub>fE" "conventions:3" "\I" "df-rules-formulas[4]" "\I") + by (fact 1) + AOT_hence 2: \\(w \ p \ w \ q) & \(w \ q \ w \ p)\ + by (metis "KBasic2:3" "vdash-properties:10") + AOT_have \\(\w \ p \ w \ q)\ and \\(\w \ q \ w \ p)\ + apply (AOT_subst (reverse) \\w \ p \ w \ q\ \w \ p \ w \ q\) + apply (simp add: "oth-class-taut:1:c") + apply (fact 2[THEN "&E"(1)]) + apply (AOT_subst (reverse) \\w \ q \ w \ p\ \w \ q \ w \ p\) + apply (simp add: "oth-class-taut:1:c") + by (fact 2[THEN "&E"(2)]) + AOT_hence \\(\w \ p) \ \w \ q\ and \\\w \ q \ \w \ p\ + using "KBasic2:2" "\E"(1) by blast+ + AOT_hence \\\w \ p \ \w \ q\ and \\\w \ q \ \w \ p\ + by (metis "KBasic:11" "\I"(1) "\I"(2) "\E"(2) "\E"(2) "raa-cor:1")+ + AOT_thus \w \ p \ w \ q\ + using "rigid-truth-at:2"[unvarify p, THEN "\E"(1), OF "log-prop-prop:2"] + by (metis "\\I" "T\" "\E"(2) "\I" "\I" "\E"(1) "rigid-truth-at:3") +next + AOT_have \\PossibleWorld(w)\ + using "\E"(1) "rigid-pw:1" PossibleWorld.\ by blast + moreover { + fix p + AOT_modally_strict { + AOT_have \PossibleWorld(w) \ (w \ p \ \w \ p)\ + using "rigid-truth-at:1" "\I" + by (metis "\E"(1)) + } + AOT_hence \\PossibleWorld(w) \ \(w \ p \ \w \ p)\ + by (rule RM) + } + ultimately AOT_have 1: \\(w \ p \ \w \ p)\ for p + by (metis "\E") + AOT_assume \w \ p \ w \ q\ + AOT_hence 0: \\(w \ p \ w \ q)\ + using "sc-eq-box-box:5"[THEN "\E", THEN "qml:2"[axiom_inst, THEN "\E"], + THEN "\E", OF "&I"] + by (metis "1") + AOT_modally_strict { + AOT_have \\p (w \ p \ p) \ ((w \ \ \ w \ \) \ (w \ (\ \ \)))\ for w \ \ + proof(safe intro!: "\I") + AOT_assume \\ p (w \ p \ p)\ + moreover AOT_assume \w \ \ \ w \ \\ + ultimately AOT_show \w \ (\ \ \)\ + by (metis "\E"(2) "\E"(6) "log-prop-prop:2" "rule-ui:1") + qed + } + AOT_hence \\\p (w \ p \ p) \ \((w \ \ \ w \ \) \ w \ (\ \ \))\ for w \ \ + by (rule "RM\") + moreover AOT_have pos: \\\p (w \ p \ p)\ + using "world:1"[THEN "\\<^sub>d\<^sub>fE", OF PossibleWorld.\] "&E" by blast + ultimately AOT_have \\((w \ p \ w \ q) \ w \ (p \ q))\ + using "\E" by blast + AOT_hence \\(w \ (p \ q))\ + by (metis 0 "KBasic2:4" "\E"(1) "\E") + AOT_thus \w \ (p \ q)\ + using "rigid-truth-at:2"[unvarify p, THEN "\E"(1), OF "log-prop-prop:2"] + by blast +qed + +AOT_theorem "conj-dist-w:5": \w \ (\\ \{\}) \ (\ \ (w \ \{\}))\ +proof(safe intro!: "\I" "\I" GEN) + AOT_assume \w \ (\\ \{\})\ + AOT_hence 0: \\w \ (\\ \{\})\ + using "rigid-truth-at:1"[unvarify p, THEN "\E"(1), OF "log-prop-prop:2"] + by blast + AOT_modally_strict { + AOT_have \\p (w \ p \ p) \ ((w \ (\\ \{\})) \ (\\ w \ \{\}))\ for w + proof(safe intro!: "\I" GEN) + AOT_assume \\p (w \ p \ p)\ + moreover AOT_assume \w \ (\\ \{\})\ + ultimately AOT_show \w \ \{\}\ for \ + by (metis "\E"(1) "\E"(2) "log-prop-prop:2" "rule-ui:1" "rule-ui:3") + qed + } + AOT_hence \\\p (w \ p \ p) \ \(w \ (\\ \{\}) \ (\\ w \ \{\}))\ for w + by (rule "RM\") + moreover AOT_have pos: \\\p (w \ p \ p)\ + using "world:1"[THEN "\\<^sub>d\<^sub>fE", OF PossibleWorld.\] "&E" by blast + ultimately AOT_have \\(w \ (\\ \{\}) \ (\\ w \ \{\}))\ using "\E" by blast + AOT_hence \\(\\ w \ \{\})\ + by (metis 0 "KBasic2:4" "\E"(1) "\E") + AOT_hence \\\ \w \ \{\}\ + by (metis "Buridan\" "\E") + AOT_thus \w \ \{\}\ for \ + using "rigid-truth-at:2"[unvarify p, THEN "\E"(1), OF "log-prop-prop:2"] + "\E"(2) by blast +next + AOT_assume \\\ w \ \{\}\ + AOT_hence \w \ \{\}\ for \ using "\E"(2) by blast + AOT_hence \\w \ \{\}\ for \ + using "rigid-truth-at:1"[unvarify p, THEN "\E"(1), OF "log-prop-prop:2"] + "&E" "&I" by blast + AOT_hence \\\ \w \ \{\}\ by (rule GEN) + AOT_hence 0: \\\\ w \ \{\}\ by (rule BF[THEN "\E"]) + AOT_modally_strict { + AOT_have \\p (w \ p \ p) \ ((\\ w \ \{\}) \ (w \ (\\ \{\})))\ for w + proof(safe intro!: "\I") + AOT_assume \\ p (w \ p \ p)\ + moreover AOT_assume \\\ w \ \{\}\ + ultimately AOT_show \w \ (\\ \{\})\ + by (metis "\E"(1) "\E"(2) "log-prop-prop:2" "rule-ui:1" + "rule-ui:3" "universal-cor") + qed + } + AOT_hence \\\p (w \ p \ p) \ \((\\ w \ \{\}) \ w \ (\\ \{\}))\ for w + by (rule "RM\") + moreover AOT_have pos: \\\p (w \ p \ p)\ + using "world:1"[THEN "\\<^sub>d\<^sub>fE", OF PossibleWorld.\] "&E" by blast + ultimately AOT_have \\((\\ w \ \{\}) \ w \ (\\ \{\}))\ + using "\E" by blast + AOT_hence \\(w \ (\\ \{\}))\ + by (metis 0 "KBasic2:4" "\E"(1) "\E") + AOT_thus \w \ (\\ \{\})\ + using "rigid-truth-at:2"[unvarify p, THEN "\E"(1), OF "log-prop-prop:2"] + by blast +qed + +AOT_theorem "conj-dist-w:6": \w \ (\\ \{\}) \ (\ \ (w \ \{\}))\ +proof(safe intro!: "\I" "\I" GEN) + AOT_assume \w \ (\\ \{\})\ + AOT_hence 0: \\w \ (\\ \{\})\ + using "rigid-truth-at:1"[unvarify p, THEN "\E"(1), OF "log-prop-prop:2"] + by blast + AOT_modally_strict { + AOT_have \\p (w \ p \ p) \ ((w \ (\\ \{\})) \ (\\ w \ \{\}))\ for w + proof(safe intro!: "\I" GEN) + AOT_assume \\p (w \ p \ p)\ + moreover AOT_assume \w \ (\\ \{\})\ + ultimately AOT_show \\ \ (w \ \{\})\ + by (metis "\E" "\I"(2) "\E"(1,2) "log-prop-prop:2" "rule-ui:1") + qed + } + AOT_hence \\\p (w \ p \ p) \ \(w \ (\\ \{\}) \ (\\ w \ \{\}))\ for w + by (rule "RM\") + moreover AOT_have pos: \\\p (w \ p \ p)\ + using "world:1"[THEN "\\<^sub>d\<^sub>fE", OF PossibleWorld.\] "&E" by blast + ultimately AOT_have \\(w \ (\\ \{\}) \ (\\ w \ \{\}))\ using "\E" by blast + AOT_hence \\(\\ w \ \{\})\ + by (metis 0 "KBasic2:4" "\E"(1) "\E") + AOT_hence \\\ \w \ \{\}\ + by (metis "BF\" "\E") + then AOT_obtain \ where \\w \ \{\}\ + using "\E"[rotated] by blast + AOT_hence \w \ \{\}\ + using "rigid-truth-at:2"[unvarify p, THEN "\E"(1), OF "log-prop-prop:2"] by blast + AOT_thus \\ \ w \ \{\}\ by (rule "\I") +next + AOT_assume \\\ w \ \{\}\ + then AOT_obtain \ where \w \ \{\}\ using "\E"[rotated] by blast + AOT_hence \\w \ \{\}\ + using "rigid-truth-at:1"[unvarify p, THEN "\E"(1), OF "log-prop-prop:2"] + "&E" "&I" by blast + AOT_hence \\\ \w \ \{\}\ + by (rule "\I") + AOT_hence 0: \\\\ w \ \{\}\ + by (metis Buridan "\E") + AOT_modally_strict { + AOT_have \\p (w \ p \ p) \ ((\\ w \ \{\}) \ (w \ (\\ \{\})))\ for w + proof(safe intro!: "\I") + AOT_assume \\ p (w \ p \ p)\ + moreover AOT_assume \\\ w \ \{\}\ + then AOT_obtain \ where \w \ \{\}\ + using "\E"[rotated] by blast + ultimately AOT_show \w \ (\\ \{\})\ + by (metis "\I"(2) "\E"(1,2) "log-prop-prop:2" "rule-ui:1") + qed + } + AOT_hence \\\p (w \ p \ p) \ \((\\ w \ \{\}) \ w \ (\\ \{\}))\ for w + by (rule "RM\") + moreover AOT_have pos: \\\p (w \ p \ p)\ + using "world:1"[THEN "\\<^sub>d\<^sub>fE", OF PossibleWorld.\] "&E" by blast + ultimately AOT_have \\((\\ w \ \{\}) \ w \ (\\ \{\}))\ + using "\E" by blast + AOT_hence \\(w \ (\\ \{\}))\ + by (metis 0 "KBasic2:4" "\E"(1) "\E") + AOT_thus \w \ (\\ \{\})\ + using "rigid-truth-at:2"[unvarify p, THEN "\E"(1), OF "log-prop-prop:2"] + by blast +qed + +AOT_theorem "conj-dist-w:7": \(w \ \p) \ \w \ p\ +proof(rule "\I") + AOT_assume \w \ \p\ + AOT_hence \\w w \ \p\ by (rule "PossibleWorld.\I") + AOT_hence \\\p\ using "fund:1"[unvarify p, OF "log-prop-prop:2", THEN "\E"(2)] + by blast + AOT_hence \\p\ + by (metis "5\" "\E") + AOT_hence 1: \\\p\ + by (metis "S5Basic:6" "\E"(1)) + AOT_have \\\w w \ p\ + by (AOT_subst (reverse) \\w w \ p\ \\p\) + (auto simp add: "fund:2" 1) + AOT_hence \\w \w \ p\ + using "fund-lem:5[world]"[THEN "\E"] by simp + AOT_thus \\w \ p\ + using "\E" "PossibleWorld.\E" by fast +qed + +AOT_theorem "conj-dist-w:8": \\w\p((\w \ p) & \w \ \p)\ +proof - + AOT_obtain r where A: r and \\\r\ + by (metis "&E"(1) "&E"(2) "\\<^sub>d\<^sub>fE" "\E" "cont-tf:1" "cont-tf-thm:1") + AOT_hence B: \\\r\ + by (metis "KBasic:11" "\E"(2)) + AOT_have \\r\ + using A "T\"[THEN "\E"] by simp + AOT_hence \\w w \ r\ + using "fund:1"[THEN "\E"(1)] by blast + then AOT_obtain w where w: \w \ r\ + using "PossibleWorld.\E"[rotated] by meson + AOT_hence \\w \ r\ + by (metis "T\" "\E"(1) "rigid-truth-at:3" "vdash-properties:10") + moreover AOT_have \\w \ \r\ + proof(rule "raa-cor:2") + AOT_assume \w \ \r\ + AOT_hence \\w w \ \r\ + by (rule "PossibleWorld.\I") + AOT_hence \\r\ + by (metis "\E"(2) "nec-dia-w:1") + AOT_thus \\r & \\r\ + using B "&I" by blast + qed + ultimately AOT_have \\w \ r & \w \ \r\ + by (rule "&I") + AOT_hence \\p (\w \ p & \w \ \p)\ + by (rule "\I") + thus ?thesis + by (rule "PossibleWorld.\I") +qed + +AOT_theorem "conj-dist-w:9": \(\w \ p) \ w \ \p\ +proof(rule "\I"; rule "raa-cor:1") + AOT_assume \\w \ p\ + AOT_hence 0: \w \ p\ + by (metis "\E"(1) "rigid-truth-at:2") + AOT_assume \\w \ \p\ + AOT_hence 1: \w \ \\p\ + using "coherent:1"[unvarify p, THEN "\E"(2), OF "log-prop-prop:2"] by blast + moreover AOT_have \w \ (\\p \ \p)\ + using "T\"[THEN "contraposition:1[1]", THEN RN] + "fund:2"[unvarify p, OF "log-prop-prop:2", THEN "\E"(1), THEN "\E"(2), + THEN "\E", rotated, OF PossibleWorld.\] + by blast + ultimately AOT_have \w \ \p\ + using "conj-dist-w:2"[unvarify p q, OF "log-prop-prop:2", OF "log-prop-prop:2", + THEN "\E"(1), THEN "\E"] + by blast + AOT_hence \w \ p & w \ \p\ using 0 "&I" by blast + AOT_thus \p & \p\ + by (metis "coherent:1" "Conjunction Simplification"(1,2) "\E"(4) + "modus-tollens:1" "raa-cor:3") +qed + +AOT_theorem "conj-dist-w:10": \\w\p((w \ \p) & \\w \ p)\ +proof - + AOT_obtain w where w: \\p (w \ p \ p)\ + using "act-world:1" "PossibleWorld.\E"[rotated] by meson + AOT_obtain r where \\r\ and \\r\ + using "cont-tf-thm:2" "cont-tf:2"[THEN "\\<^sub>d\<^sub>fE"] "&E" "\E"[rotated] by metis + AOT_hence \w \ \r\ and 0: \w \ \r\ + using w[THEN "\E"(1), OF "log-prop-prop:2", THEN "\E"(2)] by blast+ + AOT_hence \\w \ r\ using "coherent:1"[THEN "\E"(1)] by blast + AOT_hence \\\w \ r\ by (metis "\E"(4) "rigid-truth-at:2") + AOT_hence \w \ \r & \\w \ r\ using 0 "&I" by blast + AOT_hence \\p (w \ \p & \\w \ p)\ by (rule "\I") + thus ?thesis by (rule "PossibleWorld.\I") +qed + +AOT_theorem "two-worlds-exist:1": \\p(ContingentlyTrue(p)) \ \w (\Actual(w))\ +proof(rule "\I") + AOT_assume \\p ContingentlyTrue(p)\ + then AOT_obtain p where \ContingentlyTrue(p)\ + using "\E"[rotated] by blast + AOT_hence p: \p & \\p\ + by (metis "\\<^sub>d\<^sub>fE" "cont-tf:1") + AOT_hence \\w w \ \p\ + using "fund:1"[unvarify p, OF "log-prop-prop:2", THEN "\E"(1)] "&E" by blast + then AOT_obtain w where w: \w \ \p\ + using "PossibleWorld.\E"[rotated] by meson + AOT_have \\Actual(w)\ + proof(rule "raa-cor:2") + AOT_assume \Actual(w)\ + AOT_hence \w \ p\ + using p[THEN "&E"(1)] actual[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2)] + by (metis "log-prop-prop:2" "raa-cor:3" "rule-ui:1" "\E" w) + moreover AOT_have \\(w \ p)\ + by (metis "coherent:1" "\E"(4) "reductio-aa:2" w) + ultimately AOT_show \w \ p & \(w \ p)\ + using "&I" by blast + qed + AOT_thus \\w \Actual(w)\ + by (rule "PossibleWorld.\I") +qed + + +AOT_theorem "two-worlds-exist:2": \\p(ContingentlyFalse(p)) \ \w (\Actual(w))\ +proof(rule "\I") + AOT_assume \\p ContingentlyFalse(p)\ + then AOT_obtain p where \ContingentlyFalse(p)\ + using "\E"[rotated] by blast + AOT_hence p: \\p & \p\ + by (metis "\\<^sub>d\<^sub>fE" "cont-tf:2") + AOT_hence \\w w \ p\ + using "fund:1"[unvarify p, OF "log-prop-prop:2", THEN "\E"(1)] "&E" by blast + then AOT_obtain w where w: \w \ p\ + using "PossibleWorld.\E"[rotated] by meson + moreover AOT_have \\Actual(w)\ + proof(rule "raa-cor:2") + AOT_assume \Actual(w)\ + AOT_hence \w \ \p\ + using p[THEN "&E"(1)] actual[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2)] + by (metis "log-prop-prop:2" "raa-cor:3" "rule-ui:1" "\E" w) + moreover AOT_have \\(w \ p)\ + using calculation by (metis "coherent:1" "\E"(4) "reductio-aa:2") + AOT_thus \w \ p & \(w \ p)\ + using "&I" w by metis + qed + AOT_thus \\w \Actual(w)\ + by (rule "PossibleWorld.\I") +qed + +AOT_theorem "two-worlds-exist:3": \\w \Actual(w)\ + using "cont-tf-thm:1" "two-worlds-exist:1" "\E" by blast + +AOT_theorem "two-worlds-exist:4": \\w\w'(w \ w')\ +proof - + AOT_obtain w where w: \Actual(w)\ + using "act-world:2"[THEN "uniqueness:1"[THEN "\\<^sub>d\<^sub>fE"], + THEN "cqt-further:5"[THEN "\E"]] + "PossibleWorld.\E"[rotated] "&E" + by blast + moreover AOT_obtain w' where w': \\Actual(w')\ + using "two-worlds-exist:3" "PossibleWorld.\E"[rotated] by meson + AOT_have \\(w = w')\ + proof(rule "raa-cor:2") + AOT_assume \w = w'\ + AOT_thus \p & \p\ for p + using w w' "&E" by (metis "rule=E" "raa-cor:3") + qed + AOT_hence \w \ w'\ + by (metis "\\<^sub>d\<^sub>fI" "=-infix") + AOT_hence \\w' w \ w'\ + by (rule "PossibleWorld.\I") + thus ?thesis + by (rule "PossibleWorld.\I") +qed + +(* TODO: more theorems *) + +AOT_theorem "w-rel:1": \[\x \{x}]\ \ [\x w \ \{x}]\\ +proof(rule "\I") + AOT_assume \[\x \{x}]\\ + AOT_hence \\[\x \{x}]\\ + by (metis "exist-nec" "\E") + moreover AOT_have + \\[\x \{x}]\ \ \\x\y(\F([F]x \ [F]y) \ ((w \ \{x}) \ ( w \ \{y})))\ + proof (rule RM; rule "\I"; rule GEN; rule GEN; rule "\I") + AOT_modally_strict { + fix x y + AOT_assume \[\x \{x}]\\ + AOT_hence \\x\y (\F ([F]x \ [F]y) \ \(\{x} \ \{y}))\ + using "&E" "kirchner-thm-cor:1"[THEN "\E"] by blast + AOT_hence \\F ([F]x \ [F]y) \ \(\{x} \ \{y})\ + using "\E"(2) by blast + moreover AOT_assume \\F ([F]x \ [F]y)\ + ultimately AOT_have \\(\{x} \ \{y})\ + using "\E" by blast + AOT_hence \\w (w \ (\{x} \ \{y}))\ + using "fund:2"[unvarify p, OF "log-prop-prop:2", THEN "\E"(1)] by blast + AOT_hence \w \ (\{x} \ \{y})\ + using "\E"(2) using PossibleWorld.\ "\E" by blast + AOT_thus \(w \ \{x}) \ (w \ \{y})\ + using "conj-dist-w:4"[unvarify p q, OF "log-prop-prop:2", + OF "log-prop-prop:2", THEN "\E"(1)] by blast + } + qed + ultimately AOT_have \\\x\y(\F([F]x \ [F]y) \ ((w \ \{x}) \ ( w \ \{y})))\ + using "\E" by blast + AOT_thus \[\x w \ \{x}]\\ + using "kirchner-thm:1"[THEN "\E"(2)] by fast +qed + +AOT_theorem "w-rel:2": \[\x\<^sub>1...x\<^sub>n \{x\<^sub>1...x\<^sub>n}]\ \ [\x\<^sub>1...x\<^sub>n w \ \{x\<^sub>1...x\<^sub>n}]\\ +proof(rule "\I") + AOT_assume \[\x\<^sub>1...x\<^sub>n \{x\<^sub>1...x\<^sub>n}]\\ + AOT_hence \\[\x\<^sub>1...x\<^sub>n \{x\<^sub>1...x\<^sub>n}]\\ + by (metis "exist-nec" "\E") + moreover AOT_have \\[\x\<^sub>1...x\<^sub>n \{x\<^sub>1...x\<^sub>n}]\ \ \\x\<^sub>1...\x\<^sub>n\y\<^sub>1...\y\<^sub>n( + \F([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) \ ((w \ \{x\<^sub>1...x\<^sub>n}) \ ( w \ \{y\<^sub>1...y\<^sub>n})))\ + proof (rule RM; rule "\I"; rule GEN; rule GEN; rule "\I") + AOT_modally_strict { + fix x\<^sub>1x\<^sub>n y\<^sub>1y\<^sub>n + AOT_assume \[\x\<^sub>1...x\<^sub>n \{x\<^sub>1...x\<^sub>n}]\\ + AOT_hence \\x\<^sub>1...\x\<^sub>n\y\<^sub>1...\y\<^sub>n ( + \F ([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) \ \(\{x\<^sub>1...x\<^sub>n} \ \{y\<^sub>1...y\<^sub>n}))\ + using "&E" "kirchner-thm-cor:2"[THEN "\E"] by blast + AOT_hence \\F ([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) \ \(\{x\<^sub>1...x\<^sub>n} \ \{y\<^sub>1...y\<^sub>n})\ + using "\E"(2) by blast + moreover AOT_assume \\F ([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n)\ + ultimately AOT_have \\(\{x\<^sub>1...x\<^sub>n} \ \{y\<^sub>1...y\<^sub>n})\ + using "\E" by blast + AOT_hence \\w (w \ (\{x\<^sub>1...x\<^sub>n} \ \{y\<^sub>1...y\<^sub>n}))\ + using "fund:2"[unvarify p, OF "log-prop-prop:2", THEN "\E"(1)] by blast + AOT_hence \w \ (\{x\<^sub>1...x\<^sub>n} \ \{y\<^sub>1...y\<^sub>n})\ + using "\E"(2) using PossibleWorld.\ "\E" by blast + AOT_thus \(w \ \{x\<^sub>1...x\<^sub>n}) \ (w \ \{y\<^sub>1...y\<^sub>n})\ + using "conj-dist-w:4"[unvarify p q, OF "log-prop-prop:2", + OF "log-prop-prop:2", THEN "\E"(1)] by blast + } + qed + ultimately AOT_have \\\x\<^sub>1...\x\<^sub>n\y\<^sub>1...\y\<^sub>n( + \F([F]x\<^sub>1...x\<^sub>n \ [F]y\<^sub>1...y\<^sub>n) \ ((w \ \{x\<^sub>1...x\<^sub>n}) \ ( w \ \{y\<^sub>1...y\<^sub>n})))\ + using "\E" by blast + AOT_thus \[\x\<^sub>1...x\<^sub>n w \ \{x\<^sub>1...x\<^sub>n}]\\ + using "kirchner-thm:2"[THEN "\E"(2)] by fast +qed + +AOT_theorem "w-rel:3": \[\x\<^sub>1...x\<^sub>n w \ [F]x\<^sub>1...x\<^sub>n]\\ + by (rule "w-rel:2"[THEN "\E"]) "cqt:2[lambda]" + +AOT_define WorldIndexedRelation :: \\ \ \ \ \\ (\_\<^sub>_\) + "w-index": \[F]\<^sub>w =\<^sub>d\<^sub>f [\x\<^sub>1...x\<^sub>n w \ [F]x\<^sub>1...x\<^sub>n]\ + +AOT_define Rigid :: \\ \ \\ (\Rigid'(_')\) + "df-rigid-rel:1": +\Rigid(F) \\<^sub>d\<^sub>f F\ & \\x\<^sub>1...\x\<^sub>n([F]x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n)\ + +AOT_define Rigidifies :: \\ \ \ \ \\ (\Rigidifies'(_,_')\) + "df-rigid-rel:2": +\Rigidifies(F, G) \\<^sub>d\<^sub>f Rigid(F) & \x\<^sub>1...\x\<^sub>n([F]x\<^sub>1...x\<^sub>n \ [G]x\<^sub>1...x\<^sub>n)\ + +AOT_theorem "rigid-der:1": \[[F]\<^sub>w]x\<^sub>1...x\<^sub>n \ w \ [F]x\<^sub>1...x\<^sub>n\ + apply (rule "rule-id-df:2:b[2]"[where \="\ (\, \). \[\]\<^sub>\\" and + \="\(\, \). \[\x\<^sub>1...x\<^sub>n \ \ [\]x\<^sub>1...x\<^sub>n]\", + simplified, OF "w-index"]) + apply (fact "w-rel:3") + apply (rule "beta-C-meta"[THEN "\E"]) + by (fact "w-rel:3") + +AOT_theorem "rigid-der:2": \Rigid([G]\<^sub>w)\ +proof(safe intro!: "\\<^sub>d\<^sub>fI"[OF "df-rigid-rel:1"] "&I") + AOT_show \[G]\<^sub>w\\ + by (rule "rule-id-df:2:b[2]"[where \="\ (\, \). \[\]\<^sub>\\" and + \="\(\, \). \[\x\<^sub>1...x\<^sub>n \ \ [\]x\<^sub>1...x\<^sub>n]\", + simplified, OF "w-index"]) + (fact "w-rel:3")+ +next + AOT_have \\\x\<^sub>1...\x\<^sub>n ([[G]\<^sub>w]x\<^sub>1...x\<^sub>n \ \[[G]\<^sub>w]x\<^sub>1...x\<^sub>n)\ + proof(rule RN; safe intro!: "\I" GEN) + AOT_modally_strict { + AOT_have assms: \PossibleWorld(w)\ using PossibleWorld.\. + AOT_hence nec_pw_w: \\PossibleWorld(w)\ + using "\E"(1) "rigid-pw:1" by blast + fix x\<^sub>1x\<^sub>n + AOT_assume \[[G]\<^sub>w]x\<^sub>1...x\<^sub>n\ + AOT_hence \[\x\<^sub>1...x\<^sub>n w \ [G]x\<^sub>1...x\<^sub>n]x\<^sub>1...x\<^sub>n\ + using "rule-id-df:2:a[2]"[where \="\ (\, \). \[\]\<^sub>\\" and + \="\(\, \). \[\x\<^sub>1...x\<^sub>n \ \ [\]x\<^sub>1...x\<^sub>n]\", + simplified, OF "w-index", OF "w-rel:3"] + by fast + AOT_hence \w \ [G]x\<^sub>1...x\<^sub>n\ + by (metis "\\C"(1)) + AOT_hence \\w \ [G]x\<^sub>1...x\<^sub>n\ + using "rigid-truth-at:1"[unvarify p, OF "log-prop-prop:2", THEN "\E"(1)] + by blast + moreover AOT_have \\w \ [G]x\<^sub>1...x\<^sub>n \ \[\x\<^sub>1...x\<^sub>n w \ [G]x\<^sub>1...x\<^sub>n]x\<^sub>1...x\<^sub>n\ + proof (rule RM; rule "\I") + AOT_modally_strict { + AOT_assume \w \ [G]x\<^sub>1...x\<^sub>n\ + AOT_thus \[\x\<^sub>1...x\<^sub>n w \ [G]x\<^sub>1...x\<^sub>n]x\<^sub>1...x\<^sub>n\ + by (auto intro!: "\\C"(1) simp: "w-rel:3" "cqt:2") + } + qed + ultimately AOT_have 1: \\[\x\<^sub>1...x\<^sub>n w \ [G]x\<^sub>1...x\<^sub>n]x\<^sub>1...x\<^sub>n\ + using "\E" by blast + AOT_show \\[[G]\<^sub>w]x\<^sub>1...x\<^sub>n\ + by (rule "rule-id-df:2:b[2]"[where \="\ (\, \). \[\]\<^sub>\\" and + \="\(\, \). \[\x\<^sub>1...x\<^sub>n \ \ [\]x\<^sub>1...x\<^sub>n]\", + simplified, OF "w-index"]) + (auto simp: 1 "w-rel:3") + } + qed + AOT_thus \\\x\<^sub>1...\x\<^sub>n ([[G]\<^sub>w]x\<^sub>1...x\<^sub>n \ \[[G]\<^sub>w]x\<^sub>1...x\<^sub>n)\ + using "\E" by blast +qed + +AOT_theorem "rigid-der:3": \\F Rigidifies(F, G)\ +proof - + AOT_obtain w where w: \\p (w \ p \ p)\ + using "act-world:1" "PossibleWorld.\E"[rotated] by meson + show ?thesis + proof (rule "\I"(1)[where \=\\[G]\<^sub>w\\]) + AOT_show \Rigidifies([G]\<^sub>w, [G])\ + proof(safe intro!: "\\<^sub>d\<^sub>fI"[OF "df-rigid-rel:2"] "&I" GEN) + AOT_show \Rigid([G]\<^sub>w)\ + using "rigid-der:2" by blast + next + fix x\<^sub>1x\<^sub>n + AOT_have \[[G]\<^sub>w]x\<^sub>1...x\<^sub>n \ [\x\<^sub>1...x\<^sub>n w \ [G]x\<^sub>1...x\<^sub>n]x\<^sub>1...x\<^sub>n\ + proof(rule "\I"; rule "\I") + AOT_assume \[[G]\<^sub>w]x\<^sub>1...x\<^sub>n\ + AOT_thus \[\x\<^sub>1...x\<^sub>n w \ [G]x\<^sub>1...x\<^sub>n]x\<^sub>1...x\<^sub>n\ + by (rule "rule-id-df:2:a[2]" + [where \="\ (\, \). \[\]\<^sub>\\" and + \="\(\, \). \[\x\<^sub>1...x\<^sub>n \ \ [\]x\<^sub>1...x\<^sub>n]\", + simplified, OF "w-index", OF "w-rel:3"]) + next + AOT_assume \[\x\<^sub>1...x\<^sub>n w \ [G]x\<^sub>1...x\<^sub>n]x\<^sub>1...x\<^sub>n\ + AOT_thus \[[G]\<^sub>w]x\<^sub>1...x\<^sub>n\ + by (rule "rule-id-df:2:b[2]" + [where \="\ (\, \). \[\]\<^sub>\\" and + \="\(\, \). \[\x\<^sub>1...x\<^sub>n \ \ [\]x\<^sub>1...x\<^sub>n]\", + simplified, OF "w-index", OF "w-rel:3"]) + qed + also AOT_have \\ \ w \ [G]x\<^sub>1...x\<^sub>n\ + by (rule "beta-C-meta"[THEN "\E"]) + (fact "w-rel:3") + also AOT_have \\ \ [G]x\<^sub>1...x\<^sub>n\ + using w[THEN "\E"(1), OF "log-prop-prop:2"] by blast + finally AOT_show \[[G]\<^sub>w]x\<^sub>1...x\<^sub>n \ [G]x\<^sub>1...x\<^sub>n\. + qed + next + AOT_show \[G]\<^sub>w\\ + by (rule "rule-id-df:2:b[2]"[where \="\ (\, \). \[\]\<^sub>\\" + and \="\(\, \). \[\x\<^sub>1...x\<^sub>n \ \ [\]x\<^sub>1...x\<^sub>n]\", + simplified, OF "w-index"]) + (auto simp: "w-rel:3") + qed +qed + +AOT_theorem "rigid-rel-thms:1": + \\(\x\<^sub>1...\x\<^sub>n ([F]x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n)) \ \x\<^sub>1...\x\<^sub>n(\[F]x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n)\ +proof(safe intro!: "\I" "\I" GEN) + fix x\<^sub>1x\<^sub>n + AOT_assume \\\x\<^sub>1...\x\<^sub>n ([F]x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n)\ + AOT_hence \\x\<^sub>1...\x\<^sub>n \([F]x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n)\ + by (metis "\E" GEN RM "cqt-orig:3") + AOT_hence \\([F]x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n)\ + using "\E"(2) by blast + AOT_hence \\[F]x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n\ + by (metis "\E"(1) "sc-eq-box-box:1") + moreover AOT_assume \\[F]x\<^sub>1...x\<^sub>n\ + ultimately AOT_show \\[F]x\<^sub>1...x\<^sub>n\ + using "\E" by blast +next + AOT_assume \\x\<^sub>1...\x\<^sub>n (\[F]x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n)\ + AOT_hence \\[F]x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n\ for x\<^sub>1x\<^sub>n + using "\E"(2) by blast + AOT_hence \\([F]x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n)\ for x\<^sub>1x\<^sub>n + by (metis "\E"(2) "sc-eq-box-box:1") + AOT_hence 0: \\x\<^sub>1...\x\<^sub>n \([F]x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n)\ + by (rule GEN) + AOT_thus \\(\x\<^sub>1...\x\<^sub>n ([F]x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n))\ + using "BF" "vdash-properties:10" by blast +qed + +AOT_theorem "rigid-rel-thms:2": + \\(\x\<^sub>1...\x\<^sub>n ([F]x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n)) \ \x\<^sub>1...\x\<^sub>n(\[F]x\<^sub>1...x\<^sub>n \ \\[F]x\<^sub>1...x\<^sub>n)\ +proof(safe intro!: "\I" "\I") + AOT_assume \\(\x\<^sub>1...\x\<^sub>n ([F]x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n))\ + AOT_hence 0: \\x\<^sub>1...\x\<^sub>n \([F]x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n)\ + using CBF[THEN "\E"] by blast + AOT_show \\x\<^sub>1...\x\<^sub>n(\[F]x\<^sub>1...x\<^sub>n \ \\[F]x\<^sub>1...x\<^sub>n)\ + proof(rule GEN) + fix x\<^sub>1x\<^sub>n + AOT_have 1: \\([F]x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n)\ + using 0[THEN "\E"(2)]. + AOT_hence 2: \\[F]x\<^sub>1...x\<^sub>n \ [F]x\<^sub>1...x\<^sub>n\ + using "B\" "Hypothetical Syllogism" "K\" "vdash-properties:10" by blast + AOT_have \[F]x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n\ + using "exc-mid". + moreover { + AOT_assume \[F]x\<^sub>1...x\<^sub>n\ + AOT_hence \\[F]x\<^sub>1...x\<^sub>n\ + using 1[THEN "qml:2"[axiom_inst, THEN "\E"], THEN "\E"] by blast + } + moreover { + AOT_assume 3: \\[F]x\<^sub>1...x\<^sub>n\ + AOT_have \\\[F]x\<^sub>1...x\<^sub>n\ + proof(rule "raa-cor:1") + AOT_assume \\\\[F]x\<^sub>1...x\<^sub>n\ + AOT_hence \\[F]x\<^sub>1...x\<^sub>n\ + by (AOT_subst_def "conventions:5") + AOT_hence \[F]x\<^sub>1...x\<^sub>n\ using 2[THEN "\E"] by blast + AOT_thus \[F]x\<^sub>1...x\<^sub>n & \[F]x\<^sub>1...x\<^sub>n\ + using 3 "&I" by blast + qed + } + ultimately AOT_show \\[F]x\<^sub>1...x\<^sub>n \ \\[F]x\<^sub>1...x\<^sub>n\ + by (metis "\I"(1,2) "raa-cor:1") + qed +next + AOT_assume 0: \\x\<^sub>1...\x\<^sub>n(\[F]x\<^sub>1...x\<^sub>n \ \\[F]x\<^sub>1...x\<^sub>n)\ + { + fix x\<^sub>1x\<^sub>n + AOT_have \\[F]x\<^sub>1...x\<^sub>n \ \\[F]x\<^sub>1...x\<^sub>n\ using 0[THEN "\E"(2)] by blast + moreover { + AOT_assume \\[F]x\<^sub>1...x\<^sub>n\ + AOT_hence \\\[F]x\<^sub>1...x\<^sub>n\ + using "S5Basic:6"[THEN "\E"(1)] by blast + AOT_hence \\([F]x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n)\ + using "KBasic:1"[THEN "\E"] by blast + } + moreover { + AOT_assume \\\[F]x\<^sub>1...x\<^sub>n\ + AOT_hence \\([F]x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n)\ + using "KBasic:2"[THEN "\E"] by blast + } + ultimately AOT_have \\([F]x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n)\ + using "con-dis-i-e:4:b" "raa-cor:1" by blast + } + AOT_hence \\x\<^sub>1...\x\<^sub>n \([F]x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n)\ + by (rule GEN) + AOT_thus \\(\x\<^sub>1...\x\<^sub>n ([F]x\<^sub>1...x\<^sub>n \ \[F]x\<^sub>1...x\<^sub>n))\ + using BF[THEN "\E"] by fast +qed + +AOT_theorem "rigid-rel-thms:3": \Rigid(F) \ \x\<^sub>1...\x\<^sub>n (\[F]x\<^sub>1...x\<^sub>n \ \\[F]x\<^sub>1...x\<^sub>n)\ + by (AOT_subst_thm "df-rigid-rel:1"[THEN "\Df", THEN "\S"(1), OF "cqt:2"(1)]; + AOT_subst_thm "rigid-rel-thms:2") + (simp add: "oth-class-taut:3:a") + +(*<*) +end +(*>*) diff --git a/thys/AOT/AOT_RestrictedVariables.thy b/thys/AOT/AOT_RestrictedVariables.thy new file mode 100644 --- /dev/null +++ b/thys/AOT/AOT_RestrictedVariables.thy @@ -0,0 +1,437 @@ +(*<*) +theory AOT_RestrictedVariables + imports AOT_PLM + keywords "AOT_register_rigid_restricted_type" :: thy_goal + and "AOT_register_restricted_type" :: thy_goal +begin +(*>*) + +section\Restricted Variables\ + +locale AOT_restriction_condition = + fixes \ :: \'a::AOT_Term_id_2 \ \\ + assumes "res-var:2"[AOT]: \[v \ \\ \{\}]\ + assumes "res-var:3"[AOT]: \[v \ \{\} \ \\]\ + +ML\ +fun register_restricted_type (name:string, restriction:string) thy = +let +val ctxt = thy +val ctxt = setupStrictWorld ctxt +val trm = Syntax.check_term ctxt (AOT_read_term @{nonterminal \'} ctxt restriction) +val free = case (Term.add_frees trm []) of [f] => f | + _ => raise Term.TERM ("Expected single free variable.", [trm]) +val trm = Term.absfree free trm +val localeTerm = Const (\<^const_name>\AOT_restriction_condition\, dummyT) $ trm +val localeTerm = Syntax.check_term ctxt localeTerm +fun after_qed thms thy = let +val st = Interpretation.global_interpretation + (([(@{locale AOT_restriction_condition}, ((name, true), + (Expression.Named [("\", trm)], [])))], [])) [] thy +val st = Proof.refine_insert (flat thms) st +val thy = Proof.global_immediate_proof st + +val thy = Local_Theory.background_theory + (AOT_Constraints.map (Symtab.update + (name, (term_of (snd free), term_of (snd free))))) thy +val thy = Local_Theory.background_theory + (AOT_Restriction.map (Symtab.update + (name, (trm, Const (\<^const_name>\AOT_term_of_var\, dummyT))))) thy + +in thy end +in +Proof.theorem NONE after_qed [[(HOLogic.mk_Trueprop localeTerm, [])]] ctxt +end + +val _ = + Outer_Syntax.command + \<^command_keyword>\AOT_register_restricted_type\ + "Register a restricted type." + (((Parse.short_ident --| Parse.$$$ ":") -- Parse.term) >> + (Toplevel.local_theory_to_proof NONE NONE o register_restricted_type)); +\ + +locale AOT_rigid_restriction_condition = AOT_restriction_condition + + assumes rigid[AOT]: \[v \ \\(\{\} \ \\{\})]\ +begin +lemma rigid_condition[AOT]: \[v \ \(\{\} \ \\{\})]\ + using rigid[THEN "\E"(2)] RN by simp +lemma type_set_nonempty[AOT_no_atp, no_atp]: \\x . x \ { \ . [w\<^sub>0 \ \{\}]}\ + by (metis "instantiation" mem_Collect_eq "res-var:2") +end + +locale AOT_restricted_type = AOT_rigid_restriction_condition + + fixes Rep and Abs + assumes AOT_restricted_type_definition[AOT_no_atp]: + \type_definition Rep Abs { \ . [w\<^sub>0 \ \{\}]}\ +begin + +AOT_theorem restricted_var_condition: \\{\AOT_term_of_var (Rep \)\}\ +proof - + interpret type_definition Rep Abs "{ \ . [w\<^sub>0 \ \{\}]}" + using AOT_restricted_type_definition. + AOT_actually { + AOT_have \\AOT_term_of_var (Rep \)\\\ and \\{\AOT_term_of_var (Rep \)\}\ + using AOT_sem_imp Rep "res-var:3" by auto + } + moreover AOT_actually { + AOT_have \\{\} \ \\{\}\ for \ + using AOT_sem_box rigid_condition by presburger + AOT_hence \\{\} \ \\{\}\ if \\\\ for \ + by (metis AOT_model.AOT_term_of_var_cases AOT_sem_denotes that) + } + ultimately AOT_show \\{\AOT_term_of_var (Rep \)\}\ + using AOT_sem_box AOT_sem_imp by blast +qed +lemmas "\" = restricted_var_condition + +AOT_theorem GEN: assumes \for arbitrary \: \{\AOT_term_of_var (Rep \)\}\ + shows \\\ (\{\} \ \{\})\ +proof(rule GEN; rule "\I") + interpret type_definition Rep Abs "{ \ . [w\<^sub>0 \ \{\}]}" + using AOT_restricted_type_definition. + fix \ + AOT_assume \\{\}\ + AOT_hence \\<^bold>\\{\}\ + by (metis AOT_model_axiom_def AOT_sem_box AOT_sem_imp act_closure rigid_condition) + hence 0: \[w\<^sub>0 \ \{\}]\ by (metis AOT_sem_act) + { + fix \ + assume \_def: \\ = Rep \\ + AOT_have \\{\}\ + unfolding \_def + using assms by blast + } + AOT_thus \\{\}\ + using Rep_cases[simplified, OF 0] + by blast +qed +lemmas "\I" = GEN + +end + + +lemma AOT_restricted_type_intro[AOT_no_atp, no_atp]: + assumes \type_definition Rep Abs { \ . [w\<^sub>0 \ \{\}]}\ + and \AOT_rigid_restriction_condition \\ + shows \AOT_restricted_type \ Rep Abs\ + by (auto intro!: assms AOT_restricted_type_axioms.intro AOT_restricted_type.intro) + + + +ML\ +fun register_rigid_restricted_type (name:string, restriction:string) thy = +let +val ctxt = thy +val ctxt = setupStrictWorld ctxt +val trm = Syntax.check_term ctxt (AOT_read_term @{nonterminal \'} ctxt restriction) +val free = case (Term.add_frees trm []) of [f] => f + | _ => raise Term.TERM ("Expected single free variable.", [trm]) +val trm = Term.absfree free trm +val localeTerm = HOLogic.mk_Trueprop + (Const (\<^const_name>\AOT_rigid_restriction_condition\, dummyT) $ trm) +val localeTerm = Syntax.check_prop ctxt localeTerm +val int_bnd = Binding.concealed (Binding.qualify true "internal" (Binding.name name)) +val bnds = {Rep_name = Binding.qualify true name (Binding.name "Rep"), + Abs_name = Binding.qualify true "Abs" int_bnd, + type_definition_name = Binding.qualify true "type_definition" int_bnd} + +fun after_qed witts thy = let + val thms = (map (Element.conclude_witness ctxt) (flat witts)) + + val typeset = HOLogic.mk_Collect ("\", dummyT, + \<^const>\AOT_model_valid_in\ $ \<^const>\w\<^sub>0\ $ + (trm $ (Const (\<^const_name>\AOT_term_of_var\, dummyT) $ Bound 0))) + val typeset = Syntax.check_term thy typeset + val nonempty_thm = Drule.OF + (@{thm AOT_rigid_restriction_condition.type_set_nonempty}, thms) + + val ((_,st),thy) = Typedef.add_typedef {overloaded=true} + (Binding.name name, [], Mixfix.NoSyn) typeset (SOME bnds) + (fn ctxt => (Tactic.resolve_tac ctxt ([nonempty_thm]) 1)) thy + val ({rep_type = _, abs_type = _, Rep_name = Rep_name, Abs_name = Abs_name, + axiom_name = _}, + {inhabited = _, type_definition = type_definition, Rep = _, + Rep_inverse = _, Abs_inverse = _, Rep_inject = _, Abs_inject = _, + Rep_cases = _, Abs_cases = _, Rep_induct = _, Abs_induct = _}) = st + + val locale_thm = Drule.OF (@{thm AOT_restricted_type_intro}, type_definition::thms) + + val st = Interpretation.global_interpretation (([(@{locale AOT_restricted_type}, + ((name, true), (Expression.Named [ + ("\", trm), + ("Rep", Const (Rep_name, dummyT)), + ("Abs", Const (Abs_name, dummyT))], []))) + ], [])) [] thy + + val st = Proof.refine_insert [locale_thm] st + val thy = Proof.global_immediate_proof st + + val thy = Local_Theory.background_theory (AOT_Constraints.map ( + Symtab.update (name, (term_of (snd free), term_of (snd free))))) thy + val thy = Local_Theory.background_theory (AOT_Restriction.map ( + Symtab.update (name, (trm, Const (Rep_name, dummyT))))) thy + + in thy end +in +Element.witness_proof after_qed [[localeTerm]] thy +end + +val _ = + Outer_Syntax.command + \<^command_keyword>\AOT_register_rigid_restricted_type\ + "Register a restricted type." + (((Parse.short_ident --| Parse.$$$ ":") -- Parse.term) >> + (Toplevel.local_theory_to_proof NONE NONE o register_rigid_restricted_type)); +\ + +(* Generalized mechanism for "AOT_restricted_type.\I" followed by \E *) +ML\ +fun get_instantiated_allI ctxt varname thm = let +val trm = Thm.concl_of thm +val trm = case trm of (@{const Trueprop} $ (@{const AOT_model_valid_in} $ _ $ x)) => x + | _ => raise Term.TERM ("Expected simple theorem.", [trm]) +fun extractVars (Const (\<^const_name>\AOT_term_of_var\, t) $ (Const rep $ Var v)) = + (if fst (fst v) = fst varname + then [Const (\<^const_name>\AOT_term_of_var\, t) $ (Const rep $ Var v)] + else []) (* TODO: care about the index *) + | extractVars (t1 $ t2) = extractVars t1 @ extractVars t2 + | extractVars (Abs (_, _, t)) = extractVars t + | extractVars _ = [] +val vars = extractVars trm +val vartrm = hd vars +val vars = fold Term.add_vars vars [] +val var = hd vars +val trmty = (case vartrm of (Const (_, Type ("fun", [_, ty])) $ _) => ty + | _ => raise Match) +val varty = snd var +val tyname = fst (Term.dest_Type varty) +val b = tyname^".\I" (* TODO: better way to find the theorem *) +val thms = fst (Context.map_proof_result (fn ctxt => (Attrib.eval_thms ctxt + [(Facts.Named ((b,Position.none),NONE),[])], ctxt)) ctxt) +val allthm = (case thms of (thm::_) => thm + | _ => raise Fail "Unknown restricted type.") +val trm = Abs (Term.string_of_vname (fst var), trmty, Term.abstract_over (vartrm, trm)) +val trm = Thm.cterm_of (Context.proof_of ctxt) trm +val phi = hd (Term.add_vars (Thm.prop_of allthm) []) +val allthm = Drule.instantiate_normalize (TVars.empty, Vars.make [(phi,trm)]) allthm +in +allthm +end +\ + +(* TODO: unconstraining multiple variables does not work yet *) +attribute_setup "unconstrain" = + \Scan.lift (Scan.repeat1 Args.var) >> (fn args => Thm.rule_attribute [] + (fn ctxt => fn thm => + let + val thm = fold (fn arg => fn thm => thm RS get_instantiated_allI ctxt arg thm) + args thm + val thm = fold (fn _ => fn thm => thm RS @{thm "\E"(2)}) args thm + in + thm + end))\ + "Generalize a statement about restricted variables to a statement about + unrestricted variables with explicit restriction condition." + + + +context AOT_restricted_type +begin + +AOT_theorem "rule-ui": + assumes \\\(\{\} \ \{\})\ + shows \\{\AOT_term_of_var (Rep \)\}\ +proof - + AOT_have \\{\}\ if \\{\}\ for \ using assms[THEN "\E"(2), THEN "\E"] that by blast + moreover AOT_have \\{\AOT_term_of_var (Rep \)\}\ + by (auto simp: \) + ultimately show ?thesis by blast +qed +lemmas "\E" = "rule-ui" + +AOT_theorem "instantiation": + assumes \for arbitrary \: \{\AOT_term_of_var (Rep \)\} \<^bold>\ \\ and \\\ (\{\} & \{\})\ + shows \\\ +proof - + AOT_have \\{\AOT_term_of_var (Rep \)\} \ \\ for \ + using assms(1) + by (simp add: "deduction-theorem") + AOT_hence 0: \\\ (\{\} \ (\{\} \ \))\ + using GEN by simp + moreover AOT_obtain \ where \\{\} & \{\}\ using assms(2) "\E"[rotated] by blast + ultimately AOT_show \\\ using "AOT_PLM.\E"(2)[THEN "\E", THEN "\E"] "&E" by fast +qed +lemmas "\E" = "instantiation" + +AOT_theorem existential: assumes \\{\AOT_term_of_var (Rep \)\}\ + shows \\ \ (\{\} & \{\})\ + by (meson AOT_restricted_type.\ AOT_restricted_type_axioms assms + "&I" "existential:2[const_var]") +lemmas "\I" = existential +end + + +context AOT_rigid_restriction_condition +begin + +AOT_theorem "res-var-bound-reas[1]": + \\\(\{\} \ \\ \{\, \}) \ \\\\ (\{\} \ \{\, \})\ +proof(safe intro!: "\I" "\I" GEN) + fix \ \ + AOT_assume \\\ (\{\} \ \\ \{\, \})\ + AOT_hence \\{\} \ \\ \{\, \}\ using "\E"(2) by blast + moreover AOT_assume \\{\}\ + ultimately AOT_have \\\ \{\, \}\ using "\E" by blast + AOT_thus \\{\, \}\ using "\E"(2) by blast +next + fix \ \ + AOT_assume \\\\\(\{\} \ \{\, \})\ + AOT_hence \\\(\{\} \ \{\, \})\ using "\E"(2) by blast + AOT_hence \\{\} \ \{\, \}\ using "\E"(2) by blast + moreover AOT_assume \\{\}\ + ultimately AOT_show \\{\, \}\ using "\E" by blast +qed + +AOT_theorem "res-var-bound-reas[BF]": + \\\(\{\} \ \\{\}) \ \\\(\{\} \ \{\})\ +proof(safe intro!: "\I") + AOT_assume \\\(\{\} \ \\{\})\ + AOT_hence \\{\} \ \\{\}\ for \ + using "\E"(2) by blast + AOT_hence \\(\{\} \ \{\})\ for \ + by (metis "sc-eq-box-box:6" rigid_condition "vdash-properties:6") + AOT_hence \\\ \(\{\} \ \{\})\ + by (rule GEN) + AOT_thus \\\\ (\{\} \ \{\})\ + by (metis "BF" "vdash-properties:6") +qed + +AOT_theorem "res-var-bound-reas[CBF]": +\\\\(\{\} \ \{\}) \ \\(\{\} \ \\{\})\ +proof(safe intro!: "\I" GEN) + fix \ + AOT_assume \\\\ (\{\} \ \{\})\ + AOT_hence \\\ \(\{\} \ \{\})\ + by (metis "CBF" "vdash-properties:6") + AOT_hence 1: \\(\{\} \ \{\})\ + using "\E"(2) by blast + AOT_assume \\{\}\ + AOT_hence \\\{\}\ + by (metis "B\" "T\" rigid_condition "vdash-properties:6") + AOT_thus \\\{\}\ + using 1 "qml:1"[axiom_inst, THEN "\E", THEN "\E"] by blast +qed + +AOT_theorem "res-var-bound-reas[2]": +\\\ (\{\} \ \<^bold>\\{\}) \ \<^bold>\\\ (\{\} \ \{\})\ +proof(safe intro!: "\I") + AOT_assume \\\ (\{\} \ \<^bold>\\{\})\ + AOT_hence \\{\} \ \<^bold>\\{\}\ for \ + using "\E"(2) by blast + AOT_hence \\<^bold>\(\{\} \ \{\})\ for \ + by (metis "sc-eq-box-box:7" rigid_condition "vdash-properties:6") + AOT_hence \\\ \<^bold>\(\{\} \ \{\})\ + by (rule GEN) + AOT_thus \\<^bold>\\\(\{\} \ \{\})\ + by (metis "\E"(2) "logic-actual-nec:3"[axiom_inst]) +qed + + +AOT_theorem "res-var-bound-reas[3]": + \\<^bold>\\\ (\{\} \ \{\}) \ \\ (\{\} \ \<^bold>\\{\})\ +proof(safe intro!: "\I" GEN) + fix \ + AOT_assume \\<^bold>\\\ (\{\} \ \{\})\ + AOT_hence \\\ \<^bold>\(\{\} \ \{\})\ + by (metis "\E"(1) "logic-actual-nec:3"[axiom_inst]) + AOT_hence 1: \\<^bold>\(\{\} \ \{\})\ by (metis "rule-ui:3") + AOT_assume \\{\}\ + AOT_hence \\<^bold>\\{\}\ + by (metis "nec-imp-act" "qml:2"[axiom_inst] rigid_condition "\E") + AOT_thus \\<^bold>\\{\}\ + using 1 by (metis "act-cond" "\E") +qed + +AOT_theorem "res-var-bound-reas[Buridan]": + \\\ (\{\} & \\{\}) \ \\\ (\{\} & \{\})\ +proof (rule "\I") + AOT_assume \\\ (\{\} & \\{\})\ + then AOT_obtain \ where \\{\} & \\{\}\ + using "\E"[rotated] by blast + AOT_hence \\(\{\} & \{\})\ + by (metis "KBasic:11" "KBasic:3" "T\" "&I" "&E"(1) "&E"(2) + "\E"(2) "reductio-aa:1" rigid_condition "vdash-properties:6") + AOT_hence \\\ \(\{\} & \{\})\ + by (rule "\I") + AOT_thus \\\\ (\{\} & \{\})\ + by (rule Buridan[THEN "\E"]) +qed + +AOT_theorem "res-var-bound-reas[BF\]": + \\\\ (\{\} & \{\}) \ \\ (\{\} & \\{\})\ +proof(rule "\I") + AOT_assume \\\\ (\{\} & \{\})\ + AOT_hence \\\ \(\{\} & \{\})\ + using "BF\"[THEN "\E"] by blast + then AOT_obtain \ where \\(\{\} & \{\})\ + using "\E"[rotated] by blast + AOT_hence \\\{\}\ and \\\{\}\ + using "KBasic2:3" "&E" "\E" by blast+ + moreover AOT_have \\{\}\ + using calculation rigid_condition by (metis "B\" "K\" "\E") + ultimately AOT_have \\{\} & \\{\}\ + using "&I" by blast + AOT_thus \\\ (\{\} & \\{\})\ + by (rule "\I") +qed + +AOT_theorem "res-var-bound-reas[CBF\]": + \\\ (\{\} & \\{\}) \ \\\ (\{\} & \{\})\ +proof(rule "\I") + AOT_assume \\\ (\{\} & \\{\})\ + then AOT_obtain \ where \\{\} & \\{\}\ + using "\E"[rotated] by blast + AOT_hence \\\{\}\ and \\\{\}\ + using rigid_condition[THEN "qml:2"[axiom_inst, THEN "\E"], THEN "\E"] "&E" by blast+ + AOT_hence \\(\{\} & \{\})\ + by (metis "KBasic:16" "con-dis-taut:5" "\E") + AOT_hence \\\ \(\{\} & \{\})\ + by (rule "\I") + AOT_thus \\\\ (\{\} & \{\})\ + using "CBF\"[THEN "\E"] by fast +qed + +AOT_theorem "res-var-bound-reas[A-Exists:1]": + \\<^bold>\\!\ (\{\} & \{\}) \ \!\ (\{\} & \<^bold>\\{\})\ +proof(safe intro!: "\I" "\I") + AOT_assume \\<^bold>\\!\ (\{\} & \{\})\ + AOT_hence \\!\ \<^bold>\(\{\} & \{\})\ + using "A-Exists:1"[THEN "\E"(1)] by blast + AOT_hence \\!\ (\<^bold>\\{\} & \<^bold>\\{\})\ + apply(AOT_subst \\<^bold>\\{\} & \<^bold>\\{\}\ \\<^bold>\(\{\} & \{\})\ for: \) + apply (meson "Act-Basic:2" "intro-elim:3:f" "oth-class-taut:3:a") + by simp + AOT_thus \\!\ (\{\} & \<^bold>\\{\})\ + apply (AOT_subst \\{\}\ \\<^bold>\\{\}\ for: \) + using "Commutativity of \" "intro-elim:3:b" "sc-eq-fur:2" + "\E" rigid_condition by blast +next + AOT_assume \\!\ (\{\} & \<^bold>\\{\})\ + AOT_hence \\!\ (\<^bold>\\{\} & \<^bold>\\{\})\ + apply (AOT_subst \\<^bold>\\{\}\ \\{\}\ for: \) + apply (meson "sc-eq-fur:2" "\E" rigid_condition) + by simp + AOT_hence \\!\ \<^bold>\(\{\} & \{\})\ + apply (AOT_subst \\<^bold>\(\{\} & \{\})\ \\<^bold>\\{\} & \<^bold>\\{\}\ for: \) + using "Act-Basic:2" apply presburger + by simp + AOT_thus \\<^bold>\\!\ (\{\} & \{\})\ + by (metis "A-Exists:1" "intro-elim:3:b") +qed + +end + +(*<*) +end +(*>*) \ No newline at end of file diff --git a/thys/AOT/AOT_commands.ML b/thys/AOT/AOT_commands.ML new file mode 100644 --- /dev/null +++ b/thys/AOT/AOT_commands.ML @@ -0,0 +1,893 @@ +(* +This file contains modified parts of the Isabelle ML sources +which are distributed with Isabelle under the following conditions: + +ISABELLE COPYRIGHT NOTICE, LICENCE AND DISCLAIMER. + +Copyright (c) 1986-2021, + University of Cambridge, + Technische Universitaet Muenchen, + and contributors. + + All rights reserved. + +Redistribution and use in source and binary forms, with or without +modification, are permitted provided that the following conditions are +met: + +* Redistributions of source code must retain the above copyright +notice, this list of conditions and the following disclaimer. + +* Redistributions in binary form must reproduce the above copyright +notice, this list of conditions and the following disclaimer in the +documentation and/or other materials provided with the distribution. + +* Neither the name of the University of Cambridge or the Technische +Universitaet Muenchen nor the names of their contributors may be used +to endorse or promote products derived from this software without +specific prior written permission. + +THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS +IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED +TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A +PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT +OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, +SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT +LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, +DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY +THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT +(INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE +OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. +*) + +fun AOT_item_by_name name = Option.map fst (List.find (fn (_,n) => n = name) AOT_items) +fun AOT_name_of_item id = Option.map snd (List.find (fn (i,_) => id = i) AOT_items) + +val print_AOT_syntax = Attrib.setup_config_bool @{binding "show_AOT_syntax"} (K false) +local + fun AOT_map_translation b (name:string, f) = (name, fn ctxt => + if Config.get ctxt print_AOT_syntax = b + then f ctxt + else raise Match) +in +val AOT_syntax_print_translations = + map (fn (n,f:Proof.context -> term list -> term) => + AOT_map_translation true (n,f)) +val AOT_syntax_typed_print_translations + = map (fn (n,f:Proof.context -> typ -> term list -> term) => + AOT_map_translation true (n,f)) +val AOT_syntax_print_ast_translations + = map (fn (n,f:Proof.context -> Ast.ast list -> Ast.ast) => + AOT_map_translation true (n,f)) +end + +fun AOT_get_item_number name = let +val name = hd (String.fields (equal #"[") name) +val name = String.fields (equal #":") name +in case (AOT_item_by_name (hd name)) of (SOME id) => SOME ( +fold (fn sub => fn str => str ^ "." ^ sub) (tl name) (Int.toString id) +) | _ => NONE end + +fun AOT_print_item_number (name:string) = case (AOT_get_item_number name) + of SOME str => Pretty.writeln (Pretty.str ("PLM item number (" ^ str ^ ")")) + | _ => () + +fun add_AOT_print_rule AOTsyntax raw_rules thy = let + val rules = map (fn (r, s) => let + val head = Ast.head_of_rule (s,r) + in (head, fn ctxt => fn asts => + if Config.get ctxt print_AOT_syntax = AOTsyntax then + let + val orig = (Ast.mk_appl (Ast.Constant head) asts) + val normalized = Ast.normalize ctxt + (fn head' => if head = head' then [(s,r)] else []) orig + in + if orig = normalized then raise Match else normalized + end + else raise Match) + end) raw_rules +in Sign.print_ast_translation rules thy end + +local + +val trans_pat = + Scan.optional + (\<^keyword>\(\ + |-- Parse.!!! (Parse.inner_syntax Parse.name --| + \<^keyword>\)\)) "logic" + -- Parse.inner_syntax Parse.string; + +fun trans_arrow toks = + ((\<^keyword>\\\ || \<^keyword>\<=\) + >> K Syntax.Print_Rule) toks; + +val trans_line = + trans_pat -- Parse.!!! (trans_arrow |-- trans_pat) + >> (fn (left, right) => (left, right)); + +fun add_trans_rule AOTsyntax raw_rules thy = +let + fun mapPair f (x, y) = (f x, f y) + val ctxt = Proof_Context.init_global thy; + val read_root = #1 o dest_Type o + Proof_Context.read_type_name {proper = true, strict = false} ctxt; + val raw_rules = map (mapPair (fn (r, s) => + (Syntax_Phases.parse_ast_pattern ctxt (read_root r, s)))) raw_rules +in add_AOT_print_rule AOTsyntax raw_rules thy end + +val _ = + Outer_Syntax.command + \<^command_keyword>\AOT_syntax_print_translations\ + "add print translation rules for AOT syntax" + (Scan.repeat1 trans_line >> (Toplevel.theory o add_trans_rule true)); + +val _ = + Outer_Syntax.command + \<^command_keyword>\AOT_no_syntax_print_translations\ + "add AOT print translation rules for non-AOT syntax" + (Scan.repeat1 trans_line >> (Toplevel.theory o add_trans_rule false)); +in end + + + +structure AOT_Theorems = Named_Thms( + val name = @{binding "AOT"} + val description = "AOT Theorems" +) +structure AOT_Definitions = Named_Thms( + val name = @{binding "AOT_defs"} + val description = "AOT Definitions" +) +structure AOT_ProofData = Proof_Data +(type T = term option + fun init _ = NONE) +fun AOT_note_theorems thms = Local_Theory.background_theory + (Context.theory_map (fold AOT_Theorems.add_thm + (map Drule.export_without_context (flat thms)))) +fun AOT_note_definitions thms = Local_Theory.background_theory + (Context.theory_map (fold AOT_Definitions.add_thm + (map Drule.export_without_context (flat thms)))) + +structure AOT_DefinedConstants = Theory_Data ( + type T = Termtab.set + val empty = Termtab.empty + val extend = I + val merge = Termtab.merge (K true) +); +fun AOT_note_defined_constant const = + Local_Theory.background_theory (AOT_DefinedConstants.map (Termtab.insert_set const)) + +fun AOT_read_prop nonterminal ctxt prop = (let + val ctxt' = Config.put Syntax.root nonterminal ctxt + val trm = Syntax.parse_term ctxt' prop + val typ = Term.fastype_of (Syntax.check_term ctxt' trm) + val trm = if typ = @{typ prop} then trm else HOLogic.mk_Trueprop trm + in trm end) +fun AOT_read_term nonterminal ctxt prop = (let + val ctxt' = Config.put Syntax.root nonterminal ctxt + val trm = Syntax.parse_term ctxt' prop + in trm end) +fun AOT_check_prop nonterminal ctxt prop = + AOT_read_prop nonterminal ctxt prop |> Syntax.check_prop ctxt; + +let +fun close_form t = + fold_rev (fn (s, T) => fn t => HOLogic.mk_all (s, T, t)) (Term.add_frees t []) t +fun remove_case_prod (Const (\<^const_name>\case_prod\, _) $ x) = remove_case_prod x + | remove_case_prod (Const (\<^const_name>\case_unit\, _) $ x) = remove_case_prod x + | remove_case_prod (Abs (a,b,c)) = Abs (a,b,remove_case_prod c) + | remove_case_prod x = x + +fun AOT_define_id (bnd,str,mx) (lhs,rhs,trm) ctxt = +let + val bnd_str = Binding.name_of bnd + val syn_typ = Syntax.parse_typ ctxt str + + val lhs = remove_case_prod lhs + val rhs = remove_case_prod rhs + + fun filter (name,_) = not (name = \<^const_name>\case_prod\ + orelse name = \<^const_name>\case_unit\) + val (const_name, const_typ) = case (List.filter filter (Term.add_consts lhs [])) + of [const] => const + | _ => raise Term.TERM + ("Expected a single constant on the LHS of the definition.", [lhs]) + + val _ = Long_Name.base_name const_name = Long_Name.base_name bnd_str + orelse raise Term.TERM ("Left-hand side does not contain the definiens.", [lhs]) + + val (lhs_abs_vars, _) = Term.strip_abs lhs + val (rhs_abs_vars, rhs_abs_body) = Term.strip_abs rhs + val _ = lhs_abs_vars = rhs_abs_vars orelse raise Term.TERM + ("Expected the LHS and RHS to abstract over the same free variables.", [lhs, rhs]) + val body = rhs_abs_body + + val witness = fold_rev (fn (s, T) => fn t => Term.absfree (s, T) t) + lhs_abs_vars body + val witness = Syntax.check_term ctxt witness + + (* Construct the choice specification theorem. *) + val thm = + let + val cname = Long_Name.base_name const_name + val vname = if Symbol_Pos.is_identifier cname then cname else "x" + val trm = @{const Trueprop} $ HOLogic.mk_exists (vname, const_typ, + Term.abstract_over (Const (const_name, const_typ), trm)) + val cwitness = Thm.cterm_of ctxt witness + val witness_exI = Thm.instantiate' + [SOME (Thm.ctyp_of_cterm cwitness)] + [NONE,SOME cwitness] exI + val simps = [ + @{thm AOT_model_id_def}, + @{thm AOT_model_nondenoing}, + @{thm AOT_model_denotes_prod_def}, + @{thm AOT_model_denotes_unit_def}, + @{thm case_unit_Unity} + ] + val thm = (Goal.prove ctxt [] [] trm (fn _ => + resolve_tac ctxt [witness_exI] 1 + THEN simp_tac (ctxt addsimps simps) 1)) + val match = Thm.match (Thm.cprop_of thm, Thm.cterm_of ctxt trm) + in Drule.instantiate_normalize match thm end + + (* Add the choice specification and cleanup and export the resulting theorem. *) + val oldctxt = ctxt + val (thm, ctxt) = Local_Theory.background_theory_result (fn lthy => let + val lthy = lthy |> Sign.add_consts [(bnd,const_typ,Mixfix.NoSyn)] |> + Sign.syntax true Syntax.mode_default [(bnd_str,syn_typ,mx)] |> + add_AOT_print_rule true [ + (Ast.Constant bnd_str, Ast.Constant (Lexicon.mark_const const_name)) + ] + in + Choice_Specification.add_specification [("",bnd_str,false)] (lthy, thm) + |> apsnd Drule.export_without_context |> swap + end) (Proof_Context.concealed ctxt) +in + (ctxt |> Proof_Context.restore_naming oldctxt, thm) +end + +fun AOT_define_equiv (bnd,str,mx) (lhs,rhs,trm) ctxt = +let + val bnd_str = Binding.name_of bnd + val syn_typ = Syntax.parse_typ ctxt str + val (const_name, const_typ) = case (Term.add_consts lhs []) of [const] => const + | _ => raise Term.TERM + ("Expected a single constant on the LHS of the definition.", [lhs]) + + (* TODO: figure out how to properly compare the constant name with the binding *) + val _ = Long_Name.base_name const_name = Long_Name.base_name bnd_str + orelse raise Term.TERM ("Left-hand side does not contain the definiens.", [lhs]) + + (* Construct a witness for the choice specification theorem. *) + val frees = Term.add_frees trm [] + val witness = let + val w = case Term.variant_frees trm [("w", @{typ w})] of [w] => w + | _ => raise Fail "Unexpected." + in + \<^const>\AOT_model_proposition_choice\ $ + Term.absfree w (\<^const>\AOT_model_valid_in\ $ Free w $ rhs) + end + val witness = fold_rev (fn (s, T) => fn t => Term.absfree (s, T) t) + (List.rev frees) witness + + (* Construct the choice specification theorem. *) + val thm = + let + val cname = Long_Name.base_name const_name + val vname = if Symbol_Pos.is_identifier cname then cname else "x" + val trm = @{const Trueprop} $ HOLogic.mk_exists (vname, const_typ, + Term.abstract_over (Const (const_name, const_typ), close_form trm)) + val cwitness = Thm.cterm_of ctxt witness + val witness_exI = Thm.instantiate' + [SOME (Thm.ctyp_of_cterm cwitness)] + [NONE,SOME cwitness] exI + val simps = [ + @{thm AOT_model_equiv_def}, + @{thm AOT_model_proposition_choice_simp} + ] + val thm = (Goal.prove ctxt [] [] trm (fn _ => + resolve_tac ctxt [witness_exI] 1 + THEN simp_tac (ctxt addsimps simps) 1)) + val match = Thm.match (Thm.cprop_of thm, Thm.cterm_of ctxt trm) + in + Drule.instantiate_normalize match thm + end + + (* Add the choice specification and cleanup and export the resulting theorem. *) + val oldctxt = ctxt + val (thm, ctxt) = Proof_Context.concealed ctxt |> + Local_Theory.background_theory_result (fn lthy => let + fun inst_all thy (name,typ) thm = + let + val cv = Thm.global_cterm_of thy (Free (name,typ)) + val cT = Thm.global_ctyp_of thy typ + val spec' = Thm.instantiate' [SOME cT] [NONE, SOME cv] spec + in thm RS spec' end + fun remove_alls frees (thy, thm) = (thy, fold (inst_all thy) frees thm) + val lthy = lthy |> Sign.add_consts [(bnd,const_typ,Mixfix.NoSyn)] |> + Sign.syntax true Syntax.mode_default [(bnd_str,syn_typ,mx)] |> + add_AOT_print_rule true [ + (Ast.Constant bnd_str, Ast.Constant (Lexicon.mark_const const_name)) + ] + in + Choice_Specification.add_specification [("",bnd_str,false)] (lthy, thm) + |> remove_alls frees |> apsnd Drule.export_without_context |> swap + end) +in + (ctxt |> Proof_Context.restore_naming oldctxt, thm) +end + +fun AOT_define (((bnd,str,mx),(thmbind,thmattrs)),defprop) ctxt = + let + val bnd_str = Binding.name_of bnd + val syn_typ = Syntax.parse_typ ctxt str + + (* Add a generic constant and the requested syntax to a temporary context. *) + val thy = Proof_Context.theory_of ctxt + val thy' = Sign.add_consts [(bnd, @{typ 'a}, Mixfix.NoSyn)] thy + val thy' = Sign.syntax true Syntax.mode_default [(bnd_str,syn_typ,mx)] thy' + (* Try to parse the definition using the temporary context. *) + val trm = AOT_check_prop @{nonterminal AOT_prop} + (Proof_Context.init_global thy') defprop + (* Extract lhs, rhs and the full definition from the parsed proposition + and delegate. *) + val (ctxt, thm) = case trm of (Const (\<^const_name>\HOL.Trueprop\, _) $ + (trm as Const (n, _) $ lhs $ rhs)) => + if n = \<^const_name>\AOT_model_equiv_def\ then + AOT_define_equiv (bnd,str,mx) (lhs,rhs,trm) ctxt + else if n = \<^const_name>\AOT_model_id_def\ then + AOT_define_id (bnd,str,mx) (lhs,rhs,trm) ctxt + else + raise Term.TERM ("Expected AOT definition.", [trm]) + | _ => raise Term.TERM ("Expected AOT definition.", [trm]) + val thmbind = if Binding.is_empty thmbind then bnd else thmbind + val _ = AOT_print_item_number (Binding.name_of thmbind) + in + ctxt |> Local_Theory.note ((thmbind, thmattrs), [thm]) |> snd |> + AOT_note_theorems [[thm]] |> AOT_note_definitions [[thm]] |> + AOT_note_defined_constant + (Proof_Context.read_const {proper=true,strict=true} ctxt (Binding.name_of bnd)) + end +in +Outer_Syntax.local_theory +@{command_keyword AOT_define} +"AOT definition by equivalence." +(Parse.const_binding -- Parse_Spec.opt_thm_name ":" -- Parse.prop >> AOT_define) +end; + +(* this is a stripped down version of Expression.read_statement + that mainly replaces Syntax.parse_prop with AOT_read_prop + and drops locale includes *) +local + +fun mk_type T = (Logic.mk_type T, []); +fun mk_propp (p, pats) = (Type.constraint propT p, pats); + +fun dest_type (T, []) = Logic.dest_type T + | dest_type _ = raise Fail "Unexpected." +fun dest_propp (p, pats) = (p, pats); + +fun finish_fixes (parms: (string * typ) list) = map (fn (binding, _, mx) => + let val x = Binding.name_of binding + in (binding, AList.lookup (op =) parms x, mx) end); + +fun finish_elem parms (Element.Fixes fixes) = Element.Fixes (finish_fixes parms fixes) + | finish_elem _ (Element.Constrains _) = Element.Constrains [] + | finish_elem _ (Element.Assumes asms) = Element.Assumes asms + | finish_elem _ (Element.Defines defs) = Element.Defines defs + | finish_elem _ (elem as Element.Notes _) = elem + | finish_elem _ (elem as Element.Lazy_Notes _) = elem; + +fun extract_elem (Element.Fixes fixes) = map (#2 #> the_list #> map mk_type) fixes + | extract_elem (Element.Constrains csts) = map (#2 #> single #> map mk_type) csts + | extract_elem (Element.Assumes asms) = map (#2 #> map mk_propp) asms + | extract_elem (Element.Defines defs) = + map (fn (_, (t, ps)) => [mk_propp (t, ps)]) defs + | extract_elem (Element.Notes _) = [] + | extract_elem (Element.Lazy_Notes _) = []; + +fun restore_elem (Element.Fixes fixes, css) = + (fixes ~~ css) |> map (fn ((x, _, mx), cs) => + (x, cs |> map dest_type |> try hd, mx)) |> Element.Fixes + | restore_elem (Element.Constrains csts, css) = + (csts ~~ css) |> map (fn ((x, _), cs) => + (x, cs |> map dest_type |> hd)) |> Element.Constrains + | restore_elem (Element.Assumes asms, css) = + (asms ~~ css) |> map (fn ((b, _), cs) => + (b, map dest_propp cs)) |> Element.Assumes + | restore_elem (Element.Defines defs, css) = + (defs ~~ css) |> map (fn ((b, _), [c]) => + (b, dest_propp c) | _ => raise Fail "Unexpected") |> Element.Defines + | restore_elem (elem as Element.Notes _, _) = elem + | restore_elem (elem as Element.Lazy_Notes _, _) = elem; + +fun prep (_, pats) (ctxt, t :: ts) = + let + val ctxt' = Proof_Context.augment t ctxt + in + ((t, Syntax.check_props + (Proof_Context.set_mode Proof_Context.mode_pattern ctxt') pats), + (ctxt', ts)) + end + | prep _ _ = raise Fail "Unexpected" + +fun check cs ctxt = + let + val (cs', (ctxt', _)) = fold_map prep cs + (ctxt, Syntax.check_terms + (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) (map fst cs)); + in (cs', ctxt') end; + +fun check_autofix elems concl ctxt = + let + val elem_css = map extract_elem elems; + val concl_cs = (map o map) mk_propp (map snd concl); + (* Type inference *) + val (css', ctxt') = + (fold_burrow o fold_burrow) check (elem_css @ [concl_cs]) ctxt; + val (elem_css', concl_cs') = chop (length elem_css) css' |> apsnd the_single; + in + ((map restore_elem (elems ~~ elem_css'), + map fst concl ~~ concl_cs'), ctxt') + end; + +fun prepare_stmt prep_prop ctxt stmt = + (case stmt of + Element.Shows raw_shows => + raw_shows |> (map o apsnd o map) (fn (t, ps) => + (prep_prop (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) t, + map (prep_prop (Proof_Context.set_mode Proof_Context.mode_pattern ctxt)) ps)) + | Element.Obtains _ => raise Fail "unsupported"); + +fun parse_elem prep_typ prep_term ctxt = + Element.map_ctxt + {binding = I, + typ = prep_typ ctxt, + term = prep_term (Proof_Context.set_mode Proof_Context.mode_schematic ctxt), + pattern = prep_term (Proof_Context.set_mode Proof_Context.mode_pattern ctxt), + fact = I, + attrib = I}; + +fun declare_elem prep_var (Element.Fixes fixes) ctxt = + let val (vars, _) = fold_map prep_var fixes ctxt + in ctxt |> Proof_Context.add_fixes vars |> snd end + | declare_elem prep_var (Element.Constrains csts) ctxt = + ctxt |> fold_map (fn (x, T) => + prep_var (Binding.name x, SOME T, NoSyn)) csts |> snd + | declare_elem _ (Element.Assumes _) ctxt = ctxt + | declare_elem _ (Element.Defines _) ctxt = ctxt + | declare_elem _ (Element.Notes _) ctxt = ctxt + | declare_elem _ (Element.Lazy_Notes _) ctxt = ctxt; + +fun prep_full_context_statement + parse_prop prop_root elem_root raw_elems raw_stmt ctxt1 = + let + + fun prep_elem raw_elem ctxt = + let + val ctxt' = ctxt + |> Context_Position.set_visible false + |> declare_elem Proof_Context.read_var raw_elem + |> Context_Position.restore_visible ctxt; + val elems' = parse_elem Syntax.parse_typ (parse_prop elem_root) ctxt' raw_elem; + in (elems', ctxt') end; + + val fors = fold_map Proof_Context.read_var [] ctxt1 |> fst; + val ctxt2 = ctxt1 |> Proof_Context.add_fixes fors |> snd; + val ctxt3 = ctxt2 + + fun prep_stmt elems ctxt = check_autofix elems + (prepare_stmt (parse_prop prop_root) ctxt raw_stmt) ctxt; + + val ((elems', concl), ctxt4) = ctxt3 + |> fold_map prep_elem raw_elems + |-> prep_stmt; + + (* parameters from expression and elements *) + + val xs = maps (fn Element.Fixes fixes => + map (Variable.check_name o #1) fixes | _ => []) + (Element.Fixes fors :: elems'); + val (parms, _) = fold_map Proof_Context.inferred_param xs ctxt4; + + val elems'' = map (finish_elem parms) elems'; + + in (elems'', concl) end; +in + +fun read_statement prop_root elem_root raw_elems raw_stmt ctxt = + let + val (elems, concl) = prep_full_context_statement + AOT_read_prop prop_root elem_root raw_elems raw_stmt ctxt; + val ctxt' = Proof_Context.set_stmt true ctxt + val (elems, ctxt') = fold_map Element.activate elems ctxt' + val _ = Proof_Context.restore_stmt ctxt ctxt' + in (concl, elems, ctxt) end; + +end + +(* End of Expression.read_statement variant. *) + +val long_keyword = + Parse_Spec.includes >> K "" || + Parse_Spec.long_statement_keyword; + +val long_statement = + Scan.optional (Parse_Spec.opt_thm_name ":" --| Scan.ahead long_keyword) + Binding.empty_atts -- + Scan.optional Parse_Spec.includes [] --Parse_Spec.long_statement + >> (fn ((binding, includes), (elems, concl)) => + (true, binding, includes, elems, concl)) +val short_statement = + Parse_Spec.statement -- Parse_Spec.if_statement -- Parse.for_fixes + >> (fn ((shows, assumes), fixes) => + (false, Binding.empty_atts, [], [Element.Fixes fixes, Element.Assumes assumes], + Element.Shows shows)) + +fun setupStrictWorld ctxt = let +(* TODO: ideally not just a fixed name, but a variant name... *) +val (v,ctxt) = Proof_Context.add_fixes + [(Binding.make ("ws", Position.none), SOME @{typ w}, Mixfix.NoSyn)] ctxt |> + apfst the_single +in AOT_ProofData.put (SOME (Free (v, @{typ w}))) ctxt end +fun setupWeakWorld ctxt = AOT_ProofData.put (SOME @{const w\<^sub>0}) ctxt + +fun mapStmt mapTerm _ (Element.Shows x) = + Element.Shows (map (map (fn (trm, trms) => + (mapTerm trm, map mapTerm trms)) |> apsnd) x) + | mapStmt mapTerm mapTyp (Element.Obtains x) = + Element.Obtains (map ((fn (x, y) => + (map (fn (a,b,c) => (a, Option.map mapTyp b, c)) x, map mapTerm y)) |> apsnd) x) +fun mapCtxt mapTerm mapTyp ctxtElem = Element.map_ctxt + {attrib = I, binding = I, fact = I, pattern = mapTerm, term = mapTerm, typ = mapTyp} + ctxtElem + +fun AOT_theorem_cmd axiom modallyStrict long afterQed thmBinding + includes assumptions shows int ctxt = + let + val ctxt = Bundle.includes_cmd includes ctxt + val ctxt = if modallyStrict then setupStrictWorld ctxt else setupWeakWorld ctxt + val root = if axiom + then (if modallyStrict + then @{nonterminal "AOT_axiom"} + else @{nonterminal "AOT_act_axiom"}) + else @{nonterminal AOT_prop} + val (stmts,assumptions,ctxt) = read_statement + root @{nonterminal AOT_prop} assumptions shows ctxt + val _ = AOT_print_item_number (Binding.name_of (fst thmBinding)) + val _ = fold (fn ((bnd,_),_) => fn _ => + AOT_print_item_number (Binding.name_of bnd)) stmts () + in + Specification.theorem long "AOT_theorem" NONE afterQed thmBinding [] + assumptions (Element.Shows stmts) int ctxt + end + +fun AOT_theorem spec note axiom modallyStrict descr = + Outer_Syntax.local_theory_to_proof' spec ("state " ^ descr) + ((long_statement || short_statement) >> + (fn (long, binding, includes, elems, concl) => + (AOT_theorem_cmd axiom modallyStrict long + (if note then AOT_note_theorems else K I) binding includes elems concl))); + +val _ = AOT_theorem \<^command_keyword>\AOT_lemma\ false false true + "AOT modally-strict lemma"; +val _ = AOT_theorem \<^command_keyword>\AOT_theorem\ true false true + "AOT modally-strict theorem"; +val _ = AOT_theorem \<^command_keyword>\AOT_act_lemma\ false false false + "AOT modally-weak lemma"; +val _ = AOT_theorem \<^command_keyword>\AOT_act_theorem\ true false false + "AOT modally-weak theorem" +val _ = AOT_theorem \<^command_keyword>\AOT_axiom\ true true true + "AOT modally-strict axiom"; +val _ = AOT_theorem \<^command_keyword>\AOT_act_axiom\ true true false + "AOT modally-weak axiom" + +local +val structured_statement = + Parse_Spec.statement -- Parse_Spec.if_statement -- Parse.for_fixes + >> (fn ((shows, assumes), fixes) => (fixes, assumes, shows)); + +fun prep_stmt (fixes, assumes, shows) ctxt = let + val (concl, elems, _) = read_statement + @{nonterminal AOT_prop} @{nonterminal AOT_prop} + [Element.Fixes fixes, Element.Assumes assumes] (Element.Shows shows) ctxt + val (fixes, assumes) = + (fn ([Element.Fixes fixes, Element.Assumes assumes]) => (fixes, assumes) + | _ => raise Fail "Unexpected.") elems + fun mapAttr (a,[]) = (a,[]) + | mapAttr _ = raise Match (* Unimplemented *) + val assumes = map (mapAttr |> apfst) assumes + val concl = map (mapAttr |> apfst) concl +in (fixes, assumes, concl) end + +fun gen_cmd kind stmt int state = let + val (fixes, assumes, shows) = prep_stmt stmt (Proof.context_of state) + in (kind true NONE (K I) fixes assumes shows int) state end +in +val _ = + Outer_Syntax.command \<^command_keyword>\AOT_show\ + "state local AOT goal, to refine pending subgoals" + (structured_statement >> (fn stmt => + Toplevel.proof' (fn int => (gen_cmd Proof.show stmt int #> #2)))); +val _ = + Outer_Syntax.command \<^command_keyword>\AOT_thus\ "alias of \"then AOT_show\"" + (structured_statement >> (fn stmt => + Toplevel.proof' (fn int => Proof.chain #> (gen_cmd Proof.show stmt int #> #2)))); +val _ = + Outer_Syntax.command \<^command_keyword>\AOT_have\ "state local AOT goal" + (structured_statement >> (fn stmt => + Toplevel.proof' (fn int => (gen_cmd Proof.have stmt int #> #2)))); +val _ = + Outer_Syntax.command \<^command_keyword>\AOT_hence\ "alias of \"then AOT_have\"" + (structured_statement >> (fn stmt => + Toplevel.proof' (fn int => Proof.chain #> (gen_cmd Proof.have stmt int #> #2)))); +val _ = + Outer_Syntax.command \<^command_keyword>\AOT_modally_strict {\ + "begin explicit AOT modally-strict proof block" + (Scan.succeed (Toplevel.proof (fn state => (Proof.map_context (fn ctxt => let + val v = singleton (Variable.variant_frees ctxt []) ("ws", @{typ w}) |> fst + val (_,ctxt) = Proof_Context.add_fixes + [(Binding.make (v, Position.none), SOME @{typ w}, Mixfix.NoSyn)] ctxt + val ctxt = AOT_ProofData.put (SOME (Free (v, @{typ w}))) ctxt + in ctxt end) (Proof.begin_block state))))); +val _ = + Outer_Syntax.command \<^command_keyword>\AOT_actually {\ + "begin explicit AOT modally-fragile proof block" + (Scan.succeed (Toplevel.proof (fn state => (Proof.map_context (fn ctxt => let + val ctxt = AOT_ProofData.put (SOME (@{const w\<^sub>0})) ctxt + in ctxt end) (Proof.begin_block state))))); +val _ = + Outer_Syntax.command \<^command_keyword>\AOT_assume\ "assume AOT propositions" + (structured_statement >> (fn stmt => + Toplevel.proof' (fn _ => fn state => (let + val (fixes, assumes, shows) = prep_stmt stmt (Proof.context_of state) + in Proof.assume fixes (map snd assumes) shows state end)))); +val _ = + Outer_Syntax.command \<^command_keyword>\AOT_obtain\ "generalized AOT elimination" + (Parse.parbinding -- Scan.optional (Parse.vars --| Parse.where_) [] -- + structured_statement >> (fn ((a, b), stmt) => Toplevel.proof' + (fn int => fn state => (let + val ctxt = Proof.context_of state + val b = map (fn (a,b,c) => (a,Option.map (Syntax.read_typ ctxt) b, c)) b + val bnds = map (fn (a,_,_) => a) b + val (_,ctxt) = Variable.add_fixes_binding bnds ctxt + val (fixes, assumes, shows) = prep_stmt stmt ctxt + in Obtain.obtain a b fixes (map snd assumes) shows int state end)))); + +end + +structure AOT_no_atp = Named_Thms( + val name = @{binding "AOT_no_atp"} + val description = "AOT Theorem Blacklist" +) + +val _ = + Outer_Syntax.command \<^command_keyword>\AOT_sledgehammer\ + "sledgehammer restricted to AOT abstraction layer" + (Scan.succeed (Toplevel.keep_proof (fn state => let +val params = Sledgehammer_Commands.default_params (Toplevel.theory_of state) [] +val ctxt = Toplevel.context_of state +fun all_facts_of ctxt = + let + val thy = Proof_Context.theory_of ctxt; + val transfer = Global_Theory.transfer_theories thy; + val global_facts = Global_Theory.facts_of thy; + in + (Facts.dest_all (Context.Proof ctxt) false [] global_facts + |> maps Facts.selections + |> map (apsnd transfer)) + end +val facts = all_facts_of ctxt +val add_facts = filter + (fn fact => AOT_Theorems.member ctxt (snd fact)) facts |> map (fn (x,_) => (x,[])) +val del_facts = filter + (fn fact => AOT_no_atp.member ctxt (snd fact)) facts |> map (fn (x,_) => (x,[])) +val ctxt = Toplevel.proof_of state +val _ = Sledgehammer.run_sledgehammer params Sledgehammer_Prover.Normal NONE 1 + {add = add_facts, del = del_facts, only = false} ctxt +in () end))) + +val _ = + Outer_Syntax.command \<^command_keyword>\AOT_sledgehammer_only\ + "sledgehammer restricted to AOT abstraction layer" + (Scan.succeed (Toplevel.keep_proof (fn state => let +val params = Sledgehammer_Commands.default_params (Toplevel.theory_of state) [] +val ctxt = Toplevel.context_of state +fun all_facts_of ctxt = + let + val thy = Proof_Context.theory_of ctxt; + val transfer = Global_Theory.transfer_theories thy; + val local_facts = Proof_Context.facts_of ctxt; + val global_facts = Global_Theory.facts_of thy; + in + (Facts.dest_all (Context.Proof ctxt) false [global_facts] local_facts + |> maps Facts.selections + |> map (apsnd transfer) |> map fst) @ + (Facts.dest_all (Context.Proof ctxt) false [] global_facts + |> maps Facts.selections + |> map (apsnd transfer) + |> filter (AOT_Theorems.member ctxt o snd) |> map fst) + end +val facts = all_facts_of ctxt +val result = map (fn x => (x,[])) facts +val ctxt = Toplevel.proof_of state +val _ = Sledgehammer.run_sledgehammer params Sledgehammer_Prover.Normal NONE 1 + {add = result, del = [], only = true} ctxt +in () end))) + +local + +fun readTermPattern ctxt str = (let + val trm = case try (AOT_read_term @{nonterminal \'} ctxt) str of SOME x => x + | NONE => + (case try (AOT_read_term @{nonterminal \'} ctxt) str of SOME x => x + | NONE => (AOT_read_term @{nonterminal AOT_prop} ctxt) str) + val trm = Syntax.check_term ctxt trm + fun varifyTerm (Const (\<^const_name>\AOT_term_of_var\, Type ("fun", [_, t])) $ + Free (x, _)) = Var ((x, 0), t) + | varifyTerm (Free (x,t)) = Var ((x, 0), t) + | varifyTerm (x $ y) = varifyTerm x $ varifyTerm y + | varifyTerm (Abs (a, b, c)) = Abs(a, b, varifyTerm c) + | varifyTerm z = z + val trm = Term.map_types + (Term.map_type_tfree (fn (str,sort) => TVar ((str, 0), sort))) trm + val trm = varifyTerm trm +in trm end) + +fun parseCriterion ctxt (Find_Theorems.Simp crit) = + (Find_Theorems.Simp (readTermPattern ctxt crit)) + | parseCriterion ctxt (Find_Theorems.Pattern crit) = + (Find_Theorems.Pattern (readTermPattern ctxt crit)) + | parseCriterion _ (Find_Theorems.Name x) = (Find_Theorems.Name x) + | parseCriterion _ Find_Theorems.Intro = Find_Theorems.Intro + | parseCriterion _ Find_Theorems.Elim = Find_Theorems.Elim + | parseCriterion _ Find_Theorems.Dest = Find_Theorems.Dest + | parseCriterion _ Find_Theorems.Solves = Find_Theorems.Solves + + +fun pretty_criterion ctxt (b, c) = + let + fun prfx s = if b then s else "-" ^ s; + in + (case c of + Find_Theorems.Name name => Pretty.str (prfx "name: " ^ quote name) + | Find_Theorems.Intro => Pretty.str (prfx "intro") + | Find_Theorems.Elim => Pretty.str (prfx "elim") + | Find_Theorems.Dest => Pretty.str (prfx "dest") + | Find_Theorems.Solves => Pretty.str (prfx "solves") + | Find_Theorems.Simp pat => Pretty.block [Pretty.str (prfx "simp:"), Pretty.brk 1, + Pretty.quote (Syntax.pretty_term ctxt (Term.show_dummy_patterns pat))] + | Find_Theorems.Pattern pat => Pretty.enclose (prfx "\"") "\"" + [Syntax.pretty_term ctxt (Term.show_dummy_patterns pat)]) + end; + +datatype query = Criterion of (bool*string Find_Theorems.criterion) list | + Item of (int*string list) + +fun pretty_theorems ctxt opt_lim rem_dups raw_spec = +let + fun pretty_ref ctxt thmref = + let + val (name, sel) = + (case thmref of + Facts.Named ((name, _), sel) => (name, sel) + | Facts.Fact _ => raise Fail "Illegal literal fact"); + val item = case + AOT_get_item_number (Binding.name_of (Binding.qualified_name name)) + of SOME str => [Pretty.str ("("^str^")"), Pretty.str ":", Pretty.brk 1] + | _ => [] + in + item @ [Pretty.marks_str (#1 (Proof_Context.markup_extern_fact ctxt name), name), + Pretty.str (Facts.string_of_selection sel), Pretty.str ":", Pretty.brk 1] + end; + fun tailToStr delim tail = (fold (fn field => fn str => str^delim^field) tail "") + fun pretty_item (id, sub) = Pretty.str ( + "item: "^quote (Int.toString id^tailToStr "." sub)^ + " (name: "^quote ((the (AOT_name_of_item id))^tailToStr ":" sub)^")") + fun pretty_thm ctxt (thmref, thm) = + Pretty.block (pretty_ref ctxt thmref @ [Thm.pretty_thm ctxt thm]) + val (spec, prefix) = (case raw_spec of Item (item, tail) => + (case AOT_name_of_item item of SOME name => + let val fullName = name^tailToStr ":" tail + in ([(true, Find_Theorems.Name fullName)], SOME fullName) end + | _ => raise Fail "Unknown PLM item number.") + | Criterion spec => (spec, NONE)) + val criteria = map (apsnd (parseCriterion ctxt)) spec + val (opt_found, _) = Find_Theorems.find_theorems ctxt NONE (SOME 0) rem_dups criteria + val (_, theorems) = Find_Theorems.find_theorems ctxt NONE opt_found rem_dups criteria + val lim = the_default + (Options.default_int \<^system_option>\find_theorems_limit\) opt_lim; + val theorems = filter (fn (_,thm) => AOT_Theorems.member ctxt thm) theorems + val theorems = case prefix of SOME prefix => + filter (fn (Facts.Named ((name, _), _), _) => + let val unqualified = (Binding.name_of (Binding.qualified_name name)) + in + String.isPrefix prefix unqualified + andalso (String.size unqualified <= String.size prefix orelse + let + val delim = String.sub (unqualified, String.size prefix) + in delim = #":" orelse delim = #"[" end) + end + | _ => false) theorems + | _ => theorems + val found = length theorems + val theorems = drop (Int.max (found - lim, 0)) theorems + val returned = length theorems + val tally_msg = + (if found <= lim then "displaying " ^ string_of_int returned ^ " theorem(s)" + else "found " ^ string_of_int found ^ " theorem(s)" ^ + (if returned < found + then " (" ^ string_of_int returned ^ " displayed)" + else "")); + val position_markup = Position.markup (Position.thread_data ()) Markup.position; + val pretty = Pretty.block + (Pretty.fbreaks + (Pretty.mark position_markup (Pretty.keyword1 "AOT_find_theorems") :: + ((case raw_spec of Item (id, sub) => + [pretty_item (id, sub)] | _ => map (pretty_criterion ctxt) criteria)))) :: + Pretty.str "" :: + (if null theorems then [Pretty.str "found nothing"] + else + Pretty.str (tally_msg ^ ":") :: + grouped 10 Par_List.map (Pretty.item o single o pretty_thm ctxt) (rev theorems)) +in + pretty |> Pretty.fbreaks |> curry Pretty.blk 0 +end + +val options = + Scan.optional + (Parse.$$$ "(" |-- + Parse.!!! (Scan.option Parse.nat -- + Scan.optional (Parse.reserved "with_dups" >> K false) true --| Parse.$$$ ")")) + (NONE, true); + +val item_parser = (Parse.nat >> (fn n => (n,[])) || (Parse.string >> ( + fn str => let + val fields = String.fields (equal #".") str + val n = case (Int.fromString (hd fields)) of (SOME n) => n | _ => 0 + in (n, tl fields) end ))) + +val query_parser = + ((Parse.reserved "item" |-- Parse.!!! (Parse.$$$ ":" |-- item_parser) >> Item) || + (Find_Theorems.query_parser >> Criterion)) + fun find_theorems ((opt_lim, rem_dups), spec) st = + (Pretty.writeln + (pretty_theorems (AOT_ProofData.put (SOME (Var (("ws", 0), @{typ w}))) + (Proof.context_of (Find_Theorems.proof_state st))) + opt_lim rem_dups spec)) +in +val _ = + Outer_Syntax.command \<^command_keyword>\AOT_find_theorems\ + "find theorems meeting specified criteria" + (options -- query_parser >> (fn query => + Toplevel.keep (fn st => find_theorems query st))); + +end + +fun setup_AOT_no_atp thy = let +val all_facts_with_AOT_semantics = + let + val transfer = Global_Theory.transfer_theories thy; + val global_facts = Global_Theory.facts_of thy; + in + (Facts.dest_all (Context.Theory thy) false [] global_facts + |> maps Facts.selections + |> map (apsnd transfer) + |> filter (not o AOT_Theorems.member (Proof_Context.init_global thy) o snd)) + end +val all_facts_Main = + let + val transfer = Global_Theory.transfer_theories @{theory Main}; + val global_facts = Global_Theory.facts_of @{theory Main}; + in + (Facts.dest_all (Context.Theory @{theory Main}) false [] global_facts + |> maps Facts.selections + |> map (apsnd transfer)) + end +val facts = filter + (fn (elem,_) => not (List.exists (fn (elem',_) => elem = elem') all_facts_Main)) + all_facts_with_AOT_semantics +val thy = fold + (fn fact => Context.theory_map (AOT_no_atp.add_thm fact)) + (map snd facts) thy +in thy end diff --git a/thys/AOT/AOT_commands.thy b/thys/AOT/AOT_commands.thy new file mode 100644 --- /dev/null +++ b/thys/AOT/AOT_commands.thy @@ -0,0 +1,47 @@ +(*<*) +theory AOT_commands + imports AOT_model "HOL-Eisbach.Eisbach_Tools" + keywords "AOT_define" :: thy_decl + and "AOT_theorem" :: thy_goal + and "AOT_lemma" :: thy_goal + and "AOT_act_theorem" :: thy_goal + and "AOT_act_lemma" :: thy_goal + + and "AOT_axiom" :: thy_goal + and "AOT_act_axiom" :: thy_goal + + and "AOT_assume" :: prf_asm % "proof" + and "AOT_have" :: prf_goal % "proof" + and "AOT_hence" :: prf_goal % "proof" + and "AOT_modally_strict {" :: prf_open % "proof" + and "AOT_actually {" :: prf_open % "proof" + and "AOT_obtain" :: prf_asm_goal % "proof" + and "AOT_show" :: prf_asm_goal % "proof" + and "AOT_thus" :: prf_asm_goal % "proof" + + and "AOT_find_theorems" :: diag + and "AOT_sledgehammer" :: diag + and "AOT_sledgehammer_only" :: diag + and "AOT_syntax_print_translations" :: thy_decl + and "AOT_no_syntax_print_translations" :: thy_decl +begin +(*>*) + +section\Outer Syntax Commands\ + +nonterminal AOT_prop +nonterminal \ +nonterminal \' +nonterminal \ +nonterminal \' +nonterminal "AOT_axiom" +nonterminal "AOT_act_axiom" +ML_file AOT_keys.ML +ML_file AOT_commands.ML +setup\AOT_Theorems.setup\ +setup\AOT_Definitions.setup\ +setup\AOT_no_atp.setup\ + +(*<*) +end +(*>*) \ No newline at end of file diff --git a/thys/AOT/AOT_keys.ML b/thys/AOT/AOT_keys.ML new file mode 100644 --- /dev/null +++ b/thys/AOT/AOT_keys.ML @@ -0,0 +1,909 @@ +val AOT_items = [ + (1, "simple-terms"), + (2, "primitive-expressions"), + (3, "syntax"), + (4, "bnf"), + (5, "conventions2"), + (6, "df-subformula"), + (7, "df-subterm"), + (8, "enc-form-sub"), + (9, "free-in"), + (10, "open-formulas-terms"), + (11, "closures"), + (12, "term-same-type"), + (13, "identity-remark"), + (14, "substitutions"), + (15, "substitutable"), + (16, "alphabetic-variants"), + (17, "two-defs"), + (18, "conventions"), + (19, "conventions3"), + (20, "existence"), + (21, "studied-amb"), + (22, "oa"), + (23, "identity"), + (24, "=-infix"), + (25, "lang-conc-rem"), + (26, "alpha-bet-rem"), + (27, "identity-pre"), + (28, "identity-pre2"), + (29, "convention-rem"), + (30, "pre-exists"), + (31, "exp-sub"), + (32, "exp-sub2"), + (33, "exist-just"), + (34, "no-rel-prop"), + (35, "identity-rem"), + (36, "p-identity-rem"), + (37, "p-identity2-rem"), + (38, "pl"), + (39, "cqt"), + (40, "cqt-just"), + (41, "l-identity"), + (42, "m-fragile"), + (43, "logic-actual"), + (44, "logic-actual-nec"), + (45, "qml"), + (46, "qml-act"), + (47, "descriptions"), + (48, "lambda-predicates"), + (49, "safe-ext"), + (50, "nary-encoding"), + (51, "encoding"), + (52, "nocoder"), + (53, "A-objects"), + (54, "no-immediate-contra"), + (55, "cqt-expl"), + (56, "qml4-remark"), + (57, "coexist-rm"), + (58, "modus-ponens"), + (59, "theoremhood"), + (60, "Box-theoremhood"), + (61, "metarules"), + (62, "non-con-thm-thm"), + (63, "vdash-properties"), + (64, "df-noncon-thm"), + (65, "dependence"), + (66, "rule-gen"), + (67, "rule-convention"), + (68, "RN"), + (69, "preserve-GEN-RN"), + (70, "RN-w-cont"), + (71, "RN-converse-rem"), + (72, "df-rules-formulas"), + (73, "df-rules-terms"), + (74, "if-p-then-p"), + (75, "deduction-theorem"), + (76, "ded-thm-cor"), + (77, "useful-tautologies"), + (78, "dn-i-e"), + (79, "modus-tollens"), + (80, "contraposition"), + (81, "reductio-aa"), + (82, "in-el-i-e"), + (83, "exc-mid"), + (84, "non-contradiction"), + (85, "con-dis-taut"), + (86, "con-dis-i-e"), + (87, "raa-cor"), + (88, "oth-class-taut"), + (89, "intro-elim"), + (90, "rule-eq-df"), + (91, "df-simplify"), + (92, "remark-taut"), + (93, "rule-ui"), + (94, "misuse-substitution"), + (95, "cqt-orig"), + (96, "universal"), + (97, "rereplace-lem"), + (98, "rereplace-rmk"), + (99, "cqt-basic"), + (100, "universal-cor"), + (101, "existential"), + (102, "instantiation"), + (103, "cqt-further"), + (104, "log-prop-prop"), + (105, "safe-ext-rem"), + (106, "exist-nec"), + (107, "t=t-proper"), + (108, "id-rel-nec-equiv"), + (109, "df-equiv-id"), + (110, "rule=E"), + (111, "propositions-lemma"), + (112, "rem-truth"), + (113, "remark-on-tautologies"), + (114, "dr-alphabetic-rules"), + (115, "oa-exist"), + (116, "p-identity-thm2"), + (117, "id-eq"), + (118, "rule=I"), + (119, "rule-id-df-rem"), + (120, "rule-id-df"), + (121, "free-thms"), + (122, "free-thms-rem"), + (123, "ex"), + (124, "all-self="), + (125, "id-nec"), + (126, "term-out"), + (127, "uniqueness"), + (128, "uni-most"), + (129, "nec-exist-!"), + (130, "act-cond"), + (131, "nec-imp-act"), + (132, "act-conj-act"), + (133, "closure-act"), + (134, "RA"), + (135, "meta-rule"), + (136, "logic-actual-rem"), + (137, "ANeg"), + (138, "Act-Basic"), + (139, "act-quant-uniq"), + (140, "fund-cont-desc"), + (141, "hintikka"), + (142, "russell-axiom"), + (143, "!-exists"), + (144, "y-in"), + (145, "act-quant-nec"), + (146, "equi-desc-descA"), + (147, "nec-hintikka-scheme"), + (148, "equiv-desc-eq"), + (149, "equiv-desc-eq2"), + (150, "nec-russell-axiom"), + (151, "actual-desc"), + (152, "!box-desc"), + (153, "dr-alphabetic-thm"), + (154, "cqt-remark"), + (155, "taut-nec"), + (156, "RM"), + (157, "KBasic"), + (158, "rule-sub-lem"), + (159, "rule-sub-nec"), + (160, "rule-sub-remark"), + (161, "KBasic2"), + (162, "T-S5-fund"), + (163, "Act-Sub"), + (164, "S5Basic"), + (165, "derived-S5-rules"), + (166, "BFs"), + (167, "sign-S5-thm"), + (168, "exist-nec2"), + (169, "id-nec2"), + (170, "sc-eq-box-box"), + (171, "west-mmrule"), + (172, "sc-eq-fur"), + (173, "id-act"), + (174, "A-Exists"), + (175, "id-act-desc"), + (176, "pre-en-eq"), + (177, "en-eq"), + (178, "oa-facts"), + (179, "beta-C-meta"), + (180, "beta-free-rem"), + (181, "beta-C-cor"), + (182, "betaC"), + (183, "eta-variants"), + (184, "eta-conversion-lemma1"), + (185, "eta-conversion-lemma2"), + (186, "sub-des-lam"), + (187, "prop-equiv"), + (188, "prop-rem"), + (189, "relations"), + (190, "block-paradox"), + (191, "block-paradox2"), + (192, "propositions"), + (193, "pos-not-equiv-ne"), + (194, "df-relation-negation"), + (195, "rel-neg-T"), + (196, "df-rel-neg-rem"), + (197, "thm-relation-negation"), + (198, "contingent-properties"), + (199, "cont-prop-rem"), + (200, "thm-cont-prop"), + (201, "thm-noncont-e-e"), + (202, "lem-cont-e"), + (203, "thm-cont-e"), + (204, "property-facts"), + (205, "thm-cont-propos"), + (206, "thm-noncont-propos"), + (207, "no-cnac"), + (208, "pos-not-pna"), + (209, "basic-prop"), + (210, "proposition-facts"), + (211, "cont-tf"), + (212, "cont-true-cont"), + (213, "q0cf"), + (214, "q0cf-rem"), + (215, "cont-tf-thm"), + (216, "rem-conting"), + (217, "property-facts1"), + (218, "eq-not-nec"), + (219, "eqnotnec"), + (220, "oa-contingent"), + (221, "df-cont-nec"), + (222, "cont-nec-fact1"), + (223, "cont-nec-fact2"), + (224, "sixteen"), + (225, "o-objects-exist"), + (226, "partition"), + (227, "=E"), + (228, "=E-infix"), + (229, "haecceity-expanded"), + (230, "=E-simple"), + (231, "id-nec3"), + (232, "non-=E"), + (233, "thm-neg=E"), + (234, "id-nec4"), + (235, "id-act2"), + (236, "ord=Eequiv"), + (237, "ord-=E="), + (238, "ind-nec"), + (239, "ord=E"), + (240, "ord=E2"), + (241, "ordnecfail"), + (242, "ab-obey"), + (243, "encoders-are-abstract"), + (244, "denote="), + (245, "denote=rem"), + (246, "A-obj-ex"), + (247, "A-objects!"), + (248, "obj-oth"), + (249, "A-descriptions"), + (250, "can-ind"), + (251, "thm-can-terms2"), + (252, "can-ab2"), + (253, "desc-encode"), + (254, "abstraction-contingent"), + (255, "desc-nec-encode"), + (256, "Box-desc-encode"), + (257, "strict-can"), + (258, "box-phi-a"), + (259, "strict-can-rem2"), + (260, "df-null-uni"), + (261, "null-uni-uniq"), + (262, "df-null-uni-terms"), + (263, "null-uni-facts"), + (264, "null-uni-rmk2"), + (265, "aclassical"), + (266, "aclassical2"), + (267, "kirchner-thm-rem"), + (268, "kirchner-thm"), + (269, "kirchner-thm-cor"), + (270, "prop-prop1"), + (271, "prop-prop2"), + (272, "prop-indis"), + (273, "prop-in-thm"), + (274, "prop-in-f"), + (275, "prop-prop-nec"), + (276, "enc-prop-nec"), + (277, "def-omit"), + (278, "t-dfs"), + (279, "df-fragile-axiom"), + (280, "tvalues-remark"), + (281, "tv-p"), + (282, "tv-p-rem"), + (283, "p-has-!tv"), + (284, "uni-tv"), + (285, "the-tv-p"), + (286, "prop-enc"), + (287, "tv-id"), + (288, "not-can-thm"), + (289, "t-prov-can"), + (290, "T-lem"), + (291, "ab-T"), + (292, "TV-lem1"), + (293, "T-value"), + (294, "TVp-TV"), + (295, "TV-lem2"), + (296, "the-true"), + (297, "T-T-value"), + (298, "two-T"), + (299, "valueof-facts"), + (300, "q-True"), + (301, "exten-p"), + (302, "extof-e"), + (303, "ext-p-tv"), + (304, "set-remark"), + (305, "tv-v-class"), + (306, "naive-logical"), + (307, "exten-property"), + (308, "exten-remark"), + (309, "pre-LawV"), + (310, "unique-ext-of"), + (311, "mat-eq-cont1"), + (312, "membership"), + (313, "in-no-col"), + (314, "mem-exem"), + (315, "fund-thm-class"), + (316, "the-extG"), + (317, "no-imprac"), + (318, "extG-id"), + (319, "mat-eq-cont2"), + (320, "eG-not-sc"), + (321, "ext-features"), + (322, "lawV"), + (323, "Frege-notation"), + (324, "Frege-lawV"), + (325, "no-lawV-pred"), + (326, "frege-members"), + (327, "mem-exem-cor"), + (328, "extG-class"), + (329, "eG-not-sc-rem"), + (330, "res-var"), + (331, "res-var-bound"), + (332, "res-var-df"), + (333, "res-var-term"), + (334, "res-var-bound-reas"), + (335, "res-var-free"), + (336, "res-var-empty"), + (337, "extensionality"), + (338, "df-null"), + (339, "null"), + (340, "null-symbol"), + (341, "null-exists"), + (342, "df-universall"), + (343, "universal-class"), + (344, "df-unions"), + (345, "unions"), + (346, "unions2"), + (347, "unions3"), + (348, "unions-df"), + (349, "unions-prin"), + (350, "df-complements"), + (351, "complements"), + (352, "complements2"), + (353, "df-intersections"), + (354, "intersections"), + (355, "intersections2"), + (356, "intersect-df"), + (357, "int-mem"), + (358, "cond-set-comprehension"), + (359, "res-meta"), + (360, "set-comp-df"), + (361, "t-defs4"), + (362, "class-ab-prin"), + (363, "class-ab-id"), + (364, "separation"), + (365, "separation2"), + (366, "separation-ab"), + (367, "sep-ab-prin"), + (368, "ab-inter"), + (369, "collection"), + (370, "pre-singletons"), + (371, "singletons"), + (372, "single-df"), + (373, "singleton-facts"), + (374, "singleton-facts2"), + (375, "pairs"), + (376, "adjunction"), + (377, "ex-anti-ext"), + (378, "log-obj"), + (379, "df-equiv-cond"), + (380, "thm-equiv-cond"), + (381, "notate-ab"), + (382, "Frege-p"), + (383, "axiom-lines"), + (384, "lem-lines"), + (385, "df-direction"), + (386, "direct-line"), + (387, "direct-line2"), + (388, "direction-df"), + (389, "para-dir"), + (390, "directions"), + (391, "exist-dir"), + (392, "remark-shapes"), + (393, "lem-shapes"), + (394, "df-shape"), + (395, "shape-figure"), + (396, "shape-figure2"), + (397, "shape-df"), + (398, "shape-thm"), + (399, "shapes2"), + (400, "exist-shape"), + (401, "equiv-rel"), + (402, "R1"), + (403, "equiv-rest"), + (404, "lem-Rab"), + (405, "df-ab"), + (406, "ab-exists"), + (407, "ab-exists2"), + (408, "df-by-ab"), + (409, "R-ab"), + (410, "form-remark"), + (411, "form-of"), + (412, "forms-exist"), + (413, "phi-G"), + (414, "phiG-id"), + (415, "phiG-sc"), + (416, "TheFormGFormG"), + (417, "pre-sp"), + (418, "participation"), + (419, "pre-OM-weak"), + (420, "part=ex"), + (421, "om"), + (422, "counter-sp"), + (423, "forms-encode"), + (424, "sp-remark"), + (425, "forms"), + (426, "TheFormGForm"), + (427, "rem-no-par"), + (428, "platonic-being"), + (429, "form-part"), + (430, "third-man"), + (431, "tforms-remark"), + (432, "F-imp-G"), + (433, "imp-facts"), + (434, "tform-of"), + (435, "tforms-exist"), + (436, "tphi-G"), + (437, "tphiG-id"), + (438, "tphiG-sc"), + (439, "tTheFormGFormG"), + (440, "geach"), + (441, "tparticipation"), + (442, "tpre-OM-weak"), + (443, "counter-tpre-rl"), + (444, "tpart=ex"), + (445, "tpta-ex"), + (446, "tom"), + (447, "tcounter-sp"), + (448, "tforms-encode"), + (449, "tforms"), + (450, "tTheFormGForm"), + (451, "SPa-b"), + (452, "tpart-con"), + (453, "tform-syl"), + (454, "tforms-NI"), + (455, "sit-remark"), + (456, "situations"), + (457, "T-sit"), + (458, "possit-sit"), + (459, "true-in-s"), + (460, "lem1"), + (461, "lem1-rem"), + (462, "lem2"), + (463, "sit-identity"), + (464, "sit-part-whole"), + (465, "part"), + (466, "sit-identity2"), + (467, "persistent"), + (468, "pers-prop"), + (469, "df-null-trivial"), + (470, "thm-null-trivial"), + (471, "df-the-null-sit"), + (472, "null-triv-sc"), + (473, "null-triv-facts"), + (474, "cond-prop"), + (475, "pre-comp-sit"), + (476, "comp-sit"), + (477, "can-sit-desc"), + (478, "strict-sit"), + (479, "strict-can-sit"), + (480, "sit-lat"), + (481, "actual"), + (482, "act-and-not-pos"), + (483, "actual-rem"), + (484, "actual-s"), + (485, "comp"), + (486, "act-sit"), + (487, "cons"), + (488, "cons-remark"), + (489, "sit-cons"), + (490, "cons-rigid"), + (491, "pos"), + (492, "sit-pos"), + (493, "pos-cons-sit"), + (494, "sit-classical"), + (495, "rem-pos-world"), + (496, "world"), + (497, "rigid-pw"), + (498, "double-res"), + (499, "true-w"), + (500, "world-pos"), + (501, "world-cons"), + (502, "rigid-truth-at"), + (503, "max"), + (504, "world-max"), + (505, "world=maxpos"), + (506, "maxcon-rem"), + (507, "nec-impl-p"), + (508, "nec-equiv-nec-im"), + (509, "world-closed-lem"), + (510, "world-clo"), + (511, "world-closed"), + (512, "coherent"), + (513, "coh-rem"), + (514, "act-world"), + (515, "remark-w-alpha"), + (516, "pre-walpha"), + (517, "w-alpha"), + (518, "T-world"), + (519, "truth-at-alpha"), + (520, "alpha-world"), + (521, "t-at-alpha-strict"), + (522, "not-act"), + (523, "w-alpha-part"), + (524, "act-world2"), + (525, "fund-lem"), + (526, "fund"), + (527, "nec-dia-w"), + (528, "conj-dist-w"), + (529, "Lew-world-con"), + (530, "two-worlds-exist"), + (531, "non-aw"), + (532, "iterated-modalities"), + (533, "world-equiv"), + (534, "no-contradictions"), + (535, "disj-dist-models"), + (536, "world-skeptic"), + (537, "tv-of-p-at-w"), + (538, "unique-TV-w"), + (539, "T-value-at-w"), + (540, "TW-strict"), + (541, "TheTrueAtW"), + (542, "TW-rigid"), + (543, "TW-Sigma"), + (544, "cor-T-TV"), + (545, "Boxp-TwT"), + (546, "ext-at-w-F"), + (547, "unique-ext-w"), + (548, "ep-w-G"), + (549, "eG-strict"), + (550, "LawV-w"), + (551, "rem-ext-w"), + (552, "w-rel"), + (553, "w-index"), + (554, "df-rigid-rel"), + (555, "rem-on-rigid"), + (556, "rigid-der"), + (557, "rigid-rel-thms"), + (558, "w-ind-equiv"), + (559, "rem-imp-w"), + (560, "impwor"), + (561, "iw-rigid"), + (562, "iw-truth-at"), + (563, "iw-id"), + (564, "false-iw"), + (565, "iw-notmc"), + (566, "iw-p-ext"), + (567, "iw-pext-lem"), + (568, "iw-fund"), + (569, "iw-ecq-no"), + (570, "iw-ds"), + (571, "rem-times"), + (572, "fic-data"), + (573, "fic-meth"), + (574, "story"), + (575, "story-exist"), + (576, "story-remark"), + (577, "stories-truth"), + (578, "strict-char"), + (579, "strict-pre"), + (580, "s-en"), + (581, "char"), + (582, "nat-char"), + (583, "nat-id"), + (584, "fictional"), + (585, "fictional-f"), + (586, "tf-stories"), + (587, "stories-apply1"), + (588, "stories-apply2"), + (589, "stories-apply3"), + (590, "fur-con"), + (591, "pos-fic"), + (592, "leibniz-strands"), + (593, "concepts"), + (594, "con=ab"), + (595, "concept-comp"), + (596, "con-res-x"), + (597, "c-id"), + (598, "sum-pre-def"), + (599, "sum-exists"), + (600, "sum-def"), + (601, "sum-lemma"), + (602, "oplus"), + (603, "sum-property"), + (604, "cid-add"), + (605, "con"), + (606, "con-part"), + (607, "con-id"), + (608, "ic-add"), + (609, "def-ded"), + (610, "fund-leib"), + (611, "leib-23"), + (612, "prod-pre-def"), + (613, "product-exists"), + (614, "prod-def"), + (615, "prod-lemma"), + (616, "otimes"), + (617, "absorption"), + (618, "bd-lattice"), + (619, "boolean-rem1"), + (620, "distrib"), + (621, "comple-df"), + (622, "comple-exist"), + (623, "comple-def"), + (624, "comple-lemma"), + (625, "comple-laws"), + (626, "post-boole"), + (627, "concept-sub"), + (628, "ex-mereology"), + (629, "part-of-c"), + (630, "ppart-of-c"), + (631, "ppart-ax"), + (632, "null-concept"), + (633, "null-facts"), + (634, "mereo-overlap"), + (635, "overlap-thms"), + (636, "mer-sup"), + (637, "underlap"), + (638, "a-underlap"), + (639, "dodge"), + (640, "nnc"), + (641, "nnc-rigid"), + (642, "df-bottom+"), + (643, "no-bottom"), + (644, "df-atom+"), + (645, "atom-part"), + (646, "o-o"), + (647, "nover-rem"), + (648, "noverlap-facts"), + (649, "newsup"), + (650, "ndf-ex"), + (651, "concept-of-G"), + (652, "con-exists"), + (653, "concept-G"), + (654, "conG-strict"), + (655, "conG-lemma"), + (656, "sum3"), + (657, "gen-inc"), + (658, "concepts-not-properties"), + (659, "form=con"), + (660, "con-of-u"), + (661, "con-u-exist"), + (662, "concept-u"), + (663, "Fu-cont"), + (664, "con-u-not-strict"), + (665, "c-uF-Fu"), + (666, "c-uF-Fu2"), + (667, "complete"), + (668, "con-of-u-com"), + (669, "restrict-concepts"), + (670, "con-gen"), + (671, "c-allGF-allGisF"), + (672, "c-allGF-allGisF2"), + (673, "fa-c-af"), + (674, "contain-truth"), + (675, "montague-gq"), + (676, "con-nec"), + (677, "no-counterpart-theory"), + (678, "cor"), + (679, "cor-rem"), + (680, "real-at-facts"), + (681, "appears-at"), + (682, "appear-fact"), + (683, "appear-co"), + (684, "mirror"), + (685, "appear-mir"), + (686, "new-appear-fact"), + (687, "box-appears"), + (688, "new-real-at-fact"), + (689, "lem-con-u"), + (690, "ind-con"), + (691, "concept-u-ind-con"), + (692, "concept-u-ind-con-a"), + (693, "ind-con-rigid"), + (694, "ind-con-appear"), + (695, "w-c"), + (696, "w-c-applies"), + (697, "ind-con-cont-c-F"), + (698, "con-non-G"), + (699, "con-comp"), + (700, "compos"), + (701, "comp-w-c"), + (702, "compos-eq"), + (703, "df-con-u-at-w"), + (704, "con-u-w-exist"), + (705, "con-u-at-w"), + (706, "con-u-strict"), + (707, "lem-con-u-at"), + (708, "ind-con-iff-c-w-u"), + (709, "lem-con-u-alpha"), + (710, "lem-con-u-at-w"), + (711, "con-u-at-comp"), + (712, "rem-con-u-at-w"), + (713, "concepts-not-properties2"), + (714, "counterparts"), + (715, "count-part-con"), + (716, "count-cor"), + (717, "con-u-at-w-count"), + (718, "fund-thm-modal-meta"), + (719, "fund-thm-modal-bi"), + (720, "fund-thm-modal-strict"), + (721, "fund-thm-mod-rem"), + (722, "Dedekind-Peano"), + (723, "1-1-cor"), + (724, "1-1-cor-rem"), + (725, "fFG"), + (726, "eq-1-1"), + (727, "frege-theorem"), + (728, "why-EqE"), + (729, "equi"), + (730, "eq-part"), + (731, "equi-rem"), + (732, "equi-rem-thm"), + (733, "empty-approx"), + (734, "F-u"), + (735, "eqP'"), + (736, "P'-eq"), + (737, "approx-cont"), + (738, "eqE"), + (739, "apE-eqE"), + (740, "eq-part-act"), + (741, "actuallyF"), + (742, "approx-nec"), + (743, "pre-number"), + (744, "numbers"), + (745, "num-tran"), + (746, "pre-Hume"), + (747, "num-tran-rem"), + (748, "two-num-not"), + (749, "num"), + (750, "num-cont"), + (751, "num-uniq"), + (752, "num-def"), + (753, "num-can"), + (754, "numG-not-strict"), + (755, "card"), + (756, "natcard-nec"), + (757, "hume"), + (758, "hume-strict"), + (759, "unotEu"), + (760, "zero"), + (761, "zero-card"), + (762, "eq-num"), + (763, "eq-df-num"), + (764, "card-en"), + (765, "0F"), + (766, "zero="), + (767, "hered"), + (768, "ances-df"), + (769, "ances"), + (770, "anc-her"), + (771, "rem-pre-wances"), + (772, "df-1-1"), + (773, "id-d-R"), + (774, "id-R-thm"), + (775, "w-ances-df"), + (776, "w-ances"), + (777, "wances-her"), + (778, "1-1-R"), + (779, "pre-ind"), + (780, "pre-ind-rem"), + (781, "pred-rem"), + (782, "pred"), + (783, "pred-thm"), + (784, "pred-1-1"), + (785, "assume-anc"), + (786, "no-pred-0"), + (787, "assume1"), + (788, "nnumber"), + (789, "0-n"), + (790, "mod-col-num"), + (791, "0-pred"), + (792, "no-same-succ"), + (793, "induction"), + (794, "suc-num"), + (795, "pred-num"), + (796, "nat-card"), + (797, "pred-func"), + (798, "modal-axiom"), + (799, "modal-just"), + (800, "modal-lemma"), + (801, "th-succ"), + (802, "Frege's-proof"), + (803, "non-classical"), + (804, "def-suc"), + (805, "suc-fact"), + (806, "ind-gnd"), + (807, "suc-thm"), + (808, "numerals"), + (809, "prec-facts"), + (810, "num-num"), + (811, "dig-alt"), + (812, "rigid-restrict"), + (813, "lt-gt"), + (814, "lt-basic"), + (815, "lt-conv"), + (816, "lt-trans"), + (817, "0lt1"), + (818, "num-ord-lt"), + (819, "remQuine"), + (820, "no-num-pred"), + (821, "dot-id-lem"), + (822, "dot-id"), + (823, "dot=eq="), + (824, "integer"), + (825, "1-pi"), + (826, "i-prec"), + (827, "n-n-lemma"), + (828, "num-quant"), + (829, "n-n"), + (830, "num-quant-gen"), + (831, "Functions"), + (832, "total-D"), + (833, "I-thm"), + (834, "func-ex-pre"), + (835, "func-ex"), + (836, "func-nec"), + (837, "fun-not-nec"), + (838, "fcn-res-rem"), + (839, "df-fa"), + (840, "func-thm"), + (841, "eq-func-nec"), + (842, "func-rem"), + (843, "func-obs"), + (844, "function-rem"), + (845, "fmaps-r"), + (846, "fcn-from-ex"), + (847, "fcn-from-ex2"), + (848, "Rmap-res"), + (849, "fcnal-from"), + (850, "fcnl-from-ex"), + (851, "fcnl-from-ex2"), + (852, "fcnl-on"), + (853, "fcnl-on-thm"), + (854, "dom-ra"), + (855, "df-dom"), + (856, "fcn-from-re"), + (857, "fcn-from"), + (858, "fcn-sum"), + (859, "mapping-ex"), + (860, "dom-thm"), + (861, "codom-thm"), + (862, "range-thm"), + (863, "allf-do"), + (864, "R-gen-1-o"), + (865, "ra-onto"), + (866, "carProd"), + (867, "restricted-f"), + (868, "rf-fact"), + (869, "rfun-exist"), + (870, "res-fn-ex"), + (871, "null-fn-fact"), + (872, "res-fn-df"), + (873, "res-fn-fact"), + (874, "func-rest-var"), + (875, "fcn-app-pre"), + (876, "fcn-app-redef"), + (877, "n-ary-op"), + (878, "n-ary-op-rigid"), + (879, "op-lem0"), + (880, "con-fun"), + (881, "proj-fun"), + (882, "op-lem1"), + (883, "pre-comp"), + (884, "op-comp"), + (885, "op-rest-var"), + (886, "recursive"), + (887, "recur-dig"), + (888, "def-rel-add"), + (889, "rec-df-add"), + (890, "gen-recur-df"), + (891, "gen-rec-R"), + (892, "base-rec-thm"), + (893, "gen-gen-rec"), + (894, "gen-rec-thm"), + (895, "trad-recurs-dfs"), + (896, "prim-gen-recur"), + (897, "minimize"), + (898, "2ndPA-inter"), + (899, "2ndPA-der"), + (900, "multi-ord"), + (901, "inf-card"), + (902, "inf-card-exist"), + (903, "df-infty"), + (904, "infty-thms"), + (905, "inf-just"), + (906, "inf-set"), + (907, "inf-set-exist") +] diff --git a/thys/AOT/AOT_misc.thy b/thys/AOT/AOT_misc.thy new file mode 100644 --- /dev/null +++ b/thys/AOT/AOT_misc.thy @@ -0,0 +1,873 @@ +theory AOT_misc + imports AOT_NaturalNumbers +begin + +section\Miscellaneous Theorems\ + +AOT_theorem PossiblyNumbersEmptyPropertyImpliesZero: + \\Numbers(x,[\z O!z & z \\<^sub>E z]) \ x = 0\ +proof(rule "\I") + AOT_have \Rigid([\z O!z & z \\<^sub>E z])\ + proof (safe intro!: "df-rigid-rel:1"[THEN "\\<^sub>d\<^sub>fI"] "&I" "cqt:2"; + rule RN; safe intro!: GEN "\I") + AOT_modally_strict { + fix x + AOT_assume \[\z O!z & z \\<^sub>E z]x\ + AOT_hence \O!x & x \\<^sub>E x\ by (rule "\\C") + moreover AOT_have \x =\<^sub>E x\ using calculation[THEN "&E"(1)] + by (metis "ord=Eequiv:1" "vdash-properties:10") + ultimately AOT_have \x =\<^sub>E x & \x =\<^sub>E x\ + by (metis "con-dis-i-e:1" "con-dis-i-e:2:b" "intro-elim:3:a" "thm-neg=E") + AOT_thus \\[\z O!z & z \\<^sub>E z]x\ using "raa-cor:1" by blast + } + qed + AOT_hence \\\x (Numbers(x,[\z O!z & z \\<^sub>E z]) \ \Numbers(x,[\z O!z & z \\<^sub>E z]))\ + by (safe intro!: "num-cont:2"[unvarify G, THEN "\E"] "cqt:2") + AOT_hence \\x \(Numbers(x,[\z O!z & z \\<^sub>E z]) \ \Numbers(x,[\z O!z & z \\<^sub>E z]))\ + using "BFs:2"[THEN "\E"] by blast + AOT_hence \\(Numbers(x,[\z O!z & z \\<^sub>E z]) \ \Numbers(x,[\z O!z & z \\<^sub>E z]))\ + using "\E"(2) by auto + moreover AOT_assume \\Numbers(x,[\z O!z & z \\<^sub>E z])\ + ultimately AOT_have \\<^bold>\Numbers(x,[\z O!z & z \\<^sub>E z])\ + using "sc-eq-box-box:1"[THEN "\E"(1), THEN "\E", THEN "nec-imp-act"[THEN "\E"]] + by blast + AOT_hence \Numbers(x,[\z \<^bold>\[\z O!z & z \\<^sub>E z]z])\ + by (safe intro!: "eq-num:1"[unvarify G, THEN "\E"(1)] "cqt:2") + AOT_hence \x = #[\z O!z & z \\<^sub>E z]\ + by (safe intro!: "eq-num:2"[unvarify G, THEN "\E"(1)] "cqt:2") + AOT_thus \x = 0\ + using "cqt:2"(1) "rule-id-df:2:b[zero]" "rule=E" "zero:1" by blast +qed + +AOT_define Numbers' :: \\ \ \ \ \\ (\Numbers'''(_,_')\) + \Numbers'(x, G) \\<^sub>d\<^sub>f A!x & G\ & \F (x[F] \ F \\<^sub>E G)\ +AOT_theorem Numbers'equiv: \Numbers'(x,G) \ A!x & \F (x[F] \ F \\<^sub>E G)\ + by (AOT_subst_def Numbers') + (auto intro!: "\I" "\I" "&I" "cqt:2" dest: "&E") + +AOT_theorem Numbers'DistinctZeroes: + \\x\y (\Numbers'(x,[\z O!z & z \\<^sub>E z]) & \Numbers'(y,[\z O!z & z \\<^sub>E z]) & x \ y)\ +proof - + AOT_obtain w\<^sub>1 where \\w w\<^sub>1 \ w\ + using "two-worlds-exist:4" "PossibleWorld.\E"[rotated] by fast + then AOT_obtain w\<^sub>2 where distinct_worlds: \w\<^sub>1 \ w\<^sub>2\ + using "PossibleWorld.\E"[rotated] by blast + AOT_obtain x where x_prop: + \A!x & \F (x[F] \ w\<^sub>1 \ F \\<^sub>E [\z O!z & z \\<^sub>E z])\ + using "A-objects"[axiom_inst] "\E"[rotated] by fast + moreover AOT_obtain y where y_prop: + \A!y & \F (y[F] \ w\<^sub>2 \ F \\<^sub>E [\z O!z & z \\<^sub>E z])\ + using "A-objects"[axiom_inst] "\E"[rotated] by fast + moreover { + fix x w + AOT_assume x_prop: \A!x & \F (x[F] \ w \ F \\<^sub>E [\z O!z & z \\<^sub>E z])\ + AOT_have \\F w \ (x[F] \ F \\<^sub>E [\z O!z & z \\<^sub>E z])\ + proof(safe intro!: GEN "conj-dist-w:4"[unvarify p q, OF "log-prop-prop:2", + OF "log-prop-prop:2",THEN "\E"(2)] "\I" "\I") + fix F + AOT_assume \w \ x[F]\ + AOT_hence \\x[F]\ + using "fund:1"[unvarify p, OF "log-prop-prop:2", THEN "\E"(2), + OF "PossibleWorld.\I"] by blast + AOT_hence \x[F]\ + by (metis "en-eq:3[1]" "intro-elim:3:a") + AOT_thus \w \ (F \\<^sub>E [\z O!z & z \\<^sub>E z])\ + using x_prop[THEN "&E"(2), THEN "\E"(2), THEN "\E"(1)] by blast + next + fix F + AOT_assume \w \ (F \\<^sub>E [\z O!z & z \\<^sub>E z])\ + AOT_hence \x[F]\ + using x_prop[THEN "&E"(2), THEN "\E"(2), THEN "\E"(2)] by blast + AOT_hence \\x[F]\ + using "pre-en-eq:1[1]"[THEN "\E"] by blast + AOT_thus \w \ x[F]\ + using "fund:2"[unvarify p, OF "log-prop-prop:2", THEN "\E"(1)] + "PossibleWorld.\E" by fast + qed + AOT_hence \w \ \F (x[F] \ F \\<^sub>E [\z O!z & z \\<^sub>E z])\ + using "conj-dist-w:5"[THEN "\E"(2)] by fast + moreover { + AOT_have \\[\z O!z & z \\<^sub>E z]\\ + by (safe intro!: RN "cqt:2") + AOT_hence \w \ [\z O!z & z \\<^sub>E z]\\ + using "fund:2"[unvarify p, OF "log-prop-prop:2", THEN "\E"(1), + THEN "PossibleWorld.\E"] by blast + } + moreover { + AOT_have \\A!x\ + using x_prop[THEN "&E"(1)] by (metis "oa-facts:2" "\E") + AOT_hence \w \ A!x\ + using "fund:2"[unvarify p, OF "log-prop-prop:2", + THEN "\E"(1), THEN "PossibleWorld.\E"] by blast + } + ultimately AOT_have \w \ (A!x & [\z O!z & z \\<^sub>E z]\ & + \F (x[F] \ F \\<^sub>E [\z O!z & z \\<^sub>E z]))\ + using "conj-dist-w:1"[unvarify p q, OF "log-prop-prop:2", + OF "log-prop-prop:2", THEN "\E"(2), OF "&I"] by auto + AOT_hence \\w w \ (A!x & [\z O!z & z \\<^sub>E z]\ & + \F (x[F] \ F \\<^sub>E [\z O!z & z \\<^sub>E z]))\ + using "PossibleWorld.\I" by auto + AOT_hence \\(A!x & [\z O!z & z \\<^sub>E z]\ & \F (x[F] \ F \\<^sub>E [\z O!z & z \\<^sub>E z]))\ + using "fund:1"[unvarify p, OF "log-prop-prop:2", THEN "\E"(2)] by blast + AOT_hence \\Numbers'(x,[\z O!z & z \\<^sub>E z])\ + by (AOT_subst_def Numbers') + } + ultimately AOT_have \\Numbers'(x,[\z O!z & z \\<^sub>E z])\ + and \\Numbers'(y,[\z O!z & z \\<^sub>E z])\ + by auto + moreover AOT_have \x \ y\ + proof (rule "ab-obey:2"[THEN "\E"]) + AOT_have \\\\u [\z O!z & z \\<^sub>E z]u\ + proof (safe intro!: RN "raa-cor:2") + AOT_modally_strict { + AOT_assume \\u [\z O!z & z \\<^sub>E z]u\ + then AOT_obtain u where \[\z O!z & z \\<^sub>E z]u\ + using "Ordinary.\E"[rotated] by blast + AOT_hence \O!u & u \\<^sub>E u\ + by (rule "\\C") + AOT_hence \\(u =\<^sub>E u)\ + by (metis "con-dis-taut:2" "intro-elim:3:d" "modus-tollens:1" + "raa-cor:3" "thm-neg=E") + AOT_hence \u =\<^sub>E u & \u =\<^sub>E u\ + by (metis "modus-tollens:1" "ord=Eequiv:1" "raa-cor:3" Ordinary.\) + AOT_thus \p & \p\ for p + by (metis "raa-cor:1") + } + qed + AOT_hence nec_not_ex: \\w w \ \\u [\z O!z & z \\<^sub>E z]u\ + using "fund:2"[unvarify p, OF "log-prop-prop:2", THEN "\E"(1)] by blast + AOT_have \\([\y p]x \ p)\ for x p + by (safe intro!: RN "beta-C-meta"[THEN "\E"] "cqt:2") + AOT_hence \\w w \ ([\y p]x \ p)\ for x p + using "fund:2"[unvarify p, OF "log-prop-prop:2", THEN "\E"(1)] by blast + AOT_hence world_prop_beta: \\w (w \ [\y p]x \ w \ p)\ for x p + using "conj-dist-w:4"[unvarify p, OF "log-prop-prop:2", THEN "\E"(1)] + "PossibleWorld.\E" "PossibleWorld.\I" by meson + + AOT_have \\p (w\<^sub>1 \ p & \w\<^sub>2 \ p)\ + proof(rule "raa-cor:1") + AOT_assume 0: \\\p (w\<^sub>1 \ p & \w\<^sub>2 \ p)\ + AOT_have 1: \w\<^sub>1 \ p \ w\<^sub>2 \ p\ for p + proof(safe intro!: GEN "\I") + AOT_assume \w\<^sub>1 \ p\ + AOT_thus \w\<^sub>2 \ p\ + using 0 "con-dis-i-e:1" "\I"(2) "raa-cor:4" by fast + qed + moreover AOT_have \w\<^sub>2 \ p \ w\<^sub>1 \ p\ for p + proof(safe intro!: GEN "\I") + AOT_assume \w\<^sub>2 \ p\ + AOT_hence \\w\<^sub>2 \ \p\ + using "coherent:2" "intro-elim:3:a" by blast + AOT_hence \\w\<^sub>1 \ \p\ + using 1["\I" p, THEN "\E"(1), OF "log-prop-prop:2"] + by (metis "modus-tollens:1") + AOT_thus \w\<^sub>1 \ p\ + using "coherent:1" "intro-elim:3:b" "reductio-aa:1" by blast + qed + ultimately AOT_have \w\<^sub>1 \ p \ w\<^sub>2 \ p\ for p + by (metis "intro-elim:2") + AOT_hence \w\<^sub>1 = w\<^sub>2\ + using "sit-identity"[unconstrain s, THEN "\E", + OF PossibleWorld.\[THEN "world:1"[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(1)], + unconstrain s', THEN "\E", + OF PossibleWorld.\[THEN "world:1"[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(1)], + THEN "\E"(2)] GEN by fast + AOT_thus \w\<^sub>1 = w\<^sub>2 & \w\<^sub>1 = w\<^sub>2\ + using "=-infix" "\\<^sub>d\<^sub>fE" "con-dis-i-e:1" distinct_worlds by blast + qed + then AOT_obtain p where 0: \w\<^sub>1 \ p & \w\<^sub>2 \ p\ + using "\E"[rotated] by blast + AOT_have \y[\y p]\ + proof (safe intro!: y_prop[THEN "&E"(2), THEN "\E"(1), THEN "\E"(2)] "cqt:2") + AOT_show \w\<^sub>2 \ [\y p] \\<^sub>E [\z O!z & z \\<^sub>E z]\ + proof (safe intro!: "cqt:2" "empty-approx:1"[unvarify F H, THEN RN, + THEN "fund:2"[unvarify p, OF "log-prop-prop:2", THEN "\E"(1)], + THEN "PossibleWorld.\E", + THEN "conj-dist-w:2"[unvarify p q, OF "log-prop-prop:2", + OF "log-prop-prop:2", THEN "\E"(1)], + THEN "\E"] + "conj-dist-w:1"[unvarify p q, OF "log-prop-prop:2", + OF "log-prop-prop:2", THEN "\E"(2)] "&I") + AOT_have \\w\<^sub>2 \ \u [\y p]u\ + proof (rule "raa-cor:2") + AOT_assume \w\<^sub>2 \ \u [\y p]u\ + AOT_hence \\x w\<^sub>2 \ (O!x & [\y p]x)\ + by (metis "conj-dist-w:6" "intro-elim:3:a") + then AOT_obtain x where \w\<^sub>2 \ (O!x & [\y p]x)\ + using "\E"[rotated] by blast + AOT_hence \w\<^sub>2 \ [\y p]x\ + using "conj-dist-w:1"[unvarify p q, OF "log-prop-prop:2", + OF "log-prop-prop:2", THEN "\E"(1), THEN "&E"(2)] by blast + AOT_hence \w\<^sub>2 \ p\ + using world_prop_beta[THEN "PossibleWorld.\E", THEN "\E"(1)] by blast + AOT_thus \w\<^sub>2 \ p & \w\<^sub>2 \ p\ + using 0[THEN "&E"(2)] "&I" by blast + qed + AOT_thus \w\<^sub>2 \ \\u [\y p]u\ + by (safe intro!: "coherent:1"[unvarify p, OF "log-prop-prop:2", + THEN "\E"(2)]) + next + AOT_show \w\<^sub>2 \ \\v [\z O!z & z \\<^sub>E z]v\ + using nec_not_ex[THEN "PossibleWorld.\E"] by blast + qed + qed + moreover AOT_have \\x[\y p]\ + proof(rule "raa-cor:2") + AOT_assume \x[\y p]\ + AOT_hence "w\<^sub>1 \ [\y p] \\<^sub>E [\z O!z & z \\<^sub>E z]" + using x_prop[THEN "&E"(2), THEN "\E"(1), THEN "\E"(1)] + "prop-prop2:2" by blast + AOT_hence "\w\<^sub>1 \ \[\y p] \\<^sub>E [\z O!z & z \\<^sub>E z]" + using "coherent:2"[unvarify p, OF "log-prop-prop:2", THEN "\E"(1)] by blast + moreover AOT_have "w\<^sub>1 \ \([\y p] \\<^sub>E [\z O!z & z \\<^sub>E z])" + proof (safe intro!: "cqt:2" "empty-approx:2"[unvarify F H, THEN RN, + THEN "fund:2"[unvarify p, OF "log-prop-prop:2", THEN "\E"(1)], + THEN "PossibleWorld.\E", + THEN "conj-dist-w:2"[unvarify p q, OF "log-prop-prop:2", + OF "log-prop-prop:2", THEN "\E"(1)], THEN "\E"] + "conj-dist-w:1"[unvarify p q, OF "log-prop-prop:2", + OF "log-prop-prop:2", THEN "\E"(2)] "&I") + fix u + AOT_have \w\<^sub>1 \ O!u\ + using Ordinary.\[THEN RN, + THEN "fund:2"[unvarify p, OF "log-prop-prop:2", THEN "\E"(1)], + THEN "PossibleWorld.\E"] by simp + moreover AOT_have \w\<^sub>1 \ [\y p]u\ + by (safe intro!: world_prop_beta[THEN "PossibleWorld.\E", THEN "\E"(2)] + 0[THEN "&E"(1)]) + ultimately AOT_have \w\<^sub>1 \ (O!u & [\y p]u)\ + using "conj-dist-w:1"[unvarify p q, OF "log-prop-prop:2", + OF "log-prop-prop:2", THEN "\E"(2), + OF "&I"] by blast + AOT_hence \\x w\<^sub>1 \ (O!x & [\y p]x)\ + by (rule "\I") + AOT_thus \w\<^sub>1 \ \u [\y p]u\ + by (metis "conj-dist-w:6" "intro-elim:3:b") + next + AOT_show \w\<^sub>1 \ \\v [\z O!z & z \\<^sub>E z]v\ + using "PossibleWorld.\E" nec_not_ex by fastforce + qed + ultimately AOT_show \p & \p\ for p + using "raa-cor:3" by blast + qed + ultimately AOT_have \y[\y p] & \x[\y p]\ + using "&I" by blast + AOT_hence \\F (y[F] & \x[F])\ + by (metis "existential:1" "prop-prop2:2") + AOT_thus \\F (x[F] & \y[F]) \ \F (y[F] & \x[F])\ + by (rule "\I") + qed + ultimately AOT_have \\Numbers'(x,[\z O!z & z \\<^sub>E z]) & + \Numbers'(y,[\z O!z & z \\<^sub>E z]) & x \ y\ + using "&I" by blast + AOT_thus \\x\y (\Numbers'(x,[\z O!z & z \\<^sub>E z]) & + \Numbers'(y,[\z O!z & z \\<^sub>E z]) & x \ y)\ + using "\I"(2)[where \=x] "\I"(2)[where \=y] by auto +qed + +AOT_theorem restricted_identity: + \x =\<^sub>\ y \ (InDomainOf(x,\) & InDomainOf(y,\) & x = y)\ + by (auto intro!: "\I" "\I" "&I" + dest: "id-R-thm:2"[THEN "\E"] "&E" + "id-R-thm:3"[THEN "\E"] + "id-R-thm:4"[THEN "\E", OF "\I"(1), THEN "\E"(2)]) + +AOT_theorem induction': \\F ([F]0 & \n([F]n \ [F]n\<^bold>') \ \n [F]n)\ +proof(rule GEN; rule "\I") + fix F n + AOT_assume A: \[F]0 & \n([F]n \ [F]n\<^bold>')\ + AOT_have \\n\m([\]nm \ ([F]n \ [F]m))\ + proof(safe intro!: "Number.GEN" "\I") + fix n m + AOT_assume \[\]nm\ + moreover AOT_have \[\]n n\<^bold>'\ + using "suc-thm". + ultimately AOT_have m_eq_suc_n: \m = n\<^bold>'\ + using "pred-func:1"[unvarify z, OF "def-suc[den2]", THEN "\E", OF "&I"] + by blast + AOT_assume \[F]n\ + AOT_hence \[F]n\<^bold>'\ + using A[THEN "&E"(2), THEN "Number.\E", THEN "\E"] by blast + AOT_thus \[F]m\ + using m_eq_suc_n[symmetric] "rule=E" by fast + qed + AOT_thus \\n[F]n\ + using induction[THEN "\E"(2), THEN "\E", OF "&I", OF A[THEN "&E"(1)]] + by simp +qed + +AOT_define ExtensionOf :: \\ \ \ \ \\ (\ExtensionOf'(_,_')\) + "exten-property:1": \ExtensionOf(x,[G]) \\<^sub>d\<^sub>f A!x & G\ & \F(x[F] \ \z([F]z \ [G]z))\ + +AOT_define OrdinaryExtensionOf :: \\ \ \ \ \\ (\OrdinaryExtensionOf'(_,_')\) + \OrdinaryExtensionOf(x,[G]) \\<^sub>d\<^sub>f A!x & G\ & \F(x[F] \ \z(O!z \ ([F]z \ [G]z)))\ + +AOT_theorem BeingOrdinaryExtensionOfDenotes: + \[\x OrdinaryExtensionOf(x,[G])]\\ +proof(rule "safe-ext"[axiom_inst, THEN "\E", OF "&I"]) + AOT_show \[\x A!x & G\ & [\x \F(x[F] \ \z(O!z \ ([F]z \ [G]z)))]x]\\ + by "cqt:2" +next + AOT_show \\\x (A!x & G\ & [\x \F (x[F] \ \z (O!z \ ([F]z \ [G]z)))]x \ + OrdinaryExtensionOf(x,[G]))\ + proof(safe intro!: RN GEN) + AOT_modally_strict { + fix x + AOT_modally_strict { + AOT_have \[\x \F (x[F] \ \z (O!z \ ([F]z \ [G]z)))]\\ + proof (safe intro!: "Comprehension_3"[THEN "\E"] RN GEN + "\I" "\I" Ordinary.GEN) + AOT_modally_strict { + fix F H u + AOT_assume \\H \\<^sub>E F\ + AOT_hence \\u([H]u \ [F]u)\ + using eqE[THEN "\\<^sub>d\<^sub>fE", THEN "&E"(2)] "qml:2"[axiom_inst, THEN "\E"] + by blast + AOT_hence 0: \[H]u \ [F]u\ using "Ordinary.\E" by fast + { + AOT_assume \\u([F]u \ [G]u)\ + AOT_hence 1: \[F]u \ [G]u\ using "Ordinary.\E" by fast + AOT_show \[G]u\ if \[H]u\ using 0 1 "\E"(1) that by blast + AOT_show \[H]u\ if \[G]u\ using 0 1 "\E"(2) that by blast + } + { + AOT_assume \\u([H]u \ [G]u)\ + AOT_hence 1: \[H]u \ [G]u\ using "Ordinary.\E" by fast + AOT_show \[G]u\ if \[F]u\ using 0 1 "\E"(1,2) that by blast + AOT_show \[F]u\ if \[G]u\ using 0 1 "\E"(1,2) that by blast + } + } + qed + } + AOT_thus \(A!x & G\ & [\x \F (x[F] \ \z (O!z \ ([F]z \ [G]z)))]x) \ + OrdinaryExtensionOf(x,[G])\ + apply (AOT_subst_def OrdinaryExtensionOf) + apply (AOT_subst \[\x \F (x[F] \ \z (O!z \ ([F]z \ [G]z)))]x\ + \\F (x[F] \ \z (O!z \ ([F]z \ [G]z)))\) + by (auto intro!: "beta-C-meta"[THEN "\E"] simp: "oth-class-taut:3:a") + } + qed +qed + +text\Fragments of PLM's theory of Concepts.\ + +AOT_define FimpG :: \\ \ \ \ \\ (infixl \\\ 50) + "F-imp-G": \[G] \ [F] \\<^sub>d\<^sub>f F\ & G\ & \\x ([G]x \ [F]x)\ + +AOT_define concept :: \\\ (\C!\) + concepts: \C! =\<^sub>d\<^sub>f A!\ + +AOT_register_rigid_restricted_type + Concept: \C!\\ +proof + AOT_modally_strict { + AOT_have \\x A!x\ + using "o-objects-exist:2" "qml:2"[axiom_inst] "\E" by blast + AOT_thus \\x C!x\ + using "rule-id-df:1[zero]"[OF concepts, OF "oa-exist:2"] "rule=E" id_sym + by fast + } +next + AOT_modally_strict { + AOT_show \C!\ \ \\\ for \ + using "cqt:5:a"[axiom_inst, THEN "\E", THEN "&E"(2)] "\I" + by blast + } +next + AOT_modally_strict { + AOT_have \\x(A!x \ \A!x)\ + by (simp add: "oa-facts:2" GEN) + AOT_thus \\x(C!x \ \C!x)\ + using "rule-id-df:1[zero]"[OF concepts, OF "oa-exist:2"] "rule=E" id_sym + by fast + } +qed + +AOT_register_variable_names + Concept: c d e + +AOT_theorem "concept-comp:1": \\x(C!x & \F(x[F] \ \{F}))\ + using concepts[THEN "rule-id-df:1[zero]", OF "oa-exist:2", symmetric] + "A-objects"[axiom_inst] + "rule=E" by fast + +AOT_theorem "concept-comp:2": \\!x(C!x & \F(x[F] \ \{F}))\ + using concepts[THEN "rule-id-df:1[zero]", OF "oa-exist:2", symmetric] + "A-objects!" + "rule=E" by fast + +AOT_theorem "concept-comp:3": \\<^bold>\x(C!x & \F(x[F] \ \{F}))\\ + using "concept-comp:2" "A-Exists:2"[THEN "\E"(2)] "RA[2]" by blast + +AOT_theorem "concept-comp:4": + \\<^bold>\x(C!x & \F(x[F] \ \{F})) = \<^bold>\x(A!x & \F(x[F] \ \{F}))\ + using "=I"(1)[OF "concept-comp:3"] + "rule=E"[rotated] + concepts[THEN "rule-id-df:1[zero]", OF "oa-exist:2"] + by fast + +AOT_define conceptInclusion :: \\ \ \ \ \\ (infixl \\\ 100) + "con:1": \c \ d \\<^sub>d\<^sub>f \F(c[F] \ d[F])\ + + +AOT_define conceptOf :: \\ \ \ \ \\ (\ConceptOf'(_,_')\) + "concept-of-G": \ConceptOf(c,G) \\<^sub>d\<^sub>f G\ & \F (c[F] \ [G] \ [F])\ + +AOT_theorem ConceptOfOrdinaryProperty: \([H] \ O!) \ [\x ConceptOf(x,H)]\\ +proof(rule "\I") + AOT_assume \[H] \ O!\ + AOT_hence \\\x([H]x \ O!x)\ + using "F-imp-G"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + AOT_hence \\\\x([H]x \ O!x)\ + using "S5Basic:6"[THEN "\E"(1)] by blast + moreover AOT_have \\\\x([H]x \ O!x) \ + \\F\G(\(G \\<^sub>E F) \ ([H] \ [F] \ [H] \ [G]))\ + proof(rule RM; safe intro!: "\I" GEN "\I") + AOT_modally_strict { + fix F G + AOT_assume 0: \\\x([H]x \ O!x)\ + AOT_assume \\G \\<^sub>E F\ + AOT_hence 1: \\\u([G]u \ [F]u)\ + by (AOT_subst_thm eqE[THEN "\Df", THEN "\S"(1), OF "&I", + OF "cqt:2[const_var]"[axiom_inst], + OF "cqt:2[const_var]"[axiom_inst], symmetric]) + { + AOT_assume \[H] \ [F]\ + AOT_hence \\\x([H]x \ [F]x)\ + using "F-imp-G"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + moreover AOT_modally_strict { + AOT_assume \\x([H]x \ O!x)\ + moreover AOT_assume \\u([G]u \ [F]u)\ + moreover AOT_assume \\x([H]x \ [F]x)\ + ultimately AOT_have \[H]x \ [G]x\ for x + by (auto intro!: "\I" dest!: "\E"(2) dest: "\E" "\E") + AOT_hence \\x([H]x \ [G]x)\ + by (rule GEN) + } + ultimately AOT_have \\\x([H]x \ [G]x)\ + using "RN[prem]"[where + \="{\\x([H]x \ O!x)\, \\u([G]u \ [F]u)\, \\x([H]x \ [F]x)\}"] + using 0 1 by fast + AOT_thus \[H] \ [G]\ + by (AOT_subst_def "F-imp-G") + (safe intro!: "cqt:2" "&I") + } + { + AOT_assume \[H] \ [G]\ + AOT_hence \\\x([H]x \ [G]x)\ + using "F-imp-G"[THEN "\\<^sub>d\<^sub>fE"] "&E" by blast + moreover AOT_modally_strict { + AOT_assume \\x([H]x \ O!x)\ + moreover AOT_assume \\u([G]u \ [F]u)\ + moreover AOT_assume \\x([H]x \ [G]x)\ + ultimately AOT_have \[H]x \ [F]x\ for x + by (auto intro!: "\I" dest!: "\E"(2) dest: "\E" "\E") + AOT_hence \\x([H]x \ [F]x)\ + by (rule GEN) + } + ultimately AOT_have \\\x([H]x \ [F]x)\ + using "RN[prem]"[where + \="{\\x([H]x \ O!x)\, \\u([G]u \ [F]u)\, \\x([H]x \ [G]x)\}"] + using 0 1 by fast + AOT_thus \[H] \ [F]\ + by (AOT_subst_def "F-imp-G") + (safe intro!: "cqt:2" "&I") + } + } + qed + ultimately AOT_have \\\F\G(\(G \\<^sub>E F) \ ([H] \ [F] \ [H] \ [G]))\ + using "\E" by blast + AOT_hence 0: \[\x \F(x[F] \ ([H] \ [F]))]\\ + using Comprehension_3[THEN "\E"] by blast + AOT_show \[\x ConceptOf(x,H)]\\ + proof (rule "safe-ext"[axiom_inst, THEN "\E", OF "&I"]) + AOT_show \[\x C!x & [\x \F(x[F] \ ([H] \ [F]))]x]\\ by "cqt:2" + next + AOT_show \\\x (C!x & [\x \F (x[F] \ [H] \ [F])]x \ ConceptOf(x,H))\ + proof (rule "RN[prem]"[where \=\{\[\x \F(x[F] \ ([H] \ [F]))]\\}\, simplified]) + AOT_modally_strict { + AOT_assume 0: \[\x \F (x[F] \ [H] \ [F])]\\ + AOT_show \\x (C!x & [\x \F (x[F] \ [H] \ [F])]x \ ConceptOf(x,H))\ + proof(safe intro!: GEN "\I" "\I" "&I") + fix x + AOT_assume \C!x & [\x \F (x[F] \ [H] \ [F])]x\ + AOT_thus \ConceptOf(x,H)\ + by (AOT_subst_def "concept-of-G") + (auto intro!: "&I" "cqt:2" dest: "&E" "\\C") + next + fix x + AOT_assume \ConceptOf(x,H)\ + AOT_hence \C!x & (H\ & \F(x[F] \ [H] \ [F]))\ + by (AOT_subst_def (reverse) "concept-of-G") + AOT_thus \C!x\ and \[\x \F(x[F] \ [H] \ [F])]x\ + by (auto intro!: "\\C" 0 "cqt:2" dest: "&E") + qed + } + next + AOT_show \\[\x \F(x[F] \ ([H] \ [F]))]\\ + using "exist-nec"[THEN "\E"] 0 by blast + qed + qed +qed + +AOT_theorem "con-exists:1": \\c ConceptOf(c,G)\ +proof - + AOT_obtain c where \\F (c[F] \ [G] \ [F])\ + using "concept-comp:1" "Concept.\E"[rotated] by meson + AOT_hence \ConceptOf(c,G)\ + by (auto intro!: "concept-of-G"[THEN "\\<^sub>d\<^sub>fI"] "&I" "cqt:2" Concept.\) + thus ?thesis by (rule "Concept.\I") +qed + +AOT_theorem "con-exists:2": \\!c ConceptOf(c,G)\ +proof - + AOT_have \\!c \F (c[F] \ [G] \ [F])\ + using "concept-comp:2" by simp + moreover { + AOT_modally_strict { + fix x + AOT_assume \\F (x[F] \ [G] \ [F])\ + moreover AOT_have \[G] \ [G]\ + by (safe intro!: "F-imp-G"[THEN "\\<^sub>d\<^sub>fI"] "&I" "cqt:2" RN GEN "\I") + ultimately AOT_have \x[G]\ + using "\E"(2) "\E" by blast + AOT_hence \A!x\ + using "encoders-are-abstract"[THEN "\E", OF "\I"(2)] by simp + AOT_hence \C!x\ + using concepts[THEN "rule-id-df:1[zero]", OF "oa-exist:2", symmetric] + "rule=E"[rotated] + by fast + } + } + ultimately show ?thesis + by (AOT_subst \ConceptOf(c,G)\ \\F (c[F] \ [G] \ [F])\ for: c; + AOT_subst_def "concept-of-G") + (auto intro!: "\I" "\I" "&I" "cqt:2" Concept.\ dest: "&E") +qed + +AOT_theorem "con-exists:3": \\<^bold>\c ConceptOf(c,G)\\ + by (safe intro!: "A-Exists:2"[THEN "\E"(2)] "con-exists:2"[THEN "RA[2]"]) + + +AOT_define theConceptOfG :: \\ \ \\<^sub>s\ (\\<^bold>c\<^sub>_\) + "concept-G": \\<^bold>c\<^sub>G =\<^sub>d\<^sub>f \<^bold>\c ConceptOf(c, G)\ + +AOT_theorem "concept-G[den]": \\<^bold>c\<^sub>G\\ + by (auto intro!: "rule-id-df:1"[OF "concept-G"] + "t=t-proper:1"[THEN "\E"] + "con-exists:3") + + +AOT_theorem "concept-G[concept]": \C!\<^bold>c\<^sub>G\ +proof - + AOT_have \\<^bold>\(C!\<^bold>c\<^sub>G & ConceptOf(\<^bold>c\<^sub>G, G))\ + by (auto intro!: "actual-desc:2"[unvarify x, THEN "\E"] + "rule-id-df:1"[OF "concept-G"] + "concept-G[den]" + "con-exists:3") + AOT_hence \\<^bold>\C!\<^bold>c\<^sub>G\ + by (metis "Act-Basic:2" "con-dis-i-e:2:a" "intro-elim:3:a") + AOT_hence \\<^bold>\A!\<^bold>c\<^sub>G\ + using "rule-id-df:1[zero]"[OF concepts, OF "oa-exist:2"] + "rule=E" by fast + AOT_hence \A!\<^bold>c\<^sub>G\ + using "oa-facts:8"[unvarify x, THEN "\E"(2)] "concept-G[den]" by blast + thus ?thesis + using "rule-id-df:1[zero]"[OF concepts, OF "oa-exist:2", symmetric] + "rule=E" by fast +qed + +AOT_theorem "conG-strict": \\<^bold>c\<^sub>G = \<^bold>\c \F(c[F] \ [G] \ [F])\ +proof (rule "id-eq:3"[unvarify \ \ \, THEN "\E"]) + AOT_have \\\x (C!x & ConceptOf(x,G) \ C!x & \F (x[F] \ [G] \ [F]))\ + by (auto intro!: "concept-of-G"[THEN "\\<^sub>d\<^sub>fI"] RN GEN "\I" "\I" "&I" "cqt:2" + dest: "&E"; + auto dest: "\E"(2) "\E"(1,2) dest!: "&E"(2) "concept-of-G"[THEN "\\<^sub>d\<^sub>fE"]) + AOT_thus \\<^bold>c\<^sub>G = \<^bold>\c ConceptOf(c, G) & \<^bold>\c ConceptOf(c, G) = \<^bold>\c \F(c[F] \ [G] \ [F])\ + by (auto intro!: "&I" "rule-id-df:1"[OF "concept-G"] "con-exists:3" + "equiv-desc-eq:3"[THEN "\E"]) +qed(auto simp: "concept-G[den]" "con-exists:3" "concept-comp:3") + + +AOT_theorem "conG-lemma:1": \\F(\<^bold>c\<^sub>G[F] \ [G] \ [F])\ +proof(safe intro!: GEN "\I" "\I") + fix F + AOT_have \\<^bold>\\F(\<^bold>c\<^sub>G[F] \ [G] \ [F])\ + using "actual-desc:4"[THEN "\E", OF "concept-comp:3", + THEN "Act-Basic:2"[THEN "\E"(1)], + THEN "&E"(2)] + "conG-strict"[symmetric] "rule=E" by fast + AOT_hence \\<^bold>\(\<^bold>c\<^sub>G[F] \ [G] \ [F])\ + using "logic-actual-nec:3"[axiom_inst, THEN "\E"(1)] "\E"(2) + by blast + AOT_hence 0: \\<^bold>\\<^bold>c\<^sub>G[F] \ \<^bold>\[G] \ [F]\ + using "Act-Basic:5"[THEN "\E"(1)] by blast + { + AOT_assume \\<^bold>c\<^sub>G[F]\ + AOT_hence \\<^bold>\\<^bold>c\<^sub>G[F]\ + by(safe intro!: "en-eq:10[1]"[unvarify x\<^sub>1, THEN "\E"(2)] + "concept-G[den]") + AOT_hence \\<^bold>\[G] \ [F]\ + using 0[THEN "\E"(1)] by blast + AOT_hence \\<^bold>\(F\ & G\ & \\x([G]x \ [F]x))\ + by (AOT_subst_def (reverse) "F-imp-G") + AOT_hence \\<^bold>\\\x([G]x \ [F]x)\ + using "Act-Basic:2"[THEN "\E"(1)] "&E" by blast + AOT_hence \\\x([G]x \ [F]x)\ + using "qml-act:2"[axiom_inst, THEN "\E"(2)] by simp + AOT_thus \[G] \ [F]\ + by (AOT_subst_def "F-imp-G"; auto intro!: "&I" "cqt:2") + } + { + AOT_assume \[G] \ [F]\ + AOT_hence \\\x([G]x \ [F]x)\ + by (safe dest!: "F-imp-G"[THEN "\\<^sub>d\<^sub>fE"] "&E"(2)) + AOT_hence \\<^bold>\\\x([G]x \ [F]x)\ + using "qml-act:2"[axiom_inst, THEN "\E"(1)] by simp + AOT_hence \\<^bold>\(F\ & G\ & \\x([G]x \ [F]x))\ + by (auto intro!: "Act-Basic:2"[THEN "\E"(2)] "&I" "cqt:2" + intro: "RA[2]") + AOT_hence \\<^bold>\([G] \ [F])\ + by (AOT_subst_def "F-imp-G") + AOT_hence \\<^bold>\\<^bold>c\<^sub>G[F]\ + using 0[THEN "\E"(2)] by blast + AOT_thus \\<^bold>c\<^sub>G[F]\ + by(safe intro!: "en-eq:10[1]"[unvarify x\<^sub>1, THEN "\E"(1)] + "concept-G[den]") + } +qed + +AOT_theorem conH_enc_ord: + \([H] \ O!) \ \\F \G (\G \\<^sub>E F \ (\<^bold>c\<^sub>H[F] \ \<^bold>c\<^sub>H[G]))\ +proof(rule "\I") + AOT_assume 0: \[H] \ O!\ + AOT_have 0: \\([H] \ O!)\ + apply (AOT_subst_def "F-imp-G") + using 0[THEN "\\<^sub>d\<^sub>fE"[OF "F-imp-G"]] + by (auto intro!: "KBasic:3"[THEN "\E"(2)] "&I" "exist-nec"[THEN "\E"] + dest: "&E" 4[THEN "\E"]) + moreover AOT_have \\([H] \ O!) \ \\F \G (\G \\<^sub>E F \ (\<^bold>c\<^sub>H[F] \ \<^bold>c\<^sub>H[G]))\ + proof(rule RM; safe intro!: "\I" GEN) + AOT_modally_strict { + fix F G + AOT_assume \[H] \ O!\ + AOT_hence 0: \\\x ([H]x \ O!x)\ + by (safe dest!: "F-imp-G"[THEN "\\<^sub>d\<^sub>fE"] "&E"(2)) + AOT_assume 1: \\G \\<^sub>E F\ + AOT_assume \\<^bold>c\<^sub>H[F]\ + AOT_hence \[H] \ [F]\ + using "conG-lemma:1"[THEN "\E"(2), THEN "\E"(1)] by simp + AOT_hence 2: \\\x ([H]x \ [F]x)\ + by (safe dest!: "F-imp-G"[THEN "\\<^sub>d\<^sub>fE"] "&E"(2)) + AOT_modally_strict { + AOT_assume 0: \\x ([H]x \ O!x)\ + AOT_assume 1: \\x ([H]x \ [F]x)\ + AOT_assume 2: \G \\<^sub>E F\ + AOT_have \\x ([H]x \ [G]x)\ + proof(safe intro!: GEN "\I") + fix x + AOT_assume \[H]x\ + AOT_hence \O!x\ and \[F]x\ + using 0 1 "\E"(2) "\E" by blast+ + AOT_thus \[G]x\ + using 2[THEN eqE[THEN "\\<^sub>d\<^sub>fE"], THEN "&E"(2)] + "\E"(2) "\E" "\E"(2) calculation by blast + qed + } + AOT_hence \\\x ([H]x \ [G]x)\ + using "RN[prem]"[where \=\{\\x ([H]x \ O!x)\, + \\x ([H]x \ [F]x)\, + \G \\<^sub>E F\}\, simplified] 0 1 2 by fast + AOT_hence \[H] \ [G]\ + by (safe intro!: "F-imp-G"[THEN "\\<^sub>d\<^sub>fI"] "&I" "cqt:2") + AOT_hence \\<^bold>c\<^sub>H[G]\ + using "conG-lemma:1"[THEN "\E"(2), THEN "\E"(2)] by simp + } note 0 = this + AOT_modally_strict { + fix F G + AOT_assume \[H] \ O!\ + moreover AOT_assume \\G \\<^sub>E F\ + moreover AOT_have \\F \\<^sub>E G\ + by (AOT_subst \F \\<^sub>E G\ \G \\<^sub>E F\) + (auto intro!: calculation(2) + eqE[THEN "\\<^sub>d\<^sub>fI"] + "\I" "\I" "&I" "cqt:2" Ordinary.GEN + dest!: eqE[THEN "\\<^sub>d\<^sub>fE"] "&E"(2) + dest: "\E"(1,2) "Ordinary.\E") + ultimately AOT_show \(\<^bold>c\<^sub>H[F] \ \<^bold>c\<^sub>H[G])\ + using 0 "\I" "\I" by auto + } + qed + ultimately AOT_show \\\F \G (\G \\<^sub>E F \ (\<^bold>c\<^sub>H[F] \ \<^bold>c\<^sub>H[G]))\ + using "\E" by blast +qed + +AOT_theorem concept_inclusion_denotes_1: + \([H] \ O!) \ [\x \<^bold>c\<^sub>H \ x]\\ +proof(rule "\I") + AOT_assume 0: \[H] \ O!\ + AOT_show \[\x \<^bold>c\<^sub>H \ x]\\ + proof(rule "safe-ext"[axiom_inst, THEN "\E", OF "&I"]) + AOT_show \[\x C!x & \F(\<^bold>c\<^sub>H[F] \ x[F])]\\ + by (safe intro!: conjunction_denotes[THEN "\E", OF "&I"] + Comprehension_2'[THEN "\E"] + conH_enc_ord[THEN "\E", OF 0]) "cqt:2" + next + AOT_show \\\x (C!x & \F (\<^bold>c\<^sub>H[F] \ x[F]) \ \<^bold>c\<^sub>H \ x)\ + by (safe intro!: RN GEN; AOT_subst_def "con:1") + (auto intro!: "\I" "\I" "&I" "concept-G[concept]" dest: "&E") + qed +qed + +AOT_theorem concept_inclusion_denotes_2: + \([H] \ O!) \ [\x x \ \<^bold>c\<^sub>H]\\ +proof(rule "\I") + AOT_assume 0: \[H] \ O!\ + AOT_show \[\x x \ \<^bold>c\<^sub>H]\\ + proof(rule "safe-ext"[axiom_inst, THEN "\E", OF "&I"]) + AOT_show \[\x C!x & \F(x[F] \ \<^bold>c\<^sub>H[F])]\\ + by (safe intro!: conjunction_denotes[THEN "\E", OF "&I"] + Comprehension_1'[THEN "\E"] + conH_enc_ord[THEN "\E", OF 0]) "cqt:2" + next + AOT_show \\\x (C!x & \F (x[F] \ \<^bold>c\<^sub>H[F]) \ x \ \<^bold>c\<^sub>H)\ + by (safe intro!: RN GEN; AOT_subst_def "con:1") + (auto intro!: "\I" "\I" "&I" "concept-G[concept]" dest: "&E") + qed +qed + +AOT_define ThickForm :: \\ \ \ \ \\ (\FormOf'(_,_')\) + "tform-of": \FormOf(x,G) \\<^sub>d\<^sub>f A!x & G\ & \F(x[F] \ [G] \ [F])\ + +AOT_theorem FormOfOrdinaryProperty: \([H] \ O!) \ [\x FormOf(x,H)]\\ +proof(rule "\I") + AOT_assume 0: \[H] \ [O!]\ + AOT_show \[\x FormOf(x,H)]\\ + proof (rule "safe-ext"[axiom_inst, THEN "\E", OF "&I"]) + AOT_show \[\x ConceptOf(x,H)]\\ + using 0 ConceptOfOrdinaryProperty[THEN "\E"] by blast + AOT_show \\\x (ConceptOf(x,H) \ FormOf(x,H))\ + proof(safe intro!: RN GEN) + AOT_modally_strict { + fix x + AOT_modally_strict { + AOT_have \A!x \ A!x\ + by (simp add: "oth-class-taut:3:a") + AOT_hence \C!x \ A!x\ + using "rule-id-df:1[zero]"[OF concepts, OF "oa-exist:2"] + "rule=E" id_sym by fast + } + AOT_thus \ConceptOf(x,H) \ FormOf(x,H)\ + by (AOT_subst_def "tform-of"; + AOT_subst_def "concept-of-G"; + AOT_subst \C!x\ \A!x\) + (auto intro!: "\I" "\I" "&I" dest: "&E") + } + qed + qed +qed + +AOT_theorem equal_E_rigid_one_to_one: \Rigid\<^sub>1\<^sub>-\<^sub>1((=\<^sub>E))\ +proof (safe intro!: "df-1-1:2"[THEN "\\<^sub>d\<^sub>fI"] "&I" "df-1-1:1"[THEN "\\<^sub>d\<^sub>fI"] + GEN "\I" "df-rigid-rel:1"[THEN "\\<^sub>d\<^sub>fI"] "=E[denotes]") + fix x y z + AOT_assume \x =\<^sub>E z & y =\<^sub>E z\ + AOT_thus \x = y\ + by (metis "rule=E" "&E"(1) "Conjunction Simplification"(2) + "=E-simple:2" id_sym "\E") +next + AOT_have \\x\y \(x =\<^sub>E y \ \x =\<^sub>E y)\ + proof(rule GEN; rule GEN) + AOT_show \\(x =\<^sub>E y \ \x =\<^sub>E y)\ for x y + by (meson RN "deduction-theorem" "id-nec3:1" "\E"(1)) + qed + AOT_hence \\x\<^sub>1...\x\<^sub>n \([(=\<^sub>E)]x\<^sub>1...x\<^sub>n \ \[(=\<^sub>E)]x\<^sub>1...x\<^sub>n)\ + by (rule tuple_forall[THEN "\\<^sub>d\<^sub>fI"]) + AOT_thus \\\x\<^sub>1...\x\<^sub>n ([(=\<^sub>E)]x\<^sub>1...x\<^sub>n \ \[(=\<^sub>E)]x\<^sub>1...x\<^sub>n)\ + using BF[THEN "\E"] by fast +qed + +AOT_theorem equal_E_domain: \InDomainOf(x,(=\<^sub>E)) \ O!x\ +proof(safe intro!: "\I" "\I") + AOT_assume \InDomainOf(x,(=\<^sub>E))\ + AOT_hence \\y x =\<^sub>E y\ + by (metis "\\<^sub>d\<^sub>fE" "df-1-1:5") + then AOT_obtain y where \x =\<^sub>E y\ + using "\E"[rotated] by blast + AOT_thus \O!x\ + using "=E-simple:1"[THEN "\E"(1)] "&E" by blast +next + AOT_assume \O!x\ + AOT_hence \x =\<^sub>E x\ + by (metis "ord=Eequiv:1"[THEN "\E"]) + AOT_hence \\y x =\<^sub>E y\ + using "\I"(2) by fast + AOT_thus \InDomainOf(x,(=\<^sub>E))\ + by (metis "\\<^sub>d\<^sub>fI" "df-1-1:5") +qed + +AOT_theorem shared_urelement_projection_identity: + assumes \\y [\x (y[\z [R]zx])]\\ + shows \\F([F]a \ [F]b) \ [\z [R]za] = [\z [R]zb]\ +proof(rule "\I") + AOT_assume 0: \\F([F]a \ [F]b)\ + { + fix z + AOT_have \[\x (z[\z [R]zx])]\\ + using assms[THEN "\E"(2)]. + AOT_hence 1: \\x \y (\F ([F]x \ [F]y) \ \(z[\z [R]zx] \ z[\z [R]zy]))\ + using "kirchner-thm-cor:1"[THEN "\E"] + by blast + AOT_have \\(z[\z [R]za] \ z[\z [R]zb])\ + using 1[THEN "\E"(2), THEN "\E"(2), THEN "\E", OF 0] by blast + } + AOT_hence \\z \(z[\z [R]za] \ z[\z [R]zb])\ + by (rule GEN) + AOT_hence \\\z(z[\z [R]za] \ z[\z [R]zb])\ + by (rule BF[THEN "\E"]) + AOT_thus \[\z [R]za] = [\z [R]zb]\ + by (AOT_subst_def "identity:2") + (auto intro!: "&I" "cqt:2") +qed + +AOT_theorem shared_urelement_exemplification_identity: + assumes \\y [\x (y[\z [G]x])]\\ + shows \\F([F]a \ [F]b) \ ([G]a) = ([G]b)\ +proof(rule "\I") + AOT_assume 0: \\F([F]a \ [F]b)\ + { + fix z + AOT_have \[\x (z[\z [G]x])]\\ + using assms[THEN "\E"(2)]. + AOT_hence 1: \\x \y (\F ([F]x \ [F]y) \ \(z[\z [G]x] \ z[\z [G]y]))\ + using "kirchner-thm-cor:1"[THEN "\E"] + by blast + AOT_have \\(z[\z [G]a] \ z[\z [G]b])\ + using 1[THEN "\E"(2), THEN "\E"(2), THEN "\E", OF 0] by blast + } + AOT_hence \\z \(z[\z [G]a] \ z[\z [G]b])\ + by (rule GEN) + AOT_hence \\\z(z[\z [G]a] \ z[\z [G]b])\ + by (rule BF[THEN "\E"]) + AOT_hence \[\z [G]a] = [\z [G]b]\ + by (AOT_subst_def "identity:2") + (auto intro!: "&I" "cqt:2") + AOT_thus \([G]a) = ([G]b)\ + by (safe intro!: "identity:4"[THEN "\\<^sub>d\<^sub>fI"] "&I" "log-prop-prop:2") +qed + +text\The assumptions of the theorems above are derivable, if the additional + introduction rules for the upcoming extension of @{thm "cqt:2[lambda]"} + are explicitly allowed (while they are currently not part of the + abstraction layer).\ +notepad +begin + AOT_modally_strict { + AOT_have \\R\y [\x (y[\z [R]zx])]\\ + by (safe intro!: GEN "cqt:2" AOT_instance_of_cqt_2_intro_next) + AOT_have \\G\y [\x (y[\z [G]x])]\\ + by (safe intro!: GEN "cqt:2" AOT_instance_of_cqt_2_intro_next) + } +end + +end diff --git a/thys/AOT/AOT_model.thy b/thys/AOT/AOT_model.thy new file mode 100644 --- /dev/null +++ b/thys/AOT/AOT_model.thy @@ -0,0 +1,1324 @@ +(*<*) +theory AOT_model + imports Main "HOL-Cardinals.Cardinals" +begin + +declare[[typedef_overloaded]] +(*>*) + +section\References\ + +text\ +A full description of this formalization including references can be found +at @{url \http://dx.doi.org/10.17169/refubium-35141\}. + +The version of Principia Logico-Metaphysica (PLM) implemented in this formalization +can be found at @{url \http://mally.stanford.edu/principia-2021-10-13.pdf\}, while +the latest version of PLM is available at @{url \http://mally.stanford.edu/principia.pdf\}. + +\ + +section\Model for the Logic of AOT\ + +text\We introduce a primitive type for hyperintensional propositions.\ +typedecl \ + +text\To be able to model modal operators following Kripke semantics, + we introduce a primitive type for possible worlds and assert, by axiom, + that there is a surjective function mapping propositions to the + boolean-valued functions acting on possible worlds. We call the result + of applying this function to a proposition the Montague intension + of the proposition.\ +typedecl w \\The primtive type of possible worlds.\ +axiomatization AOT_model_d\ :: \\\(w\bool)\ where + d\_surj: \surj AOT_model_d\\ + +text\The axioms of PLM require the existence of a non-actual world.\ +consts w\<^sub>0 :: w \\The designated actual world.\ +axiomatization where AOT_model_nonactual_world: \\w . w \ w\<^sub>0\ + +text\Validity of a proposition in a given world can now be modelled as the result + of applying that world to the Montague intension of the proposition.\ +definition AOT_model_valid_in :: \w\\\bool\ where + \AOT_model_valid_in w \ \ AOT_model_d\ \ w\ + +text\By construction, we can choose a proposition for any given Montague intension, + s.t. the proposition is valid in a possible world iff the Montague intension + evaluates to true at that world.\ +definition AOT_model_proposition_choice :: \(w\bool) \ \\ (binder \\\<^sub>\ \ 8) + where \\\<^sub>\ w. \ w \ (inv AOT_model_d\) \\ +lemma AOT_model_proposition_choice_simp: \AOT_model_valid_in w (\\<^sub>\ w. \ w) = \ w\ + by (simp add: surj_f_inv_f[OF d\_surj] AOT_model_valid_in_def + AOT_model_proposition_choice_def) + +text\Nitpick can trivially show that there are models for the axioms above.\ +lemma \True\ nitpick[satisfy, user_axioms, expect = genuine] .. + +typedecl \ \\The primtive type of ordinary objects/urelements.\ + +text\Validating extended relation comprehension requires a large set of + special urelements. For simple models that do not validate extended + relation comprehension (and consequently the predecessor axiom in the + theory of natural numbers), it suffices to use a primitive type as @{text \}, + i.e. @{theory_text \typedecl \\}.\ +typedecl \' +typedef \ = \UNIV::((\ \ w \ bool) set \ (\ \ w \ bool) set \ \') set\ .. + +typedecl null \ \Null-urelements representing non-denoting terms.\ + +datatype \ = \\ \ | \\ \ | is_null\: null\ null \ \Type of urelements\ + +text\Urrelations are proposition-valued functions on urelements. + Urrelations are required to evaluate to necessarily false propositions for + null-urelements (note that there may be several distinct necessarily false + propositions).\ +typedef urrel = \{ \ . \ x w . \AOT_model_valid_in w (\ (null\ x)) }\ + by (rule exI[where x=\\ x . (\\<^sub>\ w . \is_null\ x)\]) + (auto simp: AOT_model_proposition_choice_simp) + +text\Abstract objects will be modelled as sets of urrelations and will + have to be mapped surjectively into the set of special urelements. + We show that any mapping from abstract objects to special urelements + has to involve at least one large set of collapsed abstract objects. + We will use this fact to extend arbitrary mappings from abstract objects + to special urelements to surjective mappings.\ +lemma \\_pigeonhole: + \ \For any arbitrary mapping @{term \\} from sets of urrelations to special + urelements, there exists an abstract object x, s.t. the cardinal of the set + of special urelements is strictly smaller than the cardinal of the set of + abstract objects that are mapped to the same urelement as x under @{term \\}.\ + \\x . |UNIV::\ set| \ x = \\ y}|\ + for \\ :: \urrel set \ \\ +proof(rule ccontr) + have card_\_set_set_bound: \|UNIV::\ set set| \o |UNIV::urrel set|\ + proof - + let ?pick = \\u s . \\<^sub>\ w . case u of (\\ s') \ s' \ s | _ \ False\ + have \\f :: \ set \ urrel . inj f\ + proof + show \inj (\s . Abs_urrel (\u . ?pick u s))\ + proof(rule injI) + fix x y + assume \Abs_urrel (\u. ?pick u x) = Abs_urrel (\u. ?pick u y)\ + hence \(\u. ?pick u x) = (\u. ?pick u y)\ + by (auto intro!: Abs_urrel_inject[THEN iffD1] + simp: AOT_model_proposition_choice_simp) + hence \AOT_model_valid_in w\<^sub>0 (?pick (\\ s) x) = + AOT_model_valid_in w\<^sub>0 (?pick (\\ s) y)\ + for s by metis + hence \(s \ x) = (s \ y)\ for s + by (auto simp: AOT_model_proposition_choice_simp) + thus \x = y\ + by blast + qed + qed + thus ?thesis + by (metis card_of_image inj_imp_surj_inv) + qed + + text\Assume, for a proof by contradiction, that there is no large collapsed set.\ + assume \\x . |UNIV::\ set| \ x = \\ y}|\ + hence A: \\x . |{y . \\ x = \\ y}| \o |UNIV::\ set|\ + by auto + have union_univ: \(\x \ range(inv \\) . {y . \\ x = \\ y}) = UNIV\ + by auto (meson f_inv_into_f range_eqI) + + text\We refute by case distinction: there is either finitely many or + infinitely many special urelements and in both cases we can derive + a contradiction from the assumption above.\ + { + text\Finite case.\ + assume finite_\_set: \finite (UNIV::\ set)\ + hence finite_collapsed: \finite {y . \\ x = \\ y}\ for x + using A card_of_ordLeq_infinite by blast + hence 0: \\x . card {y . \\ x = \\ y} \ card (UNIV::\ set)\ + by (metis A finite_\_set card_of_ordLeq inj_on_iff_card_le) + have 1: \card (range (inv \\)) \ card (UNIV::\ set)\ + using finite_\_set card_image_le by blast + hence 2: \finite (range (inv \\))\ + using finite_\_set by blast + + define n where \n = card (UNIV::urrel set set)\ + define m where \m = card (UNIV::\ set)\ + + have \n = card (\x \ range(inv \\) . {y . \\ x = \\ y})\ + unfolding n_def using union_univ by argo + also have \\ \ (\i\range (inv \\). card {y. \\ i = \\ y})\ + using card_UN_le 2 by blast + also have \\ \ (\i\range (inv \\). card (UNIV::\ set))\ + by (metis (no_types, lifting) 0 sum_mono) + also have \\ \ card (range (inv \\)) * card (UNIV::\ set)\ + using sum_bounded_above by auto + also have \\ \ card (UNIV::\ set) * card (UNIV::\ set)\ + using 1 by force + also have \\ = m*m\ + unfolding m_def by blast + finally have n_upper: \n \ m*m\. + + have \finite (\x \ range(inv \\) . {y . \\ x = \\ y})\ + using 2 finite_collapsed by blast + hence finite_\set: \finite (UNIV::urrel set set)\ + using union_univ by argo + + have \2^2^m = (2::nat)^(card (UNIV::\ set set))\ + by (metis Pow_UNIV card_Pow finite_\_set m_def) + moreover have \card (UNIV::\ set set) \ (card (UNIV::urrel set))\ + using card_\_set_set_bound + by (meson Finite_Set.finite_set card_of_ordLeq finite_\set + finite_\_set inj_on_iff_card_le) + ultimately have \2^2^m \ (2::nat)^(card (UNIV:: urrel set))\ + by simp + also have \\ = n\ + unfolding n_def + by (metis Finite_Set.finite_set Pow_UNIV card_Pow finite_\set) + finally have \2^2^m \ n\ by blast + hence \2^2^m \ m*m\ using n_upper by linarith + moreover { + have \(2::nat)^(2^m) \ (2^(m + 1))\ + by (metis Suc_eq_plus1 Suc_leI less_exp one_le_numeral power_increasing) + also have \(2^(m + 1)) = (2::nat) * 2^m\ + by auto + have \m < 2^m\ + by (simp add: less_exp) + hence \m*m < (2^m)*(2^m)\ + by (simp add: mult_strict_mono) + moreover have \\ = 2^(m+m)\ + by (simp add: power_add) + ultimately have \m*m < 2 ^ (m + m)\ by presburger + moreover have \m+m \ 2^m\ + proof (induct m) + case 0 + thus ?case by auto + next + case (Suc m) + thus ?case + by (metis Suc_leI less_exp mult_2 mult_le_mono2 power_Suc) + qed + ultimately have \m*m < 2^2^m\ + by (meson less_le_trans one_le_numeral power_increasing) + } + ultimately have False by auto + } + moreover { + text\Infinite case.\ + assume \infinite (UNIV::\ set)\ + hence Cinf\: \Cinfinite |UNIV::\ set|\ + by (simp add: cinfinite_def) + have 1: \|range (inv \\)| \o |UNIV::\ set|\ + by auto + have 2: \\i\range (inv \\). |{y . \\ i = \\ y}| \o |UNIV::\ set|\ + proof + fix i :: \urrel set\ + assume \i \ range (inv \\)\ + show \|{y . \\ i = \\ y}| \o |UNIV::\ set|\ + using A by blast + qed + have \|\ ((\i. {y. \\ i = \\ y}) ` (range (inv \\)))| \o + |Sigma (range (inv \\)) (\i. {y. \\ i = \\ y})|\ + using card_of_UNION_Sigma by blast + hence \|UNIV::urrel set set| \o + |Sigma (range (inv \\)) (\i. {y. \\ i = \\ y})|\ + using union_univ by argo + moreover have \|Sigma (range (inv \\)) (\i. {y. \\ i = \\ y})| \o |UNIV::\ set|\ + using card_of_Sigma_ordLeq_Cinfinite[OF Cinf\, OF 1, OF 2] by blast + ultimately have \|UNIV::urrel set set| \o |UNIV::\ set|\ + using ordLeq_transitive by blast + moreover { + have \|UNIV::\ set| set set|\ + by auto + moreover have \|UNIV::\ set set| \o |UNIV::urrel set|\ + using card_\_set_set_bound by blast + moreover have \|UNIV::urrel set| + by auto + ultimately have \|UNIV::\ set| + by (metis ordLess_imp_ordLeq ordLess_ordLeq_trans) + } + ultimately have False + using not_ordLeq_ordLess by blast + } + ultimately show False by blast +qed + +text\We introduce a mapping from abstract objects (i.e. sets of urrelations) to + special urelements @{text \\\\} that is surjective and distinguishes all + abstract objects that are distinguished by a (not necessarily surjective) + mapping @{text \\\'\}. @{text \\\'\} will be used to model extended relation + comprehension.\ +consts \\' :: \urrel set \ \\ +consts \\ :: \urrel set \ \\ + +specification(\\) + \\_surj: \surj \\\ + \\_\\': \\\ x = \\ y \ \\' x = \\' y\ +proof - + obtain x where x_prop: \|UNIV::\ set| \' x = \\' y}|\ + using \\_pigeonhole by blast + have \\f :: urrel set \ \ . f ` {y. \\' x = \\' y} = UNIV \ f x = \\' x\ + proof - + have \\f :: urrel set \ \ . f ` {y. \\' x = \\' y} = UNIV\ + by (simp add: x_prop card_of_ordLeq2 ordLess_imp_ordLeq) + then obtain f :: \urrel set \ \\ where \f ` {y. \\' x = \\' y} = UNIV\ + by presburger + moreover obtain a where \f a = \\' x\ and \\\' a = \\' x\ + by (smt (verit, best) calculation UNIV_I image_iff mem_Collect_eq) + ultimately have \(f (a := f x, x := f a)) ` {y. \\' x = \\' y} = UNIV \ + (f (a := f x, x := f a)) x = \\' x\ + by (auto simp: image_def) + thus ?thesis by blast + qed + then obtain f where fimage: \f ` {y. \\' x = \\' y} = UNIV\ + and fx: \f x = \\' x\ + by blast + + define \\ :: \urrel set \ \\ where + \\\ \ \ urrels . if \\' urrels = \\' x \ f urrels \ range \\' + then f urrels + else \\' urrels\ + have \surj \\\ + proof - + { + fix s :: \ + { + assume \s \ range \\'\ + hence 0: \\\' (inv \\' s) = s\ + by (meson f_inv_into_f) + { + assume \s = \\' x\ + hence \\\ x = s\ + using \\_def fx by presburger + hence \\f . \\ (f s) = s\ + by auto + } + moreover { + assume \s \ \\' x\ + hence \\\ (inv \\' s) = s\ + unfolding \\_def 0 by presburger + hence \\f . \\ (f s) = s\ + by blast + } + ultimately have \\f . \\ (f s) = s\ + by blast + } + moreover { + assume \s \ range \\'\ + moreover obtain urrels where \f urrels = s\ and \\\' x = \\' urrels\ + by (smt (verit, best) UNIV_I fimage image_iff mem_Collect_eq) + ultimately have \\\ urrels = s\ + using \\_def by presburger + hence \\f . \\ (f s) = s\ + by (meson f_inv_into_f range_eqI) + } + ultimately have \\f . \\ (f s) = s\ + by blast + } + thus ?thesis + by (metis surj_def) + qed + moreover have \\x y. \\ x = \\ y \ \\' x = \\' y\ + by (metis \\_def rangeI) + ultimately show ?thesis + by blast +qed + +text\For extended models that validate extended relation comprehension + (and consequently the predecessor axiom), we specify which + abstract objects are distinguished by @{const \\'}.\ + +definition urrel_to_\rel :: \urrel \ (\ \ w \ bool)\ where + \urrel_to_\rel \ \ r u w . AOT_model_valid_in w (Rep_urrel r (\\ u))\ +definition \rel_to_urrel :: \(\ \ w \ bool) \ urrel\ where + \\rel_to_urrel \ \ \ . Abs_urrel + (\ u . \\<^sub>\ w . case u of \\ x \ \ x w | _ \ False)\ + +definition AOT_urrel_\equiv :: \urrel \ urrel \ bool\ where + \AOT_urrel_\equiv \ \ r s . \ u v . AOT_model_valid_in v (Rep_urrel r (\\ u)) = + AOT_model_valid_in v (Rep_urrel s (\\ u))\ + +lemma urrel_\rel_quot: \Quotient3 AOT_urrel_\equiv urrel_to_\rel \rel_to_urrel\ +proof(rule Quotient3I) + show \urrel_to_\rel (\rel_to_urrel a) = a\ for a + unfolding \rel_to_urrel_def urrel_to_\rel_def + apply (rule ext) + apply (subst Abs_urrel_inverse) + by (auto simp: AOT_model_proposition_choice_simp) +next + show \AOT_urrel_\equiv (\rel_to_urrel a) (\rel_to_urrel a)\ for a + unfolding \rel_to_urrel_def AOT_urrel_\equiv_def + apply (subst (1 2) Abs_urrel_inverse) + by (auto simp: AOT_model_proposition_choice_simp) +next + show \AOT_urrel_\equiv r s = (AOT_urrel_\equiv r r \ AOT_urrel_\equiv s s \ + urrel_to_\rel r = urrel_to_\rel s)\ for r s + proof + assume \AOT_urrel_\equiv r s\ + hence \AOT_model_valid_in v (Rep_urrel r (\\ u)) = + AOT_model_valid_in v (Rep_urrel s (\\ u))\ for u v + using AOT_urrel_\equiv_def by metis + hence \urrel_to_\rel r = urrel_to_\rel s\ + unfolding urrel_to_\rel_def + by simp + thus \AOT_urrel_\equiv r r \ AOT_urrel_\equiv s s \ + urrel_to_\rel r = urrel_to_\rel s\ + unfolding AOT_urrel_\equiv_def + by auto + next + assume \AOT_urrel_\equiv r r \ AOT_urrel_\equiv s s \ + urrel_to_\rel r = urrel_to_\rel s\ + hence \AOT_model_valid_in v (Rep_urrel r (\\ u)) = + AOT_model_valid_in v (Rep_urrel s (\\ u))\ for u v + by (metis urrel_to_\rel_def) + thus \AOT_urrel_\equiv r s\ + using AOT_urrel_\equiv_def by presburger + qed +qed + +specification (\\') + \\_eq_ord_exts_all: + \\\' a = \\' b \ (\s . urrel_to_\rel s = urrel_to_\rel r \ s \ a) + \ (\ s . urrel_to_\rel s = urrel_to_\rel r \ s \ b)\ + \\_eq_ord_exts_ex: + \\\' a = \\' b \ (\ s . s \ a \ urrel_to_\rel s = urrel_to_\rel r) + \ (\s . s \ b \ urrel_to_\rel s = urrel_to_\rel r)\ +proof - + define \\_wit_intersection where + \\\_wit_intersection \ \ urrels . + {ordext . \urrel . urrel_to_\rel urrel = ordext \ urrel \ urrels}\ + define \\_wit_union where + \\\_wit_union \ \ urrels . + {ordext . \urrel\urrels . urrel_to_\rel urrel = ordext}\ + + let ?\\_wit = \\ urrels . + let ordexts = \\_wit_intersection urrels in + let ordexts' = \\_wit_union urrels in + (ordexts, ordexts', undefined)\ + define \\_wit :: \urrel set \ \\ where + \\\_wit \ \ urrels . Abs_\ (?\\_wit urrels)\ + { + fix a b :: \urrel set\ and r s + assume \\\_wit a = \\_wit b\ + hence 0: \{ordext. \urrel. urrel_to_\rel urrel = ordext \ urrel \ a} = + {ordext. \urrel. urrel_to_\rel urrel = ordext \ urrel \ b}\ + unfolding \\_wit_def Let_def + apply (subst (asm) Abs_\_inject) + by (auto simp: \\_wit_intersection_def \\_wit_union_def) + assume \urrel_to_\rel s = urrel_to_\rel r \ s \ a\ for s + hence \urrel_to_\rel r \ + {ordext. \urrel. urrel_to_\rel urrel = ordext \ urrel \ a}\ + by auto + hence \urrel_to_\rel r \ + {ordext. \urrel. urrel_to_\rel urrel = ordext \ urrel \ b}\ + using 0 by blast + moreover assume \urrel_to_\rel s = urrel_to_\rel r\ + ultimately have \s \ b\ + by blast + } + moreover { + fix a b :: \urrel set\ and s r + assume \\\_wit a = \\_wit b\ + hence 0: \{ordext. \urrel \ a. urrel_to_\rel urrel = ordext} = + {ordext. \urrel \ b. urrel_to_\rel urrel = ordext}\ + unfolding \\_wit_def + using Abs_\_inject \\_wit_union_def by auto + assume \s \ a\ + hence \urrel_to_\rel s \ {ordext. \urrel \ a. urrel_to_\rel urrel = ordext}\ + by blast + moreover assume \urrel_to_\rel s = urrel_to_\rel r\ + ultimately have \urrel_to_\rel r \ + {ordext. \urrel \ b. urrel_to_\rel urrel = ordext}\ + using "0" by argo + hence \\s. s \ b \ urrel_to_\rel s = urrel_to_\rel r\ + by blast + } + ultimately show ?thesis + by (safe intro!: exI[where x=\\_wit]; metis) +qed + +text\We enable the extended model version.\ +abbreviation (input) AOT_ExtendedModel where \AOT_ExtendedModel \ True\ + +text\Individual terms are either ordinary objects, represented by ordinary urelements, + abstract objects, modelled as sets of urrelations, or null objects, used to + represent non-denoting definite descriptions.\ +datatype \ = \\ \ | \\ \urrel set\ | is_null\: null\ null + +text\The mapping from abstract objects to urelements can be naturally + lifted to a surjective mapping from individual terms to urelements.\ +primrec \\ :: \\\\\ where + \\\ (\\ x) = \\ x\ +| \\\ (\\ x) = \\ (\\ x)\ +| \\\ (null\ x) = null\ x\ + +lemma \\_surj: \surj \\\ + using \\_surj by (metis \\.simps(1) \\.simps(2) \\.simps(3) \.exhaust surj_def) + +text\By construction if the urelement of an individual term is exemplified by + an urrelation, it cannot be a null-object.\ +lemma urrel_null_false: + assumes \AOT_model_valid_in w (Rep_urrel f (\\ x))\ + shows \\is_null\ x\ + by (metis (mono_tags, lifting) assms Rep_urrel \.collapse(3) \\.simps(3) + mem_Collect_eq) + +text\AOT requires any ordinary object to be @{emph \possibly concrete\} and that + there is an object that is not actually, but possibly concrete.\ +consts AOT_model_concrete\ :: \\ \ w \ bool\ +specification (AOT_model_concrete\) + AOT_model_\_concrete_in_some_world: + \\ w . AOT_model_concrete\ x w\ + AOT_model_contingent_object: + \\ x w . AOT_model_concrete\ x w \ \AOT_model_concrete\ x w\<^sub>0\ + by (rule exI[where x=\\_ w. w \ w\<^sub>0\]) (auto simp: AOT_model_nonactual_world) + +text\We define a type class for AOT's terms specifying the conditions under which + objects of that type denote and require the set of denoting terms to be + non-empty.\ +class AOT_Term = + fixes AOT_model_denotes :: \'a \ bool\ + assumes AOT_model_denoting_ex: \\ x . AOT_model_denotes x\ + +text\All types except the type of propositions involve non-denoting terms. We + define a refined type class for those.\ +class AOT_IncompleteTerm = AOT_Term + + assumes AOT_model_nondenoting_ex: \\ x . \AOT_model_denotes x\ + +text\Generic non-denoting term.\ +definition AOT_model_nondenoting :: \'a::AOT_IncompleteTerm\ where + \AOT_model_nondenoting \ SOME \ . \AOT_model_denotes \\ +lemma AOT_model_nondenoing: \\AOT_model_denotes (AOT_model_nondenoting)\ + using someI_ex[OF AOT_model_nondenoting_ex] + unfolding AOT_model_nondenoting_def by blast + +text\@{const AOT_model_denotes} can trivially be extended to products of types.\ +instantiation prod :: (AOT_Term, AOT_Term) AOT_Term +begin +definition AOT_model_denotes_prod :: \'a\'b \ bool\ where + \AOT_model_denotes_prod \ \(x,y) . AOT_model_denotes x \ AOT_model_denotes y\ +instance proof + show \\x::'a\'b. AOT_model_denotes x\ + by (simp add: AOT_model_denotes_prod_def AOT_model_denoting_ex) +qed +end + +text\We specify a transformation of proposition-valued functions on terms, s.t. + the result is fully determined by @{emph \regular\} terms. This will be required + for modelling n-ary relations as functions on tuples while preserving AOT's + definition of n-ary relation identity.\ +locale AOT_model_irregular_spec = + fixes AOT_model_irregular :: \('a \ \) \ 'a \ \\ + and AOT_model_regular :: \'a \ bool\ + and AOT_model_term_equiv :: \'a \ 'a \ bool\ + assumes AOT_model_irregular_false: + \\AOT_model_valid_in w (AOT_model_irregular \ x)\ + assumes AOT_model_irregular_equiv: + \AOT_model_term_equiv x y \ + AOT_model_irregular \ x = AOT_model_irregular \ y\ + assumes AOT_model_irregular_eqI: + \(\ x . AOT_model_regular x \ \ x = \ x) \ + AOT_model_irregular \ x = AOT_model_irregular \ x\ + +text\We introduce a type class for individual terms that specifies being regular, + being equivalent (i.e. conceptually @{emph \sharing urelements\}) and the + transformation on proposition-valued functions as specified above.\ +class AOT_IndividualTerm = AOT_IncompleteTerm + + fixes AOT_model_regular :: \'a \ bool\ + fixes AOT_model_term_equiv :: \'a \ 'a \ bool\ + fixes AOT_model_irregular :: \('a \ \) \ 'a \ \\ + assumes AOT_model_irregular_nondenoting: + \\AOT_model_regular x \ \AOT_model_denotes x\ + assumes AOT_model_term_equiv_part_equivp: + \equivp AOT_model_term_equiv\ + assumes AOT_model_term_equiv_denotes: + \AOT_model_term_equiv x y \ (AOT_model_denotes x = AOT_model_denotes y)\ + assumes AOT_model_term_equiv_regular: + \AOT_model_term_equiv x y \ (AOT_model_regular x = AOT_model_regular y)\ + assumes AOT_model_irregular: + \AOT_model_irregular_spec AOT_model_irregular AOT_model_regular + AOT_model_term_equiv\ + +interpretation AOT_model_irregular_spec AOT_model_irregular AOT_model_regular + AOT_model_term_equiv + using AOT_model_irregular . + +text\Our concrete type for individual terms satisfies the type class of + individual terms. + Note that all unary individuals are regular. In general, an individual term + may be a tuple and is regular, if at most one tuple element does not denote.\ +instantiation \ :: AOT_IndividualTerm +begin +definition AOT_model_term_equiv_\ :: \\ \ \ \ bool\ where + \AOT_model_term_equiv_\ \ \ x y . \\ x = \\ y\ +definition AOT_model_denotes_\ :: \\ \ bool\ where + \AOT_model_denotes_\ \ \ x . \is_null\ x\ +definition AOT_model_regular_\ :: \\ \ bool\ where + \AOT_model_regular_\ \ \ x . True\ +definition AOT_model_irregular_\ :: \(\ \ \) \ \ \ \\ where + \AOT_model_irregular_\ \ SOME \ . AOT_model_irregular_spec \ + AOT_model_regular AOT_model_term_equiv\ +instance proof + show \\x :: \. AOT_model_denotes x\ + by (rule exI[where x=\\\ undefined\]) + (simp add: AOT_model_denotes_\_def) +next + show \\x :: \. \AOT_model_denotes x\ + by (rule exI[where x=\null\ undefined\]) + (simp add: AOT_model_denotes_\_def AOT_model_regular_\_def) +next + show "\AOT_model_regular x \ \ AOT_model_denotes x" for x :: \ + by (simp add: AOT_model_regular_\_def) +next + show \equivp (AOT_model_term_equiv :: \ \ \ \ bool)\ + by (rule equivpI; rule reflpI exI sympI transpI) + (simp_all add: AOT_model_term_equiv_\_def) +next + fix x y :: \ + show \AOT_model_term_equiv x y \ AOT_model_denotes x = AOT_model_denotes y\ + by (metis AOT_model_denotes_\_def AOT_model_term_equiv_\_def \.exhaust_disc + \\.simps \.disc(1,3,5,6) is_\\_def is_\\_def is_null\_def) +next + fix x y :: \ + show \AOT_model_term_equiv x y \ AOT_model_regular x = AOT_model_regular y\ + by (simp add: AOT_model_regular_\_def) +next + have "AOT_model_irregular_spec (\ \ (x::\) . \\<^sub>\ w . False) + AOT_model_regular AOT_model_term_equiv" + by standard (auto simp: AOT_model_proposition_choice_simp) + thus \AOT_model_irregular_spec (AOT_model_irregular::(\\\) \ \ \ \) + AOT_model_regular AOT_model_term_equiv\ + unfolding AOT_model_irregular_\_def by (metis (no_types, lifting) someI_ex) +qed +end + +text\We define relations among individuals as proposition valued functions. + @{emph \Denoting\} unary relations (among @{typ \}) will match the + urrelations introduced above.\ +typedef 'a rel (\<_>\) = \UNIV::('a::AOT_IndividualTerm \ \) set\ .. +setup_lifting type_definition_rel + +text\We will use the transformation specified above to "fix" the behaviour of + functions on irregular terms when defining @{text \\\}-expressions.\ +definition fix_irregular :: \('a::AOT_IndividualTerm \ \) \ ('a \ \)\ where + \fix_irregular \ \ \ x . if AOT_model_regular x + then \ x else AOT_model_irregular \ x\ +lemma fix_irregular_denoting: + \AOT_model_denotes x \ fix_irregular \ x = \ x\ + by (meson AOT_model_irregular_nondenoting fix_irregular_def) +lemma fix_irregular_regular: + \AOT_model_regular x \ fix_irregular \ x = \ x\ + by (meson AOT_model_irregular_nondenoting fix_irregular_def) +lemma fix_irregular_irregular: + \\AOT_model_regular x \ fix_irregular \ x = AOT_model_irregular \ x\ + by (simp add: fix_irregular_def) + +text\Relations among individual terms are (potentially non-denoting) terms. + A relation denotes, if it agrees on all equivalent terms (i.e. terms sharing + urelements), is necessarily false on all non-denoting terms and is + well-behaved on irregular terms.\ +instantiation rel :: (AOT_IndividualTerm) AOT_IncompleteTerm +begin +text\\linelabel{AOT_model_denotes_rel}\ +lift_definition AOT_model_denotes_rel :: \<'a> \ bool\ is + \\ \ . (\ x y . AOT_model_term_equiv x y \ \ x = \ y) \ + (\ w x . AOT_model_valid_in w (\ x) \ AOT_model_denotes x) \ + (\ x . \AOT_model_regular x \ \ x = AOT_model_irregular \ x)\ . +instance proof + have \AOT_model_irregular (fix_irregular \) x = AOT_model_irregular \ x\ + for \ and x :: 'a + by (rule AOT_model_irregular_eqI) (simp add: fix_irregular_def) + thus \\ x :: <'a> . AOT_model_denotes x\ + by (safe intro!: exI[where x=\Abs_rel (fix_irregular (\x. \\<^sub>\ w . False))\]) + (transfer; auto simp: AOT_model_proposition_choice_simp fix_irregular_def + AOT_model_irregular_equiv AOT_model_term_equiv_regular + AOT_model_irregular_false) +next + show \\f :: <'a> . \AOT_model_denotes f\ + by (rule exI[where x=\Abs_rel (\x. \\<^sub>\ w . True)\]; + auto simp: AOT_model_denotes_rel.abs_eq AOT_model_nondenoting_ex + AOT_model_proposition_choice_simp) +qed +end + +text\Auxiliary lemmata.\ + +lemma AOT_model_term_equiv_eps: + shows \AOT_model_term_equiv (Eps (AOT_model_term_equiv \)) \\ + and \AOT_model_term_equiv \ (Eps (AOT_model_term_equiv \))\ + and \AOT_model_term_equiv \ \' \ + (Eps (AOT_model_term_equiv \)) = (Eps (AOT_model_term_equiv \'))\ + apply (metis AOT_model_term_equiv_part_equivp equivp_def someI_ex) + apply (metis AOT_model_term_equiv_part_equivp equivp_def someI_ex) + by (metis AOT_model_term_equiv_part_equivp equivp_def) + +lemma AOT_model_denotes_Abs_rel_fix_irregularI: + assumes \\ x y . AOT_model_term_equiv x y \ \ x = \ y\ + and \\ w x . AOT_model_valid_in w (\ x) \ AOT_model_denotes x\ + shows \AOT_model_denotes (Abs_rel (fix_irregular \))\ +proof - + have \AOT_model_irregular \ x = AOT_model_irregular + (\x. if AOT_model_regular x then \ x else AOT_model_irregular \ x) x\ + if \\ AOT_model_regular x\ + for x + by (rule AOT_model_irregular_eqI) auto + thus ?thesis + unfolding AOT_model_denotes_rel.rep_eq + using assms by (auto simp: AOT_model_irregular_false Abs_rel_inverse + AOT_model_irregular_equiv fix_irregular_def + AOT_model_term_equiv_regular) +qed + +lemma AOT_model_term_equiv_rel_equiv: + assumes \AOT_model_denotes x\ + and \AOT_model_denotes y\ + shows \AOT_model_term_equiv x y = (\ \ w . AOT_model_denotes \ \ + AOT_model_valid_in w (Rep_rel \ x) = AOT_model_valid_in w (Rep_rel \ y))\ +proof + assume \AOT_model_term_equiv x y\ + thus \\ \ w . AOT_model_denotes \ \ AOT_model_valid_in w (Rep_rel \ x) = + AOT_model_valid_in w (Rep_rel \ y)\ + by (simp add: AOT_model_denotes_rel.rep_eq) +next + have 0: \(AOT_model_denotes x' \ AOT_model_term_equiv x' y) = + (AOT_model_denotes y' \ AOT_model_term_equiv y' y)\ + if \AOT_model_term_equiv x' y'\ for x' y' + by (metis that AOT_model_term_equiv_denotes AOT_model_term_equiv_part_equivp + equivp_def) + assume \\ \ w . AOT_model_denotes \ \ AOT_model_valid_in w (Rep_rel \ x) = + AOT_model_valid_in w (Rep_rel \ y)\ + moreover have \AOT_model_denotes (Abs_rel (fix_irregular + (\ x . \\<^sub>\ w . AOT_model_denotes x \ AOT_model_term_equiv x y)))\ + (is "AOT_model_denotes ?r") + by (rule AOT_model_denotes_Abs_rel_fix_irregularI) + (auto simp: 0 AOT_model_denotes_rel.rep_eq Abs_rel_inverse fix_irregular_def + AOT_model_proposition_choice_simp AOT_model_irregular_false) + ultimately have \AOT_model_valid_in w (Rep_rel ?r x) = + AOT_model_valid_in w (Rep_rel ?r y)\ for w + by blast + thus \AOT_model_term_equiv x y\ + by (simp add: Abs_rel_inverse AOT_model_proposition_choice_simp + fix_irregular_denoting[OF assms(1)] AOT_model_term_equiv_part_equivp + fix_irregular_denoting[OF assms(2)] assms equivp_reflp) +qed + +text\Denoting relations among terms of type @{typ \} correspond to urrelations.\ + +definition rel_to_urrel :: \<\> \ urrel\ where + \rel_to_urrel \ \ \ . Abs_urrel (\ u . Rep_rel \ (SOME x . \\ x = u))\ +definition urrel_to_rel :: \urrel \ <\>\ where + \urrel_to_rel \ \ \ . Abs_rel (\ x . Rep_urrel \ (\\ x))\ +definition AOT_rel_equiv :: \<'a::AOT_IndividualTerm> \ <'a> \ bool\ where + \AOT_rel_equiv \ \ f g . AOT_model_denotes f \ AOT_model_denotes g \ f = g\ + +lemma urrel_quotient3: \Quotient3 AOT_rel_equiv rel_to_urrel urrel_to_rel\ +proof (rule Quotient3I) + have \(\u. Rep_urrel a (\\ (SOME x. \\ x = u))) = (\u. Rep_urrel a u)\ for a + by (rule ext) (metis (mono_tags, lifting) \\_surj surj_f_inv_f verit_sko_ex') + thus \rel_to_urrel (urrel_to_rel a) = a\ for a + by (simp add: Abs_rel_inverse rel_to_urrel_def urrel_to_rel_def + Rep_urrel_inverse) +next + show \AOT_rel_equiv (urrel_to_rel a) (urrel_to_rel a)\ for a + unfolding AOT_rel_equiv_def urrel_to_rel_def + by transfer (simp add: AOT_model_regular_\_def AOT_model_denotes_\_def + AOT_model_term_equiv_\_def urrel_null_false) +next + { + fix a + assume \\w x. AOT_model_valid_in w (a x) \ \ is_null\ x\ + hence \(\u. a (SOME x. \\ x = u)) \ + {\. \x w. \ AOT_model_valid_in w (\ (null\ x))}\ + by (simp; metis (mono_tags, lifting) \.exhaust_disc \\.simps \.disc(1,3,5) + \.disc(6) is_\\_def is_\\_def someI_ex) + } note 1 = this + { + fix r s :: \\ \ \\ + assume A: \\x y. AOT_model_term_equiv x y \ r x = r y\ + assume \\w x. AOT_model_valid_in w (r x) \ AOT_model_denotes x\ + hence 2: \(\u. r (SOME x. \\ x = u)) \ + {\. \x w. \ AOT_model_valid_in w (\ (null\ x))}\ + using 1 AOT_model_denotes_\_def by meson + assume B: \\x y. AOT_model_term_equiv x y \ s x = s y\ + assume \\w x. AOT_model_valid_in w (s x) \ AOT_model_denotes x\ + hence 3: \(\u. s (SOME x. \\ x = u)) \ + {\. \x w. \ AOT_model_valid_in w (\ (null\ x))}\ + using 1 AOT_model_denotes_\_def by meson + assume \Abs_urrel (\u. r (SOME x. \\ x = u)) = + Abs_urrel (\u. s (SOME x. \\ x = u))\ + hence 4: \r (SOME x. \\ x = u) = s (SOME x::\. \\ x = u)\ for u + unfolding Abs_urrel_inject[OF 2 3] by metis + have \r x = s x\ for x + using 4[of \\\ x\] + by (metis (mono_tags, lifting) A B AOT_model_term_equiv_\_def someI_ex) + hence \r = s\ by auto + } + thus \AOT_rel_equiv r s = (AOT_rel_equiv r r \ AOT_rel_equiv s s \ + rel_to_urrel r = rel_to_urrel s)\ for r s + unfolding AOT_rel_equiv_def rel_to_urrel_def + by transfer auto +qed + +lemma urrel_quotient: + \Quotient AOT_rel_equiv rel_to_urrel urrel_to_rel + (\x y. AOT_rel_equiv x x \ rel_to_urrel x = y)\ + using Quotient3_to_Quotient[OF urrel_quotient3] by auto + +text\Unary individual terms are always regular and equipped with encoding and + concreteness. The specification of the type class anticipates the required + properties for deriving the axiom system.\ +class AOT_UnaryIndividualTerm = + fixes AOT_model_enc :: \'a \ <'a::AOT_IndividualTerm> \ bool\ + and AOT_model_concrete :: \w \ 'a \ bool\ + assumes AOT_model_unary_regular: + \AOT_model_regular x\ \ \All unary individual terms are regular.\ + and AOT_model_enc_relid: + \AOT_model_denotes F \ + AOT_model_denotes G \ + (\ x . AOT_model_enc x F \ AOT_model_enc x G) + \ F = G\ + and AOT_model_A_objects: + \\x . AOT_model_denotes x \ + (\w. \ AOT_model_concrete w x) \ + (\F. AOT_model_denotes F \ AOT_model_enc x F = \ F)\ + and AOT_model_contingent: + \\ x w. AOT_model_concrete w x \ \ AOT_model_concrete w\<^sub>0 x\ + and AOT_model_nocoder: + \AOT_model_concrete w x \ \AOT_model_enc x F\ + and AOT_model_concrete_equiv: + \AOT_model_term_equiv x y \ + AOT_model_concrete w x = AOT_model_concrete w y\ + and AOT_model_concrete_denotes: + \AOT_model_concrete w x \ AOT_model_denotes x\ + \ \The following are properties that will only hold in the extended models.\ + and AOT_model_enc_indistinguishable_all: + \AOT_ExtendedModel \ + AOT_model_denotes a \ \(\ w . AOT_model_concrete w a) \ + AOT_model_denotes b \ \(\ w . AOT_model_concrete w b) \ + AOT_model_denotes \ \ + (\ \' . AOT_model_denotes \' \ + (\ v . AOT_model_valid_in v (Rep_rel \' a) = + AOT_model_valid_in v (Rep_rel \' b))) \ + (\ \' . AOT_model_denotes \' \ + (\ v x . \ w . AOT_model_concrete w x \ + AOT_model_valid_in v (Rep_rel \' x) = + AOT_model_valid_in v (Rep_rel \ x)) \ + AOT_model_enc a \') \ + (\ \' . AOT_model_denotes \' \ + (\ v x . \ w . AOT_model_concrete w x \ + AOT_model_valid_in v (Rep_rel \' x) = + AOT_model_valid_in v (Rep_rel \ x)) \ + AOT_model_enc b \')\ + and AOT_model_enc_indistinguishable_ex: + \AOT_ExtendedModel \ + AOT_model_denotes a \ \(\ w . AOT_model_concrete w a) \ + AOT_model_denotes b \ \(\ w . AOT_model_concrete w b) \ + AOT_model_denotes \ \ + (\ \' . AOT_model_denotes \' \ + (\ v . AOT_model_valid_in v (Rep_rel \' a) = + AOT_model_valid_in v (Rep_rel \' b))) \ + (\ \' . AOT_model_denotes \' \ AOT_model_enc a \' \ + (\ v x . (\ w . AOT_model_concrete w x) \ + AOT_model_valid_in v (Rep_rel \' x) = + AOT_model_valid_in v (Rep_rel \ x))) \ + (\ \' . AOT_model_denotes \' \ AOT_model_enc b \' \ + (\ v x . (\ w . AOT_model_concrete w x) \ + AOT_model_valid_in v (Rep_rel \' x) = + AOT_model_valid_in v (Rep_rel \ x)))\ + +text\Instantiate the class of unary individual terms for our concrete type of + individual terms @{typ \}.\ +instantiation \ :: AOT_UnaryIndividualTerm +begin + +definition AOT_model_enc_\ :: \\ \ <\> \ bool\ where + \AOT_model_enc_\ \ \ x F . + case x of \\ a \ AOT_model_denotes F \ rel_to_urrel F \ a + | _ \ False\ +primrec AOT_model_concrete_\ :: \w \ \ \ bool\ where + \AOT_model_concrete_\ w (\\ x) = AOT_model_concrete\ x w\ +| \AOT_model_concrete_\ w (\\ x) = False\ +| \AOT_model_concrete_\ w (null\ x) = False\ + +lemma AOT_meta_A_objects_\: + \\x :: \. AOT_model_denotes x \ + (\w. \ AOT_model_concrete w x) \ + (\F. AOT_model_denotes F \ AOT_model_enc x F = \ F)\ for \ + apply (rule exI[where x=\\\ {f . \ (urrel_to_rel f)}\]) + apply (simp add: AOT_model_enc_\_def AOT_model_denotes_\_def) + by (metis (no_types, lifting) AOT_rel_equiv_def urrel_quotient + Quotient_rep_abs_fold_unmap) + +instance proof + show \AOT_model_regular x\ for x :: \ + by (simp add: AOT_model_regular_\_def) +next + fix F G :: \<\>\ + assume \AOT_model_denotes F\ + moreover assume \AOT_model_denotes G\ + moreover assume \\x. AOT_model_enc x F = AOT_model_enc x G\ + moreover obtain x where \\G. AOT_model_denotes G \ AOT_model_enc x G = (F = G)\ + using AOT_meta_A_objects_\ by blast + ultimately show \F = G\ by blast +next + show \\x :: \. AOT_model_denotes x \ + (\w. \ AOT_model_concrete w x) \ + (\F. AOT_model_denotes F \ AOT_model_enc x F = \ F)\ for \ + using AOT_meta_A_objects_\ . +next + show \\ (x::\) w. AOT_model_concrete w x \ \ AOT_model_concrete w\<^sub>0 x\ + using AOT_model_concrete_\.simps(1) AOT_model_contingent_object by blast +next + show \AOT_model_concrete w x \ \ AOT_model_enc x F\ for w and x :: \ and F + by (metis AOT_model_concrete_\.simps(2) AOT_model_enc_\_def \.case_eq_if + \.collapse(2)) +next + show \AOT_model_concrete w x = AOT_model_concrete w y\ + if \AOT_model_term_equiv x y\ + for x y :: \ and w + using that by (induct x; induct y; auto simp: AOT_model_term_equiv_\_def) +next + show \AOT_model_concrete w x \ AOT_model_denotes x\ for w and x :: \ + by (metis AOT_model_concrete_\.simps(3) AOT_model_denotes_\_def \.collapse(3)) +(* Extended models only *) +next + fix \ \' :: \ and \ \' :: \<\>\ and w :: w + assume ext: \AOT_ExtendedModel\ + assume \AOT_model_denotes \\ + moreover assume \\w. AOT_model_concrete w \\ + ultimately obtain a where a_def: \\\ a = \\ + by (metis AOT_model_\_concrete_in_some_world AOT_model_concrete_\.simps(1) + AOT_model_denotes_\_def \.discI(3) \.exhaust_sel) + assume \AOT_model_denotes \'\ + moreover assume \\w. AOT_model_concrete w \'\ + ultimately obtain b where b_def: \\\ b = \'\ + by (metis AOT_model_\_concrete_in_some_world AOT_model_concrete_\.simps(1) + AOT_model_denotes_\_def \.discI(3) \.exhaust_sel) + assume \AOT_model_denotes \' \ AOT_model_valid_in w (Rep_rel \' \) = + AOT_model_valid_in w (Rep_rel \' \')\ for \' w + hence \AOT_model_valid_in w (Rep_urrel r (\\ \)) = + AOT_model_valid_in w (Rep_urrel r (\\ \'))\ for r + by (metis AOT_rel_equiv_def Abs_rel_inverse Quotient3_rel_rep + iso_tuple_UNIV_I urrel_quotient3 urrel_to_rel_def) + hence \let r = (Abs_urrel (\ u . \\<^sub>\ w . u = \\ \)) in + AOT_model_valid_in w (Rep_urrel r (\\ \)) = + AOT_model_valid_in w (Rep_urrel r (\\ \'))\ + by presburger + hence \\_eq: \\\ a = \\ b\ + unfolding Let_def + apply (subst (asm) (1 2) Abs_urrel_inverse) + using AOT_model_proposition_choice_simp a_def b_def by force+ + assume \_den: \AOT_model_denotes \\ + have \\AOT_model_valid_in w (Rep_rel \ (SOME xa. \\ xa = null\ x))\ for x w + by (metis (mono_tags, lifting) AOT_model_denotes_\_def + AOT_model_denotes_rel.rep_eq \.exhaust_disc \\.simps(1,2,3) + \AOT_model_denotes \\ \.disc(8,9) \.distinct(3) + is_\\_def is_\\_def verit_sko_ex') + moreover have \Rep_rel \ (\\ x) = Rep_rel \ (SOME y. \\ y = \\ x)\ for x + by (metis (mono_tags, lifting) AOT_model_denotes_rel.rep_eq + AOT_model_term_equiv_\_def \\.simps(1) \_den verit_sko_ex') + ultimately have \Rep_rel \ (\\ x) = Rep_urrel (rel_to_urrel \) (\\ x)\ for x + unfolding rel_to_urrel_def + by (subst Abs_urrel_inverse) auto + hence \\r . \ x . Rep_rel \ (\\ x) = Rep_urrel r (\\ x)\ + by (auto intro!: exI[where x=\rel_to_urrel \\]) + then obtain r where r_prop: \Rep_rel \ (\\ x) = Rep_urrel r (\\ x)\ for x + by blast + assume \AOT_model_denotes \' \ + (\v x. \w. AOT_model_concrete w x \ + AOT_model_valid_in v (Rep_rel \' x) = + AOT_model_valid_in v (Rep_rel \ x)) \ AOT_model_enc \ \'\ for \' + hence \AOT_model_denotes \' \ + (\v x. AOT_model_valid_in v (Rep_rel \' (\\ x)) = + AOT_model_valid_in v (Rep_rel \ (\\ x))) \ AOT_model_enc \ \'\ for \' + by (metis AOT_model_concrete_\.simps(2) AOT_model_concrete_\.simps(3) + \.exhaust_disc is_\\_def is_\\_def is_null\_def) + hence \(\v x. AOT_model_valid_in v (Rep_urrel r (\\ x)) = + AOT_model_valid_in v (Rep_rel \ (\\ x))) \ r \ a\ for r + unfolding a_def[symmetric] AOT_model_enc_\_def apply simp + by (smt (verit, best) AOT_rel_equiv_def Abs_rel_inverse Quotient3_def + \\.simps(1) iso_tuple_UNIV_I urrel_quotient3 urrel_to_rel_def) + hence \(\v x. AOT_model_valid_in v (Rep_urrel r' (\\ x)) = + AOT_model_valid_in v (Rep_urrel r (\\ x))) \ r' \ a\ for r' + unfolding r_prop. + hence \\s. urrel_to_\rel s = urrel_to_\rel r \ s \ a\ + by (metis urrel_to_\rel_def) + hence 0: \\s. urrel_to_\rel s = urrel_to_\rel r \ s \ b\ + using \\_eq_ord_exts_all \\_eq ext \\_\\' by blast + + assume \'_den: \AOT_model_denotes \'\ + assume \\w. AOT_model_concrete w x \ AOT_model_valid_in v (Rep_rel \' x) = + AOT_model_valid_in v (Rep_rel \ x)\ for v x + hence \AOT_model_valid_in v (Rep_rel \' (\\ x)) = + AOT_model_valid_in v (Rep_rel \ (\\ x))\ for v x + using AOT_model_\_concrete_in_some_world AOT_model_concrete_\.simps(1) + by presburger + hence \AOT_model_valid_in v (Rep_urrel (rel_to_urrel \') (\\ x)) = + AOT_model_valid_in v (Rep_urrel r (\\ x))\ for v x + by (smt (verit, best) AOT_rel_equiv_def Abs_rel_inverse Quotient3_def + \\.simps(1) iso_tuple_UNIV_I r_prop urrel_quotient3 urrel_to_rel_def \'_den) + hence \urrel_to_\rel (rel_to_urrel \') = urrel_to_\rel r\ + by (metis (full_types) AOT_urrel_\equiv_def Quotient3_def urrel_\rel_quot) + hence \rel_to_urrel \' \ b\ using 0 by blast + thus \AOT_model_enc \' \'\ + unfolding b_def[symmetric] AOT_model_enc_\_def by (auto simp: \'_den) +next + fix \ \' :: \ and \ \' :: \<\>\ and w :: w + assume ext: \AOT_ExtendedModel\ + assume \AOT_model_denotes \\ + moreover assume \\w. AOT_model_concrete w \\ + ultimately obtain a where a_def: \\\ a = \\ + by (metis AOT_model_\_concrete_in_some_world AOT_model_concrete_\.simps(1) + AOT_model_denotes_\_def \.discI(3) \.exhaust_sel) + assume \AOT_model_denotes \'\ + moreover assume \\w. AOT_model_concrete w \'\ + ultimately obtain b where b_def: \\\ b = \'\ + by (metis AOT_model_\_concrete_in_some_world AOT_model_concrete_\.simps(1) + AOT_model_denotes_\_def \.discI(3) \.exhaust_sel) + assume \AOT_model_denotes \' \ AOT_model_valid_in w (Rep_rel \' \) = + AOT_model_valid_in w (Rep_rel \' \')\ for \' w + hence \AOT_model_valid_in w (Rep_urrel r (\\ \)) = + AOT_model_valid_in w (Rep_urrel r (\\ \'))\ for r + by (metis AOT_rel_equiv_def Abs_rel_inverse Quotient3_rel_rep + iso_tuple_UNIV_I urrel_quotient3 urrel_to_rel_def) + hence \let r = (Abs_urrel (\ u . \\<^sub>\ w . u = \\ \)) in + AOT_model_valid_in w (Rep_urrel r (\\ \)) = + AOT_model_valid_in w (Rep_urrel r (\\ \'))\ + by presburger + hence \\_eq: \\\ a = \\ b\ + unfolding Let_def + apply (subst (asm) (1 2) Abs_urrel_inverse) + using AOT_model_proposition_choice_simp a_def b_def by force+ + assume \_den: \AOT_model_denotes \\ + have \\AOT_model_valid_in w (Rep_rel \ (SOME xa. \\ xa = null\ x))\ for x w + by (metis (mono_tags, lifting) AOT_model_denotes_\_def + AOT_model_denotes_rel.rep_eq \.exhaust_disc \\.simps(1,2,3) + \AOT_model_denotes \\ \.disc(8) \.disc(9) \.distinct(3) + is_\\_def is_\\_def verit_sko_ex') + moreover have \Rep_rel \ (\\ x) = Rep_rel \ (SOME xa. \\ xa = \\ x)\ for x + by (metis (mono_tags, lifting) AOT_model_denotes_rel.rep_eq + AOT_model_term_equiv_\_def \\.simps(1) \_den verit_sko_ex') + ultimately have \Rep_rel \ (\\ x) = Rep_urrel (rel_to_urrel \) (\\ x)\ for x + unfolding rel_to_urrel_def + by (subst Abs_urrel_inverse) auto + hence \\r . \ x . Rep_rel \ (\\ x) = Rep_urrel r (\\ x)\ + by (auto intro!: exI[where x=\rel_to_urrel \\]) + then obtain r where r_prop: \Rep_rel \ (\\ x) = Rep_urrel r (\\ x)\ for x + by blast + + assume \\\'. AOT_model_denotes \' \ + AOT_model_enc \ \' \ + (\v x. (\w. AOT_model_concrete w x) \ AOT_model_valid_in v (Rep_rel \' x) = + AOT_model_valid_in v (Rep_rel \ x))\ + then obtain \' where + \'_den: \AOT_model_denotes \'\ and + \_enc_\': \AOT_model_enc \ \'\ and + \'_prop: \\w. AOT_model_concrete w x \ + AOT_model_valid_in v (Rep_rel \' x) = + AOT_model_valid_in v (Rep_rel \ x)\ for v x + by blast + have \AOT_model_valid_in v (Rep_rel \' (\\ x)) = + AOT_model_valid_in v (Rep_rel \ (\\ x))\ for x v + by (simp add: AOT_model_\_concrete_in_some_world \'_prop) + hence 0: \AOT_urrel_\equiv (rel_to_urrel \') (rel_to_urrel \)\ + unfolding AOT_urrel_\equiv_def + by (smt (verit) AOT_rel_equiv_def Abs_rel_inverse Quotient3_def + \\.simps(1) iso_tuple_UNIV_I urrel_quotient3 urrel_to_rel_def + \_den \'_den) + have \rel_to_urrel \' \ a\ + and \urrel_to_\rel (rel_to_urrel \') = urrel_to_\rel (rel_to_urrel \)\ + apply (metis AOT_model_enc_\_def \.simps(11) \_enc_\' a_def) + by (metis Quotient3_rel 0 urrel_\rel_quot) + hence \\s. s \ b \ urrel_to_\rel s = urrel_to_\rel (rel_to_urrel \)\ + using \\_eq_ord_exts_ex \\_eq ext \\_\\' by blast + then obtain s where + s_prop: \s \ b \ urrel_to_\rel s = urrel_to_\rel (rel_to_urrel \)\ + by blast + then obtain \'' where + \''_prop: \rel_to_urrel \'' = s\ and \''_den: \AOT_model_denotes \''\ + by (metis AOT_rel_equiv_def Quotient3_def urrel_quotient3) + moreover have \AOT_model_enc \' \''\ + by (metis AOT_model_enc_\_def \''_den \''_prop \.simps(11) b_def s_prop) + moreover have \AOT_model_valid_in v (Rep_rel \'' x) = + AOT_model_valid_in v (Rep_rel \ x)\ + if \\w. AOT_model_concrete w x\ for v x + proof(insert that) + assume \\w. AOT_model_concrete w x\ + then obtain u where x_def: \x = \\ u\ + by (metis AOT_model_concrete_\.simps(2,3) \.exhaust) + show \AOT_model_valid_in v (Rep_rel \'' x) = + AOT_model_valid_in v (Rep_rel \ x)\ + unfolding x_def + by (smt (verit, best) AOT_rel_equiv_def Abs_rel_inverse Quotient3_def + \''_den \''_prop \_den \\.simps(1) iso_tuple_UNIV_I s_prop + urrel_quotient3 urrel_to_\rel_def urrel_to_rel_def) + qed + ultimately show \\\'. AOT_model_denotes \' \ AOT_model_enc \' \' \ + (\v x. (\w. AOT_model_concrete w x) \ AOT_model_valid_in v (Rep_rel \' x) = + AOT_model_valid_in v (Rep_rel \ x))\ + apply (safe intro!: exI[where x=\'']) + by auto +qed +end + +text\Products of unary individual terms and individual terms are individual terms. + A tuple is regular, if at most one element does not denote. I.e. a pair is + regular, if the first (unary) element denotes and the second is regular (i.e. + at most one of its recursive tuple elements does not denote), or the first does + not denote, but the second denotes (i.e. all its recursive tuple elements + denote).\ +instantiation prod :: (AOT_UnaryIndividualTerm, AOT_IndividualTerm) AOT_IndividualTerm +begin +definition AOT_model_regular_prod :: \'a\'b \ bool\ where + \AOT_model_regular_prod \ \ (x,y) . AOT_model_denotes x \ AOT_model_regular y \ + \AOT_model_denotes x \ AOT_model_denotes y\ +definition AOT_model_term_equiv_prod :: \'a\'b \ 'a\'b \ bool\ where + \AOT_model_term_equiv_prod \ \ (x\<^sub>1,y\<^sub>1) (x\<^sub>2,y\<^sub>2) . + AOT_model_term_equiv x\<^sub>1 x\<^sub>2 \ AOT_model_term_equiv y\<^sub>1 y\<^sub>2\ +function AOT_model_irregular_prod :: \('a\'b \ \) \ 'a\'b \ \\ where + AOT_model_irregular_proj2: \AOT_model_denotes x \ + AOT_model_irregular \ (x,y) = + AOT_model_irregular (\y. \ (SOME x' . AOT_model_term_equiv x x', y)) y\ +| AOT_model_irregular_proj1: \\AOT_model_denotes x \ AOT_model_denotes y \ + AOT_model_irregular \ (x,y) = + AOT_model_irregular (\x. \ (x, SOME y' . AOT_model_term_equiv y y')) x\ +| AOT_model_irregular_prod_generic: \\AOT_model_denotes x \ \AOT_model_denotes y \ + AOT_model_irregular \ (x,y) = + (SOME \ . AOT_model_irregular_spec \ AOT_model_regular AOT_model_term_equiv) + \ (x,y)\ + by auto blast +termination using "termination" by blast + +instance proof + obtain x :: 'a and y :: 'b where + \\AOT_model_denotes x\ and \\AOT_model_denotes y\ + by (meson AOT_model_nondenoting_ex AOT_model_denoting_ex) + thus \\x::'a\'b. \AOT_model_denotes x\ + by (auto simp: AOT_model_denotes_prod_def AOT_model_regular_prod_def) +next + show \equivp (AOT_model_term_equiv :: 'a\'b \ 'a\'b \ bool)\ + by (rule equivpI; rule reflpI sympI transpI; + simp add: AOT_model_term_equiv_prod_def AOT_model_term_equiv_part_equivp + equivp_reflp prod.case_eq_if case_prod_unfold equivp_symp) + (metis equivp_transp[OF AOT_model_term_equiv_part_equivp]) +next + show \\AOT_model_regular x \ \ AOT_model_denotes x\ for x :: \'a\'b\ + by (metis (mono_tags, lifting) AOT_model_denotes_prod_def case_prod_unfold + AOT_model_irregular_nondenoting AOT_model_regular_prod_def) +next + fix x y :: \'a\'b\ + show \AOT_model_term_equiv x y \ AOT_model_denotes x = AOT_model_denotes y\ + by (metis (mono_tags, lifting) AOT_model_denotes_prod_def case_prod_beta + AOT_model_term_equiv_denotes AOT_model_term_equiv_prod_def ) +next + fix x y :: \'a\'b\ + show \AOT_model_term_equiv x y \ AOT_model_regular x = AOT_model_regular y\ + by (induct x; induct y; + simp add: AOT_model_term_equiv_prod_def AOT_model_regular_prod_def) + (meson AOT_model_term_equiv_denotes AOT_model_term_equiv_regular) +next + interpret sp: AOT_model_irregular_spec \\\ (x::'a\'b) . \\<^sub>\ w . False\ + AOT_model_regular AOT_model_term_equiv + by (simp add: AOT_model_irregular_spec_def AOT_model_proposition_choice_simp) + have ex_spec: \\ \ :: ('a\'b \ \) \ 'a\'b \ \ . + AOT_model_irregular_spec \ AOT_model_regular AOT_model_term_equiv\ + using sp.AOT_model_irregular_spec_axioms by blast + have some_spec: \AOT_model_irregular_spec + (SOME \ :: ('a\'b \ \) \ 'a\'b \ \ . + AOT_model_irregular_spec \ AOT_model_regular AOT_model_term_equiv) + AOT_model_regular AOT_model_term_equiv\ + using someI_ex[OF ex_spec] by argo + interpret sp_some: AOT_model_irregular_spec + \SOME \ :: ('a\'b \ \) \ 'a\'b \ \ . + AOT_model_irregular_spec \ AOT_model_regular AOT_model_term_equiv\ + AOT_model_regular AOT_model_term_equiv + using some_spec by blast + show \AOT_model_irregular_spec (AOT_model_irregular :: ('a\'b \ \) \ 'a\'b \ \) + AOT_model_regular AOT_model_term_equiv\ + proof + have \\AOT_model_valid_in w (AOT_model_irregular \ (a, b))\ + for w \ and a :: 'a and b :: 'b + by (induct arbitrary: \ rule: AOT_model_irregular_prod.induct) + (auto simp: AOT_model_irregular_false sp_some.AOT_model_irregular_false) + thus "\AOT_model_valid_in w (AOT_model_irregular \ x)" for w \ and x :: \'a\'b\ + by (induct x) + next + { + fix x\<^sub>1 y\<^sub>1 :: 'a and x\<^sub>2 y\<^sub>2 :: 'b and \ :: \'a\'b\\\ + assume x\<^sub>1y\<^sub>1_equiv: \AOT_model_term_equiv x\<^sub>1 y\<^sub>1\ + moreover assume x\<^sub>2y\<^sub>2_equiv: \AOT_model_term_equiv x\<^sub>2 y\<^sub>2\ + ultimately have xy_equiv: \AOT_model_term_equiv (x\<^sub>1,x\<^sub>2) (y\<^sub>1,y\<^sub>2)\ + by (simp add: AOT_model_term_equiv_prod_def) + { + assume \AOT_model_denotes x\<^sub>1\ + moreover hence \AOT_model_denotes y\<^sub>1\ + using AOT_model_term_equiv_denotes AOT_model_term_equiv_regular + x\<^sub>1y\<^sub>1_equiv x\<^sub>2y\<^sub>2_equiv by blast + ultimately have \AOT_model_irregular \ (x\<^sub>1,x\<^sub>2) = + AOT_model_irregular \ (y\<^sub>1,y\<^sub>2)\ + using AOT_model_irregular_equiv AOT_model_term_equiv_eps(3) + x\<^sub>1y\<^sub>1_equiv x\<^sub>2y\<^sub>2_equiv by fastforce + } + moreover { + assume \~AOT_model_denotes x\<^sub>1 \ AOT_model_denotes x\<^sub>2\ + moreover hence \~AOT_model_denotes y\<^sub>1 \ AOT_model_denotes y\<^sub>2\ + by (meson AOT_model_term_equiv_denotes x\<^sub>1y\<^sub>1_equiv x\<^sub>2y\<^sub>2_equiv) + ultimately have \AOT_model_irregular \ (x\<^sub>1,x\<^sub>2) = + AOT_model_irregular \ (y\<^sub>1,y\<^sub>2)\ + using AOT_model_irregular_equiv AOT_model_term_equiv_eps(3) + x\<^sub>1y\<^sub>1_equiv x\<^sub>2y\<^sub>2_equiv by fastforce + } + moreover { + assume denotes_x: \(\AOT_model_denotes x\<^sub>1 \ \AOT_model_denotes x\<^sub>2)\ + hence denotes_y: \(\AOT_model_denotes y\<^sub>1 \ \AOT_model_denotes y\<^sub>2)\ + by (meson AOT_model_term_equiv_denotes AOT_model_term_equiv_regular + x\<^sub>1y\<^sub>1_equiv x\<^sub>2y\<^sub>2_equiv) + have eps_eq: \Eps (AOT_model_term_equiv x\<^sub>1) = Eps (AOT_model_term_equiv y\<^sub>1)\ + by (simp add: AOT_model_term_equiv_eps(3) x\<^sub>1y\<^sub>1_equiv) + have \AOT_model_irregular \ (x\<^sub>1,x\<^sub>2) = AOT_model_irregular \ (y\<^sub>1,y\<^sub>2)\ + using denotes_x denotes_y + using sp_some.AOT_model_irregular_equiv xy_equiv by auto + } + moreover { + assume denotes_x: \\AOT_model_denotes x\<^sub>1 \ AOT_model_denotes x\<^sub>2\ + hence denotes_y: \\AOT_model_denotes y\<^sub>1 \ AOT_model_denotes y\<^sub>2\ + by (meson AOT_model_term_equiv_denotes x\<^sub>1y\<^sub>1_equiv x\<^sub>2y\<^sub>2_equiv) + have eps_eq: \Eps (AOT_model_term_equiv x\<^sub>2) = Eps (AOT_model_term_equiv y\<^sub>2)\ + by (simp add: AOT_model_term_equiv_eps(3) x\<^sub>2y\<^sub>2_equiv) + have \AOT_model_irregular \ (x\<^sub>1,x\<^sub>2) = AOT_model_irregular \ (y\<^sub>1,y\<^sub>2)\ + using denotes_x denotes_y + using AOT_model_irregular_nondenoting calculation(2) by blast + } + ultimately have \AOT_model_irregular \ (x\<^sub>1,x\<^sub>2) = AOT_model_irregular \ (y\<^sub>1,y\<^sub>2)\ + using AOT_model_term_equiv_denotes AOT_model_term_equiv_regular + sp_some.AOT_model_irregular_equiv x\<^sub>1y\<^sub>1_equiv x\<^sub>2y\<^sub>2_equiv xy_equiv + by blast + } note 0 = this + show \AOT_model_term_equiv x y \ + AOT_model_irregular \ x = AOT_model_irregular \ y\ + for x y :: \'a\'b\ and \ + by (induct x; induct y; simp add: AOT_model_term_equiv_prod_def 0) + next + fix \ \ :: \'a\'b \ \\ + assume \AOT_model_regular x \ \ x = \ x\ for x + hence \\ (x, y) = \ (x, y)\ + if \AOT_model_denotes x \ AOT_model_regular y \ + \AOT_model_denotes x \ AOT_model_denotes y\ for x y + using that unfolding AOT_model_regular_prod_def by simp + hence \AOT_model_irregular \ (x,y) = AOT_model_irregular \ (x,y)\ + for x :: 'a and y :: 'b + proof (induct arbitrary: \ \ rule: AOT_model_irregular_prod.induct) + case (1 x y \) + thus ?case + apply simp + by (meson AOT_model_irregular_eqI AOT_model_irregular_nondenoting + AOT_model_term_equiv_denotes AOT_model_term_equiv_eps(1)) + next + case (2 x y \) + thus ?case + apply simp + by (meson AOT_model_irregular_nondenoting AOT_model_term_equiv_denotes + AOT_model_term_equiv_eps(1)) + next + case (3 x y \) + thus ?case + apply simp + by (metis (mono_tags, lifting) AOT_model_regular_prod_def case_prod_conv + sp_some.AOT_model_irregular_eqI surj_pair) + qed + thus \AOT_model_irregular \ x = AOT_model_irregular \ x\ for x :: \'a\'b\ + by (metis surjective_pairing) + qed +qed +end + +text\Introduction rules for term equivalence on tuple terms.\ +lemma AOT_meta_prod_equivI: + shows "\ (a::'a::AOT_UnaryIndividualTerm) x (y :: 'b::AOT_IndividualTerm) . + AOT_model_term_equiv x y \ AOT_model_term_equiv (a,x) (a,y)" + and "\ (x::'a::AOT_UnaryIndividualTerm) y (b :: 'b::AOT_IndividualTerm) . + AOT_model_term_equiv x y \ AOT_model_term_equiv (x,b) (y,b)" + unfolding AOT_model_term_equiv_prod_def + by (simp add: AOT_model_term_equiv_part_equivp equivp_reflp)+ + +text\The type of propositions are trivial instances of terms.\ + +instantiation \ :: AOT_Term +begin +definition AOT_model_denotes_\ :: \\ \ bool\ where + \AOT_model_denotes_\ \ \_. True\ +instance proof + show \\x::\. AOT_model_denotes x\ + by (simp add: AOT_model_denotes_\_def) +qed +end + +text\AOT's variables are modelled by restricting the type of terms to those terms + that denote.\ +typedef 'a AOT_var = \{ x :: 'a::AOT_Term . AOT_model_denotes x }\ + morphisms AOT_term_of_var AOT_var_of_term + by (simp add: AOT_model_denoting_ex) + +text\Simplify automatically generated theorems and rules.\ +declare AOT_var_of_term_induct[induct del] + AOT_var_of_term_cases[cases del] + AOT_term_of_var_induct[induct del] + AOT_term_of_var_cases[cases del] +lemmas AOT_var_of_term_inverse = AOT_var_of_term_inverse[simplified] + and AOT_var_of_term_inject = AOT_var_of_term_inject[simplified] + and AOT_var_of_term_induct = + AOT_var_of_term_induct[simplified, induct type: AOT_var] + and AOT_var_of_term_cases = + AOT_var_of_term_cases[simplified, cases type: AOT_var] + and AOT_term_of_var = AOT_term_of_var[simplified] + and AOT_term_of_var_cases = + AOT_term_of_var_cases[simplified, induct pred: AOT_term_of_var] + and AOT_term_of_var_induct = + AOT_term_of_var_induct[simplified, induct pred: AOT_term_of_var] + and AOT_term_of_var_inverse = AOT_term_of_var_inverse[simplified] + and AOT_term_of_var_inject = AOT_term_of_var_inject[simplified] + +text\Equivalence by definition is modelled as necessary equivalence.\ +consts AOT_model_equiv_def :: \\ \ \ \ bool\ +specification(AOT_model_equiv_def) + AOT_model_equiv_def: \AOT_model_equiv_def \ \ = (\ v . AOT_model_valid_in v \ = + AOT_model_valid_in v \)\ + by (rule exI[where x=\\ \ \ . \ v . AOT_model_valid_in v \ = + AOT_model_valid_in v \\]) simp + +text\Identity by definition is modelled as identity for denoting terms plus + co-denoting.\ +consts AOT_model_id_def :: \('b \ 'a::AOT_Term) \ ('b \ 'a) \ bool\ +specification(AOT_model_id_def) + AOT_model_id_def: \(AOT_model_id_def \ \) = (\ \ . if AOT_model_denotes (\ \) + then \ \ = \ \ + else \AOT_model_denotes (\ \))\ + by (rule exI[where x="\ \ \ . \ \ . if AOT_model_denotes (\ \) + then \ \ = \ \ + else \AOT_model_denotes (\ \)"]) + blast +text\To reduce definitions by identity without free variables to definitions + by identity with free variables acting on the unit type, we give the unit type + a trivial instantiation to @{class AOT_Term}.\ +instantiation unit :: AOT_Term +begin +definition AOT_model_denotes_unit :: \unit \ bool\ where + \AOT_model_denotes_unit \ \_. True\ +instance proof qed(simp add: AOT_model_denotes_unit_def) +end + +text\Modally-strict and modally-fragile axioms are as necessary, + resp. actually valid propositions.\ +definition AOT_model_axiom where + \AOT_model_axiom \ \ \ . \ v . AOT_model_valid_in v \\ +definition AOT_model_act_axiom where + \AOT_model_act_axiom \ \ \ . AOT_model_valid_in w\<^sub>0 \\ + +lemma AOT_model_axiomI: + assumes \\v . AOT_model_valid_in v \\ + shows \AOT_model_axiom \\ + unfolding AOT_model_axiom_def using assms .. + +lemma AOT_model_act_axiomI: + assumes \AOT_model_valid_in w\<^sub>0 \\ + shows \AOT_model_act_axiom \\ + unfolding AOT_model_act_axiom_def using assms . + +(*<*) +end +(*>*) \ No newline at end of file diff --git a/thys/AOT/AOT_semantics.thy b/thys/AOT/AOT_semantics.thy new file mode 100644 --- /dev/null +++ b/thys/AOT/AOT_semantics.thy @@ -0,0 +1,1801 @@ +(*<*) +theory AOT_semantics + imports AOT_syntax +begin +(*>*) + +section\Abstract Semantics for AOT\ + +specification(AOT_denotes) + \ \Relate object level denoting to meta-denoting. AOT's definitions of + denoting will become derivable at each type.\ + AOT_sem_denotes: \[w \ \\] = AOT_model_denotes \\ + by (rule exI[where x=\\ \ . \\<^sub>\ w . AOT_model_denotes \\]) + (simp add: AOT_model_proposition_choice_simp) + +lemma AOT_sem_var_induct[induct type: AOT_var]: + assumes AOT_denoting_term_case: \\ \ . [v \ \\] \ [v \ \{\}]\ + shows \[v \ \{\}]\ + by (simp add: AOT_denoting_term_case AOT_sem_denotes AOT_term_of_var) + +text\\linelabel{AOT_imp_spec}\ +specification(AOT_imp) + AOT_sem_imp: \[w \ \ \ \] = ([w \ \] \ [w \ \])\ + by (rule exI[where x=\\ \ \ . \\<^sub>\ w . ([w \ \] \ [w \ \])\]) + (simp add: AOT_model_proposition_choice_simp) + +specification(AOT_not) + AOT_sem_not: \[w \ \\] = (\[w \ \])\ + by (rule exI[where x=\\ \ . \\<^sub>\ w . \[w \ \]\]) + (simp add: AOT_model_proposition_choice_simp) + +text\\linelabel{AOT_box_spec}\ +specification(AOT_box) + AOT_sem_box: \[w \ \\] = (\ w . [w \ \])\ + by (rule exI[where x=\\ \ . \\<^sub>\ w . \ w . [w \ \]\]) + (simp add: AOT_model_proposition_choice_simp) + +text\\linelabel{AOT_act_spec}\ +specification(AOT_act) + AOT_sem_act: \[w \ \<^bold>\\] = [w\<^sub>0 \ \]\ + by (rule exI[where x=\\ \ . \\<^sub>\ w . [w\<^sub>0 \ \]\]) + (simp add: AOT_model_proposition_choice_simp) + +text\Derived semantics for basic defined connectives.\ +lemma AOT_sem_conj: \[w \ \ & \] = ([w \ \] \ [w \ \])\ + using AOT_conj AOT_model_equiv_def AOT_sem_imp AOT_sem_not by auto +lemma AOT_sem_equiv: \[w \ \ \ \] = ([w \ \] = [w \ \])\ + using AOT_equiv AOT_sem_conj AOT_model_equiv_def AOT_sem_imp by auto +lemma AOT_sem_disj: \[w \ \ \ \] = ([w \ \] \ [w \ \])\ + using AOT_disj AOT_model_equiv_def AOT_sem_imp AOT_sem_not by auto +lemma AOT_sem_dia: \[w \ \\] = (\ w . [w \ \])\ + using AOT_dia AOT_sem_box AOT_model_equiv_def AOT_sem_not by auto + +specification(AOT_forall) + AOT_sem_forall: \[w \ \\ \{\}] = (\ \ . [w \ \\] \ [w \ \{\}])\ + by (rule exI[where x=\\ op . \\<^sub>\ w . \ \ . [w \ \\] \ [w \ \op \\]\]) + (simp add: AOT_model_proposition_choice_simp) + +lemma AOT_sem_exists: \[w \ \\ \{\}] = (\ \ . [w \ \\] \ [w \ \{\}])\ + unfolding AOT_exists[unfolded AOT_model_equiv_def, THEN spec] + by (simp add: AOT_sem_forall AOT_sem_not) + +text\\linelabel{AOT_eq_spec}\ +specification(AOT_eq) + \ \Relate identity to denoting identity in the meta-logic. AOT's definitions + of identity will become derivable at each type.\ + AOT_sem_eq: \[w \ \ = \'] = ([w \ \\] \ [w \ \'\] \ \ = \')\ + by (rule exI[where x=\\ \ \' . \\<^sub>\ w . [w \ \\] \ [w \ \'\] \ \ = \'\]) + (simp add: AOT_model_proposition_choice_simp) + +text\\linelabel{AOT_desc_spec}\ +specification(AOT_desc) + \ \Descriptions denote, if there is a unique denoting object satisfying the + matrix in the actual world.\ + AOT_sem_desc_denotes: \[w \ \<^bold>\x(\{x})\] = (\! \ . [w\<^sub>0 \ \\] \ [w\<^sub>0 \ \{\}])\ + \ \Denoting descriptions satisfy their matrix in the actual world.\ + AOT_sem_desc_prop: \[w \ \<^bold>\x(\{x})\] \ [w\<^sub>0 \ \{\<^bold>\x(\{x})}]\ + \ \Uniqueness of denoting descriptions.\ + AOT_sem_desc_unique: \[w \ \<^bold>\x(\{x})\] \ [w \ \\] \ [w\<^sub>0 \ \{\}] \ + [w \ \<^bold>\x(\{x}) = \]\ +proof - + have \\x::'a . \AOT_model_denotes x\ + using AOT_model_nondenoting_ex + by blast + text\Note that we may choose a distinct non-denoting object for each matrix. + We do this explicitly merely to convince ourselves that our specification + can still be satisfied.\ + then obtain nondenoting :: \('a \ \) \ 'a\ where + nondenoting: \\ \ . \AOT_model_denotes (nondenoting \)\ + by fast + define desc where + \desc = (\ \ . if (\! \ . [w\<^sub>0 \ \\] \ [w\<^sub>0 \ \{\}]) + then (THE \ . [w\<^sub>0 \ \\] \ [w\<^sub>0 \ \{\}]) + else nondenoting \)\ + { + fix \ :: \'a \ \\ + assume ex1: \\! \ . [w\<^sub>0 \ \\] \ [w\<^sub>0 \ \{\}]\ + then obtain \ where x_prop: "[w\<^sub>0 \ \\] \ [w\<^sub>0 \ \{\}]" + unfolding AOT_sem_denotes by blast + moreover have "(desc \) = \" + unfolding desc_def using x_prop ex1 by fastforce + ultimately have "[w\<^sub>0 \ \desc \\\] \ [w\<^sub>0 \ \\ (desc \)\]" + by blast + } note 1 = this + moreover { + fix \ :: \'a \ \\ + assume nex1: \\! \ . [w\<^sub>0 \ \\] \ [w\<^sub>0 \ \{\}]\ + hence "(desc \) = nondenoting \" by (simp add: desc_def AOT_sem_denotes) + hence "[w \ \\desc \\\]" for w + by (simp add: AOT_sem_denotes nondenoting AOT_sem_not) + } + ultimately have desc_denotes_simp: + \[w \ \desc \\\] = (\! \ . [w\<^sub>0 \ \\] \ [w\<^sub>0 \ \{\}])\ for \ w + by (simp add: AOT_sem_denotes desc_def nondenoting) + have \(\\ w. [w \ \desc \\\] = (\!\. [w\<^sub>0 \ \\] \ [w\<^sub>0 \ \{\}])) \ + (\\ w. [w \ \desc \\\] \ [w\<^sub>0 \ \\ (desc \)\]) \ + (\\ w \. [w \ \desc \\\] \ [w \ \\] \ [w\<^sub>0 \ \{\}] \ + [w \ \desc \\ = \])\ + by (insert 1; auto simp: desc_denotes_simp AOT_sem_eq AOT_sem_denotes + desc_def nondenoting) + thus ?thesis + by (safe intro!: exI[where x=desc]; presburger) +qed + +text\\linelabel{AOT_exe_lambda_spec}\ +specification(AOT_exe AOT_lambda) + \ \Truth conditions of exemplification formulas.\ + AOT_sem_exe: \[w \ [\]\\<^sub>1...\\<^sub>n] = ([w \ \\] \ [w \ \\<^sub>1...\\<^sub>n\] \ + [w \ \Rep_rel \ \\<^sub>1\\<^sub>n\])\ + \ \\eta-conversion for denoting terms; equivalent to AOT's axiom\ + AOT_sem_lambda_eta: \[w \ \\] \ [w \ [\\\<^sub>1...\\<^sub>n [\]\\<^sub>1...\\<^sub>n] = \]\ + \ \\beta-conversion; equivalent to AOT's axiom\ + AOT_sem_lambda_beta: \[w \ [\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\] \ [w \ \\<^sub>1...\\<^sub>n\] \ + [w \ [\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\\<^sub>1...\\<^sub>n] = [w \ \{\\<^sub>1...\\<^sub>n}]\ + \ \Necessary and sufficient conditions for relations to denote. Equivalent + to a theorem of AOT and used to derive the base cases of denoting relations + (cqt.2).\ + AOT_sem_lambda_denotes: \[w \ [\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\] = + (\ v \\<^sub>1\\<^sub>n \\<^sub>1'\\<^sub>n' . [v \ \\<^sub>1...\\<^sub>n\] \ [v \ \\<^sub>1'...\\<^sub>n'\] \ + (\ \ v . [v \ \\] \ [v \ [\]\\<^sub>1...\\<^sub>n] = [v \ [\]\\<^sub>1'...\\<^sub>n']) \ + [v \ \{\\<^sub>1...\\<^sub>n}] = [v \ \{\\<^sub>1'...\\<^sub>n'}])\ + \ \Equivalent to AOT's coexistence axiom.\ + AOT_sem_lambda_coex: \[w \ [\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\] \ + (\ w \\<^sub>1\\<^sub>n . [w \ \\<^sub>1...\\<^sub>n\] \ [w \ \{\\<^sub>1...\\<^sub>n}] = [w \ \{\\<^sub>1...\\<^sub>n}]) \ + [w \ [\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\]\ + \ \Only the unary case of the following should hold, but our specification + has to range over all types. We might move @{const AOT_exe} and + @{const AOT_lambda} to type classes in the future to solve this.\ + AOT_sem_lambda_eq_prop_eq: \\[\\\<^sub>1...\\<^sub>n \]\ = \[\\\<^sub>1...\\<^sub>n \]\ \ \ = \\ + \ \The following is solely required for validating n-ary relation identity + and has the danger of implying artifactual theorems. Possibly avoidable + by moving @{const AOT_exe} and @{const AOT_lambda} to type classes.\ + AOT_sem_exe_denoting: \[w \ \\] \ AOT_exe \ \s = Rep_rel \ \s\ + \ \The following is required for validating the base cases of denoting + relations (cqt.2). A version of this meta-logical identity will + become derivable in future versions of AOT, so this will ultimately not + result in artifactual theorems.\ + AOT_sem_exe_equiv: \AOT_model_term_equiv x y \ AOT_exe \ x = AOT_exe \ y\ +proof - + have \\ x :: <'a> . \AOT_model_denotes x\ + by (rule exI[where x=\Abs_rel (\ x . \\<^sub>\ w. True)\]) + (meson AOT_model_denotes_rel.abs_eq AOT_model_nondenoting_ex + AOT_model_proposition_choice_simp) + define exe :: \<'a> \ 'a \ \\ where + \exe \ \ \ \s . if AOT_model_denotes \ + then Rep_rel \ \s + else (\\<^sub>\ w . False)\ + define lambda :: \('a\\) \ <'a>\ where + \lambda \ \ \ . if AOT_model_denotes (Abs_rel \) + then (Abs_rel \) + else + if (\ \ \' w . (AOT_model_denotes \ \ AOT_model_term_equiv \ \') \ + [w \ \\ \\] = [w \ \\ \'\]) + then + Abs_rel (fix_irregular (\ x . if AOT_model_denotes x + then \ (SOME y . AOT_model_term_equiv x y) + else (\\<^sub>\ w . False))) + else + Abs_rel \\ + have fix_irregular_denoting_simp[simp]: + \fix_irregular (\x. if AOT_model_denotes x then \ x else \ x) \ = \ \\ + if \AOT_model_denotes \\ + for \ :: 'a and \ \ + by (simp add: that fix_irregular_denoting) + have denoting_eps_cong[cong]: + \[w \ \\ (Eps (AOT_model_term_equiv \))\] = [w \ \\ \\]\ + if \AOT_model_denotes \\ + and \\ \ \'. AOT_model_denotes \ \ AOT_model_term_equiv \ \' \ + (\w. [w \ \\ \\] = [w \ \\ \'\])\ + for w :: w and \ :: 'a and \ :: \'a\\\ + using that AOT_model_term_equiv_eps(2) by blast + have exe_rep_rel: \[w \ \exe \ \\<^sub>1\\<^sub>n\] = ([w \ \\] \ [w \ \\<^sub>1...\\<^sub>n\] \ + [w \ \Rep_rel \ \\<^sub>1\\<^sub>n\])\ for w \ \\<^sub>1\\<^sub>n + by (metis AOT_model_denotes_rel.rep_eq exe_def AOT_sem_denotes + AOT_model_proposition_choice_simp) + moreover have \[w \ \\\\] \ [w \ \lambda (exe \)\ = \\\]\ for \ w + by (auto simp: Rep_rel_inverse lambda_def AOT_sem_denotes AOT_sem_eq + AOT_model_denotes_rel_def Abs_rel_inverse exe_def) + moreover have lambda_denotes_beta: + \[w \ \exe (lambda \) \ \] = [w \ \\ \\]\ + if \[w \ \lambda \\\]\ and \[w \ \\\\]\ + for \ \ w + using that unfolding exe_def AOT_sem_denotes + by (auto simp: lambda_def Abs_rel_inverse split: if_split_asm) + moreover have lambda_denotes_simp: + \[w \ \lambda \\\] = (\ v \\<^sub>1\\<^sub>n \\<^sub>1'\\<^sub>n' . [v \ \\<^sub>1...\\<^sub>n\] \ [v \ \\<^sub>1'...\\<^sub>n'\] \ + (\ \ v . [v \ \\] \ [v \ \exe \ \\<^sub>1\\<^sub>n\] = [v \ \exe \ \\<^sub>1'\\<^sub>n'\]) \ + [v \ \{\\<^sub>1...\\<^sub>n}] = [v \ \{\\<^sub>1'...\\<^sub>n'}])\ for \ w + proof + assume \[w \ \lambda \\\]\ + hence \AOT_model_denotes (lambda \)\ + unfolding AOT_sem_denotes by simp + moreover have \[w \ \\ \\] \ [w \ \\ \'\]\ + and \[w \ \\ \'\] \ [w \ \\ \\]\ + if \AOT_model_denotes \\ and \AOT_model_term_equiv \ \'\ + for w \ \' + by (metis (no_types, lifting) AOT_model_denotes_rel.abs_eq lambda_def + that calculation)+ + ultimately show \\ v \\<^sub>1\\<^sub>n \\<^sub>1'\\<^sub>n' . [v \ \\<^sub>1...\\<^sub>n\] \ [v \ \\<^sub>1'...\\<^sub>n'\] \ + (\ \ v . [v \ \\] \ [v \ \exe \ \\<^sub>1\\<^sub>n\] = [v \ \exe \ \\<^sub>1'\\<^sub>n'\]) \ + [v \ \{\\<^sub>1...\\<^sub>n}] = [v \ \{\\<^sub>1'...\\<^sub>n'}]\ + unfolding AOT_sem_denotes + by (metis (no_types) AOT_sem_denotes lambda_denotes_beta) + next + assume \\ v \\<^sub>1\\<^sub>n \\<^sub>1'\\<^sub>n' . [v \ \\<^sub>1...\\<^sub>n\] \ [v \ \\<^sub>1'...\\<^sub>n'\] \ + (\ \ v . [v \ \\] \ [v \ \exe \ \\<^sub>1\\<^sub>n\] = [v \ \exe \ \\<^sub>1'\\<^sub>n'\]) \ + [v \ \{\\<^sub>1...\\<^sub>n}] = [v \ \{\\<^sub>1'...\\<^sub>n'}]\ + hence \[w \ \\ \\] = [w \ \\ \'\]\ + if \AOT_model_denotes \ \ AOT_model_denotes \' \ AOT_model_term_equiv \ \'\ + for w \ \' + using that + by (auto simp: AOT_sem_denotes) + (meson AOT_model_term_equiv_rel_equiv AOT_sem_denotes exe_rep_rel)+ + hence \[w \ \\ \\] = [w \ \\ \'\]\ + if \AOT_model_denotes \ \ AOT_model_term_equiv \ \'\ + for w \ \' + using that AOT_model_term_equiv_denotes by blast + hence \AOT_model_denotes (lambda \)\ + by (auto simp: lambda_def Abs_rel_inverse AOT_model_denotes_rel.abs_eq + AOT_model_irregular_equiv AOT_model_term_equiv_eps(3) + AOT_model_term_equiv_regular fix_irregular_def AOT_sem_denotes + AOT_model_term_equiv_denotes AOT_model_proposition_choice_simp + AOT_model_irregular_false + split: if_split_asm + intro: AOT_model_irregular_eqI) + thus \[w \ \lambda \\\]\ + by (simp add: AOT_sem_denotes) + qed + moreover have \[w \ \lambda \\\]\ + if \[w \ \lambda \\\]\ + and \\ w \\<^sub>1\\<^sub>n . [w \ \\<^sub>1...\\<^sub>n\] \ [w \ \{\\<^sub>1...\\<^sub>n}] = [w \ \{\\<^sub>1...\\<^sub>n}]\ + for \ \ w using that unfolding lambda_denotes_simp by auto + moreover have \[w \ \\] \ exe \ \s = Rep_rel \ \s\ for \ \s w + by (simp add: exe_def AOT_sem_denotes) + moreover have \lambda (\x. p) = lambda (\x. q) \ p = q\ for p q + unfolding lambda_def + by (auto split: if_split_asm simp: Abs_rel_inject fix_irregular_def) + (meson AOT_model_irregular_nondenoting AOT_model_denoting_ex)+ + moreover have \AOT_model_term_equiv x y \ exe \ x = exe \ y\ for x y \ + unfolding exe_def + by (meson AOT_model_denotes_rel.rep_eq) + note calculation = calculation this + show ?thesis + apply (safe intro!: exI[where x=exe] exI[where x=lambda]) + using calculation apply simp_all + using lambda_denotes_simp by blast+ +qed + +lemma AOT_model_lambda_denotes: + \AOT_model_denotes (AOT_lambda \) = (\ v \ \' . + AOT_model_denotes \ \ AOT_model_denotes \' \ AOT_model_term_equiv \ \' \ + [v \ \\ \\] = [v \ \\ \'\])\ +proof(safe) + fix v and \ \' :: 'a + assume \AOT_model_denotes (AOT_lambda \)\ + hence 1: \AOT_model_denotes \\<^sub>1\\<^sub>n \ + AOT_model_denotes \\<^sub>1'\\<^sub>n' \ + (\\ v. AOT_model_denotes \ \ [v \ [\]\\<^sub>1...\\<^sub>n] = [v \ [\]\\<^sub>1'...\\<^sub>n']) \ + [v \ \{\\<^sub>1...\\<^sub>n}] = [v \ \{\\<^sub>1'...\\<^sub>n'}]\ for \\<^sub>1\\<^sub>n \\<^sub>1'\\<^sub>n' v + using AOT_sem_lambda_denotes[simplified AOT_sem_denotes] by blast + { + fix v and \\<^sub>1\\<^sub>n \\<^sub>1'\\<^sub>n' :: 'a + assume d: \AOT_model_denotes \\<^sub>1\\<^sub>n \ AOT_model_denotes \\<^sub>1'\\<^sub>n' \ + AOT_model_term_equiv \\<^sub>1\\<^sub>n \\<^sub>1'\\<^sub>n'\ + hence \\\ w. AOT_model_denotes \ \ [w \ [\]\\<^sub>1...\\<^sub>n] = [w \ [\]\\<^sub>1'...\\<^sub>n']\ + by (metis AOT_sem_exe_equiv) + hence \[v \ \{\\<^sub>1...\\<^sub>n}] = [v \ \{\\<^sub>1'...\\<^sub>n'}]\ using d 1 by auto + } + moreover assume \AOT_model_denotes \\ + moreover assume \AOT_model_denotes \'\ + moreover assume \AOT_model_term_equiv \ \'\ + ultimately show \[v \ \\ \\] \ [v \ \\ \'\]\ + and \[v \ \\ \'\] \ [v \ \\ \\]\ + by auto +next + assume 0: \\ v \ \' . AOT_model_denotes \ \ AOT_model_denotes \' \ + AOT_model_term_equiv \ \' \ [v \ \\ \\] = [v \ \\ \'\]\ + { + fix \\<^sub>1\\<^sub>n \\<^sub>1'\\<^sub>n' :: 'a + assume den: \AOT_model_denotes \\<^sub>1\\<^sub>n\ + moreover assume den': \AOT_model_denotes \\<^sub>1'\\<^sub>n'\ + assume \\\ v . AOT_model_denotes \ \ + [v \ [\]\\<^sub>1...\\<^sub>n] = [v \ [\]\\<^sub>1'...\\<^sub>n']\ + hence \\\ v . AOT_model_denotes \ \ + [v \ \Rep_rel \ \\<^sub>1\\<^sub>n\] = [v \ \Rep_rel \ \\<^sub>1'\\<^sub>n'\]\ + by (simp add: AOT_sem_denotes AOT_sem_exe den den') + hence "AOT_model_term_equiv \\<^sub>1\\<^sub>n \\<^sub>1'\\<^sub>n'" + unfolding AOT_model_term_equiv_rel_equiv[OF den, OF den'] + by argo + hence \[v \ \{\\<^sub>1...\\<^sub>n}] = [v \ \{\\<^sub>1'...\\<^sub>n'}]\ for v + using den den' 0 by blast + } + thus \AOT_model_denotes (AOT_lambda \)\ + unfolding AOT_sem_lambda_denotes[simplified AOT_sem_denotes] + by blast +qed + +specification (AOT_lambda0) + AOT_sem_lambda0: "AOT_lambda0 \ = \" + by (rule exI[where x=\\x. x\]) simp + +specification(AOT_concrete) + AOT_sem_concrete: \[w \ [E!]\] = + AOT_model_concrete w \\ + by (rule exI[where x=\AOT_var_of_term (Abs_rel + (\ x . \\<^sub>\ w . AOT_model_concrete w x))\]; + subst AOT_var_of_term_inverse) + (auto simp: AOT_model_unary_regular AOT_model_concrete_denotes + AOT_model_concrete_equiv AOT_model_regular_\_def + AOT_model_proposition_choice_simp AOT_sem_exe Abs_rel_inverse + AOT_model_denotes_rel_def AOT_sem_denotes) + +lemma AOT_sem_equiv_defI: + assumes \\ v . [v \ \] \ [v \ \]\ + and \\ v . [v \ \] \ [v \ \]\ + shows \AOT_model_equiv_def \ \\ + using AOT_model_equiv_def assms by blast + +lemma AOT_sem_id_defI: + assumes \\ \ v . [v \ \\ \\\] \ [v \ \\ \\ = \\ \\]\ + assumes \\ \ v . \[v \ \\ \\\] \ [v \ \\\ \\\]\ + shows \AOT_model_id_def \ \\ + using assms + unfolding AOT_sem_denotes AOT_sem_eq AOT_sem_not + using AOT_model_id_def[THEN iffD2] + by metis + +lemma AOT_sem_id_def2I: + assumes \\ \ \ v . [v \ \\ \ \\\] \ [v \ \\ \ \\ = \\ \ \\]\ + assumes \\ \ \ v . \[v \ \\ \ \\\] \ [v \ \\\ \ \\\]\ + shows \AOT_model_id_def (case_prod \) (case_prod \)\ + apply (rule AOT_sem_id_defI) + using assms by (auto simp: AOT_sem_conj AOT_sem_not AOT_sem_eq AOT_sem_denotes + AOT_model_denotes_prod_def) + +lemma AOT_sem_id_defE1: + assumes \AOT_model_id_def \ \\ + and \[v \ \\ \\\]\ + shows \[v \ \\ \\ = \\ \\]\ + using assms by (simp add: AOT_model_id_def AOT_sem_denotes AOT_sem_eq) + +lemma AOT_sem_id_defE2: + assumes \AOT_model_id_def \ \\ + and \\[v \ \\ \\\]\ + shows \\[v \ \\ \\\]\ + using assms by (simp add: AOT_model_id_def AOT_sem_denotes AOT_sem_eq) + +lemma AOT_sem_id_def0I: + assumes \\ v . [v \ \\] \ [v \ \ = \]\ + and \\ v . \[v \ \\] \ [v \ \\\]\ + shows \AOT_model_id_def (case_unit \) (case_unit \)\ + apply (rule AOT_sem_id_defI) + using assms + by (simp_all add: AOT_sem_conj AOT_sem_eq AOT_sem_not AOT_sem_denotes + AOT_model_denotes_unit_def case_unit_Unity) + +lemma AOT_sem_id_def0E1: + assumes \AOT_model_id_def (case_unit \) (case_unit \)\ + and \[v \ \\]\ + shows \[v \ \ = \]\ + by (metis (full_types) AOT_sem_id_defE1 assms(1) assms(2) case_unit_Unity) + +lemma AOT_sem_id_def0E2: + assumes \AOT_model_id_def (case_unit \) (case_unit \)\ + and \\[v \ \\]\ + shows \\[v \ \\]\ + by (metis AOT_sem_id_defE2 assms(1) assms(2) case_unit_Unity) + +lemma AOT_sem_id_def0E3: + assumes \AOT_model_id_def (case_unit \) (case_unit \)\ + and \[v \ \\]\ + shows \[v \ \\]\ + using AOT_sem_id_def0E1[OF assms] + by (simp add: AOT_sem_eq AOT_sem_denotes) + +lemma AOT_sem_ordinary_def_denotes: \[w \ [\x \[E!]x]\]\ + unfolding AOT_sem_denotes AOT_model_lambda_denotes + by (auto simp: AOT_sem_dia AOT_model_concrete_equiv + AOT_sem_concrete AOT_sem_denotes) +lemma AOT_sem_abstract_def_denotes: \[w \ [\x \\[E!]x]\]\ + unfolding AOT_sem_denotes AOT_model_lambda_denotes + by (auto simp: AOT_sem_dia AOT_model_concrete_equiv + AOT_sem_concrete AOT_sem_denotes AOT_sem_not) + +text\Relation identity is constructed using an auxiliary abstract projection + mechanism with suitable instantiations for @{typ \} and products.\ +class AOT_RelationProjection = + fixes AOT_sem_proj_id :: \'a::AOT_IndividualTerm \ ('a \ \) \ ('a \ \) \ \\ + assumes AOT_sem_proj_id_prop: + \[v \ \ = \'] = + [v \ \\ & \'\ & \\ (\AOT_sem_proj_id \ (\ \ . \[\]\\) (\ \ . \[\']\\)\)]\ + and AOT_sem_proj_id_refl: + \[v \ \\] \ [v \ [\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}] = [\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]] \ + [v \ \AOT_sem_proj_id \ \ \\]\ + +class AOT_UnaryRelationProjection = AOT_RelationProjection + + assumes AOT_sem_unary_proj_id: + \AOT_sem_proj_id \ \ \ = \[\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}] = [\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\\ + +instantiation \ :: AOT_UnaryRelationProjection +begin +definition AOT_sem_proj_id_\ :: \\ \ (\ \ \) \ (\ \ \) \ \\ where + \AOT_sem_proj_id_\ \ \ \ = \[\z \{z}] = [\z \{z}]\\ +instance proof + show \[v \ \ = \'] = + [v \ \\ & \'\ & \\ (\AOT_sem_proj_id \ (\ \ . \[\]\\) (\ \ . \[\']\\)\)]\ + for v and \ \' :: \<\>\ + unfolding AOT_sem_proj_id_\_def + by (simp add: AOT_sem_eq AOT_sem_conj AOT_sem_denotes AOT_sem_forall) + (metis AOT_sem_denotes AOT_model_denoting_ex AOT_sem_eq AOT_sem_lambda_eta) +next + show \AOT_sem_proj_id \ \ \ = \[\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}] = [\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\\ + for \ :: \ and \ \ + unfolding AOT_sem_proj_id_\_def .. +next + show \[v \ \AOT_sem_proj_id \ \ \\]\ + if \[v \ \\]\ and \[v \ [\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}] = [\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]]\ + for \ :: \ and v \ + using that by (simp add: AOT_sem_proj_id_\_def AOT_sem_eq) +qed +end + +instantiation prod :: + ("{AOT_UnaryRelationProjection, AOT_UnaryIndividualTerm}", AOT_RelationProjection) + AOT_RelationProjection +begin +definition AOT_sem_proj_id_prod :: \'a\'b \ ('a\'b \ \) \ ('a\'b \ \) \ \\ where + \AOT_sem_proj_id_prod \ \ (x,y) \ \ . \[\z \\ (z,y)\] = [\z \\ (z,y)\] & + \AOT_sem_proj_id y (\ a . \ (x,a)) (\ a . \ (x,a))\\\ +instance proof + text\This is the main proof that allows to derive the definition of n-ary + relation identity. We need to show that our defined projection identity + implies relation identity for relations on pairs of individual terms.\ + fix v and \ \' :: \<'a\'b>\ + have AOT_meta_proj_denotes1: \AOT_model_denotes (Abs_rel (\z. AOT_exe \ (z, \)))\ + if \AOT_model_denotes \\ for \ :: \<'a\'b>\ and \ + using that unfolding AOT_model_denotes_rel.rep_eq + apply (simp add: Abs_rel_inverse AOT_meta_prod_equivI(2) AOT_sem_denotes + that) + by (metis (no_types, lifting) AOT_meta_prod_equivI(2) AOT_model_denotes_prod_def + AOT_model_unary_regular AOT_sem_exe AOT_sem_exe_equiv case_prodD) + { + fix \ :: 'a and \ :: \<'a\'b>\ + assume \_denotes: \AOT_model_denotes \\ + assume \_denotes: \AOT_model_denotes \\ + hence \AOT_exe \ (\, x) = AOT_exe \ (\, y)\ + if \AOT_model_term_equiv x y\ for x y :: 'b + by (simp add: AOT_meta_prod_equivI(1) AOT_sem_exe_equiv that) + moreover have \AOT_model_denotes \\<^sub>1'\\<^sub>n'\ + if \[w \ [\]\ \\<^sub>1'...\\<^sub>n']\ for w \\<^sub>1'\\<^sub>n' + by (metis that AOT_model_denotes_prod_def AOT_sem_exe + AOT_sem_denotes case_prodD) + moreover { + fix x :: 'b + assume x_irregular: \\AOT_model_regular x\ + hence prod_irregular: \\AOT_model_regular (\, x)\ + by (metis (no_types, lifting) AOT_model_irregular_nondenoting + AOT_model_regular_prod_def case_prodD) + hence \(\AOT_model_denotes \ \ \AOT_model_regular x) \ + (AOT_model_denotes \ \ \AOT_model_denotes x)\ + unfolding AOT_model_regular_prod_def by blast + hence x_nonden: \\AOT_model_regular x\ + by (simp add: \_denotes) + have \Rep_rel \ (\, x) = AOT_model_irregular (Rep_rel \) (\, x)\ + using AOT_model_denotes_rel.rep_eq \_denotes prod_irregular by blast + moreover have \AOT_model_irregular (Rep_rel \) (\, x) = + AOT_model_irregular (\z. Rep_rel \ (\, z)) x\ + using \_denotes x_irregular prod_irregular x_nonden + using AOT_model_irregular_prod_generic + apply (induct arbitrary: \ x rule: AOT_model_irregular_prod.induct) + by (auto simp: \_denotes AOT_model_irregular_nondenoting + AOT_model_regular_prod_def AOT_meta_prod_equivI(2) + AOT_model_denotes_rel.rep_eq AOT_model_term_equiv_eps(1) + intro!: AOT_model_irregular_eqI) + ultimately have + \AOT_exe \ (\, x) = AOT_model_irregular (\z. AOT_exe \ (\, z)) x\ + unfolding AOT_sem_exe_denoting[simplified AOT_sem_denotes, OF \_denotes] + by auto + } + ultimately have \AOT_model_denotes (Abs_rel (\z. AOT_exe \ (\, z)))\ + by (simp add: Abs_rel_inverse AOT_model_denotes_rel.rep_eq) + } note AOT_meta_proj_denotes2 = this + { + fix \\<^sub>1'\\<^sub>n' :: 'b and \ :: \<'a\'b>\ + assume \_denotes: \AOT_model_denotes \\ + assume \_denotes: \AOT_model_denotes \\<^sub>1'\\<^sub>n'\ + hence \AOT_exe \ (x, \\<^sub>1'\\<^sub>n') = AOT_exe \ (y, \\<^sub>1'\\<^sub>n')\ + if \AOT_model_term_equiv x y\ for x y :: 'a + by (simp add: AOT_meta_prod_equivI(2) AOT_sem_exe_equiv that) + moreover have \AOT_model_denotes \\ + if \[w \ [\]\ \\<^sub>1'...\\<^sub>n']\ for w \ + by (metis that AOT_model_denotes_prod_def AOT_sem_exe + AOT_sem_denotes case_prodD) + moreover { + fix x :: 'a + assume \\AOT_model_regular x\ + hence \False\ + using AOT_model_unary_regular by blast + hence + \AOT_exe \ (x,\\<^sub>1'\\<^sub>n') = AOT_model_irregular (\z. AOT_exe \ (z,\\<^sub>1'\\<^sub>n')) x\ + unfolding AOT_sem_exe_denoting[simplified AOT_sem_denotes, OF \_denotes] + by auto + } + ultimately have \AOT_model_denotes (Abs_rel (\z. AOT_exe \ (z,\\<^sub>1'\\<^sub>n')))\ + by (simp add: Abs_rel_inverse AOT_model_denotes_rel.rep_eq) + } note AOT_meta_proj_denotes1 = this + { + assume \_denotes: \AOT_model_denotes \\ + assume \'_denotes: \AOT_model_denotes \'\ + have \_proj2_den: \AOT_model_denotes (Abs_rel (\z. Rep_rel \ (\, z)))\ + if \AOT_model_denotes \\ for \ + using that AOT_meta_proj_denotes2[OF \_denotes] + AOT_sem_exe_denoting[simplified AOT_sem_denotes,OF \_denotes] by simp + have \'_proj2_den: \AOT_model_denotes (Abs_rel (\z. Rep_rel \' (\, z)))\ + if \AOT_model_denotes \\ for \ + using that AOT_meta_proj_denotes2[OF \'_denotes] + AOT_sem_exe_denoting[simplified AOT_sem_denotes,OF \'_denotes] by simp + have \_proj1_den: \AOT_model_denotes (Abs_rel (\z. Rep_rel \ (z, \)))\ + if \AOT_model_denotes \\ for \ + using that AOT_meta_proj_denotes1[OF \_denotes] + AOT_sem_exe_denoting[simplified AOT_sem_denotes,OF \_denotes] by simp + have \'_proj1_den: \AOT_model_denotes (Abs_rel (\z. Rep_rel \' (z, \)))\ + if \AOT_model_denotes \\ for \ + using that AOT_meta_proj_denotes1[OF \'_denotes] + AOT_sem_exe_denoting[simplified AOT_sem_denotes,OF \'_denotes] by simp + { + fix \ :: 'a and \\<^sub>1'\\<^sub>n' :: 'b + assume \\ = \'\ + assume \AOT_model_denotes (\,\\<^sub>1'\\<^sub>n')\ + hence \AOT_model_denotes \\ and beta_denotes: \AOT_model_denotes \\<^sub>1'\\<^sub>n'\ + by (auto simp: AOT_model_denotes_prod_def) + have \AOT_model_denotes \[\z [\]z \\<^sub>1'...\\<^sub>n']\\ + by (rule AOT_model_lambda_denotes[THEN iffD2]) + (metis AOT_sem_exe_denoting AOT_meta_prod_equivI(2) + AOT_model_denotes_rel.rep_eq AOT_sem_denotes + AOT_sem_exe_denoting \_denotes) + moreover have \\[\z [\]z \\<^sub>1'...\\<^sub>n']\ = \[\z [\']z \\<^sub>1'...\\<^sub>n']\\ + by (simp add: \\ = \'\) + moreover have \[v \ \AOT_sem_proj_id \\<^sub>1'\\<^sub>n' (\\\<^sub>1'\\<^sub>n'. \[\]\ \\<^sub>1'...\\<^sub>n'\) + (\\\<^sub>1'\\<^sub>n'. \[\']\ \\<^sub>1'...\\<^sub>n'\)\]\ + unfolding \\ = \'\ using beta_denotes + by (rule AOT_sem_proj_id_refl[unfolded AOT_sem_denotes]; + simp add: AOT_sem_denotes AOT_sem_eq AOT_model_lambda_denotes) + (metis AOT_meta_prod_equivI(1) AOT_model_denotes_rel.rep_eq + AOT_sem_exe AOT_sem_exe_denoting \'_denotes) + ultimately have \[v \ \AOT_sem_proj_id (\,\\<^sub>1'\\<^sub>n') (\ \\<^sub>1\\<^sub>n . \[\]\\<^sub>1...\\<^sub>n\) + (\ \\<^sub>1\\<^sub>n . \[\']\\<^sub>1...\\<^sub>n\)\]\ + unfolding AOT_sem_proj_id_prod_def + by (simp add: AOT_sem_denotes AOT_sem_conj AOT_sem_eq) + } + moreover { + assume \\ \ . AOT_model_denotes \ \ + [v \ \AOT_sem_proj_id \ (\ \\<^sub>1\\<^sub>n . \[\]\\<^sub>1...\\<^sub>n\) (\ \\<^sub>1\\<^sub>n . \[\']\\<^sub>1...\\<^sub>n\)\]\ + hence 0: \AOT_model_denotes \ \ AOT_model_denotes \\<^sub>1'\\<^sub>n' \ + AOT_model_denotes \[\z [\]z \\<^sub>1'...\\<^sub>n']\ \ + AOT_model_denotes \[\z [\']z \\<^sub>1'...\\<^sub>n']\ \ + \[\z [\]z \\<^sub>1'...\\<^sub>n']\ = \[\z [\']z \\<^sub>1'...\\<^sub>n']\ \ + [v \ \AOT_sem_proj_id \\<^sub>1'\\<^sub>n' (\\\<^sub>1\\<^sub>n. \[\]\ \\<^sub>1...\\<^sub>n\) + (\\\<^sub>1\\<^sub>n. \[\']\ \\<^sub>1...\\<^sub>n\)\]\ for \ \\<^sub>1'\\<^sub>n' + unfolding AOT_sem_proj_id_prod_def + by (auto simp: AOT_sem_denotes AOT_sem_conj AOT_sem_eq + AOT_model_denotes_prod_def) + obtain \den :: 'a and \den :: 'b where + \den: \AOT_model_denotes \den\ and \den: \AOT_model_denotes \den\ + using AOT_model_denoting_ex by metis + { + fix \ :: 'a + assume \denotes: \AOT_model_denotes \\ + have 1: \[v \ \AOT_sem_proj_id \\<^sub>1'\\<^sub>n' (\\\<^sub>1'\\<^sub>n'. \[\]\ \\<^sub>1'...\\<^sub>n'\) + (\\\<^sub>1'\\<^sub>n'. \[\']\ \\<^sub>1'...\\<^sub>n'\)\]\ + if \AOT_model_denotes \\<^sub>1'\\<^sub>n'\ for \\<^sub>1'\\<^sub>n' + using that 0 using \denotes by blast + hence \[v \ \AOT_sem_proj_id \ (\z. Rep_rel \ (\, z)) + (\z. Rep_rel \' (\, z))\]\ + if \AOT_model_denotes \\ for \ + using that + unfolding AOT_sem_exe_denoting[simplified AOT_sem_denotes, OF \_denotes] + AOT_sem_exe_denoting[simplified AOT_sem_denotes, OF \'_denotes] + by blast + hence \Abs_rel (\z. Rep_rel \ (\, z)) = Abs_rel (\z. Rep_rel \' (\, z))\ + using AOT_sem_proj_id_prop[of v \Abs_rel (\z. Rep_rel \ (\, z))\ + \Abs_rel (\z. Rep_rel \' (\, z))\, + simplified AOT_sem_eq AOT_sem_conj AOT_sem_forall + AOT_sem_denotes, THEN iffD2] + \_proj2_den[OF \denotes] \'_proj2_den[OF \denotes] + unfolding AOT_sem_exe_denoting[simplified AOT_sem_denotes, OF \_denotes] + AOT_sem_exe_denoting[simplified AOT_sem_denotes, + OF \_proj2_den[OF \denotes]] + AOT_sem_exe_denoting[simplified AOT_sem_denotes, + OF \'_proj2_den[OF \denotes]] + by (metis Abs_rel_inverse UNIV_I) + hence "Rep_rel \ (\,\) = Rep_rel \' (\,\)" for \ + by (simp add: Abs_rel_inject[simplified]) meson + } note \denotes = this + { + fix \\<^sub>1'\\<^sub>n' :: 'b + assume \den: \AOT_model_denotes \\<^sub>1'\\<^sub>n'\ + have 1: \\[\z [\]z \\<^sub>1'...\\<^sub>n']\ = \[\z [\']z \\<^sub>1'...\\<^sub>n']\\ + using 0 \den AOT_model_denoting_ex by blast + hence \Abs_rel (\z. Rep_rel \ (z, \\<^sub>1'\\<^sub>n')) = + Abs_rel (\z. Rep_rel \' (z, \\<^sub>1'\\<^sub>n'))\ (is \?a = ?b\) + apply (safe intro!: AOT_sem_proj_id_prop[of v \?a\ \?b\, + simplified AOT_sem_eq AOT_sem_conj AOT_sem_forall + AOT_sem_denotes, THEN iffD2, THEN conjunct2, THEN conjunct2] + \_proj1_den[OF \den] \'_proj1_den[OF \den]) + unfolding AOT_sem_exe_denoting[simplified AOT_sem_denotes, OF \_denotes] + AOT_sem_exe_denoting[simplified AOT_sem_denotes, OF \'_denotes] + AOT_sem_exe_denoting[simplified AOT_sem_denotes, + OF \_proj1_den[OF \den]] + AOT_sem_exe_denoting[simplified AOT_sem_denotes, + OF \'_proj1_den[OF \den]] + by (subst (0 1) Abs_rel_inverse; simp?) + (metis (no_types, lifting) AOT_model_denotes_rel.abs_eq + AOT_model_lambda_denotes AOT_sem_denotes AOT_sem_eq + AOT_sem_unary_proj_id \_proj1_den[OF \den]) + hence \Rep_rel \ (x,\\<^sub>1'\\<^sub>n') = Rep_rel \' (x,\\<^sub>1'\\<^sub>n')\ for x + by (simp add: Abs_rel_inject) + metis + } note \denotes = this + { + fix \ :: 'a and \ :: 'b + assume \AOT_model_regular (\, \)\ + moreover { + assume \AOT_model_denotes \ \ AOT_model_regular \\ + hence \Rep_rel \ (\,\) = Rep_rel \' (\,\)\ + using \denotes by presburger + } + moreover { + assume \\AOT_model_denotes \ \ AOT_model_denotes \\ + hence \Rep_rel \ (\,\) = Rep_rel \' (\,\)\ + by (simp add: \denotes) + } + ultimately have \Rep_rel \ (\,\) = Rep_rel \' (\,\)\ + by (metis (no_types, lifting) AOT_model_regular_prod_def case_prodD) + } + hence \Rep_rel \ = Rep_rel \'\ + using \_denotes[unfolded AOT_model_denotes_rel.rep_eq, + THEN conjunct2, THEN conjunct2, THEN spec, THEN mp] + using \'_denotes[unfolded AOT_model_denotes_rel.rep_eq, + THEN conjunct2, THEN conjunct2, THEN spec, THEN mp] + using AOT_model_irregular_eqI[of \Rep_rel \\ \Rep_rel \'\ \(_,_)\] + using AOT_model_irregular_nondenoting by fastforce + hence \\ = \'\ + by (rule Rep_rel_inject[THEN iffD1]) + } + ultimately have \\ = \' = (\ \ . AOT_model_denotes \ \ + [v \ \AOT_sem_proj_id \ (\ \\<^sub>1\\<^sub>n . \[\]\\<^sub>1...\\<^sub>n\) + (\ \\<^sub>1\\<^sub>n . \[\']\\<^sub>1...\\<^sub>n\)\])\ + by auto + } + thus \[v \ \ = \'] = [v \ \\ & \'\ & + \\ (\AOT_sem_proj_id \ (\ \\<^sub>1\\<^sub>n . \[\]\\<^sub>1...\\<^sub>n\) (\ \\<^sub>1\\<^sub>n . \[\']\\<^sub>1...\\<^sub>n\)\)]\ + by (auto simp: AOT_sem_eq AOT_sem_denotes AOT_sem_forall AOT_sem_conj) +next + fix v and \ :: \'a\'b\\\ and \ :: \'a\'b\ + assume \[v \ \\]\ + moreover assume \[v \ [\z\<^sub>1...z\<^sub>n \\ z\<^sub>1z\<^sub>n\] = [\z\<^sub>1...z\<^sub>n \\ z\<^sub>1z\<^sub>n\]]\ + ultimately show \[v \ \AOT_sem_proj_id \ \ \\]\ + unfolding AOT_sem_proj_id_prod_def + using AOT_sem_proj_id_refl[of v "snd \" "\b. \ (fst \, b)"] + by (auto simp: AOT_sem_eq AOT_sem_conj AOT_sem_denotes + AOT_model_denotes_prod_def AOT_model_lambda_denotes + AOT_meta_prod_equivI) +qed +end + +text\Sanity-check to verify that n-ary relation identity follows.\ +lemma \[v \ \ = \'] = [v \ \\ & \'\ & \x\y([\z [\]z y] = [\z [\']z y] & + [\z [\]x z] = [\z [\']x z])]\ + for \ :: \<\\\>\ + by (auto simp: AOT_sem_proj_id_prop[of v \ \'] AOT_sem_proj_id_prod_def + AOT_sem_conj AOT_sem_denotes AOT_sem_forall AOT_sem_unary_proj_id + AOT_model_denotes_prod_def) +lemma \[v \ \ = \'] = [v \ \\ & \'\ & \x\<^sub>1\x\<^sub>2\x\<^sub>3 ( + [\z [\]z x\<^sub>2 x\<^sub>3] = [\z [\']z x\<^sub>2 x\<^sub>3] & + [\z [\]x\<^sub>1 z x\<^sub>3] = [\z [\']x\<^sub>1 z x\<^sub>3] & + [\z [\]x\<^sub>1 x\<^sub>2 z] = [\z [\']x\<^sub>1 x\<^sub>2 z])]\ + for \ :: \<\\\\\>\ + by (auto simp: AOT_sem_proj_id_prop[of v \ \'] AOT_sem_proj_id_prod_def + AOT_sem_conj AOT_sem_denotes AOT_sem_forall AOT_sem_unary_proj_id + AOT_model_denotes_prod_def) +lemma \[v \ \ = \'] = [v \ \\ & \'\ & \x\<^sub>1\x\<^sub>2\x\<^sub>3\x\<^sub>4 ( + [\z [\]z x\<^sub>2 x\<^sub>3 x\<^sub>4] = [\z [\']z x\<^sub>2 x\<^sub>3 x\<^sub>4] & + [\z [\]x\<^sub>1 z x\<^sub>3 x\<^sub>4] = [\z [\']x\<^sub>1 z x\<^sub>3 x\<^sub>4] & + [\z [\]x\<^sub>1 x\<^sub>2 z x\<^sub>4] = [\z [\']x\<^sub>1 x\<^sub>2 z x\<^sub>4] & + [\z [\]x\<^sub>1 x\<^sub>2 x\<^sub>3 z] = [\z [\']x\<^sub>1 x\<^sub>2 x\<^sub>3 z])]\ + for \ :: \<\\\\\\\>\ + by (auto simp: AOT_sem_proj_id_prop[of v \ \'] AOT_sem_proj_id_prod_def + AOT_sem_conj AOT_sem_denotes AOT_sem_forall AOT_sem_unary_proj_id + AOT_model_denotes_prod_def) + +text\n-ary Encoding is constructed using a similar mechanism as n-ary relation + identity using an auxiliary notion of projection-encoding.\ +class AOT_Enc = + fixes AOT_enc :: \'a \ <'a::AOT_IndividualTerm> \ \\ + and AOT_proj_enc :: \'a \ ('a \ \) \ \\ + assumes AOT_sem_enc_denotes: + \[v \ \AOT_enc \\<^sub>1\\<^sub>n \\] \ [v \ \\<^sub>1...\\<^sub>n\] \ [v \ \\]\ + assumes AOT_sem_enc_proj_enc: + \[v \ \AOT_enc \\<^sub>1\\<^sub>n \\] = + [v \ \\ & \AOT_proj_enc \\<^sub>1\\<^sub>n (\ \\<^sub>1\\<^sub>n. \[\]\\<^sub>1...\\<^sub>n\)\]\ + assumes AOT_sem_proj_enc_denotes: + \[v \ \AOT_proj_enc \\<^sub>1\\<^sub>n \\] \ [v \ \\<^sub>1...\\<^sub>n\]\ + assumes AOT_sem_enc_nec: + \[v \ \AOT_enc \\<^sub>1\\<^sub>n \\] \ [w \ \AOT_enc \\<^sub>1\\<^sub>n \\]\ + assumes AOT_sem_proj_enc_nec: + \[v \ \AOT_proj_enc \\<^sub>1\\<^sub>n \\] \ [w \ \AOT_proj_enc \\<^sub>1\\<^sub>n \\]\ + assumes AOT_sem_universal_encoder: + \\ \\<^sub>1\\<^sub>n. [v \ \\<^sub>1...\\<^sub>n\] \ (\ \ . [v \ \\] \ [v \ \AOT_enc \\<^sub>1\\<^sub>n \\]) \ + (\ \ . [v \ [\z\<^sub>1...z\<^sub>n \{z\<^sub>1...z\<^sub>n}]\] \ [v \ \AOT_proj_enc \\<^sub>1\\<^sub>n \\])\ + +AOT_syntax_print_translations + "_AOT_enc (_AOT_individual_term \) (_AOT_relation_term \)" <= "CONST AOT_enc \ \" + +context AOT_meta_syntax +begin +notation AOT_enc ("\<^bold>\_,_\<^bold>\") +end +context AOT_no_meta_syntax +begin +no_notation AOT_enc ("\<^bold>\_,_\<^bold>\") +end + +text\Unary encoding additionally has to satisfy the axioms of unary encoding and + the definition of property identity.\ +class AOT_UnaryEnc = AOT_UnaryIndividualTerm + + assumes AOT_sem_enc_eq: \[v \ \\ & \'\ & \\\ (\[\] \ \[\']) \ \ = \']\ + and AOT_sem_A_objects: \[v \ \x (\\[E!]x & \F (x[F] \ \{F}))]\ + and AOT_sem_unary_proj_enc: \AOT_proj_enc x \ = AOT_enc x \[\z \{z}]\\ + and AOT_sem_nocoder: \[v \ [E!]\] \ \[w \ \AOT_enc \ \\]\ + and AOT_sem_ind_eq: \([v \ \\] \ [v \ \'\] \ \ = (\')) = + (([v \ [\x \[E!]x]\] \ + [v \ [\x \[E!]x]\'] \ + (\ v \ . [v \ \\] \ [v \ [\]\] = [v \ [\]\'])) + \ ([v \ [\x \\[E!]x]\] \ + [v \ [\x \\[E!]x]\'] \ + (\ v \ . [v \ \\] \ [v \ \[\]] = [v \ \'[\]])))\ + + (* only extended models *) + and AOT_sem_enc_indistinguishable_all: + \AOT_ExtendedModel \ + [v \ [\x \\[E!]x]\] \ + [v \ [\x \\[E!]x]\'] \ + (\ \' . [v \ \'\] \ (\ w . [w \ [\']\] = [w \ [\']\'])) \ + [v \ \\] \ + (\ \' . [v \ \'\] \ (\ \\<^sub>0 . [v \ [\x \[E!]x]\\<^sub>0] \ + (\ w . [w \ [\']\\<^sub>0] = [w \ [\]\\<^sub>0])) \ [v \ \[\']]) \ + (\ \' . [v \ \'\] \ (\ \\<^sub>0 . [v \ [\x \[E!]x]\\<^sub>0] \ + (\ w . [w \ [\']\\<^sub>0] = [w \ [\]\\<^sub>0])) \ [v \ \'[\']])\ + and AOT_sem_enc_indistinguishable_ex: + \AOT_ExtendedModel \ + [v \ [\x \\[E!]x]\] \ + [v \ [\x \\[E!]x]\'] \ + (\ \' . [v \ \'\] \ (\ w . [w \ [\']\] = [w \ [\']\'])) \ + [v \ \\] \ + \ \' . [v \ \'\] \ [v \ \[\']] \ + (\ \\<^sub>0 . [v \ [\x \[E!]x]\\<^sub>0] \ + (\ w . [w \ [\']\\<^sub>0] = [w \ [\]\\<^sub>0])) \ + \ \' . [v \ \'\] \ [v \ \'[\']] \ + (\ \\<^sub>0 . [v \ [\x \[E!]x]\\<^sub>0] \ + (\ w . [w \ [\']\\<^sub>0] = [w \ [\]\\<^sub>0]))\ + +text\We specify encoding to align with the model-construction of encoding.\ +consts AOT_sem_enc_\ :: \\ \ <\> \ \\ +specification(AOT_sem_enc_\) + AOT_sem_enc_\: + \[v \ \AOT_sem_enc_\ \ \\] = + (AOT_model_denotes \ \ AOT_model_denotes \ \ AOT_model_enc \ \)\ + by (rule exI[where x=\\ \ \ . \\<^sub>\ w . AOT_model_denotes \ \ AOT_model_denotes \ \ + AOT_model_enc \ \\]) + (simp add: AOT_model_proposition_choice_simp AOT_model_enc_\_def \.case_eq_if) + +text\We show that @{typ \} satisfies the generic properties of n-ary encoding.\ +instantiation \ :: AOT_Enc +begin +definition AOT_enc_\ :: \\ \ <\> \ \\ where + \AOT_enc_\ \ AOT_sem_enc_\\ +definition AOT_proj_enc_\ :: \\ \ (\ \ \) \ \\ where + \AOT_proj_enc_\ \ \ \ \ . AOT_enc \ \[\z \\ z\]\\ +lemma AOT_enc_\_meta: + \[v \ \[\]] = (AOT_model_denotes \ \ AOT_model_denotes \ \ AOT_model_enc \ \)\ + for \::\ + using AOT_sem_enc_\ unfolding AOT_enc_\_def by auto +instance proof + fix v and \ :: \ and \ + show \[v \ \AOT_enc \ \\] \ [v \ \\] \ [v \ \\]\ + unfolding AOT_sem_denotes + using AOT_enc_\_meta by blast +next + fix v and \ :: \ and \ + show \[v \ \[\]] = [v \ \\ & \AOT_proj_enc \ (\ \'. \[\]\'\)\]\ + proof + assume enc: \[v \ \[\]]\ + hence \_denotes: \AOT_model_denotes \\ + by (simp add: AOT_enc_\_meta) + hence \_eta_denotes: \AOT_model_denotes \[\z [\]z]\\ + using AOT_sem_denotes AOT_sem_eq AOT_sem_lambda_eta by metis + show \[v \ \\ & \AOT_proj_enc \ (\ \. \[\]\\)\]\ + using AOT_sem_lambda_eta[simplified AOT_sem_denotes AOT_sem_eq, OF \_denotes] + using \_eta_denotes \_denotes + by (simp add: AOT_sem_conj AOT_sem_denotes AOT_proj_enc_\_def enc) + next + assume \[v \ \\ & \AOT_proj_enc \ (\ \. \[\]\\)\]\ + hence \_denotes: "AOT_model_denotes \" and eta_enc: "[v \ \[\z [\]z]]" + by (auto simp: AOT_sem_conj AOT_sem_denotes AOT_proj_enc_\_def) + thus \[v \ \[\]]\ + using AOT_sem_lambda_eta[simplified AOT_sem_denotes AOT_sem_eq, OF \_denotes] + by auto + qed +next + show \[v \ \AOT_proj_enc \ \\] \ [v \ \\]\ for v and \ :: \ and \ + by (simp add: AOT_enc_\_meta AOT_sem_denotes AOT_proj_enc_\_def) +next + fix v w and \ :: \ and \ + assume \[v \ \[\]]\ + thus \[w \ \[\]]\ + by (simp add: AOT_enc_\_meta) +next + fix v w and \ :: \ and \ + assume \[v \ \AOT_proj_enc \ \\]\ + thus \[w \ \AOT_proj_enc \ \\]\ + by (simp add: AOT_enc_\_meta AOT_proj_enc_\_def) +next + show \\\::\. [v \ \\] \ (\ \ . [v \ \\] \ [v \ \[\]]) \ + (\ \ . [v \ [\z \{z}]\] \ [v \ \AOT_proj_enc \ \\])\ for v + by (rule exI[where x=\\\ UNIV\]) + (simp add: AOT_sem_denotes AOT_enc_\_meta AOT_model_enc_\_def + AOT_model_denotes_\_def AOT_proj_enc_\_def) +qed +end + +text\We show that @{typ \} satisfies the properties of unary encoding.\ +instantiation \ :: AOT_UnaryEnc +begin +instance proof + fix v and \ \' :: \<\>\ + show \[v \ \\ & \'\ & \\\ (\[\] \ \[\']) \ \ = \']\ + apply (simp add: AOT_sem_forall AOT_sem_eq AOT_sem_imp AOT_sem_equiv + AOT_enc_\_meta AOT_sem_conj AOT_sem_denotes AOT_sem_box) + using AOT_meta_A_objects_\ by fastforce +next + fix v and \:: \<\> \ \\ + show \[v \ \x (\\[E!]x & \F (x[F] \ \{F}))]\ + using AOT_model_A_objects[of "\ \ . [v \ \{\}]"] + by (auto simp: AOT_sem_denotes AOT_sem_exists AOT_sem_conj AOT_sem_not + AOT_sem_dia AOT_sem_concrete AOT_enc_\_meta AOT_sem_equiv + AOT_sem_forall) +next + show \AOT_proj_enc x \ = AOT_enc x (AOT_lambda \)\ for x :: \ and \ + by (simp add: AOT_proj_enc_\_def) +next + show \[v \ [E!]\] \ \ [w \ \[\]]\ for v w and \ :: \ and \ + by (simp add: AOT_enc_\_meta AOT_sem_concrete AOT_model_nocoder) +next + fix v and \ \' :: \ + show \([v \ \\] \ [v \ \'\] \ \ = \') = + (([v \ [\x \[E!]x]\] \ + [v \ [\x \[E!]x]\'] \ + (\ v \ . [v \ \\] \ [v \ [\]\] = [v \ [\]\'])) + \ ([v \ [\x \\[E!]x]\] \ + [v \ [\x \\[E!]x]\'] \ + (\ v \ . [v \ \\] \ [v \ \[\]] = [v \ \'[\]])))\ + (is \?lhs = (?ordeq \ ?abseq)\) + proof - + { + assume 0: \[v \ \\] \ [v \ \'\] \ \ = \'\ + { + assume \is_\\ \'\ + hence \[v \ [\x \[E!]x]\']\ + apply (subst AOT_sem_lambda_beta[OF AOT_sem_ordinary_def_denotes, of v \']) + apply (simp add: "0") + apply (simp add: AOT_sem_dia) + using AOT_sem_concrete AOT_model_\_concrete_in_some_world is_\\_def by force + hence \?ordeq\ unfolding 0[THEN conjunct2, THEN conjunct2] by auto + } + moreover { + assume \is_\\ \'\ + hence \[v \ [\x \\[E!]x]\']\ + apply (subst AOT_sem_lambda_beta[OF AOT_sem_abstract_def_denotes, of v \']) + apply (simp add: "0") + apply (simp add: AOT_sem_not AOT_sem_dia) + using AOT_sem_concrete is_\\_def by force + hence \?abseq\ unfolding 0[THEN conjunct2, THEN conjunct2] by auto + } + ultimately have \?ordeq \ ?abseq\ + by (meson "0" AOT_sem_denotes AOT_model_denotes_\_def \.exhaust_disc) + } + moreover { + assume ordeq: \?ordeq\ + hence \_denotes: \[v \ \\]\ and \'_denotes: \[v \ \'\]\ + by (simp add: AOT_sem_denotes AOT_sem_exe)+ + hence \is_\\ \\ and \is_\\ \'\ + by (metis AOT_model_concrete_\.simps(2) AOT_model_denotes_\_def + AOT_sem_concrete AOT_sem_denotes AOT_sem_dia AOT_sem_lambda_beta + AOT_sem_ordinary_def_denotes \.collapse(2) \.exhaust_disc ordeq)+ + have denotes: \[v \ [\z \\\<^sub>\ w . \\ z = \\ \\]\]\ + unfolding AOT_sem_denotes AOT_model_lambda_denotes + by (simp add: AOT_model_term_equiv_\_def) + hence "[v \ [\z \\\<^sub>\ w . \\ z = \\ \\]\] = [v \ [\z \\\<^sub>\ w . \\ z = \\ \\]\']" + using ordeq by (simp add: AOT_sem_denotes) + hence \[v \ \\\\] \ [v \ \\'\\] \ \ = \'\ + unfolding AOT_sem_lambda_beta[OF denotes, OF \_denotes] + AOT_sem_lambda_beta[OF denotes, OF \'_denotes] + using \'_denotes \is_\\ \'\ \is_\\ \\ is_\\_def + AOT_model_proposition_choice_simp by force + } + moreover { + assume 0: \?abseq\ + hence \_denotes: \[v \ \\]\ and \'_denotes: \[v \ \'\]\ + by (simp add: AOT_sem_denotes AOT_sem_exe)+ + hence \\is_\\ \\ and \\is_\\ \'\ + by (metis AOT_model_\_concrete_in_some_world AOT_model_concrete_\.simps(1) + AOT_sem_concrete AOT_sem_dia AOT_sem_exe AOT_sem_lambda_beta + AOT_sem_not \.collapse(1) 0)+ + hence \is_\\ \\ and \is_\\ \'\ + by (meson AOT_sem_denotes AOT_model_denotes_\_def \.exhaust_disc + \_denotes \'_denotes)+ + then obtain x y where \_def: \\ = \\ x\ and \'_def: \\' = \\ y\ + using is_\\_def by auto + { + fix r + assume \r \ x\ + hence \[v \ \[\urrel_to_rel r\]]\ + unfolding \_def + unfolding AOT_enc_\_meta + unfolding AOT_model_enc_\_def + apply (simp add: AOT_model_denotes_\_def) + by (metis (mono_tags) AOT_rel_equiv_def Quotient_def urrel_quotient) + hence \[v \ \'[\urrel_to_rel r\]]\ + using AOT_enc_\_meta 0 by (metis AOT_sem_enc_denotes) + hence \r \ y\ + unfolding \'_def + unfolding AOT_enc_\_meta + unfolding AOT_model_enc_\_def + apply (simp add: AOT_model_denotes_\_def) + using Quotient_abs_rep urrel_quotient by fastforce + } + moreover { + fix r + assume \r \ y\ + hence \[v \ \'[\urrel_to_rel r\]]\ + unfolding \'_def + unfolding AOT_enc_\_meta + unfolding AOT_model_enc_\_def + apply (simp add: AOT_model_denotes_\_def) + by (metis (mono_tags) AOT_rel_equiv_def Quotient_def urrel_quotient) + hence \[v \ \[\urrel_to_rel r\]]\ + using AOT_enc_\_meta 0 by (metis AOT_sem_enc_denotes) + hence \r \ x\ + unfolding \_def + unfolding AOT_enc_\_meta + unfolding AOT_model_enc_\_def + apply (simp add: AOT_model_denotes_\_def) + using Quotient_abs_rep urrel_quotient by fastforce + } + ultimately have "x = y" by blast + hence \[v \ \\] \ [v \ \'\] \ \ = \'\ + using \'_def \'_denotes \_def by blast + } + ultimately show ?thesis + unfolding AOT_sem_denotes + by auto + qed +(* Only extended model *) +next + fix v and \ \' :: \ and \ \' :: \<\>\ + assume ext: \AOT_ExtendedModel\ + assume \[v \ [\x \\[E!]x]\]\ + hence \is_\\ \\ + by (metis AOT_model_\_concrete_in_some_world AOT_model_concrete_\.simps(1) + AOT_model_denotes_\_def AOT_sem_concrete AOT_sem_denotes AOT_sem_dia + AOT_sem_exe AOT_sem_lambda_beta AOT_sem_not \.collapse(1) \.exhaust_disc) + hence \_abs: \\(\ w . AOT_model_concrete w \)\ + using is_\\_def by fastforce + have \_den: \AOT_model_denotes \\ + by (simp add: AOT_model_denotes_\_def \.distinct_disc(5) \is_\\ \\) + assume \[v \ [\x \\[E!]x]\']\ + hence \is_\\ \'\ + by (metis AOT_model_\_concrete_in_some_world AOT_model_concrete_\.simps(1) + AOT_model_denotes_\_def AOT_sem_concrete AOT_sem_denotes AOT_sem_dia + AOT_sem_exe AOT_sem_lambda_beta AOT_sem_not \.collapse(1) + \.exhaust_disc) + hence \'_abs: \\(\ w . AOT_model_concrete w \')\ + using is_\\_def by fastforce + have \'_den: \AOT_model_denotes \'\ + by (meson AOT_model_denotes_\_def \.distinct_disc(6) \is_\\ \'\) + assume \[v \ \'\] \ [w \ [\']\] = [w \ [\']\']\ for \' w + hence indist: \[v \ \Rep_rel \' \\] = [v \ \Rep_rel \' \'\]\ + if \AOT_model_denotes \'\ for \' v + by (metis AOT_sem_denotes AOT_sem_exe \'_den \_den that) + assume \_enc_cond: \[v \ \'\] \ + (\ \\<^sub>0 w. [v \ [\x \[E!]x]\\<^sub>0] \ + [w \ [\']\\<^sub>0] = [w \ [\]\\<^sub>0]) \ + [v \ \[\']]\ for \' + assume \_den': \[v \ \\]\ + hence \_den: \AOT_model_denotes \\ + using AOT_sem_denotes by blast + { + fix \' :: \<\>\ + assume \'_den: \AOT_model_denotes \'\ + hence \'_den': \[v \ \'\]\ + by (simp add: AOT_sem_denotes) + assume 1: \\w. AOT_model_concrete w x \ + [v \ \Rep_rel \' x\] = [v \ \Rep_rel \ x\]\ for v x + { + fix \\<^sub>0 :: \ and w + assume \[v \ [\x \[E!]x]\\<^sub>0]\ + hence \is_\\ \\<^sub>0\ + by (smt (z3) AOT_model_concrete_\.simps(2) AOT_model_denotes_\_def + AOT_sem_concrete AOT_sem_denotes AOT_sem_dia AOT_sem_exe + AOT_sem_lambda_beta \.exhaust_disc is_\\_def) + then obtain x where x_prop: \\\<^sub>0 = \\ x\ + using is_\\_def by blast + have \\w . AOT_model_concrete w (\\ x)\ + by (simp add: AOT_model_\_concrete_in_some_world) + hence \[v \ \Rep_rel \' (\\ x)\] = [v \ \Rep_rel \ (\\ x)\]\ for v + using 1 by blast + hence \[w \ [\']\\<^sub>0] = [w \ [\]\\<^sub>0]\ unfolding x_prop + by (simp add: AOT_sem_exe AOT_sem_denotes AOT_model_denotes_\_def + \'_den \_den) + } note 2 = this + have \[v \ \[\']]\ + using \_enc_cond[OF \'_den', OF 2] + by metis + hence \AOT_model_enc \ \'\ + using AOT_enc_\_meta by blast + } note \_enc_cond = this + hence \AOT_model_denotes \' \ + (\v x. \w. AOT_model_concrete w x \ + [v \ \Rep_rel \' x\] = [v \ \Rep_rel \ x\]) \ + AOT_model_enc \ \'\ for \' + by blast + assume \'_den': \[v \ \'\]\ + hence \'_den: \AOT_model_denotes \'\ + using AOT_sem_denotes by blast + assume ord_indist: \[v \ [\x \[E!]x]\\<^sub>0] \ + [w \ [\']\\<^sub>0] = [w \ [\]\\<^sub>0]\ for \\<^sub>0 w + { + fix w and \\<^sub>0 :: \ + assume 0: \\w. AOT_model_concrete w \\<^sub>0\ + hence \[v \ [\x \[E!]x]\\<^sub>0]\ + using AOT_model_concrete_denotes AOT_sem_concrete AOT_sem_denotes AOT_sem_dia + AOT_sem_lambda_beta AOT_sem_ordinary_def_denotes by blast + hence \[w \ [\']\\<^sub>0] = [w \ [\]\\<^sub>0]\ + using ord_indist by metis + hence \[w \ \Rep_rel \' \\<^sub>0\] = [w \ \Rep_rel \ \\<^sub>0\]\ + by (metis AOT_model_concrete_denotes AOT_sem_denotes AOT_sem_exe \'_den \_den 0) + } note ord_indist = this + have \AOT_model_enc \' \'\ + using AOT_model_enc_indistinguishable_all + [OF ext, OF \_den, OF \_abs, OF \'_den, OF \'_abs, OF \_den] + indist \_enc_cond \'_den ord_indist by blast + thus \[v \ \'[\']]\ + using AOT_enc_\_meta \'_den \'_den by blast +next + fix v and \ \' :: \ and \ \' :: \<\>\ + assume ext: \AOT_ExtendedModel\ + assume \[v \ [\x \\[E!]x]\]\ + hence \is_\\ \\ + by (metis AOT_model_\_concrete_in_some_world AOT_model_concrete_\.simps(1) + AOT_model_denotes_\_def AOT_sem_concrete AOT_sem_denotes AOT_sem_dia + AOT_sem_exe AOT_sem_lambda_beta AOT_sem_not \.collapse(1) + \.exhaust_disc) + hence \_abs: \\(\ w . AOT_model_concrete w \)\ + using is_\\_def by fastforce + have \_den: \AOT_model_denotes \\ + by (simp add: AOT_model_denotes_\_def \.distinct_disc(5) \is_\\ \\) + assume \[v \ [\x \\[E!]x]\']\ + hence \is_\\ \'\ + by (metis AOT_model_\_concrete_in_some_world AOT_model_concrete_\.simps(1) + AOT_model_denotes_\_def AOT_sem_concrete AOT_sem_denotes AOT_sem_dia + AOT_sem_exe AOT_sem_lambda_beta AOT_sem_not \.collapse(1) + \.exhaust_disc) + hence \'_abs: \\(\ w . AOT_model_concrete w \')\ + using is_\\_def by fastforce + have \'_den: \AOT_model_denotes \'\ + by (meson AOT_model_denotes_\_def \.distinct_disc(6) \is_\\ \'\) + assume \[v \ \'\] \ [w \ [\']\] = [w \ [\']\']\ for \' w + hence indist: \[v \ \Rep_rel \' \\] = [v \ \Rep_rel \' \'\]\ + if \AOT_model_denotes \'\ for \' v + by (metis AOT_sem_denotes AOT_sem_exe \'_den \_den that) + assume \_den': \[v \ \\]\ + hence \_den: \AOT_model_denotes \\ + using AOT_sem_denotes by blast + assume \\\'. [v \ \'\] \ [v \ \[\']] \ + (\\\<^sub>0. [v \ [\x \[E!]x]\\<^sub>0] \ + (\w. [w \ [\']\\<^sub>0] = [w \ [\]\\<^sub>0]))\ + then obtain \' where + \'_den: \[v \ \'\]\ and + \'_enc: \[v \ \[\']]\ and + \'_prop: \\\\<^sub>0. [v \ [\x \[E!]x]\\<^sub>0] \ + (\w. [w \ [\']\\<^sub>0] = [w \ [\]\\<^sub>0])\ + by blast + have \AOT_model_denotes \'\ + using AOT_enc_\_meta \'_enc by force + moreover have \AOT_model_enc \ \'\ + using AOT_enc_\_meta \'_enc by blast + moreover have \(\w. AOT_model_concrete w \\<^sub>0) \ + [v \ \Rep_rel \' \\<^sub>0\] = [v \ \Rep_rel \ \\<^sub>0\]\ for \\<^sub>0 v + proof - + assume 0: \\w. AOT_model_concrete w \\<^sub>0\ + hence \[v \ [\x \[E!]x]\\<^sub>0]\ for v + using AOT_model_concrete_denotes AOT_sem_concrete AOT_sem_denotes AOT_sem_dia + AOT_sem_lambda_beta AOT_sem_ordinary_def_denotes by blast + hence \\w. [w \ [\']\\<^sub>0] = [w \ [\]\\<^sub>0]\ using \'_prop by blast + thus \[v \ \Rep_rel \' \\<^sub>0\] = [v \ \Rep_rel \ \\<^sub>0\]\ + by (meson "0" AOT_model_concrete_denotes AOT_sem_denotes AOT_sem_exe \_den + calculation(1)) + qed + ultimately have \\\'. AOT_model_denotes \' \ AOT_model_enc \ \' \ + (\v x. (\w. AOT_model_concrete w x) \ + [v \ \Rep_rel \' x\] = [v \ \Rep_rel \ x\])\ + by blast + hence \\\'. AOT_model_denotes \' \ AOT_model_enc \' \' \ + (\v x. (\w. AOT_model_concrete w x) \ + [v \ \Rep_rel \' x\] = [v \ \Rep_rel \ x\])\ + using AOT_model_enc_indistinguishable_ex + [OF ext, OF \_den, OF \_abs, OF \'_den, OF \'_abs, OF \_den] + indist by blast + then obtain \'' where + \''_den: \AOT_model_denotes \''\ + and \''_enc: \AOT_model_enc \' \''\ + and \''_prop: \(\w. AOT_model_concrete w x) \ + [v \ \Rep_rel \'' x\] = [v \ \Rep_rel \ x\]\ for v x + by blast + have \[v \ \''\]\ + by (simp add: AOT_sem_denotes \''_den) + moreover have \[v \ \'[\'']]\ + by (simp add: AOT_enc_\_meta \''_den \''_enc \'_den) + moreover have \[v \ [\x \[E!]x]\\<^sub>0] \ + (\w. [w \ [\'']\\<^sub>0] = [w \ [\]\\<^sub>0])\ for \\<^sub>0 + proof - + assume \[v \ [\x \[E!]x]\\<^sub>0]\ + hence \\w. AOT_model_concrete w \\<^sub>0\ + by (metis AOT_sem_concrete AOT_sem_dia AOT_sem_exe AOT_sem_lambda_beta) + thus \\w. [w \ [\'']\\<^sub>0] = [w \ [\]\\<^sub>0]\ + using \''_prop + by (metis AOT_sem_denotes AOT_sem_exe \''_den \_den) + qed + ultimately show \\\'. [v \ \'\] \ [v \ \'[\']] \ + (\\\<^sub>0. [v \ [\x \[E!]x]\\<^sub>0] \ + (\w. [w \ [\']\\<^sub>0] = [w \ [\]\\<^sub>0]))\ + by (safe intro!: exI[where x=\'']) blast+ +qed +end + +text\Define encoding for products using projection-encoding.\ +instantiation prod :: (AOT_UnaryEnc, AOT_Enc) AOT_Enc +begin +definition AOT_proj_enc_prod :: \'a\'b \ ('a\'b \ \) \ \\ where + \AOT_proj_enc_prod \ \ (\,\') \ . \\[\\ \\ (\,\')\] & + \AOT_proj_enc \' (\\. \ (\,\))\\\ +definition AOT_enc_prod :: \'a\'b \ <'a\'b> \ \\ where + \AOT_enc_prod \ \ \ \ . \\\ & \AOT_proj_enc \ (\ \\<^sub>1'\\<^sub>n'. \[\]\\<^sub>1'...\\<^sub>n'\)\\\ +instance proof + show \[v \ \\<^sub>1...\\<^sub>n[\]] \ [v \ \\<^sub>1...\\<^sub>n\] \ [v \ \\]\ + for v and \\<^sub>1\\<^sub>n :: \'a\'b\ and \ + unfolding AOT_enc_prod_def + apply (induct \\<^sub>1\\<^sub>n; simp add: AOT_sem_conj AOT_sem_denotes AOT_proj_enc_prod_def) + by (metis AOT_sem_denotes AOT_model_denotes_prod_def AOT_sem_enc_denotes + AOT_sem_proj_enc_denotes case_prodI) +next + show \[v \ \\<^sub>1...\\<^sub>n[\]] = + [v \ \\\\ & \AOT_proj_enc \\<^sub>1\\<^sub>n (\ \\<^sub>1\\<^sub>n. \[\]\\<^sub>1...\\<^sub>n\)\]\ + for v and \\<^sub>1\\<^sub>n :: \'a\'b\ and \ + unfolding AOT_enc_prod_def .. +next + show \[v \ \AOT_proj_enc \s \\] \ [v \ \\s\\]\ + for v and \s :: \'a\'b\ and \ + by (metis (mono_tags, lifting) + AOT_sem_conj AOT_sem_denotes AOT_model_denotes_prod_def + AOT_sem_enc_denotes AOT_sem_proj_enc_denotes + AOT_proj_enc_prod_def case_prod_unfold) +next + fix v w \ and \\<^sub>1\\<^sub>n :: \'a\'b\ + show \[w \ \\<^sub>1...\\<^sub>n[\]]\ if \[v \ \\<^sub>1...\\<^sub>n[\]]\ for v w \ and \\<^sub>1\\<^sub>n :: \'a\'b\ + by (metis (mono_tags, lifting) + AOT_enc_prod_def AOT_sem_enc_proj_enc AOT_sem_conj AOT_sem_denotes + AOT_sem_proj_enc_nec AOT_proj_enc_prod_def case_prod_unfold that) +next + show \[w \ \AOT_proj_enc \\<^sub>1\\<^sub>n \\]\ if \[v \ \AOT_proj_enc \\<^sub>1\\<^sub>n \\]\ + for v w \ and \\<^sub>1\\<^sub>n :: \'a\'b\ + by (metis (mono_tags, lifting) + that AOT_sem_enc_proj_enc AOT_sem_conj AOT_sem_denotes + AOT_sem_proj_enc_nec AOT_proj_enc_prod_def case_prod_unfold) +next + fix v + obtain \ :: 'a where a_prop: \[v \ \\] \ (\ \ . [v \ \\] \ [v \ \[\]])\ + using AOT_sem_universal_encoder by blast + obtain \\<^sub>1'\\<^sub>n' :: 'b where b_prop: + \[v \ \\<^sub>1'...\\<^sub>n'\] \ (\ \ . [v \ [\\\<^sub>1...\\<^sub>n \\ \\<^sub>1\\<^sub>n\]\] \ + [v \ \AOT_proj_enc \\<^sub>1'\\<^sub>n' \\])\ + using AOT_sem_universal_encoder by blast + have \AOT_model_denotes \[\\\<^sub>1...\\<^sub>n [\\\]\\<^sub>1...\\<^sub>n \\<^sub>1'...\\<^sub>n']\\ + if \AOT_model_denotes \\ for \ :: \<'a\'b>\ + unfolding AOT_model_lambda_denotes + by (metis AOT_meta_prod_equivI(2) AOT_sem_exe_equiv) + moreover have \AOT_model_denotes \[\\\<^sub>1...\\<^sub>n [\\\]\ \\<^sub>1...\\<^sub>n]\\ + if \AOT_model_denotes \\ for \ :: \<'a\'b>\ + unfolding AOT_model_lambda_denotes + by (metis AOT_meta_prod_equivI(1) AOT_sem_exe_equiv) + ultimately have 1: \[v \ \(\,\\<^sub>1'\\<^sub>n')\\]\ + and 2: \(\ \ . [v \ \\] \ [v \ \ \\<^sub>1'...\\<^sub>n'[\]])\ + using a_prop b_prop + by (auto simp: AOT_sem_denotes AOT_enc_\_meta AOT_model_enc_\_def + AOT_model_denotes_\_def AOT_model_denotes_prod_def + AOT_enc_prod_def AOT_proj_enc_prod_def AOT_sem_conj) + have \AOT_model_denotes \[\z\<^sub>1...z\<^sub>n \\ (z\<^sub>1z\<^sub>n, \\<^sub>1'\\<^sub>n')\]\\ + if \AOT_model_denotes \[\z\<^sub>1...z\<^sub>m \{z\<^sub>1...z\<^sub>m}]\\ for \ :: \'a\'b \ \\ + using that + unfolding AOT_model_lambda_denotes + by (metis (no_types, lifting) AOT_sem_denotes AOT_model_denotes_prod_def + AOT_meta_prod_equivI(2) b_prop case_prodI) + moreover have \AOT_model_denotes \[\z\<^sub>1...z\<^sub>n \\ (\, z\<^sub>1z\<^sub>n)\]\\ + if \AOT_model_denotes \[\z\<^sub>1...z\<^sub>m \{z\<^sub>1...z\<^sub>m}]\\ for \ :: \'a\'b \ \\ + using that + unfolding AOT_model_lambda_denotes + by (metis (no_types, lifting) AOT_sem_denotes AOT_model_denotes_prod_def + AOT_meta_prod_equivI(1) a_prop case_prodI) + ultimately have 3: + \[v \ \(\,\\<^sub>1'\\<^sub>n')\\] \ (\ \ . [v \ [\z\<^sub>1...z\<^sub>n \{z\<^sub>1...z\<^sub>n}]\] \ + [v \ \AOT_proj_enc (\,\\<^sub>1'\\<^sub>n') \\])\ + using a_prop b_prop + by (auto simp: AOT_sem_denotes AOT_enc_\_meta AOT_model_enc_\_def + AOT_model_denotes_\_def AOT_enc_prod_def AOT_proj_enc_prod_def + AOT_sem_conj AOT_model_denotes_prod_def) + show \\\\<^sub>1\\<^sub>n::'a\'b. [v \ \\<^sub>1...\\<^sub>n\] \ (\ \ . [v \ \\] \ [v \ \\<^sub>1...\\<^sub>n[\]]) \ + (\ \ . [v \ [\z\<^sub>1...z\<^sub>n \\ z\<^sub>1z\<^sub>n\]\] \ + [v \ \AOT_proj_enc \\<^sub>1\\<^sub>n \\])\ + apply (rule exI[where x=\(\,\\<^sub>1'\\<^sub>n')\]) using 1 2 3 by blast +qed +end + +text\Sanity-check to verify that n-ary encoding follows.\ +lemma \[v \ \\<^sub>1\\<^sub>2[\]] = [v \ \\ & \\<^sub>1[\\ [\]\\\<^sub>2] & \\<^sub>2[\\ [\]\\<^sub>1\]]\ + for \\<^sub>1 :: "'a::AOT_UnaryEnc" and \\<^sub>2 :: "'b::AOT_UnaryEnc" + by (simp add: AOT_sem_conj AOT_enc_prod_def AOT_proj_enc_prod_def + AOT_sem_unary_proj_enc) +lemma \[v \ \\<^sub>1\\<^sub>2\\<^sub>3[\]] = + [v \ \\ & \\<^sub>1[\\ [\]\\\<^sub>2\\<^sub>3] & \\<^sub>2[\\ [\]\\<^sub>1\\\<^sub>3] & \\<^sub>3[\\ [\]\\<^sub>1\\<^sub>2\]]\ + for \\<^sub>1 \\<^sub>2 \\<^sub>3 :: "'a::AOT_UnaryEnc" + by (simp add: AOT_sem_conj AOT_enc_prod_def AOT_proj_enc_prod_def + AOT_sem_unary_proj_enc) + +lemma AOT_sem_vars_denote: \[v \ \\<^sub>1...\\<^sub>n\]\ + by induct simp + +text\Combine the introduced type classes and register them as + type constraints for individual terms.\ +class AOT_\s = AOT_IndividualTerm + AOT_RelationProjection + AOT_Enc +class AOT_\ = AOT_\s + AOT_UnaryIndividualTerm + + AOT_UnaryRelationProjection + AOT_UnaryEnc + +instance \ :: AOT_\ by standard +instance prod :: (AOT_\, AOT_\s) AOT_\s by standard + +AOT_register_type_constraints + Individual: \_::AOT_\\ \_::AOT_\s\ and + Relation: \<_::AOT_\s>\ + +text\We define semantic predicates to capture the conditions of cqt.2 (i.e. + the base cases of denoting terms) on matrices of @{text \}-expressions.\ +definition AOT_instance_of_cqt_2 :: \('a::AOT_\s \ \) \ bool\ where + \AOT_instance_of_cqt_2 \ \ \ . \ x y . AOT_model_denotes x \ AOT_model_denotes y \ + AOT_model_term_equiv x y \ \ x = \ y\ +definition AOT_instance_of_cqt_2_exe_arg :: \('a::AOT_\s \ 'b::AOT_\s) \ bool\ where + \AOT_instance_of_cqt_2_exe_arg \ \ \ . \ x y . + AOT_model_denotes x \ AOT_model_denotes y \ AOT_model_term_equiv x y \ + AOT_model_term_equiv (\ x) (\ y)\ + +text\@{text \}-expressions with a matrix that satisfies our predicate denote.\ +lemma AOT_sem_cqt_2: + assumes \AOT_instance_of_cqt_2 \\ + shows \[v \ [\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n}]\]\ + using assms + by (metis AOT_instance_of_cqt_2_def AOT_model_lambda_denotes AOT_sem_denotes) + +syntax AOT_instance_of_cqt_2 :: \id_position \ AOT_prop\ + ("INSTANCE'_OF'_CQT'_2'(_')") + +text\Prove introduction rules for the predicates that match the natural language + restrictions of the axiom.\ +named_theorems AOT_instance_of_cqt_2_intro +lemma AOT_instance_of_cqt_2_intros_const[AOT_instance_of_cqt_2_intro]: + \AOT_instance_of_cqt_2 (\\. \)\ + by (simp add: AOT_instance_of_cqt_2_def AOT_sem_denotes AOT_model_lambda_denotes) +lemma AOT_instance_of_cqt_2_intros_not[AOT_instance_of_cqt_2_intro]: + assumes \AOT_instance_of_cqt_2 \\ + shows \AOT_instance_of_cqt_2 (\\. \\\{\}\)\ + using assms + by (metis (no_types, lifting) AOT_instance_of_cqt_2_def) +lemma AOT_instance_of_cqt_2_intros_imp[AOT_instance_of_cqt_2_intro]: + assumes \AOT_instance_of_cqt_2 \\ and \AOT_instance_of_cqt_2 \\ + shows \AOT_instance_of_cqt_2 (\\. \\{\} \ \{\}\)\ + using assms + by (auto simp: AOT_instance_of_cqt_2_def AOT_sem_denotes + AOT_model_lambda_denotes AOT_sem_imp) +lemma AOT_instance_of_cqt_2_intros_box[AOT_instance_of_cqt_2_intro]: + assumes \AOT_instance_of_cqt_2 \\ + shows \AOT_instance_of_cqt_2 (\\. \\\{\}\)\ + using assms + by (auto simp: AOT_instance_of_cqt_2_def AOT_sem_denotes + AOT_model_lambda_denotes AOT_sem_box) +lemma AOT_instance_of_cqt_2_intros_act[AOT_instance_of_cqt_2_intro]: + assumes \AOT_instance_of_cqt_2 \\ + shows \AOT_instance_of_cqt_2 (\\. \\<^bold>\\{\}\)\ + using assms + by (auto simp: AOT_instance_of_cqt_2_def AOT_sem_denotes + AOT_model_lambda_denotes AOT_sem_act) +lemma AOT_instance_of_cqt_2_intros_diamond[AOT_instance_of_cqt_2_intro]: + assumes \AOT_instance_of_cqt_2 \\ + shows \AOT_instance_of_cqt_2 (\\. \\\{\}\)\ + using assms + by (auto simp: AOT_instance_of_cqt_2_def AOT_sem_denotes + AOT_model_lambda_denotes AOT_sem_dia) +lemma AOT_instance_of_cqt_2_intros_conj[AOT_instance_of_cqt_2_intro]: + assumes \AOT_instance_of_cqt_2 \\ and \AOT_instance_of_cqt_2 \\ + shows \AOT_instance_of_cqt_2 (\\. \\{\} & \{\}\)\ + using assms + by (auto simp: AOT_instance_of_cqt_2_def AOT_sem_denotes + AOT_model_lambda_denotes AOT_sem_conj) +lemma AOT_instance_of_cqt_2_intros_disj[AOT_instance_of_cqt_2_intro]: + assumes \AOT_instance_of_cqt_2 \\ and \AOT_instance_of_cqt_2 \\ + shows \AOT_instance_of_cqt_2 (\\. \\{\} \ \{\}\)\ + using assms + by (auto simp: AOT_instance_of_cqt_2_def AOT_sem_denotes + AOT_model_lambda_denotes AOT_sem_disj) +lemma AOT_instance_of_cqt_2_intros_equib[AOT_instance_of_cqt_2_intro]: + assumes \AOT_instance_of_cqt_2 \\ and \AOT_instance_of_cqt_2 \\ + shows \AOT_instance_of_cqt_2 (\\. \\{\} \ \{\}\)\ + using assms + by (auto simp: AOT_instance_of_cqt_2_def AOT_sem_denotes + AOT_model_lambda_denotes AOT_sem_equiv) +lemma AOT_instance_of_cqt_2_intros_forall[AOT_instance_of_cqt_2_intro]: + assumes \\ \ . AOT_instance_of_cqt_2 (\ \)\ + shows \AOT_instance_of_cqt_2 (\\. \\\ \{\,\}\)\ + using assms + by (auto simp: AOT_instance_of_cqt_2_def AOT_sem_denotes + AOT_model_lambda_denotes AOT_sem_forall) +lemma AOT_instance_of_cqt_2_intros_exists[AOT_instance_of_cqt_2_intro]: + assumes \\ \ . AOT_instance_of_cqt_2 (\ \)\ + shows \AOT_instance_of_cqt_2 (\\. \\\ \{\,\}\)\ + using assms + by (auto simp: AOT_instance_of_cqt_2_def AOT_sem_denotes + AOT_model_lambda_denotes AOT_sem_exists) +lemma AOT_instance_of_cqt_2_intros_exe_arg_self[AOT_instance_of_cqt_2_intro]: + \AOT_instance_of_cqt_2_exe_arg (\x. x)\ + unfolding AOT_instance_of_cqt_2_exe_arg_def AOT_instance_of_cqt_2_def + AOT_sem_lambda_denotes + by (auto simp: AOT_model_term_equiv_part_equivp equivp_reflp AOT_sem_denotes) +lemma AOT_instance_of_cqt_2_intros_exe_arg_const[AOT_instance_of_cqt_2_intro]: + \AOT_instance_of_cqt_2_exe_arg (\x. \)\ + unfolding AOT_instance_of_cqt_2_exe_arg_def AOT_instance_of_cqt_2_def + by (auto simp: AOT_model_term_equiv_part_equivp equivp_reflp + AOT_sem_denotes AOT_sem_lambda_denotes) +lemma AOT_instance_of_cqt_2_intros_exe_arg_fst[AOT_instance_of_cqt_2_intro]: + \AOT_instance_of_cqt_2_exe_arg fst\ + unfolding AOT_instance_of_cqt_2_exe_arg_def AOT_instance_of_cqt_2_def + by (simp add: AOT_model_term_equiv_prod_def case_prod_beta) +lemma AOT_instance_of_cqt_2_intros_exe_arg_snd[AOT_instance_of_cqt_2_intro]: + \AOT_instance_of_cqt_2_exe_arg snd\ + unfolding AOT_instance_of_cqt_2_exe_arg_def AOT_instance_of_cqt_2_def + by (simp add: AOT_model_term_equiv_prod_def AOT_sem_denotes AOT_sem_lambda_denotes) +lemma AOT_instance_of_cqt_2_intros_exe_arg_Pair[AOT_instance_of_cqt_2_intro]: + assumes \AOT_instance_of_cqt_2_exe_arg \\ and \AOT_instance_of_cqt_2_exe_arg \\ + shows \AOT_instance_of_cqt_2_exe_arg (\\. Pair (\ \) (\ \))\ + using assms + unfolding AOT_instance_of_cqt_2_exe_arg_def AOT_instance_of_cqt_2_def + AOT_sem_denotes AOT_sem_lambda_denotes AOT_model_term_equiv_prod_def + AOT_model_denotes_prod_def + by auto +lemma AOT_instance_of_cqt_2_intros_desc[AOT_instance_of_cqt_2_intro]: + assumes \\z :: 'a::AOT_\. AOT_instance_of_cqt_2 (\ z)\ + shows \AOT_instance_of_cqt_2_exe_arg (\ \ :: 'b::AOT_\ . \\<^bold>\z(\{z,\})\)\ +proof - + have 0: \\ \ \'. AOT_model_denotes \ \ AOT_model_denotes \' \ + AOT_model_term_equiv \ \' \ + \ z \ = \ z \'\ for z + using assms + unfolding AOT_instance_of_cqt_2_def + AOT_sem_denotes AOT_model_lambda_denotes by force + { + fix \ \' + have \\\<^bold>\z(\{z,\})\ = \\<^bold>\z(\{z,\'})\\ + if \AOT_model_denotes \ \ AOT_model_denotes \' \ AOT_model_term_equiv \ \'\ + using 0[OF that] + by auto + moreover have \AOT_model_term_equiv x x\ for x :: \'a::AOT_\\ + by (metis AOT_instance_of_cqt_2_exe_arg_def + AOT_instance_of_cqt_2_intros_exe_arg_const + AOT_model_A_objects AOT_model_term_equiv_denotes + AOT_model_term_equiv_eps(1)) + ultimately have \AOT_model_term_equiv \\<^bold>\z(\{z,\})\ \\<^bold>\z(\{z,\'})\\ + if \AOT_model_denotes \ \ AOT_model_denotes \' \ AOT_model_term_equiv \ \'\ + using that by simp + } + thus ?thesis using 0 + unfolding AOT_instance_of_cqt_2_exe_arg_def + by simp +qed + +lemma AOT_instance_of_cqt_2_intros_exe_const[AOT_instance_of_cqt_2_intro]: + assumes \AOT_instance_of_cqt_2_exe_arg \s\ + shows \AOT_instance_of_cqt_2 (\x :: 'b::AOT_\s. AOT_exe \ (\s x))\ + using assms + unfolding AOT_instance_of_cqt_2_def AOT_sem_denotes AOT_model_lambda_denotes + AOT_sem_disj AOT_sem_conj + AOT_sem_not AOT_sem_box AOT_sem_act AOT_instance_of_cqt_2_exe_arg_def + AOT_sem_equiv AOT_sem_imp AOT_sem_forall AOT_sem_exists AOT_sem_dia + by (auto intro!: AOT_sem_exe_equiv) +lemma AOT_instance_of_cqt_2_intros_exe_lam[AOT_instance_of_cqt_2_intro]: + assumes \\ y . AOT_instance_of_cqt_2 (\x. \ x y)\ + and \AOT_instance_of_cqt_2_exe_arg \s\ + shows \AOT_instance_of_cqt_2 (\\\<^sub>1\\<^sub>n :: 'b::AOT_\s. + \[\\\<^sub>1...\\<^sub>n \{\\<^sub>1...\\<^sub>n,\\<^sub>1...\\<^sub>n}]\\s \\<^sub>1\\<^sub>n\\)\ +proof - + { + fix x y :: 'b + assume \AOT_model_denotes x\ + moreover assume \AOT_model_denotes y\ + moreover assume \AOT_model_term_equiv x y\ + moreover have 1: \\ x = \ y\ + using assms calculation unfolding AOT_instance_of_cqt_2_def by blast + ultimately have \AOT_exe (AOT_lambda (\ x)) (\s x) = + AOT_exe (AOT_lambda (\ y)) (\s y)\ + unfolding 1 + apply (safe intro!: AOT_sem_exe_equiv) + by (metis AOT_instance_of_cqt_2_exe_arg_def assms(2)) + } + thus ?thesis + unfolding AOT_instance_of_cqt_2_def + AOT_instance_of_cqt_2_exe_arg_def + by blast +qed +lemma AOT_instance_of_cqt_2_intro_prod[AOT_instance_of_cqt_2_intro]: + assumes \\ x . AOT_instance_of_cqt_2 (\ x)\ + and \\ x . AOT_instance_of_cqt_2 (\ z . \ z x)\ + shows \AOT_instance_of_cqt_2 (\(x,y) . \ x y)\ + using assms unfolding AOT_instance_of_cqt_2_def + by (auto simp add: AOT_model_lambda_denotes AOT_sem_denotes + AOT_model_denotes_prod_def + AOT_model_term_equiv_prod_def) + +text\The following are already derivable semantically, but not yet added + to @{attribute AOT_instance_of_cqt_2_intro}. They will be added with the + next planned extension of axiom cqt:2.\ +named_theorems AOT_instance_of_cqt_2_intro_next +definition AOT_instance_of_cqt_2_enc_arg :: \('a::AOT_\s \ 'b::AOT_\s) \ bool\ where + \AOT_instance_of_cqt_2_enc_arg \ \ \ . \ x y z . + AOT_model_denotes x \ AOT_model_denotes y \ AOT_model_term_equiv x y \ + AOT_enc (\ x) z = AOT_enc (\ y) z\ +definition AOT_instance_of_cqt_2_enc_rel :: \('a::AOT_\s \ <'b::AOT_\s>) \ bool\ where + \AOT_instance_of_cqt_2_enc_rel \ \ \ . \ x y z . + AOT_model_denotes x \ AOT_model_denotes y \ AOT_model_term_equiv x y \ + AOT_enc z (\ x) = AOT_enc z (\ y)\ +lemma AOT_instance_of_cqt_2_intros_enc[AOT_instance_of_cqt_2_intro_next]: + assumes \AOT_instance_of_cqt_2_enc_rel \\ and \AOT_instance_of_cqt_2_enc_arg \s\ + shows \AOT_instance_of_cqt_2 (\x . AOT_enc (\s x) \[\\ x\]\)\ + using assms + unfolding AOT_instance_of_cqt_2_def AOT_sem_denotes AOT_model_lambda_denotes + AOT_instance_of_cqt_2_enc_rel_def AOT_sem_not AOT_sem_box AOT_sem_act + AOT_sem_dia AOT_sem_conj AOT_sem_disj AOT_sem_equiv AOT_sem_imp + AOT_sem_forall AOT_sem_exists AOT_instance_of_cqt_2_enc_arg_def + by fastforce+ +lemma AOT_instance_of_cqt_2_enc_arg_intro_const[AOT_instance_of_cqt_2_intro_next]: + \AOT_instance_of_cqt_2_enc_arg (\x. c)\ + unfolding AOT_instance_of_cqt_2_enc_arg_def by simp +lemma AOT_instance_of_cqt_2_enc_arg_intro_desc[AOT_instance_of_cqt_2_intro_next]: + assumes \\z :: 'a::AOT_\. AOT_instance_of_cqt_2 (\ z)\ + shows \AOT_instance_of_cqt_2_enc_arg (\ \ :: 'b::AOT_\ . \\<^bold>\z(\{z,\})\)\ +proof - + have 0: \\ \ \'. AOT_model_denotes \ \ AOT_model_denotes \' \ + AOT_model_term_equiv \ \' \ + \ z \ = \ z \'\ for z + using assms + unfolding AOT_instance_of_cqt_2_def + AOT_sem_denotes AOT_model_lambda_denotes by force + { + fix \ \' + have \\\<^bold>\z(\{z,\})\ = \\<^bold>\z(\{z,\'})\\ + if \AOT_model_denotes \ \ AOT_model_denotes \' \ AOT_model_term_equiv \ \'\ + using 0[OF that] + by auto + } + thus ?thesis using 0 + unfolding AOT_instance_of_cqt_2_enc_arg_def by meson +qed +lemma AOT_instance_of_cqt_2_enc_rel_intro[AOT_instance_of_cqt_2_intro_next]: + assumes \\ \ . AOT_instance_of_cqt_2 (\\' :: 'b::AOT_\s . \ \ \')\ + assumes \\ \' . AOT_instance_of_cqt_2 (\\ :: 'a::AOT_\s . \ \ \')\ + shows \AOT_instance_of_cqt_2_enc_rel (\\ :: 'a::AOT_\s. AOT_lambda (\\'. \ \ \'))\ +proof - + { + fix x y :: 'a and z ::'b + assume \AOT_model_term_equiv x y\ + moreover assume \AOT_model_denotes x\ + moreover assume \AOT_model_denotes y\ + ultimately have \\ x = \ y\ + using assms unfolding AOT_instance_of_cqt_2_def by blast + hence \AOT_enc z (AOT_lambda (\ x)) = AOT_enc z (AOT_lambda (\ y))\ + by simp + } + thus ?thesis + unfolding AOT_instance_of_cqt_2_enc_rel_def by auto +qed + +text\Further restrict unary individual variables to type @{typ \} (rather + than class @{class AOT_\} only) and define being ordinary and being abstract.\ +AOT_register_type_constraints + Individual: \\\ \_::AOT_\s\ + +AOT_define AOT_ordinary :: \\\ (\O!\) \O! =\<^sub>d\<^sub>f [\x \E!x]\ +declare AOT_ordinary[AOT del, AOT_defs del] +AOT_define AOT_abstract :: \\\ (\A!\) \A! =\<^sub>d\<^sub>f [\x \\E!x]\ +declare AOT_abstract[AOT del, AOT_defs del] + +context AOT_meta_syntax +begin +notation AOT_ordinary ("\<^bold>O\<^bold>!") +notation AOT_abstract ("\<^bold>A\<^bold>!") +end +context AOT_no_meta_syntax +begin +no_notation AOT_ordinary ("\<^bold>O\<^bold>!") +no_notation AOT_abstract ("\<^bold>A\<^bold>!") +end + +no_translations + "_AOT_concrete" => "CONST AOT_term_of_var (CONST AOT_concrete)" +parse_translation\ +[(\<^syntax_const>\_AOT_concrete\, fn _ => fn [] => + Const (\<^const_name>\AOT_term_of_var\, dummyT) + $ Const (\<^const_name>\AOT_concrete\, \<^typ>\<\> AOT_var\))] +\ + +text\Auxiliary lemmata.\ +lemma AOT_sem_ordinary: "\O!\ = \[\x \E!x]\" + using AOT_ordinary[THEN AOT_sem_id_def0E1] AOT_sem_ordinary_def_denotes + by (auto simp: AOT_sem_eq) +lemma AOT_sem_abstract: "\A!\ = \[\x \\E!x]\" + using AOT_abstract[THEN AOT_sem_id_def0E1] AOT_sem_abstract_def_denotes + by (auto simp: AOT_sem_eq) +lemma AOT_sem_ordinary_denotes: \[w \ O!\]\ + by (simp add: AOT_sem_ordinary AOT_sem_ordinary_def_denotes) +lemma AOT_meta_abstract_denotes: \[w \ A!\]\ + by (simp add: AOT_sem_abstract AOT_sem_abstract_def_denotes) +lemma AOT_model_abstract_\\: \\ a . \ = \\ a\ if \[v \ A!\]\ + using that[unfolded AOT_sem_abstract, simplified + AOT_meta_abstract_denotes[unfolded AOT_sem_abstract, THEN AOT_sem_lambda_beta, + OF that[simplified AOT_sem_exe, THEN conjunct2, THEN conjunct1]]] + apply (simp add: AOT_sem_not AOT_sem_dia AOT_sem_concrete) + by (metis AOT_model_\_concrete_in_some_world AOT_model_concrete_\.simps(1) + AOT_model_denotes_\_def AOT_sem_denotes AOT_sem_exe \.exhaust_disc + is_\\_def is_\\_def that) +lemma AOT_model_ordinary_\\: \\ a . \ = \\ a\ if \[v \ O!\]\ + using that[unfolded AOT_sem_ordinary, simplified + AOT_sem_ordinary_denotes[unfolded AOT_sem_ordinary, THEN AOT_sem_lambda_beta, + OF that[simplified AOT_sem_exe, THEN conjunct2, THEN conjunct1]]] + apply (simp add: AOT_sem_dia AOT_sem_concrete) + by (metis AOT_model_concrete_\.simps(2) AOT_model_concrete_\.simps(3) + \.exhaust_disc is_\\_def is_\\_def is_null\_def) +lemma AOT_model_\\_ordinary: \[v \ O!\\\ x\]\ + by (metis AOT_model_abstract_\\ AOT_model_denotes_\_def AOT_sem_abstract + AOT_sem_denotes AOT_sem_ind_eq AOT_sem_ordinary \.disc(7) \.distinct(1)) +lemma AOT_model_\\_ordinary: \[v \ A!\\\ x\]\ + by (metis AOT_model_denotes_\_def AOT_model_ordinary_\\ AOT_sem_abstract + AOT_sem_denotes AOT_sem_ind_eq AOT_sem_ordinary \.disc(8) \.distinct(1)) +AOT_theorem prod_denotesE: assumes \\(\\<^sub>1,\\<^sub>2)\\\ shows \\\<^sub>1\ & \\<^sub>2\\ + using assms by (simp add: AOT_sem_denotes AOT_sem_conj AOT_model_denotes_prod_def) +declare prod_denotesE[AOT del] +AOT_theorem prod_denotesI: assumes \\\<^sub>1\ & \\<^sub>2\\ shows \\(\\<^sub>1,\\<^sub>2)\\\ + using assms by (simp add: AOT_sem_denotes AOT_sem_conj AOT_model_denotes_prod_def) +declare prod_denotesI[AOT del] + + +text\Prepare the derivation of the additional axioms that are validated by + our extended models.\ +locale AOT_ExtendedModel = + assumes AOT_ExtendedModel: \AOT_ExtendedModel\ +begin +lemma AOT_sem_indistinguishable_ord_enc_all: + assumes \_den: \[v \ \\]\ + assumes Ax: \[v \ A!x]\ + assumes Ay: \[v \ A!y]\ + assumes indist: \[v \ \F \([F]x \ [F]y)]\ + shows + \[v \ \G(\z(O!z \ \([G]z \ [\]z)) \ x[G])] = + [v \ \G(\z(O!z \ \([G]z \ [\]z)) \ y[G])]\ +proof - + have 0: \[v \ [\x \\[E!]x]x]\ + using Ax by (simp add: AOT_sem_abstract) + have 1: \[v \ [\x \\[E!]x]y]\ + using Ay by (simp add: AOT_sem_abstract) + { + assume \[v \ \G(\z (O!z \ \([G]z \ [\]z)) \ x[G])]\ + hence 3: \[v \ \G(\z([\x \[E!]x]z \ \([G]z \ [\]z)) \ x[G])]\ + by (simp add: AOT_sem_ordinary) + { + fix \' :: \<\>\ + assume 1: \[v \ \'\]\ + assume 2: \[v \ [\x \[E!]x]z \ \([\']z \ [\]z)]\ for z + have \[v \ x[\']]\ + using 3 + by (auto simp: AOT_sem_forall AOT_sem_imp AOT_sem_box AOT_sem_denotes) + (metis (no_types, lifting) 1 2 AOT_term_of_var_cases AOT_sem_box + AOT_sem_denotes AOT_sem_imp) + } note 3 = this + fix \' :: \<\>\ + assume \_den: \[v \ \'\]\ + assume 4: \[v \ \z (O!z \ \([\']z \ [\]z))]\ + { + fix \\<^sub>0 + assume \[v \ [\x \[E!]x]\\<^sub>0]\ + hence \[v \ O!\\<^sub>0]\ + using AOT_sem_ordinary by metis + moreover have \[v \ \\<^sub>0\]\ + using calculation by (simp add: AOT_sem_exe) + ultimately have \[v \ \([\']\\<^sub>0 \ [\]\\<^sub>0)]\ + using 4 by (auto simp: AOT_sem_forall AOT_sem_imp) + } note 4 = this + { + fix \'' :: \<\>\ + assume \[v \ \''\]\ + moreover assume \[w \ [\'']\\<^sub>0] = [w \ [\']\\<^sub>0]\ if \[v \ [\x \E!x]\\<^sub>0]\ for \\<^sub>0 w + ultimately have 5: \[v \ x[\'']]\ + using 4 3 + by (auto simp: AOT_sem_imp AOT_sem_equiv AOT_sem_box) + } note 5 = this + have \[v \ y[\']]\ + apply (rule AOT_sem_enc_indistinguishable_all[OF AOT_ExtendedModel]) + apply (fact 0) + by (auto simp: 5 0 1 \_den indist[simplified AOT_sem_forall + AOT_sem_box AOT_sem_equiv]) + } + moreover { + { + assume \[v \ \G(\z (O!z \ \([G]z \ [\]z)) \ y[G])]\ + hence 3: \[v \ \G(\z ([\x \[E!]x]z \ \([G]z \ [\]z)) \ y[G])]\ + by (simp add: AOT_sem_ordinary) + { + fix \' :: \<\>\ + assume 1: \[v \ \'\]\ + assume 2: \[v \ [\x \[E!]x]z \ \([\']z \ [\]z)]\ for z + have \[v \ y[\']]\ + using 3 + apply (simp add: AOT_sem_forall AOT_sem_imp AOT_sem_box AOT_sem_denotes) + by (metis (no_types, lifting) 1 2 AOT_model.AOT_term_of_var_cases + AOT_sem_box AOT_sem_denotes AOT_sem_imp) + } note 3 = this + fix \' :: \<\>\ + assume \_den: \[v \ \'\]\ + assume 4: \[v \ \z (O!z \ \([\']z \ [\]z))]\ + { + fix \\<^sub>0 + assume \[v \ [\x \[E!]x]\\<^sub>0]\ + hence \[v \ O!\\<^sub>0]\ + using AOT_sem_ordinary by metis + moreover have \[v \ \\<^sub>0\]\ + using calculation by (simp add: AOT_sem_exe) + ultimately have \[v \ \([\']\\<^sub>0 \ [\]\\<^sub>0)]\ + using 4 by (auto simp: AOT_sem_forall AOT_sem_imp) + } note 4 = this + { + fix \'' :: \<\>\ + assume \[v \ \''\]\ + moreover assume \[w \ [\'']\\<^sub>0] = [w \ [\']\\<^sub>0]\ if \[v \ [\x \E!x]\\<^sub>0]\ for w \\<^sub>0 + ultimately have \[v \ y[\'']]\ + using 3 4 by (auto simp: AOT_sem_imp AOT_sem_equiv AOT_sem_box) + } note 5 = this + have \[v \ x[\']]\ + apply (rule AOT_sem_enc_indistinguishable_all[OF AOT_ExtendedModel]) + apply (fact 1) + by (auto simp: 5 0 1 \_den indist[simplified AOT_sem_forall + AOT_sem_box AOT_sem_equiv]) + } + } + ultimately show \[v \ \G (\z (O!z \ \([G]z \ [\]z)) \ x[G])] = + [v \ \G (\z (O!z \ \([G]z \ [\]z)) \ y[G])]\ + by (auto simp: AOT_sem_forall AOT_sem_imp) +qed + +lemma AOT_sem_indistinguishable_ord_enc_ex: + assumes \_den: \[v \ \\]\ + assumes Ax: \[v \ A!x]\ + assumes Ay: \[v \ A!y]\ + assumes indist: \[v \ \F \([F]x \ [F]y)]\ + shows \[v \ \G(\z (O!z \ \([G]z \ [\]z)) & x[G])] = + [v \ \G(\z(O!z \ \([G]z \ [\]z)) & y[G])]\ +proof - + have Aux: \[v \ [\x \[E!]x]\] = ([v \ [\x \[E!]x]\] \ [v \ \\])\ for v \ + using AOT_sem_exe by blast + AOT_modally_strict { + fix x y + AOT_assume \_den: \[\]\\ + AOT_assume 2: \\F \([F]x \ [F]y)\ + AOT_assume \A!x\ + AOT_hence 0: \[\x \\[E!]x]x\ + by (simp add: AOT_sem_abstract) + AOT_assume \A!y\ + AOT_hence 1: \[\x \\[E!]x]y\ + by (simp add: AOT_sem_abstract) + { + AOT_assume \\G(\z (O!z \ \([G]z \ [\]z)) & x[G])\ + then AOT_obtain \' + where \'_den: \\'\\ + and \'_indist: \\z (O!z \ \([\']z \ [\]z))\ + and x_enc_\': \x[\']\ + by (meson AOT_sem_conj AOT_sem_exists) + { + fix \\<^sub>0 + AOT_assume \[\x \[E!]x]\\<^sub>0\ + AOT_hence \\([\']\\<^sub>0 \ [\]\\<^sub>0)\ + using \'_indist + by (auto simp: AOT_sem_exe AOT_sem_imp AOT_sem_exists AOT_sem_conj + AOT_sem_ordinary AOT_sem_forall) + } note 3 = this + AOT_have \\z ([\x \[E!]x]z \ \([\']z \ [\]z))\ + using \'_indist by (simp add: AOT_sem_ordinary) + AOT_obtain \'' where + \''_den: \\''\\ and + \''_indist: \[\x \[E!]x]\\<^sub>0 \ \([\'']\\<^sub>0 \ [\]\\<^sub>0)\ and + y_enc_\'': \y[\'']\ for \\<^sub>0 + using AOT_sem_enc_indistinguishable_ex[OF AOT_ExtendedModel, + OF 0, OF 1, rotated, OF \_den, + OF exI[where x=\'], OF conjI, OF \'_den, OF conjI, + OF x_enc_\', OF allI, OF impI, + OF 3[simplified AOT_sem_box AOT_sem_equiv], simplified, OF + 2[simplified AOT_sem_forall AOT_sem_equiv AOT_sem_box, + THEN spec, THEN mp, THEN spec], simplified] + unfolding AOT_sem_imp AOT_sem_box AOT_sem_equiv by blast + { + AOT_have \\''\\ + and \\x ([\x \[E!]x]x \ \([\'']x \ [\]x))\ + and \y[\'']\ + apply (simp add: \''_den) + apply (simp add: AOT_sem_forall \''_indist) + by (simp add: y_enc_\'') + } note 2 = this + AOT_have \\G(\z (O!z \ \([G]z \ [\]z)) & y[G])\ + apply (simp add: AOT_sem_exists AOT_sem_ordinary + AOT_sem_imp AOT_sem_box AOT_sem_forall AOT_sem_equiv AOT_sem_conj) + using 2[simplified AOT_sem_box AOT_sem_equiv AOT_sem_imp AOT_sem_forall] + by blast + } + } note 0 = this + AOT_modally_strict { + { + fix x y + AOT_assume \_den: \[\]\\ + moreover AOT_assume \\F \([F]x \ [F]y)\ + moreover AOT_have \\F \([F]y \ [F]x)\ + using calculation(2) + by (auto simp: AOT_sem_forall AOT_sem_box AOT_sem_equiv) + moreover AOT_assume \A!x\ + moreover AOT_assume \A!y\ + ultimately AOT_have \\G (\z (O!z \ \([G]z \ [\]z)) & x[G]) \ + \G (\z (O!z \ \([G]z \ [\]z)) & y[G])\ + using 0 by (auto simp: AOT_sem_equiv) + } + have 1: \[v \ \F \([F]y \ [F]x)]\ + using indist + by (auto simp: AOT_sem_forall AOT_sem_box AOT_sem_equiv) + thus \[v \ \G (\z (O!z \ \([G]z \ [\]z)) & x[G])] = + [v \ \G (\z (O!z \ \([G]z \ [\]z)) & y[G])]\ + using assms + by (auto simp: AOT_sem_imp AOT_sem_conj AOT_sem_equiv 0) + } +qed +end + + +(* Collect all theorems that are not in Main and not declared [AOT] + and store them in a blacklist. *) +setup\setup_AOT_no_atp\ +bundle AOT_no_atp begin declare AOT_no_atp[no_atp] end +(* Can be used as: "including AOT_no_atp sledgehammer" or + "sledgehammer(del: AOT_no_atp) *) + +(*<*) +end +(*>*) \ No newline at end of file diff --git a/thys/AOT/AOT_syntax.ML b/thys/AOT/AOT_syntax.ML new file mode 100644 --- /dev/null +++ b/thys/AOT/AOT_syntax.ML @@ -0,0 +1,537 @@ +fun AOT_binder_trans thy bnd syntaxConst = + (Lexicon.mark_const (Sign.full_name thy bnd), + K (fn trms => Term.list_comb (Const (syntaxConst, dummyT),trms))) + +datatype AOT_VariableKind = AOT_Variable of (term*term) option | AOT_MetaVariable +structure AOT_VariablePrefix = Theory_Data ( + type T = (AOT_VariableKind*string) Symtab.table + val empty = Symtab.empty + val extend = I + (* TODO: probably better to remove conflicts than to ignore them *) + val merge = Symtab.merge (K true) +); +structure AOT_PremiseSetPrefix = Theory_Data ( + type T = unit Symtab.table + val empty = Symtab.empty + val extend = I + val merge = Symtab.merge (K true) +); +structure AOT_Constraints = Theory_Data ( + type T = (term*term) Symtab.table + val empty = Symtab.empty + val extend = I + val merge = Symtab.merge (fn ((x,y),(x',y')) => x = x' andalso y = y') +) +structure AOT_Restriction = Theory_Data ( + type T = (term*term) Symtab.table + val empty = Symtab.empty + val extend = I + val merge = Symtab.merge (fn ((x,y),(x',y')) => x = x' andalso y = y') +) + +fun AOT_IsPremiseSetPrefix ctxt = Local_Theory.raw_theory_result + (fn thy => (AOT_PremiseSetPrefix.get thy, thy)) ctxt + |> fst |> Symtab.lookup #> Option.isSome + +fun term_of_sort S = + let + val class = Syntax.const o Lexicon.mark_class; + fun classes [c] = class c + | classes (c :: cs) = Syntax.const "_classes" $ class c $ classes cs + | classes _ = raise Fail "Unexpected."; + in + if S = dummyS then Syntax.const "_dummy_sort" + else + (case S of + [] => Syntax.const "_topsort" + | [c] => class c + | cs => Syntax.const "_sort" $ classes cs) + end +fun term_of (Type (a, Ts)) = + Term.list_comb (Syntax.const (Lexicon.mark_type a), map term_of Ts) + | term_of (TFree ("'_dummy_",sort)) = + (Const ("_dummy_ofsort", dummyT) $ term_of_sort sort) + | term_of (t as TFree _) = (@{print} t; raise Term.TYPE ("", [t], [])) + | term_of (TVar _) = raise Fail ""; + +fun fetchTermCategory ctxt = Local_Theory.raw_theory_result (fn thy => + (Symtab.lookup (AOT_VariablePrefix.get thy), thy)) ctxt |> fst +fun maybeGetConstraint ctxt unary name = Local_Theory.raw_theory_result (fn thy => + ((if unary then Option.map fst else Option.map snd) + (Symtab.lookup (AOT_Constraints.get thy) name), thy)) ctxt |> fst +fun getConstraint ctxt unary name = + (case maybeGetConstraint ctxt unary name of SOME c => c | + _ => raise Fail ("Unknown type category: " ^ name)) + +fun fetchTermConstraint ctxt name unary = + Local_Theory.raw_theory_result (fn thy => + (Option.map (fn (meta, category) => (meta, getConstraint ctxt unary category)) + ((Symtab.lookup o AOT_VariablePrefix.get) thy (hd (Symbol.explode name))), thy) +) ctxt |> fst + +fun register_constraint (name:string, (unaryConstraint,naryConstraint)) thy = ( +let +fun trmOf constr = term_of (Syntax.parse_typ (Proof_Context.init_global thy) constr) +val unaryConstraint = trmOf unaryConstraint +val naryConstraint = (case naryConstraint of + (SOME constraint) => trmOf constraint + | _ => unaryConstraint +) +in +AOT_Constraints.map (Symtab.update (name, (unaryConstraint, naryConstraint))) thy +end +) + +fun register_variable_name meta (category, prefices) thy = +let +val restr = (Symtab.lookup (AOT_Restriction.get thy) category) +val kind = if meta then AOT_MetaVariable else AOT_Variable restr +in + fold (fn prefix => AOT_VariablePrefix.map + (Symtab.update (prefix, (kind, category)))) prefices thy +end +val _ = + Outer_Syntax.command \<^command_keyword>\AOT_register_variable_names\ + "Register variable names for type categories." + (Parse.and_list1 ((Parse.short_ident --| Parse.$$$ ":" ) + -- Scan.repeat1 Parse.short_ident) + >> (Toplevel.theory o (fold (register_variable_name false)))); +val _ = + Outer_Syntax.command \<^command_keyword>\AOT_register_metavariable_names\ + "Register meta-variable names for type categories." + (Parse.and_list1 ((Parse.short_ident --| Parse.$$$ ":") + -- Scan.repeat1 Parse.short_ident) + >> (Toplevel.theory o (fold (register_variable_name true)))); +val _ = + Outer_Syntax.command \<^command_keyword>\AOT_register_premise_set_names\ + "Register names for premise sets." + (Scan.repeat1 Parse.short_ident + >> (Toplevel.theory o fold + (fn prefix => AOT_PremiseSetPrefix.map (Symtab.update (prefix,()))))); +val _ = + Outer_Syntax.command \<^command_keyword>\AOT_register_type_constraints\ + "Register constraints for term types." + (Parse.and_list1 ((Parse.short_ident --| Parse.$$$ ":") + -- (Parse.typ -- Scan.option Parse.typ)) + >> (Toplevel.theory o fold register_constraint)); + + +fun decode_pos str = + case (Term_Position.decode str) of SOME pos => pos + | _ => raise Fail "expected position" +fun unconstrain_var + (Ast.Appl [Ast.Constant "_constrain", Ast.Variable name, Ast.Variable pos]) = + (name, decode_pos pos) + | unconstrain_var ast = raise Ast.AST + ("Expected position constrained variable.", [ast]) +fun make_constrained_var sx = + (Ast.Appl [Ast.Constant "_constrain", Ast.Variable (Symbol_Pos.implode sx), + Ast.Variable (Term_Position.encode + (Position.range_position (Symbol_Pos.range sx)))]) +fun implode_pos x = (Symbol_Pos.implode_range (Symbol_Pos.range x) x) |> + (fn (x,y) => (x,Position.range_position y)) +fun splitFormulaParts x = x |> unconstrain_var |> Symbol_Pos.explode |> + Scan.finite Symbol_Pos.stopper (Scan.repeat ( + (Scan.one (Symbol_Pos.symbol #> Symbol.is_letter) -- + (((Scan.repeat (Symbol_Pos.$$ "\<^sub>" -- + (Scan.one (Symbol_Pos.symbol #> Symbol.is_digit)) >> + (fn (x,y) => [x,y])) >> List.concat) + -- (Scan.repeat (Symbol_Pos.$$ "'"))) >> (fn (x,y) => x@y))))) +fun parseFormulaParts x = (case splitFormulaParts x of + (parts,[]) => parts |> map (fn (x,y) => implode_pos (x::y)) + | _ => raise Ast.AST ("Expected one or more variable or term names.", [x])) +fun foldAppl const = List.rev #> (fn list => fold (fn a => fn b => + (Ast.mk_appl (Ast.Constant const) [a,b])) (tl list) (hd list)) +fun dropConstraints (Const ("_constrain", _) $ x $ _) = dropConstraints x + | dropConstraints (Const ("_constrainAbs", _) $ x $ _) = dropConstraints x + | dropConstraints (Abs (a, b, x)) = Abs (a, b, dropConstraints x) + | dropConstraints (x$y) = dropConstraints x $ dropConstraints y + | dropConstraints x = x + +local +fun constrain (name, pos) = Ast.mk_appl (Ast.Constant "_constrain") + [Ast.Variable name, Ast.Variable (Term_Position.encode pos)] +in +fun AOT_split_exe_vars [x] = x |> parseFormulaParts |> map constrain |> + map (fn x => Ast.mk_appl (Ast.Constant "_AOT_term_var") [x]) |> + foldAppl "_AOT_exe_args" +fun AOT_split_lambda_args [x] = x |> parseFormulaParts |> map constrain |> + map (fn x => Ast.mk_appl (Ast.Constant "_AOT_var") [x]) |> + foldAppl \<^const_syntax>\Pair\ +fun AOT_check_var [x] = x |> parseFormulaParts |> map constrain |> + (fn [x] => Ast.mk_appl (Ast.Constant "_AOT_var") [x] + | _ => raise Ast.AST ("Expected single variable.", [x])) +end + +fun parseVar unary ctxt [var as Const ("_constrain", _) $ Free (x,_) $ Free _] = + Const ("_constrain", dummyT) $ var $ (case fetchTermConstraint ctxt x unary of + SOME (AOT_MetaVariable,_) => raise Term.TERM + ("Expected variable prefix, but got metavariable prefix.", [var]) + | SOME (AOT_Variable _, constraint) => constraint + | _ => raise Term.TERM ("Unknown variable prefix.", [var])) + | parseVar _ _ var = raise Term.TERM ("Expected constrained free variable.", var) + +fun constrainTrm ctxt forceMeta unary (Free (var, _)) = (fn trm => + case fetchTermConstraint ctxt var unary of + SOME (AOT_MetaVariable,constraint) => + Const ("_constrain", dummyT) $ trm $ constraint + | SOME (AOT_Variable restr, constraint) => + if forceMeta then Const ("_constrain", dummyT) $ trm $ constraint + else Const ("_constrain", dummyT) $ + (Const (\<^const_name>\AOT_term_of_var\, dummyT) $ + (case restr of SOME (_,r) => r $ trm | _ => trm)) $ + constraint + | _ => raise Term.TERM ("Unknown variable or metavariable prefix.", [trm])) + | constrainTrm _ _ _ (Bound _) = (fn var => var) + | constrainTrm _ _ _ trm = raise Term.TERM + ("Expected free or bound variable.", [trm]) +fun isPremiseVar ctxt (Free (var, _)) = + AOT_IsPremiseSetPrefix ctxt (hd (Symbol.explode var)) + | isPremiseVar _ _ = false +fun getVarConstraint ctxt unary (Free (var, _)) = + (case fetchTermConstraint ctxt var unary of + SOME (AOT_MetaVariable,_) => NONE + | SOME (AOT_Variable Rep_term,_) => Option.map fst Rep_term + | _ => NONE) + | getVarConstraint _ _ _ = NONE +fun getVarConstraints ctxt (Const (\<^syntax_const>\_AOT_term_var\, _) $ v) = + (case (getVarConstraint ctxt true (dropConstraints v)) of SOME c => [(c,v)] + | _ => []) + | getVarConstraints ctxt (Const ("_AOT_term_vars", _) $ v) = + (case (getVarConstraint ctxt true (dropConstraints v)) of SOME c => [(c,v)] + | _ => []) + | getVarConstraints _ (Const (\<^syntax_const>\_AOT_verbatim\, _) $ _) = [] + | getVarConstraints ctxt (x $ y) = + getVarConstraints ctxt x @ getVarConstraints ctxt y + | getVarConstraints ctxt (Abs (_,_,z)) = getVarConstraints ctxt z + | getVarConstraints _ _ = [] +fun processFreesForceMeta forceMeta premiseVars ctxt + (Const (\<^syntax_const>\_AOT_term_var\, _) $ v) = ( + if isPremiseVar ctxt (dropConstraints v) + then (dropConstraints v, if List.find (fn x => x = v) premiseVars = NONE + then v::premiseVars else premiseVars) + else (constrainTrm ctxt forceMeta true (dropConstraints v) v, premiseVars)) + | processFreesForceMeta forceMeta premiseVars ctxt + (Const ("_AOT_term_vars", _) $ v) = (if isPremiseVar ctxt (dropConstraints v) + then (v, if List.find (fn x => x = v) premiseVars = NONE + then v::premiseVars else premiseVars) + else (constrainTrm ctxt forceMeta false (dropConstraints v) v, premiseVars) + ) + | processFreesForceMeta _ premiseVars _ + (Const (\<^syntax_const>\_AOT_verbatim\, _) $ v) = (v, premiseVars) + | processFreesForceMeta forceMeta premiseVars ctxt (x $ y) = let + val (x, premiseVars) = processFreesForceMeta forceMeta premiseVars ctxt x + val (y, premiseVars) = processFreesForceMeta forceMeta premiseVars ctxt y + in (x $ y, premiseVars) end + | processFreesForceMeta forceMeta premiseVars ctxt (Abs (x,y,z)) = let + val (z, premiseVars) = processFreesForceMeta forceMeta premiseVars ctxt z + in (Abs (x,y,z), premiseVars) end + | processFreesForceMeta _ premiseVars _ x = (x, premiseVars) +fun processFrees ctxt trm = + (case processFreesForceMeta false [] ctxt trm of (r,[]) => r + | _ => raise Term.TERM ("No premise set expected in term.", [trm])) +fun processFreesAlwaysMeta ctxt trm = + (case processFreesForceMeta true [] ctxt trm of (r,[]) => r + | _ => raise Term.TERM ("No premise set expected in term.", [trm])) +val processFreesAndPremises = processFreesForceMeta false [] + +local +fun makeArgList (Const (\<^syntax_const>\_AOT_exe_args\, _) $ y $ z) = + makeArgList y @ makeArgList z + | makeArgList t = [t] +fun makePairs (x::[]) = x + | makePairs (x::xs) = Const (\<^const_syntax>\Pair\, dummyT) $ x $ makePairs xs +fun makeExeArgs y = makePairs (makeArgList y) +in +fun foldPremises world (Const (\<^syntax_const>\_AOT_premises\, _) $ p1 $ p2) y = + @{const "Pure.imp"} $ (p1 $ world) $ foldPremises world p2 y +| foldPremises world x y = + @{const "Pure.imp"} $ (x $ world) $ + HOLogic.mk_Trueprop (@{const AOT_model_valid_in} $ world $ y) +fun parseExe ctxt [x,y] = (Const (\<^const_syntax>\AOT_exe\, dummyT) $ x $ makeExeArgs y) +fun parseEnc ctxt [x,y] = (Const ("AOT_enc", dummyT) $ makeExeArgs x $ y) +fun parseEquivDef ctxt [lhs,rhs] = + let + val constraints = getVarConstraints ctxt lhs + fun collectConstraints c [] = c + | collectConstraints NONE ((x,y)::xs) = collectConstraints (SOME (x $ y)) xs + | collectConstraints (SOME c) ((x,y)::xs) = + collectConstraints (SOME (Const ("AOT_conj", dummyT) $ c $ (x $ y))) xs + val rhs = (case collectConstraints NONE constraints + of SOME c => Const ("AOT_conj", dummyT) $ c $ rhs + | _ => rhs) + in + HOLogic.mk_Trueprop (\<^const>\AOT_model_equiv_def\ $ processFreesAlwaysMeta ctxt lhs $ + processFreesAlwaysMeta ctxt rhs) + end + | parseEquivDef _ terms = raise Term.TERM ("Expected definition arguments.", terms) +fun parseIdDef ctxt [lhs, rhs] = + let + val lhs = processFreesAlwaysMeta ctxt lhs + val rhs = processFreesAlwaysMeta ctxt rhs + fun add_frees (Free _) frees = frees + | add_frees (Const _) frees = frees + | add_frees (Free _ $ args) frees = Term.add_frees args frees + | add_frees (Const _ $ args) frees = Term.add_frees args frees + | add_frees (args $ args') frees = + Term.add_frees args' (Term.add_frees args frees) + | add_frees trm _ = raise Term.TERM ("Expected definition term.", [trm]) + val lhs' = dropConstraints lhs + val rhs' = dropConstraints rhs + val frees = add_frees lhs' [] + val _ = frees = add_frees rhs' frees orelse + raise Term.TERM ("Invalid free variables on RHS.", [lhs,rhs]) + fun mkabs trm = if frees = [] + then Const (\<^const_name>\case_unit\, dummyT) $ trm + else fold_rev + (fn (s, T) => fn t => Const (\<^const_name>\case_prod\, dummyT) $ + Term.absfree (s, T) t) + (List.rev (tl frees)) (Term.absfree (hd frees) trm) + val lhs_abs = mkabs lhs + val rhs_abs = mkabs rhs + in + (Const ("_constrain", dummyT) $ + Const (\<^const_name>\AOT_model_id_def\, dummyT) $ + (Const (\<^type_syntax>\fun\, dummyT) $ + (Const (\<^type_syntax>\fun\, dummyT) $ Const (\<^type_syntax>\dummy\, dummyT) $ + (getConstraint ctxt false "Term")) $ + (Const (\<^type_syntax>\dummy\, dummyT))) + ) + $ lhs_abs $ rhs_abs + end + | parseIdDef _ terms = raise Term.TERM ("Expected definition arguments.", terms) +end + +fun parseEllipseList constName _ [s,e] = + let + val (start_name, start_pos) = unconstrain_var s + val (end_name, end_pos) = unconstrain_var e + val _ = let val h = hd (Symbol.explode start_name) in + if (h = hd (Symbol.explode end_name)) + then h else raise Ast.AST ("Invalid ellipses.", [s,e]) + end + val name = (Symbol_Pos.explode (start_name, start_pos)) @ + (Symbol_Pos.explode (end_name, end_pos)) + in + Ast.mk_appl (Ast.Constant constName) [make_constrained_var name] + end + | parseEllipseList _ _ _ = raise Fail "Invalid ellipse parsing." + +datatype PrintVarKind = SingleVariable of string | + Ellipses of string*string | Verbatim of string + +fun printVarKind name = let +fun splitFormulaParts x = x |> Symbol.explode |> + Scan.finite Symbol.stopper (Scan.repeat ( + (Scan.one (Symbol.is_letter) -- + (((Scan.repeat ($$ "\<^sub>" -- (Scan.one (Symbol.is_char)) >> + (fn (x,y) => [x,y])) >> List.concat ) + -- (Scan.repeat ($$ "'"))) >> (fn (x,y) => x@y))))) +val parts = splitFormulaParts (Name.clean name) +val isSingleVariableName = case parts of + ([_],[]) => true | _ => false +(* TODO: ellipses handling is very fragile *) +val (isEllipses,s,e) = case parts + of ([(n,s),(m,e)],[]) => (n = m, n^String.concat s, m^String.concat e) + | _ => (false,"","") +in +if isSingleVariableName then SingleVariable name else +if isEllipses then Ellipses (s,e) +else Verbatim name +end + +local +fun addFunct (x,f) g = (x, fn y => g (f y)) +fun unconstrain (Ast.Appl (Ast.Constant "_constrain"::x::tl)) = + addFunct (unconstrain x) (fn x => Ast.Appl (Ast.Constant "_constrain"::x::tl)) + | unconstrain (Ast.Appl (Ast.Constant "_free"::[x])) = + addFunct (unconstrain x) (fn x => Ast.Appl (Ast.Constant "_free"::[x])) + | unconstrain (Ast.Appl (Ast.Constant "_bound"::[x])) = + addFunct (unconstrain x) (fn x => Ast.Appl (Ast.Constant "_bound"::[x])) + | unconstrain (Ast.Appl (Ast.Constant "_var"::[x])) = + addFunct (unconstrain x) (fn x => Ast.Appl (Ast.Constant "_var"::[x])) + | unconstrain trm = (trm, fn x => x) +fun isDefinedConst ctxt name = let + val unmarkedName = Lexicon.unmark {case_class = fn str => NONE, + case_type = fn name => NONE, + case_const = fn name => SOME name, + case_fixed = fn name => NONE, + case_default = fn name => SOME name} name + val cons = Option.mapPartial (fn name => try (Proof_Context.read_const + {proper = true, strict = true} ctxt) name) unmarkedName + val defined = case cons of + SOME cons => + Termtab.defined (AOT_DefinedConstants.get (Proof_Context.theory_of ctxt)) cons + orelse (case cons of Const (name,_) => name = \<^const_name>\AOT_concrete\ + | _ => false) + | _ => false + in defined end +in +val AOT_print_individual_term = (fn ctxt => + (fn [trm as Ast.Appl (Ast.Constant \<^const_syntax>\AOT_term_of_var\::_)] => trm + | [trm as Ast.Appl (Ast.Constant \<^syntax_const>\_AOT_desc\::_)] => trm + | [trm as Ast.Appl (Ast.Constant \<^syntax_const>\_AOT_free_var_ellipse\::_)] => trm + | [trm as Ast.Constant _] => trm + | [trm] => (case unconstrain trm + of (Ast.Variable name,c) => (case printVarKind name + of SingleVariable x => c (Ast.Variable name) + | Ellipses (x,y) => + (Ast.mk_appl (Ast.Constant \<^syntax_const>\_AOT_exe_arg_ellipse\) + [c (Ast.Variable x), c (Ast.Variable y)]) + | _ => Ast.mk_appl (Ast.Constant \<^syntax_const>\_AOT_quoted\) [trm]) + | (Ast.Constant name,c) => + if isDefinedConst ctxt name + then c (Ast.Constant name) + else Ast.mk_appl (Ast.Constant \<^syntax_const>\_AOT_quoted\) [trm] + | (trm' as Ast.Appl (Ast.Constant name::_),c) => + if isDefinedConst ctxt name + then c trm' + else Ast.mk_appl (Ast.Constant \<^syntax_const>\_AOT_quoted\) [trm] + | _ => Ast.mk_appl (Ast.Constant \<^syntax_const>\_AOT_quoted\) [trm]) + | trms => Ast.mk_appl (Ast.Constant \<^syntax_const>\_AOT_quoted\) trms)) +val AOT_print_relation_term = (fn ctxt => + (fn [Ast.Appl [Ast.Constant \<^const_syntax>\AOT_term_of_var\, + Ast.Constant \<^const_syntax>\AOT_concrete\]] => + Ast.Constant \<^syntax_const>\_AOT_concrete\ + | [trm as Ast.Appl (Ast.Constant \<^const_syntax>\AOT_term_of_var\::_)] => + Ast.mk_appl (Ast.Constant \<^syntax_const>\_explicitRelation\) [trm] + | [trm as Ast.Appl (Ast.Constant \<^syntax_const>\_AOT_lambda\::_)] => trm + | [trm as Ast.Appl (Ast.Constant \<^const_syntax>\AOT_lambda\::_)] => trm + | [trm] => (case unconstrain trm + of (Ast.Variable name,c) => + (case printVarKind name + of SingleVariable _ => + (Ast.mk_appl (Ast.Constant \<^syntax_const>\_explicitRelation\) + [c (Ast.Variable name)]) + | _ => Ast.mk_appl (Ast.Constant \<^syntax_const>\_AOT_quoted\) [trm]) + | (Ast.Constant name,c) => + if isDefinedConst ctxt name + then c (Ast.Constant name) + else Ast.mk_appl (Ast.Constant \<^syntax_const>\_AOT_quoted\) [trm] + | (trm' as Ast.Appl (Ast.Constant name::_),c) => + if isDefinedConst ctxt name + then (Ast.mk_appl (Ast.Constant \<^syntax_const>\_explicitRelation\) [c trm']) + else Ast.mk_appl (Ast.Constant \<^syntax_const>\_AOT_quoted\) [trm] + | _ => Ast.mk_appl (Ast.Constant \<^syntax_const>\_AOT_quoted\) [trm]) + | trms => Ast.mk_appl (Ast.Constant \<^syntax_const>\_AOT_quoted\) trms)) +val AOT_print_generic_term = (fn ctxt => + (fn [Ast.Appl [Ast.Constant \<^const_syntax>\AOT_term_of_var\, + Ast.Constant \<^const_syntax>\AOT_concrete\]] => + Ast.Constant \<^syntax_const>\_AOT_concrete\ + | [trm as Ast.Appl (Ast.Constant \<^const_syntax>\AOT_term_of_var\::_)] => +(* Ast.mk_appl (Ast.Constant \<^syntax_const>\_explicitRelation\) [trm] *) + trm + | [trm as Ast.Appl (Ast.Constant \<^syntax_const>\_AOT_desc\::_)] => trm + | [trm as Ast.Appl (Ast.Constant \<^syntax_const>\_AOT_free_var_ellipse\::_)] => trm + | [trm as Ast.Appl (Ast.Constant \<^syntax_const>\_AOT_lambda\::_)] => trm + | [trm as Ast.Appl (Ast.Constant \<^const_syntax>\AOT_lambda\::_)] => trm + | [trm as Ast.Appl (Ast.Constant "_AOT_raw_appl"::_)] => trm + | [trm] => (case unconstrain trm + of (Ast.Variable name,c) => + (case printVarKind name + of SingleVariable x => c (Ast.Variable name) + | Ellipses (x,y) => + (Ast.mk_appl (Ast.Constant \<^syntax_const>\_AOT_exe_arg_ellipse\) + [c (Ast.Variable x), c (Ast.Variable y)]) + | _ => Ast.mk_appl (Ast.Constant \<^syntax_const>\_AOT_quoted\) [trm] + ) + | (Ast.Constant name,c) => + if isDefinedConst ctxt name + then c (Ast.Constant name) + else Ast.mk_appl (Ast.Constant \<^syntax_const>\_AOT_quoted\) [trm] + | (trm' as Ast.Appl (Ast.Constant name::_),c) => + (if isDefinedConst ctxt name + then c trm' + else Ast.mk_appl (Ast.Constant \<^syntax_const>\_AOT_quoted\) [trm]) + | _ => Ast.mk_appl (Ast.Constant \<^syntax_const>\_AOT_quoted\) [trm]) + | trms => Ast.mk_appl (Ast.Constant \<^syntax_const>\_AOT_quoted\) trms)) +end + +fun AOT_preserve_binder_abs_tr' constName syntaxConst + (ellipseConst,includesSyntaxConst) restrConnect = (constName, fn ctxt => fn terms => +let +val term_opt = case terms of Abs (name, T, trm)::trms => +let +val trm = case printVarKind name of SingleVariable name => + let + val optBody = case fetchTermCategory ctxt (hd (Symbol.explode (name))) + of SOME (AOT_Variable _, category) => + let + val (restr, _) = Local_Theory.raw_theory_result + (fn thy => (Symtab.lookup (AOT_Restriction.get thy) category, thy)) ctxt + in + case restr of SOME restr => + (case trm of (Const (c,_) $ x $ trm) => + if (c = restrConnect orelse Lexicon.unmark_const c = restrConnect) + then + if Term.could_unify (Abs ("x", dummyT, x), + Abs ("x", dummyT, Term.betapply (fst restr,(Bound 0)))) then + SOME trm + else + NONE + else NONE | _ => NONE) + | _ => NONE + end + | _ => NONE + val terms = case optBody of SOME trm => Abs (name, T, trm)::trms | _ => trms + in + snd (Syntax_Trans.preserve_binder_abs_tr' constName syntaxConst) ctxt terms + end + | Ellipses (s,e) => + let + val body = Term.subst_bound (Const (\<^syntax_const>\_AOT_free_var_ellipse\, dummyT) $ + Syntax_Trans.mark_bound_body (s,dummyT) $ + Syntax_Trans.mark_bound_body (e,dummyT), + trm) + in + if includesSyntaxConst then + list_comb (Syntax.const ellipseConst $ Syntax_Trans.mark_bound_abs (s,dummyT) $ + Syntax_Trans.mark_bound_abs (e,dummyT) $ body, trms) + else + list_comb (Syntax.const syntaxConst $ + (Syntax.const ellipseConst $ Syntax_Trans.mark_bound_abs (s,dummyT) $ + Syntax_Trans.mark_bound_abs (e,dummyT)) $ body, trms) + end + | Verbatim _ => (* TODO *) + snd (Syntax_Trans.preserve_binder_abs_tr' constName syntaxConst) ctxt terms +in SOME trm end +| _ => NONE +in +case term_opt of SOME trm => trm | _ => +snd (Syntax_Trans.preserve_binder_abs_tr' constName syntaxConst) ctxt terms +end +) + +fun AOT_restricted_binder const connect = + fn ctxt => (fn [a, b] => Ast.mk_appl (Ast.Constant const) [ +let +val b = case a of (Ast.Appl [Ast.Constant "_AOT_var", var]) => ( +case fetchTermCategory ctxt (hd (Symbol.explode (fst (unconstrain_var var)))) +of SOME (AOT_Variable _, category) => + let + val (restr, _) = Local_Theory.raw_theory_result + (fn thy => (Symtab.lookup (AOT_Restriction.get thy) category, thy)) ctxt + in + case restr of SOME _ => Ast.mk_appl (Ast.Constant connect) + [Ast.mk_appl (Ast.mk_appl (Ast.Constant "_AOT_restriction") + [Ast.Constant category]) [a], b] | _ => b end | _ => b) | _ => b +in + Ast.mk_appl (Ast.Constant "_abs") [a,b] +end] | _ => raise Match) + +fun parseDDOT ctxt _ = + let + val trm = Proof_Context.get_fact_single ctxt + (Facts.named (Long_Name.localN ^ Long_Name.separator ^ Auto_Bind.thisN)) + val trm = Thm.concl_of trm + fun mapTerms (Free (x,typ)) = + (case List.rev (String.explode x) of #"_" :: #"_" :: tl => + Free (String.implode (List.rev tl), typ) | _ => Free (x,typ)) + | mapTerms x = x + val trm = Term.map_aterms mapTerms trm + fun readThisRHS (Const ("HOL.Trueprop", _) $ + (Const ("AOT_model.AOT_model_valid_in", _) $ _ $ (Const _ $ _ $ rhs))) = rhs + | readThisRHS _ = raise Term.TERM ("Could not expand ... from term.", [trm]) + in + readThisRHS trm + end \ No newline at end of file diff --git a/thys/AOT/AOT_syntax.thy b/thys/AOT/AOT_syntax.thy new file mode 100644 --- /dev/null +++ b/thys/AOT/AOT_syntax.thy @@ -0,0 +1,648 @@ +(*<*) +theory AOT_syntax + imports AOT_commands + keywords "AOT_register_variable_names" :: thy_decl + and "AOT_register_metavariable_names" :: thy_decl + and "AOT_register_premise_set_names" :: thy_decl + and "AOT_register_type_constraints" :: thy_decl + abbrevs "actually" = "\<^bold>\" + and "neccessarily" = "\" + and "possibly" = "\" + and "the" = "\<^bold>\" + and "lambda" = "[\]" + and "being such that" = "[\ ]" + and "forall" = "\" + and "exists" = "\" + and "equivalent" = "\" + and "not" = "\" + and "implies" = "\" + and "equal" = "=" + and "by definition" = "\<^sub>d\<^sub>f" + and "df" = "\<^sub>d\<^sub>f" + and "denotes" = "\" +begin +(*>*) + +section\Approximation of the Syntax of PLM\ + +locale AOT_meta_syntax +begin +notation AOT_model_valid_in ("\<^bold>[_ \<^bold>\ _\<^bold>]") +notation AOT_model_axiom ("\<^bold>\\<^bold>[_\<^bold>]") +notation AOT_model_act_axiom ("\<^bold>\\<^bold>[_\<^bold>]") +end +locale AOT_no_meta_syntax +begin +no_notation AOT_model_valid_in ("\<^bold>[_ \<^bold>\ _\<^bold>]") +no_notation AOT_model_axiom ("\<^bold>\\<^bold>[_\<^bold>]") +no_notation AOT_model_act_axiom ("\<^bold>\\<^bold>[_\<^bold>]") +end + +consts AOT_denotes :: \'a::AOT_Term \ \\ + AOT_imp :: \[\, \] \ \\ + AOT_not :: \\ \ \\ + AOT_box :: \\ \ \\ + AOT_act :: \\ \ \\ + AOT_forall :: \('a::AOT_Term \ \) \ \\ + AOT_eq :: \'a::AOT_Term \ 'a::AOT_Term \ \\ + AOT_desc :: \('a::AOT_UnaryIndividualTerm \ \) \ 'a\ + AOT_exe :: \<'a::AOT_IndividualTerm> \ 'a \ \\ + AOT_lambda :: \('a::AOT_IndividualTerm \ \) \ <'a>\ + AOT_lambda0 :: \\ \ \\ + AOT_concrete :: \<'a::AOT_UnaryIndividualTerm> AOT_var\ + +nonterminal \\<^sub>s and \ and \0 and \ and exe_arg and exe_args + and lambda_args and desc and free_var and free_vars + and AOT_props and AOT_premises and AOT_world_relative_prop + +syntax "_AOT_process_frees" :: \\ \ \'\ ("_") + "_AOT_verbatim" :: \any \ \\ (\\_\\) + "_AOT_verbatim" :: \any \ \\ (\\_\\) + "_AOT_quoted" :: \\' \ any\ (\\_\\) + "_AOT_quoted" :: \\' \ any\ (\\_\\) + "" :: \\ \ \\ (\'(_')\) + "_AOT_process_frees" :: \\ \ \'\ ("_") + "" :: \\\<^sub>s \ \\ ("_") + "" :: \\ \ \\ ("_") + "" :: \\ \ \\ ("'(_')") + "_AOT_term_var" :: \id_position \ \\ ("_") + "_AOT_term_var" :: \id_position \ \\ ("_") + "_AOT_exe_vars" :: \id_position \ exe_arg\ ("_") + "_AOT_lambda_vars" :: \id_position \ lambda_args\ ("_") + "_AOT_var" :: \id_position \ \\ ("_") + "_AOT_vars" :: \id_position \ any\ + "_AOT_verbatim" :: \any \ \\ (\\_\\) + "_AOT_valid" :: \w \ \' \ bool\ (\[_ \ _]\) + "_AOT_denotes" :: \\ \ \\ (\_\\) + "_AOT_imp" :: \[\, \] \ \\ (infixl \\\ 25) + "_AOT_not" :: \\ \ \\ (\~_\ [50] 50) + "_AOT_not" :: \\ \ \\ (\\_\ [50] 50) + "_AOT_box" :: \\ \ \\ (\\_\ [49] 54) + "_AOT_act" :: \\ \ \\ (\\<^bold>\_\ [49] 54) + "_AOT_all" :: \\ \ \ \ \\ (\\_ _\ [1,40]) +syntax (input) + "_AOT_all_ellipse" + :: \id_position \ id_position \ \ \ \\ (\\_...\_ _\ [1,40]) +syntax (output) + "_AOT_all_ellipse" + :: \id_position \ id_position \ \ \ \\ (\\_...\_'(_')\ [1,40]) +syntax + "_AOT_eq" :: \[\, \] \ \\ (infixl \=\ 50) + "_AOT_desc" :: \\ \ \ \ desc\ ("\<^bold>\__" [1,1000]) + "" :: \desc \ \\<^sub>s\ ("_") + "_AOT_lambda" :: \lambda_args \ \ \ \\ (\[\_ _]\) + "_explicitRelation" :: \\ \ \\ ("[_]") + "" :: \\\<^sub>s \ exe_arg\ ("_") + "" :: \exe_arg \ exe_args\ ("_") + "_AOT_exe_args" :: \exe_arg \ exe_args \ exe_args\ ("__") + "_AOT_exe_arg_ellipse" :: \id_position \ id_position \ exe_arg\ ("_..._") + "_AOT_lambda_arg_ellipse" + :: \id_position \ id_position \ lambda_args\ ("_..._") + "_AOT_term_ellipse" :: \id_position \ id_position \ \\ ("_..._") + "_AOT_exe" :: \\ \ exe_args \ \\ (\__\) + "_AOT_enc" :: \exe_args \ \ \ \\ (\__\) + "_AOT_lambda0" :: \\ \ \0\ (\[\ _]\) + "" :: \\0 \ \\ ("_") + "" :: \\0 \ \\ ("_") + "_AOT_concrete" :: \\\ (\E!\) + "" :: \any \ exe_arg\ ("\_\") + "" :: \desc \ free_var\ ("_") + "" :: \\ \ free_var\ ("_") + "_AOT_appl" :: \id_position \ free_vars \ \\ ("_'{_'}") + "_AOT_appl" :: \id_position \ free_vars \ \\ ("_'{_'}") + "_AOT_appl" :: \id_position \ free_vars \ free_vars\ ("_'{_'}") + "_AOT_appl" :: \id_position \ free_vars \ free_vars\ ("_'{_'}") + "_AOT_term_var" :: \id_position \ free_var\ ("_") + "" :: \any \ free_var\ ("\_\") + "" :: \free_var \ free_vars\ ("_") + "_AOT_args" :: \free_var \ free_vars \ free_vars\ ("_,_") + "_AOT_free_var_ellipse" :: \id_position \ id_position \ free_var\ ("_..._") +syntax "_AOT_premises" + :: \AOT_world_relative_prop \ AOT_premises \ AOT_premises\ (infixr \,\ 3) + "_AOT_world_relative_prop" :: "\ \ AOT_world_relative_prop" ("_") + "" :: "AOT_world_relative_prop \ AOT_premises" ("_") + "_AOT_prop" :: \AOT_world_relative_prop \ AOT_prop\ (\_\) + "" :: \AOT_prop \ AOT_props\ (\_\) + "_AOT_derivable" :: "AOT_premises \ \' \ AOT_prop" (infixl \\<^bold>\\ 2) + "_AOT_nec_derivable" :: "AOT_premises \ \' \ AOT_prop" (infixl \\<^bold>\\<^sub>\\ 2) + "_AOT_theorem" :: "\' \ AOT_prop" (\\<^bold>\ _\) + "_AOT_nec_theorem" :: "\' \ AOT_prop" (\\<^bold>\\<^sub>\ _\) + "_AOT_equiv_def" :: \\ \ \ \ AOT_prop\ (infixl \\\<^sub>d\<^sub>f\ 3) + "_AOT_axiom" :: "\' \ AOT_axiom" (\_\) + "_AOT_act_axiom" :: "\' \ AOT_act_axiom" (\_\) + "_AOT_axiom" :: "\' \ AOT_prop" (\_ \ \\<^sub>\\) + "_AOT_act_axiom" :: "\' \ AOT_prop" (\_ \ \\) + "_AOT_id_def" :: \\ \ \ \ AOT_prop\ (infixl \=\<^sub>d\<^sub>f\ 3) + "_AOT_for_arbitrary" + :: \id_position \ AOT_prop \ AOT_prop\ (\for arbitrary _: _\ [1000,1] 1) +syntax (output) "_lambda_args" :: \any \ patterns \ patterns\ ("__") + +translations + "[w \ \]" => "CONST AOT_model_valid_in w \" + +AOT_syntax_print_translations + "[w \ \]" <= "CONST AOT_model_valid_in w \" + +ML_file AOT_syntax.ML + +AOT_register_type_constraints + Individual: \_::AOT_UnaryIndividualTerm\ \_::AOT_IndividualTerm\ and + Proposition: \ and + Relation: \<_::AOT_IndividualTerm>\ and + Term: \_::AOT_Term\ + +AOT_register_variable_names + Individual: x y z \ \ a b c d and + Proposition: p q r s and + Relation: F G H P Q R S and + Term: \ \ \ \ + +AOT_register_metavariable_names + Individual: \ and + Proposition: \ \ \ \ \ \ \ and + Relation: \ and + Term: \ \ + +AOT_register_premise_set_names \ \ \ + +parse_ast_translation\[ + (\<^syntax_const>\_AOT_var\, K AOT_check_var), + (\<^syntax_const>\_AOT_exe_vars\, K AOT_split_exe_vars), + (\<^syntax_const>\_AOT_lambda_vars\, K AOT_split_lambda_args) +]\ + +translations + "_AOT_denotes \" => "CONST AOT_denotes \" + "_AOT_imp \ \" => "CONST AOT_imp \ \" + "_AOT_not \" => "CONST AOT_not \" + "_AOT_box \" => "CONST AOT_box \" + "_AOT_act \" => "CONST AOT_act \" + "_AOT_eq \ \'" => "CONST AOT_eq \ \'" + "_AOT_lambda0 \" => "CONST AOT_lambda0 \" + "_AOT_concrete" => "CONST AOT_term_of_var (CONST AOT_concrete)" + "_AOT_lambda \ \" => "CONST AOT_lambda (_abs \ \)" + "_explicitRelation \" => "\" + +AOT_syntax_print_translations + "_AOT_lambda (_lambda_args x y) \" <= "CONST AOT_lambda (_abs (_pattern x y) \)" + "_AOT_lambda (_lambda_args x y) \" <= "CONST AOT_lambda (_abs (_patterns x y) \)" + "_AOT_lambda x \" <= "CONST AOT_lambda (_abs x \)" + "_lambda_args x (_lambda_args y z)" <= "_lambda_args x (_patterns y z)" + "_lambda_args (x y z)" <= "_lambda_args (_tuple x (_tuple_arg (_tuple y z)))" + + +AOT_syntax_print_translations + "_AOT_imp \ \" <= "CONST AOT_imp \ \" + "_AOT_not \" <= "CONST AOT_not \" + "_AOT_box \" <= "CONST AOT_box \" + "_AOT_act \" <= "CONST AOT_act \" + "_AOT_all \ \" <= "CONST AOT_forall (_abs \ \)" + "_AOT_all \ \" <= "CONST AOT_forall (\\. \)" + "_AOT_eq \ \'" <= "CONST AOT_eq \ \'" + "_AOT_desc x \" <= "CONST AOT_desc (_abs x \)" + "_AOT_desc x \" <= "CONST AOT_desc (\x. \)" + "_AOT_lambda0 \" <= "CONST AOT_lambda0 \" + "_AOT_concrete" <= "CONST AOT_term_of_var (CONST AOT_concrete)" + +translations + "_AOT_appl \ (_AOT_args a b)" => "_AOT_appl (\ a) b" + "_AOT_appl \ a" => "\ a" + + +parse_translation\ +[ + (\<^syntax_const>\_AOT_var\, parseVar true), + (\<^syntax_const>\_AOT_vars\, parseVar false), + (\<^syntax_const>\_AOT_valid\, fn ctxt => fn [w,x] => + \<^const>\AOT_model_valid_in\ $ w $ x), + (\<^syntax_const>\_AOT_quoted\, fn ctxt => fn [x] => x), + (\<^syntax_const>\_AOT_process_frees\, fn ctxt => fn [x] => processFrees ctxt x), + (\<^syntax_const>\_AOT_world_relative_prop\, fn ctxt => fn [x] => let + val (x, premises) = processFreesAndPremises ctxt x + val (world::formulas) = Variable.variant_frees ctxt [x] + (("v", dummyT)::(map (fn _ => ("\", dummyT)) premises)) + val term = HOLogic.mk_Trueprop + (@{const AOT_model_valid_in} $ Free world $ processFrees ctxt x) + val term = fold (fn (premise,form) => fn trm => + @{const "Pure.imp"} $ + HOLogic.mk_Trueprop + (Const (\<^const_name>\Set.member\, dummyT) $ Free form $ premise) $ + (Term.absfree (Term.dest_Free (dropConstraints premise)) trm $ Free form) + ) (ListPair.zipEq (premises,formulas)) term + val term = fold (fn (form) => fn trm => + Const (\<^const_name>\Pure.all\, dummyT) $ + (Term.absfree form trm) + ) formulas term + val term = Term.absfree world term + in term end), + (\<^syntax_const>\_AOT_prop\, fn ctxt => fn [x] => let + val world = case (AOT_ProofData.get ctxt) of SOME w => w + | _ => raise Fail "Expected world to be stored in the proof state." + in x $ world end), + (\<^syntax_const>\_AOT_theorem\, fn ctxt => fn [x] => + HOLogic.mk_Trueprop (@{const AOT_model_valid_in} $ @{const w\<^sub>0} $ x)), + (\<^syntax_const>\_AOT_axiom\, fn ctxt => fn [x] => + HOLogic.mk_Trueprop (@{const AOT_model_axiom} $ x)), + (\<^syntax_const>\_AOT_act_axiom\, fn ctxt => fn [x] => + HOLogic.mk_Trueprop (@{const AOT_model_act_axiom} $ x)), + (\<^syntax_const>\_AOT_nec_theorem\, fn ctxt => fn [trm] => let + val world = singleton (Variable.variant_frees ctxt [trm]) ("v", @{typ w}) + val trm = HOLogic.mk_Trueprop (@{const AOT_model_valid_in} $ Free world $ trm) + val trm = Term.absfree world trm + val trm = Const (\<^const_name>\Pure.all\, dummyT) $ trm + in trm end), + (\<^syntax_const>\_AOT_derivable\, fn ctxt => fn [x,y] => let + val world = case (AOT_ProofData.get ctxt) of SOME w => w + | _ => raise Fail "Expected world to be stored in the proof state." + in foldPremises world x y end), + (\<^syntax_const>\_AOT_nec_derivable\, fn ctxt => fn [x,y] => let + in Const (\<^const_name>\Pure.all\, dummyT) $ + Abs ("v", dummyT, foldPremises (Bound 0) x y) end), + (\<^syntax_const>\_AOT_for_arbitrary\, fn ctxt => fn [_ $ var $ pos,trm] => let + val trm = Const (\<^const_name>\Pure.all\, dummyT) $ + (Const ("_constrainAbs", dummyT) $ Term.absfree (Term.dest_Free var) trm $ pos) + in trm end), + (\<^syntax_const>\_AOT_equiv_def\, parseEquivDef), + (\<^syntax_const>\_AOT_exe\, parseExe), + (\<^syntax_const>\_AOT_enc\, parseEnc) +] +\ + +parse_ast_translation\ +[ + (\<^syntax_const>\_AOT_exe_arg_ellipse\, parseEllipseList "_AOT_term_vars"), + (\<^syntax_const>\_AOT_lambda_arg_ellipse\, parseEllipseList "_AOT_vars"), + (\<^syntax_const>\_AOT_free_var_ellipse\, parseEllipseList "_AOT_term_vars"), + (\<^syntax_const>\_AOT_term_ellipse\, parseEllipseList "_AOT_term_vars"), + (\<^syntax_const>\_AOT_all_ellipse\, fn ctx => fn [a,b,c] => + Ast.mk_appl (Ast.Constant \<^const_name>\AOT_forall\) [ + Ast.mk_appl (Ast.Constant "_abs") [parseEllipseList "_AOT_vars" ctx [a,b],c] + ]) +] +\ + +syntax (output) + "_AOT_individual_term" :: \'a \ tuple_args\ ("_") + "_AOT_individual_terms" :: \tuple_args \ tuple_args \ tuple_args\ ("__") + "_AOT_relation_term" :: \'a \ \\ + "_AOT_any_term" :: \'a \ \\ + + +print_ast_translation\AOT_syntax_print_ast_translations[ + (\<^syntax_const>\_AOT_individual_term\, AOT_print_individual_term), + (\<^syntax_const>\_AOT_relation_term\, AOT_print_relation_term), + (\<^syntax_const>\_AOT_any_term\, AOT_print_generic_term) +]\ + +AOT_syntax_print_translations + "_AOT_individual_terms (_AOT_individual_term x) (_AOT_individual_terms (_tuple y z))" + <= "_AOT_individual_terms (_tuple x (_tuple_args y z))" + "_AOT_individual_terms (_AOT_individual_term x) (_AOT_individual_term y)" + <= "_AOT_individual_terms (_tuple x (_tuple_arg y))" + "_AOT_individual_terms (_tuple x y)" <= "_AOT_individual_term (_tuple x y)" + "_AOT_exe (_AOT_relation_term \) (_AOT_individual_term \)" <= "CONST AOT_exe \ \" + "_AOT_denotes (_AOT_any_term \)" <= "CONST AOT_denotes \" + +AOT_define AOT_conj :: \[\, \] \ \\ (infixl \&\ 35) \\ & \ \\<^sub>d\<^sub>f \(\ \ \\)\ +declare "AOT_conj"[AOT del, AOT_defs del] +AOT_define AOT_disj :: \[\, \] \ \\ (infixl \\\ 35) \\ \ \ \\<^sub>d\<^sub>f \\ \ \\ +declare "AOT_disj"[AOT del, AOT_defs del] +AOT_define AOT_equiv :: \[\, \] \ \\ (infix \\\ 20) \\ \ \ \\<^sub>d\<^sub>f (\ \ \) & (\ \ \)\ +declare "AOT_equiv"[AOT del, AOT_defs del] +AOT_define AOT_dia :: \\ \ \\ (\\_\ [49] 54) \\\ \\<^sub>d\<^sub>f \\\\\ +declare "AOT_dia"[AOT del, AOT_defs del] + +context AOT_meta_syntax +begin +notation AOT_dia ("\<^bold>\_" [49] 54) +notation AOT_conj (infixl \\<^bold>&\ 35) +notation AOT_disj (infixl \\<^bold>\\ 35) +notation AOT_equiv (infixl \\<^bold>\\ 20) +end +context AOT_no_meta_syntax +begin +no_notation AOT_dia ("\<^bold>\_" [49] 54) +no_notation AOT_conj (infixl \\<^bold>&\ 35) +no_notation AOT_disj (infixl \\<^bold>\\ 35) +no_notation AOT_equiv (infixl \\<^bold>\\ 20) +end + + +print_translation \ +AOT_syntax_print_translations + [ + AOT_preserve_binder_abs_tr' + \<^const_syntax>\AOT_forall\ + \<^syntax_const>\_AOT_all\ + (\<^syntax_const>\_AOT_all_ellipse\, true) + \<^const_name>\AOT_imp\, + AOT_binder_trans @{theory} @{binding "AOT_forall_binder"} \<^syntax_const>\_AOT_all\, + Syntax_Trans.preserve_binder_abs_tr' + \<^const_syntax>\AOT_desc\ + \<^syntax_const>\_AOT_desc\, + AOT_binder_trans @{theory} @{binding "AOT_desc_binder"} \<^syntax_const>\_AOT_desc\, + AOT_preserve_binder_abs_tr' + \<^const_syntax>\AOT_lambda\ + \<^syntax_const>\_AOT_lambda\ + (\<^syntax_const>\_AOT_lambda_arg_ellipse\, false) + \<^const_name>\undefined\, + AOT_binder_trans + @{theory} + @{binding "AOT_lambda_binder"} + \<^syntax_const>\_AOT_lambda\ + ] +\ + +parse_translation\ +[(\<^syntax_const>\_AOT_id_def\, parseIdDef)] +\ + +parse_ast_translation\[ + (\<^syntax_const>\_AOT_all\, + AOT_restricted_binder \<^const_name>\AOT_forall\ \<^const_name>\AOT_imp\), + (\<^syntax_const>\_AOT_desc\, + AOT_restricted_binder \<^const_name>\AOT_desc\ \<^const_name>\AOT_conj\) +]\ + +AOT_define AOT_exists :: \\ \ \ \ \\ \\AOT_exists \\ \\<^sub>d\<^sub>f \\\ \\{\}\ +declare AOT_exists[AOT del, AOT_defs del] +syntax "_AOT_exists" :: \\ \ \ \ \\ ("\_ _" [1,40]) + +AOT_syntax_print_translations + "_AOT_exists \ \" <= "CONST AOT_exists (_abs \ \)" + "_AOT_exists \ \" <= "CONST AOT_exists (\\. \)" + +parse_ast_translation\ +[(\<^syntax_const>\_AOT_exists\, + AOT_restricted_binder \<^const_name>\AOT_exists\ \<^const_name>\AOT_conj\)] +\ + +context AOT_meta_syntax +begin +notation AOT_exists (binder "\<^bold>\" 8) +end +context AOT_no_meta_syntax +begin +no_notation AOT_exists (binder "\<^bold>\" 8) +end + + +syntax (input) + "_AOT_exists_ellipse" :: \id_position \ id_position \ \ \ \\ (\\_...\_ _\ [1,40]) +syntax (output) + "_AOT_exists_ellipse" :: \id_position \ id_position \ \ \ \\ (\\_...\_ '(_')\ [1,40]) +parse_ast_translation\[(\<^syntax_const>\_AOT_exists_ellipse\, fn ctx => fn [a,b,c] => + Ast.mk_appl (Ast.Constant "AOT_exists") + [Ast.mk_appl (Ast.Constant "_abs") [parseEllipseList "_AOT_vars" ctx [a,b],c]])]\ +print_translation\AOT_syntax_print_translations [ + AOT_preserve_binder_abs_tr' + \<^const_syntax>\AOT_exists\ + \<^syntax_const>\_AOT_exists\ + (\<^syntax_const>\_AOT_exists_ellipse\,true) \<^const_name>\AOT_conj\, + AOT_binder_trans + @{theory} + @{binding "AOT_exists_binder"} + \<^syntax_const>\_AOT_exists\ +]\ + + + +syntax "_AOT_DDDOT" :: "\" ("...") +syntax "_AOT_DDDOT" :: "\" ("\") +parse_translation\[(\<^syntax_const>\_AOT_DDDOT\, parseDDOT)]\ + +print_translation\AOT_syntax_print_translations +[(\<^const_syntax>\Pure.all\, fn ctxt => fn [Abs (_, _, + Const (\<^const_syntax>\HOL.Trueprop\, _) $ + (Const (\<^const_syntax>\AOT_model_valid_in\, _) $ Bound 0 $ y))] => let + val y = (Const (\<^syntax_const>\_AOT_process_frees\, dummyT) $ y) + in (Const (\<^syntax_const>\_AOT_nec_theorem\, dummyT) $ y) end +| [p as Abs (name, _, + Const (\<^const_syntax>\HOL.Trueprop\, _) $ + (Const (\<^const_syntax>\AOT_model_valid_in\, _) $ w $ y))] +=> (Const (\<^syntax_const>\_AOT_for_arbitrary\, dummyT) $ + (Const ("_bound", dummyT) $ Free (name, dummyT)) $ + (Term.betapply (p, (Const ("_bound", dummyT) $ Free (name, dummyT))))) +), + + (\<^const_syntax>\AOT_model_valid_in\, fn ctxt => + fn [w as (Const ("_free", _) $ Free (v, _)), y] => let + val is_world = (case (AOT_ProofData.get ctxt) + of SOME (Free (w, _)) => Name.clean w = Name.clean v | _ => false) + val y = (Const (\<^syntax_const>\_AOT_process_frees\, dummyT) $ y) + in if is_world then y else Const (\<^syntax_const>\_AOT_valid\, dummyT) $ w $ y end + | [Const (\<^const_syntax>\w\<^sub>0\, _), y] => let + val y = (Const (\<^syntax_const>\_AOT_process_frees\, dummyT) $ y) + in case (AOT_ProofData.get ctxt) of SOME (Const (\<^const_name>\w\<^sub>0\, _)) => y | + _ => Const (\<^syntax_const>\_AOT_theorem\, dummyT) $ y end + | [Const ("_var", _) $ _, y] => let + val y = (Const (\<^syntax_const>\_AOT_process_frees\, dummyT) $ y) + in Const (\<^syntax_const>\_AOT_nec_theorem\, dummyT) $ y end + ), + (\<^const_syntax>\AOT_model_axiom\, fn ctxt => fn [trm] => + Const (\<^syntax_const>\_AOT_axiom\, dummyT) $ + (Const (\<^syntax_const>\_AOT_process_frees\, dummyT) $ trm)), + (\<^const_syntax>\AOT_model_act_axiom\, fn ctxt => fn [trm] => + Const (\<^syntax_const>\_AOT_axiom\, dummyT) $ + (Const (\<^syntax_const>\_AOT_process_frees\, dummyT) $ trm)), +(\<^syntax_const>\_AOT_process_frees\, fn _ => fn [t] => let + fun mapAppls (x as Const ("_free", _) $ + Free (_, Type ("_ignore_type", [Type ("fun", _)]))) + = (Const ("_AOT_raw_appl", dummyT) $ x) + | mapAppls (x as Const ("_free", _) $ Free (_, Type ("fun", _))) + = (Const ("_AOT_raw_appl", dummyT) $ x) + | mapAppls (x as Const ("_var", _) $ + Var (_, Type ("_ignore_type", [Type ("fun", _)]))) + = (Const ("_AOT_raw_appl", dummyT) $ x) + | mapAppls (x as Const ("_var", _) $ Var (_, Type ("fun", _))) + = (Const ("_AOT_raw_appl", dummyT) $ x) + | mapAppls (x $ y) = mapAppls x $ mapAppls y + | mapAppls (Abs (x,y,z)) = Abs (x,y, mapAppls z) + | mapAppls x = x + in mapAppls t end +) +] +\ + +print_ast_translation\AOT_syntax_print_ast_translations +let +fun handleTermOfVar x kind name = ( +let +val _ = case kind of "_free" => () | "_var" => () | "_bound" => () | _ => raise Match +in + case printVarKind name + of (SingleVariable name) => Ast.Appl [Ast.Constant kind, Ast.Variable name] + | (Ellipses (s, e)) => Ast.Appl [Ast.Constant "_AOT_free_var_ellipse", + Ast.Appl [Ast.Constant kind, Ast.Variable s], + Ast.Appl [Ast.Constant kind, Ast.Variable e] + ] + | Verbatim name => Ast.mk_appl (Ast.Constant "_AOT_quoted") + [Ast.mk_appl (Ast.Constant "_AOT_term_of_var") [x]] +end +) +fun termOfVar ctxt (Ast.Appl [Ast.Constant "_constrain", + x as Ast.Appl [Ast.Constant kind, Ast.Variable name], _]) = termOfVar ctxt x + | termOfVar ctxt (x as Ast.Appl [Ast.Constant kind, Ast.Variable name]) + = handleTermOfVar x kind name + | termOfVar ctxt (x as Ast.Appl [Ast.Constant rep, y]) = ( +let +val (restr,_) = Local_Theory.raw_theory_result (fn thy => ( +let +val restrs = Symtab.dest (AOT_Restriction.get thy) +val restr = List.find (fn (n,(_,Const (c,t))) => ( + c = rep orelse c = Lexicon.unmark_const rep) | _ => false) restrs +in +(restr,thy) +end +)) ctxt +in + case restr of SOME r => Ast.Appl [Ast.Constant (\<^const_syntax>\AOT_term_of_var\), y] + | _ => raise Match +end) + +in +[(\<^const_syntax>\AOT_term_of_var\, fn ctxt => fn [x] => termOfVar ctxt x), +("_AOT_raw_appl", fn ctxt => fn t::a::args => let +fun applyTermOfVar (t as Ast.Appl (Ast.Constant \<^const_syntax>\AOT_term_of_var\::[x])) + = (case try (termOfVar ctxt) x of SOME y => y | _ => t) + | applyTermOfVar y = (case try (termOfVar ctxt) y of SOME x => x | _ => y) +val ts = fold (fn a => fn b => Ast.mk_appl (Ast.Constant \<^syntax_const>\_AOT_args\) + [b,applyTermOfVar a]) args (applyTermOfVar a) +in Ast.mk_appl (Ast.Constant \<^syntax_const>\_AOT_appl\) [t,ts] end)] +end +\ + +context AOT_meta_syntax +begin +notation AOT_denotes ("_\<^bold>\") +notation AOT_imp (infixl "\<^bold>\" 25) +notation AOT_not ("\<^bold>\_" [50] 50) +notation AOT_box ("\<^bold>\_" [49] 54) +notation AOT_act ("\<^bold>\_" [49] 54) +notation AOT_forall (binder "\<^bold>\" 8) +notation AOT_eq (infixl "\<^bold>=" 50) +notation AOT_desc (binder "\<^bold>\" 100) +notation AOT_lambda (binder "\<^bold>\" 100) +notation AOT_lambda0 ("\<^bold>[\<^bold>\ _\<^bold>]") +notation AOT_exe ("\<^bold>\_,_\<^bold>\") +notation AOT_model_equiv_def (infixl "\<^bold>\\<^sub>d\<^sub>f" 10) +notation AOT_model_id_def (infixl "\<^bold>=\<^sub>d\<^sub>f" 10) +notation AOT_term_of_var ("\<^bold>\_\<^bold>\") +notation AOT_concrete ("\<^bold>E\<^bold>!") +end +context AOT_no_meta_syntax +begin +no_notation AOT_denotes ("_\<^bold>\") +no_notation AOT_imp (infixl "\<^bold>\" 25) +no_notation AOT_not ("\<^bold>\_" [50] 50) +no_notation AOT_box ("\<^bold>\_" [49] 54) +no_notation AOT_act ("\<^bold>\_" [49] 54) +no_notation AOT_forall (binder "\<^bold>\" 8) +no_notation AOT_eq (infixl "\<^bold>=" 50) +no_notation AOT_desc (binder "\<^bold>\" 100) +no_notation AOT_lambda (binder "\<^bold>\" 100) +no_notation AOT_lambda0 ("\<^bold>[\<^bold>\ _\<^bold>]") +no_notation AOT_exe ("\<^bold>\_,_\<^bold>\") +no_notation AOT_model_equiv_def (infixl "\<^bold>\\<^sub>d\<^sub>f" 10) +no_notation AOT_model_id_def (infixl "\<^bold>=\<^sub>d\<^sub>f" 10) +no_notation AOT_term_of_var ("\<^bold>\_\<^bold>\") +no_notation AOT_concrete ("\<^bold>E\<^bold>!") +end + +bundle AOT_syntax +begin +declare[[show_AOT_syntax=true, show_question_marks=false, eta_contract=false]] +end + +bundle AOT_no_syntax +begin +declare[[show_AOT_syntax=false, show_question_marks=true]] +end + +parse_translation\ +[("_AOT_restriction", fn ctxt => fn [Const (name,_)] => +let +val (restr, ctxt) = ctxt |> Local_Theory.raw_theory_result + (fn thy => (Option.map fst (Symtab.lookup (AOT_Restriction.get thy) name), thy)) +val restr = case restr of SOME x => x + | _ => raise Fail ("Unknown restricted type: " ^ name) +in restr end +)] +\ + +print_translation\ +AOT_syntax_print_translations +[ + (\<^const_syntax>\AOT_model_equiv_def\, fn ctxt => fn [x,y] => + Const (\<^syntax_const>\_AOT_equiv_def\, dummyT) $ + (Const (\<^syntax_const>\_AOT_process_frees\, dummyT) $ x) $ + (Const (\<^syntax_const>\_AOT_process_frees\, dummyT) $ y)) +] +\ + +print_translation\ +AOT_syntax_print_translations [ +(\<^const_syntax>\AOT_model_id_def\, fn ctxt => + fn [lhs as Abs (lhsName, lhsTy, lhsTrm), rhs as Abs (rhsName, rhsTy, rhsTrm)] => + let + val (name,_) = Name.variant lhsName + (Term.declare_term_names rhsTrm (Term.declare_term_names lhsTrm Name.context)); + val lhs = Term.betapply (lhs, Const ("_bound", dummyT) $ Free (name, lhsTy)) + val rhs = Term.betapply (rhs, Const ("_bound", dummyT) $ Free (name, rhsTy)) + in + Const (\<^const_syntax>\AOT_model_id_def\, dummyT) $ lhs $ rhs + end + | [Const (\<^const_syntax>\case_prod\, _) $ lhs, + Const (\<^const_syntax>\case_prod\, _) $ rhs] => + Const (\<^const_syntax>\AOT_model_id_def\, dummyT) $ lhs $ rhs + | [Const (\<^const_syntax>\case_unit\, _) $ lhs, + Const (\<^const_syntax>\case_unit\, _) $ rhs] => + Const (\<^const_syntax>\AOT_model_id_def\, dummyT) $ lhs $ rhs + | [x, y] => + Const (\<^syntax_const>\_AOT_id_def\, dummyT) $ + (Const (\<^syntax_const>\_AOT_process_frees\, dummyT) $ x) $ + (Const (\<^syntax_const>\_AOT_process_frees\, dummyT) $ y) +)]\ + +text\Special marker for printing propositions as theorems + and for pretty-printing AOT terms.\ +definition print_as_theorem :: \\ \ bool\ where + \print_as_theorem \ \ \ . \v . [v \ \]\ +lemma print_as_theoremI: + assumes \\ v . [v \ \]\ + shows \print_as_theorem \\ + using assms by (simp add: print_as_theorem_def) +attribute_setup print_as_theorem = + \Scan.succeed (Thm.rule_attribute [] + (K (fn thm => thm RS @{thm print_as_theoremI})))\ + "Print as theorem." +print_translation\AOT_syntax_print_translations [ + (\<^const_syntax>\print_as_theorem\, fn ctxt => fn [x] => + (Const (\<^syntax_const>\_AOT_process_frees\, dummyT) $ x)) +]\ + +definition print_term :: \'a \ 'a\ where \print_term \ \ x . x\ +syntax "_AOT_print_term" :: \\ \ 'a\ (\AOT'_TERM[_]\) +translations + "_AOT_print_term \" => "CONST print_term (_AOT_process_frees \)" +print_translation\AOT_syntax_print_translations [ + (\<^const_syntax>\print_term\, fn ctxt => fn [x] => + (Const (\<^syntax_const>\_AOT_process_frees\, dummyT) $ x)) +]\ + + +(* To enable meta syntax: *) +(* interpretation AOT_meta_syntax. *) +(* To disable meta syntax: *) +interpretation AOT_no_meta_syntax. + +(* To enable AOT syntax (takes precedence over meta syntax; + can be done locally using "including" or "include"): *) +unbundle AOT_syntax +(* To disable AOT syntax (restoring meta syntax or no syntax; + can be done locally using "including" or "include"): *) +(* unbundle AOT_no_syntax *) + +(*<*) +end +(*>*) diff --git a/thys/AOT/ROOT b/thys/AOT/ROOT new file mode 100644 --- /dev/null +++ b/thys/AOT/ROOT @@ -0,0 +1,23 @@ +chapter AFP + +session "AOT" (AFP) = "HOL-Cardinals" + + options [show_question_marks = false, timeout = 600, names_short = true] + sessions + "HOL-Cardinals" + "HOL-Eisbach" + theories + AOT_model + AOT_commands + AOT_syntax + AOT_semantics + AOT_Definitions + AOT_Axioms + AOT_PLM + AOT_BasicLogicalObjects + AOT_RestrictedVariables + AOT_ExtendedRelationComprehension + AOT_PossibleWorlds + AOT_NaturalNumbers + AOT_misc + document_files + "root.tex" diff --git a/thys/AOT/document/root.tex b/thys/AOT/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/AOT/document/root.tex @@ -0,0 +1,105 @@ +\documentclass[a4paper,enabledeprecatedfontcommands,abstract=on,twoside=true]{article} +\usepackage{isabelle,isabellesym} +\usepackage{url} +\usepackage{xr} +\usepackage[usenames]{color} +\usepackage{csquotes} +\usepackage{graphicx} +\usepackage{geometry} +\usepackage{epigraph} +\usepackage{tabularx} +\usepackage{array} +\usepackage[all]{nowidow} +\usepackage[stable]{footmisc} +\usepackage[ngerman,english]{babel} +\newcommand{\embeddedstyle}[1]{{\color{blue}#1}} +\newcommand{\bigbox}{} +\newcommand{\linelabel}[1]{} + +\setcounter{secnumdepth}{2} +\setcounter{tocdepth}{1} + +% further packages required for unusual symbols (see also +% isabellesym.sty), use only when needed + +\usepackage{amssymb} + %for \, \, \, \, \, \, + %\, \, \, \, \, + %\, \, \ +\usepackage{amsthm} + +\usepackage{amsmath} + +%\usepackage{eurosym} + %for \ + +%\usepackage[only,bigsqcap]{stmaryrd} + %for \ + +%\usepackage{eufrak} + %for \ ... \, \ ... \ (also included in amssymb) + +%\usepackage{textcomp} + %for \, \, \, \, \, + %\ + +% this should be the last package used +\usepackage{pdfsetup} +\definecolor{linkcolor}{rgb}{0,0,0} +\definecolor{citecolor}{rgb}{0,0,0} +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +% for uniform font size +\renewcommand{\isastyleminor}{\isastyle} + +\begin{document} + +% sane default for proof documents +\parindent 0pt\parskip 0.5ex + + +\begin{abstract} +We utilize and extend the method of \emph{shallow semantic embeddings} (SSEs) in classical higher-order logic (HOL) to construct a custom theorem proving environment +for \emph{abstract objects theory} (AOT) on the basis of Isabelle/HOL. + +SSEs are a means for universal logical reasoning by translating a target logic to HOL using a representation of its semantics. +AOT is a foundational metaphysical theory, developed by Edward Zalta, that explains the objects presupposed by the sciences as \emph{abstract objects} that reify property patterns. In particular, AOT aspires to provide a philosphically grounded basis for the construction and analysis of the objects of mathematics. + +We can support this claim by verifying Uri Nodelman's and Edward Zalta's reconstruction of Frege's theorem: we can confirm that the Dedekind-Peano postulates for natural numbers are consistently derivable in AOT using Frege's method. Furthermore, we can suggest and discuss generalizations and variants of the construction and can thereby provide theoretical insights into, and contribute to the philosophical justification of, the construction. + +In the process, we can demonstrate that our method allows for a nearly transparent exchange of results between traditional pen-and-paper-based reasoning and the computerized implementation, which in turn can retain the automation mechanisms available for Isabelle/HOL. + +During our work, we could significantly contribute to the evolution of our target theory itself, while simultaneously solving the technical challenge of using an SSE to implement a theory that is based on +logical foundations that significantly differ from the meta-logic HOL. + +In general, our results demonstrate the fruitfulness of the practice of Computational Metaphysics, i.e. the application of computational methods to metaphysical questions and theories. + +A full description of this formalization including references can be found at \url{http://dx.doi.org/10.17169/refubium-35141}. + +The version of Principia Logico-Metaphysica (PLM) implemented in this formalization +can be found at \url{http://mally.stanford.edu/principia-2021-10-13.pdf}, while +the latest version of PLM is available at \url{http://mally.stanford.edu/principia.pdf}. +\end{abstract} + +\cleardoublepage + +\tableofcontents + +\cleardoublepage + +% generated text of all theories + +\newgeometry{margin=1in} + +\input{session} + +\restoregeometry + +\end{document} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff --git a/thys/ROOTS b/thys/ROOTS --- a/thys/ROOTS +++ b/thys/ROOTS @@ -1,724 +1,725 @@ ADS_Functor AI_Planning_Languages_Semantics AODV 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 +AOT 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 Balog_Szemeredi_Gowers BDD BD_Security_Compositional BNF_CC BNF_Operations BTree Banach_Steinhaus Belief_Revision Bell_Numbers_Spivey BenOr_Kozen_Reif Berlekamp_Zassenhaus Bernoulli Bertrands_Postulate Bicategory BinarySearchTree 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 CISC-Kernel Cook_Levin CRYSTALS-Kyber CRDT CSP_RefTK 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 Cauchy Cayley_Hamilton Certification_Monads Chandy_Lamport CHERI-C_Memory_Model 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 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 CoreC++ Core_DOM Core_SC_DOM Correctness_Algebras Cotangent_PFD_Formula Count_Complex_Roots CryptHOL CryptoBasedCompositionalProperties Cubic_Quartic_Equations DFS_Framework DOM_Components 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 Digit_Expansions Dijkstra_Shortest_Path Diophantine_Eqns_Lin_Hom Dirichlet_L Dirichlet_Series DiscretePricing Discrete_Summation DiskPaxos Dominance_CHK DPRM_Theorem DynamicArchitectures Dynamic_Tables E_Transcendental Echelon_Form EdmondsKarp_Maxflow Efficient-Mergesort Elliptic_Curves_Group_Law Encodability_Process_Calculi Epistemic_Logic Equivalence_Relation_Enumeration Ergodic_Theory Error_Function Euler_MacLaurin Euler_Partition Eval_FO Example-Submission 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 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 Flow_Networks Floyd_Warshall Flyspeck-Tame FocusStreamsCaseStudies Forcing Formal_Puiseux_Series Formal_SSA Formula_Derivatives Foundation_of_geometry Fourier FO_Theory_Rewriting 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 GoedelGod Goedel_HFSet_Semantic Goedel_HFSet_Semanticless Goedel_Incompleteness Goodstein_Lambda GraphMarkingIBP Graph_Saturation Graph_Theory 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 Hoare_Time Hood_Melville_Queue HotelKeyCards Huffman Hybrid_Logic Hybrid_Multi_Lane_Spatial_Logic Hybrid_Systems_VCs HyperCTL Hyperdual IEEE_Floating_Point IFC_Tracking IMAP-CRDT IMO2019 IMP2 IMP2_Binary_Heap IMP_Compiler IMP_Compiler_Reuse 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 Isabelle_C Isabelle_Marries_Dirac Isabelle_Meta_Model IsaNet 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 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 Lowe_Ontological_Argument Lower_Semicontinuous Lp LP_Duality 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 Maximum_Segment_Sum Max-Card-Matching Median_Method Median_Of_Medians_Selection Menger Mereology Mersenne_Primes Metalogic_ProofChecker MiniML MiniSail Minimal_SSA Minkowskis_Theorem Minsky_Machines Modal_Logics_for_NTS Modular_Assembly_Kit_Security Modular_arithmetic_LLL_and_HNF_algorithms Monad_Memo_DP Monad_Normalisation MonoBoolTranAlgebra MonoidalCategory Monomorphic_Monad MuchAdoAboutTwo Multiset_Ordering_NPC Multi_Party_Computation Multirelations 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 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 Package_logic PAL PAPP_Impossibility PCF PLM POPLmark-deBruijn PSemigroupsConvolution Padic_Ints Padic_Field Pairing_Heap Paraconsistency Parity_Game Partial_Function_MR Partial_Order_Reduction Password_Authentication_Protocol Pell Perfect-Number-Thm Perron_Frobenius Physical_Quantities Pi_Calculus Pi_Transcendental Planarity_Certificates Pluennecke_Ruzsa_Inequality Poincare_Bendixson Poincare_Disc 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 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 QHLProver QR_Decomposition Quantales 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_Disjoint_Set_Forests Relational_Forests Relational_Method Relational_Minimum_Spanning_Trees Relational_Paths 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 SC_DOM_Components SDS_Impossibility SIFPL SIFUM_Type_Systems SPARCv8 Safe_Distance Safe_OCL Safe_Range_RC Saturation_Framework Saturation_Framework_Extensions Sauer_Shelah_Lemma SCC_Bloemen_Sequential Schutz_Spacetime 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_Firewall Simplex Simplicial_complexes_and_boolean_functions SimplifiedOntologicalArgument Skew_Heap Skip_Lists Slicing Sliding_Window_Algorithm Smith_Normal_Form Smooth_Manifolds Sophomores_Dream Solidity Sort_Encodings Source_Coding_Theorem SpecCheck Special_Function_Bounds Splay_Tree Sqrt_Babylonian Stable_Matching Stalnaker_Logic 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 Strong_Security Sturm_Sequences Sturm_Tarski Stuttering_Equivalence Subresultants Subset_Boolean_Algebras SumSquares Sunflowers SuperCalc Surprise_Paradox Symmetric_Polynomials Syntax_Independent_Logic Synthetic_Completeness Szemeredi_Regularity Szpilrajn TESL_Language TLA Tail_Recursive_Functions Tarskis_Geometry Taylor_Models Three_Circles Timed_Automata Topological_Semantics Topology TortoiseHare Transcendence_Series_Hancl_Rucki Transformer_Semantics Transition_Systems_and_Automata Transitive-Closure Transitive-Closure-II Transitive_Models Treaps Tree-Automata Tree_Decomposition Triangle Trie Turans_Graph_Theorem Twelvefold_Way 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 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 VYDRA_MDL 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 Zeta_3_Irrational Zeta_Function pGCL