Theory AOT_PLM
theory AOT_PLM
imports AOT_Axioms
begin
section‹The Deductive System PLM›
text‹\label{PLM: 9}›
unbundle AOT_no_atp
subsection‹Primitive Rule of PLM: Modus Ponens›
text‹\label{PLM: 9.1}›
AOT_theorem "modus-ponens":
assumes ‹φ› and ‹φ → ψ›
shows ‹ψ›
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 ‹❙⊢⇩□ φ›
shows ‹❙⊢ φ›
using assms by simp
AOT_theorem "vdash-properties:1[1]":
assumes ‹φ ∈ Λ›
shows ‹❙⊢ φ›
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 ‹φ ∈ Λ⇩□›
shows ‹❙⊢⇩□ φ›
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 ‹❙⊢⇩□ φ›
shows ‹Γ ❙⊢ φ›
using assms by blast
AOT_theorem "vdash-properties:5":
assumes ‹Γ⇩1 ❙⊢ φ› and ‹Γ⇩2 ❙⊢ φ → ψ›
shows ‹Γ⇩1, Γ⇩2 ❙⊢ ψ›
using MP assms by blast
AOT_theorem "vdash-properties:6":
assumes ‹φ› and ‹φ → ψ›
shows ‹ψ›
using MP assms by blast
AOT_theorem "vdash-properties:8":
assumes ‹Γ ❙⊢ φ› and ‹φ ❙⊢ ψ›
shows ‹Γ ❙⊢ ψ›
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 ‹∀α φ{α}›
using assms by (metis AOT_var_of_term_inverse AOT_sem_denotes AOT_sem_forall)
lemmas GEN = "rule-gen"
AOT_theorem "RN[prem]":
assumes ‹Γ ❙⊢⇩□ φ›
shows ‹□Γ ❙⊢⇩□ □φ›
by (meson AOT_sem_box assms image_iff)
AOT_theorem RN:
assumes ‹❙⊢⇩□ φ›
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 ‹φ ≡⇩d⇩f ψ›
shows ‹φ → ψ›
using assms
by (auto simp: assms AOT_model_axiomI AOT_model_equiv_def AOT_sem_imp)
AOT_axiom "df-rules-formulas[2]":
assumes ‹φ ≡⇩d⇩f ψ›
shows ‹ψ → φ›
using assms
by (auto simp: AOT_model_axiomI AOT_model_equiv_def AOT_sem_imp)
AOT_theorem "df-rules-formulas[3]":
assumes ‹φ ≡⇩d⇩f ψ›
shows ‹φ → ψ›
using "df-rules-formulas[1]"[axiom_inst, OF assms].
AOT_theorem "df-rules-formulas[4]":
assumes ‹φ ≡⇩d⇩f ψ›
shows ‹ψ → φ›
using "df-rules-formulas[2]"[axiom_inst, OF assms].
AOT_axiom "df-rules-terms[1]":
assumes ‹τ{α⇩1...α⇩n} =⇩d⇩f σ{α⇩1...α⇩n}›
shows ‹(σ{τ⇩1...τ⇩n}↓ → τ{τ⇩1...τ⇩n} = σ{τ⇩1...τ⇩n}) &
(¬σ{τ⇩1...τ⇩n}↓ → ¬τ{τ⇩1...τ⇩n}↓)›
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 ‹τ =⇩d⇩f σ›
shows ‹(σ↓ → τ = σ) & (¬σ↓ → ¬τ↓)›
by (metis "df-rules-terms[1]" case_unit_Unity assms)
AOT_theorem "df-rules-terms[3]":
assumes ‹τ{α⇩1...α⇩n} =⇩d⇩f σ{α⇩1...α⇩n}›
shows ‹(σ{τ⇩1...τ⇩n}↓ → τ{τ⇩1...τ⇩n} = σ{τ⇩1...τ⇩n}) &
(¬σ{τ⇩1...τ⇩n}↓ → ¬τ{τ⇩1...τ⇩n}↓)›
using "df-rules-terms[1]"[axiom_inst, OF assms].
AOT_theorem "df-rules-terms[4]":
assumes ‹τ =⇩d⇩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 ‹φ ❙⊢ ψ›
shows ‹φ → ψ›
using assms by (simp add: AOT_sem_imp)
lemmas CP = "deduction-theorem"
lemmas "→I" = "deduction-theorem"
AOT_theorem "ded-thm-cor:1":
assumes ‹Γ⇩1 ❙⊢ φ → ψ› and ‹Γ⇩2 ❙⊢ ψ → χ›
shows ‹Γ⇩1, Γ⇩2 ❙⊢ φ → χ›
using "→E" "→I" assms by blast
AOT_theorem "ded-thm-cor:2":
assumes ‹Γ⇩1 ❙⊢ φ → (ψ → χ)› and ‹Γ⇩2 ❙⊢ ψ›
shows ‹Γ⇩1, Γ⇩2 ❙⊢ φ → χ›
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 ‹¬φ ❙⊢ ¬ψ› and ‹¬φ ❙⊢ ψ›
shows ‹φ›
using "→I" "¬¬E" MT(2) assms by blast
AOT_theorem "reductio-aa:2":
assumes ‹φ ❙⊢ ¬ψ› and ‹φ ❙⊢ ψ›
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 ‹¬φ ❙⊢ ψ & ¬ψ›
shows ‹φ›
using "&E" "∨E"(3) "∨I"(2) RAA(2) assms by blast
AOT_theorem "raa-cor:2":
assumes ‹φ ❙⊢ ψ & ¬ψ›
shows ‹¬φ›
using "raa-cor:1" assms by blast
AOT_theorem "raa-cor:3":
assumes ‹φ› and ‹¬ψ ❙⊢ ¬φ›
shows ‹ψ›
using RAA assms by blast
AOT_theorem "raa-cor:4":
assumes ‹¬φ› and ‹¬ψ ❙⊢ φ›
shows ‹ψ›
using RAA assms by blast
AOT_theorem "raa-cor:5":
assumes ‹φ› and ‹ψ ❙⊢ ¬φ›
shows ‹¬ψ›
using RAA assms by blast
AOT_theorem "raa-cor:6":
assumes ‹¬φ› and ‹ψ ❙⊢ φ›
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 ‹φ ≡⇩d⇩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 ‹φ ≡⇩d⇩f ψ› and ‹φ›
shows ‹ψ›
using "≡Df" "≡E"(1) assms by blast
lemmas "≡⇩d⇩fE" = "rule-eq-df:2"
AOT_theorem "rule-eq-df:3":
assumes ‹φ ≡⇩d⇩f ψ› and ‹ψ›
shows ‹φ›
using "≡Df" "≡E"(2) assms by blast
lemmas "≡⇩d⇩fI" = "rule-eq-df:3"
AOT_theorem "df-simplify:1":
assumes ‹φ ≡ (ψ & χ)› and ‹ψ›
shows ‹φ ≡ χ›
by (metis "&E"(2) "&I" "≡E"(1, 2) "≡I" "→I" assms)
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 ‹φ{[λν⇩1...ν⇩n ψ{ν⇩1...ν⇩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} → φ{[λν⇩1...ν⇩n ψ{ν⇩1...ν⇩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
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."
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": ‹(∀α⇩1...∀α⇩n(φ{α⇩1...α⇩n})) → φ{α⇩1...α⇩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":
‹(∀α⇩1...∀α⇩n (φ{α⇩1...α⇩n} → ψ{α⇩1...α⇩n})) →
((∀α⇩1...∀α⇩n φ{α⇩1...α⇩n}) → (∀α⇩1...∀α⇩n ψ{α⇩1...α⇩n}))›
using "cqt:3"[axiom_inst] by auto
AOT_theorem "cqt-basic:15":
‹(∀α⇩1...∀α⇩n (φ → ψ{α⇩1...α⇩n})) → (φ → (∀α⇩1...∀α⇩n ψ{α⇩1...α⇩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 "≡⇩d⇩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 ‹φ{[λν⇩1...ν⇩n ψ{ν⇩1...ν⇩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 β: φ{β} ❙⊢ ψ› and ‹∃α φ{α}›
shows ‹ψ›
by (metis (no_types, lifting) "≡⇩d⇩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))
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 "≡⇩d⇩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
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 "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩fE"[OF "identity:4"[of φ ψ]] "&E" by blast
qed
}
next
AOT_modally_strict {
fix φ ψ
AOT_show ‹φ = ψ → ψ↓›
proof(rule "→I")
AOT_assume ‹φ = ψ›
AOT_thus ‹ψ↓› using "≡⇩d⇩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 τ⇩1 τ⇩1' :: 'a and τ⇩2 τ⇩2' :: 'b
AOT_assume ‹«(τ⇩1, τ⇩2)» = «(τ⇩1', τ⇩2')»›
AOT_hence ‹(τ⇩1 = τ⇩1') & (τ⇩2 = τ⇩2')› by (metis "≡⇩d⇩fE" tuple_identity_1)
AOT_hence ‹τ⇩1↓› and ‹τ⇩2↓›
using "t=t-proper:1" "&E" "vdash-properties:10" by blast+
AOT_thus ‹«(τ⇩1, τ⇩2)»↓› by (metis "≡⇩d⇩fI" "&I" tuple_denotes)
qed
}
next
AOT_modally_strict {
fix τ τ' :: ‹'a×'b›
AOT_show ‹τ = τ' → τ'↓›
proof (induct τ; induct τ'; rule "→I")
fix τ⇩1 τ⇩1' :: 'a and τ⇩2 τ⇩2' :: 'b
AOT_assume ‹«(τ⇩1, τ⇩2)» = «(τ⇩1', τ⇩2')»›
AOT_hence ‹(τ⇩1 = τ⇩1') & (τ⇩2 = τ⇩2')› by (metis "≡⇩d⇩fE" tuple_identity_1)
AOT_hence ‹τ⇩1'↓› and ‹τ⇩2'↓›
using "t=t-proper:2" "&E" "vdash-properties:10" by blast+
AOT_thus ‹«(τ⇩1', τ⇩2')»↓› by (metis "≡⇩d⇩fI" "&I" tuple_denotes)
qed
}
qed
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⇩1...∀x⇩n ([Π]x⇩1...x⇩n ≡ [Π']x⇩1...x⇩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⇩1...∀x⇩n ([F]x⇩1...x⇩n ≡ [F]x⇩1...x⇩n)) →
□∀x⇩1...∀x⇩n ([F]x⇩1...x⇩n ≡ [G]x⇩1...x⇩n)))›
apply (rule GEN)+ using "l-identity"[axiom_inst] by force
ultimately AOT_have ‹Π = Π' → ((□∀x⇩1...∀x⇩n ([Π]x⇩1...x⇩n ≡ [Π]x⇩1...x⇩n)) →
□∀x⇩1...∀x⇩n ([Π]x⇩1...x⇩n ≡ [Π']x⇩1...x⇩n))›
using "∀E"(1) by blast
AOT_hence ‹(□∀x⇩1...∀x⇩n ([Π]x⇩1...x⇩n ≡ [Π]x⇩1...x⇩n)) →
□∀x⇩1...∀x⇩n ([Π]x⇩1...x⇩n ≡ [Π']x⇩1...x⇩n)›
using assumption "→E" by blast
moreover AOT_have ‹□∀x⇩1...∀x⇩n ([Π]x⇩1...x⇩n ≡ [Π]x⇩1...x⇩n)›
by (simp add: RN "oth-class-taut:3:a" "universal-cor")
ultimately AOT_show ‹□∀x⇩1...∀x⇩n ([Π]x⇩1...x⇩n ≡ [Π']x⇩1...x⇩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⇩1([λx [F]xy⇩1] = [λx [G]xy⇩1] & [λx [F]y⇩1x] = [λx [G]y⇩1x])›
proof -
AOT_have ‹F = G ≡ F↓ & G↓ &
∀y⇩1([λx [F]xy⇩1] = [λx [G]xy⇩1] & [λx [F]y⇩1x] = [λx [G]y⇩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⇩1∀y⇩2([λx [F]xy⇩1y⇩2] = [λx [G]xy⇩1y⇩2] &
[λx [F]y⇩1xy⇩2] = [λx [G]y⇩1xy⇩2] &
[λx [F]y⇩1y⇩2x] = [λx [G]y⇩1y⇩2x])›
proof -
AOT_have ‹F = G ≡ F↓ & G↓ & ∀y⇩1∀y⇩2([λx [F]xy⇩1y⇩2] = [λx [G]xy⇩1y⇩2] &
[λx [F]y⇩1xy⇩2] = [λx [G]y⇩1xy⇩2] &
[λx [F]y⇩1y⇩2x] = [λx [G]y⇩1y⇩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⇩1∀y⇩2∀y⇩3([λx [F]xy⇩1y⇩2y⇩3] = [λx [G]xy⇩1y⇩2y⇩3] &
[λx [F]y⇩1xy⇩2y⇩3] = [λx [G]y⇩1xy⇩2y⇩3] &
[λx [F]y⇩1y⇩2xy⇩3] = [λx [G]y⇩1y⇩2xy⇩3] &
[λx [F]y⇩1y⇩2y⇩3x] = [λx [G]y⇩1y⇩2y⇩3x])›
proof -
AOT_have ‹F = G ≡ F↓ & G↓ & ∀y⇩1∀y⇩2∀y⇩3([λx [F]xy⇩1y⇩2y⇩3] = [λx [G]xy⇩1y⇩2y⇩3] &
[λx [F]y⇩1xy⇩2y⇩3] = [λx [G]y⇩1xy⇩2y⇩3] &
[λx [F]y⇩1y⇩2xy⇩3] = [λx [G]y⇩1y⇩2xy⇩3] &
[λx [F]y⇩1y⇩2y⇩3x] = [λx [G]y⇩1y⇩2y⇩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⇩1...∀x⇩n «AOT_sem_proj_id x⇩1x⇩n (λ τ . «[F]τ») (λ τ . «[G]τ»)»›
proof -
AOT_have ‹F = G ≡ F↓ & G↓ &
∀x⇩1...∀x⇩n «AOT_sem_proj_id x⇩1x⇩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⇩1...x⇩n [F]x⇩1...x⇩n] = F›
by (simp add: "lambda-predicates:3"[axiom_inst])
AOT_have ‹[λx⇩1...x⇩n [F]x⇩1...x⇩n]↓›
by "cqt:2[lambda]"
AOT_hence ‹[λx⇩1...x⇩n [F]x⇩1...x⇩n] = [λx⇩1...x⇩n [F]x⇩1...x⇩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 τ⇩1 :: 'a and τ⇩2 :: 'b
AOT_assume ‹«(τ⇩1,τ⇩2)»↓›
AOT_hence ‹τ⇩1↓› and ‹τ⇩2↓›
using "≡⇩d⇩fE" "&E" tuple_denotes by blast+
AOT_hence ‹τ⇩1 = τ⇩1› and ‹τ⇩2 = τ⇩2›
using "id-eq:1"[unvarify α] by blast+
AOT_thus ‹«(τ⇩1, τ⇩2)» = «(τ⇩1, τ⇩2)»›
by (metis "≡⇩d⇩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 "[λν⇩1...ν⇩n φ{ν⇩1...ν⇩n}] = [λν⇩1...ν⇩n φ{ν⇩1...ν⇩n}]"
proof -
AOT_have ‹∀α (α = α)›
by (rule GEN) (metis "id-eq:1")
moreover AOT_have ‹[λν⇩1...ν⇩n φ{ν⇩1...ν⇩n}]↓›
using assms by (rule "cqt:2[lambda]"[axiom_inst])
ultimately AOT_show ‹[λν⇩1...ν⇩n φ{ν⇩1...ν⇩n}] = [λν⇩1...ν⇩n φ{ν⇩1...ν⇩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 ‹τ{α⇩1...α⇩n} =⇩d⇩f σ{α⇩1...α⇩n}› and ‹σ{τ⇩1...τ⇩n}↓›
shows ‹τ{τ⇩1...τ⇩n} = σ{τ⇩1...τ⇩n}›
proof -
AOT_have ‹σ{τ⇩1...τ⇩n}↓ → τ{τ⇩1...τ⇩n} = σ{τ⇩1...τ⇩n}›
using "df-rules-terms[3]" assms(1) "&E" by blast
AOT_thus ‹τ{τ⇩1...τ⇩n} = σ{τ⇩1...τ⇩n}›
using assms(2) "→E" by blast
qed
AOT_theorem "rule-id-df:1[zero]":
assumes ‹τ =⇩d⇩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 ‹τ{α⇩1...α⇩n} =⇩d⇩f σ{α⇩1...α⇩n}› and ‹σ{τ⇩1...τ⇩n}↓› and ‹φ{τ{τ⇩1...τ⇩n}}›
shows ‹φ{σ{τ⇩1...τ⇩n}}›
proof -
AOT_have ‹τ{τ⇩1...τ⇩n} = σ{τ⇩1...τ⇩n}› using "rule-id-df:1" assms(1,2) by blast
AOT_thus ‹φ{σ{τ⇩1...τ⇩n}}› using assms(3) "rule=E" by blast
qed
AOT_theorem "rule-id-df:2:a[2]":
assumes ‹τ{«(α⇩1,α⇩2)»} =⇩d⇩f σ{«(α⇩1,α⇩2)»}›
and ‹σ{«(τ⇩1,τ⇩2)»}↓›
and ‹φ{τ{«(τ⇩1,τ⇩2)»}}›
shows ‹φ{σ{«(τ⇩1::'a::AOT_Term_id_2,τ⇩2::'b::AOT_Term_id_2)»}}›
proof -
AOT_have ‹τ{«(τ⇩1,τ⇩2)»} = σ{«(τ⇩1,τ⇩2)»}›
using "rule-id-df:1" assms(1,2) by auto
AOT_thus ‹φ{σ{«(τ⇩1,τ⇩2)»}}› using assms(3) "rule=E" by blast
qed
AOT_theorem "rule-id-df:2:a[zero]":
assumes ‹τ =⇩d⇩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 "=⇩d⇩fE" = "rule-id-df:2:a" "rule-id-df:2:a[zero]"
AOT_theorem "rule-id-df:2:b":
assumes ‹τ{α⇩1...α⇩n} =⇩d⇩f σ{α⇩1...α⇩n}› and ‹σ{τ⇩1...τ⇩n}↓› and ‹φ{σ{τ⇩1...τ⇩n}}›
shows ‹φ{τ{τ⇩1...τ⇩n}}›
proof -
AOT_have ‹τ{τ⇩1...τ⇩n} = σ{τ⇩1...τ⇩n}›
using "rule-id-df:1" assms(1,2) by blast
AOT_hence ‹σ{τ⇩1...τ⇩n} = τ{τ⇩1...τ⇩n}›
using "rule=E" "=I"(1) "t=t-proper:1" "→E" by fast
AOT_thus ‹φ{τ{τ⇩1...τ⇩n}}› using assms(3) "rule=E" by blast
qed
AOT_theorem "rule-id-df:2:b[2]":
assumes ‹τ{«(α⇩1,α⇩2)»} =⇩d⇩f σ{«(α⇩1,α⇩2)»}›
and ‹σ{«(τ⇩1,τ⇩2)»}↓›
and ‹φ{σ{«(τ⇩1,τ⇩2)»}}›
shows ‹φ{τ{«(τ⇩1::'a::AOT_Term_id_2,τ⇩2::'b::AOT_Term_id_2)»}}›
proof -
AOT_have ‹τ{«(τ⇩1,τ⇩2)»} = σ{«(τ⇩1,τ⇩2)»}›
using "=I"(1) "rule-id-df:2:a[2]" RAA(1) assms(1,2) "→I" by metis
AOT_hence ‹σ{«(τ⇩1,τ⇩2)»} = τ{«(τ⇩1,τ⇩2)»}›
using "rule=E" "=I"(1) "t=t-proper:1" "→E" by fast
AOT_thus ‹φ{τ{«(τ⇩1,τ⇩2)»}}› using assms(3) "rule=E" by blast
qed
AOT_theorem "rule-id-df:2:b[zero]":
assumes ‹τ =⇩d⇩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 "=⇩d⇩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 ‹∃β (β = [λν⇩1...ν⇩n φ{ν⇩1...ν⇩n}])›
by (meson "=I"(3) assms "cqt:2[lambda]"[axiom_inst] "existential:1")
AOT_theorem "free-thms:4[rel]":
‹([Π]κ⇩1...κ⇩n ∨ κ⇩1...κ⇩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]":
‹([Π]κ⇩1...κ⇩n ∨ κ⇩1...κ⇩n[Π]) → ∃β⇩1...∃β⇩n (β⇩1...β⇩n = κ⇩1...κ⇩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]":
‹([Π]κ⇩1κ⇩2 ∨ κ⇩1κ⇩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]":
‹([Π]κ⇩1κ⇩2 ∨ κ⇩1κ⇩2[Π]) → ∃β (β = κ⇩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]":
‹([Π]κ⇩1κ⇩2 ∨ κ⇩1κ⇩2[Π]) → ∃β (β = κ⇩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]":
‹([Π]κ⇩1κ⇩2κ⇩3 ∨ κ⇩1κ⇩2κ⇩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]":
‹([Π]κ⇩1κ⇩2κ⇩3 ∨ κ⇩1κ⇩2κ⇩3[Π]) → ∃β (β = κ⇩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]":
‹([Π]κ⇩1κ⇩2κ⇩3 ∨ κ⇩1κ⇩2κ⇩3[Π]) → ∃β (β = κ⇩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]":
‹([Π]κ⇩1κ⇩2κ⇩3 ∨ κ⇩1κ⇩2κ⇩3[Π]) → ∃β (β = κ⇩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]":
‹([Π]κ⇩1κ⇩2κ⇩3κ⇩4 ∨ κ⇩1κ⇩2κ⇩3κ⇩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]":
‹([Π]κ⇩1κ⇩2κ⇩3κ⇩4 ∨ κ⇩1κ⇩2κ⇩3κ⇩4[Π]) → ∃β (β = κ⇩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]":
‹([Π]κ⇩1κ⇩2κ⇩3κ⇩4 ∨ κ⇩1κ⇩2κ⇩3κ⇩4[Π]) → ∃β (β = κ⇩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]":
‹([Π]κ⇩1κ⇩2κ⇩3κ⇩4 ∨ κ⇩1κ⇩2κ⇩3κ⇩4[Π]) → ∃β (β = κ⇩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]":
‹([Π]κ⇩1κ⇩2κ⇩3κ⇩4 ∨ κ⇩1κ⇩2κ⇩3κ⇩4[Π]) → ∃β (β = κ⇩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
AOT_theorem "term-out:4":
‹(φ{β} & ∀α(φ{α} → α = β)) ≡ ∀α(φ{α} ≡ α = β)›
using "term-out:3" .
AOT_define AOT_exists_unique :: ‹α ⇒ φ ⇒ φ› "uniqueness:1":
‹«AOT_exists_unique φ» ≡⇩d⇩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 "❙∃❙!" 20)
end
context AOT_no_meta_syntax
begin
no_notation AOT_exists_unique (binder "❙∃❙!" 20)
end
AOT_theorem "uniqueness:2": ‹∃!α φ{α} ≡ ∃α∀β(φ{β} ≡ β = α)›
proof(rule "≡I"; rule "→I")
AOT_assume ‹∃!α φ{α}›
AOT_hence ‹∃α (φ{α} & ∀β (φ{β} → β = α))›
using "uniqueness:1" "≡⇩d⇩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" "≡⇩d⇩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" "≡⇩d⇩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" "≡⇩d⇩fI" "∃I" by fast
qed
subsection‹The Theory of Actuality and Descriptions›
text‹\label{PLM: 9.8}›
AOT_theorem "act-cond": ‹❙𝒜(φ → ψ) → (❙𝒜φ → ❙𝒜ψ)›
using "→I" "≡E"(1) "logic-actual-nec:2"[axiom_inst] by blast
AOT_theorem "nec-imp-act": ‹□φ → ❙𝒜φ›
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": ‹❙𝒜(❙𝒜φ → φ)›
using "→I" "≡E"(2) "logic-actual-nec:2"[axiom_inst]
"logic-actual-nec:4"[axiom_inst] by blast
AOT_theorem "act-conj-act:2": ‹❙𝒜(φ → ❙𝒜φ)›
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": ‹(❙𝒜φ & ❙𝒜ψ) → ❙𝒜(φ & ψ)›
proof -
AOT_have ‹□(φ → (ψ → (φ & ψ)))›
by (rule RN) (fact Adjunction)
AOT_hence ‹❙𝒜(φ → (ψ → (φ & ψ)))›
using "nec-imp-act" "→E" by blast
AOT_hence ‹❙𝒜φ → ❙𝒜(ψ → (φ & ψ))›
using "act-cond" "→E" by blast
moreover AOT_have ‹❙𝒜(ψ → (φ & ψ)) → (❙𝒜ψ → ❙𝒜(φ & ψ))›
by (fact "act-cond")
ultimately AOT_have ‹❙𝒜φ → (❙𝒜ψ → ❙𝒜(φ & ψ))›
using "→I" "→E" by metis
AOT_thus ‹(❙𝒜φ & ❙𝒜ψ) → ❙𝒜(φ & ψ)›
by (metis Importation "→E")
qed
AOT_theorem "act-conj-act:4": ‹❙𝒜(❙𝒜φ ≡ φ)›
proof -
AOT_have ‹(❙𝒜(❙𝒜φ → φ) & ❙𝒜(φ → ❙𝒜φ)) → ❙𝒜((❙𝒜φ → φ) & (φ → ❙𝒜φ))›
by (fact "act-conj-act:3")
moreover AOT_have ‹❙𝒜(❙𝒜φ → φ) & ❙𝒜(φ → ❙𝒜φ)›
using "&I" "act-conj-act:1" "act-conj-act:2" by simp
ultimately AOT_have ζ: ‹❙𝒜((❙𝒜φ → φ) & (φ → ❙𝒜φ))›
using "→E" by blast
AOT_have ‹❙𝒜(((❙𝒜φ → φ) & (φ → ❙𝒜φ)) → (❙𝒜φ ≡ φ))›
using "conventions:3"[THEN "df-rules-formulas[2]",
THEN act_closure, axiom_inst] by blast
AOT_hence ‹❙𝒜((❙𝒜φ → φ) & (φ → ❙𝒜φ)) → ❙𝒜(❙𝒜φ ≡ φ)›
using "act-cond" "→E" by blast
AOT_thus ‹❙𝒜(❙𝒜φ ≡ φ)› using ζ "→E" by blast
qed
inductive arbitrary_actualization for φ where
‹arbitrary_actualization φ «❙𝒜φ»›
| ‹arbitrary_actualization φ «❙𝒜ψ»› 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(❙𝒜φ ≡ φ, ❙𝒜(❙𝒜φ ≡ φ))›
using AOT_PLM.arbitrary_actualization.intros by metis
AOT_have ‹ARBITRARY_ACTUALIZATION(❙𝒜φ ≡ φ, ❙𝒜❙𝒜(❙𝒜φ ≡ φ))›
using AOT_PLM.arbitrary_actualization.intros by metis
AOT_have ‹ARBITRARY_ACTUALIZATION(❙𝒜φ ≡ φ, ❙𝒜❙𝒜❙𝒜(❙𝒜φ ≡ φ))›
using AOT_PLM.arbitrary_actualization.intros by metis
}
end
AOT_theorem "closure-act:1":
assumes ‹ARBITRARY_ACTUALIZATION(❙𝒜φ ≡ φ, ψ)›
shows ‹ψ›
using assms proof(induct)
case 1
AOT_show ‹❙𝒜(❙𝒜φ ≡ φ)›
by (simp add: "act-conj-act:4")
next
case (2 ψ)
AOT_thus ‹❙𝒜ψ›
by (metis arbitrary_actualization.simps "≡E"(1)
"logic-actual-nec:4"[axiom_inst])
qed
AOT_theorem "closure-act:2": ‹∀α ❙𝒜(❙𝒜φ{α} ≡ φ{α})›
by (simp add: "act-conj-act:4" "∀I")
AOT_theorem "closure-act:3": ‹❙𝒜∀α ❙𝒜(❙𝒜φ{α} ≡ φ{α})›
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": ‹❙𝒜∀α⇩1...∀α⇩n ❙𝒜(❙𝒜φ{α⇩1...α⇩n} ≡ φ{α⇩1...α⇩n})›
using "closure-act:3" .
AOT_act_theorem "RA[1]":
assumes ‹❙⊢ φ›
shows ‹❙⊢ ❙𝒜φ›
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 ‹❙⊢⇩□ φ›
shows ‹❙⊢⇩□ ❙𝒜φ›
using RN assms "nec-imp-act" "vdash-properties:5" by blast
AOT_theorem "RA[3]":
assumes ‹Γ ❙⊢⇩□ φ›
shows ‹❙𝒜Γ ❙⊢⇩□ ❙𝒜φ›
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.›
oops
AOT_act_theorem "ANeg:1": ‹¬❙𝒜φ ≡ ¬φ›
by (simp add: "RA[1]" "contraposition:1[1]" "deduction-theorem"
"≡I" "logic-actual"[act_axiom_inst])
AOT_act_theorem "ANeg:2": ‹¬❙𝒜¬φ ≡ φ›
using "ANeg:1" "≡I" "≡E"(5) "useful-tautologies:1"
"useful-tautologies:2" by blast
AOT_theorem "Act-Basic:1": ‹❙𝒜φ ∨ ❙𝒜¬φ›
by (meson "∨I"(1,2) "≡E"(2) "logic-actual-nec:1"[axiom_inst] "raa-cor:1")
AOT_theorem "Act-Basic:2": ‹❙𝒜(φ & ψ) ≡ (❙𝒜φ & ❙𝒜ψ)›
proof (rule "≡I"; rule "→I")
AOT_assume ‹❙𝒜(φ & ψ)›
moreover AOT_have ‹❙𝒜((φ & ψ) → φ)›
by (simp add: "RA[2]" "Conjunction Simplification"(1))
moreover AOT_have ‹❙𝒜((φ & ψ) → ψ)›
by (simp add: "RA[2]" "Conjunction Simplification"(2))
ultimately AOT_show ‹❙𝒜φ & ❙𝒜ψ›
using "act-cond"[THEN "→E", THEN "→E"] "&I" by metis
next
AOT_assume ‹❙𝒜φ & ❙𝒜ψ›
AOT_thus ‹❙𝒜(φ & ψ)›
using "act-conj-act:3" "vdash-properties:6" by blast
qed
AOT_theorem "Act-Basic:3": ‹❙𝒜(φ ≡ ψ) ≡ (❙𝒜(φ → ψ) & ❙𝒜(ψ → φ))›
proof (rule "≡I"; rule "→I")
AOT_assume ‹❙𝒜(φ ≡ ψ)›
moreover AOT_have ‹❙𝒜((φ ≡ ψ) → (φ → ψ))›
by (simp add: "RA[2]" "deduction-theorem" "≡E"(1))
moreover AOT_have ‹❙𝒜((φ ≡ ψ) → (ψ → φ))›
by (simp add: "RA[2]" "deduction-theorem" "≡E"(2))
ultimately AOT_show ‹❙𝒜(φ → ψ) & ❙𝒜(ψ → φ)›
using "act-cond"[THEN "→E", THEN "→E"] "&I" by metis
next
AOT_assume ‹❙𝒜(φ → ψ) & ❙𝒜(ψ → φ)›
AOT_hence ‹❙𝒜((φ → ψ) & (ψ → φ))›
by (metis "act-conj-act:3" "vdash-properties:10")
moreover AOT_have ‹❙𝒜(((φ → ψ) & (ψ → φ)) → (φ ≡ ψ))›
by (simp add: "conventions:3" "RA[2]" "df-rules-formulas[2]"
"vdash-properties:1[2]")
ultimately AOT_show ‹❙𝒜(φ ≡ ψ)›
using "act-cond"[THEN "→E", THEN "→E"] by metis
qed
AOT_theorem "Act-Basic:4": ‹(❙𝒜(φ → ψ) & ❙𝒜(ψ → φ)) ≡ (❙𝒜φ ≡ ❙𝒜ψ)›
proof (rule "≡I"; rule "→I")
AOT_assume 0: ‹❙𝒜(φ → ψ) & ❙𝒜(ψ → φ)›
AOT_show ‹❙𝒜φ ≡ ❙𝒜ψ›
using 0 "&E" "act-cond"[THEN "→E", THEN "→E"] "≡I" "→I" by metis
next
AOT_assume ‹❙𝒜φ ≡ ❙𝒜ψ›
AOT_thus ‹❙𝒜(φ → ψ) & ❙𝒜(ψ → φ)›
by (metis "→I" "logic-actual-nec:2"[axiom_inst] "≡E"(1,2) "&I")
qed
AOT_theorem "Act-Basic:5": ‹❙𝒜(φ ≡ ψ) ≡ (❙𝒜φ ≡ ❙𝒜ψ)›
using "Act-Basic:3" "Act-Basic:4" "≡E"(5) by blast
AOT_theorem "Act-Basic:6": ‹❙𝒜φ ≡ □❙𝒜φ›
by (simp add: "≡I" "qml:2"[axiom_inst] "qml-act:1"[axiom_inst])
AOT_theorem "Act-Basic:7": ‹❙𝒜□φ → □❙𝒜φ›
by (metis "Act-Basic:6" "→I" "→E" "≡E"(1,2) "nec-imp-act"
"qml-act:2"[axiom_inst])
AOT_theorem "Act-Basic:8": ‹□φ → □❙𝒜φ›
using "Hypothetical Syllogism" "nec-imp-act" "qml-act:1"[axiom_inst] by blast
AOT_theorem "Act-Basic:9": ‹❙𝒜(φ ∨ ψ) ≡ (❙𝒜φ ∨ ❙𝒜ψ)›
proof (rule "≡I"; rule "→I")
AOT_assume ‹❙𝒜(φ ∨ ψ)›
AOT_thus ‹❙𝒜φ ∨ ❙𝒜ψ›
proof (rule "raa-cor:3")
AOT_assume ‹¬(❙𝒜φ ∨ ❙𝒜ψ)›
AOT_hence ‹¬❙𝒜φ & ¬❙𝒜ψ›
by (metis "≡E"(1) "oth-class-taut:5:d")
AOT_hence ‹❙𝒜¬φ & ❙𝒜¬ψ›
using "logic-actual-nec:1"[axiom_inst, THEN "≡E"(2)] "&E" "&I" by metis
AOT_hence ‹❙𝒜(¬φ & ¬ψ)›
using "≡E" "Act-Basic:2" by metis
moreover AOT_have ‹❙𝒜((¬φ & ¬ψ) ≡ ¬(φ ∨ ψ))›
using "RA[2]" "≡E"(6) "oth-class-taut:3:a" "oth-class-taut:5:d" by blast
moreover AOT_have ‹❙𝒜(¬φ & ¬ψ) ≡ ❙𝒜(¬(φ ∨ ψ))›
using calculation(2) by (metis "Act-Basic:5" "≡E"(1))
ultimately AOT_have ‹❙𝒜(¬(φ ∨ ψ))› using "≡E" by blast
AOT_thus ‹¬❙𝒜(φ ∨ ψ)›
using "logic-actual-nec:1"[axiom_inst, THEN "≡E"(1)] by auto
qed
next
AOT_assume ‹❙𝒜φ ∨ ❙𝒜ψ›
AOT_thus ‹❙𝒜(φ ∨ ψ)›
by (meson "RA[2]" "act-cond" "∨I"(1) "∨E"(1) "Disjunction Addition"(1,2))
qed
AOT_theorem "Act-Basic:10": ‹❙𝒜∃α φ{α} ≡ ∃α ❙𝒜φ{α}›
proof -
AOT_have θ: ‹¬❙𝒜∀α ¬φ{α} ≡ ¬∀α ❙𝒜¬φ{α}›
by (rule "oth-class-taut:4:b"[THEN "≡E"(1)])
(metis "logic-actual-nec:3"[axiom_inst])
AOT_have ξ: ‹¬∀α ❙𝒜¬φ{α} ≡ ¬∀α ¬❙𝒜φ{α}›
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 ‹❙𝒜(∃α φ{α}) ≡ ❙𝒜(¬∀α ¬φ{α})›
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 ‹… ≡ ¬❙𝒜∀α ¬φ{α}›
by (simp add: "logic-actual-nec:1" "vdash-properties:1[2]")
also AOT_have ‹… ≡ ¬∀α ❙𝒜 ¬φ{α}› using θ by blast
also AOT_have ‹… ≡ ¬∀α ¬❙𝒜 φ{α}› using ξ by blast
also AOT_have ‹… ≡ ∃α ❙𝒜 φ{α}›
using "conventions:4"[THEN "≡Df"] by (metis "≡E"(6) "oth-class-taut:3:a")
finally AOT_show ‹❙𝒜∃α φ{α} ≡ ∃α ❙𝒜φ{α}› .
qed
AOT_theorem "Act-Basic:11":
‹❙𝒜∀α(φ{α} ≡ ψ{α}) ≡ ∀α(❙𝒜φ{α} ≡ ❙𝒜ψ{α})›
proof(rule "≡I"; rule "→I")
AOT_assume ‹❙𝒜∀α(φ{α} ≡ ψ{α})›
AOT_hence ‹∀α❙𝒜(φ{α} ≡ ψ{α})›
using "logic-actual-nec:3"[axiom_inst, THEN "≡E"(1)] by blast
AOT_hence ‹❙𝒜(φ{α} ≡ ψ{α})› for α using "∀E" by blast
AOT_hence ‹❙𝒜φ{α} ≡ ❙𝒜ψ{α}› for α by (metis "Act-Basic:5" "≡E"(1))
AOT_thus ‹∀α(❙𝒜φ{α} ≡ ❙𝒜ψ{α})› by (rule "∀I")
next
AOT_assume ‹∀α(❙𝒜φ{α} ≡ ❙𝒜ψ{α})›
AOT_hence ‹❙𝒜φ{α} ≡ ❙𝒜ψ{α}› for α using "∀E" by blast
AOT_hence ‹❙𝒜(φ{α} ≡ ψ{α})› for α by (metis "Act-Basic:5" "≡E"(2))
AOT_hence ‹∀α ❙𝒜(φ{α} ≡ ψ{α})› by (rule "∀I")
AOT_thus ‹❙𝒜∀α(φ{α} ≡ ψ{α})›
using "logic-actual-nec:3"[axiom_inst, THEN "≡E"(2)] by fast
qed
AOT_act_theorem "act-quant-uniq":
‹∀β(❙𝒜φ{β} ≡ β = α) ≡ ∀β(φ{β} ≡ β = α)›
proof(rule "≡I"; rule "→I")
AOT_assume ‹∀β(❙𝒜φ{β} ≡ β = α)›
AOT_hence ‹❙𝒜φ{β} ≡ β = α› 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 ‹❙𝒜φ{β} ≡ β = α› for β
using "≡I" "→I" "RA[1]" "≡E"(1,2) "logic-actual"[act_axiom_inst] "→E"
by metis
AOT_thus ‹∀β(❙𝒜φ{β} ≡ β = α)› by (rule "∀I")
qed
AOT_act_theorem "fund-cont-desc": ‹x = ❙ιx(φ{x}) ≡ ∀z(φ{z} ≡ z = x)›
using descriptions[axiom_inst] "act-quant-uniq" "≡E"(5) by fast
AOT_act_theorem hintikka: ‹x = ❙ι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":
‹ψ{❙ιx φ{x}} ≡ ∃x(φ{x} & ∀z(φ{z} → z = x) & ψ{x})›
proof -
AOT_have b: ‹∀x (x = ❙ιx φ{x} ≡ (φ{x} & ∀z(φ{z} → z = x)))›
using hintikka "∀I" by fast
show ?thesis
proof(rule "≡I"; rule "→I")
AOT_assume c: ‹ψ{❙ιx φ{x}}›
AOT_hence d: ‹❙ιx φ{x}↓›
using ψ_denotes_asm by blast
AOT_hence ‹∃y (y = ❙ιx φ{x})›
by (metis "rule=I:1" "existential:1")
then AOT_obtain a where a_def: ‹a = ❙ιx φ{x}›
using "instantiation"[rotated] by blast
moreover AOT_have ‹a = ❙ι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 = ❙ιx φ{x} → ❙ιx φ{x} = a›
by (rule "∀E"(1)[where τ="«❙ι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 = ❙ι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 = ❙ιx φ{x}› using h "≡E" by blast
AOT_thus ‹ψ{❙ι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": ‹❙ιx φ{x}↓ ≡ ∃!x φ{x}›
proof(rule "≡I"; rule "→I")
AOT_assume ‹❙ιx φ{x}↓›
AOT_hence ‹∃y (y = ❙ιx φ{x})› by (metis "rule=I:1" "existential:1")
then AOT_obtain a where ‹a = ❙ι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 "≡⇩d⇩fI"] by blast
next
AOT_assume ‹∃!x φ{x}›
AOT_hence ‹∃x (φ{x} & ∀z (φ{z} → z = x))›
using "uniqueness:1"[THEN "≡⇩d⇩fE"] by blast
then AOT_obtain b where ‹φ{b} & ∀z (φ{z} → z = b)›
using "instantiation"[rotated] by blast
AOT_hence ‹b = ❙ιx φ{x}›
using hintikka "≡E" by blast
AOT_thus ‹❙ιx φ{x}↓›
by (metis "t=t-proper:2" "vdash-properties:6")
qed
AOT_act_theorem "!-exists:2": ‹∃y(y=❙ιx φ{x}) ≡ ∃!x φ{x}›
using "!-exists:1" "free-thms:1" "≡E"(6) by blast
AOT_act_theorem "y-in:1": ‹x = ❙ιx φ{x} → φ{x}›
using "&E"(1) "→I" hintikka "≡E"(1) by blast
AOT_act_theorem "y-in:2": ‹z = ❙ιx φ{x} → φ{z}› using "y-in:1".
AOT_act_theorem "y-in:3": ‹❙ιx φ{x}↓ → φ{❙ιx φ{x}}›
proof(rule "→I")
AOT_assume ‹❙ιx φ{x}↓›
AOT_hence ‹∃y (y = ❙ιx φ{x})›
by (metis "rule=I:1" "existential:1")
then AOT_obtain a where ‹a = ❙ιx φ{x}›
using "instantiation"[rotated] by blast
moreover AOT_have ‹φ{a}›
using calculation hintikka "≡E"(1) "&E" by blast
ultimately AOT_show ‹φ{❙ιx φ{x}}› using "rule=E" by blast
qed
AOT_act_theorem "y-in:4": ‹∃y (y = ❙ιx φ{x}) → φ{❙ιx φ{x}}›
using "y-in:3"[THEN "→E"] "free-thms:1"[THEN "≡E"(2)] "→I" by blast
AOT_theorem "act-quant-nec":
‹∀β (❙𝒜φ{β} ≡ β = α) ≡ ∀β(❙𝒜❙𝒜φ{β} ≡ β = α)›
proof(rule "≡I"; rule "→I")
AOT_assume ‹∀β (❙𝒜φ{β} ≡ β = α)›
AOT_hence ‹❙𝒜φ{β} ≡ β = α› for β using "∀E" by blast
AOT_hence ‹❙𝒜❙𝒜φ{β} ≡ β = α› for β
by (metis "Act-Basic:5" "act-conj-act:4" "≡E"(1) "≡E"(5))
AOT_thus ‹∀β(❙𝒜❙𝒜φ{β} ≡ β = α)›
by (rule "∀I")
next
AOT_assume ‹∀β(❙𝒜❙𝒜φ{β} ≡ β = α)›
AOT_hence ‹❙𝒜❙𝒜φ{β} ≡ β = α› for β using "∀E" by blast
AOT_hence ‹❙𝒜φ{β} ≡ β = α› for β
by (metis "Act-Basic:5" "act-conj-act:4" "≡E"(1) "≡E"(6))
AOT_thus ‹∀β (❙𝒜φ{β} ≡ β = α)›
by (rule "∀I")
qed
AOT_theorem "equi-desc-descA:1": ‹x = ❙ιx φ{x} ≡ x = ❙ιx(❙𝒜φ{x})›
proof -
AOT_have ‹x = ❙ιx φ{x} ≡ ∀z (❙𝒜φ{z} ≡ z = x)›
using descriptions[axiom_inst] by blast
also AOT_have ‹... ≡ ∀z (❙𝒜❙𝒜φ{z} ≡ z = x)›
proof(rule "≡I"; rule "→I"; rule "∀I")
AOT_assume ‹∀z (❙𝒜φ{z} ≡ z = x)›
AOT_hence ‹❙𝒜φ{a} ≡ a = x› for a
using "∀E" by blast
AOT_thus ‹❙𝒜❙𝒜φ{a} ≡ a = x› for a
by (metis "Act-Basic:5" "act-conj-act:4" "≡E"(1) "≡E"(5))
next
AOT_assume ‹∀z (❙𝒜❙𝒜φ{z} ≡ z = x)›
AOT_hence ‹❙𝒜❙𝒜φ{a} ≡ a = x› for a
using "∀E" by blast
AOT_thus ‹❙𝒜φ{a} ≡ a = x› for a
by (metis "Act-Basic:5" "act-conj-act:4" "≡E"(1) "≡E"(6))
qed
also AOT_have ‹... ≡ x = ❙ιx(❙𝒜φ{x})›
using "Commutativity of ≡"[THEN "≡E"(1)] descriptions[axiom_inst] by fast
finally show ?thesis .
qed
AOT_theorem "equi-desc-descA:2": ‹❙ιx φ{x}↓ → ❙ιx φ{x} = ❙ιx(❙𝒜φ{x})›
proof(rule "→I")
AOT_assume ‹❙ιx φ{x}↓›
AOT_hence ‹∃y (y = ❙ιx φ{x})›
by (metis "rule=I:1" "existential:1")
then AOT_obtain a where ‹a = ❙ιx φ{x}›
using "instantiation"[rotated] by blast
moreover AOT_have ‹a = ❙ιx(❙𝒜φ{x})›
using calculation "equi-desc-descA:1"[THEN "≡E"(1)] by blast
ultimately AOT_show ‹❙ιx φ{x} = ❙ιx(❙𝒜φ{x})›
using "rule=E" by fast
qed
AOT_theorem "nec-hintikka-scheme":
‹x = ❙ιx φ{x} ≡ ❙𝒜φ{x} & ∀z(❙𝒜φ{z} → z = x)›
proof -
AOT_have ‹x = ❙ιx φ{x} ≡ ∀z(❙𝒜φ{z} ≡ z = x)›
using descriptions[axiom_inst] by blast
also AOT_have ‹… ≡ (❙𝒜φ{x} & ∀z(❙𝒜φ{z} → z = x))›
using "Commutativity of ≡"[THEN "≡E"(1)] "term-out:3" by fast
finally show ?thesis.
qed
AOT_theorem "equiv-desc-eq:1":
‹❙𝒜∀x(φ{x} ≡ ψ{x}) → ∀x (x = ❙ιx φ{x} ≡ x = ❙ιx ψ{x})›
proof(rule "→I"; rule "∀I")
fix β
AOT_assume ‹❙𝒜∀x(φ{x} ≡ ψ{x})›
AOT_hence ‹❙𝒜(φ{x} ≡ ψ{x})› for x
using "logic-actual-nec:3"[axiom_inst, THEN "≡E"(1)] "∀E"(2) by blast
AOT_hence 0: ‹❙𝒜φ{x} ≡ ❙𝒜ψ{x}› for x
by (metis "Act-Basic:5" "≡E"(1))
AOT_have ‹β = ❙ιx φ{x} ≡ ❙𝒜φ{β} & ∀z(❙𝒜φ{z} → z = β)›
using "nec-hintikka-scheme" by blast
also AOT_have ‹... ≡ ❙𝒜ψ{β} & ∀z(❙𝒜ψ{z} → z = β)›
proof (rule "≡I"; rule "→I")
AOT_assume 1: ‹❙𝒜φ{β} & ∀z(❙𝒜φ{z} → z = β)›
AOT_hence ‹❙𝒜φ{z} → z = β› for z
using "&E" "∀E" by blast
AOT_hence ‹❙𝒜ψ{z} → z = β› for z
using 0 "≡E" "→I" "→E" by metis
AOT_hence ‹∀z(❙𝒜ψ{z} → z = β)›
using "∀I" by fast
moreover AOT_have ‹❙𝒜ψ{β}›
using "&E" 0[THEN "≡E"(1)] 1 by blast
ultimately AOT_show ‹❙𝒜ψ{β} & ∀z(❙𝒜ψ{z} → z = β)›
using "&I" by blast
next
AOT_assume 1: ‹❙𝒜ψ{β} & ∀z(❙𝒜ψ{z} → z = β)›
AOT_hence ‹❙𝒜ψ{z} → z = β› for z
using "&E" "∀E" by blast
AOT_hence ‹❙𝒜φ{z} → z = β› for z
using 0 "≡E" "→I" "→E" by metis
AOT_hence ‹∀z(❙𝒜φ{z} → z = β)›
using "∀I" by fast
moreover AOT_have ‹❙𝒜φ{β}›
using "&E" 0[THEN "≡E"(2)] 1 by blast
ultimately AOT_show ‹❙𝒜φ{β} & ∀z(❙𝒜φ{z} → z = β)›
using "&I" by blast
qed
also AOT_have ‹... ≡ β = ❙ιx ψ{x}›
using "Commutativity of ≡"[THEN "≡E"(1)] "nec-hintikka-scheme" by blast
finally AOT_show ‹β = ❙ιx φ{x} ≡ β = ❙ιx ψ{x}› .
qed
AOT_theorem "equiv-desc-eq:2":
‹❙ιx φ{x}↓ & ❙𝒜∀x(φ{x} ≡ ψ{x}) → ❙ιx φ{x} = ❙ιx ψ{x}›
proof(rule "→I")
AOT_assume ‹❙ιx φ{x}↓ & ❙𝒜∀x(φ{x} ≡ ψ{x})›
AOT_hence 0: ‹∃y (y = ❙ιx φ{x})› and
1: ‹∀x (x = ❙ιx φ{x} ≡ x = ❙ιx ψ{x})›
using "&E" "free-thms:1"[THEN "≡E"(1)] "equiv-desc-eq:1" "→E" by blast+
then AOT_obtain a where ‹a = ❙ιx φ{x}›
using "instantiation"[rotated] by blast
moreover AOT_have ‹a = ❙ιx ψ{x}›
using calculation 1 "∀E" "≡E"(1) by fast
ultimately AOT_show ‹❙ιx φ{x} = ❙ιx ψ{x}›
using "rule=E" by fast
qed
AOT_theorem "equiv-desc-eq:3":
‹❙ιx φ{x}↓ & □∀x(φ{x} ≡ ψ{x}) → ❙ιx φ{x} = ❙ιx ψ{x}›
using "→I" "equiv-desc-eq:2"[THEN "→E", OF "&I"] "&E"
"nec-imp-act"[THEN "→E"] by metis
AOT_theorem "equiv-desc-eq:4": ‹❙ιx φ{x}↓ → □❙ιx φ{x}↓›
proof(rule "→I")
AOT_assume ‹❙ιx φ{x}↓›
AOT_hence ‹∃y (y = ❙ιx φ{x})›
by (metis "rule=I:1" "existential:1")
then AOT_obtain a where ‹a = ❙ιx φ{x}›
using "instantiation"[rotated] by blast
AOT_thus ‹□❙ιx φ{x}↓›
using "ex:2:a" "rule=E" by fast
qed
AOT_theorem "equiv-desc-eq:5": ‹❙ιx φ{x}↓ → ∃y □(y = ❙ιx φ{x})›
proof(rule "→I")
AOT_assume ‹❙ιx φ{x}↓›
AOT_hence ‹∃y (y = ❙ιx φ{x})›
by (metis "rule=I:1" "existential:1")
then AOT_obtain a where ‹a = ❙ιx φ{x}›
using "instantiation"[rotated] by blast
AOT_hence ‹□(a = ❙ιx φ{x})›
by (metis "id-nec:2" "vdash-properties:10")
AOT_thus ‹∃y □(y = ❙ιx φ{x})›
by (rule "∃I")
qed
AOT_act_theorem "equiv-desc-eq2:1":
‹∀x (φ{x} ≡ ψ{x}) → ∀x (x = ❙ιx φ{x} ≡ x = ❙ι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":
‹❙ιx φ{x}↓ & ∀x (φ{x} ≡ ψ{x}) → ❙ιx φ{x} = ❙ι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":
‹ψ{❙ιx φ{x}} ≡ ∃x(❙𝒜φ{x} & ∀z(❙𝒜φ{z} → z = x) & ψ{x})›
proof -
AOT_have b: ‹∀x (x = ❙ιx φ{x} ≡ (❙𝒜φ{x} & ∀z(❙𝒜φ{z} → z = x)))›
using "nec-hintikka-scheme" "∀I" by fast
show ?thesis
proof(rule "≡I"; rule "→I")
AOT_assume c: ‹ψ{❙ιx φ{x}}›
AOT_hence d: ‹❙ιx φ{x}↓›
using ψ_denotes_asm by blast
AOT_hence ‹∃y (y = ❙ιx φ{x})›
by (metis "rule=I:1" "existential:1")
then AOT_obtain a where a_def: ‹a = ❙ιx φ{x}›
using "instantiation"[rotated] by blast
moreover AOT_have ‹a = ❙ι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 = ❙ιx φ{x} → ❙ιx φ{x} = a›
by (rule "∀E"(1)[where τ="«❙ι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 ‹❙𝒜φ{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 = ❙ι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 = ❙ιx φ{x}›
using h "≡E" by blast
AOT_thus ‹ψ{❙ιx φ{x}}›
using j "rule=E" by blast
qed
qed
end
AOT_theorem "actual-desc:1": ‹❙ιx φ{x}↓ ≡ ∃!x ❙𝒜φ{x}›
proof (rule "≡I"; rule "→I")
AOT_assume ‹❙ιx φ{x}↓›
AOT_hence ‹∃y (y = ❙ιx φ{x})›
by (metis "rule=I:1" "existential:1")
then AOT_obtain a where ‹a = ❙ιx φ{x}›
using "instantiation"[rotated] by blast
moreover AOT_have ‹a = ❙ιx φ{x} ≡ ∀z(❙𝒜φ{z} ≡ z = a)›
using descriptions[axiom_inst] by blast
ultimately AOT_have ‹∀z(❙𝒜φ{z} ≡ z = a)›
using "≡E" by blast
AOT_hence ‹∃x∀z(❙𝒜φ{z} ≡ z = x)› by (rule "∃I")
AOT_thus ‹∃!x ❙𝒜φ{x}›
using "uniqueness:2"[THEN "≡E"(2)] by fast
next
AOT_assume ‹∃!x ❙𝒜φ{x}›
AOT_hence ‹∃x∀z(❙𝒜φ{z} ≡ z = x)›
using "uniqueness:2"[THEN "≡E"(1)] by fast
then AOT_obtain a where ‹∀z(❙𝒜φ{z} ≡ z = a)›
using "instantiation"[rotated] by blast
moreover AOT_have ‹a = ❙ιx φ{x} ≡ ∀z(❙𝒜φ{z} ≡ z = a)›
using descriptions[axiom_inst] by blast
ultimately AOT_have ‹a = ❙ιx φ{x}›
using "≡E" by blast
AOT_thus ‹❙ιx φ{x}↓›
by (metis "t=t-proper:2" "vdash-properties:6")
qed
AOT_theorem "actual-desc:2": ‹x = ❙ιx φ{x} → ❙𝒜φ{x}›
using "&E"(1) "contraposition:1[2]" "≡E"(1) "nec-hintikka-scheme"
"reductio-aa:2" "vdash-properties:9" by blast
AOT_theorem "actual-desc:3": ‹z = ❙ιx φ{x} → ❙𝒜φ{z}›
using "actual-desc:2".
AOT_theorem "actual-desc:4": ‹❙ιx φ{x}↓ → ❙𝒜φ{❙ιx φ{x}}›
proof(rule "→I")
AOT_assume ‹❙ιx φ{x}↓›
AOT_hence ‹∃y (y = ❙ιx φ{x})› by (metis "rule=I:1" "existential:1")
then AOT_obtain a where ‹a = ❙ιx φ{x}› using "instantiation"[rotated] by blast
AOT_thus ‹❙𝒜φ{❙ιx φ{x}}›
using "actual-desc:2" "rule=E" "→E" by fast
qed
AOT_theorem "actual-desc:5": ‹❙ιx φ{x} = ❙ιx ψ{x} → ❙𝒜∀x(φ{x} ≡ ψ{x})›
proof(rule "→I")
AOT_assume 0: ‹❙ιx φ{x} = ❙ιx ψ{x}›
AOT_hence φ_down: ‹❙ιx φ{x}↓› and ψ_down: ‹❙ιx ψ{x}↓›
using "t=t-proper:1" "t=t-proper:2" "vdash-properties:6" by blast+
AOT_hence ‹∃y (y = ❙ιx φ{x})› and ‹∃y (y = ❙ιx ψ{x})›
by (metis "rule=I:1" "existential:1")+
then AOT_obtain a and b where a_eq: ‹a = ❙ιx φ{x}› and b_eq: ‹b = ❙ιx ψ{x}›
using "instantiation"[rotated] by metis
AOT_have ‹∀α∀β (α = β → β = α)›
by (rule "∀I"; rule "∀I"; rule "id-eq:2")
AOT_hence ‹∀β (❙ιx φ{x} = β → β = ❙ιx φ{x})›
using "∀E" φ_down by blast
AOT_hence ‹❙ιx φ{x} = ❙ιx ψ{x} → ❙ιx ψ{x} = ❙ιx φ{x}›
using "∀E" ψ_down by blast
AOT_hence 1: ‹❙ιx ψ{x} = ❙ιx φ{x}› using 0
"→E" by blast
AOT_have ‹❙𝒜φ{x} ≡ ❙𝒜ψ{x}› for x
proof(rule "≡I"; rule "→I")
AOT_assume ‹❙𝒜φ{x}›
moreover AOT_have ‹❙𝒜φ{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 = ❙ιx φ{x}›
using a_eq "rule=E" by blast
AOT_hence ‹x = ❙ιx ψ{x}›
using 0 "rule=E" by blast
AOT_thus ‹❙𝒜ψ{x}›
by (metis "actual-desc:3" "vdash-properties:6")
next
AOT_assume ‹❙𝒜ψ{x}›
moreover AOT_have ‹❙𝒜ψ{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 = ❙ιx ψ{x}›
using b_eq "rule=E" by blast
AOT_hence ‹x = ❙ιx φ{x}›
using 1 "rule=E" by blast
AOT_thus ‹❙𝒜φ{x}›
by (metis "actual-desc:3" "vdash-properties:6")
qed
AOT_hence ‹❙𝒜(φ{x} ≡ ψ{x})› for x
by (metis "Act-Basic:5" "≡E"(2))
AOT_hence ‹∀x ❙𝒜(φ{x} ≡ ψ{x})›
by (rule "∀I")
AOT_thus ‹❙𝒜∀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 = ❙ιx φ{x} → φ{y})›
proof(rule "→I")
AOT_assume ‹∃!x □φ{x}›
AOT_hence ζ: ‹∃x (□φ{x} & ∀z (□φ{z} → z = x))›
using "uniqueness:1"[THEN "≡⇩d⇩fE"] by blast
then AOT_obtain b where θ: ‹□φ{b} & ∀z (□φ{z} → z = b)›
using "instantiation"[rotated] by blast
AOT_show ‹∀y (y = ❙ιx φ{x} → φ{y})›
proof(rule GEN; rule "→I")
fix y
AOT_assume ‹y = ❙ιx φ{x}›
AOT_hence ‹❙𝒜φ{y} & ∀z (❙𝒜φ{z} → z = y)›
using "nec-hintikka-scheme"[THEN "≡E"(1)] by blast
AOT_hence ‹❙𝒜φ{b} → b = y›
using "&E" "∀E" by blast
moreover AOT_have ‹❙𝒜φ{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 = ❙ι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 = ❙ιx φ{x} → φ{y})›
using "!box-desc:1" "→E" by blast
qed
AOT_theorem "dr-alphabetic-thm": ‹❙ιν φ{ν}↓ → ❙ιν φ{ν} = ❙ιμ φ{μ}›
by (simp add: "rule=I:1" "→I")
subsection‹The Theory of Necessity›
text‹\label{PLM: 9.9}›
AOT_theorem "RM:1[prem]":
assumes ‹Γ ❙⊢⇩□ φ → ψ›
shows ‹□Γ ❙⊢⇩□ □φ → □ψ›
proof -
AOT_have ‹□Γ ❙⊢⇩□ □(φ → ψ)›
using "RN[prem]" assms by blast
AOT_thus ‹□Γ ❙⊢⇩□ □φ → □ψ›
by (metis "qml:1"[axiom_inst] "→E")
qed
AOT_theorem "RM:1":
assumes ‹❙⊢⇩□ φ → ψ›
shows ‹❙⊢⇩□ □φ → □ψ›
using "RM:1[prem]" assms by blast
lemmas RM = "RM:1"
AOT_theorem "RM:2[prem]":
assumes ‹Γ ❙⊢⇩□ φ → ψ›
shows ‹□Γ ❙⊢⇩□ ◇φ → ◇ψ›
proof -
AOT_have ‹Γ ❙⊢⇩□ ¬ψ → ¬φ›
using assms
by (simp add: "contraposition:1[1]")
AOT_hence ‹□Γ ❙⊢⇩□ □¬ψ → □¬φ›
using "RM:1[prem]" by blast
AOT_thus ‹□Γ ❙⊢⇩□ ◇φ → ◇ψ›
by (meson "≡⇩d⇩fE" "≡⇩d⇩fI" "conventions:5" "→I" "modus-tollens:1")
qed
AOT_theorem "RM:2":
assumes ‹❙⊢⇩□ φ → ψ›
shows ‹❙⊢⇩□ ◇φ → ◇ψ›
using "RM:2[prem]" assms by blast
lemmas "RM◇" = "RM:2"
AOT_theorem "RM:3[prem]":
assumes ‹Γ ❙⊢⇩□ φ ≡ ψ›
shows ‹□Γ ❙⊢⇩□ □φ ≡ □ψ›
proof -
AOT_have ‹Γ ❙⊢⇩□ φ → ψ› and ‹Γ ❙⊢⇩□ ψ → φ›
using assms "≡E" "→I" by metis+
AOT_hence ‹□Γ ❙⊢⇩□ □φ → □ψ› and ‹□Γ ❙⊢⇩□ □ψ → □φ›
using "RM:1[prem]" by metis+
AOT_thus ‹□Γ ❙⊢⇩□ □φ ≡ □ψ›
by (simp add: "≡I")
qed
AOT_theorem "RM:3":
assumes ‹❙⊢⇩□ φ ≡ ψ›
shows ‹❙⊢⇩□ □φ ≡ □ψ›
using "RM:3[prem]" assms by blast
lemmas RE = "RM:3"
AOT_theorem "RM:4[prem]":
assumes ‹Γ ❙⊢⇩□ φ ≡ ψ›
shows ‹□Γ ❙⊢⇩□ ◇φ ≡ ◇ψ›
proof -
AOT_have ‹Γ ❙⊢⇩□ φ → ψ› and ‹Γ ❙⊢⇩□ ψ → φ›
using assms "≡E" "→I" by metis+
AOT_hence ‹□Γ ❙⊢⇩□ ◇φ → ◇ψ› and ‹□Γ ❙⊢⇩□ ◇ψ → ◇φ›
using "RM:2[prem]" by metis+
AOT_thus ‹□Γ ❙⊢⇩□ ◇φ ≡ ◇ψ›
by (simp add: "≡I")
qed
AOT_theorem "RM:4":
assumes ‹❙⊢⇩□ φ ≡ ψ›
shows ‹❙⊢⇩□ ◇φ ≡ ◇ψ›
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 "≡⇩d⇩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 "≡⇩d⇩fI" "conventions:5" "KBasic:10" "≡E"(3) by blast
next
AOT_show ‹¬□φ› if ‹◇¬φ›
using "≡⇩d⇩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 ‹φ → ψ ❙⊢⇩□ φ → ψ› by blast
AOT_hence ‹□(φ → ψ) ❙⊢⇩□ ◇φ → ◇ψ›
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 ‹❙⊢⇩□ □(ψ ≡ χ)›
shows ‹❙⊢⇩□ ¬ψ ≡ ¬χ›
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 ‹❙⊢⇩□ □(ψ ≡ χ)›
shows ‹❙⊢⇩□ (ψ → Θ) ≡ (χ → Θ)›
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 ‹❙⊢⇩□ □(ψ ≡ χ)›
shows ‹❙⊢⇩□ (Θ → ψ) ≡ (Θ → χ)›
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 α: ❙⊢⇩□ □(ψ{α} ≡ χ{α})›
shows ‹❙⊢⇩□ ∀α ψ{α} ≡ ∀α χ{α}›
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 ‹❙⊢⇩□ □(ψ ≡ χ)›
shows ‹❙⊢⇩□ [λ ψ] ≡ [λ χ]›
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 ‹❙⊢⇩□ □(ψ ≡ χ)›
shows ‹❙⊢⇩□ ❙𝒜ψ ≡ ❙𝒜χ›
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 ‹❙⊢⇩□ □(ψ ≡ χ)›
shows ‹❙⊢⇩□ □ψ ≡ □χ›
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_𝗈 (λ φ. «❙𝒜Θ{φ}»)›
| 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 ‹❙⊢⇩□ ψ ≡ χ›
using cond unfolding AOT_subst_cond_𝗈_def by blast
ultimately AOT_show ‹❙⊢⇩□ φ{ψ} ≡ φ{χ}›
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 (λ φ. «❙𝒜«Θ φ»»)›
| 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 ‹❙⊢⇩□ «φ ψ» ≡ «φ χ»›
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 ‹❙⊢⇩□ □(Θ{α, «ψ»} ≡ Θ{α, «χ»})› 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)
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
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 "":
assumes ‹❙⊢⇩□ A!x ≡ ¬◇E!x› and ‹¬A!x›
shows ‹¬¬◇E!x›
by (AOT_subst (reverse) ‹¬◇E!x› ‹A!x›)
(auto simp: assms)
AOT_theorem "":
assumes ‹❙⊢⇩□ A!x ≡ ¬◇E!x› and ‹¬¬◇E!x›
shows ‹¬A!x›
by (AOT_subst ‹A!x› ‹¬◇E!x›)
(auto simp: assms)
AOT_theorem "":
assumes ‹❙⊢⇩□ [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 "":
assumes ‹❙⊢⇩□ [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 "":
assumes ‹for arbitrary x: ❙⊢⇩□ 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 "":
assumes ‹for arbitrary x: ❙⊢⇩□ 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 "":
assumes ‹❙⊢⇩□ ¬¬[P]x ≡ [P]x› and ‹❙𝒜¬¬[P]x›
shows ‹❙𝒜[P]x›
by (AOT_subst_thm (reverse) assms(1)) (simp add: assms(2))
AOT_theorem "":
assumes ‹❙⊢⇩□ ¬¬[P]x ≡ [P]x› and ‹❙𝒜[P]x›
shows ‹❙𝒜¬¬[P]x›
by (AOT_subst_thm assms(1)) (simp add: assms(2))
AOT_theorem "":
assumes ‹❙⊢⇩□ (φ → ψ) ≡ (¬ψ → ¬φ)› and ‹□(φ → ψ)›
shows ‹□(¬ψ → ¬φ)›
by (AOT_subst_thm (reverse) assms(1)) (simp add: assms(2))
AOT_theorem "":
assumes ‹❙⊢⇩□ (φ → ψ) ≡ (¬ψ → ¬φ)› and ‹□(¬ψ → ¬φ)›
shows ‹□(φ → ψ)›
by (AOT_subst_thm assms(1)) (simp add: assms(2))
AOT_theorem "":
assumes ‹❙⊢⇩□ ψ ≡ χ› and ‹□(φ → ψ)›
shows ‹□(φ → χ)›
by (AOT_subst_thm (reverse) assms(1)) (simp add: assms(2))
AOT_theorem "":
assumes ‹❙⊢⇩□ ψ ≡ χ› and ‹□(φ → χ)›
shows ‹□(φ → ψ)›
by (AOT_subst_thm assms(1)) (simp add: assms(2))
AOT_theorem "":
assumes ‹❙⊢⇩□ φ ≡ ¬¬φ› and ‹□(φ → φ)›
shows ‹□(¬¬φ → φ)›
by (AOT_subst_thm (reverse) assms(1)) (simp add: assms(2))
AOT_theorem "":
assumes ‹❙⊢⇩□ φ ≡ ¬¬φ› 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 "≡⇩d⇩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": ‹❙𝒜φ ≡ ¬❙𝒜¬φ›
by (AOT_subst ‹❙𝒜¬φ› ‹¬❙𝒜φ›)
(auto simp: "logic-actual-nec:1"[axiom_inst] "oth-class-taut:3:b")
AOT_theorem "Act-Sub:2": ‹◇φ ≡ ❙𝒜◇φ›
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": ‹❙𝒜φ → ◇φ›
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": ‹❙𝒜φ ≡ ◇❙𝒜φ›
proof (rule "≡I"; rule "→I")
AOT_assume ‹❙𝒜φ›
AOT_thus ‹◇❙𝒜φ› using "T◇" "vdash-properties:10" by blast
next
AOT_assume ‹◇❙𝒜φ›
AOT_hence ‹¬□¬❙𝒜φ›
using "≡⇩d⇩fE" "conventions:5" by blast
AOT_hence ‹¬□❙𝒜¬φ›
by (AOT_subst ‹❙𝒜¬φ› ‹¬❙𝒜φ›)
(simp add: "logic-actual-nec:1"[axiom_inst])
AOT_thus ‹❙𝒜φ›
using "Act-Basic:1" "Act-Basic:6" "∨E"(3) "≡E"(4)
"reductio-aa:1" by blast
qed
AOT_theorem "Act-Sub:5": ‹◇❙𝒜φ → ❙𝒜◇φ›
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 ‹Γ ❙⊢⇩□ ◇φ → ψ›
shows ‹□Γ ❙⊢⇩□ φ → □ψ›
proof -
AOT_have ‹□Γ ❙⊢⇩□ □◇φ → □ψ›
using assms by (rule "RM:1[prem]")
AOT_thus ‹□Γ ❙⊢⇩□ φ → □ψ›
using "B" "Hypothetical Syllogism" by blast
qed
AOT_theorem "derived-S5-rules:2":
assumes ‹Γ ❙⊢⇩□ φ → □ψ›
shows ‹□Γ ❙⊢⇩□ ◇φ → ψ›
proof -
AOT_have ‹□Γ ❙⊢⇩□ ◇φ → ◇□ψ›
using assms by (rule "RM:2[prem]")
AOT_thus ‹□Γ ❙⊢⇩□ ◇φ → ψ›
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 "≡⇩d⇩fE" "conventions:5" by blast
AOT_hence ‹¬□∀α ¬φ{α}›
apply (AOT_subst ‹∀α ¬φ{α}› ‹¬(∃α φ{α})›)
using "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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 ‹□φ{τ}, □τ↓ ❙⊢⇩□ □∃α φ{α}› for τ
proof -
AOT_have ‹φ{τ}, τ↓ ❙⊢⇩□ ∃α φ{α}› using "existential:1" by blast
AOT_thus ‹□φ{τ}, □τ↓ ❙⊢⇩□ □∃α φ{α}›
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 ‹❙⊢⇩□ φ' & ψ' → χ'›
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 ‹❙⊢⇩□ φ' & ψ' → χ'›
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": ‹□(φ → □φ) → ((φ → ❙𝒜ψ) → ❙𝒜(φ → ψ))›
proof (rule "→I"; rule "→I"; rule "raa-cor:1")
AOT_assume ‹¬❙𝒜(φ → ψ)›
AOT_hence ‹❙𝒜¬(φ → ψ)›
by (metis "Act-Basic:1" "∨E"(2))
AOT_hence ‹❙𝒜(φ & ¬ψ)›
by (AOT_subst ‹φ & ¬ψ› ‹¬(φ → ψ)›)
(meson "Commutativity of ≡" "≡E"(1) "oth-class-taut:1:b")
AOT_hence ‹❙𝒜φ› and 2: ‹❙𝒜¬ψ›
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 ‹φ → ❙𝒜ψ›
ultimately AOT_have ‹❙𝒜ψ›
using "→E" by blast
moreover AOT_have ‹¬❙𝒜ψ›
using 2 by (meson "Act-Sub:1" "≡E"(4) "raa-cor:3")
ultimately AOT_show ‹❙𝒜ψ & ¬❙𝒜ψ›
using "&I" by blast
qed
AOT_theorem "sc-eq-fur:1": ‹◇❙𝒜φ ≡ □❙𝒜φ›
using "Act-Basic:6" "Act-Sub:4" "≡E"(6) by blast
AOT_theorem "sc-eq-fur:2": ‹□(φ → □φ) → (❙𝒜φ ≡ φ)›
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} → ❙ι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 "≡⇩d⇩fE"]] by blast
moreover AOT_have ‹□φ{a}›
using calculation A "∀E"(2) "qml:2"[axiom_inst] "→E" "&E"(1) by blast
AOT_hence ‹❙𝒜φ{a}›
using "nec-imp-act" "→E" by blast
moreover AOT_have ‹∀y (❙𝒜φ{y} → y = a)›
proof (rule "∀I"; rule "→I")
fix b
AOT_assume ‹❙𝒜φ{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 ‹❙𝒜φ{a} & ∀y (❙𝒜φ{y} → y = a)›
using "&I" by blast
AOT_hence ‹∃x (❙𝒜φ{x} & ∀y (❙𝒜φ{y} → y = x))›
using "∃I" by fast
AOT_hence ‹∃!x ❙𝒜φ{x}›
using "uniqueness:1"[THEN "≡⇩d⇩fI"] by fast
AOT_thus ‹❙ιx φ{x}↓›
using "actual-desc:1"[THEN "≡E"(2)] by blast
qed
AOT_theorem "sc-eq-fur:4":
‹□∀x (φ{x} → □φ{x}) → (x = ❙ι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: ‹❙𝒜φ{α} ≡ φ{α}› for α
using "sc-eq-fur:2" "∀E" "→E" by fast
AOT_show ‹x = ❙ιx φ{x} ≡ (φ{x} & ∀z (φ{z} → z = x))›
proof (rule "≡I"; rule "→I")
AOT_assume ‹x = ❙ιx φ{x}›
AOT_hence B: ‹❙𝒜φ{x} & ∀z (❙𝒜φ{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 ‹❙𝒜φ{x} & ∀z (❙𝒜φ{z} → z = x)›
proof(rule "&I"; (rule "∀I"; rule "→I")?)
AOT_show ‹❙𝒜φ{x}›
using B[THEN "&E"(1)] A[THEN "≡E"(2)] by blast
next
AOT_show ‹b = x› if ‹❙𝒜φ{b}› for b
using A[THEN "≡E"(1)] that
B[THEN "&E"(2), THEN "∀E"(2), THEN "→E"] by blast
qed
AOT_thus ‹x = ❙ιx φ{x}›
using "nec-hintikka-scheme"[THEN "≡E"(2)] by blast
qed
qed
AOT_theorem "id-act:1": ‹α = β ≡ ❙𝒜α = β›
by (meson "Act-Sub:3" "Hypothetical Syllogism"
"id-nec2:1" "id-nec:2" "≡I" "nec-imp-act")
AOT_theorem "id-act:2": ‹α ≠ β ≡ ❙𝒜α ≠ β›
proof (AOT_subst ‹α ≠ β› ‹¬(α = β)›)
AOT_modally_strict {
AOT_show ‹α ≠ β ≡ ¬(α = β)›
by (simp add: "=-infix" "≡Df")
}
next
AOT_show ‹¬(α = β) ≡ ❙𝒜¬(α = β)›
proof (safe intro!: "≡I" "→I")
AOT_assume ‹¬α = β›
AOT_hence ‹¬❙𝒜α = β› using "id-act:1" "≡E"(3) by blast
AOT_thus ‹❙𝒜¬α = β›
using "¬¬E" "Act-Sub:1" "≡E"(3) by blast
next
AOT_assume ‹❙𝒜¬α = β›
AOT_hence ‹¬❙𝒜α = β›
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": ‹❙𝒜∃!α φ{α} ≡ ∃!α ❙𝒜φ{α}›
proof -
AOT_have ‹❙𝒜∃!α φ{α} ≡ ❙𝒜∃α∀β (φ{β} ≡ β = α)›
by (AOT_subst ‹∃!α φ{α}› ‹∃α∀β (φ{β} ≡ β = α)›)
(auto simp add: "oth-class-taut:3:a" "uniqueness:2")
also AOT_have ‹… ≡ ∃α ❙𝒜∀β (φ{β} ≡ β = α)›
by (simp add: "Act-Basic:10")
also AOT_have ‹… ≡ ∃α∀β ❙𝒜(φ{β} ≡ β = α)›
by (AOT_subst ‹❙𝒜∀β (φ{β} ≡ β = α)› ‹∀β ❙𝒜(φ{β} ≡ β = α)› for: α)
(auto simp: "logic-actual-nec:3"[axiom_inst] "oth-class-taut:3:a")
also AOT_have ‹… ≡ ∃α∀β (❙𝒜φ{β} ≡ ❙𝒜β = α)›
by (AOT_subst (reverse) ‹❙𝒜φ{β} ≡ ❙𝒜β = α›
‹❙𝒜(φ{β} ≡ β = α)› for: α β :: 'a)
(auto simp: "Act-Basic:5" "cqt-further:7")
also AOT_have ‹… ≡ ∃α∀β (❙𝒜φ{β} ≡ β = α)›
by (AOT_subst (reverse) ‹❙𝒜β = α› ‹β = α› for: α β :: 'a)
(auto simp: "id-act:1" "cqt-further:7")
also AOT_have ‹... ≡ ∃!α ❙𝒜φ{α}›
using "uniqueness:2" "Commutativity of ≡"[THEN "≡E"(1)] by fast
finally show ?thesis.
qed
AOT_theorem "A-Exists:2": ‹❙ιx φ{x}↓ ≡ ❙𝒜∃!x φ{x}›
by (AOT_subst ‹❙𝒜∃!x φ{x}› ‹∃!x ❙𝒜φ{x}›)
(auto simp: "actual-desc:1" "A-Exists:1")
AOT_theorem "id-act-desc:1": ‹❙ιx (x = y)↓›
proof(rule "existence:1"[THEN "≡⇩d⇩fI"]; rule "∃I")
AOT_show ‹[λx E!x → E!x]❙ιx (x = y)›
proof (rule "russell-axiom[exe,1].nec-russell-axiom"[THEN "≡E"(2)];
rule "∃I"; (rule "&I")+)
AOT_show ‹❙𝒜y = y› by (simp add: "RA[2]" "id-eq:1")
next
AOT_show ‹∀z (❙𝒜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 = ❙ι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⇩1[F] → □x⇩1[F]›
by (simp add: encoding "vdash-properties:1[2]")
AOT_theorem "pre-en-eq:1[2]": ‹x⇩1x⇩2[F] → □x⇩1x⇩2[F]›
proof (rule "→I")
AOT_assume ‹x⇩1x⇩2[F]›
AOT_hence ‹x⇩1[λy [F]yx⇩2]› and ‹x⇩2[λy [F]x⇩1y]›
using "nary-encoding[2]"[axiom_inst, THEN "≡E"(1)] "&E" by blast+
moreover AOT_have ‹[λy [F]yx⇩2]↓› by "cqt:2"
moreover AOT_have ‹[λy [F]x⇩1y]↓› by "cqt:2"
ultimately AOT_have ‹□x⇩1[λy [F]yx⇩2]› and ‹□x⇩2[λy [F]x⇩1y]›
using encoding[axiom_inst, unvarify F] "→E" "&I" by blast+
note A = this
AOT_hence ‹□(x⇩1[λy [F]yx⇩2] & x⇩2[λy [F]x⇩1y])›
using "KBasic:3"[THEN "≡E"(2)] "&I" by blast
AOT_thus ‹□x⇩1x⇩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⇩1x⇩2x⇩3[F] → □x⇩1x⇩2x⇩3[F]›
proof (rule "→I")
AOT_assume ‹x⇩1x⇩2x⇩3[F]›
AOT_hence ‹x⇩1[λy [F]yx⇩2x⇩3]›
and ‹x⇩2[λy [F]x⇩1yx⇩3]›
and ‹x⇩3[λy [F]x⇩1x⇩2y]›
using "nary-encoding[3]"[axiom_inst, THEN "≡E"(1)] "&E" by blast+
moreover AOT_have ‹[λy [F]yx⇩2x⇩3]↓› by "cqt:2"
moreover AOT_have ‹[λy [F]x⇩1yx⇩3]↓› by "cqt:2"
moreover AOT_have ‹[λy [F]x⇩1x⇩2y]↓› by "cqt:2"
ultimately AOT_have ‹□x⇩1[λy [F]yx⇩2x⇩3]›
and ‹□x⇩2[λy [F]x⇩1yx⇩3]›
and ‹□x⇩3[λy [F]x⇩1x⇩2y]›
using encoding[axiom_inst, unvarify F] "→E" by blast+
note A = this
AOT_have B: ‹□(x⇩1[λy [F]yx⇩2x⇩3] & x⇩2[λy [F]x⇩1yx⇩3] & x⇩3[λy [F]x⇩1x⇩2y])›
by (rule "KBasic:3"[THEN "≡E"(2)] "&I" A)+
AOT_thus ‹□x⇩1x⇩2x⇩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⇩1x⇩2x⇩3x⇩4[F] → □x⇩1x⇩2x⇩3x⇩4[F]›
proof (rule "→I")
AOT_assume ‹x⇩1x⇩2x⇩3x⇩4[F]›
AOT_hence ‹x⇩1[λy [F]yx⇩2x⇩3x⇩4]›
and ‹x⇩2[λy [F]x⇩1yx⇩3x⇩4]›
and ‹x⇩3[λy [F]x⇩1x⇩2yx⇩4]›
and ‹x⇩4[λy [F]x⇩1x⇩2x⇩3y]›
using "nary-encoding[4]"[axiom_inst, THEN "≡E"(1)] "&E" by metis+
moreover AOT_have ‹[λy [F]yx⇩2x⇩3x⇩4]↓› by "cqt:2"
moreover AOT_have ‹[λy [F]x⇩1yx⇩3x⇩4]↓› by "cqt:2"
moreover AOT_have ‹[λy [F]x⇩1x⇩2yx⇩4]↓› by "cqt:2"
moreover AOT_have ‹[λy [F]x⇩1x⇩2x⇩3y]↓› by "cqt:2"
ultimately AOT_have ‹□x⇩1[λy [F]yx⇩2x⇩3x⇩4]›
and ‹□x⇩2[λy [F]x⇩1yx⇩3x⇩4]›
and ‹□x⇩3[λy [F]x⇩1x⇩2yx⇩4]›
and ‹□x⇩4[λy [F]x⇩1x⇩2x⇩3y]›
using "→E" encoding[axiom_inst, unvarify F] by blast+
note A = this
AOT_have B: ‹□(x⇩1[λy [F]yx⇩2x⇩3x⇩4] &
x⇩2[λy [F]x⇩1yx⇩3x⇩4] &
x⇩3[λy [F]x⇩1x⇩2yx⇩4] &
x⇩4[λy [F]x⇩1x⇩2x⇩3y])›
by (rule "KBasic:3"[THEN "≡E"(2)] "&I" A)+
AOT_thus ‹□x⇩1x⇩2x⇩3x⇩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⇩1[F] → □¬x⇩1[F]›
proof (rule "→I"; rule "raa-cor:1")
AOT_assume ‹¬□¬x⇩1[F]›
AOT_hence ‹◇x⇩1[F]›
by (rule "conventions:5"[THEN "≡⇩d⇩fI"])
AOT_hence ‹x⇩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⇩1[F]›
ultimately AOT_show ‹x⇩1[F] & ¬x⇩1[F]› by (rule "&I")
qed
AOT_theorem "pre-en-eq:2[2]": ‹¬x⇩1x⇩2[F] → □¬x⇩1x⇩2[F]›
proof (rule "→I"; rule "raa-cor:1")
AOT_assume ‹¬□¬x⇩1x⇩2[F]›
AOT_hence ‹◇x⇩1x⇩2[F]›
by (rule "conventions:5"[THEN "≡⇩d⇩fI"])
AOT_hence ‹x⇩1x⇩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⇩1x⇩2[F]›
ultimately AOT_show ‹x⇩1x⇩2[F] & ¬x⇩1x⇩2[F]› by (rule "&I")
qed
AOT_theorem "pre-en-eq:2[3]": ‹¬x⇩1x⇩2x⇩3[F] → □¬x⇩1x⇩2x⇩3[F]›
proof (rule "→I"; rule "raa-cor:1")
AOT_assume ‹¬□¬x⇩1x⇩2x⇩3[F]›
AOT_hence ‹◇x⇩1x⇩2x⇩3[F]›
by (rule "conventions:5"[THEN "≡⇩d⇩fI"])
AOT_hence ‹x⇩1x⇩2x⇩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⇩1x⇩2x⇩3[F]›
ultimately AOT_show ‹x⇩1x⇩2x⇩3[F] & ¬x⇩1x⇩2x⇩3[F]› by (rule "&I")
qed
AOT_theorem "pre-en-eq:2[4]": ‹¬x⇩1x⇩2x⇩3x⇩4[F] → □¬x⇩1x⇩2x⇩3x⇩4[F]›
proof (rule "→I"; rule "raa-cor:1")
AOT_assume ‹¬□¬x⇩1x⇩2x⇩3x⇩4[F]›
AOT_hence ‹◇x⇩1x⇩2x⇩3x⇩4[F]›
by (rule "conventions:5"[THEN "≡⇩d⇩fI"])
AOT_hence ‹x⇩1x⇩2x⇩3x⇩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⇩1x⇩2x⇩3x⇩4[F]›
ultimately AOT_show ‹x⇩1x⇩2x⇩3x⇩4[F] & ¬x⇩1x⇩2x⇩3x⇩4[F]› by (rule "&I")
qed
AOT_theorem "en-eq:1[1]": ‹◇x⇩1[F] ≡ □x⇩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⇩1x⇩2[F] ≡ □x⇩1x⇩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⇩1x⇩2x⇩3[F] ≡ □x⇩1x⇩2x⇩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⇩1x⇩2x⇩3x⇩4[F] ≡ □x⇩1x⇩2x⇩3x⇩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⇩1[F] ≡ □x⇩1[F]›
by (simp add: "≡I" "pre-en-eq:1[1]" "qml:2"[axiom_inst])
AOT_theorem "en-eq:2[2]": ‹x⇩1x⇩2[F] ≡ □x⇩1x⇩2[F]›
by (simp add: "≡I" "pre-en-eq:1[2]" "qml:2"[axiom_inst])
AOT_theorem "en-eq:2[3]": ‹x⇩1x⇩2x⇩3[F] ≡ □x⇩1x⇩2x⇩3[F]›
by (simp add: "≡I" "pre-en-eq:1[3]" "qml:2"[axiom_inst])
AOT_theorem "en-eq:2[4]": ‹x⇩1x⇩2x⇩3x⇩4[F] ≡ □x⇩1x⇩2x⇩3x⇩4[F]›
by (simp add: "≡I" "pre-en-eq:1[4]" "qml:2"[axiom_inst])
AOT_theorem "en-eq:3[1]": ‹◇x⇩1[F] ≡ x⇩1[F]›
using "T◇" "derived-S5-rules:2"[OF "pre-en-eq:1[1]"] "≡I" by blast
AOT_theorem "en-eq:3[2]": ‹◇x⇩1x⇩2[F] ≡ x⇩1x⇩2[F]›
using "T◇" "derived-S5-rules:2"[OF "pre-en-eq:1[2]"] "≡I" by blast
AOT_theorem "en-eq:3[3]": ‹◇x⇩1x⇩2x⇩3[F] ≡ x⇩1x⇩2x⇩3[F]›
using "T◇" "derived-S5-rules:2"[OF "pre-en-eq:1[3]"] "≡I" by blast
AOT_theorem "en-eq:3[4]": ‹◇x⇩1x⇩2x⇩3x⇩4[F] ≡ x⇩1x⇩2x⇩3x⇩4[F]›
using "T◇" "derived-S5-rules:2"[OF "pre-en-eq:1[4]"] "≡I" by blast
AOT_theorem "en-eq:4[1]":
‹(x⇩1[F] ≡ y⇩1[G]) ≡ (□x⇩1[F] ≡ □y⇩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⇩1x⇩2[F] ≡ y⇩1y⇩2[G]) ≡ (□x⇩1x⇩2[F] ≡ □y⇩1y⇩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⇩1x⇩2x⇩3[F] ≡ y⇩1y⇩2y⇩3[G]) ≡ (□x⇩1x⇩2x⇩3[F] ≡ □y⇩1y⇩2y⇩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⇩1x⇩2x⇩3x⇩4[F] ≡ y⇩1y⇩2y⇩3y⇩4[G]) ≡ (□x⇩1x⇩2x⇩3x⇩4[F] ≡ □y⇩1y⇩2y⇩3y⇩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⇩1[F] ≡ y⇩1[G]) ≡ (□x⇩1[F] ≡ □y⇩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⇩1x⇩2[F] ≡ y⇩1y⇩2[G]) ≡ (□x⇩1x⇩2[F] ≡ □y⇩1y⇩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⇩1x⇩2x⇩3[F] ≡ y⇩1y⇩2y⇩3[G]) ≡ (□x⇩1x⇩2x⇩3[F] ≡ □y⇩1y⇩2y⇩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⇩1x⇩2x⇩3x⇩4[F] ≡ y⇩1y⇩2y⇩3y⇩4[G]) ≡ (□x⇩1x⇩2x⇩3x⇩4[F] ≡ □y⇩1y⇩2y⇩3y⇩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⇩1[F] ≡ y⇩1[G]) ≡ □(x⇩1[F] ≡ y⇩1[G])›
using "en-eq:5[1]"[symmetric] "en-eq:4[1]" "≡E"(5) by fast
AOT_theorem "en-eq:6[2]":
‹(x⇩1x⇩2[F] ≡ y⇩1y⇩2[G]) ≡ □(x⇩1x⇩2[F] ≡ y⇩1y⇩2[G])›
using "en-eq:5[2]"[symmetric] "en-eq:4[2]" "≡E"(5) by fast
AOT_theorem "en-eq:6[3]":
‹(x⇩1x⇩2x⇩3[F] ≡ y⇩1y⇩2y⇩3[G]) ≡ □(x⇩1x⇩2x⇩3[F] ≡ y⇩1y⇩2y⇩3[G])›
using "en-eq:5[3]"[symmetric] "en-eq:4[3]" "≡E"(5) by fast
AOT_theorem "en-eq:6[4]":
‹(x⇩1x⇩2x⇩3x⇩4[F] ≡ y⇩1y⇩2y⇩3y⇩4[G]) ≡ □(x⇩1x⇩2x⇩3x⇩4[F] ≡ y⇩1y⇩2y⇩3y⇩4[G])›
using "en-eq:5[4]"[symmetric] "en-eq:4[4]" "≡E"(5) by fast
AOT_theorem "en-eq:7[1]": ‹¬x⇩1[F] ≡ □¬x⇩1[F]›
using "pre-en-eq:2[1]" "qml:2"[axiom_inst] "≡I" by blast
AOT_theorem "en-eq:7[2]": ‹¬x⇩1x⇩2[F] ≡ □¬x⇩1x⇩2[F]›
using "pre-en-eq:2[2]" "qml:2"[axiom_inst] "≡I" by blast
AOT_theorem "en-eq:7[3]": ‹¬x⇩1x⇩2x⇩3[F] ≡ □¬x⇩1x⇩2x⇩3[F]›
using "pre-en-eq:2[3]" "qml:2"[axiom_inst] "≡I" by blast
AOT_theorem "en-eq:7[4]": ‹¬x⇩1x⇩2x⇩3x⇩4[F] ≡ □¬x⇩1x⇩2x⇩3x⇩4[F]›
using "pre-en-eq:2[4]" "qml:2"[axiom_inst] "≡I" by blast
AOT_theorem "en-eq:8[1]": ‹◇¬x⇩1[F] ≡ ¬x⇩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⇩1x⇩2[F] ≡ ¬x⇩1x⇩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⇩1x⇩2x⇩3[F] ≡ ¬x⇩1x⇩2x⇩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⇩1x⇩2x⇩3x⇩4[F] ≡ ¬x⇩1x⇩2x⇩3x⇩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⇩1[F] ≡ □¬x⇩1[F]›
using "en-eq:7[1]" "en-eq:8[1]" "≡E"(5) by blast
AOT_theorem "en-eq:9[2]": ‹◇¬x⇩1x⇩2[F] ≡ □¬x⇩1x⇩2[F]›
using "en-eq:7[2]" "en-eq:8[2]" "≡E"(5) by blast
AOT_theorem "en-eq:9[3]": ‹◇¬x⇩1x⇩2x⇩3[F] ≡ □¬x⇩1x⇩2x⇩3[F]›
using "en-eq:7[3]" "en-eq:8[3]" "≡E"(5) by blast
AOT_theorem "en-eq:9[4]": ‹◇¬x⇩1x⇩2x⇩3x⇩4[F] ≡ □¬x⇩1x⇩2x⇩3x⇩4[F]›
using "en-eq:7[4]" "en-eq:8[4]" "≡E"(5) by blast
AOT_theorem "en-eq:10[1]": ‹❙𝒜x⇩1[F] ≡ x⇩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]": ‹❙𝒜x⇩1x⇩2[F] ≡ x⇩1x⇩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]": ‹❙𝒜x⇩1x⇩2x⇩3[F] ≡ x⇩1x⇩2x⇩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]": ‹❙𝒜x⇩1x⇩2x⇩3x⇩4[F] ≡ x⇩1x⇩2x⇩3x⇩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 "=⇩d⇩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 "=⇩d⇩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 "=⇩d⇩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 "=⇩d⇩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 ≡ ❙𝒜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 ≡ ❙𝒜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":
‹[λμ⇩1...μ⇩n φ{μ⇩1...μ⇩n, ν⇩1...ν⇩n}]↓ →
([λμ⇩1...μ⇩n φ{μ⇩1...μ⇩n, ν⇩1...ν⇩n}]ν⇩1...ν⇩n ≡ φ{ν⇩1...ν⇩n, ν⇩1...ν⇩n})›
using "lambda-predicates:2"[axiom_inst] by blast
AOT_theorem "beta-C-cor:1":
‹(∀ν⇩1...∀ν⇩n([λμ⇩1...μ⇩n φ{μ⇩1...μ⇩n, ν⇩1...ν⇩n}]↓)) →
∀ν⇩1...∀ν⇩n ([λμ⇩1...μ⇩n φ{μ⇩1...μ⇩n, ν⇩1...ν⇩n}]ν⇩1...ν⇩n ≡ φ{ν⇩1...ν⇩n, ν⇩1...ν⇩n})›
apply (rule "cqt-basic:14"[where 'a='a, THEN "→E"])
using "beta-C-meta" "∀I" by fast
AOT_theorem "beta-C-cor:2":
‹[λμ⇩1...μ⇩n φ{μ⇩1...μ⇩n}]↓ →
∀ν⇩1...∀ν⇩n ([λμ⇩1...μ⇩n φ{μ⇩1...μ⇩n}]ν⇩1...ν⇩n ≡ φ{ν⇩1...ν⇩n})›
apply (rule "→I"; rule "∀I")
using "beta-C-meta"[THEN "→E"] by fast
theorem "beta-C-cor:3":
assumes ‹⋀ν⇩1ν⇩n. AOT_instance_of_cqt_2 (φ (AOT_term_of_var ν⇩1ν⇩n))›
shows ‹[v ⊨ ∀ν⇩1...∀ν⇩n ([λμ⇩1...μ⇩n φ{ν⇩1...ν⇩n, μ⇩1...μ⇩n}]ν⇩1...ν⇩n ≡
φ{ν⇩1...ν⇩n, ν⇩1...ν⇩n})]›
using "cqt:2[lambda]"[axiom_inst, OF assms]
"beta-C-cor:1"[THEN "→E"] "∀I" by fast
AOT_theorem "betaC:1:a": ‹[λμ⇩1...μ⇩n φ{μ⇩1...μ⇩n}]κ⇩1...κ⇩n ❙⊢⇩□ φ{κ⇩1...κ⇩n}›
proof -
AOT_modally_strict {
AOT_assume ‹[λμ⇩1...μ⇩n φ{μ⇩1...μ⇩n}]κ⇩1...κ⇩n›
moreover AOT_have ‹[λμ⇩1...μ⇩n φ{μ⇩1...μ⇩n}]↓› and ‹κ⇩1...κ⇩n↓›
using calculation "cqt:5:a"[axiom_inst, THEN "→E"] "&E" by blast+
ultimately AOT_show ‹φ{κ⇩1...κ⇩n}›
using "beta-C-cor:2"[THEN "→E", THEN "∀E"(1), THEN "≡E"(1)] by blast
}
qed
AOT_theorem "betaC:1:b": ‹¬φ{κ⇩1...κ⇩n} ❙⊢⇩□ ¬[λμ⇩1...μ⇩n φ{μ⇩1...μ⇩n}]κ⇩1...κ⇩n›
using "betaC:1:a" "raa-cor:3" by blast
lemmas "β→C" = "betaC:1:a" "betaC:1:b"
AOT_theorem "betaC:2:a":
‹[λμ⇩1...μ⇩n φ{μ⇩1...μ⇩n}]↓, κ⇩1...κ⇩n↓, φ{κ⇩1...κ⇩n} ❙⊢⇩□
[λμ⇩1...μ⇩n φ{μ⇩1...μ⇩n}]κ⇩1...κ⇩n›
proof -
AOT_modally_strict {
AOT_assume 1: ‹[λμ⇩1...μ⇩n φ{μ⇩1...μ⇩n}]↓›
and 2: ‹κ⇩1...κ⇩n↓›
and 3: ‹φ{κ⇩1...κ⇩n}›
AOT_hence ‹[λμ⇩1...μ⇩n φ{μ⇩1...μ⇩n}]κ⇩1...κ⇩n›
using "beta-C-cor:2"[THEN "→E", OF 1, THEN "∀E"(1), THEN "≡E"(2)]
by blast
}
AOT_thus ‹[λμ⇩1...μ⇩n φ{μ⇩1...μ⇩n}]↓, κ⇩1...κ⇩n↓, φ{κ⇩1...κ⇩n} ❙⊢⇩□
[λμ⇩1...μ⇩n φ{μ⇩1...μ⇩n}]κ⇩1...κ⇩n›
by blast
qed
AOT_theorem "betaC:2:b":
‹[λμ⇩1...μ⇩n φ{μ⇩1...μ⇩n}]↓, κ⇩1...κ⇩n↓, ¬[λμ⇩1...μ⇩n φ{μ⇩1...μ⇩n}]κ⇩1...κ⇩n ❙⊢⇩□
¬φ{κ⇩1...κ⇩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⇩1...x⇩n [Π]x⇩1...x⇩n] = Π›
using "lambda-predicates:3"[axiom_inst] "∀I" "∀E"(1) "→I" by fast
AOT_theorem "eta-conversion-lemma1:2": ‹Π↓ → [λν⇩1...ν⇩n [Π]ν⇩1...ν⇩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⇩1...z⇩n χ{z⇩1...z⇩n, ❙ιx φ{x}}]↓ & ❙ιx φ{x} = ❙ιx ψ{x} →
[λz⇩1...z⇩n χ{z⇩1...z⇩n, ❙ιx φ{x}}] = [λz⇩1...z⇩n χ{z⇩1...z⇩n, ❙ιx ψ{x}}]›
proof(rule "→I")
AOT_assume A: ‹[λz⇩1...z⇩n χ{z⇩1...z⇩n, ❙ιx φ{x}}]↓ & ❙ιx φ{x} = ❙ιx ψ{x}›
AOT_show ‹[λz⇩1...z⇩n χ{z⇩1...z⇩n, ❙ιx φ{x}}] = [λz⇩1...z⇩n χ{z⇩1...z⇩n, ❙ιx ψ{x}}]›
using "rule=E"[where φ="λ τ . «[λz⇩1...z⇩n χ{z⇩1...z⇩n, ❙ιx φ{x}}] =
[λz⇩1...z⇩n χ{z⇩1...z⇩n, τ}]»",
OF "=I"(1)[OF A[THEN "&E"(1)]], OF A[THEN "&E"(2)]]
by blast
qed
AOT_theorem "sub-des-lam:2":
‹❙ιx φ{x} = ❙ιx ψ{x} → χ{❙ιx φ{x}} = χ{❙ιx ψ{x}}› for χ :: ‹κ ⇒ 𝗈›
using "rule=E"[where φ="λ τ . «χ{❙ι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⇩1...∀x⇩n ([F]x⇩1...x⇩n ≡ φ{x⇩1...x⇩n})›
apply (rule "∃I"(1)[where τ="«[λx⇩1...x⇩n φ{x⇩1...x⇩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]❙ι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]❙ιy(y = x & ∃H (x[H] & ¬[H]x)) ≡ ∃H (x[H] & ¬[H]x))›
proof(rule GEN)
fix x
AOT_have A: ‹[G]❙ιy (y = x & ∃H (x[H] & ¬[H]x)) ≡
∃!y (y = x & ∃H (x[H] & ¬[H]x))›
proof(rule "≡I"; rule "→I")
AOT_assume ‹[G]❙ιy (y = x & ∃H (x[H] & ¬[H]x))›
AOT_hence ‹❙ι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 "≡⇩d⇩fE", OF A] "&E" "∃E"[rotated] by blast
AOT_have a_3: ‹[G]a›
using antecedant "∀E" by blast
AOT_show ‹[G]❙ι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 "≡⇩d⇩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 "≡⇩d⇩fI"] "&I" "∃I"(2) by fast
qed
finally AOT_show ‹([G]❙ιy(y = x & ∃H (x[H] & ¬[H]x)) ≡ ∃H (x[H] & ¬[H]x))› .
qed
AOT_assume A: ‹[λx [G]❙ιy (y = x & ∃H (x[H] & ¬[H]x))]↓›
AOT_have θ: ‹∀x ([λx [G]❙ιy (y = x & ∃H (x[H] & ¬[H]x))]x ≡
[G]❙ιy(y = x & ∃H (x[H] & ¬[H]x)))›
using "beta-C-meta"[THEN "→E", OF A] "∀I" by fast
AOT_have ‹∀x ([λx [G]❙ι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 ❙𝒜[G]x → ¬[λx [G]❙ι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]❙ι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: ‹❙𝒜[G]❙ιy (y = x & ∃H (x[H] & ¬[H]x)) ≡
∃!y ❙𝒜(y = x & ∃H (x[H] & ¬[H]x))›
proof(rule "≡I"; rule "→I")
AOT_assume ‹❙𝒜[G]❙ιy (y = x & ∃H (x[H] & ¬[H]x))›
moreover AOT_have ‹□([G]❙ιy (y = x & ∃H (x[H] & ¬[H]x)) →
□❙ιy (y = x & ∃H (x[H] & ¬[H]x))↓)›
proof(rule RN; rule "→I")
AOT_modally_strict {
AOT_assume ‹[G]❙ιy (y = x & ∃H (x[H] & ¬[H]x))›
AOT_hence ‹❙ι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 "exist-nec"[THEN "→E"] by blast
}
qed
ultimately AOT_have ‹❙𝒜□❙ιy (y = x & ∃H (x[H] & ¬[H]x))↓›
using "act-cond"[THEN "→E", THEN "→E"] "nec-imp-act"[THEN "→E"] by blast
AOT_hence ‹❙ιy (y = x & ∃H (x[H] & ¬[H]x))↓›
using "Act-Sub:3" "B◇" "vdash-properties:10" by blast
AOT_thus ‹∃!y ❙𝒜(y = x & ∃H (x[H] & ¬[H]x))›
using "actual-desc: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 "≡⇩d⇩fE", OF A] "&E" "∃E"[rotated] by blast
AOT_have a_3: ‹❙𝒜[G]a›
using antecedant "∀E" by blast
moreover AOT_have ‹a = ❙ι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 ‹❙𝒜[G]❙ιy (y = x & ∃H (x[H] & ¬[H]x))›
using "rule=E" by fast
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 "≡⇩d⇩fE", OF A] "&E" "∃E"[rotated] by blast
AOT_thus ‹❙𝒜∃H (x[H] & ¬[H]x)›
using "Act-Basic:2"[THEN "≡E"(1), THEN "&E"(2)] by blast
next
AOT_assume ‹❙𝒜∃H (x[H] & ¬[H]x)›
AOT_hence ‹❙𝒜x = x & ❙𝒜∃H (x[H] & ¬[H]x)›
using "id-eq:1" "&I" "RA[2]" by blast
AOT_hence ‹❙𝒜(x = x & ∃H (x[H] & ¬[H]x))›
using "act-conj-act:3" "Act-Basic:2" "≡E" by blast
moreover AOT_have ‹∀z (❙𝒜(z = x & ∃H (x[H] & ¬[H]x)) → z = x)›
proof(safe intro!: GEN "→I")
fix z
AOT_assume ‹❙𝒜(z = x & ∃H (x[H] & ¬[H]x))›
AOT_hence ‹❙𝒜(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 ❙𝒜(y = x & ∃H (x[H] & ¬[H]x))›
using "uniqueness:1"[THEN "≡⇩d⇩fI"] "&I" "∃I"(2) by fast
qed
finally AOT_show ‹(❙𝒜[G]❙ιy(y = x & ∃H (x[H] & ¬[H]x)) ≡ ❙𝒜∃H (x[H] & ¬[H]x))›.
qed
AOT_assume A: ‹[λx [G]❙ιy (y = x & ∃H (x[H] & ¬[H]x))]↓›
AOT_hence ‹❙𝒜[λx [G]❙ιy (y = x & ∃H (x[H] & ¬[H]x))]↓›
using "exist-nec" "→E" "nec-imp-act"[THEN "→E"] by blast
AOT_hence ‹❙𝒜([λx [G]❙ιy (y = x & ∃H (x[H] & ¬[H]x))]↓ &
∀x ([G]❙ι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 ‹❙𝒜([λx [G]❙ιy (y = x & ∃H (x[H] & ¬[H]x))]↓ &
∀x ([G]❙ιy(y = x & ∃H (x[H] & ¬[H]x)) ≡ ∃H (x[H] & ¬[H]x)))
→ ❙𝒜∃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]❙ιy (y = x & ∃H (x[H] & ¬[H]x))]↓ &
∀x ([G]❙ι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]❙ι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]❙ι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]❙ι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]❙ι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 ‹❙𝒜∃p (p & ¬p)›
using "→E" by blast
AOT_hence ‹∃p ❙𝒜(p & ¬p)›
by (metis "Act-Basic:10" "intro-elim:3:a")
then AOT_obtain p where ‹❙𝒜(p & ¬p)›
using "∃E"[rotated] by blast
moreover AOT_have ‹¬❙𝒜(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]❙ι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]❙ιy (y = x & ∃H (x[H] & ¬[H]x))]↓)›
using "block-paradox2:1" "∀I" by fast
ultimately AOT_show ‹¬[λx [λx ∀p (p →p)]❙ι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⇩1...∀x⇩n ([F]x⇩1...x⇩n ≡ [G]x⇩1...x⇩n)) → F ≠ G›
proof (rule "→I")
AOT_assume ‹◇¬∀x⇩1...∀x⇩n ([F]x⇩1...x⇩n ≡ [G]x⇩1...x⇩n)›
AOT_hence ‹¬□∀x⇩1...∀x⇩n ([F]x⇩1...x⇩n ≡ [G]x⇩1...x⇩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 "≡⇩d⇩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⇩1...∀x⇩n ([F]x⇩1...x⇩n ≡ [G]x⇩1...x⇩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 :: "Π ⇒ Π" ("_⇧-")
"df-relation-negation": "[F]⇧- =⇩d⇩f [λx⇩1...x⇩n ¬[F]x⇩1...x⇩n]"
nonterminal φneg
syntax "" :: "φneg ⇒ τ" ("_")
syntax "" :: "φneg ⇒ φ" ("'(_')")
AOT_define relation_negation_0 :: ‹φ ⇒ φneg› ("'(_')⇧-")
"df-relation-negation[zero]": "(p)⇧- =⇩d⇩f [λ ¬p]"
AOT_theorem "rel-neg-T:1": ‹[λx⇩1...x⇩n ¬[Π]x⇩1...x⇩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": ‹[Π]⇧- = [λx⇩1...x⇩n ¬[Π]x⇩1...x⇩n]›
using "=I"(1)[OF "rel-neg-T:1"]
by (rule "=⇩d⇩fI"(1)[OF "df-relation-negation", OF "rel-neg-T:1"])
AOT_theorem "rel-neg-T:2[zero]": ‹(φ)⇧- = [λ ¬φ]›
using "=I"(1)[OF "rel-neg-T:1[zero]"]
by (rule "=⇩d⇩fI"(1)[OF "df-relation-negation[zero]", OF "rel-neg-T:1[zero]"])
AOT_theorem "rel-neg-T:3": ‹[Π]⇧-↓›
using "=⇩d⇩fI"(1)[OF "df-relation-negation", OF "rel-neg-T:1"]
"rel-neg-T:1" by blast
AOT_theorem "rel-neg-T:3[zero]": ‹(φ)⇧-↓›
using "log-prop-prop:2" by blast
AOT_theorem "thm-relation-negation:1": ‹[F]⇧-x⇩1...x⇩n ≡ ¬[F]x⇩1...x⇩n›
proof -
AOT_have ‹[F]⇧-x⇩1...x⇩n ≡ [λx⇩1...x⇩n ¬[F]x⇩1...x⇩n]x⇩1...x⇩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⇩1...x⇩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]⇧-x⇩1...x⇩n ≡ [F]x⇩1...x⇩n›
apply (AOT_subst ‹[F]x⇩1...x⇩n› ‹¬¬[F]x⇩1...x⇩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)⇧-) ≡ ¬p›
proof -
AOT_have ‹(p)⇧- = [λ ¬p]› using "rel-neg-T:2[zero]" by blast
AOT_hence ‹((p)⇧-) ≡ [λ ¬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)⇧-)) ≡ 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]⇧-›
proof -
AOT_have ‹¬([F] = [F]⇧-)›
proof (rule RAA(2))
AOT_show ‹[F]x⇩1...x⇩n → [F]x⇩1...x⇩n› for x⇩1x⇩n
using "if-p-then-p".
next
AOT_assume ‹[F] = [F]⇧-›
AOT_hence ‹[F]⇧- = [F]› using id_sym by blast
AOT_hence ‹[F]x⇩1...x⇩n ≡ ¬[F]x⇩1...x⇩n› for x⇩1x⇩n
using "rule=E" "thm-relation-negation:1" by fast
AOT_thus ‹¬([F]x⇩1...x⇩n → [F]x⇩1...x⇩n)› for x⇩1x⇩n
using "≡E" RAA by metis
qed
thus ?thesis
using "≡⇩d⇩fI" "=-infix" by blast
qed
AOT_theorem "thm-relation-negation:6": ‹p ≠ (p)⇧-›
proof -
AOT_have ‹¬(p = (p)⇧-)›
proof (rule RAA(2))
AOT_show ‹p → p›
using "if-p-then-p".
next
AOT_assume ‹p = (p)⇧-›
AOT_hence ‹(p)⇧- = 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 "≡⇩d⇩fI" "=-infix" by blast
qed
AOT_theorem "thm-relation-negation:7": ‹(p)⇧- = (¬p)›
apply (rule "df-relation-negation[zero]"[THEN "=⇩d⇩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)⇧- = (q)⇧-›
proof(rule "→I")
AOT_assume ‹p = q›
AOT_hence ‹(¬p) = (¬q)› using "thm-relation-negation:8" "→E" by blast
AOT_thus ‹(p)⇧- = (q)⇧-›
using "thm-relation-negation:7" id_sym id_trans by metis
qed
AOT_define Necessary :: ‹Π ⇒ φ› ("Necessary'(_')")
"contingent-properties:1":
‹Necessary([F]) ≡⇩d⇩f □∀x⇩1...∀x⇩n [F]x⇩1...x⇩n›
AOT_define Necessary0 :: ‹φ ⇒ φ› ("Necessary0'(_')")
"contingent-properties:1[zero]":
‹Necessary0(p) ≡⇩d⇩f □p›
AOT_define Impossible :: ‹Π ⇒ φ› ("Impossible'(_')")
"contingent-properties:2":
‹Impossible([F]) ≡⇩d⇩f F↓ & □∀x⇩1...∀x⇩n ¬[F]x⇩1...x⇩n›
AOT_define Impossible0 :: ‹φ ⇒ φ› ("Impossible0'(_')")
"contingent-properties:2[zero]":
‹Impossible0(p) ≡⇩d⇩f □¬p›
AOT_define NonContingent :: ‹Π ⇒ φ› ("NonContingent'(_')")
"contingent-properties:3":
‹NonContingent([F]) ≡⇩d⇩f Necessary([F]) ∨ Impossible([F])›
AOT_define NonContingent0 :: ‹φ ⇒ φ› ("NonContingent0'(_')")
"contingent-properties:3[zero]":
‹NonContingent0(p) ≡⇩d⇩f Necessary0(p) ∨ Impossible0(p)›
AOT_define Contingent :: ‹Π ⇒ φ› ("Contingent'(_')")
"contingent-properties:4":
‹Contingent([F]) ≡⇩d⇩f F↓ & ¬(Necessary([F]) ∨ Impossible([F]))›
AOT_define Contingent0 :: ‹φ ⇒ φ› ("Contingent0'(_')")
"contingent-properties:4[zero]":
‹Contingent0(p) ≡⇩d⇩f ¬(Necessary0(p) ∨ Impossible0(p))›
AOT_theorem "thm-cont-prop:1": ‹NonContingent([F]) ≡ NonContingent([F]⇧-)›
proof (rule "≡I"; rule "→I")
AOT_assume ‹NonContingent([F])›
AOT_hence ‹Necessary([F]) ∨ Impossible([F])›
using "≡⇩d⇩fE"[OF "contingent-properties:3"] by blast
moreover {
AOT_assume ‹Necessary([F])›
AOT_hence ‹□(∀x⇩1...∀x⇩n [F]x⇩1...x⇩n)›
using "≡⇩d⇩fE"[OF "contingent-properties:1"] by blast
moreover AOT_modally_strict {
AOT_assume ‹∀x⇩1...∀x⇩n [F]x⇩1...x⇩n›
AOT_hence ‹[F]x⇩1...x⇩n› for x⇩1x⇩n using "∀E" by blast
AOT_hence ‹¬[F]⇧-x⇩1...x⇩n› for x⇩1x⇩n
by (meson "≡E"(6) "oth-class-taut:3:a"
"thm-relation-negation:2" "≡E"(1))
AOT_hence ‹∀x⇩1...∀x⇩n ¬[F]⇧-x⇩1...x⇩n› using "∀I" by fast
}
ultimately AOT_have ‹□(∀x⇩1...∀x⇩n ¬[F]⇧-x⇩1...x⇩n)›
using "RN[prem]"[where Γ="{«∀x⇩1...∀x⇩n [F]x⇩1...x⇩n»}", simplified] by blast
AOT_hence ‹Impossible([F]⇧-)›
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⇩1...∀x⇩n ¬[F]x⇩1...x⇩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⇩1...∀x⇩n ¬[F]x⇩1...x⇩n›
AOT_hence ‹¬[F]x⇩1...x⇩n› for x⇩1x⇩n using "∀E" by blast
AOT_hence ‹[F]⇧-x⇩1...x⇩n› for x⇩1x⇩n
by (meson "≡E"(6) "oth-class-taut:3:a"
"thm-relation-negation:1" "≡E"(1))
AOT_hence ‹∀x⇩1...∀x⇩n [F]⇧-x⇩1...x⇩n› using "∀I" by fast
}
ultimately AOT_have ‹□(∀x⇩1...∀x⇩n [F]⇧-x⇩1...x⇩n)›
using "RN[prem]"[where Γ="{«∀x⇩1...∀x⇩n ¬[F]x⇩1...x⇩n»}"] by blast
AOT_hence ‹Necessary([F]⇧-)›
using "≡⇩d⇩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 "≡⇩d⇩fI"[OF "contingent-properties:3"] by blast
next
AOT_assume ‹NonContingent([F]⇧-)›
AOT_hence ‹Necessary([F]⇧-) ∨ Impossible([F]⇧-)›
using "≡⇩d⇩fE"[OF "contingent-properties:3"] by blast
moreover {
AOT_assume ‹Necessary([F]⇧-)›
AOT_hence ‹□(∀x⇩1...∀x⇩n [F]⇧-x⇩1...x⇩n)›
using "≡⇩d⇩fE"[OF "contingent-properties:1"] by blast
moreover AOT_modally_strict {
AOT_assume ‹∀x⇩1...∀x⇩n [F]⇧-x⇩1...x⇩n›
AOT_hence ‹[F]⇧-x⇩1...x⇩n› for x⇩1x⇩n using "∀E" by blast
AOT_hence ‹¬[F]x⇩1...x⇩n› for x⇩1x⇩n
by (meson "≡E"(6) "oth-class-taut:3:a"
"thm-relation-negation:1" "≡E"(2))
AOT_hence ‹∀x⇩1...∀x⇩n ¬[F]x⇩1...x⇩n› using "∀I" by fast
}
ultimately AOT_have ‹□∀x⇩1...∀x⇩n ¬[F]x⇩1...x⇩n›
using "RN[prem]"[where Γ="{«∀x⇩1...∀x⇩n [F]⇧-x⇩1...x⇩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]⇧-)›
AOT_hence ‹□(∀x⇩1...∀x⇩n ¬[F]⇧-x⇩1...x⇩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⇩1...∀x⇩n ¬[F]⇧-x⇩1...x⇩n›
AOT_hence ‹¬[F]⇧-x⇩1...x⇩n› for x⇩1x⇩n using "∀E" by blast
AOT_hence ‹[F]x⇩1...x⇩n› for x⇩1x⇩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⇩1...∀x⇩n [F]x⇩1...x⇩n› using "∀I" by fast
}
ultimately AOT_have ‹□(∀x⇩1...∀x⇩n [F]x⇩1...x⇩n)›
using "RN[prem]"[where Γ="{«∀x⇩1...∀x⇩n ¬[F]⇧-x⇩1...x⇩n»}"] by blast
AOT_hence ‹Necessary([F])›
using "≡⇩d⇩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 "≡⇩d⇩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]⇧-)› 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]⇧-x & ◇∃x [F]x›
by (AOT_subst ‹[F]⇧-x› ‹¬[F]x› for: x)
(auto simp: "thm-relation-negation:1" "oth-class-taut:3:a")
also AOT_have ‹... ≡ ◇∃x [F]⇧-x & ◇∃x ¬[F]⇧-x›
by (AOT_subst (reverse) ‹[F]x› ‹¬[F]⇧-x› for: x)
(auto simp: "thm-relation-negation:2" "oth-class-taut:3:a")
also AOT_have ‹... ≡ Contingent([F]⇧-)›
using 1[OF "rel-neg-T:3", symmetric] by blast
finally show ?thesis.
qed
AOT_define concrete_if_concrete :: ‹Π› ("L")
L_def: ‹L =⇩d⇩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 "=⇩d⇩fI"(2)[OF L_def])
apply "cqt:2[lambda]"
by (rule "contingent-properties:1"[THEN "≡⇩d⇩fI", OF 0])
qed
AOT_theorem "thm-noncont-e-e:2": ‹Impossible([L]⇧-)›
proof -
AOT_modally_strict {
fix x
AOT_have 0: ‹∀F (¬[F]⇧-x ≡ [F]x)›
using "thm-relation-negation:2" GEN by fast
AOT_have ‹¬[λx E!x → E!x]⇧-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]⇧-x›
using "≡E" by blast
}
AOT_hence 0: ‹□∀x ¬[λx E!x → E!x]⇧-x›
using RN GEN by fast
show ?thesis
apply (rule "=⇩d⇩fI"(2)[OF L_def])
apply "cqt:2[lambda]"
apply (rule "contingent-properties:2"[THEN "≡⇩d⇩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 "≡⇩d⇩fI", OF "∨I"(1)])
AOT_theorem "thm-noncont-e-e:4": ‹NonContingent([L]⇧-)›
proof -
AOT_have 0: ‹∀F (NonContingent([F]) ≡ NonContingent([F]⇧-))›
using "thm-cont-prop:1" "∀I" by fast
moreover AOT_have 1: ‹L↓›
by (rule "=⇩d⇩fI"(2)[OF L_def]) "cqt:2[lambda]"+
AOT_show ‹NonContingent([L]⇧-)›
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]⇧-›
using "thm-relation-negation:5" GEN by fast
moreover AOT_have ‹L↓›
by (rule "=⇩d⇩fI"(2)[OF L_def]) "cqt:2[lambda]"+
ultimately AOT_have ‹L ≠ [L]⇧-›
using "∀E" by blast
}
AOT_thus ‹L ≠ [L]⇧- & NonContingent(L) & NonContingent([L]⇧-)›
using "thm-noncont-e-e:3" "thm-noncont-e-e:4" "&I" by metis
next
AOT_show ‹[L]⇧-↓›
using "rel-neg-T:3" by blast
next
AOT_show ‹L↓›
by (rule "=⇩d⇩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]⇧-x & ◇¬[F]⇧-x)›
proof -
AOT_have ‹◇∃x ([F]x & ◇¬[F]x) ≡ ◇∃x (¬[F]x & ◇[F]x)›
using "lem-cont-e:1".
also AOT_have ‹… ≡ ◇∃x ([F]⇧-x & ◇¬[F]⇧-x)›
apply (AOT_subst ‹¬[F]⇧-x› ‹[F]x› for: x)
apply (simp add: "thm-relation-negation:2")
apply (AOT_subst ‹[F]⇧-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 & ¬❙𝒜E!x)›
using "qml:4"[axiom_inst] "BF◇"[THEN "→E"] by blast
then AOT_obtain a where ‹◇(E!a & ¬❙𝒜E!a)›
using "∃E"[rotated] by blast
AOT_hence θ: ‹◇E!a & ◇¬❙𝒜E!a›
using "KBasic2:3"[THEN "→E"] by blast
AOT_have ξ: ‹◇E!a & ◇❙𝒜¬E!a›
by (AOT_subst ‹❙𝒜¬E!a› ‹¬❙𝒜E!a›)
(auto simp: "logic-actual-nec:1"[axiom_inst] θ)
AOT_have ζ: ‹◇E!a & ❙𝒜¬E!a›
by (AOT_subst ‹❙𝒜¬E!a› ‹◇❙𝒜¬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!]⇧-)›
proof -
AOT_have ‹∀F (Contingent([«F::<κ>»]) ≡ Contingent([F]⇧-))›
using "thm-cont-prop:3" GEN by fast
AOT_hence ‹Contingent([E!]) ≡ Contingent([E!]⇧-)›
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]⇧-›
using "thm-relation-negation:5" GEN by fast
AOT_hence ‹[E!] ≠ [E!]⇧-›
using "∀E" by fast
AOT_thus ‹Contingent([E!]) & Contingent([E!]⇧-) & [E!] ≠ [E!]⇧-›
using "thm-cont-e:5" "thm-cont-e:6" "&I" by metis
next
AOT_show ‹E!⇧-↓›
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 "≡⇩d⇩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 "≡⇩d⇩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]⇧- & L ≠ E! & L ≠ E!⇧- & [L]⇧- ≠ [E!]⇧- & E! ≠ [E!]⇧-›
proof -
AOT_have noneqI: ‹Π ≠ Π'› if ‹φ{Π}› and ‹¬φ{Π'}› for φ and Π Π' :: ‹<κ>›
apply (rule "=-infix"[THEN "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩fE"] by blast
qed
show ?thesis
proof (safe intro!: "&I")
AOT_show ‹L ≠ [L]⇧-›
apply (rule "=⇩d⇩fI"(2)[OF L_def])
apply "cqt:2[lambda]"
apply (rule "∀E"(1)[where φ="λ Π . «Π ≠ [Π]⇧-»"])
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!⇧-›
apply (rule noneqI)
using "thm-noncont-e-e:3" apply fast
apply (rule not_noncontingent_if_contingent)
apply (rule "∀E"(1)[
where φ="λ Π . «Contingent([Π]) ≡ Contingent([Π]⇧-)»",
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]⇧- ≠ E!⇧-›
apply (rule noneqI)
using "thm-noncont-e-e:4" apply fast
apply (rule not_noncontingent_if_contingent)
apply (rule "∀E"(1)[
where φ="λ Π . «Contingent([Π]) ≡ Contingent([Π]⇧-)»",
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!⇧-›
apply (rule "=⇩d⇩fI"(2)[OF L_def])
apply "cqt:2[lambda]"
apply (rule "∀E"(1)[where φ="λ Π . «Π ≠ [Π]⇧-»"])
apply (rule GEN) apply (fact AOT)
by "cqt:2"
qed
qed
AOT_theorem "thm-cont-propos:1":
‹NonContingent0(p) ≡ NonContingent0(((p)⇧-))›
proof(rule "≡I"; rule "→I")
AOT_assume ‹NonContingent0(p)›
AOT_hence ‹Necessary0(p) ∨ Impossible0(p)›
using "contingent-properties:3[zero]"[THEN "≡⇩d⇩fE"] by blast
moreover {
AOT_assume ‹Necessary0(p)›
AOT_hence 1: ‹□p›
using "contingent-properties:1[zero]"[THEN "≡⇩d⇩fE"] by blast
AOT_have ‹□¬((p)⇧-)›
by (AOT_subst ‹¬((p)⇧-)› ‹p›)
(auto simp add: 1 "thm-relation-negation:4")
AOT_hence ‹Impossible0(((p)⇧-))›
by (rule "contingent-properties:2[zero]"[THEN "≡⇩d⇩fI"])
}
moreover {
AOT_assume ‹Impossible0(p)›
AOT_hence 1: ‹□¬p›
by (rule "contingent-properties:2[zero]"[THEN "≡⇩d⇩fE"])
AOT_have ‹□((p)⇧-)›
by (AOT_subst ‹((p)⇧-)› ‹¬p›)
(auto simp: 1 "thm-relation-negation:3")
AOT_hence ‹Necessary0(((p)⇧-))›
by (rule "contingent-properties:1[zero]"[THEN "≡⇩d⇩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 "≡⇩d⇩fI"] by blast
next
AOT_assume ‹NonContingent0(((p)⇧-))›
AOT_hence ‹Necessary0(((p)⇧-)) ∨ Impossible0(((p)⇧-))›
using "contingent-properties:3[zero]"[THEN "≡⇩d⇩fE"] by blast
moreover {
AOT_assume ‹Impossible0(((p)⇧-))›
AOT_hence 1: ‹□¬((p)⇧-)›
by (rule "contingent-properties:2[zero]"[THEN "≡⇩d⇩fE"])
AOT_have ‹□p›
by (AOT_subst (reverse) ‹p› ‹¬((p)⇧-)›)
(auto simp: 1 "thm-relation-negation:4")
AOT_hence ‹Necessary0(p)›
using "contingent-properties:1[zero]"[THEN "≡⇩d⇩fI"] by blast
}
moreover {
AOT_assume ‹Necessary0(((p)⇧-))›
AOT_hence 1: ‹□((p)⇧-)›
by (rule "contingent-properties:1[zero]"[THEN "≡⇩d⇩fE"])
AOT_have ‹□¬p›
by (AOT_subst (reverse) ‹¬p› ‹((p)⇧-)›)
(auto simp: 1 "thm-relation-negation:3")
AOT_hence ‹Impossible0(p)›
by (rule "contingent-properties:2[zero]"[THEN "≡⇩d⇩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 "≡⇩d⇩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)⇧-))›
proof -
AOT_have ‹Contingent0(p) ≡ ◇p & ◇¬p› using "thm-cont-propos:2".
also AOT_have ‹… ≡ ◇¬p & ◇p› by (fact AOT)
also AOT_have ‹… ≡ ◇((p)⇧-) & ◇p›
by (AOT_subst ‹((p)⇧-)› ‹¬p›)
(auto simp: "thm-relation-negation:3" "oth-class-taut:3:a")
also AOT_have ‹… ≡ ◇((p)⇧-) & ◇¬((p)⇧-)›
by (AOT_subst ‹¬((p)⇧-)› ‹p›)
(auto simp: "thm-relation-negation:4" "oth-class-taut:3:a")
also AOT_have ‹… ≡ Contingent0(((p)⇧-))›
using "thm-cont-propos:2"[symmetric] by blast
finally show ?thesis.
qed
AOT_define noncontingent_prop :: ‹φ› ("p⇩0")
p⇩0_def: "(p⇩0) =⇩d⇩f (∀x (E!x → E!x))"
AOT_theorem "thm-noncont-propos:1": ‹Necessary0((p⇩0))›
proof(rule "contingent-properties:1[zero]"[THEN "≡⇩d⇩fI"])
AOT_show ‹□(p⇩0)›
apply (rule "=⇩d⇩fI"(2)[OF p⇩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⇩0)⇧-))›
proof(rule "contingent-properties:2[zero]"[THEN "≡⇩d⇩fI"])
AOT_show ‹□¬((p⇩0)⇧-)›
apply (AOT_subst ‹((p⇩0)⇧-)› ‹¬p⇩0›)
using "thm-relation-negation:3" GEN "∀E"(1)[rotated, OF "log-prop-prop:2"]
apply fast
apply (AOT_subst (reverse) ‹¬¬p⇩0› ‹p⇩0›)
apply (simp add: "oth-class-taut:3:b")
apply (rule "=⇩d⇩fI"(2)[OF p⇩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⇩0))›
apply(rule "contingent-properties:3[zero]"[THEN "≡⇩d⇩fI"])
using "thm-noncont-propos:1" "∨I" by blast
AOT_theorem "thm-noncont-propos:4": ‹NonContingent0(((p⇩0)⇧-))›
apply(rule "contingent-properties:3[zero]"[THEN "≡⇩d⇩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: ‹φ ≠ (φ)⇧-› for φ
using "thm-relation-negation:6" "∀I"
"∀E"(1)[rotated, OF "log-prop-prop:2"] by fast
AOT_thus ‹NonContingent0((p⇩0)) & NonContingent0(((p⇩0)⇧-)) & (p⇩0) ≠ (p⇩0)⇧-›
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 & ¬❙𝒜E!x)›
proof(rule "raa-cor:2")
AOT_assume ‹∃x(E!x & ¬❙𝒜E!x)›
then AOT_obtain a where a: ‹E!a & ¬❙𝒜E!a›
using "∃E"[rotated] by blast
AOT_hence ‹❙𝒜¬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": ‹¬❙𝒜∃x (E!x & ¬❙𝒜E!x)›
proof(rule "raa-cor:2")
AOT_assume ‹❙𝒜∃x (E!x & ¬❙𝒜E!x)›
AOT_hence ‹∃x ❙𝒜(E!x & ¬❙𝒜E!x)›
using "Act-Basic:10"[THEN "≡E"(1)] by blast
then AOT_obtain a where ‹❙𝒜(E!a & ¬❙𝒜E!a)›
using "∃E"[rotated] by blast
AOT_hence 1: ‹❙𝒜E!a & ❙𝒜¬❙𝒜E!a›
using "Act-Basic:2"[THEN "≡E"(1)] by blast
AOT_hence ‹¬❙𝒜❙𝒜E!a›
using "&E"(2) "logic-actual-nec:1"[axiom_inst, THEN "≡E"(1)] by blast
AOT_hence ‹¬❙𝒜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 & ¬❙𝒜E!x)›
proof (rule RAA(1))
AOT_show ‹¬❙𝒜∃x (E!x & ¬❙𝒜E!x)›
using "pos-not-pna:1" by blast
next
AOT_assume ‹¬◇¬∃x (E!x & ¬❙𝒜E!x)›
AOT_hence ‹□∃x (E!x & ¬❙𝒜E!x)›
using "KBasic:12"[THEN "≡E"(2)] by blast
AOT_thus ‹❙𝒜∃x (E!x & ¬❙𝒜E!x)›
using "nec-imp-act"[THEN "→E"] by blast
qed
AOT_theorem "pos-not-pna:3": ‹∃x (◇E!x & ¬❙𝒜E!x)›
proof -
AOT_obtain a where ‹◇(E!a & ¬❙𝒜E!a)›
using "qml:4"[axiom_inst] "BF◇"[THEN "→E"] "∃E"[rotated] by blast
AOT_hence θ: ‹◇E!a› and ξ: ‹◇¬❙𝒜E!a›
using "KBasic2:3"[THEN "→E"] "&E" by blast+
AOT_have ‹¬□❙𝒜E!a›
using ξ "KBasic:11"[THEN "≡E"(2)] by blast
AOT_hence ‹¬❙𝒜E!a›
using "Act-Basic:6"[THEN "oth-class-taut:4:b"[THEN "≡E"(1)],
THEN "≡E"(2)] by blast
AOT_hence ‹◇E!a & ¬❙𝒜E!a› using θ "&I" by blast
thus ?thesis using "∃I" by fast
qed
AOT_define contingent_prop :: φ ("q⇩0")
q⇩0_def: ‹(q⇩0) =⇩d⇩f (∃x (E!x & ¬❙𝒜E!x))›
AOT_theorem q⇩0_prop: ‹◇q⇩0 & ◇¬q⇩0›
apply (rule "=⇩d⇩fI"(2)[OF q⇩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⇩0))›
proof(rule "contingent-properties:4[zero]"[THEN "≡⇩d⇩fI"])
AOT_have ‹¬Necessary0((q⇩0)) & ¬Impossible0((q⇩0))›
proof (rule "&I";
rule "=⇩d⇩fI"(2)[OF q⇩0_def];
(rule "log-prop-prop:2" | rule "raa-cor:2"))
AOT_assume ‹Necessary0(∃x (E!x & ¬❙𝒜E!x))›
AOT_hence ‹□∃x (E!x & ¬❙𝒜E!x)›
using "contingent-properties:1[zero]"[THEN "≡⇩d⇩fE"] by blast
AOT_hence ‹❙𝒜∃x (E!x & ¬❙𝒜E!x)›
using "Act-Basic:8"[THEN "→E"] "qml:2"[axiom_inst, THEN "→E"] by blast
AOT_thus ‹❙𝒜∃x (E!x & ¬❙𝒜E!x) & ¬❙𝒜∃x (E!x & ¬❙𝒜E!x)›
using "pos-not-pna:1" "&I" by blast
next
AOT_assume ‹Impossible0(∃x (E!x & ¬❙𝒜E!x))›
AOT_hence ‹□¬(∃x (E!x & ¬❙𝒜E!x))›
using "contingent-properties:2[zero]"[THEN "≡⇩d⇩fE"] by blast
AOT_hence ‹¬◇(∃x (E!x & ¬❙𝒜E!x))›
using "KBasic2:1"[THEN "≡E"(1)] by blast
AOT_thus ‹◇(∃x (E!x & ¬❙𝒜E!x)) & ¬◇(∃x (E!x & ¬❙𝒜E!x))›
using "qml:4"[axiom_inst] "&I" by blast
qed
AOT_thus ‹¬(Necessary0((q⇩0)) ∨ Impossible0((q⇩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⇩0)⇧-))›
apply (AOT_subst ‹((q⇩0)⇧-)› ‹¬q⇩0›)
apply (insert "thm-relation-negation:3" "∀I"
"∀E"(1)[rotated, OF "log-prop-prop:2"]; fast)
apply (rule "contingent-properties:4[zero]"[THEN "≡⇩d⇩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 "≡⇩d⇩fE"])
apply (rule "=⇩d⇩fE"(2)[OF q⇩0_def])
apply (rule "log-prop-prop:2")
apply (rule q⇩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 "≡⇩d⇩fE"])
by (rule q⇩0_prop[THEN "&E"(2)])
AOT_theorem "basic-prop:4":
‹∃p∃q (p ≠ q & Contingent0(p) & Contingent0(q))›
proof(rule "∃I")+
AOT_have 0: ‹φ ≠ (φ)⇧-› for φ
using "thm-relation-negation:6" "∀I"
"∀E"(1)[rotated, OF "log-prop-prop:2"] by fast
AOT_show ‹(q⇩0) ≠ (q⇩0)⇧- & Contingent0(q⇩0) & Contingent0(((q⇩0)⇧-))›
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 "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩fE"] 1 "&I" by blast
qed
AOT_theorem "proposition-facts:3":
‹(p⇩0) ≠ (p⇩0)⇧- & (p⇩0) ≠ (q⇩0) & (p⇩0) ≠ (q⇩0)⇧- & (p⇩0)⇧- ≠ (q⇩0)⇧- & (q⇩0) ≠ (q⇩0)⇧-›
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(((φ)⇧-))› 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 "≡⇩d⇩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) ≡⇩d⇩f p & ◇¬p›
AOT_define ContingentlyFalse :: ‹φ ⇒ φ› ("ContingentlyFalse'(_')")
"cont-tf:2": ‹ContingentlyFalse(p) ≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩fI"] by blast
qed
AOT_theorem "cont-true-cont:3":
‹ContingentlyTrue((p)) ≡ ContingentlyFalse(((p)⇧-))›
proof(rule "≡I"; rule "→I")
AOT_assume ‹ContingentlyTrue((p))›
AOT_hence 0: ‹p & ◇¬p› using "cont-tf:1"[THEN "≡⇩d⇩fE"] by blast
AOT_have 1: ‹ContingentlyFalse(¬p)›
apply (rule "cont-tf:2"[THEN "≡⇩d⇩fI"])
apply (AOT_subst (reverse) ‹¬¬p› p)
by (auto simp: "oth-class-taut:3:b" 0)
AOT_show ‹ContingentlyFalse(((p)⇧-))›
apply (AOT_subst ‹((p)⇧-)› ‹¬p›)
by (auto simp: "thm-relation-negation:3" 1)
next
AOT_assume 1: ‹ContingentlyFalse(((p)⇧-))›
AOT_have ‹ContingentlyFalse(¬p)›
by (AOT_subst (reverse) ‹¬p› ‹((p)⇧-)›)
(auto simp: "thm-relation-negation:3" 1)
AOT_hence ‹¬¬p & ◇¬p› using "cont-tf:2"[THEN "≡⇩d⇩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 "≡⇩d⇩fI"] by blast
qed
AOT_theorem "cont-true-cont:4":
‹ContingentlyFalse((p)) ≡ ContingentlyTrue(((p)⇧-))›
proof(rule "≡I"; rule "→I")
AOT_assume ‹ContingentlyFalse(p)›
AOT_hence 0: ‹¬p & ◇p›
using "cont-tf:2"[THEN "≡⇩d⇩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 "≡⇩d⇩fI"])
AOT_show ‹ContingentlyTrue(((p)⇧-))›
by (AOT_subst ‹((p)⇧-)› ‹¬p›)
(auto simp: "thm-relation-negation:3" 1)
next
AOT_assume 1: ‹ContingentlyTrue(((p)⇧-))›
AOT_have ‹ContingentlyTrue(¬p)›
by (AOT_subst (reverse) ‹¬p› ‹((p)⇧-)›)
(auto simp add: "thm-relation-negation:3" 1)
AOT_hence 2: ‹¬p & ◇¬¬p› using "cont-tf:1"[THEN "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩fE"] "&E" by blast
AOT_hence 1: ‹¬□¬p›
using "conventions:5"[THEN "≡⇩d⇩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 "≡⇩d⇩fE"] by blast
AOT_thus ‹□¬p & ¬□¬p› using 1 "&I" by blast
qed
AOT_act_theorem "q0cf:1": ‹ContingentlyFalse(q⇩0)›
apply (rule "cont-tf:2"[THEN "≡⇩d⇩fI"])
apply (rule "=⇩d⇩fI"(2)[OF q⇩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⇩0)⇧-))›
apply (rule "cont-tf:1"[THEN "≡⇩d⇩fI"])
apply (rule "=⇩d⇩fI"(2)[OF q⇩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 & ¬❙𝒜E!x))› ‹∃x (E!x & ¬❙𝒜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⇩0›
AOT_hence ‹q⇩0 & ◇¬q⇩0› using q⇩0_prop[THEN "&E"(2)] "&I" by blast
AOT_thus ‹ContingentlyTrue(q⇩0)›
by (rule "cont-tf:1"[THEN "≡⇩d⇩fI"])
next
AOT_assume ‹¬q⇩0›
AOT_hence ‹¬q⇩0 & ◇q⇩0› using q⇩0_prop[THEN "&E"(1)] "&I" by blast
AOT_hence ‹ContingentlyFalse(q⇩0)›
by (rule "cont-tf:2"[THEN "≡⇩d⇩fI"])
AOT_thus ‹ContingentlyTrue(((q⇩0)⇧-))›
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⇩0›
AOT_hence ‹q⇩0 & ◇¬q⇩0› using q⇩0_prop[THEN "&E"(2)] "&I" by blast
AOT_hence ‹ContingentlyTrue(q⇩0)›
by (rule "cont-tf:1"[THEN "≡⇩d⇩fI"])
AOT_thus ‹ContingentlyFalse(((q⇩0)⇧-))›
by (rule "cont-true-cont:3"[unvarify p,
OF "log-prop-prop:2", THEN "≡E"(1)])
next
AOT_assume ‹¬q⇩0›
AOT_hence ‹¬q⇩0 & ◇q⇩0› using q⇩0_prop[THEN "&E"(1)] "&I" by blast
AOT_thus ‹ContingentlyFalse(q⇩0)›
by (rule "cont-tf:2"[THEN "≡⇩d⇩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⇩1 where ‹ContingentlyTrue((p⇩1))›
using "cont-tf-thm:1" "∃E"[rotated] by blast
AOT_hence 1: ‹p⇩1 & ◇¬p⇩1› using "cont-tf:1"[THEN "≡⇩d⇩fE"] by blast
AOT_modally_strict {
AOT_have ‹for arbitrary p: ❙⊢⇩□ ([λz p]x ≡ p)›
by (rule "beta-C-cor:3"[THEN "∀E"(2)]) cqt_2_lambda_inst_prover
AOT_hence ‹for arbitrary p: ❙⊢⇩□ □ ([λz p]x ≡ p)›
by (rule RN)
AOT_hence ‹∀p □([λz p]x ≡ p)› using GEN by fast
AOT_hence ‹□([λz p⇩1]x ≡ p⇩1)› using "∀E" by fast
} note 2 = this
AOT_hence ‹□([λz p⇩1]x ≡ p⇩1)› using "∀E" by blast
AOT_hence ‹[λz p⇩1]x›
using 1[THEN "&E"(1)] "qml:2"[axiom_inst, THEN "→E"] "≡E"(2) by blast
moreover AOT_have ‹◇¬[λz p⇩1]x›
using 2[THEN "qml:2"[axiom_inst, THEN "→E"]]
apply (AOT_subst ‹[λz p⇩1]x› ‹p⇩1›)
using 1[THEN "&E"(2)] by blast
ultimately AOT_have ‹[λz p⇩1]x & ◇¬[λz p⇩1]x› using "&I" by blast
AOT_hence ‹∃x ([λz p⇩1]x & ◇¬[λz p⇩1]x)› using "∃I"(2) by fast
moreover AOT_have ‹[λz p⇩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⇩1 where ‹ContingentlyFalse((p⇩1))›
using "cont-tf-thm:2" "∃E"[rotated] by blast
AOT_hence 1: ‹¬p⇩1 & ◇p⇩1› using "cont-tf:2"[THEN "≡⇩d⇩fE"] by blast
AOT_modally_strict {
AOT_have ‹for arbitrary p: ❙⊢⇩□ ([λz p]x ≡ p)›
by (rule "beta-C-cor:3"[THEN "∀E"(2)]) cqt_2_lambda_inst_prover
AOT_hence ‹for arbitrary p: ❙⊢⇩□ (¬[λz p]x ≡ ¬p)›
using "oth-class-taut:4:b" "≡E" by blast
AOT_hence ‹for arbitrary p: ❙⊢⇩□ □(¬[λz p]x ≡ ¬p)›
by (rule RN)
AOT_hence ‹∀p □(¬[λz p]x ≡ ¬p)› using GEN by fast
AOT_hence ‹□(¬[λz p⇩1]x ≡ ¬p⇩1)› using "∀E" by fast
} note 2 = this
AOT_hence ‹□(¬[λz p⇩1]x ≡ ¬p⇩1)› using "∀E" by blast
AOT_hence 3: ‹¬[λz p⇩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: ❙⊢⇩□ ([λz p]x ≡ p)›
by (rule "beta-C-cor:3"[THEN "∀E"(2)]) cqt_2_lambda_inst_prover
AOT_hence ‹for arbitrary p: ❙⊢⇩□ □([λz p]x ≡ p)›
by (rule RN)
AOT_hence ‹∀p □([λz p]x ≡ p)› using GEN by fast
AOT_hence ‹□([λz p⇩1]x ≡ p⇩1)› using "∀E" by fast
} note 4 = this
AOT_have ‹◇[λz p⇩1]x›
using 4[THEN "qml:2"[axiom_inst, THEN "→E"]]
apply (AOT_subst ‹[λz p⇩1]x› ‹p⇩1›)
using 1[THEN "&E"(2)] by blast
AOT_hence ‹¬[λz p⇩1]x & ◇[λz p⇩1]x› using 3 "&I" by blast
AOT_hence ‹∃x (¬[λz p⇩1]x & ◇[λz p⇩1]x)› using "∃I"(2) by fast
moreover AOT_have ‹[λz p⇩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 "=⇩d⇩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⇩1 where ‹ContingentlyTrue(p⇩1)›
using "cont-tf-thm:1" "∃E"[rotated] by blast
AOT_hence ‹p⇩1 & ◇¬p⇩1› using "cont-tf:1"[THEN "≡⇩d⇩fE"] by blast
AOT_hence ‹∀x ([L]x ≡ [λz p⇩1]x) & ◇¬∀x([L]x ≡ [λz p⇩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 "=⇩d⇩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⇩1 where ‹ContingentlyFalse(p⇩1)›
using "cont-tf-thm:2" "∃E"[rotated] by blast
AOT_hence ‹¬p⇩1 & ◇p⇩1› using "cont-tf:2"[THEN "≡⇩d⇩fE"] by blast
AOT_hence ‹¬∀x ([L]x ≡ [λz p⇩1]x) & ◇∀x([L]x ≡ [λz p⇩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 "=⇩d⇩fI"(2)[OF L_def]) "cqt:2[lambda]"+
qed
AOT_theorem "eqnotnec:3": ‹∃F∃G(❙𝒜¬∀x([F]x ≡ [G]x) & ◇∀x([F]x ≡ [G]x))›
proof-
AOT_have ‹¬❙𝒜q⇩0›
apply (rule "=⇩d⇩fI"(2)[OF q⇩0_def])
apply (fact "log-prop-prop:2")
by (fact AOT)
AOT_hence ‹❙𝒜¬q⇩0›
using "logic-actual-nec:1"[axiom_inst, THEN "≡E"(2)] by blast
AOT_hence ‹❙𝒜¬∀x ([L]x ≡ [λz q⇩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⇩0]x)›
using eqnotnec_123_Aux_ξ'[THEN "→E"] q⇩0_prop[THEN "&E"(1)] by blast
ultimately AOT_have ‹❙𝒜¬∀x ([L]x ≡ [λz q⇩0]x) & ◇∀x ([L]x ≡ [λz q⇩0]x)›
using "&I" by blast
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 "=⇩d⇩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: ‹❙⊢⇩□ ψ → ∀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: ‹❙⊢⇩□ ψ → ∀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:
‹❙⊢⇩□ ◇¬ψ → ◇¬∀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⇩1 where p⇩1_prop: ‹p⇩1 & ◇¬p⇩1›
using "cont-tf-thm:1" "∃E"[rotated]
"cont-tf:1"[THEN "≡⇩d⇩fE"] by blast
{
AOT_assume 1: ‹□∀x([F]x ≡ [λz [F]z & p⇩1]x)›
AOT_have 2: ‹∀x([F]x ≡ [λz [F]z & p⇩1 ∨ ¬p⇩1]x)›
using Aux_B[THEN "→E", OF p⇩1_prop[THEN "&E"(1)]].
AOT_have ‹◇¬∀x([λz [F]z & p⇩1]x ≡ [λz [F]z & p⇩1 ∨ ¬p⇩1]x)›
using Aux_C[THEN "→E", OF p⇩1_prop[THEN "&E"(2)]].
AOT_hence 3: ‹◇¬∀x([F]x ≡ [λz [F]z & p⇩1 ∨ ¬p⇩1]x)›
using Aux_D[THEN "→E", OF 1, THEN "≡E"(1)] by blast
AOT_hence ‹∀x([F]x ≡ [λz [F]z & p⇩1 ∨ ¬p⇩1]x) &
◇¬∀x([F]x ≡ [λz [F]z & p⇩1 ∨ ¬p⇩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⇩1]x)›
AOT_hence ‹◇¬∀x([F]x ≡ [λz [F]z & p⇩1]x)›
using "KBasic:11"[THEN "≡E"(1)] by blast
AOT_hence ‹∀x ([F]x ≡ [λz [F]z & p⇩1]x) & ◇¬∀x([F]x ≡ [λz [F]z & p⇩1]x)›
using Aux_A[THEN "→E", OF p⇩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: ‹❙⊢⇩□ ◇ψ → ◇∀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: ‹❙⊢⇩□ ◇ψ → ◇∀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: ‹❙⊢⇩□ ¬ψ → ¬∀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⇩1 where p⇩1_prop: ‹¬p⇩1 & ◇p⇩1›
using "cont-tf-thm:2" "∃E"[rotated] "cont-tf:2"[THEN "≡⇩d⇩fE"] by blast
{
AOT_assume 1: ‹∀x([F]x ≡ [λz [F]z & p⇩1]x)›
AOT_have 2: ‹◇∀x([F]x ≡ [λz [F]z & p⇩1 ∨ ¬p⇩1]x)›
using Aux_B[THEN "→E", OF p⇩1_prop[THEN "&E"(2)]].
AOT_have ‹¬∀x([λz [F]z & p⇩1]x ≡ [λz [F]z & p⇩1 ∨ ¬p⇩1]x)›
using Aux_C[THEN "→E", OF p⇩1_prop[THEN "&E"(1)]].
AOT_hence 3: ‹¬∀x([F]x ≡ [λz [F]z & p⇩1 ∨ ¬p⇩1]x)›
using Aux_D[THEN "→E", OF 1, THEN "≡E"(1)] by blast
AOT_hence ‹¬∀x([F]x ≡ [λz [F]z & p⇩1 ∨ ¬p⇩1]x) &
◇∀x([F]x ≡ [λz [F]z & p⇩1 ∨ ¬p⇩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⇩1]x)›
AOT_hence ‹¬∀x([F]x ≡ [λz [F]z & p⇩1]x)›
using "KBasic:11"[THEN "≡E"(1)] by blast
AOT_hence ‹¬∀x ([F]x ≡ [λz [F]z & p⇩1]x) &
◇∀x([F]x ≡ [λz [F]z & p⇩1]x)›
using Aux_A[THEN "→E", OF p⇩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(❙𝒜¬∀x([F]x ≡ [G]x) & ◇∀x([F]x ≡ [G]x))›
proof(rule GEN)
fix F
AOT_have Aux_A: ‹❙⊢⇩□ ◇ψ → ◇∀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: ‹❙⊢⇩□ ◇ψ → ◇∀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:
‹❙⊢⇩□ ❙𝒜¬ψ → ❙𝒜¬∀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 ‹❙𝒜(∀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 ‹❙𝒜∀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 "act-cond"[THEN "→E"] by blast
AOT_hence 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 ψ
by (auto intro!: "→I" "Act-Basic:5"[THEN "≡E"(1)] dest!: "→E")
AOT_have ‹¬❙𝒜q⇩0›
apply (rule "=⇩d⇩fI"(2)[OF q⇩0_def])
apply (fact "log-prop-prop:2")
by (fact AOT)
AOT_hence q⇩0_prop_1: ‹❙𝒜¬q⇩0›
using "logic-actual-nec:1"[axiom_inst, THEN "≡E"(2)] by blast
{
AOT_assume 1: ‹❙𝒜∀x([F]x ≡ [λz [F]z & q⇩0]x)›
AOT_have 2: ‹◇∀x([F]x ≡ [λz [F]z & q⇩0 ∨ ¬q⇩0]x)›
using Aux_B[THEN "→E", OF q⇩0_prop[THEN "&E"(1)]].
AOT_have ‹❙𝒜¬∀x([λz [F]z & q⇩0]x ≡ [λz [F]z & q⇩0 ∨ ¬q⇩0]x)›
using Aux_C[THEN "→E", OF q⇩0_prop_1].
AOT_hence 3: ‹❙𝒜¬∀x([F]x ≡ [λz [F]z & q⇩0 ∨ ¬q⇩0]x)›
using Aux_D[THEN "→E", OF 1, THEN "≡E"(1)] by blast
AOT_hence ‹❙𝒜¬∀x([F]x ≡ [λz [F]z & q⇩0 ∨ ¬q⇩0]x) &
◇∀x([F]x ≡ [λz [F]z & q⇩0 ∨ ¬q⇩0]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 & q⇩0]x)›
AOT_hence ‹❙𝒜¬∀x([F]x ≡ [λz [F]z & q⇩0]x)›
using "logic-actual-nec:1"[axiom_inst, THEN "≡E"(2)] by blast
AOT_hence ‹❙𝒜¬∀x ([F]x ≡ [λz [F]z & q⇩0]x) & ◇∀x([F]x ≡ [λz [F]z & q⇩0]x)›
using Aux_A[THEN "→E", OF q⇩0_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 "oa-contingent:1": ‹O! ≠ A!›
proof(rule "≡⇩d⇩fI"[OF "=-infix"]; rule "raa-cor:2")
fix x
AOT_assume 1: ‹O! = A!›
AOT_hence ‹[λx ◇E!x] = A!›
by (rule "=⇩d⇩fE"(2)[OF AOT_ordinary, rotated]) "cqt:2[lambda]"
AOT_hence ‹[λx ◇E!x] = [λx ¬◇E!x]›
by (rule "=⇩d⇩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 "=⇩d⇩fE"(2)[OF AOT_ordinary])
apply "cqt:2[lambda]"
apply argo
apply (rule "=⇩d⇩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 "=⇩d⇩fI"(2)[OF AOT_abstract])
apply "cqt:2[lambda]"
apply argo
apply (rule "=⇩d⇩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 "=⇩d⇩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 "=⇩d⇩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!⇧-x ≡ ¬A!⇧-x›
proof -
AOT_have ‹O!x ≡ ¬A!x›
using "oa-contingent:2" by blast
also AOT_have ‹… ≡ A!⇧-x›
using "thm-relation-negation:1"[symmetric, unvarify F, OF "oa-exist:2"].
finally AOT_have 1: ‹O!x ≡ A!⇧-x›.
AOT_have ‹A!x ≡ ¬O!x›
using "oa-contingent:3" by blast
also AOT_have ‹… ≡ O!⇧-x›
using "thm-relation-negation:1"[symmetric, unvarify F, OF "oa-exist:1"].
finally AOT_have 2: ‹A!x ≡ O!⇧-x›.
AOT_show ‹O!⇧-x ≡ ¬A!⇧-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!⇧- ≠ A!⇧-›
proof (rule "=-infix"[THEN "≡⇩d⇩fI"]; rule "raa-cor:2")
AOT_assume 1: ‹O!⇧- = A!⇧-›
fix x
AOT_have ‹A!⇧-x ≡ O!⇧-x›
apply (rule "rule=E"[rotated, OF 1])
by (fact "oth-class-taut:3:a")
AOT_hence ‹A!⇧-x ≡ ¬A!⇧-x›
using "oa-contingent:7" "≡E" by fast
AOT_thus ‹(A!⇧-x ≡ ¬A!⇧-x) & ¬(A!⇧-x ≡ ¬A!⇧-x)›
using "oth-class-taut:3:c" "&I" by blast
qed
AOT_theorem "oa-contingent:8": ‹Contingent(O!⇧-)›
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!⇧-)›
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]) ≡⇩d⇩f Contingent([F]) & ∀x (◇[F]x → □[F]x)›
AOT_theorem "cont-nec-fact1:1":
‹WeaklyContingent([F]) ≡ WeaklyContingent([F]⇧-)›
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]⇧-) & ∀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]⇧-) & ∀x (◇[F]⇧-x → □[F]⇧-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]⇧-x›
AOT_have ‹◇¬[F]x›
by (AOT_subst (reverse) ‹¬[F]x› ‹[F]⇧-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]⇧-x›
proof (rule "raa-cor:1")
AOT_assume 3: ‹¬□[F]⇧-x›
AOT_have ‹¬□¬[F]x›
by (AOT_subst (reverse) ‹¬[F]x› ‹[F]⇧-x›)
(auto simp add: "thm-relation-negation:1" 3)
AOT_hence ‹◇[F]x›
using "conventions:5"[THEN "≡⇩d⇩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]⇧-x → □[F]⇧-x)›
AOT_assume 1: ‹◇[F]x›
AOT_have ‹◇¬[F]⇧-x›
by (AOT_subst ‹¬[F]⇧-x› ‹[F]x›)
(auto simp: "thm-relation-negation:2" 1)
AOT_hence 2: ‹¬□[F]⇧-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]⇧-x›
by (AOT_subst ‹¬[F]⇧-x› ‹[F]x›)
(auto simp add: "thm-relation-negation:2" 3)
AOT_hence ‹◇[F]⇧-x›
using "conventions:5"[THEN "≡⇩d⇩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
qed
also AOT_have ‹… ≡ WeaklyContingent([F]⇧-)›
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 "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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 & ¬❙𝒜E!x)› using "qml:4"[axiom_inst].
AOT_hence ‹∃x ◇(E!x & ¬❙𝒜E!x)› using "BF◇"[THEN "→E"] by blast
then AOT_obtain a where ‹◇(E!a & ¬❙𝒜E!a)› using "∃E"[rotated] by blast
AOT_hence 1: ‹◇E!a & ◇¬❙𝒜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 ‹❙𝒜E!a› using "nec-imp-act"[THEN "→E"] by blast
AOT_hence ‹□❙𝒜E!a› using "qml-act:1"[axiom_inst, THEN "→E"] by blast
moreover AOT_have ‹¬□❙𝒜E!a›
using "KBasic:11"[THEN "≡E"(2)] 1[THEN "&E"(2)] by meson
ultimately AOT_have ‹□❙𝒜E!a & ¬□❙𝒜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 "≡⇩d⇩fE"]].
AOT_theorem "cont-nec-fact2:5": ‹O! ≠ E! & O! ≠ E!⇧- & O! ≠ L & O! ≠ L⇧-›
proof -
AOT_have 1: ‹L↓›
by (rule "=⇩d⇩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 "=⇩d⇩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 "=⇩d⇩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!⇧- & A! ≠ L & A! ≠ L⇧-›
proof -
AOT_have 1: ‹L↓›
by (rule "=⇩d⇩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 "=⇩d⇩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 "=⇩d⇩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 :: ‹φ ⇒ φ› ("❙Δ_" [49] 54)
‹❙Δp ≡⇩d⇩f □p ∨ (¬❙𝒜p & ◇p)›
AOT_theorem sixteen:
shows ‹∃F⇩1∃F⇩2∃F⇩3∃F⇩4∃F⇩5∃F⇩6∃F⇩7∃F⇩8∃F⇩9∃F⇩1⇩0∃F⇩1⇩1∃F⇩1⇩2∃F⇩1⇩3∃F⇩1⇩4∃F⇩1⇩5∃F⇩1⇩6 (
«F⇩1::<κ>» ≠ F⇩2 & F⇩1 ≠ F⇩3 & F⇩1 ≠ F⇩4 & F⇩1 ≠ F⇩5 & F⇩1 ≠ F⇩6 & F⇩1 ≠ F⇩7 &
F⇩1 ≠ F⇩8 & F⇩1 ≠ F⇩9 & F⇩1 ≠ F⇩1⇩0 & F⇩1 ≠ F⇩1⇩1 & F⇩1 ≠ F⇩1⇩2 & F⇩1 ≠ F⇩1⇩3 &
F⇩1 ≠ F⇩1⇩4 & F⇩1 ≠ F⇩1⇩5 & F⇩1 ≠ F⇩1⇩6 &
F⇩2 ≠ F⇩3 & F⇩2 ≠ F⇩4 & F⇩2 ≠ F⇩5 & F⇩2 ≠ F⇩6 & F⇩2 ≠ F⇩7 & F⇩2 ≠ F⇩8 &
F⇩2 ≠ F⇩9 & F⇩2 ≠ F⇩1⇩0 & F⇩2 ≠ F⇩1⇩1 & F⇩2 ≠ F⇩1⇩2 & F⇩2 ≠ F⇩1⇩3 & F⇩2 ≠ F⇩1⇩4 &
F⇩2 ≠ F⇩1⇩5 & F⇩2 ≠ F⇩1⇩6 &
F⇩3 ≠ F⇩4 & F⇩3 ≠ F⇩5 & F⇩3 ≠ F⇩6 & F⇩3 ≠ F⇩7 & F⇩3 ≠ F⇩8 & F⇩3 ≠ F⇩9 & F⇩3 ≠ F⇩1⇩0 &
F⇩3 ≠ F⇩1⇩1 & F⇩3 ≠ F⇩1⇩2 & F⇩3 ≠ F⇩1⇩3 & F⇩3 ≠ F⇩1⇩4 & F⇩3 ≠ F⇩1⇩5 & F⇩3 ≠ F⇩1⇩6 &
F⇩4 ≠ F⇩5 & F⇩4 ≠ F⇩6 & F⇩4 ≠ F⇩7 & F⇩4 ≠ F⇩8 & F⇩4 ≠ F⇩9 & F⇩4 ≠ F⇩1⇩0 & F⇩4 ≠ F⇩1⇩1 &
F⇩4 ≠ F⇩1⇩2 & F⇩4 ≠ F⇩1⇩3 & F⇩4 ≠ F⇩1⇩4 & F⇩4 ≠ F⇩1⇩5 & F⇩4 ≠ F⇩1⇩6 &
F⇩5 ≠ F⇩6 & F⇩5 ≠ F⇩7 & F⇩5 ≠ F⇩8 & F⇩5 ≠ F⇩9 & F⇩5 ≠ F⇩1⇩0 & F⇩5 ≠ F⇩1⇩1 & F⇩5 ≠ F⇩1⇩2 &
F⇩5 ≠ F⇩1⇩3 & F⇩5 ≠ F⇩1⇩4 & F⇩5 ≠ F⇩1⇩5 & F⇩5 ≠ F⇩1⇩6 &
F⇩6 ≠ F⇩7 & F⇩6 ≠ F⇩8 & F⇩6 ≠ F⇩9 & F⇩6 ≠ F⇩1⇩0 & F⇩6 ≠ F⇩1⇩1 & F⇩6 ≠ F⇩1⇩2 & F⇩6 ≠ F⇩1⇩3 &
F⇩6 ≠ F⇩1⇩4 & F⇩6 ≠ F⇩1⇩5 & F⇩6 ≠ F⇩1⇩6 &
F⇩7 ≠ F⇩8 & F⇩7 ≠ F⇩9 & F⇩7 ≠ F⇩1⇩0 & F⇩7 ≠ F⇩1⇩1 & F⇩7 ≠ F⇩1⇩2 & F⇩7 ≠ F⇩1⇩3 & F⇩7 ≠ F⇩1⇩4 &
F⇩7 ≠ F⇩1⇩5 & F⇩7 ≠ F⇩1⇩6 &
F⇩8 ≠ F⇩9 & F⇩8 ≠ F⇩1⇩0 & F⇩8 ≠ F⇩1⇩1 & F⇩8 ≠ F⇩1⇩2 & F⇩8 ≠ F⇩1⇩3 & F⇩8 ≠ F⇩1⇩4 & F⇩8 ≠ F⇩1⇩5 &
F⇩8 ≠ F⇩1⇩6 &
F⇩9 ≠ F⇩1⇩0 & F⇩9 ≠ F⇩1⇩1 & F⇩9 ≠ F⇩1⇩2 & F⇩9 ≠ F⇩1⇩3 & F⇩9 ≠ F⇩1⇩4 & F⇩9 ≠ F⇩1⇩5 & F⇩9 ≠ F⇩1⇩6 &
F⇩1⇩0 ≠ F⇩1⇩1 & F⇩1⇩0 ≠ F⇩1⇩2 & F⇩1⇩0 ≠ F⇩1⇩3 & F⇩1⇩0 ≠ F⇩1⇩4 & F⇩1⇩0 ≠ F⇩1⇩5 & F⇩1⇩0 ≠ F⇩1⇩6 &
F⇩1⇩1 ≠ F⇩1⇩2 & F⇩1⇩1 ≠ F⇩1⇩3 & F⇩1⇩1 ≠ F⇩1⇩4 & F⇩1⇩1 ≠ F⇩1⇩5 & F⇩1⇩1 ≠ F⇩1⇩6 &
F⇩1⇩2 ≠ F⇩1⇩3 & F⇩1⇩2 ≠ F⇩1⇩4 & F⇩1⇩2 ≠ F⇩1⇩5 & F⇩1⇩2 ≠ F⇩1⇩6 &
F⇩1⇩3 ≠ F⇩1⇩4 & F⇩1⇩3 ≠ F⇩1⇩5 & F⇩1⇩3 ≠ F⇩1⇩6 &
F⇩1⇩4 ≠ F⇩1⇩5 & F⇩1⇩4 ≠ F⇩1⇩6 &
F⇩1⇩5 ≠ F⇩1⇩6)›
proof -
AOT_have Delta_pos: ‹❙Δφ → ◇φ› for φ
proof(rule "→I")
AOT_assume ‹❙Δφ›
AOT_hence ‹□φ ∨ (¬❙𝒜φ & ◇φ)›
using "≡⇩d⇩fE"[OF necessary_or_contingently_false] by blast
moreover {
AOT_assume ‹□φ›
AOT_hence ‹◇φ›
by (metis "B◇" "T◇" "vdash-properties:10")
}
moreover {
AOT_assume ‹¬❙𝒜φ & ◇φ›
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: ‹¬❙Δφ› if ‹❙𝒜φ› and ‹¬□φ› for φ
using "≡⇩d⇩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: ‹¬❙Δφ› if ‹❙𝒜φ› and ‹◇¬φ› for φ
using "KBasic:11" act_and_not_nec_not_delta "≡E"(2) that(1,2) by blast
AOT_have impossible_delta: ‹¬❙Δφ› if ‹¬◇φ› for φ
using Delta_pos "modus-tollens:1" that by blast
AOT_have not_act_and_pos_delta: ‹❙Δφ› if ‹¬❙𝒜φ› and ‹◇φ› for φ
by (meson "≡⇩d⇩fI" "&I" "∨I"(2) necessary_or_contingently_false that(1,2))
AOT_have nec_delta: ‹❙Δφ› if ‹□φ› for φ
using "≡⇩d⇩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 & ¬❙𝒜[E!]b›
using "pos-not-pna:3" using "∃E"[rotated] by blast
AOT_have b_ord: ‹[O!]b›
proof(rule "=⇩d⇩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⇧-]x› for x
using "thm-noncont-e-e:2" "contingent-properties:2"[THEN "≡⇩d⇩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 "≡⇩d⇩fE"]
CBF[THEN "→E"] "∀E" by blast
AOT_have act_ord_b: ‹❙𝒜[O!]b›
using b_ord "≡E"(1) "oa-facts:7" by blast
AOT_have delta_ord_b: ‹❙Δ[O!]b›
by (meson "≡⇩d⇩fI" b_ord "∨I"(1) necessary_or_contingently_false
"oa-facts:1" "→E")
AOT_have not_act_ord_a: ‹¬❙𝒜[O!]a›
by (meson a_prop "≡E"(1) "≡E"(3) "oa-contingent:3" "oa-facts:7")
AOT_have not_delta_ord_a: ‹¬❙Δ[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: ‹¬❙𝒜[A!]b›
by (meson b_ord "≡E"(1) "≡E"(3) "oa-contingent:2" "oa-facts:8")
AOT_have not_delta_abs_b: ‹¬❙Δ[A!]b›
proof(rule "raa-cor:2")
AOT_assume ‹❙Δ[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: ‹❙𝒜[A!]a›
using a_prop "≡E"(1) "oa-facts:8" by blast
AOT_have delta_abs_a: ‹❙Δ[A!]a›
by (metis "≡⇩d⇩fI" a_prop "oa-facts:2" "→E" "∨I"(1)
necessary_or_contingently_false)
AOT_have not_act_concrete_b: ‹¬❙𝒜[E!]b›
using b_prop "&E"(2) by blast
AOT_have delta_concrete_b: ‹❙Δ[E!]b›
proof (rule "≡⇩d⇩fI"[OF necessary_or_contingently_false];
rule "∨I"(2); rule "&I")
AOT_show ‹¬❙𝒜[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: ‹¬❙𝒜[E!]a›
proof (rule "raa-cor:2")
AOT_assume ‹❙𝒜[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 "=⇩d⇩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: ‹¬❙Δ[E!]a›
proof (rule "raa-cor:2")
AOT_assume ‹❙Δ[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 "=⇩d⇩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: ‹¬❙𝒜q⇩0›
by (meson "log-prop-prop:2" "pos-not-pna:1"
q⇩0_def "reductio-aa:1" "rule-id-df:2:a[zero]")
AOT_have delta_q_zero: ‹❙Δq⇩0›
proof(rule "≡⇩d⇩fI"[OF necessary_or_contingently_false];
rule "∨I"(2); rule "&I")
AOT_show ‹¬❙𝒜q⇩0› using not_act_q_zero.
AOT_show ‹◇q⇩0› by (meson "&E"(1) q⇩0_prop)
qed
AOT_have act_not_q_zero: ‹❙𝒜¬q⇩0›
using "Act-Basic:1" "∨E"(2) not_act_q_zero by blast
AOT_have not_delta_not_q_zero: ‹¬❙Δ¬q⇩0›
using "≡⇩d⇩fE" "conventions:5" "Act-Basic:1" act_and_not_nec_not_delta
"&E"(1) "∨E"(2) not_act_q_zero q⇩0_prop by blast
AOT_have ‹[L⇧-]↓› by (simp add: "rel-neg-T:3")
moreover AOT_have ‹¬❙𝒜[L⇧-]b & ¬❙Δ[L⇧-]b & ¬❙𝒜[L⇧-]a & ¬❙Δ[L⇧-]a›
proof (safe intro!: "&I")
AOT_show ‹¬❙𝒜[L⇧-]b›
by (meson "≡E"(1) "logic-actual-nec:1"[axiom_inst] "nec-imp-act"
nec_not_L_neg "→E")
AOT_show ‹¬❙Δ[L⇧-]b›
by (meson Delta_pos "KBasic2:1" "≡E"(1)
"modus-tollens:1" nec_not_L_neg)
AOT_show ‹¬❙𝒜[L⇧-]a›
by (meson "≡E"(1) "logic-actual-nec:1"[axiom_inst]
"nec-imp-act" nec_not_L_neg "→E")
AOT_show ‹¬❙Δ[L⇧-]a›
using Delta_pos "KBasic2:1" "≡E"(1) "modus-tollens:1"
nec_not_L_neg by blast
qed
ultimately AOT_obtain F⇩0 where ‹¬❙𝒜[F⇩0]b & ¬❙Δ[F⇩0]b & ¬❙𝒜[F⇩0]a & ¬❙Δ[F⇩0]a›
using "∃I"(1)[rotated, THEN "∃E"[rotated]] by fastforce
AOT_hence ‹¬❙𝒜[F⇩0]b› and ‹¬❙Δ[F⇩0]b› and ‹¬❙𝒜[F⇩0]a› and ‹¬❙Δ[F⇩0]a›
using "&E" by blast+
note props = this
let ?Π = "«[λy [A!]y & q⇩0]»"
AOT_modally_strict {
AOT_have ‹[«?Π»]↓› by "cqt:2[lambda]"
} note 1 = this
moreover AOT_have ‹¬❙𝒜[«?Π»]b & ¬❙Δ[«?Π»]b & ¬❙𝒜[«?Π»]a & ❙Δ[«?Π»]a›
proof (safe intro!: "&I"; AOT_subst ‹[λy A!y & q⇩0]x› ‹A!x & q⇩0› for: x)
AOT_show ‹¬❙𝒜([A!]b & q⇩0)›
using "Act-Basic:2" "&E"(1) "≡E"(1) not_act_abs_b "raa-cor:3" by blast
next AOT_show ‹¬❙Δ([A!]b & q⇩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 ‹¬❙𝒜([A!]a & q⇩0)›
using "Act-Basic:2" "&E"(2) "≡E"(1) not_act_q_zero
"raa-cor:3" by blast
next AOT_show ‹❙Δ([A!]a & q⇩0)›
proof (rule not_act_and_pos_delta)
AOT_show ‹¬❙𝒜([A!]a & q⇩0)›
using "Act-Basic:2" "&E"(2) "≡E"(4) not_act_q_zero
"raa-cor:3" by blast
next AOT_show ‹◇([A!]a & q⇩0)›
by (metis "&I" "→E" Delta_pos "KBasic:16" "&E"(1) delta_abs_a
"≡E"(1) "oa-facts:6" q⇩0_prop)
qed
qed(auto simp: "beta-C-meta"[THEN "→E", OF 1])
ultimately AOT_obtain F⇩1 where ‹¬❙𝒜[F⇩1]b & ¬❙Δ[F⇩1]b & ¬❙𝒜[F⇩1]a & ❙Δ[F⇩1]a›
using "∃I"(1)[rotated, THEN "∃E"[rotated]] by fastforce
AOT_hence ‹¬❙𝒜[F⇩1]b› and ‹¬❙Δ[F⇩1]b› and ‹¬❙𝒜[F⇩1]a› and ‹❙Δ[F⇩1]a›
using "&E" by blast+
note props = props this
let ?Π = "«[λy [A!]y & ¬q⇩0]»"
AOT_modally_strict {
AOT_have ‹[«?Π»]↓› by "cqt:2[lambda]"
} note 1 = this
moreover AOT_have ‹¬❙𝒜[«?Π»]b & ¬❙Δ[«?Π»]b & ❙𝒜[«?Π»]a & ¬❙Δ[«?Π»]a›
proof (safe intro!: "&I"; AOT_subst ‹[λy A!y & ¬q⇩0]x› ‹A!x & ¬q⇩0› for: x)
AOT_show ‹¬❙𝒜([A!]b & ¬q⇩0)›
using "Act-Basic:2" "&E"(1) "≡E"(1) not_act_abs_b "raa-cor:3" by blast
next AOT_show ‹¬❙Δ([A!]b & ¬q⇩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 ‹❙𝒜([A!]a & ¬q⇩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⇩0)›
proof (rule act_and_not_nec_not_delta)
AOT_show ‹❙𝒜([A!]a & ¬q⇩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⇩0)›
by (metis "KBasic2:1" "KBasic:3" "&E"(1) "&E"(2) "≡E"(4)
q⇩0_prop "raa-cor:3")
qed
qed(auto simp: "beta-C-meta"[THEN "→E", OF 1])
ultimately AOT_obtain F⇩2 where ‹¬❙𝒜[F⇩2]b & ¬❙Δ[F⇩2]b & ❙𝒜[F⇩2]a & ¬❙Δ[F⇩2]a›
using "∃I"(1)[rotated, THEN "∃E"[rotated]] by fastforce
AOT_hence ‹¬❙𝒜[F⇩2]b› and ‹¬❙Δ[F⇩2]b› and ‹❙𝒜[F⇩2]a› and ‹¬❙Δ[F⇩2]a›
using "&E" by blast+
note props = props this
AOT_have abstract_prop: ‹¬❙𝒜[A!]b & ¬❙Δ[A!]b & ❙𝒜[A!]a & ❙Δ[A!]a›
using act_abs_a "&I" delta_abs_a not_act_abs_b not_delta_abs_b
by presburger
then AOT_obtain F⇩3 where ‹¬❙𝒜[F⇩3]b & ¬❙Δ[F⇩3]b & ❙𝒜[F⇩3]a & ❙Δ[F⇩3]a›
using "∃I"(1)[rotated, THEN "∃E"[rotated]] "oa-exist:2" by fastforce
AOT_hence ‹¬❙𝒜[F⇩3]b› and ‹¬❙Δ[F⇩3]b› and ‹❙𝒜[F⇩3]a› and ‹❙Δ[F⇩3]a›
using "&E" by blast+
note props = props this
AOT_have ‹¬❙𝒜[E!]b & ❙Δ[E!]b & ¬❙𝒜[E!]a & ¬❙Δ[E!]a›
by (meson "&I" delta_concrete_b not_act_concrete_a
not_act_concrete_b not_delta_concrete_a)
then AOT_obtain F⇩4 where ‹¬❙𝒜[F⇩4]b & ❙Δ[F⇩4]b & ¬❙𝒜[F⇩4]a & ¬❙Δ[F⇩4]a›
using "∃I"(1)[rotated, THEN "∃E"[rotated]]
by fastforce
AOT_hence ‹¬❙𝒜[F⇩4]b› and ‹❙Δ[F⇩4]b› and ‹¬❙𝒜[F⇩4]a› and ‹¬❙Δ[F⇩4]a›
using "&E" by blast+
note props = props this
AOT_modally_strict {
AOT_have ‹[λy q⇩0]↓› by "cqt:2[lambda]"
} note 1 = this
moreover AOT_have ‹¬❙𝒜[λy q⇩0]b & ❙Δ[λy q⇩0]b & ¬❙𝒜[λy q⇩0]a & ❙Δ[λy q⇩0]a›
by (safe intro!: "&I"; AOT_subst ‹[λy q⇩0]b› ‹q⇩0› for: b)
(auto simp: not_act_q_zero delta_q_zero "beta-C-meta"[THEN "→E", OF 1])
ultimately AOT_obtain F⇩5 where ‹¬❙𝒜[F⇩5]b & ❙Δ[F⇩5]b & ¬❙𝒜[F⇩5]a & ❙Δ[F⇩5]a›
using "∃I"(1)[rotated, THEN "∃E"[rotated]]
by fastforce
AOT_hence ‹¬❙𝒜[F⇩5]b› and ‹❙Δ[F⇩5]b› and ‹¬❙𝒜[F⇩5]a› and ‹❙Δ[F⇩5]a›
using "&E" by blast+
note props = props this
let ?Π = "«[λy [E!]y ∨ ([A!]y & ¬q⇩0)]»"
AOT_modally_strict {
AOT_have ‹[«?Π»]↓› by "cqt:2[lambda]"
} note 1 = this
moreover AOT_have ‹¬❙𝒜[«?Π»]b & ❙Δ[«?Π»]b & ❙𝒜[«?Π»]a & ¬❙Δ[«?Π»]a›
proof(safe intro!: "&I";
AOT_subst ‹[λy E!y ∨ (A!y & ¬q⇩0)]x› ‹E!x ∨ (A!x & ¬q⇩0)› for: x)
AOT_have ‹❙𝒜¬([A!]b & ¬q⇩0)›
by (metis "Act-Basic:1" "Act-Basic:2" abstract_prop "&E"(1) "∨E"(2)
"≡E"(1) "raa-cor:3")
moreover AOT_have ‹¬❙𝒜[E!]b›
using b_prop "&E"(2) by blast
ultimately AOT_have 2: ‹❙𝒜(¬[E!]b & ¬([A!]b & ¬q⇩0))›
by (metis "Act-Basic:2" "Act-Sub:1" "&I" "≡E"(3) "raa-cor:1")
AOT_have ‹❙𝒜¬([E!]b ∨ ([A!]b & ¬q⇩0))›
by (AOT_subst ‹¬([E!]b ∨ ([A!]b & ¬q⇩0))› ‹¬[E!]b & ¬([A!]b & ¬q⇩0)›)
(auto simp: "oth-class-taut:5:d" 2)
AOT_thus ‹¬❙𝒜([E!]b ∨ ([A!]b & ¬q⇩0))›
by (metis "¬¬I" "Act-Sub:1" "≡E"(4))
next
AOT_show ‹❙Δ([E!]b ∨ ([A!]b & ¬q⇩0))›
proof (rule not_act_and_pos_delta)
AOT_show ‹¬❙𝒜([E!]b ∨ ([A!]b & ¬q⇩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⇩0))›
using "KBasic2:2" b_prop "&E"(1) "∨I"(1) "≡E"(3) "raa-cor:3" by blast
qed
next AOT_show ‹❙𝒜([E!]a ∨ ([A!]a & ¬q⇩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 ‹¬❙Δ([E!]a ∨ ([A!]a & ¬q⇩0))›
proof (rule act_and_not_nec_not_delta)
AOT_show ‹❙𝒜([E!]a ∨ ([A!]a & ¬q⇩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 "≡⇩d⇩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⇩0)›
by (metis "KBasic2:1" "KBasic:11" "KBasic:3"
"&E"(1,2) "≡E"(1) q⇩0_prop "raa-cor:3")
ultimately AOT_have ‹◇(¬[E!]a & ¬([A!]a & ¬q⇩0))›
by (metis "KBasic:16" "&I" "vdash-properties:10")
AOT_hence ‹◇¬([E!]a ∨ ([A!]a & ¬q⇩0))›
by (metis "RE◇" "≡E"(2) "oth-class-taut:5:d")
AOT_thus ‹¬□([E!]a ∨ ([A!]a & ¬q⇩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⇩6 where ‹¬❙𝒜[F⇩6]b & ❙Δ[F⇩6]b & ❙𝒜[F⇩6]a & ¬❙Δ[F⇩6]a›
using "∃I"(1)[rotated, THEN "∃E"[rotated]] by fastforce
AOT_hence ‹¬❙𝒜[F⇩6]b› and ‹❙Δ[F⇩6]b› and ‹❙𝒜[F⇩6]a› and ‹¬❙Δ[F⇩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 ‹¬❙𝒜[«?Π»]b & ❙Δ[«?Π»]b & ❙𝒜[«?Π»]a & ❙Δ[«?Π»]a›
proof(safe intro!: "&I"; AOT_subst ‹[λy A!y ∨ E!y]x› ‹A!x ∨ E!x› for: x)
AOT_show ‹¬❙𝒜([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)›
proof (rule not_act_and_pos_delta)
AOT_show ‹¬❙𝒜([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 ‹❙𝒜([A!]a ∨ [E!]a)›
by (meson "Act-Basic:9" act_abs_a "∨I"(1) "≡E"(2))
next AOT_show ‹❙Δ([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⇩7 where ‹¬❙𝒜[F⇩7]b & ❙Δ[F⇩7]b & ❙𝒜[F⇩7]a & ❙Δ[F⇩7]a›
using "∃I"(1)[rotated, THEN "∃E"[rotated]] by fastforce
AOT_hence ‹¬❙𝒜[F⇩7]b› and ‹❙Δ[F⇩7]b› and ‹❙𝒜[F⇩7]a› and ‹❙Δ[F⇩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 ‹❙𝒜[«?Π»]b & ¬❙Δ[«?Π»]b & ¬❙𝒜[«?Π»]a & ¬❙Δ[«?Π»]a›
proof(safe intro!: "&I"; AOT_subst ‹[λy O!y & ¬E!y]x› ‹O!x & ¬E!x› for: x)
AOT_show ‹❙𝒜([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 ‹¬❙Δ([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 ‹¬❙𝒜([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 ‹¬❙Δ([O!]a & ¬[E!]a)›
by (rule impossible_delta)
qed(auto simp: "beta-C-meta"[THEN "→E", OF 1])
ultimately AOT_obtain F⇩8 where ‹❙𝒜[F⇩8]b & ¬❙Δ[F⇩8]b & ¬❙𝒜[F⇩8]a & ¬❙Δ[F⇩8]a›
using "∃I"(1)[rotated, THEN "∃E"[rotated]] by fastforce
AOT_hence ‹❙𝒜[F⇩8]b› and ‹¬❙Δ[F⇩8]b› and ‹¬❙𝒜[F⇩8]a› and ‹¬❙Δ[F⇩8]a›
using "&E" by blast+
note props = props this
let ?Π = "«[λy ¬[E!]y & ([O!]y ∨ q⇩0)]»"
AOT_modally_strict {
AOT_have ‹[«?Π»]↓› by "cqt:2[lambda]"
} note 1 = this
moreover AOT_have ‹❙𝒜[«?Π»]b & ¬❙Δ[«?Π»]b & ¬❙𝒜[«?Π»]a & ❙Δ[«?Π»]a›
proof(safe intro!: "&I";
AOT_subst ‹[λy ¬E!y & (O!y ∨ q⇩0)]x› ‹¬E!x & (O!x ∨ q⇩0)› for: x)
AOT_show ‹❙𝒜(¬[E!]b & ([O!]b ∨ q⇩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⇩0))›
proof (rule act_and_pos_not_not_delta)
AOT_show ‹❙𝒜(¬[E!]b & ([O!]b ∨ q⇩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⇩0))›
proof (AOT_subst ‹¬(¬[E!]b & ([O!]b ∨ q⇩0))› ‹[E!]b ∨ ¬([O!]b ∨ q⇩0)›)
AOT_modally_strict {
AOT_show ‹¬(¬[E!]b & ([O!]b ∨ q⇩0)) ≡ [E!]b ∨ ¬([O!]b ∨ q⇩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⇩0))›
using "KBasic2:2" b_prop "&E"(1) "∨I"(1) "≡E"(3)
"raa-cor:3" by blast
qed
qed
next
AOT_show ‹¬❙𝒜(¬[E!]a & ([O!]a ∨ q⇩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 ‹❙Δ(¬[E!]a & ([O!]a ∨ q⇩0))›
proof (rule not_act_and_pos_delta)
AOT_show ‹¬❙𝒜(¬[E!]a & ([O!]a ∨ q⇩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⇩0)›
by (metis "KBasic2:2" "&E"(1) "∨I"(2) "≡E"(3) q⇩0_prop "raa-cor:3")
ultimately AOT_show ‹◇(¬[E!]a & ([O!]a ∨ q⇩0))›
by (metis "KBasic:16" "&I" "vdash-properties:10")
qed
qed(auto simp: "beta-C-meta"[THEN "→E", OF 1])
ultimately AOT_obtain F⇩9 where ‹❙𝒜[F⇩9]b & ¬❙Δ[F⇩9]b & ¬❙𝒜[F⇩9]a & ❙Δ[F⇩9]a›
using "∃I"(1)[rotated, THEN "∃E"[rotated]] by fastforce
AOT_hence ‹❙𝒜[F⇩9]b› and ‹¬❙Δ[F⇩9]b› and ‹¬❙𝒜[F⇩9]a› and ‹❙Δ[F⇩9]a›
using "&E" by blast+
note props = props this
AOT_modally_strict {
AOT_have ‹[λy ¬q⇩0]↓› by "cqt:2[lambda]"
} note 1 = this
moreover AOT_have ‹❙𝒜[λy ¬q⇩0]b & ¬❙Δ[λy ¬q⇩0]b & ❙𝒜[λy ¬q⇩0]a & ¬❙Δ[λy ¬q⇩0]a›
by (safe intro!: "&I"; AOT_subst ‹[λy ¬q⇩0]x› ‹¬q⇩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⇩1⇩0 where ‹❙𝒜[F⇩1⇩0]b & ¬❙Δ[F⇩1⇩0]b & ❙𝒜[F⇩1⇩0]a & ¬❙Δ[F⇩1⇩0]a›
using "∃I"(1)[rotated, THEN "∃E"[rotated]] by fastforce
AOT_hence ‹❙𝒜[F⇩1⇩0]b› and ‹¬❙Δ[F⇩1⇩0]b› and ‹❙𝒜[F⇩1⇩0]a› and ‹¬❙Δ[F⇩1⇩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 ‹❙𝒜[λy ¬[E!]y]b & ¬❙Δ[λy ¬[E!]y]b &
❙𝒜[λy ¬[E!]y]a & ❙Δ[λy ¬[E!]y]a›
proof (safe intro!: "&I"; AOT_subst ‹[λy ¬[E!]y]x› ‹¬[E!]x› for: x)
AOT_show ‹❙𝒜¬[E!]b›
using "Act-Basic:1" "∨E"(2) not_act_concrete_b by blast
next AOT_show ‹¬❙Δ¬[E!]b›
using "≡⇩d⇩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 ‹❙𝒜¬[E!]a›
using "Act-Basic:1" "∨E"(2) not_act_concrete_a by blast
next AOT_show ‹❙Δ¬[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⇩1⇩1 where ‹❙𝒜[F⇩1⇩1]b & ¬❙Δ[F⇩1⇩1]b & ❙𝒜[F⇩1⇩1]a & ❙Δ[F⇩1⇩1]a›
using "∃I"(1)[rotated, THEN "∃E"[rotated]] by fastforce
AOT_hence ‹❙𝒜[F⇩1⇩1]b› and ‹¬❙Δ[F⇩1⇩1]b› and ‹❙𝒜[F⇩1⇩1]a› and ‹❙Δ[F⇩1⇩1]a›
using "&E" by blast+
note props = props this
AOT_have ‹❙𝒜[O!]b & ❙Δ[O!]b & ¬❙𝒜[O!]a & ¬❙Δ[O!]a›
by (simp add: act_ord_b "&I" delta_ord_b not_act_ord_a not_delta_ord_a)
then AOT_obtain F⇩1⇩2 where ‹❙𝒜[F⇩1⇩2]b & ❙Δ[F⇩1⇩2]b & ¬❙𝒜[F⇩1⇩2]a & ¬❙Δ[F⇩1⇩2]a›
using "oa-exist:1" "∃I"(1)[rotated, THEN "∃E"[rotated]] by fastforce
AOT_hence ‹❙𝒜[F⇩1⇩2]b› and ‹❙Δ[F⇩1⇩2]b› and ‹¬❙𝒜[F⇩1⇩2]a› and ‹¬❙Δ[F⇩1⇩2]a›
using "&E" by blast+
note props = props this
let ?Π = "«[λy [O!]y ∨ q⇩0]»"
AOT_modally_strict {
AOT_have ‹[«?Π»]↓› by "cqt:2[lambda]"
} note 1 = this
moreover AOT_have ‹❙𝒜[«?Π»]b & ❙Δ[«?Π»]b & ¬❙𝒜[«?Π»]a & ❙Δ[«?Π»]a›
proof (safe intro!: "&I"; AOT_subst ‹[λy O!y ∨ q⇩0]x› ‹O!x ∨ q⇩0› for: x)
AOT_show ‹❙𝒜([O!]b ∨ q⇩0)›
by (meson "Act-Basic:9" act_ord_b "∨I"(1) "≡E"(2))
next AOT_show ‹❙Δ([O!]b ∨ q⇩0)›
by (meson "KBasic:15" b_ord "∨I"(1) nec_delta "oa-facts:1" "→E")
next AOT_show ‹¬❙𝒜([O!]a ∨ q⇩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⇩0)›
proof (rule not_act_and_pos_delta)
AOT_show ‹¬❙𝒜([O!]a ∨ q⇩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⇩0)›
using "KBasic2:2" "&E"(1) "∨I"(2) "≡E"(2) q⇩0_prop by blast
qed
qed(auto simp: "beta-C-meta"[THEN "→E", OF 1])
ultimately AOT_obtain F⇩1⇩3 where ‹❙𝒜[F⇩1⇩3]b & ❙Δ[F⇩1⇩3]b & ¬❙𝒜[F⇩1⇩3]a & ❙Δ[F⇩1⇩3]a›
using "∃I"(1)[rotated, THEN "∃E"[rotated]] by fastforce
AOT_hence ‹❙𝒜[F⇩1⇩3]b› and ‹❙Δ[F⇩1⇩3]b› and ‹¬❙𝒜[F⇩1⇩3]a› and ‹❙Δ[F⇩1⇩3]a›
using "&E" by blast+
note props = props this
let ?Π = "«[λy [O!]y ∨ ¬q⇩0]»"
AOT_modally_strict {
AOT_have ‹[«?Π»]↓› by "cqt:2[lambda]"
} note 1 = this
moreover AOT_have ‹❙𝒜[«?Π»]b & ❙Δ[«?Π»]b & ❙𝒜[«?Π»]a & ¬❙Δ[«?Π»]a›
proof (safe intro!: "&I"; AOT_subst ‹[λy O!y ∨ ¬q⇩0]x› ‹O!x ∨ ¬q⇩0› for: x)
AOT_show ‹❙𝒜([O!]b ∨ ¬q⇩0)›
by (meson "Act-Basic:9" act_not_q_zero "∨I"(2) "≡E"(2))
next AOT_show ‹❙Δ([O!]b ∨ ¬q⇩0)›
by (meson "KBasic:15" b_ord "∨I"(1) nec_delta "oa-facts:1" "→E")
next AOT_show ‹❙𝒜([O!]a ∨ ¬q⇩0)›
by (meson "Act-Basic:9" act_not_q_zero "∨I"(2) "≡E"(2))
next AOT_show ‹¬❙Δ([O!]a ∨ ¬q⇩0)›
proof(rule act_and_pos_not_not_delta)
AOT_show ‹❙𝒜([O!]a ∨ ¬q⇩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⇩0›
by (meson "&E"(1) q⇩0_prop)
ultimately AOT_have 2: ‹◇(¬[O!]a & q⇩0)›
by (metis "KBasic:16" "&I" "vdash-properties:10")
AOT_show ‹◇¬([O!]a ∨ ¬q⇩0)›
proof (AOT_subst (reverse) ‹¬([O!]a ∨ ¬q⇩0)› ‹¬[O!]a & q⇩0›)
AOT_modally_strict {
AOT_show ‹¬[O!]a & q⇩0 ≡ ¬([O!]a ∨ ¬q⇩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⇩0)›
using "2" by blast
qed
qed
qed(auto simp: "beta-C-meta"[THEN "→E", OF 1])
ultimately AOT_obtain F⇩1⇩4 where ‹❙𝒜[F⇩1⇩4]b & ❙Δ[F⇩1⇩4]b & ❙𝒜[F⇩1⇩4]a & ¬❙Δ[F⇩1⇩4]a›
using "∃I"(1)[rotated, THEN "∃E"[rotated]] by fastforce
AOT_hence ‹❙𝒜[F⇩1⇩4]b› and ‹❙Δ[F⇩1⇩4]b› and ‹❙𝒜[F⇩1⇩4]a› and ‹¬❙Δ[F⇩1⇩4]a›
using "&E" by blast+
note props = props this
AOT_have ‹[L]↓›
by (rule "=⇩d⇩fI"(2)[OF L_def]) "cqt:2[lambda]"+
moreover AOT_have ‹❙𝒜[L]b & ❙Δ[L]b & ❙𝒜[L]a & ❙Δ[L]a›
proof (safe intro!: "&I")
AOT_show ‹❙𝒜[L]b›
by (meson nec_L "nec-imp-act" "vdash-properties:10")
next AOT_show ‹❙Δ[L]b› using nec_L nec_delta by blast
next AOT_show ‹❙𝒜[L]a› by (meson nec_L "nec-imp-act" "→E")
next AOT_show ‹❙Δ[L]a› using nec_L nec_delta by blast
qed
ultimately AOT_obtain F⇩1⇩5 where ‹❙𝒜[F⇩1⇩5]b & ❙Δ[F⇩1⇩5]b & ❙𝒜[F⇩1⇩5]a & ❙Δ[F⇩1⇩5]a›
using "∃I"(1)[rotated, THEN "∃E"[rotated]] by fastforce
AOT_hence ‹❙𝒜[F⇩1⇩5]b› and ‹❙Δ[F⇩1⇩5]b› and ‹❙𝒜[F⇩1⇩5]a› and ‹❙Δ[F⇩1⇩5]a›
using "&E" by blast+
note props = props this
show ?thesis
by (rule "∃I"(2)[where β=F⇩0]; rule "∃I"(2)[where β=F⇩1];
rule "∃I"(2)[where β=F⇩2]; rule "∃I"(2)[where β=F⇩3];
rule "∃I"(2)[where β=F⇩4]; rule "∃I"(2)[where β=F⇩5];
rule "∃I"(2)[where β=F⇩6]; rule "∃I"(2)[where β=F⇩7];
rule "∃I"(2)[where β=F⇩8]; rule "∃I"(2)[where β=F⇩9];
rule "∃I"(2)[where β=F⇩1⇩0]; rule "∃I"(2)[where β=F⇩1⇩1];
rule "∃I"(2)[where β=F⇩1⇩2]; rule "∃I"(2)[where β=F⇩1⇩3];
rule "∃I"(2)[where β=F⇩1⇩4]; rule "∃I"(2)[where β=F⇩1⇩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 & ¬❙𝒜[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 "=⇩d⇩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 "=⇩d⇩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 :: ‹Π› ("'(=⇩E')")
"=E": ‹(=⇩E) =⇩d⇩f [λxy O!x & O!y & □∀F ([F]x ≡ [F]y)]›
syntax "_AOT_eq_E_infix" :: ‹τ ⇒ τ ⇒ φ› (infixl "=⇩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]": ‹[(=⇩E)]↓›
by (rule "=⇩d⇩fI"(2)[OF "=E"]) "cqt:2[lambda]"+
AOT_theorem "=E-simple:1": ‹x =⇩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 "=⇩d⇩fI"(2)[OF "=E"]; "cqt:2[lambda]"?)
using "beta-C-meta"[THEN "→E", OF 1, unvarify ν⇩1ν⇩n, of "(_,_)",
OF tuple_denotes[THEN "≡⇩d⇩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 =⇩E y → x = y›
proof (rule "→I")
AOT_assume ‹x =⇩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 "≡⇩d⇩fI"[OF "identity:1"] "∨I" by blast
qed
AOT_theorem "id-nec3:1": ‹x =⇩E y ≡ □(x =⇩E y)›
proof (rule "≡I"; rule "→I")
AOT_assume ‹x =⇩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 =⇩E y)›
using "=E-simple:1"
by (AOT_subst ‹x =⇩E y› ‹O!x & O!y & □∀F ([F]x ≡ [F]y)›) auto
next
AOT_assume ‹□(x =⇩E y)›
AOT_thus ‹x =⇩E y› using "qml:2"[axiom_inst, THEN "→E"] by blast
qed
AOT_theorem "id-nec3:2": ‹◇(x =⇩E y) ≡ x =⇩E y›
by (meson "RE◇" "S5Basic:2" "id-nec3:1" "≡E"(1,5) "Commutativity of ≡")
AOT_theorem "id-nec3:3": ‹◇(x =⇩E y) ≡ □(x =⇩E y)›
by (meson "id-nec3:1" "id-nec3:2" "≡E"(5))
syntax "_AOT_non_eq_E" :: ‹Π› ("'(≠⇩E')")
translations
(Π) "(≠⇩E)" == (Π) "(=⇩E)⇧-"
syntax "_AOT_non_eq_E_infix" :: ‹τ ⇒ τ ⇒ φ› (infixl "≠⇩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 ≠⇩E y ≡ ¬(x =⇩E y)›
proof -
AOT_have θ: ‹[λx⇩1...x⇩2 ¬(=⇩E)x⇩1...x⇩2]↓› by "cqt:2"
AOT_have ‹x ≠⇩E y ≡ [λx⇩1...x⇩2 ¬(=⇩E)x⇩1...x⇩2]xy›
by (rule "=⇩d⇩fI"(1)[OF "df-relation-negation", OF θ])
(meson "oth-class-taut:3:a")
also AOT_have ‹… ≡ ¬(=⇩E)xy›
by (safe intro!: "beta-C-meta"[THEN "→E", unvarify ν⇩1ν⇩n] "cqt:2"
tuple_denotes[THEN "≡⇩d⇩fI"] "&I")
finally show ?thesis.
qed
AOT_theorem "id-nec4:1": ‹x ≠⇩E y ≡ □(x ≠⇩E y)›
proof -
AOT_have ‹x ≠⇩E y ≡ ¬(x =⇩E y)› using "thm-neg=E".
also AOT_have ‹… ≡ ¬◇(x =⇩E y)›
by (meson "id-nec3:2" "≡E"(1) "Commutativity of ≡" "oth-class-taut:4:b")
also AOT_have ‹… ≡ □¬(x =⇩E y)›
by (meson "KBasic2:1" "≡E"(2) "Commutativity of ≡")
also AOT_have ‹… ≡ □(x ≠⇩E y)›
by (AOT_subst (reverse) ‹¬(x =⇩E y)› ‹x ≠⇩E y›)
(auto simp: "thm-neg=E" "oth-class-taut:3:a")
finally show ?thesis.
qed
AOT_theorem "id-nec4:2": ‹◇(x ≠⇩E y) ≡ (x ≠⇩E y)›
by (meson "RE◇" "S5Basic:2" "id-nec4:1" "≡E"(2,5) "Commutativity of ≡")
AOT_theorem "id-nec4:3": ‹◇(x ≠⇩E y) ≡ □(x ≠⇩E y)›
by (meson "id-nec4:1" "id-nec4:2" "≡E"(5))
AOT_theorem "id-act2:1": ‹x =⇩E y ≡ ❙𝒜x =⇩E y›
by (meson "Act-Basic:5" "Act-Sub:2" "RA[2]" "id-nec3:2" "≡E"(1,6))
AOT_theorem "id-act2:2": ‹x ≠⇩E y ≡ ❙𝒜x ≠⇩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 =⇩E x›
proof (rule "→I")
AOT_assume 1: ‹O!x›
AOT_show ‹x =⇩E x›
apply (rule "=⇩d⇩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 =⇩E y → y =⇩E x›
proof(rule CP)
AOT_assume 1: ‹x =⇩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 =⇩E x› using "ord=Eequiv:1" "→E" by blast
AOT_thus ‹y =⇩E x› using "rule=E"[rotated, OF 2] by fast
qed
AOT_theorem "ord=Eequiv:3": ‹(x =⇩E y & y =⇩E z) → x =⇩E z›
proof (rule CP)
AOT_assume 1: ‹x =⇩E y & y =⇩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 =⇩E x›
using 1[THEN "&E"(1)] "&E"(1) "=E-simple:1" "≡E"(1)
"ord=Eequiv:1" "→E" by blast
ultimately AOT_show ‹x =⇩E z›
using "rule=E" by fast
qed
AOT_theorem "ord-=E=:1": ‹(O!x ∨ O!y) → □(x = y ≡ x =⇩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 =⇩E y)›
proof (rule "→I"; rule "≡I"; rule "→I")
AOT_assume ‹O!x›
AOT_hence ‹x =⇩E x› by (metis "ord=Eequiv:1" "→E")
moreover AOT_assume ‹x = y›
ultimately AOT_show ‹x =⇩E y› using "rule=E" by fast
next
AOT_assume ‹x =⇩E y›
AOT_thus ‹x = y› by (metis "=E-simple:2" "→E")
qed
}
AOT_hence ‹□O!x → □(x = y ≡ x =⇩E y)› by (metis "RM:1")
}
ultimately AOT_have ‹□(x = y ≡ x =⇩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 =⇩E y)›
proof (rule "→I"; rule "≡I"; rule "→I")
AOT_assume ‹O!y›
AOT_hence ‹y =⇩E y› by (metis "ord=Eequiv:1" "→E")
moreover AOT_assume ‹x = y›
ultimately AOT_show ‹x =⇩E y› using "rule=E" id_sym by fast
next
AOT_assume ‹x =⇩E y›
AOT_thus ‹x = y› by (metis "=E-simple:2" "→E")
qed
}
AOT_hence ‹□O!y → □(x = y ≡ x =⇩E y)› by (metis "RM:1")
}
ultimately AOT_have ‹□(x = y ≡ x =⇩E y)› using "→E" by blast
}
ultimately AOT_show ‹□(x = y ≡ x =⇩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 =⇩E y]↓› by "cqt:2[lambda]"
next
AOT_assume ‹O!y›
AOT_hence 1: ‹□(x = y ≡ x =⇩E y)› for x
using "ord-=E=:1" "→E" "∨I" by blast
AOT_have ‹□(x =⇩E y ≡ x = y)› for x
by (AOT_subst ‹x =⇩E y ≡ x = y› ‹x = y ≡ x =⇩E y›)
(auto simp add: "Commutativity of ≡" 1)
AOT_hence ‹∀x □(x =⇩E y ≡ x = y)› by (rule GEN)
AOT_thus ‹□∀x (x =⇩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 =⇩E y]↓› by "cqt:2[lambda]"
next
AOT_show ‹□∀x∀y ([O!]x & [O!]y & x =⇩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 =⇩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 =⇩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 =⇩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 =⇩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 =⇩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 =⇩E x] ≠ [λz z =⇩E y])›
proof (rule "→I"; rule "≡I"; rule "→I";
rule "≡⇩d⇩fI"[OF "=-infix"]; rule "raa-cor:2")
AOT_assume 0: ‹O!x & O!y›
AOT_assume ‹x ≠ y›
AOT_hence 1: ‹¬(x = y)› using "≡⇩d⇩fE"[OF "=-infix"] by blast
AOT_assume ‹[λz z =⇩E x] = [λz z =⇩E y]›
moreover AOT_have ‹[λz z =⇩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 =⇩E y]x› using "rule=E" by fast
AOT_hence ‹x =⇩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 =⇩E x] ≠ [λz z =⇩E y]›
AOT_hence 0: ‹¬([λz z =⇩E x] = [λz z =⇩E y])›
using "≡⇩d⇩fE"[OF "=-infix"] by blast
AOT_have ‹[λz z =⇩E x]↓› by "cqt:2[lambda]"
AOT_hence ‹[λz z =⇩E x] = [λz z =⇩E x]›
by (metis "rule=I:1")
moreover AOT_assume ‹x = y›
ultimately AOT_have ‹[λz z =⇩E x] = [λz z =⇩E y]›
using "rule=E" by fast
AOT_thus ‹[λz z =⇩E x] = [λz z =⇩E y] & ¬([λz z =⇩E x] = [λz z =⇩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 "≡⇩d⇩fI"[OF "=-infix"]; rule "raa-cor:2")
AOT_assume 0: ‹O!x & O!y›
AOT_assume ‹x ≠ y›
AOT_hence 1: ‹¬(x = y)› using "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩fE"]; "cqt:2")
AOT_theorem "denote=:2": ‹∀G∃x⇩1...∃x⇩n x⇩1...x⇩n[H]›
by (rule GEN; rule "existence:2"[THEN "≡⇩d⇩fE"]; "cqt:2")
AOT_theorem "denote=:2[2]": ‹∀G∃x⇩1∃x⇩2 x⇩1x⇩2[H]›
by (rule GEN; rule "existence:2[2]"[THEN "≡⇩d⇩fE"]; "cqt:2")
AOT_theorem "denote=:2[3]": ‹∀G∃x⇩1∃x⇩2∃x⇩3 x⇩1x⇩2x⇩3[H]›
by (rule GEN; rule "existence:2[3]"[THEN "≡⇩d⇩fE"]; "cqt:2")
AOT_theorem "denote=:2[4]": ‹∀G∃x⇩1∃x⇩2∃x⇩3∃x⇩4 x⇩1x⇩2x⇩3x⇩4[H]›
by (rule GEN; rule "existence:2[4]"[THEN "≡⇩d⇩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⇩1...∃x⇩n x⇩1...x⇩n[Π]) ≡ ∃H (H = Π)›
using "existence:2" "free-thms:1" "≡E"(6) "≡Df" by blast
AOT_theorem "denote=:4[2]": ‹(∃x⇩1∃x⇩2 x⇩1x⇩2[Π]) ≡ ∃H (H = Π)›
using "existence:2[2]" "free-thms:1" "≡E"(6) "≡Df" by blast
AOT_theorem "denote=:4[3]": ‹(∃x⇩1∃x⇩2∃x⇩3 x⇩1x⇩2x⇩3[Π]) ≡ ∃H (H = Π)›
using "existence:2[3]" "free-thms:1" "≡E"(6) "≡Df" by blast
AOT_theorem "denote=:4[4]": ‹(∃x⇩1∃x⇩2∃x⇩3∃x⇩4 x⇩1x⇩2x⇩3x⇩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 "≡⇩d⇩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": ‹❙ι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 = ❙ιx(A!x & ∀F (x[F] ≡ φ{F})) → (A!y & ∀F (y[F] ≡ φ{F}))›
using "y-in:2" by blast
AOT_theorem "can-ab2": ‹y = ❙ιx(A!x & ∀F (x[F] ≡ φ{F})) → A!y›
proof(rule "→I")
AOT_assume ‹y = ❙ιx(A!x & ∀F (x[F] ≡ φ{F}))›
AOT_hence ‹❙𝒜(A!y & ∀F (y[F] ≡ φ{F}))›
using "actual-desc:2"[THEN "→E"] by blast
AOT_hence ‹❙𝒜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": ‹❙ιx(A!x & ∀F (x[F] ≡ φ{F}))[F] ≡ φ{F}›
proof -
AOT_have ‹❙ιx(A!x & ∀F (x[F] ≡ φ{F}))↓›
by (simp add: "A-descriptions")
AOT_hence ‹A!❙ιx(A!x & ∀F (x[F] ≡ φ{F})) &
∀F(❙ιx(A!x & ∀F (x[F] ≡ φ{F}))[F] ≡ φ{F})›
using "y-in:3"[THEN "→E"] by blast
AOT_thus ‹❙ιx(A!x & ∀F (x[F] ≡ φ{F}))[F] ≡ φ{F}›
using "&E" "∀E" by blast
qed
AOT_act_theorem "desc-encode:2": ‹❙ιx(A!x & ∀F (x[F] ≡ φ{F}))[G] ≡ φ{G}›
using "desc-encode:1".
AOT_theorem "desc-nec-encode:1":
‹❙ιx (A!x & ∀F (x[F] ≡ φ{F}))[F] ≡ ❙𝒜φ{F}›
proof -
AOT_have 0: ‹❙ιx(A!x & ∀F (x[F] ≡ φ{F}))↓›
by (simp add: "A-descriptions")
AOT_hence ‹❙𝒜(A!❙ιx(A!x & ∀F (x[F] ≡ φ{F})) &
∀F(❙ιx(A!x & ∀F (x[F] ≡ φ{F}))[F] ≡ φ{F}))›
using "actual-desc:4"[THEN "→E"] by blast
AOT_hence ‹❙𝒜∀F (❙ιx(A!x & ∀F (x[F] ≡ φ{F}))[F] ≡ φ{F})›
using "Act-Basic:2" "&E"(2) "≡E"(1) by blast
AOT_hence ‹∀F ❙𝒜(❙ιx(A!x & ∀F (x[F] ≡ φ{F}))[F] ≡ φ{F})›
using "≡E"(1) "logic-actual-nec:3" "vdash-properties:1[2]" by blast
AOT_hence ‹❙𝒜(❙ιx(A!x & ∀F (x[F] ≡ φ{F}))[F] ≡ φ{F})›
using "∀E" by blast
AOT_hence ‹❙𝒜❙ιx(A!x & ∀F (x[F] ≡ φ{F}))[F] ≡ ❙𝒜φ{F}›
using "Act-Basic:5" "≡E"(1) by blast
AOT_thus ‹❙ιx(A!x & ∀F (x[F] ≡ φ{F}))[F] ≡ ❙𝒜φ{F}›
using "en-eq:10[1]"[unvarify x⇩1, OF 0] "≡E"(6) by blast
qed
AOT_theorem "desc-nec-encode:2":
‹❙ιx (A!x & ∀F (x[F] ≡ φ{F}))[G] ≡ ❙𝒜φ{G}›
using "desc-nec-encode:1".
AOT_theorem "Box-desc-encode:1": ‹□φ{G} → ❙ι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} → □(❙ι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} → □(❙ιx(A!x & ∀F (x[F] ≡ φ{G}))[G] ≡ φ{G})›
proof (rule RM; rule "→I")
AOT_modally_strict {
AOT_assume 1: ‹□φ{G}›
AOT_hence ‹❙ι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 ‹❙ιx(A!x & ∀F (x[F] ≡ φ{G}))[G] ≡ φ{G}›
using "→I" "≡I" by simp
}
qed
ultimately AOT_show ‹□(❙ι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 ‹❙⊢⇩□ ∀α (φ{α} → □φ{α})›
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 = ❙ιx(A!x & ∀F (x[F] ≡ φ{F})) → (A!y & ∀F (y[F] ≡ φ{F}))›
proof(rule "→I")
AOT_assume ‹y = ❙ιx(A!x & ∀F (x[F] ≡ φ{F}))›
AOT_hence ‹❙𝒜(A!y & ∀F (y[F] ≡ φ{F}))›
using "actual-desc:2"[THEN "→E"] by fast
AOT_hence abs: ‹❙𝒜A!y› and ‹❙𝒜∀F (y[F] ≡ φ{F})›
using "Act-Basic:2" "&E" "≡E"(1) by blast+
AOT_hence ‹∀F ❙𝒜(y[F] ≡ φ{F})›
by (metis "≡E"(1) "logic-actual-nec:3" "vdash-properties:1[2]")
AOT_hence ‹❙𝒜(y[F] ≡ φ{F})› for F
using "∀E" by blast
AOT_hence ‹❙𝒜y[F] ≡ ❙𝒜φ{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 ‹❙ι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) ≡⇩d⇩f A!x & ¬∃F x[F]›
AOT_define Universal :: ‹τ ⇒ φ› ("Universal'(_')")
"df-null-uni:2": ‹Universal(x) ≡⇩d⇩f A!x & ∀F x[F]›
AOT_theorem "null-uni-uniq:1": ‹∃!x Null(x)›
proof (rule "uniqueness:1"[THEN "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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": ‹❙ιx Null(x)↓›
using "A-Exists:2" "RA[2]" "≡E"(2) "null-uni-uniq:1" by blast
AOT_theorem "null-uni-uniq:4": ‹❙ιx Universal(x)↓›
using "A-Exists:2" "RA[2]" "≡E"(2) "null-uni-uniq:2" by blast
AOT_define Null_object :: ‹κ⇩s› (‹a⇩∅›)
"df-null-uni-terms:1": ‹a⇩∅ =⇩d⇩f ❙ιx Null(x)›
AOT_define Universal_object :: ‹κ⇩s› (‹a⇩V›)
"df-null-uni-terms:2": ‹a⇩V =⇩d⇩f ❙ι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 "≡⇩d⇩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 "≡⇩d⇩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⇩∅)›
apply (rule "=⇩d⇩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⇩V)›
apply (rule "=⇩d⇩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⇩∅ ≠ a⇩V›
proof (rule "=⇩d⇩fI"(2)[OF "df-null-uni-terms:1", OF "null-uni-uniq:3"];
rule "=⇩d⇩fI"(2)[OF "df-null-uni-terms:2", OF "null-uni-uniq:4"];
rule "≡⇩d⇩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: ‹❙𝒜Null(x)›
by (metis "nec-imp-act" "null-uni-facts:1" "→E")
AOT_assume ‹❙ιx Null(x) = ❙ιx Universal(x)›
AOT_hence ‹❙𝒜∀x(Null(x) ≡ Universal(x))›
using "actual-desc:5"[THEN "→E"] by blast
AOT_hence ‹∀x ❙𝒜(Null(x) ≡ Universal(x))›
by (metis "≡E"(1) "logic-actual-nec:3" "vdash-properties:1[2]")
AOT_hence ‹❙𝒜Null(x) ≡ ❙𝒜Universal(x)›
using "Act-Basic:5" "≡E"(1) "rule-ui:3" by blast
AOT_hence ‹❙𝒜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 "≡⇩d⇩fE"[OF "df-null-uni:2"] "&E" by metis
moreover AOT_have ‹¬∃F x[F]›
using nullx "≡⇩d⇩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⇩∅ = ❙ιx(A!x & ∀F (x[F] ≡ F ≠ F))›
proof (rule "ab-obey:1"[unvarify x y, THEN "→E", THEN "→E"])
AOT_show ‹❙ιx([A!]x & ∀F (x[F] ≡ F ≠ F))↓›
by (simp add: "A-descriptions")
next
AOT_show ‹a⇩∅↓›
by (rule "=⇩d⇩fI"(2)[OF "df-null-uni-terms:1", OF "null-uni-uniq:3"])
(simp add: "null-uni-uniq:3")
next
AOT_have ‹❙ιx([A!]x & ∀F (x[F] ≡ F ≠ F))↓›
by (simp add: "A-descriptions")
AOT_hence 1: ‹❙ιx([A!]x & ∀F (x[F] ≡ F ≠ F)) = ❙ιx([A!]x & ∀F (x[F] ≡ F ≠ F))›
using "rule=I:1" by blast
AOT_show ‹[A!]a⇩∅ & [A!]❙ιx([A!]x & ∀F (x[F] ≡ F ≠ F))›
apply (rule "=⇩d⇩fI"(2)[OF "df-null-uni-terms:1", OF "null-uni-uniq:3"];
rule "&I")
apply (meson "≡⇩d⇩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⇩∅[F] ≡ ❙ιx([A!]x & ∀F (x[F] ≡ F ≠ F))[F])›
proof (rule GEN)
fix F
AOT_have ‹¬a⇩∅[F]›
by (rule "=⇩d⇩fI"(2)[OF "df-null-uni-terms:1", OF "null-uni-uniq:3"])
(metis (no_types, lifting) "≡⇩d⇩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 ‹¬❙ιx([A!]x & ∀F (x[F] ≡ F ≠ F))[F]›
proof(rule "raa-cor:2")
AOT_assume 0: ‹❙ιx([A!]x & ∀F (x[F] ≡ F ≠ F))[F]›
AOT_hence ‹❙𝒜(F ≠ F)›
using "desc-nec-encode:2"[THEN "≡E"(1), OF 0] by blast
moreover AOT_have ‹¬❙𝒜(F ≠ F)›
using "≡⇩d⇩fE" "id-act:2" "id-eq:1" "≡E"(2)
"=-infix" "raa-cor:3" by blast
ultimately AOT_show ‹❙𝒜(F ≠ F) & ¬❙𝒜(F ≠ F)› by (rule "&I")
qed
ultimately AOT_show ‹a⇩∅[F] ≡ ❙ι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⇩V = ❙ιx(A!x & ∀F (x[F] ≡ F = F))›
proof (rule "ab-obey:1"[unvarify x y, THEN "→E", THEN "→E"])
AOT_show ‹❙ιx([A!]x & ∀F (x[F] ≡ F = F))↓›
by (simp add: "A-descriptions")
next
AOT_show ‹a⇩V↓›
by (rule "=⇩d⇩fI"(2)[OF "df-null-uni-terms:2", OF "null-uni-uniq:4"])
(simp add: "null-uni-uniq:4")
next
AOT_have ‹❙ιx([A!]x & ∀F (x[F] ≡ F = F))↓›
by (simp add: "A-descriptions")
AOT_hence 1: ‹❙ιx([A!]x & ∀F (x[F] ≡ F = F)) = ❙ιx([A!]x & ∀F (x[F] ≡ F = F))›
using "rule=I:1" by blast
AOT_show ‹[A!]a⇩V & [A!]❙ιx([A!]x & ∀F (x[F] ≡ F = F))›
apply (rule "=⇩d⇩fI"(2)[OF "df-null-uni-terms:2", OF "null-uni-uniq:4"];
rule "&I")
apply (meson "≡⇩d⇩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⇩V[F] ≡ ❙ιx([A!]x & ∀F (x[F] ≡ F = F))[F])›
proof (rule GEN)
fix F
AOT_have ‹a⇩V[F]›
apply (rule "=⇩d⇩fI"(2)[OF "df-null-uni-terms:2", OF "null-uni-uniq:4"])
using "≡⇩d⇩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 ‹❙ι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⇩V[F] ≡ ❙ι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 "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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⇩1...x⇩n φ{x⇩1...x⇩n}]↓ ≡ □∀x⇩1...∀x⇩n∀y⇩1...∀y⇩n
(∀F([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n) → (φ{x⇩1...x⇩n} ≡ φ{y⇩1...y⇩n}))›
proof(rule "≡I"; rule "→I")
AOT_assume ‹[λx⇩1...x⇩n φ{x⇩1...x⇩n}]↓›
AOT_hence ‹□[λx⇩1...x⇩n φ{x⇩1...x⇩n}]↓› by (metis "exist-nec" "→E")
moreover AOT_have ‹□[λx⇩1...x⇩n φ{x⇩1...x⇩n}]↓ → □∀x⇩1...∀x⇩n∀y⇩1...∀y⇩n
(∀F([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n) → (φ{x⇩1...x⇩n} ≡ φ{y⇩1...y⇩n}))›
proof (rule "RM:1"; rule "→I"; rule GEN; rule GEN; rule "→I")
AOT_modally_strict {
fix x⇩1x⇩n y⇩1y⇩n :: ‹'a AOT_var›
AOT_assume 0: ‹[λx⇩1...x⇩n φ{x⇩1...x⇩n}]↓›
moreover AOT_assume ‹∀F([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n)›
ultimately AOT_have ‹[λx⇩1...x⇩n φ{x⇩1...x⇩n}]x⇩1...x⇩n ≡
[λx⇩1...x⇩n φ{x⇩1...x⇩n}]y⇩1...y⇩n›
using "∀E" by blast
AOT_thus ‹(φ{x⇩1...x⇩n} ≡ φ{y⇩1...y⇩n})›
using "beta-C-meta"[THEN "→E", OF 0] "≡E"(6) by meson
}
qed
ultimately AOT_show ‹□∀x⇩1...∀x⇩n∀y⇩1...∀y⇩n(
∀F([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n) → (φ{x⇩1...x⇩n} ≡ φ{y⇩1...y⇩n})
)›
using "→E" by blast
next
AOT_have ‹
□(∀x⇩1...∀x⇩n∀y⇩1...∀y⇩n
(∀F([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n) → (φ{x⇩1...x⇩n} ≡ φ{y⇩1...y⇩n})))
→ □∀y⇩1...∀y⇩n
((∃x⇩1...∃x⇩n(∀F([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n) & φ{x⇩1...x⇩n})) ≡
φ{y⇩1...y⇩n})›
proof(rule "RM:1"; rule "→I"; rule GEN)
AOT_modally_strict {
AOT_assume ‹∀x⇩1...∀x⇩n∀y⇩1...∀y⇩n
(∀F([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n) → (φ{x⇩1...x⇩n} ≡ φ{y⇩1...y⇩n}))›
AOT_hence indisc: ‹φ{x⇩1...x⇩n} ≡ φ{y⇩1...y⇩n}›
if ‹∀F([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n)› for x⇩1x⇩n y⇩1y⇩n
using "∀E"(2) "→E" that by blast
AOT_show ‹(∃x⇩1...∃x⇩n(∀F([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n) & φ{x⇩1...x⇩n})) ≡
φ{y⇩1...y⇩n}› for y⇩1y⇩n
proof (rule "raa-cor:1")
AOT_assume ‹¬((∃x⇩1...∃x⇩n(∀F([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n) & φ{x⇩1...x⇩n})) ≡
φ{y⇩1...y⇩n})›
AOT_hence ‹((∃x⇩1...∃x⇩n(∀F([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n)
& φ{x⇩1...x⇩n}))
& ¬φ{y⇩1...y⇩n}) ∨
(¬(∃x⇩1...∃x⇩n(∀F([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n) & φ{x⇩1...x⇩n}))
& φ{y⇩1...y⇩n})›
using "≡E"(1) "oth-class-taut:4:h" by blast
moreover {
AOT_assume 0: ‹(∃x⇩1...∃x⇩n(∀F([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n) & φ{x⇩1...x⇩n}))
& ¬φ{y⇩1...y⇩n}›
AOT_obtain a⇩1a⇩n where ‹∀F([F]a⇩1...a⇩n ≡ [F]y⇩1...y⇩n) & φ{a⇩1...a⇩n}›
using "∃E"[rotated, OF 0[THEN "&E"(1)]] by blast
AOT_hence ‹φ{y⇩1...y⇩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⇩1...∃x⇩n(∀F([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n) & φ{x⇩1...x⇩n}))
& φ{y⇩1...y⇩n}›
AOT_hence ‹∀x⇩1...∀x⇩n ¬(∀F([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n) & φ{x⇩1...x⇩n})›
using "&E"(1) "cqt-further:4" "→E" by blast
AOT_hence ‹¬(∀F([F]y⇩1...y⇩n ≡ [F]y⇩1...y⇩n) & φ{y⇩1...y⇩n})›
using "∀E" by blast
AOT_hence ‹¬∀F([F]y⇩1...y⇩n ≡ [F]y⇩1...y⇩n) ∨ ¬φ{y⇩1...y⇩n}›
using "≡E"(1) "oth-class-taut:5:c" by blast
moreover AOT_have ‹∀F([F]y⇩1...y⇩n ≡ [F]y⇩1...y⇩n)›
by (simp add: "oth-class-taut:3:a" "universal-cor")
ultimately AOT_have ‹¬φ{y⇩1...y⇩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⇩1...∀x⇩n∀y⇩1...∀y⇩n
(∀F([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n) → (φ{x⇩1...x⇩n} ≡ φ{y⇩1...y⇩n}))›
ultimately AOT_have ‹□∀y⇩1...∀y⇩n
((∃x⇩1...∃x⇩n(∀F([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n) & φ{x⇩1...x⇩n})) ≡
φ{y⇩1...y⇩n})›
using "→E" by blast
AOT_thus ‹[λx⇩1...x⇩n φ{x⇩1...x⇩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⇩1...x⇩n φ{x⇩1...x⇩n}]↓ → ∀x⇩1...∀x⇩n∀y⇩1...∀y⇩n
(∀F([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n) → □(φ{x⇩1...x⇩n} ≡ φ{y⇩1...y⇩n}))›
proof(rule "→I"; rule GEN; rule GEN; rule "→I")
fix x⇩1x⇩n y⇩1y⇩n
AOT_assume ‹[λx⇩1...x⇩n φ{x⇩1...x⇩n}]↓›
AOT_hence 0: ‹□∀x⇩1...∀x⇩n∀y⇩1...∀y⇩n
(∀F ([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n) → (φ{x⇩1...x⇩n} ≡ φ{y⇩1...y⇩n}))›
by (rule "kirchner-thm:2"[THEN "≡E"(1)])
AOT_have ‹∀x⇩1...∀x⇩n∀y⇩1...∀y⇩n
□(∀F ([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n) → (φ{x⇩1...x⇩n} ≡ φ{y⇩1...y⇩n}))›
proof(rule GEN; rule GEN)
fix x⇩1x⇩n y⇩1y⇩n
AOT_show ‹□(∀F ([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n) → (φ{x⇩1...x⇩n} ≡ φ{y⇩1...y⇩n}))›
apply (rule "RM:1"[THEN "→E", rotated, OF 0]; rule "→I")
using "∀E" by blast
qed
AOT_hence ‹∀y⇩1...∀y⇩n □(∀F ([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n) →
(φ{x⇩1...x⇩n} ≡ φ{y⇩1...y⇩n}))›
using "∀E" by blast
AOT_hence ‹□(∀F ([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n) → (φ{x⇩1...x⇩n} ≡ φ{y⇩1...y⇩n}))›
using "∀E" by blast
AOT_hence ‹□(∀F ([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n) → (φ{x⇩1...x⇩n} ≡ φ{y⇩1...y⇩n}))›
using "∀E" by blast
AOT_hence 0: ‹□∀F ([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n) → □(φ{x⇩1...x⇩n} ≡ φ{y⇩1...y⇩n})›
using "qml:1"[axiom_inst] "vdash-properties:6" by blast
moreover AOT_assume ‹∀F([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n)›
moreover AOT_have ‹[λx⇩1...x⇩n □∀F ([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n)]↓› by "cqt:2"
ultimately AOT_have ‹[λx⇩1...x⇩n □∀F ([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n)]x⇩1...x⇩n ≡
[λx⇩1...x⇩n □∀F ([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n)]y⇩1...y⇩n›
using "∀E" by blast
moreover AOT_have ‹[λx⇩1...x⇩n □∀F ([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n)]y⇩1...y⇩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⇩1...x⇩n □∀F ([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n)]x⇩1...x⇩n›
using "≡E"(2) by blast
AOT_hence ‹□∀F ([F]x⇩1...x⇩n ≡ [F]y⇩1...y⇩n)›
using "β→C"(1) by blast
AOT_thus ‹□(φ{x⇩1...x⇩n} ≡ φ{y⇩1...y⇩n})› using "→E" 0 by blast
qed
subsection‹Propositional Properties›
text‹\label{PLM: 9.12}›
AOT_define propositional :: ‹Π ⇒ φ› (‹Propositional'(_')›)
"prop-prop1": ‹Propositional([F]) ≡⇩d⇩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 "≡⇩d⇩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]) ≡⇩d⇩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 "≡⇩d⇩fE"[OF "prop-prop1"] by blast
then AOT_obtain p where Π_def: ‹Π = [λy p]› using "∃E"[rotated] by blast
AOT_show ‹Indiscriminate([Π])›
proof (rule "≡⇩d⇩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⇩1...∀x⇩n [F]x⇩1...x⇩n›
using "≡⇩d⇩fE"[OF "contingent-properties:1"] by blast
AOT_show ‹Indiscriminate([F])›
by (rule "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩fE" "conventions:5" "o-objects-exist:5" "reductio-aa:1")
qed
AOT_theorem "prop-in-f:3:b": ‹¬Indiscriminate([E!]⇧-)›
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 "≡⇩d⇩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 "≡⇩d⇩fE" "conventions:5" "reductio-aa:1")
AOT_have ‹A!a›
apply (rule "=⇩d⇩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 "≡⇩d⇩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 "≡⇩d⇩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!⇧-)›
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