diff --git a/Foundation.lean b/Foundation.lean index c3d0e6f61..61150a69b 100644 --- a/Foundation.lean +++ b/Foundation.lean @@ -98,6 +98,7 @@ import Foundation.Modal.Kripke.Logic.GrzPoint3 import Foundation.Modal.Kripke.Logic.K4M import Foundation.Modal.Kripke.Logic.K4Point2 import Foundation.Modal.Kripke.Logic.K4Point3 +import Foundation.Modal.Kripke.Logic.K4n import Foundation.Modal.Kripke.Logic.KHen import Foundation.Modal.Kripke.Logic.KT4B import Foundation.Modal.Kripke.Logic.KTc diff --git a/Foundation/Modal/Axioms.lean b/Foundation/Modal/Axioms.lean index 2f3bd2d2d..57f3b2447 100644 --- a/Foundation/Modal/Axioms.lean +++ b/Foundation/Modal/Axioms.lean @@ -28,6 +28,8 @@ protected abbrev P : F := ∼(□⊥) /-- Axiom for transivity -/ protected abbrev Four := □φ ➝ □□φ +protected abbrev FourN (n : ℕ+) (φ : F) := □^[n]φ ➝ □^[(n + 1)]φ + /-- Axiom for euclidean -/ protected abbrev Five := ◇φ ➝ □◇φ diff --git a/Foundation/Modal/Entailment/Basic.lean b/Foundation/Modal/Entailment/Basic.lean index 5e17d1b56..5c5e877c2 100644 --- a/Foundation/Modal/Entailment/Basic.lean +++ b/Foundation/Modal/Entailment/Basic.lean @@ -240,6 +240,24 @@ def axiomFour'! (h : 𝓢 ⊢! □φ) : 𝓢 ⊢! □□φ := ⟨axiomFour' h.so end +class HasAxiomFourN (n) (𝓢 : S) where + FourN (φ : F) : 𝓢 ⊢ Axioms.FourN n φ + +section + +variable [HasAxiomFourN n 𝓢] + +def axiomFourN : 𝓢 ⊢ □^[n]φ ➝ □^[(n + 1)]φ := by apply HasAxiomFourN.FourN +@[simp] lemma axiomFourN! : 𝓢 ⊢! □^[n]φ ➝ □^[(n + 1)]φ := ⟨axiomFourN⟩ + +variable [Entailment.Minimal 𝓢] + +instance (Γ : FiniteContext F 𝓢) : HasAxiomFourN n Γ := ⟨fun _ ↦ FiniteContext.of axiomFourN⟩ +instance (Γ : Context F 𝓢) : HasAxiomFourN n Γ := ⟨fun _ ↦ Context.of axiomFourN⟩ + +end + + class HasAxiomFive [Dia F] (𝓢 : S) where Five (φ : F) : 𝓢 ⊢ Axioms.Five φ @@ -542,6 +560,7 @@ instance [Entailment.HasAxiomT 𝓢] : Entailment.HasAxiomGeach ⟨0, 0, 1, instance [Entailment.HasAxiomB 𝓢] : Entailment.HasAxiomGeach ⟨0, 1, 0, 1⟩ 𝓢 := ⟨fun _ => axiomB⟩ instance [Entailment.HasAxiomD 𝓢] : Entailment.HasAxiomGeach ⟨0, 0, 1, 1⟩ 𝓢 := ⟨fun _ => axiomD⟩ instance [Entailment.HasAxiomFour 𝓢] : Entailment.HasAxiomGeach ⟨0, 2, 1, 0⟩ 𝓢 := ⟨fun _ => axiomFour⟩ +instance [Entailment.HasAxiomFourN n 𝓢] : HasAxiomGeach ⟨0, n + 1, n, 0⟩ 𝓢 := ⟨fun _ ↦ axiomFourN⟩ instance [Entailment.HasAxiomFive 𝓢] : Entailment.HasAxiomGeach ⟨1, 1, 0, 1⟩ 𝓢 := ⟨fun _ => axiomFive⟩ instance [Entailment.HasAxiomTc 𝓢] : Entailment.HasAxiomGeach ⟨0, 1, 0, 0⟩ 𝓢 := ⟨fun _ => axiomTc⟩ instance [Entailment.HasAxiomPoint2 𝓢] : Entailment.HasAxiomGeach ⟨1, 1, 1, 1⟩ 𝓢 := ⟨fun _ => axiomPoint2⟩ diff --git a/Foundation/Modal/Hilbert/WellKnown.lean b/Foundation/Modal/Hilbert/WellKnown.lean index bb572b2e7..3468e6416 100644 --- a/Foundation/Modal/Hilbert/WellKnown.lean +++ b/Foundation/Modal/Hilbert/WellKnown.lean @@ -70,6 +70,21 @@ instance [H.HasFour] : Entailment.HasAxiomFour H.logic where . use (λ b => if (HasFour.p H) = b then φ else (.atom b)); simp; +class HasFourN (H : Hilbert α) (n : ℕ+) where + p : α + mem_FourN : Axioms.FourN n (.atom p) ∈ H.axioms := by tauto; + +instance [H.HasFourN n] : Entailment.HasAxiomFourN n H.logic where + FourN φ := by + constructor; + apply maxm; + use Axioms.FourN n (.atom (HasFourN.p H n)); + constructor; + . exact HasFourN.mem_FourN; + . use (λ b => if (HasFourN.p H n) = b then φ else (.atom b)); + simp; + + class HasFive (H : Hilbert α) where p : α mem_Five : Axioms.Five (.atom p) ∈ H.axioms := by tauto; @@ -375,6 +390,11 @@ instance : (Hilbert.K4).HasK where p := 0; q := 1; instance : (Hilbert.K4).HasFour where p := 0 instance : Entailment.K4 (Logic.K4) where +protected abbrev Hilbert.K4n (n : ℕ+) : Hilbert ℕ := ⟨{Axioms.K (.atom 0) (.atom 1), Axioms.FourN n (.atom 0)}⟩ +protected abbrev Logic.K4n (n : ℕ+) := Hilbert.K4n n |>.logic +instance : (Hilbert.K4n n).HasK where p := 0; q := 1; +instance : (Hilbert.K4n n).HasFourN n where p := 0; + protected abbrev Hilbert.K4M : Hilbert ℕ := ⟨{Axioms.K (.atom 0) (.atom 1), Axioms.Four (.atom 0), Axioms.M (.atom 0)}⟩ protected abbrev Logic.K4M := Hilbert.K4M.logic diff --git a/Foundation/Modal/Kripke/AxiomFourN.lean b/Foundation/Modal/Kripke/AxiomFourN.lean new file mode 100644 index 000000000..673d5d9d0 --- /dev/null +++ b/Foundation/Modal/Kripke/AxiomFourN.lean @@ -0,0 +1,59 @@ +import Foundation.Modal.Kripke.AxiomGeach +import Foundation.Modal.Kripke.Filtration + + +namespace LO.Modal + +open Kripke + +namespace Kripke + +namespace Frame + +variable {F : Frame} {n : ℕ+} + +class IsWeakTransitive (F : Kripke.Frame) (n : ℕ+) where + weak_trans : ∀ ⦃x y : F⦄, x ≺^[n + 1] y → x ≺^[n] y + +instance [F.IsPiecewiseStronglyConnected] : F.IsPiecewiseConnected := inferInstance + +lemma weak_trans [F.IsWeakTransitive n] : ∀ {x y : F.World}, x ≺^[n + 1] y → x ≺^[n] y := by + apply IsWeakTransitive.weak_trans + +instance [F.IsGeachConvergent ⟨0, n + 1, n, 0⟩] : F.IsWeakTransitive n := ⟨by + rintro x y ⟨u, Rxu, Ruy⟩; + have : ∀ u, x ≺ u → u ≺^[n] y → x ≺^[n] y := by + simpa using @IsGeachConvergent.gconv (g := ⟨0, n + 1, n, 0⟩) (F := F) _ x x y; + apply this u Rxu Ruy; +⟩ +instance [F.IsWeakTransitive n] : F.IsGeachConvergent ⟨0, n + 1, n, 0⟩ := ⟨by + suffices ∀ ⦃x y z : F⦄, x = y → ∀ (x_1 : F), x ≺ x_1 → x_1 ≺^[n] z → y ≺^[n] z by simpa using this; + rintro x _ y rfl u Rxw Rwz; + apply IsWeakTransitive.weak_trans; + use u; +⟩ + +end Frame + + +variable {F : Kripke.Frame} {n : ℕ+} + +lemma validate_AxiomFourN_of_weakTransitive [F.IsWeakTransitive n] : F ⊧ (Axioms.FourN n (.atom 0)) := validate_axiomGeach_of_isGeachConvergent (g := ⟨0, n + 1, n, 0⟩) + +namespace Canonical + +variable {S} [Entailment (Formula ℕ) S] +variable {𝓢 : S} [Entailment.Consistent 𝓢] [Entailment.K 𝓢] + +open Formula.Kripke +open Entailment +open MaximalConsistentTableau +open canonicalModel + +instance [Entailment.HasAxiomFourN n 𝓢] : (canonicalFrame 𝓢).IsWeakTransitive n := by infer_instance; + +end Canonical + +end Kripke + +end LO.Modal diff --git a/Foundation/Modal/Kripke/Logic/K4n.lean b/Foundation/Modal/Kripke/Logic/K4n.lean new file mode 100644 index 000000000..b1f0139dd --- /dev/null +++ b/Foundation/Modal/Kripke/Logic/K4n.lean @@ -0,0 +1,177 @@ +import Foundation.Modal.Kripke.Logic.K4 +import Foundation.Modal.Kripke.AxiomFourN +import Foundation.Modal.Kripke.Filtration +import Foundation.Vorspiel.Fin.Supplemental +import Foundation.Modal.Logic.Extension + +namespace LO.Modal + +open Kripke +open Hilbert.Kripke + +namespace Kripke + +protected abbrev FrameClass.K4n (n : ℕ+) : FrameClass := { F | F.IsWeakTransitive n } + +end Kripke + +namespace Hilbert + +variable {n : ℕ+} + +namespace K4n.Kripke + +instance : Sound (Logic.K4n n) (FrameClass.K4n n) := instSound_of_validates_axioms $ by + apply FrameClass.Validates.withAxiomK; + rintro F F_trans φ rfl; + replace F_trans := Set.mem_setOf_eq.mp F_trans; + apply validate_AxiomFourN_of_weakTransitive; + +instance : Entailment.Consistent (Logic.K4n n) := + consistent_of_sound_frameclass (FrameClass.K4n n) $ by + use whitepoint; + apply Set.mem_setOf_eq.mpr; + constructor; + induction n <;> + . simp [whitepoint]; + +instance : Canonical (Logic.K4n n) (FrameClass.K4n n) := ⟨by + apply Set.mem_setOf_eq.mpr; + infer_instance; +⟩ + +instance : Complete (Logic.K4n n) (FrameClass.K4n n) := inferInstance + +end K4n.Kripke + + +end Hilbert + + +namespace Logic + +open Formula +open Entailment +open Kripke + +namespace K4n + +lemma Kripke.eq_K4n_logic (n : ℕ+) : Logic.K4n n = (Kripke.FrameClass.K4n n).logic := eq_hilbert_logic_frameClass_logic + +lemma eq_K4_K4n_one : Logic.K4n 1 = Logic.K4 := rfl + + +variable {n m : ℕ+} + +abbrev counterframe (n : ℕ+) : Kripke.Frame := ⟨Fin (n + 2), λ ⟨x, _⟩ ⟨y, _⟩ => y = min (x + 1) (n + 1)⟩ + +abbrev counterframe.last : counterframe n |>.World := ⟨n + 1, by omega⟩ + +lemma counterframe.iff_rel_from_zero {m : ℕ} {x : counterframe n} : (0 : counterframe n) ≺^[m] x ↔ x.1 = min m (n + 1) := by + induction m generalizing x with + | zero => + simp; + tauto; + | succ m ih => + constructor; + . intro R0x; + obtain ⟨u, R0u, Rux⟩ := HRel.iterate.forward.mp R0x; + have := ih.mp R0u; + simp_all; + . rintro h; + apply HRel.iterate.forward.mpr; + by_cases hm : m ≤ n + 1; + . use ⟨m, by omega⟩; + constructor; + . apply ih.mpr; + simpa; + . simp_all; + . use x; + constructor; + . apply ih.mpr; + omega; + . omega; + +@[simp] +lemma counterframe.refl_last : (last : counterframe n) ≺ last := by simp [Frame.Rel']; + +instance : Frame.IsWeakTransitive (counterframe n) (n + 1) := by + constructor; + intro x y Rxy; + by_cases ey : y = counterframe.last; + . subst ey; + sorry; + . sorry; + +instance succ_strictlyWeakerThan : Logic.K4n (n + 1) ⪱ Logic.K4n n := by + constructor; + . apply Hilbert.weakerThan_of_provable_axioms; + rintro φ (rfl | rfl); + . simp; + . suffices (Hilbert.K4n n).logic ⊢! □□^[n](.atom 0) ➝ □□^[(n + 1)](.atom 0) by + simpa [Axioms.FourN, PNat.add_coe, PNat.val_ofNat, Box.multibox_succ]; + apply imply_box_distribute'!; + exact axiomFourN!; + . apply Entailment.not_weakerThan_iff.mpr; + use Axioms.FourN n (.atom 0); + constructor; + . simp; + . apply Sound.not_provable_of_countermodel (𝓜 := FrameClass.K4n (n + 1)) + apply not_validOnFrameClass_of_exists_frame; + use (counterframe n); + constructor; + . apply Set.mem_setOf_eq.mpr; + infer_instance; + . apply ValidOnFrame.not_of_exists_valuation_world; + use (λ w a => w ≠ counterframe.last), 0; + apply Satisfies.not_imp_def.mpr; + constructor; + . apply Satisfies.multibox_def.mpr; + intro y R0y; + replace R0y := counterframe.iff_rel_from_zero.mp R0y; + simp only [Semantics.Realize, Satisfies, counterframe.last, ne_eq]; + aesop; + . apply Satisfies.multibox_def.not.mpr; + push_neg; + use counterframe.last; + constructor; + . apply counterframe.iff_rel_from_zero.mpr; + simp; + . simp [Semantics.Realize, Satisfies]; + +instance add_strictlyWeakerThan : Logic.K4n (n + m) ⪱ Logic.K4n n := by + induction m with + | one => infer_instance; + | succ m ih => + trans Logic.K4n (n + m); + . rw [(show n + (m + 1) = n + m + 1 by rfl)]; + infer_instance; + . apply ih; + +instance strictlyWeakerThan_of_lt (hnm : n < m) : Logic.K4n m ⪱ Logic.K4n n := by + convert add_strictlyWeakerThan (n := n) (m := m - n); + exact PNat.add_sub_of_lt hnm |>.symm; + +instance not_equiv_of_ne (hnm : n ≠ m) : ¬(Logic.K4n n ≊ Logic.K4n m) := by + rcases lt_trichotomy n m with h | rfl | h; + . by_contra!; + apply strictlyWeakerThan_of_lt h |>.notWT; + exact this.le; + . contradiction; + . by_contra!; + apply strictlyWeakerThan_of_lt h |>.notWT; + exact this.symm.le; + +end K4n + +instance : Infinite { L : Logic ℕ // Entailment.Consistent L } := Infinite.of_injective (β := ℕ+) (λ n => ⟨Logic.K4n n, inferInstance⟩) $ by + intro i j; + simp only [Subtype.mk.injEq]; + contrapose!; + intro h; + apply Logic.iff_equal_provable_equiv.not.mpr; + apply K4n.not_equiv_of_ne h; + +end Logic + +end LO.Modal diff --git a/Foundation/Vorspiel/Fin/Supplemental.lean b/Foundation/Vorspiel/Fin/Supplemental.lean index 262d7744b..168def2a3 100644 --- a/Foundation/Vorspiel/Fin/Supplemental.lean +++ b/Foundation/Vorspiel/Fin/Supplemental.lean @@ -12,7 +12,17 @@ lemma isEmpty_embedding_lt (hn : n > m) : IsEmpty (Fin n ↪ Fin m) := by apply Function.Embedding.isEmpty_of_card_lt; simpa; -variable {n : ℕ} [NeZero n] {i : Fin n} +variable {n : ℕ} {i : Fin n} + +@[simp] +lemma lt_last : n < Fin.last (n + 1) := by + induction n with + | zero => simp; + | succ n ih => simp; + +section + +variable [NeZero n] def last' : Fin n := ⟨n - 1, Nat.sub_one_lt'⟩ @@ -21,4 +31,6 @@ lemma lt_last' : i ≤ Fin.last' := by apply Nat.le_sub_one_of_lt; apply Fin.is_lt; +end + end Fin