From 8a92a5293c6c8edcb73b7a7c280a663f1ea3e8a1 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Sun, 1 Oct 2023 20:40:31 +0300 Subject: [PATCH 01/63] first syntax and semantics draft --- src/modal-logic/modal-logic-syntax.lagda.md | 248 ++++++++++++++++++++ 1 file changed, 248 insertions(+) create mode 100644 src/modal-logic/modal-logic-syntax.lagda.md diff --git a/src/modal-logic/modal-logic-syntax.lagda.md b/src/modal-logic/modal-logic-syntax.lagda.md new file mode 100644 index 0000000000..5bc10c121f --- /dev/null +++ b/src/modal-logic/modal-logic-syntax.lagda.md @@ -0,0 +1,248 @@ +# Modal logic syntax + +```agda +module modal-logic.modal-logic-syntax where +``` + +
Imports + +```agda +-- open import foundation.contractible-types +-- open import foundation.dependent-pair-types +-- open import foundation.equivalences +-- open import foundation.identity-types +-- open import foundation.propositions +-- open import foundation.sets +open import foundation.booleans +open import foundation.binary-relations +open import foundation.decidable-types +open import foundation.sets +open import foundation.powersets +open import foundation-core.propositions +open import foundation.universe-levels +open import foundation.unit-type + +open import foundation-core.coproduct-types +open import foundation-core.empty-types +open import foundation-core.propositions +open import foundation-core.subtypes +open import foundation-core.negation +open import foundation-core.identity-types + +open import elementary-number-theory.natural-numbers +``` + +
+ +## Idea + +TODO + +## Definition + +**Syntax** + +```agda +data formula : UU lzero where + var : ℕ → formula + ⊥ : formula + _⇒_ : formula → formula → formula + □_ : formula → formula + +~_ : formula → formula +~ a = a ⇒ ⊥ + +_∨_ : formula → formula → formula +a ∨ b = (~ a) ⇒ b + +_∧_ : formula → formula → formula +a ∧ b = ~ (a ⇒ (~ b)) + +⊤ : formula +⊤ = ~ ⊥ + +◇_ : formula → formula +◇ a = ~ (□ (~ a)) +``` + +**Tautology** + +```agda +valuate-cl : (ℕ → bool) → formula → bool +valuate-cl v (var x) = v x +valuate-cl v ⊥ = false +valuate-cl v (a ⇒ b) = implication-bool (valuate-cl v a) (valuate-cl v b) +valuate-cl v (□ _) = false + +tautology : formula → UU lzero +tautology a = ∀ v → type-prop-bool (valuate-cl v a) +``` + +**Examples** + +```agda +module tautology-example where + A = var 0 + B = var 1 + C = var 2 + + ex1 : tautology ⊤ + ex1 v = star + + ex2 : ¬ (tautology ⊥) + ex2 h = h (λ _ → false) + + ex3 : tautology (A ⇒ A) + ex3 v with v 0 + ... | true = star + ... | false = star + + ex4 : tautology ((A ⇒ B) ∨ (B ⇒ A)) + ex4 v with v 0 | v 1 + ... | true | true = star + ... | true | false = star + ... | false | true = star + ... | false | false = star + + ex5 : tautology ((A ⇒ B) ⇒ ((B ⇒ C) ⇒ (A ⇒ C))) + ex5 v with v 0 | v 1 | v 2 + ... | true | true | true = star + ... | true | true | false = star + ... | true | false | true = star + ... | true | false | false = star + ... | false | true | true = star + ... | false | true | false = star + ... | false | false | true = star + ... | false | false | false = star + + ex6 : ¬ (tautology A) + ex6 h = h (λ _ → false) +``` + +**Tautology is decidable** + +```agda +is-decidable-tautology : (a : formula) → is-decidable (tautology a) +is-decidable-tautology (var x) = inr (λ h → h (λ _ → false)) +is-decidable-tautology ⊥ = inr (λ h → h (λ _ → false)) +is-decidable-tautology (a ⇒ b) = {! !} +is-decidable-tautology (□ a) = inr (λ h → h (λ _ → false)) +``` + +**Semantics** + +```agda +data ⊢_ : formula → UU lzero where + ax : {a : formula} → tautology a → ⊢ a + mp : {a b : formula} → ⊢ (a ⇒ b) → ⊢ a → ⊢ b + nec : {a : formula} → ⊢ (□ a) → ⊢ a + +record kripke-frame {l : Level} (w : UU l) : UU (lsuc l) where + constructor frame + field + R : Relation l w + +open kripke-frame public + +module sets where + module _ + {l1 : Level} + (l2 : Level) + (w : UU l1) + where + + private + l = lsuc (l1 ⊔ lsuc l2) + W = powerset l2 w + + record kripke-model : UU l where + constructor model + field + f : kripke-frame w + V : ℕ → W + open kripke-model public + + module _ + {l1 l2 : Level} + {w : UU l1} + where + + private + l = lsuc (l1 ⊔ lsuc l2) + + data _,_⊨_ (M : kripke-model l2 w) (x : w) : formula → UU l where + ⊨-var : {n : ℕ} → is-in-subtype (V M n) x → M , x ⊨ (var n) + ⊨-left̄⇒ : {a b : formula} → M , x ⊨ (~ a) → M , x ⊨ (a ⇒ b) -- strict positivity + ⊨-right⇒ : {a b : formula} → M , x ⊨ b → M , x ⊨ (a ⇒ b) + open _,_⊨_ public + + _,_⊭_ : kripke-model l2 w → w → formula → UU l + M , x ⊭ a = ¬ (M , x ⊨ a) + + _⊨M_ : kripke-model l2 w → formula → UU l + M ⊨M a = (x : w) → M , x ⊨ a + + _⊭M_ : kripke-model l2 w → formula → UU l + M ⊭M a = ¬ (M ⊨M a) + + module _ + {l : Level} + (w : UU l) + where + + _⊨F_ : kripke-frame w → formula → UUω + F ⊨F a = (l2 : Level) (M : kripke-model l2 w) → (M ⊨M a) + + -- _⊭F_ : kripke-frame w → formula → UUω + -- F ⊭F a = ¬ (F ⊨F a) -- Why I cannot write this? + + +module functions where + module _ + {l : Level} + (w : UU l) + where + + private + -- bool or type? + + -- in-subset : {l : Level} → UU l → UU l + -- in-subset A = A → bool + + in-subset : {l : Level} → UU l → UU (l ⊔ lsuc lzero) + in-subset A = A → UU lzero + + record kripke-model : UU (lsuc l) where + constructor model + field + F : kripke-frame w + V : ℕ → in-subset w + open kripke-model public + + data _,_⊨_ (M : kripke-model) (x : w) : formula → UU l where + ⊨-var : {n : ℕ} → V M n x → M , x ⊨ (var n) + ⊨-left̄⇒ : {a b : formula} → M , x ⊨ (~ a) → M , x ⊨ (a ⇒ b) -- strict positivity + ⊨-right⇒ : {a b : formula} → M , x ⊨ b → M , x ⊨ (a ⇒ b) + open _,_⊨_ public + + _,_⊭_ : kripke-model → w → formula → UU l + M , x ⊭ a = ¬ (M , x ⊨ a) + + _⊨M_ : kripke-model → formula → UU l + M ⊨M a = ∀ x → M , x ⊨ a + + _⊭M_ : kripke-model → formula → UU l + M ⊭M a = ¬ (M ⊨M a) + + _⊨F_ : kripke-frame w → formula → UU (l ⊔ lsuc lzero) + F ⊨F a = ∀ V → model F V ⊨M a + + _⊭F_ : kripke-frame w → formula → UU (l ⊔ lsuc lzero) + F ⊭F a = ¬ (F ⊨F a) + + _⊨C_ : in-subset (kripke-frame w) → formula → UU (lsuc l) + C ⊨C a = ∀ F → C F → F ⊨F a + + _⊭C_ : in-subset (kripke-frame w) → formula → UU (lsuc l) + C ⊭C a = ¬ (C ⊨C a) +``` From 63b0e15e4130b6f6f271d38fb5a69bfd0f2eaa94 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Sun, 8 Oct 2023 14:04:51 +0300 Subject: [PATCH 02/63] add soundness --- src/modal-logic/modal-logic-syntax.lagda.md | 353 +++++++++++++------- 1 file changed, 233 insertions(+), 120 deletions(-) diff --git a/src/modal-logic/modal-logic-syntax.lagda.md b/src/modal-logic/modal-logic-syntax.lagda.md index 5bc10c121f..09cd06e6d7 100644 --- a/src/modal-logic/modal-logic-syntax.lagda.md +++ b/src/modal-logic/modal-logic-syntax.lagda.md @@ -7,21 +7,23 @@ module modal-logic.modal-logic-syntax where
Imports ```agda --- open import foundation.contractible-types --- open import foundation.dependent-pair-types --- open import foundation.equivalences --- open import foundation.identity-types --- open import foundation.propositions --- open import foundation.sets +open import foundation.contractible-types +open import foundation.equivalences +open import foundation.identity-types +open import foundation.propositions +open import foundation.sets open import foundation.booleans open import foundation.binary-relations +open import foundation.cartesian-product-types open import foundation.decidable-types -open import foundation.sets +open import foundation.dependent-pair-types open import foundation.powersets -open import foundation-core.propositions open import foundation.universe-levels open import foundation.unit-type +open import foundation.action-on-identifications-functions +open import foundation.logical-equivalences +open import foundation-core.transport-along-identifications open import foundation-core.coproduct-types open import foundation-core.empty-types open import foundation-core.propositions @@ -40,6 +42,18 @@ TODO ## Definition +```agda +-- I'am not found in lib +record Lift {a : Level} (l : Level) (A : UU a) : UU (a ⊔ l) where + constructor lift + field lower : A +open Lift + +is-decidable-type-prop-bool : (b : bool) → is-decidable (type-prop-bool b) +is-decidable-type-prop-bool true = inl star +is-decidable-type-prop-bool false = inr (λ ()) +``` + **Syntax** ```agda @@ -69,13 +83,21 @@ a ∧ b = ~ (a ⇒ (~ b)) ```agda valuate-cl : (ℕ → bool) → formula → bool -valuate-cl v (var x) = v x +valuate-cl v (var n) = v n valuate-cl v ⊥ = false valuate-cl v (a ⇒ b) = implication-bool (valuate-cl v a) (valuate-cl v b) valuate-cl v (□ _) = false +data without-boxes : formula → UU lzero where + without-boxes-var : (n : ℕ) → without-boxes (var n) + without-boxes-bot : without-boxes ⊥ + without-boxes-imp : {a b : formula} → + without-boxes a → + without-boxes b → + without-boxes (a ⇒ b) + tautology : formula → UU lzero -tautology a = ∀ v → type-prop-bool (valuate-cl v a) +tautology a = without-boxes a × ∀ v → type-prop-bool (valuate-cl v a) ``` **Examples** @@ -87,46 +109,103 @@ module tautology-example where C = var 2 ex1 : tautology ⊤ - ex1 v = star + ex1 = wb , λ _ → star + where + wb : without-boxes ⊤ + wb = without-boxes-imp without-boxes-bot without-boxes-bot ex2 : ¬ (tautology ⊥) - ex2 h = h (λ _ → false) + ex2 h = pr2 h λ _ → false ex3 : tautology (A ⇒ A) - ex3 v with v 0 - ... | true = star - ... | false = star + ex3 = wb , val + where + wb : without-boxes (A ⇒ A) + wb = (without-boxes-imp (without-boxes-var 0) (without-boxes-var 0)) + val : ∀ v → type-prop-bool (valuate-cl v (A ⇒ A)) + val v with v 0 + ... | true = star + ... | false = star ex4 : tautology ((A ⇒ B) ∨ (B ⇒ A)) - ex4 v with v 0 | v 1 - ... | true | true = star - ... | true | false = star - ... | false | true = star - ... | false | false = star + ex4 = wb , val + where + wb : without-boxes ((A ⇒ B) ∨ (B ⇒ A)) + wb = without-boxes-imp + (without-boxes-imp + (without-boxes-imp (without-boxes-var 0) (without-boxes-var 1)) + without-boxes-bot) + (without-boxes-imp (without-boxes-var 1) (without-boxes-var 0)) + val : ∀ v → type-prop-bool (valuate-cl v ((A ⇒ B) ∨ (B ⇒ A))) + val v with v 0 | v 1 + ... | true | true = star + ... | true | false = star + ... | false | true = star + ... | false | false = star ex5 : tautology ((A ⇒ B) ⇒ ((B ⇒ C) ⇒ (A ⇒ C))) - ex5 v with v 0 | v 1 | v 2 - ... | true | true | true = star - ... | true | true | false = star - ... | true | false | true = star - ... | true | false | false = star - ... | false | true | true = star - ... | false | true | false = star - ... | false | false | true = star - ... | false | false | false = star + ex5 = wb , val + where + wb : without-boxes ((A ⇒ B) ⇒ ((B ⇒ C) ⇒ (A ⇒ C))) + wb = without-boxes-imp + (without-boxes-imp + (without-boxes-var 0) + (without-boxes-var 1)) + (without-boxes-imp + (without-boxes-imp + (without-boxes-var 1) + (without-boxes-var 2)) + (without-boxes-imp + (without-boxes-var 0) + (without-boxes-var 2))) + val : ∀ v → type-prop-bool (valuate-cl v ((A ⇒ B) ⇒ ((B ⇒ C) ⇒ (A ⇒ C)))) + val v with v 0 | v 1 | v 2 + ... | true | true | true = star + ... | true | true | false = star + ... | true | false | true = star + ... | true | false | false = star + ... | false | true | true = star + ... | false | true | false = star + ... | false | false | true = star + ... | false | false | false = star ex6 : ¬ (tautology A) - ex6 h = h (λ _ → false) + ex6 h = pr2 h (λ _ → false) + + ex7 : ¬ (tautology (⊥ ⇒ (□ A))) + ex7 (without-boxes-imp _ () , _) ``` **Tautology is decidable** ```agda +without-boxes-inversion : {a b : formula} → + without-boxes (a ⇒ b) → + without-boxes a × without-boxes b +without-boxes-inversion (without-boxes-imp wa wb) = wa , wb + +is-decidable-without-boxes : (a : formula) → is-decidable (without-boxes a) +is-decidable-without-boxes (var n) = inl (without-boxes-var n) +is-decidable-without-boxes ⊥ = inl without-boxes-bot +is-decidable-without-boxes (a ⇒ b) + with is-decidable-without-boxes a | is-decidable-without-boxes b +... | inr wa | _ = inr λ w → wa (pr1 (without-boxes-inversion w)) +... | inl wa | inr wb = inr λ w → wb (pr2 (without-boxes-inversion w)) +... | inl wa | inl wb = inl (without-boxes-imp wa wb) +is-decidable-without-boxes (□ a) = inr (λ ()) + +is-decidable-tautology-v : (a : formula) → is-decidable (∀ v → type-prop-bool (valuate-cl v a)) +is-decidable-tautology-v (var x) = inr λ h → h (λ _ → false) +is-decidable-tautology-v ⊥ = inr λ h → h (λ _ → false) +is-decidable-tautology-v (a ⇒ b) = {! !} +is-decidable-tautology-v (□ a) = inr λ h → h (λ _ → false) + is-decidable-tautology : (a : formula) → is-decidable (tautology a) -is-decidable-tautology (var x) = inr (λ h → h (λ _ → false)) -is-decidable-tautology ⊥ = inr (λ h → h (λ _ → false)) -is-decidable-tautology (a ⇒ b) = {! !} -is-decidable-tautology (□ a) = inr (λ h → h (λ _ → false)) +is-decidable-tautology a + with is-decidable-without-boxes a | is-decidable-tautology-v a +... | inr wa | _ = inr (λ x → wa (pr1 x)) +... | inl wa | inr t = inr (λ x → t (pr2 x)) +... | inl wa | inl t = inl (wa , t) ``` **Semantics** @@ -134,8 +213,9 @@ is-decidable-tautology (□ a) = inr (λ h → h (λ _ → false)) ```agda data ⊢_ : formula → UU lzero where ax : {a : formula} → tautology a → ⊢ a + k : {a b : formula} → ⊢ ((□ (a ⇒ b)) ⇒ ((□ a) ⇒ (□ b))) mp : {a b : formula} → ⊢ (a ⇒ b) → ⊢ a → ⊢ b - nec : {a : formula} → ⊢ (□ a) → ⊢ a + nec : {a : formula} → ⊢ a → ⊢ (□ a) record kripke-frame {l : Level} (w : UU l) : UU (lsuc l) where constructor frame @@ -144,105 +224,138 @@ record kripke-frame {l : Level} (w : UU l) : UU (lsuc l) where open kripke-frame public -module sets where - module _ - {l1 : Level} - (l2 : Level) - (w : UU l1) - where +private +-- bool or type? - private - l = lsuc (l1 ⊔ lsuc l2) - W = powerset l2 w + in-subset : {l : Level} → UU l → UU l + in-subset A = A → bool - record kripke-model : UU l where - constructor model - field - f : kripke-frame w - V : ℕ → W - open kripke-model public + -- in-subset : {l : Level} → UU l → UU (l ⊔ lsuc lzero) + -- in-subset A = A → UU lzero - module _ - {l1 l2 : Level} - {w : UU l1} - where +module _ + {l : Level} + (w : UU l) + -- (w : UU lzero) + where - private - l = lsuc (l1 ⊔ lsuc l2) + record kripke-model : UU (lsuc l) where + -- record kripke-model : UU₁ where + constructor model + field + F : kripke-frame w + V : ℕ → in-subset w + open kripke-model public - data _,_⊨_ (M : kripke-model l2 w) (x : w) : formula → UU l where - ⊨-var : {n : ℕ} → is-in-subtype (V M n) x → M , x ⊨ (var n) - ⊨-left̄⇒ : {a b : formula} → M , x ⊨ (~ a) → M , x ⊨ (a ⇒ b) -- strict positivity - ⊨-right⇒ : {a b : formula} → M , x ⊨ b → M , x ⊨ (a ⇒ b) - open _,_⊨_ public + _,_⊨_ : kripke-model → w → formula → UU l + M , x ⊨ var n = Lift _ (type-prop-bool (V M n x)) + M , x ⊨ ⊥ = Lift _ empty + M , x ⊨ (a ⇒ b) = ¬ (M , x ⊨ a) + M , x ⊨ b + M , x ⊨ (□ a) = ∀ y → R (F M) x y → M , y ⊨ a - _,_⊭_ : kripke-model l2 w → w → formula → UU l - M , x ⊭ a = ¬ (M , x ⊨ a) + _,_⊭_ : kripke-model → w → formula → UU l + M , x ⊭ a = ¬ (M , x ⊨ a) - _⊨M_ : kripke-model l2 w → formula → UU l - M ⊨M a = (x : w) → M , x ⊨ a + _⊨M_ : kripke-model → formula → UU l + M ⊨M a = ∀ x → M , x ⊨ a - _⊭M_ : kripke-model l2 w → formula → UU l - M ⊭M a = ¬ (M ⊨M a) + _⊭M_ : kripke-model → formula → UU l + M ⊭M a = ¬ (M ⊨M a) - module _ - {l : Level} - (w : UU l) - where + _⊨F_ : kripke-frame w → formula → UU l + F ⊨F a = ∀ V → model F V ⊨M a - _⊨F_ : kripke-frame w → formula → UUω - F ⊨F a = (l2 : Level) (M : kripke-model l2 w) → (M ⊨M a) + _⊭F_ : kripke-frame w → formula → UU l + F ⊭F a = ¬ (F ⊨F a) - -- _⊭F_ : kripke-frame w → formula → UUω - -- F ⊭F a = ¬ (F ⊨F a) -- Why I cannot write this? + _⊨C_ : in-subset (kripke-frame w) → formula → UU (lsuc l) + C ⊨C a = ∀ F → type-prop-bool (C F) → F ⊨F a + _⊭C_ : in-subset (kripke-frame w) → formula → UU (lsuc l) + C ⊭C a = ¬ (C ⊨C a) +``` -module functions where - module _ - {l : Level} - (w : UU l) - where - - private - -- bool or type? - - -- in-subset : {l : Level} → UU l → UU l - -- in-subset A = A → bool - - in-subset : {l : Level} → UU l → UU (l ⊔ lsuc lzero) - in-subset A = A → UU lzero - - record kripke-model : UU (lsuc l) where - constructor model - field - F : kripke-frame w - V : ℕ → in-subset w - open kripke-model public - - data _,_⊨_ (M : kripke-model) (x : w) : formula → UU l where - ⊨-var : {n : ℕ} → V M n x → M , x ⊨ (var n) - ⊨-left̄⇒ : {a b : formula} → M , x ⊨ (~ a) → M , x ⊨ (a ⇒ b) -- strict positivity - ⊨-right⇒ : {a b : formula} → M , x ⊨ b → M , x ⊨ (a ⇒ b) - open _,_⊨_ public - - _,_⊭_ : kripke-model → w → formula → UU l - M , x ⊭ a = ¬ (M , x ⊨ a) - - _⊨M_ : kripke-model → formula → UU l - M ⊨M a = ∀ x → M , x ⊨ a - - _⊭M_ : kripke-model → formula → UU l - M ⊭M a = ¬ (M ⊨M a) - - _⊨F_ : kripke-frame w → formula → UU (l ⊔ lsuc lzero) - F ⊨F a = ∀ V → model F V ⊨M a +**Force without boxes** - _⊭F_ : kripke-frame w → formula → UU (l ⊔ lsuc lzero) - F ⊭F a = ¬ (F ⊨F a) +```agda + force-without-boxes : {a : formula} + → without-boxes a + → {M : kripke-model} + → (x : w) + → type-prop-bool (valuate-cl (λ n → V M n x) a) ↔ M , x ⊨ a + force-without-boxes (without-boxes-var n) x = (λ val → lift val) , λ f → lower f + force-without-boxes without-boxes-bot x = (λ ()) , (λ ()) + force-without-boxes ab@{a ⇒ b} (without-boxes-imp wa wb) {M} x = prove→ , prove← + where + v = λ n → V M n x + + prove→ : type-prop-bool (valuate-cl v ab) → M , x ⊨ ab + prove→ val with valuate-cl v a in eqa | valuate-cl v b in eqb + ... | _ | true = inr (pr1 (force-without-boxes wb x) (tr type-prop-bool (inv eqb) _)) + ... | false | false = inl λ fa → tr type-prop-bool eqa (pr2 (force-without-boxes wa x) fa) + + prove← : M , x ⊨ ab → type-prop-bool (valuate-cl v ab) + prove← (inl fab) = {! !} + prove← (inr fab) with pr2 (force-without-boxes wb x) fab in eqt + ... | ttt = {! ttt !} + -- force-without-boxes (without-boxes-var n) x val = lift val + -- force-without-boxes {a ⇒ b} (without-boxes-imp wa wb) {M} x val with valuate-cl (λ n → V M n x) a in eqa | valuate-cl (λ n → V M n x) b in eqb + -- ... | true | true = inr (force-without-boxes wb x (tr type-prop-bool (inv eqb) val)) + -- ... | false | _ = inl {! !} -- inl (λ fa → tr type-prop-bool eqa {! force-without-boxes !}) +``` - _⊨C_ : in-subset (kripke-frame w) → formula → UU (lsuc l) - C ⊨C a = ∀ F → C F → F ⊨F a +**Soundness** - _⊭C_ : in-subset (kripke-frame w) → formula → UU (lsuc l) - C ⊭C a = ¬ (C ⊨C a) +```agda + -- all-frames : in-subset (kripke-frame w) + -- all-frames = λ _ → true + + is-classical : kripke-model → UU l + is-classical M = ∀ a x → is-decidable (M , x ⊨ a) + + imp-force : {a b : formula} {M : kripke-model} {x : w} + → M , x ⊨ (a ⇒ b) + → M , x ⊨ a + → M , x ⊨ b + imp-force (inl fab) fa = ex-falso (fab fa) + imp-force (inr fab) fa = fab + + tautology-canonical : {a : formula} + → tautology a + → (M : kripke-model) + → M ⊨M a + tautology-canonical {var n} (_ , t) M x = ex-falso (t (λ _ → false)) + tautology-canonical {⊥} (_ , t) M x = ex-falso (t (λ _ → false)) + tautology-canonical {a ⇒ b} (without-boxes-imp wa wb , t) M x = + {! !} + tautology-canonical {□ a} () M x + + soundness : {a : formula} + → ⊢ a + → (M : kripke-model) + → is-classical M + → M ⊨M a + soundness (mp dab da) M dec x with soundness dab M dec x + ... | inl f = ex-falso (f (soundness da M dec x)) + ... | inr f = f + soundness (ax t) M dec = {! !} + soundness {ab ⇒ (a ⇒ b)} k M dec x with dec ab x + ... | inr fab = inl fab + ... | inl fab with dec a x + ... | inr fa = inr (inl fa) + ... | inl fa = inr (inr (λ y r → imp-force (fab y r) (fa y r))) + soundness (nec d) M dec x y _ = soundness d M dec y + + -- soundness : {a : formula} → ⊢ a → all-frames ⊨C a + -- soundness {a} (mp dab da) F _ V x with soundness dab F star V x + -- ... | inl d = ex-falso (d (soundness da F star V x)) + -- ... | inr d = d + -- soundness {a} (ax t) = tautology-canonical t + -- -- soundness {var n} (ax (_ , t)) = ex-falso (t λ _ → false) + -- -- soundness {⊥} (ax (_ , t)) = ex-falso (t λ _ → false) + -- -- soundness {a ⇒ b} (ax (without-boxes-imp wa wb , t)) F _ V x + -- -- = {! !} + -- soundness {(□ (a ⇒ b)) ⇒ ((□ a) ⇒ (□ b))} k F _ V + -- = {! !} + -- soundness {□ a} (nec d) F _ V x y _ = soundness d F star V y ``` From 5f89c974ec2bc1ea398e85ac7ca1adb935ba70e2 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Sun, 8 Oct 2023 17:46:05 +0300 Subject: [PATCH 03/63] - Change Hilbert system - Add index for variables in formulas - Add soundness for decidable valuation --- src/modal-logic/modal-logic-syntax.lagda.md | 398 +++++++------------- 1 file changed, 130 insertions(+), 268 deletions(-) diff --git a/src/modal-logic/modal-logic-syntax.lagda.md b/src/modal-logic/modal-logic-syntax.lagda.md index 09cd06e6d7..5ccf220c4c 100644 --- a/src/modal-logic/modal-logic-syntax.lagda.md +++ b/src/modal-logic/modal-logic-syntax.lagda.md @@ -7,6 +7,7 @@ module modal-logic.modal-logic-syntax where
Imports ```agda +open import foundation.double-negation open import foundation.contractible-types open import foundation.equivalences open import foundation.identity-types @@ -52,310 +53,171 @@ open Lift is-decidable-type-prop-bool : (b : bool) → is-decidable (type-prop-bool b) is-decidable-type-prop-bool true = inl star is-decidable-type-prop-bool false = inr (λ ()) -``` - -**Syntax** - -```agda -data formula : UU lzero where - var : ℕ → formula - ⊥ : formula - _⇒_ : formula → formula → formula - □_ : formula → formula -~_ : formula → formula -~ a = a ⇒ ⊥ +LEM : (l : Level) → UU (lsuc l) +LEM l = (T : UU l) → is-decidable T -_∨_ : formula → formula → formula -a ∨ b = (~ a) ⇒ b - -_∧_ : formula → formula → formula -a ∧ b = ~ (a ⇒ (~ b)) - -⊤ : formula -⊤ = ~ ⊥ +-- bool or type? +in-subset : {l : Level} → UU l → UU l +in-subset A = A → bool -◇_ : formula → formula -◇ a = ~ (□ (~ a)) +-- in-subset : {l : Level} → UU l → UU (l ⊔ lsuc lzero) +-- in-subset A = A → UU lzero ``` -**Tautology** +**Syntax** ```agda -valuate-cl : (ℕ → bool) → formula → bool -valuate-cl v (var n) = v n -valuate-cl v ⊥ = false -valuate-cl v (a ⇒ b) = implication-bool (valuate-cl v a) (valuate-cl v b) -valuate-cl v (□ _) = false - -data without-boxes : formula → UU lzero where - without-boxes-var : (n : ℕ) → without-boxes (var n) - without-boxes-bot : without-boxes ⊥ - without-boxes-imp : {a b : formula} → - without-boxes a → - without-boxes b → - without-boxes (a ⇒ b) - -tautology : formula → UU lzero -tautology a = without-boxes a × ∀ v → type-prop-bool (valuate-cl v a) -``` +module _ + {l : Level} + (i : UU l) + where -**Examples** + data formula : UU l where + var : i → formula + ⊥ : formula + _⇒_ : formula → formula → formula + □ : formula → formula -```agda -module tautology-example where - A = var 0 - B = var 1 - C = var 2 +module _ + {l : Level} + {i : UU l} + where - ex1 : tautology ⊤ - ex1 = wb , λ _ → star - where - wb : without-boxes ⊤ - wb = without-boxes-imp without-boxes-bot without-boxes-bot + ~ : formula i → formula i + ~ a = a ⇒ ⊥ - ex2 : ¬ (tautology ⊥) - ex2 h = pr2 h λ _ → false + _∨_ : formula i → formula i → formula i + a ∨ b = (~ a) ⇒ b - ex3 : tautology (A ⇒ A) - ex3 = wb , val - where - wb : without-boxes (A ⇒ A) - wb = (without-boxes-imp (without-boxes-var 0) (without-boxes-var 0)) - val : ∀ v → type-prop-bool (valuate-cl v (A ⇒ A)) - val v with v 0 - ... | true = star - ... | false = star - - ex4 : tautology ((A ⇒ B) ∨ (B ⇒ A)) - ex4 = wb , val - where - wb : without-boxes ((A ⇒ B) ∨ (B ⇒ A)) - wb = without-boxes-imp - (without-boxes-imp - (without-boxes-imp (without-boxes-var 0) (without-boxes-var 1)) - without-boxes-bot) - (without-boxes-imp (without-boxes-var 1) (without-boxes-var 0)) - val : ∀ v → type-prop-bool (valuate-cl v ((A ⇒ B) ∨ (B ⇒ A))) - val v with v 0 | v 1 - ... | true | true = star - ... | true | false = star - ... | false | true = star - ... | false | false = star - - ex5 : tautology ((A ⇒ B) ⇒ ((B ⇒ C) ⇒ (A ⇒ C))) - ex5 = wb , val - where - wb : without-boxes ((A ⇒ B) ⇒ ((B ⇒ C) ⇒ (A ⇒ C))) - wb = without-boxes-imp - (without-boxes-imp - (without-boxes-var 0) - (without-boxes-var 1)) - (without-boxes-imp - (without-boxes-imp - (without-boxes-var 1) - (without-boxes-var 2)) - (without-boxes-imp - (without-boxes-var 0) - (without-boxes-var 2))) - val : ∀ v → type-prop-bool (valuate-cl v ((A ⇒ B) ⇒ ((B ⇒ C) ⇒ (A ⇒ C)))) - val v with v 0 | v 1 | v 2 - ... | true | true | true = star - ... | true | true | false = star - ... | true | false | true = star - ... | true | false | false = star - ... | false | true | true = star - ... | false | true | false = star - ... | false | false | true = star - ... | false | false | false = star - - ex6 : ¬ (tautology A) - ex6 h = pr2 h (λ _ → false) - - ex7 : ¬ (tautology (⊥ ⇒ (□ A))) - ex7 (without-boxes-imp _ () , _) -``` + _∧_ : formula i → formula i → formula i + a ∧ b = ~ (a ⇒ (~ b)) -**Tautology is decidable** + ⊤ : formula i + ⊤ = ~ ⊥ -```agda -without-boxes-inversion : {a b : formula} → - without-boxes (a ⇒ b) → - without-boxes a × without-boxes b -without-boxes-inversion (without-boxes-imp wa wb) = wa , wb - -is-decidable-without-boxes : (a : formula) → is-decidable (without-boxes a) -is-decidable-without-boxes (var n) = inl (without-boxes-var n) -is-decidable-without-boxes ⊥ = inl without-boxes-bot -is-decidable-without-boxes (a ⇒ b) - with is-decidable-without-boxes a | is-decidable-without-boxes b -... | inr wa | _ = inr λ w → wa (pr1 (without-boxes-inversion w)) -... | inl wa | inr wb = inr λ w → wb (pr2 (without-boxes-inversion w)) -... | inl wa | inl wb = inl (without-boxes-imp wa wb) -is-decidable-without-boxes (□ a) = inr (λ ()) - -is-decidable-tautology-v : (a : formula) → is-decidable (∀ v → type-prop-bool (valuate-cl v a)) -is-decidable-tautology-v (var x) = inr λ h → h (λ _ → false) -is-decidable-tautology-v ⊥ = inr λ h → h (λ _ → false) -is-decidable-tautology-v (a ⇒ b) = {! !} -is-decidable-tautology-v (□ a) = inr λ h → h (λ _ → false) - -is-decidable-tautology : (a : formula) → is-decidable (tautology a) -is-decidable-tautology a - with is-decidable-without-boxes a | is-decidable-tautology-v a -... | inr wa | _ = inr (λ x → wa (pr1 x)) -... | inl wa | inr t = inr (λ x → t (pr2 x)) -... | inl wa | inl t = inl (wa , t) + ◇ : formula i → formula i + ◇ a = ~ (□ (~ a)) + + data ⊢ : formula i → UU l where + ax-k : {a b : formula i} → ⊢ (a ⇒ (b ⇒ a)) + ax-s : {a b c : formula i} → ⊢ ((a ⇒ (b ⇒ c)) ⇒ ((a ⇒ b) ⇒ (a ⇒ c))) + ax-dn : {a : formula i} → ⊢ ((~ (~ a)) ⇒ a) + ax-n : {a b : formula i} → ⊢ ((□ (a ⇒ b)) ⇒ ((□ a) ⇒ (□ b))) + mp : {a b : formula i} → ⊢ (a ⇒ b) → ⊢ a → ⊢ b + nec : {a : formula i} → ⊢ a → ⊢ (□ a) ``` **Semantics** ```agda -data ⊢_ : formula → UU lzero where - ax : {a : formula} → tautology a → ⊢ a - k : {a b : formula} → ⊢ ((□ (a ⇒ b)) ⇒ ((□ a) ⇒ (□ b))) - mp : {a b : formula} → ⊢ (a ⇒ b) → ⊢ a → ⊢ b - nec : {a : formula} → ⊢ a → ⊢ (□ a) - -record kripke-frame {l : Level} (w : UU l) : UU (lsuc l) where - constructor frame - field - R : Relation l w - -open kripke-frame public - -private --- bool or type? - - in-subset : {l : Level} → UU l → UU l - in-subset A = A → bool - - -- in-subset : {l : Level} → UU l → UU (l ⊔ lsuc lzero) - -- in-subset A = A → UU lzero - module _ - {l : Level} - (w : UU l) - -- (w : UU lzero) + {l1 : Level} + (w : UU l1) where - record kripke-model : UU (lsuc l) where - -- record kripke-model : UU₁ where - constructor model + record kripke-frame : UU (lsuc l1) where + constructor frame field - F : kripke-frame w - V : ℕ → in-subset w - open kripke-model public + R : Relation l1 w - _,_⊨_ : kripke-model → w → formula → UU l - M , x ⊨ var n = Lift _ (type-prop-bool (V M n x)) - M , x ⊨ ⊥ = Lift _ empty - M , x ⊨ (a ⇒ b) = ¬ (M , x ⊨ a) + M , x ⊨ b - M , x ⊨ (□ a) = ∀ y → R (F M) x y → M , y ⊨ a + open kripke-frame public - _,_⊭_ : kripke-model → w → formula → UU l - M , x ⊭ a = ¬ (M , x ⊨ a) + module _ + {l2 : Level} + (i : UU l2) + where - _⊨M_ : kripke-model → formula → UU l - M ⊨M a = ∀ x → M , x ⊨ a + private + l = l1 ⊔ l2 - _⊭M_ : kripke-model → formula → UU l - M ⊭M a = ¬ (M ⊨M a) + record kripke-model : UU (lsuc l) where + constructor model + field + F : kripke-frame + V : i → in-subset w - _⊨F_ : kripke-frame w → formula → UU l - F ⊨F a = ∀ V → model F V ⊨M a + open kripke-model public - _⊭F_ : kripke-frame w → formula → UU l - F ⊭F a = ¬ (F ⊨F a) + _,_⊨_ : kripke-model → w → formula i → UU l + M , x ⊨ var n = Lift _ (type-prop-bool (V M n x)) + M , x ⊨ ⊥ = Lift _ empty + M , x ⊨ (a ⇒ b) = ¬ (M , x ⊨ a) + M , x ⊨ b + M , x ⊨ (□ a) = ∀ y → R (F M) x y → M , y ⊨ a - _⊨C_ : in-subset (kripke-frame w) → formula → UU (lsuc l) - C ⊨C a = ∀ F → type-prop-bool (C F) → F ⊨F a + _,_⊭_ : kripke-model → w → formula i → UU l + M , x ⊭ a = ¬ (M , x ⊨ a) - _⊭C_ : in-subset (kripke-frame w) → formula → UU (lsuc l) - C ⊭C a = ¬ (C ⊨C a) -``` + _⊨M_ : kripke-model → formula i → UU l + M ⊨M a = ∀ x → M , x ⊨ a -**Force without boxes** + _⊭M_ : kripke-model → formula i → UU l + M ⊭M a = ¬ (M ⊨M a) -```agda - force-without-boxes : {a : formula} - → without-boxes a - → {M : kripke-model} - → (x : w) - → type-prop-bool (valuate-cl (λ n → V M n x) a) ↔ M , x ⊨ a - force-without-boxes (without-boxes-var n) x = (λ val → lift val) , λ f → lower f - force-without-boxes without-boxes-bot x = (λ ()) , (λ ()) - force-without-boxes ab@{a ⇒ b} (without-boxes-imp wa wb) {M} x = prove→ , prove← - where - v = λ n → V M n x - - prove→ : type-prop-bool (valuate-cl v ab) → M , x ⊨ ab - prove→ val with valuate-cl v a in eqa | valuate-cl v b in eqb - ... | _ | true = inr (pr1 (force-without-boxes wb x) (tr type-prop-bool (inv eqb) _)) - ... | false | false = inl λ fa → tr type-prop-bool eqa (pr2 (force-without-boxes wa x) fa) - - prove← : M , x ⊨ ab → type-prop-bool (valuate-cl v ab) - prove← (inl fab) = {! !} - prove← (inr fab) with pr2 (force-without-boxes wb x) fab in eqt - ... | ttt = {! ttt !} - -- force-without-boxes (without-boxes-var n) x val = lift val - -- force-without-boxes {a ⇒ b} (without-boxes-imp wa wb) {M} x val with valuate-cl (λ n → V M n x) a in eqa | valuate-cl (λ n → V M n x) b in eqb - -- ... | true | true = inr (force-without-boxes wb x (tr type-prop-bool (inv eqb) val)) - -- ... | false | _ = inl {! !} -- inl (λ fa → tr type-prop-bool eqa {! force-without-boxes !}) + _⊨F_ : kripke-frame → formula i → UU l + F ⊨F a = ∀ V → model F V ⊨M a + + _⊭F_ : kripke-frame → formula i → UU l + F ⊭F a = ¬ (F ⊨F a) + + _⊨C_ : in-subset kripke-frame → formula i → UU (lsuc l1 ⊔ l2) + C ⊨C a = ∀ F → type-prop-bool (C F) → F ⊨F a + + _⊭C_ : in-subset kripke-frame → formula i → UU (lsuc l1 ⊔ l2) + C ⊭C a = ¬ (C ⊨C a) ``` **Soundness** ```agda - -- all-frames : in-subset (kripke-frame w) - -- all-frames = λ _ → true - - is-classical : kripke-model → UU l - is-classical M = ∀ a x → is-decidable (M , x ⊨ a) - - imp-force : {a b : formula} {M : kripke-model} {x : w} - → M , x ⊨ (a ⇒ b) - → M , x ⊨ a - → M , x ⊨ b - imp-force (inl fab) fa = ex-falso (fab fa) - imp-force (inr fab) fa = fab - - tautology-canonical : {a : formula} - → tautology a - → (M : kripke-model) - → M ⊨M a - tautology-canonical {var n} (_ , t) M x = ex-falso (t (λ _ → false)) - tautology-canonical {⊥} (_ , t) M x = ex-falso (t (λ _ → false)) - tautology-canonical {a ⇒ b} (without-boxes-imp wa wb , t) M x = - {! !} - tautology-canonical {□ a} () M x - - soundness : {a : formula} - → ⊢ a - → (M : kripke-model) - → is-classical M - → M ⊨M a - soundness (mp dab da) M dec x with soundness dab M dec x - ... | inl f = ex-falso (f (soundness da M dec x)) - ... | inr f = f - soundness (ax t) M dec = {! !} - soundness {ab ⇒ (a ⇒ b)} k M dec x with dec ab x - ... | inr fab = inl fab - ... | inl fab with dec a x - ... | inr fa = inr (inl fa) - ... | inl fa = inr (inr (λ y r → imp-force (fab y r) (fa y r))) - soundness (nec d) M dec x y _ = soundness d M dec y - - -- soundness : {a : formula} → ⊢ a → all-frames ⊨C a - -- soundness {a} (mp dab da) F _ V x with soundness dab F star V x - -- ... | inl d = ex-falso (d (soundness da F star V x)) - -- ... | inr d = d - -- soundness {a} (ax t) = tautology-canonical t - -- -- soundness {var n} (ax (_ , t)) = ex-falso (t λ _ → false) - -- -- soundness {⊥} (ax (_ , t)) = ex-falso (t λ _ → false) - -- -- soundness {a ⇒ b} (ax (without-boxes-imp wa wb , t)) F _ V x - -- -- = {! !} - -- soundness {(□ (a ⇒ b)) ⇒ ((□ a) ⇒ (□ b))} k F _ V - -- = {! !} - -- soundness {□ a} (nec d) F _ V x y _ = soundness d F star V y + all-frames : in-subset kripke-frame + all-frames = λ _ → true + + is-classical : kripke-model → UU l + is-classical M = ∀ a x → is-decidable (M , x ⊨ a) + + imp-force : {a b : formula i} {M : kripke-model} {x : w} + → M , x ⊨ (a ⇒ b) + → M , x ⊨ a + → M , x ⊨ b + imp-force (inl fab) fa = ex-falso (fab fa) + imp-force (inr fab) fa = fab + + soundness : {a : formula i} + → ⊢ a + → (M : kripke-model) + → is-classical M + → M ⊨M a + soundness (mp dab da) M dec x with soundness dab M dec x + ... | inl f = ex-falso (f (soundness da M dec x)) + ... | inr f = f + soundness {a ⇒ (b ⇒ a)} ax-k M dec x with dec a x + ... | inl fna = inr (inr fna) + ... | inr fa = inl fa + soundness {abc ⇒ (ab ⇒ ac)} ax-s M dec x with dec abc x + ... | inr fabc = inl fabc + ... | inl (inl fna) = inr (inr (inl fna)) + ... | inl (inr (inr fc)) = inr (inr (inr fc)) + ... | inl (inr (inl fnb)) with dec ab x + ... | inr fab = inr (inl fab) + ... | inl (inl fna) = inr (inr (inl fna)) + ... | inl (inr fb) = ex-falso (fnb fb) + soundness {((a ⇒ ⊥) ⇒ ⊥) ⇒ a} ax-dn M dec x with dec a x + ... | inl fa = inr fa + ... | inr fna with dec (~ (~ a)) x + ... | inl (inl fnna) = ex-falso (fnna (inl fna)) + ... | inr fnnna = inl (fnnna) + soundness {ab ⇒ (a ⇒ b)} ax-n M dec x with dec ab x + ... | inr fab = inl fab + ... | inl fab with dec a x + ... | inr fa = inr (inl fa) + ... | inl fa = inr (inr (λ y r → imp-force (fab y r) (fa y r))) + soundness (nec d) M dec x y _ = soundness d M dec y + + soundness-LEM : {a : formula i} + → LEM l + → ⊢ a + → (M : kripke-model) + → M ⊨M a + soundness-LEM lem d M = soundness d M (λ b x → lem _) ``` From 49f2cf6abb84b007c3dfa4b7b0d527f396750214 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Sun, 8 Oct 2023 18:40:15 +0300 Subject: [PATCH 04/63] Kripke fram should be not empty --- src/modal-logic/modal-logic-syntax.lagda.md | 2 ++ 1 file changed, 2 insertions(+) diff --git a/src/modal-logic/modal-logic-syntax.lagda.md b/src/modal-logic/modal-logic-syntax.lagda.md index 5ccf220c4c..0ab77a7da3 100644 --- a/src/modal-logic/modal-logic-syntax.lagda.md +++ b/src/modal-logic/modal-logic-syntax.lagda.md @@ -23,6 +23,7 @@ open import foundation.universe-levels open import foundation.unit-type open import foundation.action-on-identifications-functions open import foundation.logical-equivalences +open import foundation.inhabited-types open import foundation-core.transport-along-identifications open import foundation-core.coproduct-types @@ -120,6 +121,7 @@ module _ constructor frame field R : Relation l1 w + frame-inhabited : is-inhabited w open kripke-frame public From 517fc9179ebb2263a53e97565be237db596a98e4 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Mon, 9 Oct 2023 01:10:02 +0300 Subject: [PATCH 05/63] Refactor soundness --- src/modal-logic/modal-logic-syntax.lagda.md | 208 +++++++++----------- 1 file changed, 90 insertions(+), 118 deletions(-) diff --git a/src/modal-logic/modal-logic-syntax.lagda.md b/src/modal-logic/modal-logic-syntax.lagda.md index 0ab77a7da3..887a9ec23a 100644 --- a/src/modal-logic/modal-logic-syntax.lagda.md +++ b/src/modal-logic/modal-logic-syntax.lagda.md @@ -7,33 +7,21 @@ module modal-logic.modal-logic-syntax where
Imports ```agda -open import foundation.double-negation -open import foundation.contractible-types open import foundation.equivalences open import foundation.identity-types open import foundation.propositions -open import foundation.sets -open import foundation.booleans open import foundation.binary-relations -open import foundation.cartesian-product-types open import foundation.decidable-types open import foundation.dependent-pair-types -open import foundation.powersets open import foundation.universe-levels -open import foundation.unit-type -open import foundation.action-on-identifications-functions -open import foundation.logical-equivalences open import foundation.inhabited-types - -open import foundation-core.transport-along-identifications -open import foundation-core.coproduct-types -open import foundation-core.empty-types -open import foundation-core.propositions -open import foundation-core.subtypes -open import foundation-core.negation -open import foundation-core.identity-types - -open import elementary-number-theory.natural-numbers +open import foundation.coproduct-types +open import foundation.function-types +open import foundation.empty-types +open import foundation.unit-type +open import foundation.negation +open import foundation.double-negation +open import foundation.propositional-truncations ```
@@ -46,24 +34,24 @@ TODO ```agda -- I'am not found in lib -record Lift {a : Level} (l : Level) (A : UU a) : UU (a ⊔ l) where - constructor lift - field lower : A -open Lift +module _ + {a : Level} (l : Level) + where + + record Lift (A : UU a) : UU (a ⊔ l) where + constructor lift + field lower : A + open Lift public -is-decidable-type-prop-bool : (b : bool) → is-decidable (type-prop-bool b) -is-decidable-type-prop-bool true = inl star -is-decidable-type-prop-bool false = inr (λ ()) + Lift-Prop : Prop a → Prop (a ⊔ l) + Lift-Prop A = + Lift (type-Prop A) , + is-prop-equiv' + (lift , is-equiv-is-invertible lower (λ _ → refl) (λ _ → refl)) + (is-prop-type-Prop A) LEM : (l : Level) → UU (lsuc l) LEM l = (T : UU l) → is-decidable T - --- bool or type? -in-subset : {l : Level} → UU l → UU l -in-subset A = A → bool - --- in-subset : {l : Level} → UU l → UU (l ⊔ lsuc lzero) --- in-subset A = A → UU lzero ``` **Syntax** @@ -122,104 +110,88 @@ module _ field R : Relation l1 w frame-inhabited : is-inhabited w - open kripke-frame public - module _ - {l2 : Level} - (i : UU l2) - where - - private - l = l1 ⊔ l2 - - record kripke-model : UU (lsuc l) where - constructor model - field - F : kripke-frame - V : i → in-subset w - - open kripke-model public +module _ + {l1 l2 : Level} + (w : UU l1) (i : UU l2) + (l3 : Level) + where - _,_⊨_ : kripke-model → w → formula i → UU l - M , x ⊨ var n = Lift _ (type-prop-bool (V M n x)) - M , x ⊨ ⊥ = Lift _ empty - M , x ⊨ (a ⇒ b) = ¬ (M , x ⊨ a) + M , x ⊨ b - M , x ⊨ (□ a) = ∀ y → R (F M) x y → M , y ⊨ a + record kripke-model : UU (lsuc l1 ⊔ l2 ⊔ lsuc l3) where + constructor model + field + F : kripke-frame w + V : i → w → UU l3 + open kripke-model public - _,_⊭_ : kripke-model → w → formula i → UU l - M , x ⊭ a = ¬ (M , x ⊨ a) - _⊨M_ : kripke-model → formula i → UU l - M ⊨M a = ∀ x → M , x ⊨ a +module _ + {l1 l2 l3 : Level} + {w : UU l1} {i : UU l2} + where - _⊭M_ : kripke-model → formula i → UU l - M ⊭M a = ¬ (M ⊨M a) + private + l = l1 ⊔ l2 ⊔ l3 - _⊨F_ : kripke-frame → formula i → UU l - F ⊨F a = ∀ V → model F V ⊨M a + _,_⊨_ : kripke-model w i l3 → w → formula i → UU l + M , x ⊨ var n = Lift l (V M n x) + M , x ⊨ ⊥ = Lift l empty + M , x ⊨ (a ⇒ b) = M , x ⊨ a → M , x ⊨ b + M , x ⊨ (□ a) = ∀ y → R (F M) x y → M , y ⊨ a - _⊭F_ : kripke-frame → formula i → UU l - F ⊭F a = ¬ (F ⊨F a) + _,_⊭_ : kripke-model w i l3 → w → formula i → UU l + M , x ⊭ a = ¬ (M , x ⊨ a) - _⊨C_ : in-subset kripke-frame → formula i → UU (lsuc l1 ⊔ l2) - C ⊨C a = ∀ F → type-prop-bool (C F) → F ⊨F a + _⊨M_ : kripke-model w i l3 → formula i → UU l + M ⊨M a = ∀ x → M , x ⊨ a - _⊭C_ : in-subset kripke-frame → formula i → UU (lsuc l1 ⊔ l2) - C ⊭C a = ¬ (C ⊨C a) + _⊭M_ : kripke-model w i l3 → formula i → UU l + M ⊭M a = ¬ (M ⊨M a) ``` **Soundness** ```agda - all-frames : in-subset kripke-frame - all-frames = λ _ → true - - is-classical : kripke-model → UU l - is-classical M = ∀ a x → is-decidable (M , x ⊨ a) - - imp-force : {a b : formula i} {M : kripke-model} {x : w} - → M , x ⊨ (a ⇒ b) - → M , x ⊨ a - → M , x ⊨ b - imp-force (inl fab) fa = ex-falso (fab fa) - imp-force (inr fab) fa = fab - - soundness : {a : formula i} - → ⊢ a - → (M : kripke-model) - → is-classical M - → M ⊨M a - soundness (mp dab da) M dec x with soundness dab M dec x - ... | inl f = ex-falso (f (soundness da M dec x)) - ... | inr f = f - soundness {a ⇒ (b ⇒ a)} ax-k M dec x with dec a x - ... | inl fna = inr (inr fna) - ... | inr fa = inl fa - soundness {abc ⇒ (ab ⇒ ac)} ax-s M dec x with dec abc x - ... | inr fabc = inl fabc - ... | inl (inl fna) = inr (inr (inl fna)) - ... | inl (inr (inr fc)) = inr (inr (inr fc)) - ... | inl (inr (inl fnb)) with dec ab x - ... | inr fab = inr (inl fab) - ... | inl (inl fna) = inr (inr (inl fna)) - ... | inl (inr fb) = ex-falso (fnb fb) - soundness {((a ⇒ ⊥) ⇒ ⊥) ⇒ a} ax-dn M dec x with dec a x - ... | inl fa = inr fa - ... | inr fna with dec (~ (~ a)) x - ... | inl (inl fnna) = ex-falso (fnna (inl fna)) - ... | inr fnnna = inl (fnnna) - soundness {ab ⇒ (a ⇒ b)} ax-n M dec x with dec ab x - ... | inr fab = inl fab - ... | inl fab with dec a x - ... | inr fa = inr (inl fa) - ... | inl fa = inr (inr (λ y r → imp-force (fab y r) (fa y r))) - soundness (nec d) M dec x y _ = soundness d M dec y - - soundness-LEM : {a : formula i} - → LEM l - → ⊢ a - → (M : kripke-model) - → M ⊨M a - soundness-LEM lem d M = soundness d M (λ b x → lem _) + is-classical : kripke-model w i l3 → UU l + is-classical M = ∀ a x → is-decidable (M , x ⊨ a) + + soundness : {a : formula i} + → ⊢ a + → (M : kripke-model w i l3) + → is-classical M + → M ⊨M a + soundness ax-k M dec x = λ fa _ → fa + soundness ax-s M dec x = λ fabc fab fa → fabc fa (fab fa) + soundness {_ ⇒ a} ax-dn M dec x with dec a x + ... | inl fa = λ _ → fa + ... | inr nfa = λ fnna → ex-falso (lower (fnna (λ fa → lift (nfa fa)))) + soundness ax-n M dec x = λ fab fa y r → fab y r (fa y r) + soundness (mp dab db) M dec x = soundness dab M dec x (soundness db M dec x) + soundness (nec d) M dec x y _ = soundness d M dec y + + +double-negation-LEM : {l : Level} → ((A : UU l) → ¬¬ A → A) → LEM l +double-negation-LEM dn A = dn _ λ ndec → ndec (inr (λ a → ndec (inl a))) + +lift-dn : {l1 l2 : Level} {A : UU l1} → ¬¬ A → (Lift l2 A → Lift l2 empty) → Lift l2 empty +lift-dn dn nA = lift (dn (λ a → lower (nA (lift a)))) + +required-LEM : ({l1 l2 l3 : Level} (i : UU l2) (a : formula i) + → ⊢ a + → (w : UU l1) + → (M : kripke-model w i l3) + → M ⊨M a) + → {l : Level} → LEM l +required-LEM sound {l} = + double-negation-LEM λ A nnA → lower (sound unit b ax-dn unit (M A) star (lift-dn nnA)) + where + a = var star + b = (~ (~ a)) ⇒ a + + f : kripke-frame unit + f = frame (λ _ _ → empty) (unit-trunc-Prop star) + + M : UU l → kripke-model unit unit l + M A = model f (λ _ _ → A) ``` From 60531733bd79d0ce2f7305778485b29aa74b3f2e Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Mon, 9 Oct 2023 16:46:03 +0300 Subject: [PATCH 06/63] Add finite model --- src/modal-logic/modal-logic-syntax.lagda.md | 45 ++++++++++++++++++++- 1 file changed, 43 insertions(+), 2 deletions(-) diff --git a/src/modal-logic/modal-logic-syntax.lagda.md b/src/modal-logic/modal-logic-syntax.lagda.md index 887a9ec23a..7a01334cef 100644 --- a/src/modal-logic/modal-logic-syntax.lagda.md +++ b/src/modal-logic/modal-logic-syntax.lagda.md @@ -7,6 +7,7 @@ module modal-logic.modal-logic-syntax where
Imports ```agda +open import foundation.cartesian-product-types open import foundation.equivalences open import foundation.identity-types open import foundation.propositions @@ -22,6 +23,10 @@ open import foundation.unit-type open import foundation.negation open import foundation.double-negation open import foundation.propositional-truncations + +open import univalent-combinatorics.finite-types +open import univalent-combinatorics.decidable-dependent-function-types +-- open import univalent-combinatorics.dependent-function-types ```
@@ -44,12 +49,16 @@ module _ open Lift public Lift-Prop : Prop a → Prop (a ⊔ l) - Lift-Prop A = - Lift (type-Prop A) , + pr1 (Lift-Prop A) = Lift (type-Prop A) + pr2 (Lift-Prop A) = is-prop-equiv' (lift , is-equiv-is-invertible lower (λ _ → refl) (λ _ → refl)) (is-prop-type-Prop A) + is-decidable-Lift : {A : UU a} → is-decidable A → is-decidable (Lift A) + is-decidable-Lift (inl a) = inl (lift a) + is-decidable-Lift (inr na) = inr (na ∘ lower) + LEM : (l : Level) → UU (lsuc l) LEM l = (T : UU l) → is-decidable T ``` @@ -125,6 +134,9 @@ module _ V : i → w → UU l3 open kripke-model public + finite-model : UU (lsuc l1 ⊔ l2 ⊔ lsuc l3) + finite-model = kripke-model × is-finite w + module _ {l1 l2 l3 : Level} @@ -170,6 +182,35 @@ module _ soundness (mp dab db) M dec x = soundness dab M dec x (soundness db M dec x) soundness (nec d) M dec x y _ = soundness d M dec y + finite-is-classical : ((M , _) : finite-model w i l3) + → (∀ n x → is-decidable (V M n x)) + → (∀ x y → is-decidable (R (F M) x y)) + → is-classical M + finite-is-classical (M , is-fin) dec-val dec-r (var n) x = + is-decidable-Lift _ (dec-val n x) + finite-is-classical (M , is-fin) dec-val dec-r ⊥ x = + inr (λ ()) + finite-is-classical (M , is-fin) dec-val dec-r (a ⇒ b) x + with finite-is-classical (M , is-fin) dec-val dec-r a x + ... | inr nfa = inl (λ fa → ex-falso (nfa fa)) + ... | inl fa with finite-is-classical (M , is-fin) dec-val dec-r b x + ... | inr nfb = inr (λ fab → nfb (fab fa)) + ... | inl fb = inl (λ _ → fb) + finite-is-classical (M , is-fin) dec-val dec-r (□ a) x = + is-decidable-Π-is-finite is-fin + (λ y → + is-decidable-function-type (dec-r x y) + ((finite-is-classical (M , is-fin) dec-val dec-r a y))) + + finite-soundness : ((M , _) : finite-model w i l3) + → (∀ n x → is-decidable (V M n x)) + → (∀ x y → is-decidable (R (F M) x y)) + → {a : formula i} + → ⊢ a + → M ⊨M a + finite-soundness (M , is-fin) dec-val dec-r d = + soundness d M (finite-is-classical ((M , is-fin)) dec-val dec-r) + double-negation-LEM : {l : Level} → ((A : UU l) → ¬¬ A → A) → LEM l double-negation-LEM dn A = dn _ λ ndec → ndec (inr (λ a → ndec (inl a))) From d1744d9b502dfc0df69b2e000aada86ff5ec94f9 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Mon, 9 Oct 2023 16:51:20 +0300 Subject: [PATCH 07/63] simplify finite-is-classical proof --- src/modal-logic/modal-logic-syntax.lagda.md | 10 ++++------ 1 file changed, 4 insertions(+), 6 deletions(-) diff --git a/src/modal-logic/modal-logic-syntax.lagda.md b/src/modal-logic/modal-logic-syntax.lagda.md index 7a01334cef..ac6390d986 100644 --- a/src/modal-logic/modal-logic-syntax.lagda.md +++ b/src/modal-logic/modal-logic-syntax.lagda.md @@ -190,12 +190,10 @@ module _ is-decidable-Lift _ (dec-val n x) finite-is-classical (M , is-fin) dec-val dec-r ⊥ x = inr (λ ()) - finite-is-classical (M , is-fin) dec-val dec-r (a ⇒ b) x - with finite-is-classical (M , is-fin) dec-val dec-r a x - ... | inr nfa = inl (λ fa → ex-falso (nfa fa)) - ... | inl fa with finite-is-classical (M , is-fin) dec-val dec-r b x - ... | inr nfb = inr (λ fab → nfb (fab fa)) - ... | inl fb = inl (λ _ → fb) + finite-is-classical (M , is-fin) dec-val dec-r (a ⇒ b) x = + is-decidable-function-type + (finite-is-classical (M , is-fin) dec-val dec-r a x) + (finite-is-classical (M , is-fin) dec-val dec-r b x) finite-is-classical (M , is-fin) dec-val dec-r (□ a) x = is-decidable-Π-is-finite is-fin (λ y → From 1f8f53fc8cb75b9ed2a1c67070edb06bee4953fe Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Sun, 22 Oct 2023 18:59:26 +0300 Subject: [PATCH 08/63] - Rewrite valuate to Propositions - Add priorities --- src/modal-logic/modal-logic-syntax.lagda.md | 177 ++++++++++---------- 1 file changed, 89 insertions(+), 88 deletions(-) diff --git a/src/modal-logic/modal-logic-syntax.lagda.md b/src/modal-logic/modal-logic-syntax.lagda.md index ac6390d986..3f58eb2764 100644 --- a/src/modal-logic/modal-logic-syntax.lagda.md +++ b/src/modal-logic/modal-logic-syntax.lagda.md @@ -8,7 +8,6 @@ module modal-logic.modal-logic-syntax where ```agda open import foundation.cartesian-product-types -open import foundation.equivalences open import foundation.identity-types open import foundation.propositions open import foundation.binary-relations @@ -23,10 +22,12 @@ open import foundation.unit-type open import foundation.negation open import foundation.double-negation open import foundation.propositional-truncations +open import foundation.raising-universe-levels +open import foundation.decidable-propositions +open import foundation.law-of-excluded-middle open import univalent-combinatorics.finite-types open import univalent-combinatorics.decidable-dependent-function-types --- open import univalent-combinatorics.dependent-function-types ```
@@ -37,32 +38,6 @@ TODO ## Definition -```agda --- I'am not found in lib -module _ - {a : Level} (l : Level) - where - - record Lift (A : UU a) : UU (a ⊔ l) where - constructor lift - field lower : A - open Lift public - - Lift-Prop : Prop a → Prop (a ⊔ l) - pr1 (Lift-Prop A) = Lift (type-Prop A) - pr2 (Lift-Prop A) = - is-prop-equiv' - (lift , is-equiv-is-invertible lower (λ _ → refl) (λ _ → refl)) - (is-prop-type-Prop A) - - is-decidable-Lift : {A : UU a} → is-decidable A → is-decidable (Lift A) - is-decidable-Lift (inl a) = inl (lift a) - is-decidable-Lift (inr na) = inr (na ∘ lower) - -LEM : (l : Level) → UU (lsuc l) -LEM l = (T : UU l) → is-decidable T -``` - **Syntax** ```agda @@ -71,39 +46,52 @@ module _ (i : UU l) where + infixr 4 _⇒_ + infixr 7 □_ + data formula : UU l where var : i → formula ⊥ : formula _⇒_ : formula → formula → formula - □ : formula → formula + □_ : formula → formula module _ {l : Level} {i : UU l} where - ~ : formula i → formula i + infixr 7 ~_ + infixr 7 ~~_ + infixr 5 _∨_ + infixr 6 _∧_ + infixr 7 ◇_ + infixr 2 ⊢_ + + ~_ : formula i → formula i ~ a = a ⇒ ⊥ + ~~_ : formula i → formula i + ~~ a = ~ ~ a + _∨_ : formula i → formula i → formula i - a ∨ b = (~ a) ⇒ b + a ∨ b = ~ a ⇒ b _∧_ : formula i → formula i → formula i - a ∧ b = ~ (a ⇒ (~ b)) + a ∧ b = ~ a ⇒ ~ b ⊤ : formula i ⊤ = ~ ⊥ - ◇ : formula i → formula i - ◇ a = ~ (□ (~ a)) + ◇_ : formula i → formula i + ◇ a = ~ □ ~ a - data ⊢ : formula i → UU l where - ax-k : {a b : formula i} → ⊢ (a ⇒ (b ⇒ a)) - ax-s : {a b c : formula i} → ⊢ ((a ⇒ (b ⇒ c)) ⇒ ((a ⇒ b) ⇒ (a ⇒ c))) - ax-dn : {a : formula i} → ⊢ ((~ (~ a)) ⇒ a) - ax-n : {a b : formula i} → ⊢ ((□ (a ⇒ b)) ⇒ ((□ a) ⇒ (□ b))) - mp : {a b : formula i} → ⊢ (a ⇒ b) → ⊢ a → ⊢ b - nec : {a : formula i} → ⊢ a → ⊢ (□ a) + data ⊢_ : formula i → UU l where + ax-k : {a b : formula i} → ⊢ a ⇒ b ⇒ a + ax-s : {a b c : formula i} → ⊢ (a ⇒ b ⇒ c) ⇒ (a ⇒ b) ⇒ (a ⇒ c) + ax-dn : {a : formula i} → ⊢ ~~ a ⇒ a + ax-n : {a b : formula i} → ⊢ □ (a ⇒ b) ⇒ □ a ⇒ □ b + mp : {a b : formula i} → ⊢ a ⇒ b → ⊢ a → ⊢ b + nec : {a : formula i} → ⊢ a → ⊢ □ a ``` **Semantics** @@ -131,7 +119,7 @@ module _ constructor model field F : kripke-frame w - V : i → w → UU l3 + V : i → w → Prop l3 open kripke-model public finite-model : UU (lsuc l1 ⊔ l2 ⊔ lsuc l3) @@ -146,50 +134,60 @@ module _ private l = l1 ⊔ l2 ⊔ l3 - _,_⊨_ : kripke-model w i l3 → w → formula i → UU l - M , x ⊨ var n = Lift l (V M n x) - M , x ⊨ ⊥ = Lift l empty - M , x ⊨ (a ⇒ b) = M , x ⊨ a → M , x ⊨ b - M , x ⊨ (□ a) = ∀ y → R (F M) x y → M , y ⊨ a + infix 2 _⊨_ + infix 2 _⊭_ + infix 2 _⊨M_ + infix 2 _⊭M_ - _,_⊭_ : kripke-model w i l3 → w → formula i → UU l - M , x ⊭ a = ¬ (M , x ⊨ a) + _⊨_ : kripke-model w i l3 × w → formula i → Prop l + M , x ⊨ var n = raise-Prop l (V M n x) + M , x ⊨ ⊥ = raise-empty-Prop l + M , x ⊨ a ⇒ b = hom-Prop (M , x ⊨ a) (M , x ⊨ b) + M , x ⊨ □ a = Π-Prop w λ y -> function-Prop (R (F M) x y) (M , y ⊨ a) - _⊨M_ : kripke-model w i l3 → formula i → UU l - M ⊨M a = ∀ x → M , x ⊨ a + _⊭_ : kripke-model w i l3 × w → formula i → Prop l + M , x ⊭ a = neg-Prop (M , x ⊨ a) - _⊭M_ : kripke-model w i l3 → formula i → UU l - M ⊭M a = ¬ (M ⊨M a) + _⊨M_ : kripke-model w i l3 → formula i → Prop l + M ⊨M a = Π-Prop w λ x → M , x ⊨ a + + _⊭M_ : kripke-model w i l3 → formula i → Prop l + M ⊭M a = neg-Prop (M ⊨M a) ``` **Soundness** ```agda + is-classical-Prop : kripke-model w i l3 → Prop l + is-classical-Prop M = Π-Prop (formula i) (λ a → Π-Prop w (λ x → is-decidable-Prop (M , x ⊨ a))) + is-classical : kripke-model w i l3 → UU l - is-classical M = ∀ a x → is-decidable (M , x ⊨ a) + is-classical = type-Prop ∘ is-classical-Prop - soundness : {a : formula i} + classical-model : UU (lsuc l1 ⊔ l2 ⊔ lsuc l3) + classical-model = Σ (kripke-model w i l3) is-classical + + soundness : ((M , _) : classical-model) + → {a : formula i} → ⊢ a - → (M : kripke-model w i l3) - → is-classical M - → M ⊨M a - soundness ax-k M dec x = λ fa _ → fa - soundness ax-s M dec x = λ fabc fab fa → fabc fa (fab fa) - soundness {_ ⇒ a} ax-dn M dec x with dec a x + → type-Prop (M ⊨M a) + soundness _ ax-k x = λ fa _ → fa + soundness _ ax-s x = λ fabc fab fa → fabc fa (fab fa) + soundness (_ , cl) {_ ⇒ a} ax-dn x with cl a x ... | inl fa = λ _ → fa - ... | inr nfa = λ fnna → ex-falso (lower (fnna (λ fa → lift (nfa fa)))) - soundness ax-n M dec x = λ fab fa y r → fab y r (fa y r) - soundness (mp dab db) M dec x = soundness dab M dec x (soundness db M dec x) - soundness (nec d) M dec x y _ = soundness d M dec y + ... | inr nfa = λ fnna → ex-falso (map-inv-raise (fnna (λ fa → map-raise (nfa fa)))) + soundness _ ax-n x = λ fab fa y r → fab y r (fa y r) + soundness CM (mp dab da) x = soundness CM dab x (soundness CM da x) + soundness CM (nec d) x y _ = soundness CM d y finite-is-classical : ((M , _) : finite-model w i l3) - → (∀ n x → is-decidable (V M n x)) + → (∀ n x → type-Prop (is-decidable-Prop (V M n x))) → (∀ x y → is-decidable (R (F M) x y)) → is-classical M finite-is-classical (M , is-fin) dec-val dec-r (var n) x = - is-decidable-Lift _ (dec-val n x) + is-decidable-raise _ _ (dec-val n x) finite-is-classical (M , is-fin) dec-val dec-r ⊥ x = - inr (λ ()) + inr map-inv-raise finite-is-classical (M , is-fin) dec-val dec-r (a ⇒ b) x = is-decidable-function-type (finite-is-classical (M , is-fin) dec-val dec-r a x) @@ -201,36 +199,39 @@ module _ ((finite-is-classical (M , is-fin) dec-val dec-r a y))) finite-soundness : ((M , _) : finite-model w i l3) - → (∀ n x → is-decidable (V M n x)) + → (∀ n x → type-Prop (is-decidable-Prop (V M n x))) → (∀ x y → is-decidable (R (F M) x y)) → {a : formula i} → ⊢ a - → M ⊨M a - finite-soundness (M , is-fin) dec-val dec-r d = - soundness d M (finite-is-classical ((M , is-fin)) dec-val dec-r) + → type-Prop (M ⊨M a) + finite-soundness FM@(M , is-fin) dec-val dec-r = + soundness (M , finite-is-classical FM dec-val dec-r) -double-negation-LEM : {l : Level} → ((A : UU l) → ¬¬ A → A) → LEM l -double-negation-LEM dn A = dn _ λ ndec → ndec (inr (λ a → ndec (inl a))) +double-negation-LEM : {l : Level} → ((P : Prop l) → ¬¬ (type-Prop P) → (type-Prop P)) → LEM l +double-negation-LEM dnP P = dnP (is-decidable-Prop P) (λ ndec → ndec (inr (λ p → ndec (inl p)))) -lift-dn : {l1 l2 : Level} {A : UU l1} → ¬¬ A → (Lift l2 A → Lift l2 empty) → Lift l2 empty -lift-dn dn nA = lift (dn (λ a → lower (nA (lift a)))) +raise-dn : {l1 l2 : Level} {A : UU l1} → ¬¬ A → (raise l2 A → raise-empty l2) → raise-empty l2 +raise-dn dnA rnA = map-raise (dnA λ a → map-inv-raise (rnA (map-raise a))) -required-LEM : ({l1 l2 l3 : Level} (i : UU l2) (a : formula i) +required-LEM : ({l1 l2 l3 : Level} (w : UU l1) (i : UU l2) (a : formula i) + (M : kripke-model w i l3) + → {a : formula i} → ⊢ a - → (w : UU l1) - → (M : kripke-model w i l3) - → M ⊨M a) + → type-Prop (M ⊨M a)) → {l : Level} → LEM l -required-LEM sound {l} = - double-negation-LEM λ A nnA → lower (sound unit b ax-dn unit (M A) star (lift-dn nnA)) +required-LEM sound {l} = double-negation-LEM required-double-negation where - a = var star - b = (~ (~ a)) ⇒ a + required-double-negation : (P : Prop l) → ¬¬ (type-Prop P) → type-Prop P + required-double-negation P dnP = map-inv-raise (sound unit unit b (M P) (ax-dn {a = a}) star (raise-dn dnP)) + where + + a = var star + b = ~~ a ⇒ a - f : kripke-frame unit - f = frame (λ _ _ → empty) (unit-trunc-Prop star) + f : kripke-frame unit + f = frame (λ _ _ → empty) (unit-trunc-Prop star) - M : UU l → kripke-model unit unit l - M A = model f (λ _ _ → A) + M : Prop l → kripke-model unit unit l + M P = model f (λ _ _ → P) ``` From 661adc60dc4ebe53c796006e6ab3468f04cadb52 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Sun, 22 Oct 2023 19:56:32 +0300 Subject: [PATCH 09/63] Small fixes, and refactor kripke-frame --- src/modal-logic/modal-logic-syntax.lagda.md | 46 ++++++++++----------- 1 file changed, 23 insertions(+), 23 deletions(-) diff --git a/src/modal-logic/modal-logic-syntax.lagda.md b/src/modal-logic/modal-logic-syntax.lagda.md index 3f58eb2764..5299b6cb22 100644 --- a/src/modal-logic/modal-logic-syntax.lagda.md +++ b/src/modal-logic/modal-logic-syntax.lagda.md @@ -77,7 +77,7 @@ module _ a ∨ b = ~ a ⇒ b _∧_ : formula i → formula i → formula i - a ∧ b = ~ a ⇒ ~ b + a ∧ b = ~ (a ⇒ ~ b) ⊤ : formula i ⊤ = ~ ⊥ @@ -99,36 +99,34 @@ module _ ```agda module _ {l1 : Level} - (w : UU l1) + ((w , _ ) : Inhabited-Type l1) where record kripke-frame : UU (lsuc l1) where constructor frame field R : Relation l1 w - frame-inhabited : is-inhabited w open kripke-frame public module _ {l1 l2 : Level} - (w : UU l1) (i : UU l2) + (W@(w , _) : Inhabited-Type l1) (i : UU l2) (l3 : Level) where record kripke-model : UU (lsuc l1 ⊔ l2 ⊔ lsuc l3) where constructor model field - F : kripke-frame w + F : kripke-frame W V : i → w → Prop l3 open kripke-model public finite-model : UU (lsuc l1 ⊔ l2 ⊔ lsuc l3) finite-model = kripke-model × is-finite w - module _ {l1 l2 l3 : Level} - {w : UU l1} {i : UU l2} + {W@(w , _) : Inhabited-Type l1} {i : UU l2} where private @@ -139,33 +137,33 @@ module _ infix 2 _⊨M_ infix 2 _⊭M_ - _⊨_ : kripke-model w i l3 × w → formula i → Prop l + _⊨_ : kripke-model W i l3 × w → formula i → Prop l M , x ⊨ var n = raise-Prop l (V M n x) M , x ⊨ ⊥ = raise-empty-Prop l M , x ⊨ a ⇒ b = hom-Prop (M , x ⊨ a) (M , x ⊨ b) M , x ⊨ □ a = Π-Prop w λ y -> function-Prop (R (F M) x y) (M , y ⊨ a) - _⊭_ : kripke-model w i l3 × w → formula i → Prop l + _⊭_ : kripke-model W i l3 × w → formula i → Prop l M , x ⊭ a = neg-Prop (M , x ⊨ a) - _⊨M_ : kripke-model w i l3 → formula i → Prop l + _⊨M_ : kripke-model W i l3 → formula i → Prop l M ⊨M a = Π-Prop w λ x → M , x ⊨ a - _⊭M_ : kripke-model w i l3 → formula i → Prop l + _⊭M_ : kripke-model W i l3 → formula i → Prop l M ⊭M a = neg-Prop (M ⊨M a) ``` **Soundness** ```agda - is-classical-Prop : kripke-model w i l3 → Prop l + is-classical-Prop : kripke-model W i l3 → Prop l is-classical-Prop M = Π-Prop (formula i) (λ a → Π-Prop w (λ x → is-decidable-Prop (M , x ⊨ a))) - is-classical : kripke-model w i l3 → UU l + is-classical : kripke-model W i l3 → UU l is-classical = type-Prop ∘ is-classical-Prop classical-model : UU (lsuc l1 ⊔ l2 ⊔ lsuc l3) - classical-model = Σ (kripke-model w i l3) is-classical + classical-model = Σ (kripke-model W i l3) is-classical soundness : ((M , _) : classical-model) → {a : formula i} @@ -180,7 +178,7 @@ module _ soundness CM (mp dab da) x = soundness CM dab x (soundness CM da x) soundness CM (nec d) x y _ = soundness CM d y - finite-is-classical : ((M , _) : finite-model w i l3) + finite-is-classical : ((M , _) : finite-model W i l3) → (∀ n x → type-Prop (is-decidable-Prop (V M n x))) → (∀ x y → is-decidable (R (F M) x y)) → is-classical M @@ -198,7 +196,7 @@ module _ is-decidable-function-type (dec-r x y) ((finite-is-classical (M , is-fin) dec-val dec-r a y))) - finite-soundness : ((M , _) : finite-model w i l3) + finite-soundness : ((M , _) : finite-model W i l3) → (∀ n x → type-Prop (is-decidable-Prop (V M n x))) → (∀ x y → is-decidable (R (F M) x y)) → {a : formula i} @@ -214,8 +212,8 @@ double-negation-LEM dnP P = dnP (is-decidable-Prop P) (λ ndec → ndec (inr (λ raise-dn : {l1 l2 : Level} {A : UU l1} → ¬¬ A → (raise l2 A → raise-empty l2) → raise-empty l2 raise-dn dnA rnA = map-raise (dnA λ a → map-inv-raise (rnA (map-raise a))) -required-LEM : ({l1 l2 l3 : Level} (w : UU l1) (i : UU l2) (a : formula i) - (M : kripke-model w i l3) +required-LEM : ({l1 l2 l3 : Level} (W : Inhabited-Type l1) (i : UU l2) (a : formula i) + (M : kripke-model W i l3) → {a : formula i} → ⊢ a → type-Prop (M ⊨M a)) @@ -223,15 +221,17 @@ required-LEM : ({l1 l2 l3 : Level} (w : UU l1) (i : UU l2) (a : formula i) required-LEM sound {l} = double-negation-LEM required-double-negation where required-double-negation : (P : Prop l) → ¬¬ (type-Prop P) → type-Prop P - required-double-negation P dnP = map-inv-raise (sound unit unit b (M P) (ax-dn {a = a}) star (raise-dn dnP)) + required-double-negation P dnP = map-inv-raise (sound W unit b (M P) (ax-dn {a = a}) star (raise-dn dnP)) where - a = var star b = ~~ a ⇒ a - f : kripke-frame unit - f = frame (λ _ _ → empty) (unit-trunc-Prop star) + W : Inhabited-Type lzero + W = unit , unit-trunc-Prop star + + f : kripke-frame W + f = frame (λ _ _ → empty) - M : Prop l → kripke-model unit unit l + M : Prop l → kripke-model W unit l M P = model f (λ _ _ → P) ``` From 3b3ba68743d9eb9b0e88ef5f013f60541d8da817 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Mon, 23 Oct 2023 00:00:20 +0300 Subject: [PATCH 10/63] Refactor modal logic --- src/modal-logic.lagda.md | 7 ++ src/modal-logic/modal-logic-syntax.lagda.md | 107 ++++++++++++-------- 2 files changed, 71 insertions(+), 43 deletions(-) create mode 100644 src/modal-logic.lagda.md diff --git a/src/modal-logic.lagda.md b/src/modal-logic.lagda.md new file mode 100644 index 0000000000..051f8a4419 --- /dev/null +++ b/src/modal-logic.lagda.md @@ -0,0 +1,7 @@ +# Modal logic + +```agda +module modal-logic where + +open import modal-logic.modal-logic-syntax public +``` diff --git a/src/modal-logic/modal-logic-syntax.lagda.md b/src/modal-logic/modal-logic-syntax.lagda.md index 5299b6cb22..b2bb268be6 100644 --- a/src/modal-logic/modal-logic-syntax.lagda.md +++ b/src/modal-logic/modal-logic-syntax.lagda.md @@ -7,27 +7,27 @@ module modal-logic.modal-logic-syntax where
Imports ```agda -open import foundation.cartesian-product-types -open import foundation.identity-types -open import foundation.propositions open import foundation.binary-relations +open import foundation.cartesian-product-types +open import foundation.coproduct-types +open import foundation.decidable-propositions open import foundation.decidable-types open import foundation.dependent-pair-types -open import foundation.universe-levels -open import foundation.inhabited-types -open import foundation.coproduct-types -open import foundation.function-types +open import foundation.double-negation open import foundation.empty-types -open import foundation.unit-type +open import foundation.function-types +open import foundation.identity-types +open import foundation.inhabited-types +open import foundation.law-of-excluded-middle open import foundation.negation -open import foundation.double-negation open import foundation.propositional-truncations +open import foundation.propositions open import foundation.raising-universe-levels -open import foundation.decidable-propositions -open import foundation.law-of-excluded-middle +open import foundation.unit-type +open import foundation.universe-levels -open import univalent-combinatorics.finite-types open import univalent-combinatorics.decidable-dependent-function-types +open import univalent-combinatorics.finite-types ```
@@ -99,7 +99,7 @@ module _ ```agda module _ {l1 : Level} - ((w , _ ) : Inhabited-Type l1) + ((w , _) : Inhabited-Type l1) where record kripke-frame : UU (lsuc l1) where @@ -157,7 +157,8 @@ module _ ```agda is-classical-Prop : kripke-model W i l3 → Prop l - is-classical-Prop M = Π-Prop (formula i) (λ a → Π-Prop w (λ x → is-decidable-Prop (M , x ⊨ a))) + is-classical-Prop M = + Π-Prop (formula i) (λ a → Π-Prop w (λ x → is-decidable-Prop (M , x ⊨ a))) is-classical : kripke-model W i l3 → UU l is-classical = type-Prop ∘ is-classical-Prop @@ -165,63 +166,83 @@ module _ classical-model : UU (lsuc l1 ⊔ l2 ⊔ lsuc l3) classical-model = Σ (kripke-model W i l3) is-classical - soundness : ((M , _) : classical-model) - → {a : formula i} - → ⊢ a - → type-Prop (M ⊨M a) + soundness : + ((M , _) : classical-model) + {a : formula i} → + ⊢ a → + type-Prop (M ⊨M a) soundness _ ax-k x = λ fa _ → fa soundness _ ax-s x = λ fabc fab fa → fabc fa (fab fa) soundness (_ , cl) {_ ⇒ a} ax-dn x with cl a x ... | inl fa = λ _ → fa - ... | inr nfa = λ fnna → ex-falso (map-inv-raise (fnna (λ fa → map-raise (nfa fa)))) + ... | inr nfa = + λ fnna → ex-falso (map-inv-raise (fnna (λ fa → map-raise (nfa fa)))) soundness _ ax-n x = λ fab fa y r → fab y r (fa y r) soundness CM (mp dab da) x = soundness CM dab x (soundness CM da x) soundness CM (nec d) x y _ = soundness CM d y - finite-is-classical : ((M , _) : finite-model W i l3) - → (∀ n x → type-Prop (is-decidable-Prop (V M n x))) - → (∀ x y → is-decidable (R (F M) x y)) - → is-classical M + finite-is-classical : + ((M , _) : finite-model W i l3) → + (∀ n x → type-Prop (is-decidable-Prop (V M n x))) → + (∀ x y → is-decidable (R (F M) x y)) → + is-classical M finite-is-classical (M , is-fin) dec-val dec-r (var n) x = is-decidable-raise _ _ (dec-val n x) finite-is-classical (M , is-fin) dec-val dec-r ⊥ x = inr map-inv-raise finite-is-classical (M , is-fin) dec-val dec-r (a ⇒ b) x = is-decidable-function-type - (finite-is-classical (M , is-fin) dec-val dec-r a x) - (finite-is-classical (M , is-fin) dec-val dec-r b x) + ( finite-is-classical (M , is-fin) dec-val dec-r a x) + ( finite-is-classical (M , is-fin) dec-val dec-r b x) finite-is-classical (M , is-fin) dec-val dec-r (□ a) x = is-decidable-Π-is-finite is-fin - (λ y → + ( λ y → is-decidable-function-type (dec-r x y) ((finite-is-classical (M , is-fin) dec-val dec-r a y))) - finite-soundness : ((M , _) : finite-model W i l3) - → (∀ n x → type-Prop (is-decidable-Prop (V M n x))) - → (∀ x y → is-decidable (R (F M) x y)) - → {a : formula i} - → ⊢ a - → type-Prop (M ⊨M a) + finite-soundness : + ((M , _) : finite-model W i l3) → + (∀ n x → type-Prop (is-decidable-Prop (V M n x))) → + (∀ x y → is-decidable (R (F M) x y)) → + {a : formula i} → + ⊢ a → + type-Prop (M ⊨M a) finite-soundness FM@(M , is-fin) dec-val dec-r = soundness (M , finite-is-classical FM dec-val dec-r) +double-negation-LEM : + {l : Level} → + ((P : Prop l) → + ¬¬ (type-Prop P) → + (type-Prop P)) → + LEM l +double-negation-LEM dnP P = + dnP (is-decidable-Prop P) (λ ndec → ndec (inr (λ p → ndec (inl p)))) -double-negation-LEM : {l : Level} → ((P : Prop l) → ¬¬ (type-Prop P) → (type-Prop P)) → LEM l -double-negation-LEM dnP P = dnP (is-decidable-Prop P) (λ ndec → ndec (inr (λ p → ndec (inl p)))) - -raise-dn : {l1 l2 : Level} {A : UU l1} → ¬¬ A → (raise l2 A → raise-empty l2) → raise-empty l2 +raise-dn : + {l1 l2 : Level} + {A : UU l1} → + ¬¬ A → + (raise l2 A → raise-empty l2) → + raise-empty l2 raise-dn dnA rnA = map-raise (dnA λ a → map-inv-raise (rnA (map-raise a))) -required-LEM : ({l1 l2 l3 : Level} (W : Inhabited-Type l1) (i : UU l2) (a : formula i) - (M : kripke-model W i l3) - → {a : formula i} - → ⊢ a - → type-Prop (M ⊨M a)) - → {l : Level} → LEM l +required-LEM : + ( {l1 l2 l3 : Level} + (W : Inhabited-Type l1) + (i : UU l2) + (a : formula i) + (M : kripke-model W i l3) → + {a : formula i} → + ⊢ a → + type-Prop (M ⊨M a)) → + {l : Level} → + LEM l required-LEM sound {l} = double-negation-LEM required-double-negation where required-double-negation : (P : Prop l) → ¬¬ (type-Prop P) → type-Prop P - required-double-negation P dnP = map-inv-raise (sound W unit b (M P) (ax-dn {a = a}) star (raise-dn dnP)) + required-double-negation P dnP = + map-inv-raise (sound W unit b (M P) (ax-dn {a = a}) star (raise-dn dnP)) where a = var star b = ~~ a ⇒ a From 957e11261c0a4294254c28db769935eefd50ed31 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Mon, 23 Oct 2023 11:55:53 +0300 Subject: [PATCH 11/63] - Refactor - Add me to contributors --- CONTRIBUTORS.toml | 5 + src/modal-logic/modal-logic-syntax.lagda.md | 155 ++++++++++---------- 2 files changed, 79 insertions(+), 81 deletions(-) diff --git a/CONTRIBUTORS.toml b/CONTRIBUTORS.toml index f011a114e6..a2f7281d57 100644 --- a/CONTRIBUTORS.toml +++ b/CONTRIBUTORS.toml @@ -242,3 +242,8 @@ github = "UlrikBuchholtz" displayName = "Garrett Figueroa" usernames = [ "Garrett Figueroa", "djspacewhale" ] github = "djspacewhale" + +[[contributors]] +displayName = "Viktor Yudov" +usernames = [ "Viktor Yudov" ] +github = "spcfox" diff --git a/src/modal-logic/modal-logic-syntax.lagda.md b/src/modal-logic/modal-logic-syntax.lagda.md index b2bb268be6..6338a55df2 100644 --- a/src/modal-logic/modal-logic-syntax.lagda.md +++ b/src/modal-logic/modal-logic-syntax.lagda.md @@ -38,12 +38,11 @@ TODO ## Definition -**Syntax** +### Syntax ```agda module _ - {l : Level} - (i : UU l) + {l : Level} (i : UU l) where infixr 4 _⇒_ @@ -56,8 +55,7 @@ module _ □_ : formula → formula module _ - {l : Level} - {i : UU l} + {l : Level} {i : UU l} where infixr 7 ~_ @@ -86,49 +84,48 @@ module _ ◇ a = ~ □ ~ a data ⊢_ : formula i → UU l where - ax-k : {a b : formula i} → ⊢ a ⇒ b ⇒ a - ax-s : {a b c : formula i} → ⊢ (a ⇒ b ⇒ c) ⇒ (a ⇒ b) ⇒ (a ⇒ c) - ax-dn : {a : formula i} → ⊢ ~~ a ⇒ a - ax-n : {a b : formula i} → ⊢ □ (a ⇒ b) ⇒ □ a ⇒ □ b + ax-k : (a b : formula i) → ⊢ a ⇒ b ⇒ a + ax-s : (a b c : formula i) → ⊢ (a ⇒ b ⇒ c) ⇒ (a ⇒ b) ⇒ (a ⇒ c) + ax-dn : (a : formula i) → ⊢ ~~ a ⇒ a + ax-n : (a b : formula i) → ⊢ □ (a ⇒ b) ⇒ □ a ⇒ □ b mp : {a b : formula i} → ⊢ a ⇒ b → ⊢ a → ⊢ b nec : {a : formula i} → ⊢ a → ⊢ □ a ``` -**Semantics** +### Semantics ```agda module _ - {l1 : Level} - ((w , _) : Inhabited-Type l1) + {l1 : Level} ((w , _) : Inhabited-Type l1) where record kripke-frame : UU (lsuc l1) where constructor frame field - R : Relation l1 w + frame-relation : Relation l1 w open kripke-frame public module _ - {l1 l2 : Level} - (W@(w , _) : Inhabited-Type l1) (i : UU l2) - (l3 : Level) + {l1 l2 : Level} (W@(w , _) : Inhabited-Type l1) (i : UU l2) (l3 : Level) where record kripke-model : UU (lsuc l1 ⊔ l2 ⊔ lsuc l3) where constructor model field - F : kripke-frame W - V : i → w → Prop l3 + model-frame : kripke-frame W + model-valuate : i → w → Prop l3 open kripke-model public finite-model : UU (lsuc l1 ⊔ l2 ⊔ lsuc l3) finite-model = kripke-model × is-finite w module _ - {l1 l2 l3 : Level} - {W@(w , _) : Inhabited-Type l1} {i : UU l2} + {l1 l2 l3 : Level} {W@(w , _) : Inhabited-Type l1} {i : UU l2} where + model-relation : kripke-model W i l3 → Relation l1 w + model-relation = frame-relation ∘ model-frame + private l = l1 ⊔ l2 ⊔ l3 @@ -138,24 +135,32 @@ module _ infix 2 _⊭M_ _⊨_ : kripke-model W i l3 × w → formula i → Prop l - M , x ⊨ var n = raise-Prop l (V M n x) + M , x ⊨ var n = raise-Prop l (model-valuate M n x) M , x ⊨ ⊥ = raise-empty-Prop l M , x ⊨ a ⇒ b = hom-Prop (M , x ⊨ a) (M , x ⊨ b) - M , x ⊨ □ a = Π-Prop w λ y -> function-Prop (R (F M) x y) (M , y ⊨ a) + M , x ⊨ □ a = + Π-Prop w (λ y -> function-Prop (model-relation M x y) (M , y ⊨ a)) _⊭_ : kripke-model W i l3 × w → formula i → Prop l M , x ⊭ a = neg-Prop (M , x ⊨ a) _⊨M_ : kripke-model W i l3 → formula i → Prop l - M ⊨M a = Π-Prop w λ x → M , x ⊨ a + M ⊨M a = Π-Prop w (λ x → M , x ⊨ a) _⊭M_ : kripke-model W i l3 → formula i → Prop l M ⊭M a = neg-Prop (M ⊨M a) ``` -**Soundness** +### Soundness ```agda +module _ + {l1 l2 l3 : Level} {W@(w , _) : Inhabited-Type l1} {i : UU l2} + where + + private + l = l1 ⊔ l2 ⊔ l3 + is-classical-Prop : kripke-model W i l3 → Prop l is-classical-Prop M = Π-Prop (formula i) (λ a → Π-Prop w (λ x → is-decidable-Prop (M , x ⊨ a))) @@ -171,44 +176,38 @@ module _ {a : formula i} → ⊢ a → type-Prop (M ⊨M a) - soundness _ ax-k x = λ fa _ → fa - soundness _ ax-s x = λ fabc fab fa → fabc fa (fab fa) - soundness (_ , cl) {_ ⇒ a} ax-dn x with cl a x - ... | inl fa = λ _ → fa - ... | inr nfa = - λ fnna → ex-falso (map-inv-raise (fnna (λ fa → map-raise (nfa fa)))) - soundness _ ax-n x = λ fab fa y r → fab y r (fa y r) + soundness _ (ax-k _ _) x fa _ = fa + soundness _ (ax-s _ _ _) x fabc fab fa = fabc fa (fab fa) + soundness (_ , cl) (ax-dn a) x fdna with cl a x + ... | inl fa = fa + ... | inr nfa = ex-falso (map-inv-raise (fdna (λ fa → map-raise (nfa fa)))) + soundness _ (ax-n _ _) x fab fa y r = fab y r (fa y r) soundness CM (mp dab da) x = soundness CM dab x (soundness CM da x) soundness CM (nec d) x y _ = soundness CM d y - finite-is-classical : - ((M , _) : finite-model W i l3) → - (∀ n x → type-Prop (is-decidable-Prop (V M n x))) → - (∀ x y → is-decidable (R (F M) x y)) → - is-classical M - finite-is-classical (M , is-fin) dec-val dec-r (var n) x = - is-decidable-raise _ _ (dec-val n x) - finite-is-classical (M , is-fin) dec-val dec-r ⊥ x = - inr map-inv-raise - finite-is-classical (M , is-fin) dec-val dec-r (a ⇒ b) x = - is-decidable-function-type - ( finite-is-classical (M , is-fin) dec-val dec-r a x) - ( finite-is-classical (M , is-fin) dec-val dec-r b x) - finite-is-classical (M , is-fin) dec-val dec-r (□ a) x = - is-decidable-Π-is-finite is-fin - ( λ y → - is-decidable-function-type (dec-r x y) - ((finite-is-classical (M , is-fin) dec-val dec-r a y))) - - finite-soundness : - ((M , _) : finite-model W i l3) → - (∀ n x → type-Prop (is-decidable-Prop (V M n x))) → - (∀ x y → is-decidable (R (F M) x y)) → - {a : formula i} → - ⊢ a → - type-Prop (M ⊨M a) - finite-soundness FM@(M , is-fin) dec-val dec-r = - soundness (M , finite-is-classical FM dec-val dec-r) + module _ + ((M , w-is-finite) : finite-model W i l3) + (dec-val : ∀ n x → type-Prop (is-decidable-Prop (model-valuate M n x))) + (dec-r : ∀ x y → is-decidable (model-relation M x y)) + where + + finite-is-classical : is-classical M + finite-is-classical (var n) x = + is-decidable-raise _ _ (dec-val n x) + finite-is-classical ⊥ x = + inr map-inv-raise + finite-is-classical (a ⇒ b) x = + is-decidable-function-type + ( finite-is-classical a x) + ( finite-is-classical b x) + finite-is-classical (□ a) x = + is-decidable-Π-is-finite w-is-finite + ( λ y → + is-decidable-function-type (dec-r x y) + ( finite-is-classical a y)) + + finite-soundness : {a : formula i} → ⊢ a → type-Prop (M ⊨M a) + finite-soundness = soundness (M , finite-is-classical) double-negation-LEM : {l : Level} → @@ -223,36 +222,30 @@ raise-dn : {l1 l2 : Level} {A : UU l1} → ¬¬ A → - (raise l2 A → raise-empty l2) → - raise-empty l2 -raise-dn dnA rnA = map-raise (dnA λ a → map-inv-raise (rnA (map-raise a))) + (raise l2 A → raise-empty l2) → raise-empty l2 +raise-dn dnA rnA = map-raise (dnA (λ a → map-inv-raise (rnA (map-raise a)))) required-LEM : ( {l1 l2 l3 : Level} (W : Inhabited-Type l1) - (i : UU l2) - (a : formula i) + {i : UU l2} (M : kripke-model W i l3) → {a : formula i} → ⊢ a → type-Prop (M ⊨M a)) → - {l : Level} → + (l : Level) → LEM l -required-LEM sound {l} = double-negation-LEM required-double-negation +required-LEM sound l = double-negation-LEM required-double-negation where - required-double-negation : (P : Prop l) → ¬¬ (type-Prop P) → type-Prop P - required-double-negation P dnP = - map-inv-raise (sound W unit b (M P) (ax-dn {a = a}) star (raise-dn dnP)) - where - a = var star - b = ~~ a ⇒ a - - W : Inhabited-Type lzero - W = unit , unit-trunc-Prop star - - f : kripke-frame W - f = frame (λ _ _ → empty) - - M : Prop l → kripke-model W unit l - M P = model f (λ _ _ → P) + required-double-negation : (P : Prop l) → ¬¬ (type-Prop P) → type-Prop P + required-double-negation P dnP = + map-inv-raise + ( sound + ( unit , unit-trunc-Prop star) + ( model + ( frame (λ _ _ → empty)) + ( λ _ _ → P)) + ( ax-dn (var star)) + star + ( raise-dn dnP)) ``` From 341456aba0bcd401d4ba2bb2e33578c2dd8f6eb0 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Mon, 23 Oct 2023 12:44:16 +0300 Subject: [PATCH 12/63] Change operators levels and associativity --- src/modal-logic/modal-logic-syntax.lagda.md | 40 ++++++++++----------- 1 file changed, 20 insertions(+), 20 deletions(-) diff --git a/src/modal-logic/modal-logic-syntax.lagda.md b/src/modal-logic/modal-logic-syntax.lagda.md index 6338a55df2..9621c8d1f9 100644 --- a/src/modal-logic/modal-logic-syntax.lagda.md +++ b/src/modal-logic/modal-logic-syntax.lagda.md @@ -45,8 +45,8 @@ module _ {l : Level} (i : UU l) where - infixr 4 _⇒_ - infixr 7 □_ + infixr 6 _⇒_ + infixr 25 □_ data formula : UU l where var : i → formula @@ -58,12 +58,12 @@ module _ {l : Level} {i : UU l} where - infixr 7 ~_ - infixr 7 ~~_ - infixr 5 _∨_ - infixr 6 _∧_ - infixr 7 ◇_ - infixr 2 ⊢_ + infixr 25 ~_ + infixr 25 ~~_ + infixl 10 _∨_ + infixl 15 _∧_ + infixr 25 ◇_ + infix 5 ⊢_ ~_ : formula i → formula i ~ a = a ⇒ ⊥ @@ -129,23 +129,23 @@ module _ private l = l1 ⊔ l2 ⊔ l3 - infix 2 _⊨_ - infix 2 _⊭_ - infix 2 _⊨M_ - infix 2 _⊭M_ + infix 5 _⊨_ + infix 5 _⊭_ + infix 5 _⊨M_ + infix 5 _⊭M_ _⊨_ : kripke-model W i l3 × w → formula i → Prop l - M , x ⊨ var n = raise-Prop l (model-valuate M n x) - M , x ⊨ ⊥ = raise-empty-Prop l - M , x ⊨ a ⇒ b = hom-Prop (M , x ⊨ a) (M , x ⊨ b) - M , x ⊨ □ a = - Π-Prop w (λ y -> function-Prop (model-relation M x y) (M , y ⊨ a)) + (M , x) ⊨ var n = raise-Prop l (model-valuate M n x) + (M , x) ⊨ ⊥ = raise-empty-Prop l + (M , x) ⊨ a ⇒ b = hom-Prop ((M , x) ⊨ a) ((M , x) ⊨ b) + (M , x) ⊨ □ a = + Π-Prop w (λ y -> function-Prop (model-relation M x y) ((M , y) ⊨ a)) _⊭_ : kripke-model W i l3 × w → formula i → Prop l - M , x ⊭ a = neg-Prop (M , x ⊨ a) + (M , x) ⊭ a = neg-Prop ((M , x) ⊨ a) _⊨M_ : kripke-model W i l3 → formula i → Prop l - M ⊨M a = Π-Prop w (λ x → M , x ⊨ a) + M ⊨M a = Π-Prop w (λ x → (M , x) ⊨ a) _⊭M_ : kripke-model W i l3 → formula i → Prop l M ⊭M a = neg-Prop (M ⊨M a) @@ -163,7 +163,7 @@ module _ is-classical-Prop : kripke-model W i l3 → Prop l is-classical-Prop M = - Π-Prop (formula i) (λ a → Π-Prop w (λ x → is-decidable-Prop (M , x ⊨ a))) + Π-Prop (formula i) (λ a → Π-Prop w (λ x → is-decidable-Prop ((M , x) ⊨ a))) is-classical : kripke-model W i l3 → UU l is-classical = type-Prop ∘ is-classical-Prop From 9db88caeebe2778645ce0bc0594bfda68aad52a4 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Mon, 23 Oct 2023 13:11:14 +0300 Subject: [PATCH 13/63] Refactor universe levels --- src/modal-logic/modal-logic-syntax.lagda.md | 70 ++++++++++++--------- 1 file changed, 39 insertions(+), 31 deletions(-) diff --git a/src/modal-logic/modal-logic-syntax.lagda.md b/src/modal-logic/modal-logic-syntax.lagda.md index 9621c8d1f9..ab9d7947b3 100644 --- a/src/modal-logic/modal-logic-syntax.lagda.md +++ b/src/modal-logic/modal-logic-syntax.lagda.md @@ -96,58 +96,61 @@ module _ ```agda module _ - {l1 : Level} ((w , _) : Inhabited-Type l1) + {l1 : Level} ((w , _) : Inhabited-Type l1) (l2 : Level) where - record kripke-frame : UU (lsuc l1) where + record kripke-frame : UU (l1 ⊔ lsuc l2) where constructor frame field - frame-relation : Relation l1 w + frame-relation : Relation l2 w open kripke-frame public module _ - {l1 l2 : Level} (W@(w , _) : Inhabited-Type l1) (i : UU l2) (l3 : Level) + {l1 : Level} (W@(w , _) : Inhabited-Type l1) + (l2 : Level) + {l3 : Level} (i : UU l3) + (l4 : Level) where - record kripke-model : UU (lsuc l1 ⊔ l2 ⊔ lsuc l3) where + record kripke-model : UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4) where constructor model field - model-frame : kripke-frame W - model-valuate : i → w → Prop l3 + model-frame : kripke-frame W l2 + model-valuate : i → w → Prop l4 open kripke-model public - finite-model : UU (lsuc l1 ⊔ l2 ⊔ lsuc l3) + finite-model : UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4) finite-model = kripke-model × is-finite w module _ - {l1 l2 l3 : Level} {W@(w , _) : Inhabited-Type l1} {i : UU l2} + {l1 l2 l3 l4 : Level} {W@(w , _) : Inhabited-Type l1} {i : UU l3} where - model-relation : kripke-model W i l3 → Relation l1 w + model-relation : kripke-model W l2 i l4 → Relation l2 w model-relation = frame-relation ∘ model-frame private - l = l1 ⊔ l2 ⊔ l3 + l = l1 ⊔ l2 ⊔ l4 infix 5 _⊨_ infix 5 _⊭_ infix 5 _⊨M_ infix 5 _⊭M_ - _⊨_ : kripke-model W i l3 × w → formula i → Prop l + _⊨_ : kripke-model W l2 i l4 × w → formula i → Prop l (M , x) ⊨ var n = raise-Prop l (model-valuate M n x) (M , x) ⊨ ⊥ = raise-empty-Prop l (M , x) ⊨ a ⇒ b = hom-Prop ((M , x) ⊨ a) ((M , x) ⊨ b) (M , x) ⊨ □ a = Π-Prop w (λ y -> function-Prop (model-relation M x y) ((M , y) ⊨ a)) - _⊭_ : kripke-model W i l3 × w → formula i → Prop l + _⊭_ : kripke-model W l2 i l4 × w → formula i → Prop l (M , x) ⊭ a = neg-Prop ((M , x) ⊨ a) - _⊨M_ : kripke-model W i l3 → formula i → Prop l + _⊨M_ : kripke-model W l2 i l4 → formula i → Prop l M ⊨M a = Π-Prop w (λ x → (M , x) ⊨ a) - _⊭M_ : kripke-model W i l3 → formula i → Prop l + _⊭M_ : kripke-model W l2 i l4 → formula i → Prop l M ⊭M a = neg-Prop (M ⊨M a) ``` @@ -155,21 +158,21 @@ module _ ```agda module _ - {l1 l2 l3 : Level} {W@(w , _) : Inhabited-Type l1} {i : UU l2} + {l1 l2 l3 l4 : Level} {W@(w , _) : Inhabited-Type l1} {i : UU l3} where private - l = l1 ⊔ l2 ⊔ l3 + l = l1 ⊔ l2 ⊔ l3 ⊔ l4 - is-classical-Prop : kripke-model W i l3 → Prop l + is-classical-Prop : kripke-model W l2 i l4 → Prop l is-classical-Prop M = Π-Prop (formula i) (λ a → Π-Prop w (λ x → is-decidable-Prop ((M , x) ⊨ a))) - is-classical : kripke-model W i l3 → UU l + is-classical : kripke-model W l2 i l4 → UU l is-classical = type-Prop ∘ is-classical-Prop - classical-model : UU (lsuc l1 ⊔ l2 ⊔ lsuc l3) - classical-model = Σ (kripke-model W i l3) is-classical + classical-model : UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4) + classical-model = Σ (kripke-model W l2 i l4) is-classical soundness : ((M , _) : classical-model) @@ -186,7 +189,7 @@ module _ soundness CM (nec d) x y _ = soundness CM d y module _ - ((M , w-is-finite) : finite-model W i l3) + ((M , w-is-finite) : finite-model W l2 i l4) (dec-val : ∀ n x → type-Prop (is-decidable-Prop (model-valuate M n x))) (dec-r : ∀ x y → is-decidable (model-relation M x y)) where @@ -225,17 +228,22 @@ raise-dn : (raise l2 A → raise-empty l2) → raise-empty l2 raise-dn dnA rnA = map-raise (dnA (λ a → map-inv-raise (rnA (map-raise a)))) -required-LEM : - ( {l1 l2 l3 : Level} - (W : Inhabited-Type l1) - {i : UU l2} - (M : kripke-model W i l3) → - {a : formula i} → - ⊢ a → - type-Prop (M ⊨M a)) → +full-soundness : UUω +full-soundness = + {l1 l2 l3 l4 : Level} + (W : Inhabited-Type l1) + {i : UU l3} + (M : kripke-model W l2 i l4) + {a : formula i} → + ⊢ a → + type-Prop (M ⊨M a) + +full-soundness-required-LEM : + full-soundness → (l : Level) → LEM l -required-LEM sound l = double-negation-LEM required-double-negation +full-soundness-required-LEM sound l = + double-negation-LEM required-double-negation where required-double-negation : (P : Prop l) → ¬¬ (type-Prop P) → type-Prop P required-double-negation P dnP = From 30056ba0a6cfef1cabc92bc5d990a542a46b5e02 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Mon, 23 Oct 2023 13:23:16 +0300 Subject: [PATCH 14/63] Small fixes --- src/modal-logic/modal-logic-syntax.lagda.md | 13 ++++++------- 1 file changed, 6 insertions(+), 7 deletions(-) diff --git a/src/modal-logic/modal-logic-syntax.lagda.md b/src/modal-logic/modal-logic-syntax.lagda.md index ab9d7947b3..157d59d54a 100644 --- a/src/modal-logic/modal-logic-syntax.lagda.md +++ b/src/modal-logic/modal-logic-syntax.lagda.md @@ -140,7 +140,7 @@ module _ _⊨_ : kripke-model W l2 i l4 × w → formula i → Prop l (M , x) ⊨ var n = raise-Prop l (model-valuate M n x) (M , x) ⊨ ⊥ = raise-empty-Prop l - (M , x) ⊨ a ⇒ b = hom-Prop ((M , x) ⊨ a) ((M , x) ⊨ b) + (M , x) ⊨ a ⇒ b = implication-Prop ((M , x) ⊨ a) ((M , x) ⊨ b) (M , x) ⊨ □ a = Π-Prop w (λ y -> function-Prop (model-relation M x y) ((M , y) ⊨ a)) @@ -221,12 +221,13 @@ double-negation-LEM : double-negation-LEM dnP P = dnP (is-decidable-Prop P) (λ ndec → ndec (inr (λ p → ndec (inl p)))) -raise-dn : +raise-double-negation : {l1 l2 : Level} {A : UU l1} → ¬¬ A → (raise l2 A → raise-empty l2) → raise-empty l2 -raise-dn dnA rnA = map-raise (dnA (λ a → map-inv-raise (rnA (map-raise a)))) +raise-double-negation dnA rnA = + map-raise (dnA (λ a → map-inv-raise (rnA (map-raise a)))) full-soundness : UUω full-soundness = @@ -250,10 +251,8 @@ full-soundness-required-LEM sound l = map-inv-raise ( sound ( unit , unit-trunc-Prop star) - ( model - ( frame (λ _ _ → empty)) - ( λ _ _ → P)) + ( model (frame (λ _ _ → empty)) (λ _ _ → P)) ( ax-dn (var star)) star - ( raise-dn dnP)) + ( raise-double-negation dnP)) ``` From 0d223691f536bd9aaf44d8e4373eaef9f4f1f53e Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Tue, 24 Oct 2023 01:56:02 +0300 Subject: [PATCH 15/63] Refactor --- src/modal-logic.lagda.md | 5 +- src/modal-logic/K-soundness.lagda.md | 142 +++++++++++ src/modal-logic/K-syntax.lagda.md | 34 +++ src/modal-logic/formulas.lagda.md | 63 +++++ src/modal-logic/kripke-semantics.lagda.md | 123 ++++++++++ src/modal-logic/modal-logic-syntax.lagda.md | 258 -------------------- 6 files changed, 366 insertions(+), 259 deletions(-) create mode 100644 src/modal-logic/K-soundness.lagda.md create mode 100644 src/modal-logic/K-syntax.lagda.md create mode 100644 src/modal-logic/formulas.lagda.md create mode 100644 src/modal-logic/kripke-semantics.lagda.md delete mode 100644 src/modal-logic/modal-logic-syntax.lagda.md diff --git a/src/modal-logic.lagda.md b/src/modal-logic.lagda.md index 051f8a4419..4592c16d0c 100644 --- a/src/modal-logic.lagda.md +++ b/src/modal-logic.lagda.md @@ -3,5 +3,8 @@ ```agda module modal-logic where -open import modal-logic.modal-logic-syntax public +open import modal-logic.K-soundness public +open import modal-logic.K-syntax public +open import modal-logic.formulas public +open import modal-logic.kripke-semantics public ``` diff --git a/src/modal-logic/K-soundness.lagda.md b/src/modal-logic/K-soundness.lagda.md new file mode 100644 index 0000000000..b333cb354d --- /dev/null +++ b/src/modal-logic/K-soundness.lagda.md @@ -0,0 +1,142 @@ +# Modal logic K soundness + +```agda +module modal-logic.K-soundness where +``` + +
Imports + +```agda +open import foundation.coproduct-types +open import foundation.decidable-propositions +open import foundation.decidable-types +open import foundation.dependent-pair-types +open import foundation.double-negation +open import foundation.empty-types +open import foundation.equivalences +open import foundation.function-types +open import foundation.functoriality-function-types +open import foundation.inhabited-types +open import foundation.law-of-excluded-middle +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.raising-universe-levels +open import foundation.sets +open import foundation.unit-type +open import foundation.universe-levels + +open import modal-logic.K-syntax +open import modal-logic.formulas +open import modal-logic.kripke-semantics + +open import univalent-combinatorics.decidable-dependent-function-types +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l1 l2 l3 l4 : Level} {w : Inhabited-Type l1} {i : Set l3} + where + + private + l = l1 ⊔ l2 ⊔ l3 ⊔ l4 + + soundness : + ((M , _) : decidable-model w l2 i l4) + {a : formula i} → + ⊢ a → + type-Prop (M ⊨M a) + soundness _ (ax-k _ _) x fa _ = fa + soundness _ (ax-s _ _ _) x fabc fab fa = fabc fa (fab fa) + soundness (_ , cl) (ax-dn a) x fdna with cl a x + ... | inl fa = fa + ... | inr nfa = raise-ex-falso _ (fdna (map-raise ∘ nfa)) + soundness _ (ax-n _ _) x fab fa y r = fab y r (fa y r) + soundness CM (mp dab da) x = soundness CM dab x (soundness CM da x) + soundness CM (nec d) x y _ = soundness CM d y + + module _ + ((M , w-is-finite) : finite-model w l2 i l4) + (dec-val : ∀ n x → is-decidable (type-Prop (model-valuate M n x))) + (dec-r : ∀ x y → is-decidable (model-relation M x y)) + where + + finite-model-is-decidable : is-decidable-model M + finite-model-is-decidable (var n) x = + is-decidable-raise _ _ (dec-val n x) + finite-model-is-decidable ⊥ x = + inr map-inv-raise + finite-model-is-decidable (a ⇒ b) x = + is-decidable-function-type + ( finite-model-is-decidable a x) + ( finite-model-is-decidable b x) + finite-model-is-decidable (□ a) x = + is-decidable-Π-is-finite w-is-finite + ( λ y → + is-decidable-function-type (dec-r x y) + ( finite-model-is-decidable a y)) + + finite-soundness : {a : formula i} → ⊢ a → type-Prop (M ⊨M a) + finite-soundness = soundness (M , finite-model-is-decidable) + +-- TODO: move to foundations +double-negation-LEM : + {l : Level} → + ((P : Prop l) → ¬¬ (type-Prop P) → (type-Prop P)) → + LEM l +double-negation-LEM dnP P = + dnP (is-decidable-Prop P) (λ ndec → ndec (inr (ndec ∘ inl))) + +-- TODO: move to raising +compute-raise-function : + (l : Level) {l1 l2 : Level} (A : UU l1) (B : UU l2) → + (A → B) ≃ (raise l A → raise l B) +compute-raise-function _ _ _ = + equiv-function-type _ (compute-raise _ _) (compute-raise _ _) + +raise-double-negation : + (l : Level) {l1 : Level} {A : UU l1} → + ¬¬ A → + (raise l A → raise-empty l) → raise-empty l +raise-double-negation l = + map-equiv-function-type _ (compute-raise-function _ _ _) (compute-raise _ _) +-- 2 +-- map-equiv-function-type _ +-- ( equiv-function-type _ (compute-raise _ _) (compute-raise _ _)) +-- ( compute-raise _ _) +-- 3 +-- raise-double-negation dnA rnA = +-- map-raise (dnA (map-inv-raise ∘ rnA ∘ map-raise)) + +full-soundness : UUω +full-soundness = + {l1 l2 l3 l4 : Level} + (W : Inhabited-Type l1) + (i : Set l3) + (M : kripke-model W l2 i l4) + {a : formula i} → + ⊢ a → + type-Prop (M ⊨M a) + +full-soundness-required-LEM : full-soundness → (l : Level) → LEM l +full-soundness-required-LEM sound l = + double-negation-LEM required-double-negation + where + required-double-negation : (P : Prop l) → ¬¬ (type-Prop P) → type-Prop P + required-double-negation P dnP = + map-inv-raise + ( sound + ( unit , unit-trunc-Prop star) + unit-Set + ( model (frame (λ _ _ → empty-Prop)) (λ _ _ → P)) + ( ax-dn (var star)) + star + ( raise-double-negation _ dnP)) +``` diff --git a/src/modal-logic/K-syntax.lagda.md b/src/modal-logic/K-syntax.lagda.md new file mode 100644 index 0000000000..4706f195ed --- /dev/null +++ b/src/modal-logic/K-syntax.lagda.md @@ -0,0 +1,34 @@ +# Modal logic K syntax + +```agda +module modal-logic.K-syntax where +``` + +
Imports + +```agda +open import foundation.sets +open import foundation.universe-levels + +open import modal-logic.formulas +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +infix 5 ⊢_ + +data ⊢_ {l : Level} {i : Set l} : formula i → UU l where + ax-k : (a b : formula i) → ⊢ a ⇒ b ⇒ a + ax-s : (a b c : formula i) → ⊢ (a ⇒ b ⇒ c) ⇒ (a ⇒ b) ⇒ (a ⇒ c) + ax-dn : (a : formula i) → ⊢ ~~ a ⇒ a + ax-n : (a b : formula i) → ⊢ □ (a ⇒ b) ⇒ □ a ⇒ □ b + mp : {a b : formula i} → ⊢ a ⇒ b → ⊢ a → ⊢ b + nec : {a : formula i} → ⊢ a → ⊢ □ a +``` diff --git a/src/modal-logic/formulas.lagda.md b/src/modal-logic/formulas.lagda.md new file mode 100644 index 0000000000..6d7f1571ae --- /dev/null +++ b/src/modal-logic/formulas.lagda.md @@ -0,0 +1,63 @@ +# Modal logic formulas + +```agda +module modal-logic.formulas where +``` + +
Imports + +```agda +open import foundation.sets +open import foundation.universe-levels +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l : Level} (i : Set l) + where + + infixr 6 _⇒_ + infixr 25 □_ + + data formula : UU l where + var : type-Set i → formula + ⊥ : formula + _⇒_ : formula → formula → formula + □_ : formula → formula + +module _ + {l : Level} {i : Set l} + where + + infixr 25 ~_ + infixr 25 ~~_ + infixl 10 _∨_ + infixl 15 _∧_ + infixr 25 ◇_ + + ~_ : formula i → formula i + ~ a = a ⇒ ⊥ + + ~~_ : formula i → formula i + ~~ a = ~ ~ a + + _∨_ : formula i → formula i → formula i + a ∨ b = ~ a ⇒ b + + _∧_ : formula i → formula i → formula i + a ∧ b = ~ (a ⇒ ~ b) + + ⊤ : formula i + ⊤ = ~ ⊥ + + ◇_ : formula i → formula i + ◇ a = ~ □ ~ a +``` diff --git a/src/modal-logic/kripke-semantics.lagda.md b/src/modal-logic/kripke-semantics.lagda.md new file mode 100644 index 0000000000..7ee03a4553 --- /dev/null +++ b/src/modal-logic/kripke-semantics.lagda.md @@ -0,0 +1,123 @@ +# Krikpe semantics + +```agda +module modal-logic.kripke-semantics where +``` + +
Imports + +```agda +open import foundation.binary-relations +open import foundation.cartesian-product-types +open import foundation.decidable-propositions +open import foundation.dependent-pair-types +open import foundation.empty-types +open import foundation.function-types +open import foundation.inhabited-types +open import foundation.negation +open import foundation.propositions +open import foundation.raising-universe-levels +open import foundation.sets +open import foundation.universe-levels + +open import modal-logic.formulas + +open import univalent-combinatorics.finite-types +``` + +
+ +## Idea + +TODO + +## Definition + +### Semantics + +```agda +module _ + {l1 : Level} (w : Inhabited-Type l1) (l2 : Level) + where + + record kripke-frame : UU (l1 ⊔ lsuc l2) where + constructor frame + field + frame-relation : Relation-Prop l2 (type-Inhabited-Type w) + open kripke-frame public + +module _ + {l1 : Level} (w : Inhabited-Type l1) + (l2 : Level) + {l3 : Level} (i : Set l3) + (l4 : Level) + where + + record kripke-model : UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4) where + constructor model + field + model-frame : kripke-frame w l2 + model-valuate : type-Set i → type-Inhabited-Type w → Prop l4 + open kripke-model public + + finite-model : UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4) + finite-model = kripke-model × is-finite (type-Inhabited-Type w) + +module _ + {l1 l2 l3 l4 : Level} {w : Inhabited-Type l1} {i : Set l3} + where + + model-relation : kripke-model w l2 i l4 → Relation l2 (type-Inhabited-Type w) + model-relation = type-Relation-Prop ∘ frame-relation ∘ model-frame + + private + l = l1 ⊔ l2 ⊔ l4 + + infix 5 _⊨_ + infix 5 _⊭_ + infix 5 _⊨M_ + infix 5 _⊭M_ + + _⊨_ : kripke-model w l2 i l4 × (type-Inhabited-Type w) → formula i → Prop l + (M , x) ⊨ var n = raise-Prop l (model-valuate M n x) + (M , x) ⊨ ⊥ = raise-empty-Prop l + (M , x) ⊨ a ⇒ b = implication-Prop ((M , x) ⊨ a) ((M , x) ⊨ b) + (M , x) ⊨ □ a = + Π-Prop + ( type-Inhabited-Type w) + ( λ y -> function-Prop (model-relation M x y) ((M , y) ⊨ a)) + + _⊭_ : kripke-model w l2 i l4 × (type-Inhabited-Type w) → formula i → Prop l + (M , x) ⊭ a = neg-Prop ((M , x) ⊨ a) + + _⊨M_ : kripke-model w l2 i l4 → formula i → Prop l + M ⊨M a = Π-Prop (type-Inhabited-Type w) (λ x → (M , x) ⊨ a) + + _⊭M_ : kripke-model w l2 i l4 → formula i → Prop l + M ⊭M a = neg-Prop (M ⊨M a) + +module _ + {l1 l2 l3 l4 : Level} {w : Inhabited-Type l1} {i : Set l3} + where + + private + l = l1 ⊔ l2 ⊔ l3 ⊔ l4 + + is-decidable-model-Prop : kripke-model w l2 i l4 → Prop l + is-decidable-model-Prop M = + Π-Prop + ( formula i) + ( λ a → + Π-Prop (type-Inhabited-Type w) (λ x → is-decidable-Prop ((M , x) ⊨ a))) + + is-decidable-model : kripke-model w l2 i l4 → UU l + is-decidable-model = type-Prop ∘ is-decidable-model-Prop + +decidable-model : + {l1 : Level} (w : Inhabited-Type l1) + (l2 : Level) + {l3 : Level} (i : Set l3) + (l4 : Level) → + UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4) +decidable-model w l2 i l4 = Σ (kripke-model w l2 i l4) is-decidable-model +``` diff --git a/src/modal-logic/modal-logic-syntax.lagda.md b/src/modal-logic/modal-logic-syntax.lagda.md deleted file mode 100644 index 157d59d54a..0000000000 --- a/src/modal-logic/modal-logic-syntax.lagda.md +++ /dev/null @@ -1,258 +0,0 @@ -# Modal logic syntax - -```agda -module modal-logic.modal-logic-syntax where -``` - -
Imports - -```agda -open import foundation.binary-relations -open import foundation.cartesian-product-types -open import foundation.coproduct-types -open import foundation.decidable-propositions -open import foundation.decidable-types -open import foundation.dependent-pair-types -open import foundation.double-negation -open import foundation.empty-types -open import foundation.function-types -open import foundation.identity-types -open import foundation.inhabited-types -open import foundation.law-of-excluded-middle -open import foundation.negation -open import foundation.propositional-truncations -open import foundation.propositions -open import foundation.raising-universe-levels -open import foundation.unit-type -open import foundation.universe-levels - -open import univalent-combinatorics.decidable-dependent-function-types -open import univalent-combinatorics.finite-types -``` - -
- -## Idea - -TODO - -## Definition - -### Syntax - -```agda -module _ - {l : Level} (i : UU l) - where - - infixr 6 _⇒_ - infixr 25 □_ - - data formula : UU l where - var : i → formula - ⊥ : formula - _⇒_ : formula → formula → formula - □_ : formula → formula - -module _ - {l : Level} {i : UU l} - where - - infixr 25 ~_ - infixr 25 ~~_ - infixl 10 _∨_ - infixl 15 _∧_ - infixr 25 ◇_ - infix 5 ⊢_ - - ~_ : formula i → formula i - ~ a = a ⇒ ⊥ - - ~~_ : formula i → formula i - ~~ a = ~ ~ a - - _∨_ : formula i → formula i → formula i - a ∨ b = ~ a ⇒ b - - _∧_ : formula i → formula i → formula i - a ∧ b = ~ (a ⇒ ~ b) - - ⊤ : formula i - ⊤ = ~ ⊥ - - ◇_ : formula i → formula i - ◇ a = ~ □ ~ a - - data ⊢_ : formula i → UU l where - ax-k : (a b : formula i) → ⊢ a ⇒ b ⇒ a - ax-s : (a b c : formula i) → ⊢ (a ⇒ b ⇒ c) ⇒ (a ⇒ b) ⇒ (a ⇒ c) - ax-dn : (a : formula i) → ⊢ ~~ a ⇒ a - ax-n : (a b : formula i) → ⊢ □ (a ⇒ b) ⇒ □ a ⇒ □ b - mp : {a b : formula i} → ⊢ a ⇒ b → ⊢ a → ⊢ b - nec : {a : formula i} → ⊢ a → ⊢ □ a -``` - -### Semantics - -```agda -module _ - {l1 : Level} ((w , _) : Inhabited-Type l1) (l2 : Level) - where - - record kripke-frame : UU (l1 ⊔ lsuc l2) where - constructor frame - field - frame-relation : Relation l2 w - open kripke-frame public - -module _ - {l1 : Level} (W@(w , _) : Inhabited-Type l1) - (l2 : Level) - {l3 : Level} (i : UU l3) - (l4 : Level) - where - - record kripke-model : UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4) where - constructor model - field - model-frame : kripke-frame W l2 - model-valuate : i → w → Prop l4 - open kripke-model public - - finite-model : UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4) - finite-model = kripke-model × is-finite w - -module _ - {l1 l2 l3 l4 : Level} {W@(w , _) : Inhabited-Type l1} {i : UU l3} - where - - model-relation : kripke-model W l2 i l4 → Relation l2 w - model-relation = frame-relation ∘ model-frame - - private - l = l1 ⊔ l2 ⊔ l4 - - infix 5 _⊨_ - infix 5 _⊭_ - infix 5 _⊨M_ - infix 5 _⊭M_ - - _⊨_ : kripke-model W l2 i l4 × w → formula i → Prop l - (M , x) ⊨ var n = raise-Prop l (model-valuate M n x) - (M , x) ⊨ ⊥ = raise-empty-Prop l - (M , x) ⊨ a ⇒ b = implication-Prop ((M , x) ⊨ a) ((M , x) ⊨ b) - (M , x) ⊨ □ a = - Π-Prop w (λ y -> function-Prop (model-relation M x y) ((M , y) ⊨ a)) - - _⊭_ : kripke-model W l2 i l4 × w → formula i → Prop l - (M , x) ⊭ a = neg-Prop ((M , x) ⊨ a) - - _⊨M_ : kripke-model W l2 i l4 → formula i → Prop l - M ⊨M a = Π-Prop w (λ x → (M , x) ⊨ a) - - _⊭M_ : kripke-model W l2 i l4 → formula i → Prop l - M ⊭M a = neg-Prop (M ⊨M a) -``` - -### Soundness - -```agda -module _ - {l1 l2 l3 l4 : Level} {W@(w , _) : Inhabited-Type l1} {i : UU l3} - where - - private - l = l1 ⊔ l2 ⊔ l3 ⊔ l4 - - is-classical-Prop : kripke-model W l2 i l4 → Prop l - is-classical-Prop M = - Π-Prop (formula i) (λ a → Π-Prop w (λ x → is-decidable-Prop ((M , x) ⊨ a))) - - is-classical : kripke-model W l2 i l4 → UU l - is-classical = type-Prop ∘ is-classical-Prop - - classical-model : UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4) - classical-model = Σ (kripke-model W l2 i l4) is-classical - - soundness : - ((M , _) : classical-model) - {a : formula i} → - ⊢ a → - type-Prop (M ⊨M a) - soundness _ (ax-k _ _) x fa _ = fa - soundness _ (ax-s _ _ _) x fabc fab fa = fabc fa (fab fa) - soundness (_ , cl) (ax-dn a) x fdna with cl a x - ... | inl fa = fa - ... | inr nfa = ex-falso (map-inv-raise (fdna (λ fa → map-raise (nfa fa)))) - soundness _ (ax-n _ _) x fab fa y r = fab y r (fa y r) - soundness CM (mp dab da) x = soundness CM dab x (soundness CM da x) - soundness CM (nec d) x y _ = soundness CM d y - - module _ - ((M , w-is-finite) : finite-model W l2 i l4) - (dec-val : ∀ n x → type-Prop (is-decidable-Prop (model-valuate M n x))) - (dec-r : ∀ x y → is-decidable (model-relation M x y)) - where - - finite-is-classical : is-classical M - finite-is-classical (var n) x = - is-decidable-raise _ _ (dec-val n x) - finite-is-classical ⊥ x = - inr map-inv-raise - finite-is-classical (a ⇒ b) x = - is-decidable-function-type - ( finite-is-classical a x) - ( finite-is-classical b x) - finite-is-classical (□ a) x = - is-decidable-Π-is-finite w-is-finite - ( λ y → - is-decidable-function-type (dec-r x y) - ( finite-is-classical a y)) - - finite-soundness : {a : formula i} → ⊢ a → type-Prop (M ⊨M a) - finite-soundness = soundness (M , finite-is-classical) - -double-negation-LEM : - {l : Level} → - ((P : Prop l) → - ¬¬ (type-Prop P) → - (type-Prop P)) → - LEM l -double-negation-LEM dnP P = - dnP (is-decidable-Prop P) (λ ndec → ndec (inr (λ p → ndec (inl p)))) - -raise-double-negation : - {l1 l2 : Level} - {A : UU l1} → - ¬¬ A → - (raise l2 A → raise-empty l2) → raise-empty l2 -raise-double-negation dnA rnA = - map-raise (dnA (λ a → map-inv-raise (rnA (map-raise a)))) - -full-soundness : UUω -full-soundness = - {l1 l2 l3 l4 : Level} - (W : Inhabited-Type l1) - {i : UU l3} - (M : kripke-model W l2 i l4) - {a : formula i} → - ⊢ a → - type-Prop (M ⊨M a) - -full-soundness-required-LEM : - full-soundness → - (l : Level) → - LEM l -full-soundness-required-LEM sound l = - double-negation-LEM required-double-negation - where - required-double-negation : (P : Prop l) → ¬¬ (type-Prop P) → type-Prop P - required-double-negation P dnP = - map-inv-raise - ( sound - ( unit , unit-trunc-Prop star) - ( model (frame (λ _ _ → empty)) (λ _ _ → P)) - ( ax-dn (var star)) - star - ( raise-double-negation dnP)) -``` From 21d74ce12ee2986e6402eca7cccf0365ffb47b92 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Thu, 26 Oct 2023 13:44:28 +0300 Subject: [PATCH 16/63] Add logics as sets of formulas --- .../intersections-subtypes.lagda.md | 96 +++++++++- src/modal-logic.lagda.md | 5 + src/modal-logic/axioms.lagda.md | 165 ++++++++++++++++++ src/modal-logic/formulas.lagda.md | 114 +++++++++++- src/modal-logic/kripke-semantics.lagda.md | 49 +++++- src/modal-logic/logic-syntax.lagda.md | 129 ++++++++++++++ src/modal-logic/modal-logic-IK.lagda.md | 60 +++++++ src/modal-logic/modal-logic-K.lagda.md | 50 ++++++ src/modal-logic/soundness.lagda.md | 139 +++++++++++++++ 9 files changed, 799 insertions(+), 8 deletions(-) create mode 100644 src/modal-logic/axioms.lagda.md create mode 100644 src/modal-logic/logic-syntax.lagda.md create mode 100644 src/modal-logic/modal-logic-IK.lagda.md create mode 100644 src/modal-logic/modal-logic-K.lagda.md create mode 100644 src/modal-logic/soundness.lagda.md diff --git a/src/foundation/intersections-subtypes.lagda.md b/src/foundation/intersections-subtypes.lagda.md index 3ff9cc3f3b..3bd50b13f4 100644 --- a/src/foundation/intersections-subtypes.lagda.md +++ b/src/foundation/intersections-subtypes.lagda.md @@ -10,13 +10,14 @@ module foundation.intersections-subtypes where open import foundation.conjunction open import foundation.decidable-subtypes open import foundation.dependent-pair-types +open import foundation.identity-types open import foundation.large-locale-of-subtypes open import foundation.powersets +open import foundation.subtypes open import foundation.universe-levels open import foundation-core.decidable-propositions open import foundation-core.propositions -open import foundation-core.subtypes open import order-theory.greatest-lower-bounds-large-posets ``` @@ -85,3 +86,96 @@ module _ {I : UU l2} (P : I → subtype l3 X) → subtype (l2 ⊔ l3) X intersection-family-of-subtypes {I} P x = Π-Prop I (λ i → P i x) ``` + +## + +```agda +-- It's too simple =) +-- module _ +-- {l1 l2 l3 : Level} {X : UU l1} (P : subtype l2 X) (Q : subtype l3 X) +-- where + +-- is-commutative-subtype-intersection : +-- intersection-subtype P Q ⊆ intersection-subtype Q P +-- is-commutative-subtype-intersection x (in-P , in-Q) = in-Q , in-P + +module _ + {l1 l2 l3 : Level} {X : UU l1} (P : subtype l2 X) (Q : subtype l3 X) + where + + subtype-intersection-left : intersection-subtype P Q ⊆ P + subtype-intersection-left _ = pr1 + + subtype-intersection-right : intersection-subtype P Q ⊆ Q + subtype-intersection-right _ = pr2 + + subtype-both-intersection : + {l4 : Level} (S : subtype l4 X) → + S ⊆ P → S ⊆ Q → S ⊆ intersection-subtype P Q + pr1 (subtype-both-intersection S S-sub-P S-sub-Q x in-S) = S-sub-P x in-S + pr2 (subtype-both-intersection S S-sub-P S-sub-Q x in-S) = S-sub-Q x in-S + + intersection-subtype-left-subtype : P ⊆ Q → P ⊆ intersection-subtype P Q + intersection-subtype-left-subtype P-sub-Q = + subtype-both-intersection P (refl-leq-subtype P) P-sub-Q + + intersection-subtype-right-subtype : Q ⊆ P → Q ⊆ intersection-subtype P Q + intersection-subtype-right-subtype Q-sub-P = + subtype-both-intersection Q Q-sub-P (refl-leq-subtype Q) + +module _ + {l1 l2 : Level} {X : UU l1} (P : subtype l2 X) (Q : subtype l2 X) + where + + intersection-subtype-left : + P ⊆ Q → intersection-subtype P Q = P + intersection-subtype-left P-sub-Q = + antisymmetric-leq-subtype _ _ + ( subtype-intersection-left P Q) + ( intersection-subtype-left-subtype P Q P-sub-Q) + + intersection-subtype-right : + Q ⊆ P → intersection-subtype P Q = Q + intersection-subtype-right Q-sub-P = + antisymmetric-leq-subtype _ _ + ( subtype-intersection-right P Q) + ( intersection-subtype-right-subtype P Q Q-sub-P) + +module _ + {l1 l2 : Level} {X : UU l1} (P : subtype l2 X) + where + + is-reflexivity-intersection : intersection-subtype P P = P + is-reflexivity-intersection = + antisymmetric-leq-subtype _ _ + ( subtype-intersection-left P P) + ( subtype-both-intersection + ( P) + ( P) + ( P) + ( refl-leq-subtype P) + ( refl-leq-subtype P)) + +module _ + {l1 l2 l3 : Level} {X : UU l1} (P : subtype l2 X) (Q : subtype l3 X) + where + + is-commutative-subtype-intersection : + intersection-subtype P Q ⊆ intersection-subtype Q P + is-commutative-subtype-intersection = + subtype-both-intersection Q P + ( intersection-subtype P Q) + ( subtype-intersection-right P Q) + ( subtype-intersection-left P Q) + +module _ + {l1 l2 l3 : Level} {X : UU l1} (P : subtype l2 X) (Q : subtype l3 X) + where + + is-commutative-intersection : + intersection-subtype P Q = intersection-subtype Q P + is-commutative-intersection = + antisymmetric-leq-subtype _ _ + ( is-commutative-subtype-intersection P Q) + ( is-commutative-subtype-intersection Q P) +``` diff --git a/src/modal-logic.lagda.md b/src/modal-logic.lagda.md index 4592c16d0c..7e7973fde8 100644 --- a/src/modal-logic.lagda.md +++ b/src/modal-logic.lagda.md @@ -5,6 +5,11 @@ module modal-logic where open import modal-logic.K-soundness public open import modal-logic.K-syntax public +open import modal-logic.axioms public open import modal-logic.formulas public open import modal-logic.kripke-semantics public +open import modal-logic.logic-syntax public +open import modal-logic.modal-logic-IK public +open import modal-logic.modal-logic-K public +open import modal-logic.soundness public ``` diff --git a/src/modal-logic/axioms.lagda.md b/src/modal-logic/axioms.lagda.md new file mode 100644 index 0000000000..18040eb947 --- /dev/null +++ b/src/modal-logic/axioms.lagda.md @@ -0,0 +1,165 @@ +# Modal logic axioms + +```agda +module modal-logic.axioms where +``` + +
Imports + +```agda +open import foundation.conjunction +open import foundation.coproduct-types +open import foundation.dependent-pair-types +open import foundation.empty-types +open import foundation.function-types +open import foundation.identity-types +open import foundation.inhabited-types +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.raising-universe-levels +open import foundation.sets +open import foundation.subtypes +open import foundation.unit-type +open import foundation.universe-levels + +open import modal-logic.formulas +open import modal-logic.kripke-semantics +open import modal-logic.logic-syntax +open import modal-logic.soundness +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l : Level} (i : Set l) + where + + ax-k : formulas l i + ax-k (a₁ ⇒ b ⇒ a₂) = Id-formula-Prop a₁ a₂ + {-# CATCHALL #-} + ax-k _ = raise-empty-Prop l + + ax-s : formulas l i + ax-s ((a₁ ⇒ b₁ ⇒ c₁) ⇒ (a₂ ⇒ b₂) ⇒ a₃ ⇒ c₂) = + Id-formula-Prop a₁ a₂ ∧ + Id-formula-Prop a₂ a₃ ∧ + Id-formula-Prop b₁ b₂ ∧ + Id-formula-Prop c₁ c₂ + {-# CATCHALL #-} + ax-s _ = raise-empty-Prop l + + ax-n : formulas l i + ax-n (□ (a₁ ⇒ b₁) ⇒ □ a₂ ⇒ □ b₂) = + Id-formula-Prop a₁ a₂ ∧ + Id-formula-Prop b₁ b₂ + {-# CATCHALL #-} + ax-n _ = raise-empty-Prop l + + ax-dn : formulas l i + ax-dn (((a₁ ⇒ ⊥) ⇒ ⊥) ⇒ a₂) = Id-formula-Prop a₁ a₂ + {-# CATCHALL #-} + ax-dn _ = raise-empty-Prop l + +module _ + {l1 l2 : Level} + (i : Set l1) + (w : Inhabited-Type l2) + (l3 l4 : Level) + where + + ax-k-soundness : + (a : formula i) → + is-in-subtype (ax-k i) a → + type-Prop (all-models w l3 i l4 ⊨C a) + ax-k-soundness (a₁ ⇒ b ⇒ a₁) refl M in-class x fa _ = fa + ax-k-soundness (var _) (map-raise ()) + ax-k-soundness ⊥ (map-raise ()) + ax-k-soundness (□ _) (map-raise ()) + ax-k-soundness (_ ⇒ var _) (map-raise ()) + ax-k-soundness (_ ⇒ ⊥) (map-raise ()) + ax-k-soundness (_ ⇒ □ _) (map-raise ()) + + ax-n-soundness : + (a : formula i) → + is-in-subtype (ax-n i) a → + type-Prop (all-models w l3 i l4 ⊨C a) + ax-n-soundness + (□ (a₁ ⇒ b₁) ⇒ □ a₁ ⇒ □ b₁) + (refl , refl) + M in-class x fab fa y r = fab y r (fa y r) + ax-n-soundness (var _) (map-raise ()) + ax-n-soundness ⊥ (map-raise ()) + ax-n-soundness (□ _) (map-raise ()) + ax-n-soundness (var _ ⇒ _) (map-raise ()) + ax-n-soundness (⊥ ⇒ _) (map-raise ()) + ax-n-soundness ((_ ⇒ _) ⇒ _) (map-raise ()) + ax-n-soundness (□ var _ ⇒ _) (map-raise ()) + ax-n-soundness (□ ⊥ ⇒ _) (map-raise ()) + ax-n-soundness (□ □ _ ⇒ _) (map-raise ()) + ax-n-soundness (□ (_ ⇒ _) ⇒ var _) (map-raise ()) + ax-n-soundness (□ (_ ⇒ _) ⇒ ⊥) (map-raise ()) + ax-n-soundness (□ (_ ⇒ _) ⇒ □ _) (map-raise ()) + ax-n-soundness (□ (_ ⇒ _) ⇒ var _ ⇒ _) (map-raise ()) + ax-n-soundness (□ (_ ⇒ _) ⇒ ⊥ ⇒ _) (map-raise ()) + ax-n-soundness (□ (_ ⇒ _) ⇒ (_ ⇒ _) ⇒ _) (map-raise ()) + ax-n-soundness (□ (_ ⇒ _) ⇒ □ _ ⇒ var _) (map-raise ()) + ax-n-soundness (□ (_ ⇒ _) ⇒ □ _ ⇒ ⊥) (map-raise ()) + ax-n-soundness (□ (_ ⇒ _) ⇒ □ _ ⇒ _ ⇒ _) (map-raise ()) + + ax-s-soundness : + (a : formula i) → + is-in-subtype (ax-s i) a → + type-Prop (all-models w l3 i l4 ⊨C a) + ax-s-soundness + ((a ⇒ b ⇒ c) ⇒ (a ⇒ b) ⇒ a ⇒ c) + (refl , refl , refl , refl) + M in-class x fabc fab fa = fabc fa (fab fa) + ax-s-soundness (var _) (map-raise ()) + ax-s-soundness ⊥ (map-raise ()) + ax-s-soundness (□ _) (map-raise ()) + ax-s-soundness (var _ ⇒ _) (map-raise ()) + ax-s-soundness (⊥ ⇒ _) (map-raise ()) + ax-s-soundness (□ _ ⇒ _) (map-raise ()) + ax-s-soundness ((_ ⇒ var _) ⇒ _) (map-raise ()) + ax-s-soundness ((_ ⇒ ⊥) ⇒ _) (map-raise ()) + ax-s-soundness ((_ ⇒ □ _) ⇒ _) (map-raise ()) + ax-s-soundness ((_ ⇒ _ ⇒ _) ⇒ var _) (map-raise ()) + ax-s-soundness ((_ ⇒ _ ⇒ _) ⇒ ⊥) (map-raise ()) + ax-s-soundness ((_ ⇒ _ ⇒ _) ⇒ □ _) (map-raise ()) + ax-s-soundness ((_ ⇒ _ ⇒ _) ⇒ var _ ⇒ _) (map-raise ()) + ax-s-soundness ((_ ⇒ _ ⇒ _) ⇒ ⊥ ⇒ _) (map-raise ()) + ax-s-soundness ((_ ⇒ _ ⇒ _) ⇒ □ _ ⇒ _) (map-raise ()) + ax-s-soundness ((_ ⇒ _ ⇒ _) ⇒ (_ ⇒ _) ⇒ var _) (map-raise ()) + ax-s-soundness ((_ ⇒ _ ⇒ _) ⇒ (_ ⇒ _) ⇒ ⊥) (map-raise ()) + ax-s-soundness ((_ ⇒ _ ⇒ _) ⇒ (_ ⇒ _) ⇒ □ _) (map-raise ()) + + ax-dn-soundness : + (a : formula i) → + is-in-subtype (ax-dn i) a → + type-Prop (decidable-models w l3 i l4 ⊨C a) + ax-dn-soundness (((a ⇒ ⊥) ⇒ ⊥) ⇒ a) refl M is-dec x f with is-dec a x + ... | inl fa = fa + ... | inr fna = raise-ex-falso _ (f (λ fa → map-raise (fna fa))) + ax-dn-soundness (var _) (map-raise ()) + ax-dn-soundness ⊥ (map-raise ()) + ax-dn-soundness (□ _) (map-raise ()) + ax-dn-soundness (var _ ⇒ _) (map-raise ()) + ax-dn-soundness (⊥ ⇒ _) (map-raise ()) + ax-dn-soundness (□ _ ⇒ _) (map-raise ()) + ax-dn-soundness ((var _ ⇒ _) ⇒ _) (map-raise ()) + ax-dn-soundness ((⊥ ⇒ _) ⇒ _) (map-raise ()) + ax-dn-soundness ((□ _ ⇒ _) ⇒ _) (map-raise ()) + ax-dn-soundness (((_ ⇒ var _) ⇒ _) ⇒ _) (map-raise ()) + ax-dn-soundness (((_ ⇒ _ ⇒ _) ⇒ _) ⇒ _) (map-raise ()) + ax-dn-soundness (((_ ⇒ □ _) ⇒ _) ⇒ _) (map-raise ()) + ax-dn-soundness (((_ ⇒ ⊥) ⇒ var _) ⇒ _) (map-raise ()) + ax-dn-soundness (((_ ⇒ ⊥) ⇒ _ ⇒ _) ⇒ _) (map-raise ()) + ax-dn-soundness (((_ ⇒ ⊥) ⇒ □ _) ⇒ _) (map-raise ()) +``` diff --git a/src/modal-logic/formulas.lagda.md b/src/modal-logic/formulas.lagda.md index 6d7f1571ae..30c8771c53 100644 --- a/src/modal-logic/formulas.lagda.md +++ b/src/modal-logic/formulas.lagda.md @@ -7,7 +7,19 @@ module modal-logic.formulas where
Imports ```agda +open import foundation.action-on-identifications-functions +open import foundation.cartesian-product-types +open import foundation.contractible-types +open import foundation.dependent-pair-types +open import foundation.empty-types +open import foundation.equality-cartesian-product-types +open import foundation.equivalences +open import foundation.identity-types +open import foundation.propositions +open import foundation.raising-universe-levels open import foundation.sets +open import foundation.transport-along-identifications +open import foundation.unit-type open import foundation.universe-levels ``` @@ -39,8 +51,8 @@ module _ infixr 25 ~_ infixr 25 ~~_ - infixl 10 _∨_ - infixl 15 _∧_ + -- infixl 10 _∨_ + -- infixl 15 _∧_ infixr 25 ◇_ ~_ : formula i → formula i @@ -49,11 +61,11 @@ module _ ~~_ : formula i → formula i ~~ a = ~ ~ a - _∨_ : formula i → formula i → formula i - a ∨ b = ~ a ⇒ b + -- _∨_ : formula i → formula i → formula i + -- a ∨ b = ~ a ⇒ b - _∧_ : formula i → formula i → formula i - a ∧ b = ~ (a ⇒ ~ b) + -- _∧_ : formula i → formula i → formula i + -- a ∧ b = ~ (a ⇒ ~ b) ⊤ : formula i ⊤ = ~ ⊥ @@ -61,3 +73,93 @@ module _ ◇_ : formula i → formula i ◇ a = ~ □ ~ a ``` + +## Properties + +### The formulas are a set + +```agda +module _ + {l : Level} (i : Set l) + where + + Eq-formula : formula i → formula i → UU l + Eq-formula (var n) (var m) = n = m + Eq-formula (var _) ⊥ = raise-empty l + Eq-formula (var _) (_ ⇒ _) = raise-empty l + Eq-formula (var -) (□ _) = raise-empty l + Eq-formula ⊥ (var _) = raise-empty l + Eq-formula ⊥ ⊥ = raise-unit l + Eq-formula ⊥ (_ ⇒ _) = raise-empty l + Eq-formula ⊥ (□ _) = raise-empty l + Eq-formula (_ ⇒ _) (var _) = raise-empty l + Eq-formula (_ ⇒ _) ⊥ = raise-empty l + Eq-formula (a ⇒ b) (c ⇒ d) = (Eq-formula a c) × (Eq-formula b d) + Eq-formula (_ ⇒ _) (□ _) = raise-empty l + Eq-formula (□ _) (var _) = raise-empty l + Eq-formula (□ _) ⊥ = raise-empty l + Eq-formula (□ _) (_ ⇒ _) = raise-empty l + Eq-formula (□ a) (□ c) = Eq-formula a c + + refl-Eq-formula : (a : formula i) → Eq-formula a a + refl-Eq-formula (var n) = refl + refl-Eq-formula ⊥ = raise-star + refl-Eq-formula (a ⇒ b) = (refl-Eq-formula a) , (refl-Eq-formula b) + refl-Eq-formula (□ a) = refl-Eq-formula a + + Eq-eq-formula : {a b : formula i} → a = b → Eq-formula a b + Eq-eq-formula {a} refl = refl-Eq-formula a + + eq-Eq-formula : {a b : formula i} → Eq-formula a b → a = b + eq-Eq-formula {var _} {var _} refl = refl + eq-Eq-formula {var _} {⊥} (map-raise ()) + eq-Eq-formula {var _} {_ ⇒ _} (map-raise ()) + eq-Eq-formula {var _} {□ _} (map-raise ()) + eq-Eq-formula {⊥} {var _} (map-raise ()) + eq-Eq-formula {⊥} {⊥} _ = refl + eq-Eq-formula {⊥} {_ ⇒ _} (map-raise ()) + eq-Eq-formula {⊥} {□ _} (map-raise ()) + eq-Eq-formula {_ ⇒ _} {var _} (map-raise ()) + eq-Eq-formula {_ ⇒ _} {⊥} (map-raise ()) + eq-Eq-formula {a ⇒ b} {c ⇒ d} (eq1 , eq2) = + ap (λ (x , y) → x ⇒ y) (eq-pair (eq-Eq-formula eq1) (eq-Eq-formula eq2)) + eq-Eq-formula {_ ⇒ _} {□ _} (map-raise ()) + eq-Eq-formula {□ _} {var _} (map-raise ()) + eq-Eq-formula {□ _} {⊥} (map-raise ()) + eq-Eq-formula {□ _} {_ ⇒ _} (map-raise ()) + eq-Eq-formula {□ _} {□ _} eq = ap □_ (eq-Eq-formula eq) + + is-prop-Eq-formula : (a b : formula i) → is-prop (Eq-formula a b) + is-prop-Eq-formula (var n) (var m) = is-prop-type-Prop (Id-Prop i n m) + is-prop-Eq-formula (var _) ⊥ = is-prop-raise-empty + is-prop-Eq-formula (var _) (_ ⇒ _) = is-prop-raise-empty + is-prop-Eq-formula (var -) (□ _) = is-prop-raise-empty + is-prop-Eq-formula ⊥ (var _) = is-prop-raise-empty + is-prop-Eq-formula ⊥ ⊥ = is-prop-raise-unit + is-prop-Eq-formula ⊥ (_ ⇒ _) = is-prop-raise-empty + is-prop-Eq-formula ⊥ (□ _) = is-prop-raise-empty + is-prop-Eq-formula (_ ⇒ _) (var _) = is-prop-raise-empty + is-prop-Eq-formula (_ ⇒ _) ⊥ = is-prop-raise-empty + is-prop-Eq-formula (a ⇒ b) (c ⇒ d) = + is-prop-prod (is-prop-Eq-formula a c) (is-prop-Eq-formula b d) + is-prop-Eq-formula (_ ⇒ _) (□ _) = is-prop-raise-empty + is-prop-Eq-formula (□ _) (var _) = is-prop-raise-empty + is-prop-Eq-formula (□ _) ⊥ = is-prop-raise-empty + is-prop-Eq-formula (□ _) (_ ⇒ _) = is-prop-raise-empty + is-prop-Eq-formula (□ a) (□ c) = is-prop-Eq-formula a c + + is-set-formula : is-set (formula i) + is-set-formula = + is-set-prop-in-id + ( Eq-formula) + ( is-prop-Eq-formula) + ( refl-Eq-formula) + ( λ _ _ → eq-Eq-formula) + + formula-Set : Set l + pr1 formula-Set = formula i + pr2 formula-Set = is-set-formula + +Id-formula-Prop : {l : Level} {i : Set l} → formula i → formula i → Prop l +Id-formula-Prop = Id-Prop (formula-Set _) +``` diff --git a/src/modal-logic/kripke-semantics.lagda.md b/src/modal-logic/kripke-semantics.lagda.md index 7ee03a4553..d22d97e942 100644 --- a/src/modal-logic/kripke-semantics.lagda.md +++ b/src/modal-logic/kripke-semantics.lagda.md @@ -10,6 +10,7 @@ module modal-logic.kripke-semantics where open import foundation.binary-relations open import foundation.cartesian-product-types open import foundation.decidable-propositions +open import foundation.decidable-types open import foundation.dependent-pair-types open import foundation.empty-types open import foundation.function-types @@ -18,9 +19,12 @@ open import foundation.negation open import foundation.propositions open import foundation.raising-universe-levels open import foundation.sets +open import foundation.subtypes +open import foundation.unit-type open import foundation.universe-levels open import modal-logic.formulas +open import modal-logic.logic-syntax open import univalent-combinatorics.finite-types ``` @@ -63,6 +67,12 @@ module _ finite-model : UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4) finite-model = kripke-model × is-finite (type-Inhabited-Type w) + model-class : (l5 : Level) → UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ lsuc l5) + model-class l5 = subtype l5 kripke-model + + all-models : model-class lzero + all-models _ = unit-Prop + module _ {l1 l2 l3 l4 : Level} {w : Inhabited-Type l1} {i : Set l3} where @@ -85,7 +95,7 @@ module _ (M , x) ⊨ □ a = Π-Prop ( type-Inhabited-Type w) - ( λ y -> function-Prop (model-relation M x y) ((M , y) ⊨ a)) + ( λ y → function-Prop (model-relation M x y) ((M , y) ⊨ a)) _⊭_ : kripke-model w l2 i l4 × (type-Inhabited-Type w) → formula i → Prop l (M , x) ⊭ a = neg-Prop ((M , x) ⊨ a) @@ -96,6 +106,43 @@ module _ _⊭M_ : kripke-model w l2 i l4 → formula i → Prop l M ⊭M a = neg-Prop (M ⊨M a) + decidable-class : model-class w l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4) + decidable-class M = Π-Prop (formula i) (λ a → M ⊨M a) + +module _ + {l1 l2 l3 l4 l5 : Level} {w : Inhabited-Type l1} {i : Set l3} + where + + _⊨C_ : + model-class w l2 i l4 l5 → + formula i → + Prop (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ l5) + C ⊨C a = + Π-Prop + ( kripke-model w l2 i l4) + ( λ M → function-Prop (is-in-subtype C M) (M ⊨M a)) + + class-modal-logic : + model-class w l2 i l4 l5 → + formulas (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ l5) i + class-modal-logic C a = C ⊨C a + +module _ + {l1 : Level} (w : Inhabited-Type l1) + (l2 : Level) + {l3 : Level} (i : Set l3) + (l4 : Level) + where + + decidable-models : model-class w l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4) + decidable-models M = + Π-Prop + ( formula i) + ( λ a → + ( Π-Prop + ( type-Inhabited-Type w) + ( λ x → is-decidable-Prop ((M , x) ⊨ a)))) + module _ {l1 l2 l3 l4 : Level} {w : Inhabited-Type l1} {i : Set l3} where diff --git a/src/modal-logic/logic-syntax.lagda.md b/src/modal-logic/logic-syntax.lagda.md new file mode 100644 index 0000000000..35f9733276 --- /dev/null +++ b/src/modal-logic/logic-syntax.lagda.md @@ -0,0 +1,129 @@ +# Modal logic syntax + +```agda +module modal-logic.logic-syntax where +``` + +
Imports + +```agda +open import foundation.conjunction +open import foundation.dependent-pair-types +open import foundation.empty-types +open import foundation.function-types +open import foundation.identity-types +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.sets +open import foundation.subtypes +open import foundation.unit-type +open import foundation.universe-levels + +open import modal-logic.formulas +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l1 : Level} (l2 : Level) (i : Set l1) + where + + formulas : UU (l1 ⊔ lsuc l2) + formulas = subtype l2 (formula i) + +module _ + {l1 l2 : Level} {i : Set l1} + where + + infix 5 _⊢_ + + data _⊢_ (axioms : formulas l2 i) : formula i → UU (l1 ⊔ l2) where + ax : {a : formula i} → is-in-subtype axioms a → axioms ⊢ a + mp : {a b : formula i} → axioms ⊢ a ⇒ b → axioms ⊢ a → axioms ⊢ b + nec : {a : formula i} → axioms ⊢ a → axioms ⊢ □ a + + modal-logic : formulas l2 i → formulas (l1 ⊔ l2) i + modal-logic axioms a = trunc-Prop (axioms ⊢ a) + +module _ + {l1 : Level} {i : Set l1} + where + + axioms-subset-modal-logic : + {l2 : Level} (axioms : formulas l2 i) → axioms ⊆ modal-logic axioms + axioms-subset-modal-logic _ a H = unit-trunc-Prop (ax H) + + modal-logic-closed : + {l2 : Level} {axioms : formulas l2 i} {a : formula i} → + modal-logic axioms ⊢ a → + type-Prop (modal-logic axioms a) + modal-logic-closed (ax x) = x + modal-logic-closed {axioms = axioms} {a} (mp tdab tda) = + apply-twice-universal-property-trunc-Prop + ( modal-logic-closed tdab) + ( modal-logic-closed tda) + ( modal-logic axioms a) + (λ dab da → unit-trunc-Prop (mp dab da)) + modal-logic-closed {axioms = axioms} {a} (nec d) = + apply-universal-property-trunc-Prop + ( modal-logic-closed d) + ( modal-logic axioms a) + ( unit-trunc-Prop ∘ nec) + + subset-double-modal-logic : + {l2 : Level} + (axioms : formulas l2 i) → + modal-logic (modal-logic axioms) ⊆ modal-logic axioms + subset-double-modal-logic axioms a = + map-universal-property-trunc-Prop + ( modal-logic axioms a) + ( modal-logic-closed) + + modal-logic-idempotent : + {l2 : Level} + (axioms : formulas l2 i) → + modal-logic axioms = modal-logic (modal-logic axioms) + modal-logic-idempotent axioms = + antisymmetric-leq-subtype _ _ + ( axioms-subset-modal-logic (modal-logic axioms)) + ( subset-double-modal-logic axioms) + +module _ + {l1 l2 l3 : Level} {i : Set l1} {ax₁ : formulas l2 i} {ax₂ : formulas l3 i} + (leq : ax₁ ⊆ ax₂) + where + + deduction-monotic : {a : formula i} → ax₁ ⊢ a → ax₂ ⊢ a + deduction-monotic (ax x) = ax (leq _ x) + deduction-monotic (mp dab da) = + mp (deduction-monotic dab) (deduction-monotic da) + deduction-monotic (nec d) = nec (deduction-monotic d) + + modal-logic-monotic : modal-logic ax₁ ⊆ modal-logic ax₂ + modal-logic-monotic a = + map-universal-property-trunc-Prop + ( modal-logic ax₂ a) + ( unit-trunc-Prop ∘ deduction-monotic) + +module _ + {l1 l2 l3 : Level} {i : Set l1} {ax₁ : formulas l2 i} {ax₂ : formulas l3 i} + where + + -- TODO: change name + modal-logic-CHOOSE-NAME : + ax₁ ⊆ modal-logic ax₂ → modal-logic ax₁ ⊆ modal-logic ax₂ + modal-logic-CHOOSE-NAME leq = + transitive-leq-subtype + ( modal-logic ax₁) + ( modal-logic (modal-logic ax₂)) + ( modal-logic ax₂) + ( subset-double-modal-logic ax₂) + ( modal-logic-monotic leq) +``` diff --git a/src/modal-logic/modal-logic-IK.lagda.md b/src/modal-logic/modal-logic-IK.lagda.md new file mode 100644 index 0000000000..c24aee5138 --- /dev/null +++ b/src/modal-logic/modal-logic-IK.lagda.md @@ -0,0 +1,60 @@ +# Modal logic IK + +```agda +module modal-logic.modal-logic-IK where +``` + +
Imports + +```agda +open import foundation.inhabited-types +open import foundation.intersections-subtypes +open import foundation.sets +open import foundation.unions-subtypes +open import foundation.universe-levels + +open import modal-logic.axioms +open import modal-logic.formulas +open import modal-logic.kripke-semantics +open import modal-logic.logic-syntax +open import modal-logic.soundness +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l : Level} (i : Set l) + where + + modal-logic-IK : formulas l i + modal-logic-IK = + modal-logic (union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i))) + +module _ + {l1 l2 : Level} + (i : Set l1) + (w : Inhabited-Type l2) + (l3 l4 : Level) + where + + soundness-IK : soundness (modal-logic-IK i) (all-models w l3 i l4) + soundness-IK = + soundness-modal-logic-union-same-class + ( ax-k i) + ( union-subtype (ax-s i) (ax-n i)) + ( all-models w l3 i l4) + ( ax-k-soundness i w l3 l4) + ( soundness-union-same-class + ( ax-s i) + ( ax-n i) + ( all-models w l3 i l4) + ( ax-s-soundness i w l3 l4) + ( ax-n-soundness i w l3 l4)) +``` diff --git a/src/modal-logic/modal-logic-K.lagda.md b/src/modal-logic/modal-logic-K.lagda.md new file mode 100644 index 0000000000..33c4a87ce6 --- /dev/null +++ b/src/modal-logic/modal-logic-K.lagda.md @@ -0,0 +1,50 @@ +# Modal logic K + +```agda +module modal-logic.modal-logic-K where +``` + +
Imports + +```agda +open import foundation.inhabited-types +open import foundation.intersections-subtypes +open import foundation.sets +open import foundation.unions-subtypes +open import foundation.universe-levels + +open import modal-logic.axioms +open import modal-logic.formulas +open import modal-logic.kripke-semantics +open import modal-logic.logic-syntax +open import modal-logic.modal-logic-IK +open import modal-logic.soundness +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l : Level} (i : Set l) + where + + modal-logic-K : formulas l i + modal-logic-K = modal-logic (union-subtype (modal-logic-IK i) (ax-dn i)) + +module _ + {l1 l2 l3 : Level} + (i : Set l1) (axioms : formulas l2 i) + (w : Inhabited-Type l3) + {l4 l5 : Level} + where + + -- soundness-K : soundness (modal-logic-K i) (decidable-models w l4 i l5) + -- soundness-K = + -- {! !} +``` diff --git a/src/modal-logic/soundness.lagda.md b/src/modal-logic/soundness.lagda.md new file mode 100644 index 0000000000..63827ec4d2 --- /dev/null +++ b/src/modal-logic/soundness.lagda.md @@ -0,0 +1,139 @@ +# Modal logic soundness + +```agda +module modal-logic.soundness where +``` + +
Imports + +```agda +open import foundation.dependent-pair-types +open import foundation.inhabited-types +open import foundation.intersections-subtypes +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.sets +open import foundation.subtypes +open import foundation.transport-along-identifications +open import foundation.unions-subtypes +open import foundation.universe-levels + +open import foundation-core.coproduct-types + +open import modal-logic.formulas +open import modal-logic.kripke-semantics +open import modal-logic.logic-syntax +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l1 l2 l3 l4 l5 l6 : Level} + {i : Set l1} (logic : formulas l2 i) + {w : Inhabited-Type l3} (C : model-class w l4 i l5 l6) + where + + soundness : UU (l1 ⊔ l2 ⊔ l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6) + soundness = logic ⊆ (class-modal-logic C) + +module _ + {l1 l2 l3 l4 l5 l6 : Level} + {i : Set l1} {axioms : formulas l2 i} + {w : Inhabited-Type l3} (C : model-class w l4 i l5 l6) + where + + soundness-axioms : + ((a : formula i) → is-in-subtype axioms a → type-Prop (C ⊨C a)) → + {a : formula i} → axioms ⊢ a → type-Prop (C ⊨C a) + soundness-axioms H (ax x) = H _ x + soundness-axioms H (mp dab da) M in-class x = + soundness-axioms H dab M in-class x (soundness-axioms H da M in-class x) + soundness-axioms H (nec d) M in-class _ y _ = + soundness-axioms H d M in-class y + + soundness-deduction : + ((a : formula i) → is-in-subtype axioms a → type-Prop (C ⊨C a)) → + soundness (modal-logic axioms) C + soundness-deduction H a = + map-universal-property-trunc-Prop (C ⊨C a) (soundness-axioms H) +``` + +## Properties + +```agda +module _ + {l1 l2 l3 l4 l5 l6 l7 l8 : Level} + {i : Set l1} (ax₁ : formulas l2 i) (ax₂ : formulas l3 i) + {w : Inhabited-Type l4} + (C₁ : model-class w l5 i l6 l7) (C₂ : model-class w l5 i l6 l8) + (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) + where + + forces-in-intersection : + (M : kripke-model w l5 i l6) → + is-in-subtype (intersection-subtype C₁ C₂) M → + (a : formula i) → + is-in-subtype ax₁ a + is-in-subtype ax₂ a → + type-Prop (M ⊨M a) + forces-in-intersection M in-class a (inl d) = + sound₁ a d M (subtype-intersection-left C₁ C₂ M in-class) + forces-in-intersection M in-class a (inr d) = + sound₂ a d M (subtype-intersection-right C₁ C₂ M in-class) + + soundness-union : + soundness (union-subtype ax₁ ax₂) (intersection-subtype C₁ C₂) + soundness-union a is-ax M in-class = + apply-universal-property-trunc-Prop + ( is-ax) + ( M ⊨M a) + ( forces-in-intersection M in-class a) + + soundness-modal-logic-union : + soundness (modal-logic (union-subtype ax₁ ax₂)) (intersection-subtype C₁ C₂) + soundness-modal-logic-union = + soundness-deduction (intersection-subtype C₁ C₂) soundness-union + +module _ + {l1 l2 l3 l4 l5 l6 l7 : Level} + {i : Set l1} (ax₁ : formulas l2 i) (ax₂ : formulas l3 i) + {w : Inhabited-Type l4} + (C₁ : model-class w l5 i l6 l7) (C₂ : model-class w l5 i l6 l7) + (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) + where + + soundness-modal-logic-union-subclass-left : + C₁ ⊆ C₂ → + soundness (modal-logic (union-subtype ax₁ ax₂)) C₁ + soundness-modal-logic-union-subclass-left C₁-sub-C₂ = + tr + ( soundness (modal-logic (union-subtype ax₁ ax₂))) + ( intersection-subtype-left C₁ C₂ C₁-sub-C₂) + ( soundness-modal-logic-union ax₁ ax₂ C₁ C₂ sound₁ sound₂) + +module _ + {l1 l2 l3 l4 l5 l6 l7 : Level} + {i : Set l1} (ax₁ : formulas l2 i) (ax₂ : formulas l3 i) + {w : Inhabited-Type l4} + (C : model-class w l5 i l6 l7) + (sound₁ : soundness ax₁ C) (sound₂ : soundness ax₂ C) + where + + soundness-union-same-class : soundness (union-subtype ax₁ ax₂) C + soundness-union-same-class = + tr + ( soundness (union-subtype ax₁ ax₂)) + ( is-reflexivity-intersection C) + ( soundness-union ax₁ ax₂ C C sound₁ sound₂) + + soundness-modal-logic-union-same-class : + soundness (modal-logic (union-subtype ax₁ ax₂)) C + soundness-modal-logic-union-same-class = + soundness-deduction C soundness-union-same-class +``` From 8ecf60ef184f6e57d7c22623a2db2283b23c48ef Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Thu, 26 Oct 2023 13:49:26 +0300 Subject: [PATCH 17/63] Remove old K definition --- src/modal-logic/K-soundness.lagda.md | 142 --------------------------- src/modal-logic/K-syntax.lagda.md | 34 ------- 2 files changed, 176 deletions(-) delete mode 100644 src/modal-logic/K-soundness.lagda.md delete mode 100644 src/modal-logic/K-syntax.lagda.md diff --git a/src/modal-logic/K-soundness.lagda.md b/src/modal-logic/K-soundness.lagda.md deleted file mode 100644 index b333cb354d..0000000000 --- a/src/modal-logic/K-soundness.lagda.md +++ /dev/null @@ -1,142 +0,0 @@ -# Modal logic K soundness - -```agda -module modal-logic.K-soundness where -``` - -
Imports - -```agda -open import foundation.coproduct-types -open import foundation.decidable-propositions -open import foundation.decidable-types -open import foundation.dependent-pair-types -open import foundation.double-negation -open import foundation.empty-types -open import foundation.equivalences -open import foundation.function-types -open import foundation.functoriality-function-types -open import foundation.inhabited-types -open import foundation.law-of-excluded-middle -open import foundation.propositional-truncations -open import foundation.propositions -open import foundation.raising-universe-levels -open import foundation.sets -open import foundation.unit-type -open import foundation.universe-levels - -open import modal-logic.K-syntax -open import modal-logic.formulas -open import modal-logic.kripke-semantics - -open import univalent-combinatorics.decidable-dependent-function-types -``` - -
- -## Idea - -TODO - -## Definition - -```agda -module _ - {l1 l2 l3 l4 : Level} {w : Inhabited-Type l1} {i : Set l3} - where - - private - l = l1 ⊔ l2 ⊔ l3 ⊔ l4 - - soundness : - ((M , _) : decidable-model w l2 i l4) - {a : formula i} → - ⊢ a → - type-Prop (M ⊨M a) - soundness _ (ax-k _ _) x fa _ = fa - soundness _ (ax-s _ _ _) x fabc fab fa = fabc fa (fab fa) - soundness (_ , cl) (ax-dn a) x fdna with cl a x - ... | inl fa = fa - ... | inr nfa = raise-ex-falso _ (fdna (map-raise ∘ nfa)) - soundness _ (ax-n _ _) x fab fa y r = fab y r (fa y r) - soundness CM (mp dab da) x = soundness CM dab x (soundness CM da x) - soundness CM (nec d) x y _ = soundness CM d y - - module _ - ((M , w-is-finite) : finite-model w l2 i l4) - (dec-val : ∀ n x → is-decidable (type-Prop (model-valuate M n x))) - (dec-r : ∀ x y → is-decidable (model-relation M x y)) - where - - finite-model-is-decidable : is-decidable-model M - finite-model-is-decidable (var n) x = - is-decidable-raise _ _ (dec-val n x) - finite-model-is-decidable ⊥ x = - inr map-inv-raise - finite-model-is-decidable (a ⇒ b) x = - is-decidable-function-type - ( finite-model-is-decidable a x) - ( finite-model-is-decidable b x) - finite-model-is-decidable (□ a) x = - is-decidable-Π-is-finite w-is-finite - ( λ y → - is-decidable-function-type (dec-r x y) - ( finite-model-is-decidable a y)) - - finite-soundness : {a : formula i} → ⊢ a → type-Prop (M ⊨M a) - finite-soundness = soundness (M , finite-model-is-decidable) - --- TODO: move to foundations -double-negation-LEM : - {l : Level} → - ((P : Prop l) → ¬¬ (type-Prop P) → (type-Prop P)) → - LEM l -double-negation-LEM dnP P = - dnP (is-decidable-Prop P) (λ ndec → ndec (inr (ndec ∘ inl))) - --- TODO: move to raising -compute-raise-function : - (l : Level) {l1 l2 : Level} (A : UU l1) (B : UU l2) → - (A → B) ≃ (raise l A → raise l B) -compute-raise-function _ _ _ = - equiv-function-type _ (compute-raise _ _) (compute-raise _ _) - -raise-double-negation : - (l : Level) {l1 : Level} {A : UU l1} → - ¬¬ A → - (raise l A → raise-empty l) → raise-empty l -raise-double-negation l = - map-equiv-function-type _ (compute-raise-function _ _ _) (compute-raise _ _) --- 2 --- map-equiv-function-type _ --- ( equiv-function-type _ (compute-raise _ _) (compute-raise _ _)) --- ( compute-raise _ _) --- 3 --- raise-double-negation dnA rnA = --- map-raise (dnA (map-inv-raise ∘ rnA ∘ map-raise)) - -full-soundness : UUω -full-soundness = - {l1 l2 l3 l4 : Level} - (W : Inhabited-Type l1) - (i : Set l3) - (M : kripke-model W l2 i l4) - {a : formula i} → - ⊢ a → - type-Prop (M ⊨M a) - -full-soundness-required-LEM : full-soundness → (l : Level) → LEM l -full-soundness-required-LEM sound l = - double-negation-LEM required-double-negation - where - required-double-negation : (P : Prop l) → ¬¬ (type-Prop P) → type-Prop P - required-double-negation P dnP = - map-inv-raise - ( sound - ( unit , unit-trunc-Prop star) - unit-Set - ( model (frame (λ _ _ → empty-Prop)) (λ _ _ → P)) - ( ax-dn (var star)) - star - ( raise-double-negation _ dnP)) -``` diff --git a/src/modal-logic/K-syntax.lagda.md b/src/modal-logic/K-syntax.lagda.md deleted file mode 100644 index 4706f195ed..0000000000 --- a/src/modal-logic/K-syntax.lagda.md +++ /dev/null @@ -1,34 +0,0 @@ -# Modal logic K syntax - -```agda -module modal-logic.K-syntax where -``` - -
Imports - -```agda -open import foundation.sets -open import foundation.universe-levels - -open import modal-logic.formulas -``` - -
- -## Idea - -TODO - -## Definition - -```agda -infix 5 ⊢_ - -data ⊢_ {l : Level} {i : Set l} : formula i → UU l where - ax-k : (a b : formula i) → ⊢ a ⇒ b ⇒ a - ax-s : (a b c : formula i) → ⊢ (a ⇒ b ⇒ c) ⇒ (a ⇒ b) ⇒ (a ⇒ c) - ax-dn : (a : formula i) → ⊢ ~~ a ⇒ a - ax-n : (a b : formula i) → ⊢ □ (a ⇒ b) ⇒ □ a ⇒ □ b - mp : {a b : formula i} → ⊢ a ⇒ b → ⊢ a → ⊢ b - nec : {a : formula i} → ⊢ a → ⊢ □ a -``` From bb2f31bd83ca9dc845b76d47a87d588e1bf3aa03 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Thu, 26 Oct 2023 14:01:50 +0300 Subject: [PATCH 18/63] Fix index file --- src/modal-logic.lagda.md | 2 -- 1 file changed, 2 deletions(-) diff --git a/src/modal-logic.lagda.md b/src/modal-logic.lagda.md index 7e7973fde8..d5092c6f45 100644 --- a/src/modal-logic.lagda.md +++ b/src/modal-logic.lagda.md @@ -3,8 +3,6 @@ ```agda module modal-logic where -open import modal-logic.K-soundness public -open import modal-logic.K-syntax public open import modal-logic.axioms public open import modal-logic.formulas public open import modal-logic.kripke-semantics public From 41d20ab7d01d6179ffea46c5702092e1a5e545d1 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Fri, 27 Oct 2023 00:23:35 +0300 Subject: [PATCH 19/63] Add K soundness --- .../intersections-subtypes.lagda.md | 83 ++++++++++++------- src/foundation/subtypes.lagda.md | 6 +- src/foundation/unions-subtypes.lagda.md | 14 ++++ src/modal-logic/axioms.lagda.md | 26 ++---- src/modal-logic/kripke-semantics.lagda.md | 4 + src/modal-logic/modal-logic-K.lagda.md | 24 +++++- src/modal-logic/soundness.lagda.md | 45 ++++++++-- 7 files changed, 143 insertions(+), 59 deletions(-) diff --git a/src/foundation/intersections-subtypes.lagda.md b/src/foundation/intersections-subtypes.lagda.md index 3bd50b13f4..00891864f6 100644 --- a/src/foundation/intersections-subtypes.lagda.md +++ b/src/foundation/intersections-subtypes.lagda.md @@ -14,6 +14,7 @@ open import foundation.identity-types open import foundation.large-locale-of-subtypes open import foundation.powersets open import foundation.subtypes +open import foundation.transport-along-identifications open import foundation.universe-levels open import foundation-core.decidable-propositions @@ -87,7 +88,7 @@ module _ intersection-family-of-subtypes {I} P x = Π-Prop I (λ i → P i x) ``` -## +## TODO: Change title ```agda -- It's too simple =) @@ -123,30 +124,31 @@ module _ intersection-subtype-right-subtype Q-sub-P = subtype-both-intersection Q Q-sub-P (refl-leq-subtype Q) -module _ - {l1 l2 : Level} {X : UU l1} (P : subtype l2 X) (Q : subtype l2 X) - where - - intersection-subtype-left : - P ⊆ Q → intersection-subtype P Q = P - intersection-subtype-left P-sub-Q = - antisymmetric-leq-subtype _ _ - ( subtype-intersection-left P Q) - ( intersection-subtype-left-subtype P Q P-sub-Q) + equiv-intersection-subtype-left : + P ⊆ Q → equiv-subtypes (intersection-subtype P Q) P + equiv-intersection-subtype-left P-sub-Q = + equiv-antisymmetric-leq-subtype + ( intersection-subtype P Q) + ( P) + ( subtype-intersection-left) + ( intersection-subtype-left-subtype P-sub-Q) - intersection-subtype-right : - Q ⊆ P → intersection-subtype P Q = Q - intersection-subtype-right Q-sub-P = - antisymmetric-leq-subtype _ _ - ( subtype-intersection-right P Q) - ( intersection-subtype-right-subtype P Q Q-sub-P) + equiv-intersection-subtype-right : + Q ⊆ P → equiv-subtypes (intersection-subtype P Q) Q + equiv-intersection-subtype-right Q-sub-P = + equiv-antisymmetric-leq-subtype + ( intersection-subtype P Q) + ( Q) + ( subtype-intersection-right) + ( intersection-subtype-right-subtype Q-sub-P) module _ - {l1 l2 : Level} {X : UU l1} (P : subtype l2 X) + {l1 : Level} {X : UU l1} where - is-reflexivity-intersection : intersection-subtype P P = P - is-reflexivity-intersection = + is-reflexivity-intersection : + {l2 : Level} (P : subtype l2 X) → intersection-subtype P P = P + is-reflexivity-intersection P = antisymmetric-leq-subtype _ _ ( subtype-intersection-left P P) ( subtype-both-intersection @@ -156,26 +158,47 @@ module _ ( refl-leq-subtype P) ( refl-leq-subtype P)) -module _ - {l1 l2 l3 : Level} {X : UU l1} (P : subtype l2 X) (Q : subtype l3 X) - where - is-commutative-subtype-intersection : + {l2 l3 : Level} (P : subtype l2 X) (Q : subtype l3 X) → intersection-subtype P Q ⊆ intersection-subtype Q P - is-commutative-subtype-intersection = + is-commutative-subtype-intersection P Q = subtype-both-intersection Q P ( intersection-subtype P Q) ( subtype-intersection-right P Q) ( subtype-intersection-left P Q) -module _ - {l1 l2 l3 : Level} {X : UU l1} (P : subtype l2 X) (Q : subtype l3 X) - where - is-commutative-intersection : + {l2 l3 : Level} (P : subtype l2 X) (Q : subtype l3 X) → intersection-subtype P Q = intersection-subtype Q P - is-commutative-intersection = + is-commutative-intersection P Q = antisymmetric-leq-subtype _ _ ( is-commutative-subtype-intersection P Q) ( is-commutative-subtype-intersection Q P) + + intersection-subtype-left-sublevels : + {l2 : Level} (l3 : Level) (P : subtype (l2 ⊔ l3) X) (Q : subtype l2 X) → + P ⊆ Q → intersection-subtype P Q = P + intersection-subtype-left-sublevels _ P Q P-sub-Q = + antisymmetric-leq-subtype _ _ + ( subtype-intersection-left P Q) + ( intersection-subtype-left-subtype P Q P-sub-Q) + + intersection-subtype-right-sublevels : + {l2 : Level} (l3 : Level) (P : subtype l2 X) (Q : subtype (l2 ⊔ l3) X) → + Q ⊆ P → intersection-subtype P Q = Q + intersection-subtype-right-sublevels l3 P Q Q-sub-P = + tr + ( _= Q) + ( is-commutative-intersection Q P) + ( intersection-subtype-left-sublevels l3 Q P Q-sub-P) + + intersection-subtype-left : + {l2 : Level} (P Q : subtype l2 X) → + P ⊆ Q → intersection-subtype P Q = P + intersection-subtype-left = intersection-subtype-left-sublevels lzero + + intersection-subtype-right : + {l2 : Level} (P Q : subtype l2 X) → + Q ⊆ P → intersection-subtype P Q = Q + intersection-subtype-right = intersection-subtype-right-sublevels lzero ``` diff --git a/src/foundation/subtypes.lagda.md b/src/foundation/subtypes.lagda.md index b181a9a10d..c91fc2e0d4 100644 --- a/src/foundation/subtypes.lagda.md +++ b/src/foundation/subtypes.lagda.md @@ -143,9 +143,13 @@ module _ {l1 : Level} {A : UU l1} where + equiv-subtypes : + {l2 l3 : Level} (P : subtype l2 A) (Q : subtype l3 A) → UU (l1 ⊔ l2 ⊔ l3) + equiv-subtypes P Q = (x : A) → is-in-subtype P x ≃ is-in-subtype Q x + equiv-antisymmetric-leq-subtype : {l2 l3 : Level} (P : subtype l2 A) (Q : subtype l3 A) → P ⊆ Q → Q ⊆ P → - (x : A) → is-in-subtype P x ≃ is-in-subtype Q x + equiv-subtypes P Q equiv-antisymmetric-leq-subtype P Q H K x = equiv-iff-is-prop ( is-prop-is-in-subtype P x) diff --git a/src/foundation/unions-subtypes.lagda.md b/src/foundation/unions-subtypes.lagda.md index f1bd0aee75..18f2c00960 100644 --- a/src/foundation/unions-subtypes.lagda.md +++ b/src/foundation/unions-subtypes.lagda.md @@ -88,3 +88,17 @@ module _ ( union-family-of-subtypes B x) ( λ (i , q) → unit-trunc-Prop (i , H i x q)) ``` + +## TODO: Change title + +```agda +module _ + {l1 l2 l3 : Level} {X : UU l1} (P : subtype l2 X) (Q : subtype l3 X) + where + + subtype-union-left : P ⊆ union-subtype P Q + subtype-union-left x = inl-disj-Prop (P x) (Q x) + + subtype-union-right : Q ⊆ union-subtype P Q + subtype-union-right x = inr-disj-Prop (P x) (Q x) +``` diff --git a/src/modal-logic/axioms.lagda.md b/src/modal-logic/axioms.lagda.md index 18040eb947..deb6dda928 100644 --- a/src/modal-logic/axioms.lagda.md +++ b/src/modal-logic/axioms.lagda.md @@ -74,10 +74,7 @@ module _ (l3 l4 : Level) where - ax-k-soundness : - (a : formula i) → - is-in-subtype (ax-k i) a → - type-Prop (all-models w l3 i l4 ⊨C a) + ax-k-soundness : soundness (ax-k i) (all-models w l3 i l4) ax-k-soundness (a₁ ⇒ b ⇒ a₁) refl M in-class x fa _ = fa ax-k-soundness (var _) (map-raise ()) ax-k-soundness ⊥ (map-raise ()) @@ -86,14 +83,12 @@ module _ ax-k-soundness (_ ⇒ ⊥) (map-raise ()) ax-k-soundness (_ ⇒ □ _) (map-raise ()) - ax-n-soundness : - (a : formula i) → - is-in-subtype (ax-n i) a → - type-Prop (all-models w l3 i l4 ⊨C a) + ax-n-soundness : soundness (ax-n i) (all-models w l3 i l4) ax-n-soundness (□ (a₁ ⇒ b₁) ⇒ □ a₁ ⇒ □ b₁) (refl , refl) - M in-class x fab fa y r = fab y r (fa y r) + M in-class x fab fa y r = + fab y r (fa y r) ax-n-soundness (var _) (map-raise ()) ax-n-soundness ⊥ (map-raise ()) ax-n-soundness (□ _) (map-raise ()) @@ -113,14 +108,12 @@ module _ ax-n-soundness (□ (_ ⇒ _) ⇒ □ _ ⇒ ⊥) (map-raise ()) ax-n-soundness (□ (_ ⇒ _) ⇒ □ _ ⇒ _ ⇒ _) (map-raise ()) - ax-s-soundness : - (a : formula i) → - is-in-subtype (ax-s i) a → - type-Prop (all-models w l3 i l4 ⊨C a) + ax-s-soundness : soundness (ax-s i) (all-models w l3 i l4) ax-s-soundness ((a ⇒ b ⇒ c) ⇒ (a ⇒ b) ⇒ a ⇒ c) (refl , refl , refl , refl) - M in-class x fabc fab fa = fabc fa (fab fa) + M in-class x fabc fab fa = + fabc fa (fab fa) ax-s-soundness (var _) (map-raise ()) ax-s-soundness ⊥ (map-raise ()) ax-s-soundness (□ _) (map-raise ()) @@ -140,10 +133,7 @@ module _ ax-s-soundness ((_ ⇒ _ ⇒ _) ⇒ (_ ⇒ _) ⇒ ⊥) (map-raise ()) ax-s-soundness ((_ ⇒ _ ⇒ _) ⇒ (_ ⇒ _) ⇒ □ _) (map-raise ()) - ax-dn-soundness : - (a : formula i) → - is-in-subtype (ax-dn i) a → - type-Prop (decidable-models w l3 i l4 ⊨C a) + ax-dn-soundness : soundness (ax-dn i) (decidable-models w l3 i l4) ax-dn-soundness (((a ⇒ ⊥) ⇒ ⊥) ⇒ a) refl M is-dec x f with is-dec a x ... | inl fa = fa ... | inr fna = raise-ex-falso _ (f (λ fa → map-raise (fna fa))) diff --git a/src/modal-logic/kripke-semantics.lagda.md b/src/modal-logic/kripke-semantics.lagda.md index d22d97e942..c96da45d68 100644 --- a/src/modal-logic/kripke-semantics.lagda.md +++ b/src/modal-logic/kripke-semantics.lagda.md @@ -73,6 +73,10 @@ module _ all-models : model-class lzero all-models _ = unit-Prop + all-models-is-biggest-class : + {l : Level} (C : model-class l) → C ⊆ all-models + all-models-is-biggest-class _ _ _ = star + module _ {l1 l2 l3 l4 : Level} {w : Inhabited-Type l1} {i : Set l3} where diff --git a/src/modal-logic/modal-logic-K.lagda.md b/src/modal-logic/modal-logic-K.lagda.md index 33c4a87ce6..eafb8e030f 100644 --- a/src/modal-logic/modal-logic-K.lagda.md +++ b/src/modal-logic/modal-logic-K.lagda.md @@ -10,6 +10,7 @@ module modal-logic.modal-logic-K where open import foundation.inhabited-types open import foundation.intersections-subtypes open import foundation.sets +open import foundation.subtypes open import foundation.unions-subtypes open import foundation.universe-levels @@ -37,6 +38,15 @@ module _ modal-logic-K : formulas l i modal-logic-K = modal-logic (union-subtype (modal-logic-IK i) (ax-dn i)) + IK-subtype-K : modal-logic-IK i ⊆ modal-logic-K + IK-subtype-K = + transitive-leq-subtype + ( modal-logic-IK i) + ( union-subtype (modal-logic-IK i) (ax-dn i)) + ( modal-logic-K) + ( axioms-subset-modal-logic _) + ( subtype-union-left (modal-logic-IK i) (ax-dn i)) + module _ {l1 l2 l3 : Level} (i : Set l1) (axioms : formulas l2 i) @@ -44,7 +54,15 @@ module _ {l4 l5 : Level} where - -- soundness-K : soundness (modal-logic-K i) (decidable-models w l4 i l5) - -- soundness-K = - -- {! !} + soundness-K : soundness (modal-logic-K i) (decidable-models w l4 i l5) + soundness-K = + soundness-modal-logic-union-subclass-right-sublevels + ( modal-logic-IK i) + ( ax-dn i) + ( l1 ⊔ l3 ⊔ l4 ⊔ l5) + ( all-models w l4 i l5) + ( decidable-models w l4 i l5) + ( soundness-IK i w l4 l5) + ( ax-dn-soundness i w l4 l5) + ( all-models-is-biggest-class w l4 i l5 (decidable-models w l4 i l5)) ``` diff --git a/src/modal-logic/soundness.lagda.md b/src/modal-logic/soundness.lagda.md index 63827ec4d2..0864ba85d6 100644 --- a/src/modal-logic/soundness.lagda.md +++ b/src/modal-logic/soundness.lagda.md @@ -104,18 +104,49 @@ module _ {l1 l2 l3 l4 l5 l6 l7 : Level} {i : Set l1} (ax₁ : formulas l2 i) (ax₂ : formulas l3 i) {w : Inhabited-Type l4} - (C₁ : model-class w l5 i l6 l7) (C₂ : model-class w l5 i l6 l7) - (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) where + soundness-modal-logic-union-subclass-left-sublevels : + (l8 : Level) + (C₁ : model-class w l5 i l6 (l7 ⊔ l8)) (C₂ : model-class w l5 i l6 l7) + (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) → + C₁ ⊆ C₂ → + soundness (modal-logic (union-subtype ax₁ ax₂)) C₁ + soundness-modal-logic-union-subclass-left-sublevels + l8 C₁ C₂ sound₁ sound₂ C₁-sub-C₂ = + tr + ( soundness (modal-logic (union-subtype ax₁ ax₂))) + ( intersection-subtype-left-sublevels l8 C₁ C₂ C₁-sub-C₂) + ( soundness-modal-logic-union ax₁ ax₂ C₁ C₂ sound₁ sound₂) + + soundness-modal-logic-union-subclass-right-sublevels : + (l8 : Level) + (C₁ : model-class w l5 i l6 l7) (C₂ : model-class w l5 i l6 (l7 ⊔ l8)) + (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) → + C₂ ⊆ C₁ → + soundness (modal-logic (union-subtype ax₁ ax₂)) C₂ + soundness-modal-logic-union-subclass-right-sublevels + l8 C₁ C₂ sound₁ sound₂ C₂-sub-C₁ = + tr + ( soundness (modal-logic (union-subtype ax₁ ax₂))) + ( intersection-subtype-right-sublevels l8 C₁ C₂ C₂-sub-C₁) + ( soundness-modal-logic-union ax₁ ax₂ C₁ C₂ sound₁ sound₂) + soundness-modal-logic-union-subclass-left : + (C₁ : model-class w l5 i l6 l7) (C₂ : model-class w l5 i l6 l7) + (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) → C₁ ⊆ C₂ → soundness (modal-logic (union-subtype ax₁ ax₂)) C₁ - soundness-modal-logic-union-subclass-left C₁-sub-C₂ = - tr - ( soundness (modal-logic (union-subtype ax₁ ax₂))) - ( intersection-subtype-left C₁ C₂ C₁-sub-C₂) - ( soundness-modal-logic-union ax₁ ax₂ C₁ C₂ sound₁ sound₂) + soundness-modal-logic-union-subclass-left = + soundness-modal-logic-union-subclass-left-sublevels lzero + + soundness-modal-logic-union-subclass-right : + (C₁ : model-class w l5 i l6 l7) (C₂ : model-class w l5 i l6 l7) + (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) → + C₂ ⊆ C₁ → + soundness (modal-logic (union-subtype ax₁ ax₂)) C₂ + soundness-modal-logic-union-subclass-right = + soundness-modal-logic-union-subclass-right-sublevels lzero module _ {l1 l2 l3 l4 l5 l6 l7 : Level} From 03d2142dd89fb06f3b1e0398f388858a1b086dcf Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Fri, 27 Oct 2023 10:57:12 +0300 Subject: [PATCH 20/63] Refactor kripke-models --- src/modal-logic/axioms.lagda.md | 2 +- src/modal-logic/kripke-semantics.lagda.md | 87 +++++++++-------------- src/modal-logic/modal-logic-IK.lagda.md | 3 +- src/modal-logic/modal-logic-K.lagda.md | 2 +- src/modal-logic/soundness.lagda.md | 10 +-- 5 files changed, 42 insertions(+), 62 deletions(-) diff --git a/src/modal-logic/axioms.lagda.md b/src/modal-logic/axioms.lagda.md index deb6dda928..34c3953b2c 100644 --- a/src/modal-logic/axioms.lagda.md +++ b/src/modal-logic/axioms.lagda.md @@ -70,7 +70,7 @@ module _ module _ {l1 l2 : Level} (i : Set l1) - (w : Inhabited-Type l2) + (w : UU l2) (l3 l4 : Level) where diff --git a/src/modal-logic/kripke-semantics.lagda.md b/src/modal-logic/kripke-semantics.lagda.md index c96da45d68..07e3f9c51d 100644 --- a/src/modal-logic/kripke-semantics.lagda.md +++ b/src/modal-logic/kripke-semantics.lagda.md @@ -41,31 +41,31 @@ TODO ```agda module _ - {l1 : Level} (w : Inhabited-Type l1) (l2 : Level) + {l1 : Level} (w : UU l1) (l2 : Level) where - record kripke-frame : UU (l1 ⊔ lsuc l2) where - constructor frame - field - frame-relation : Relation-Prop l2 (type-Inhabited-Type w) - open kripke-frame public + kripke-frame : UU (l1 ⊔ lsuc l2) + kripke-frame = is-inhabited w × Relation-Prop l2 w module _ - {l1 : Level} (w : Inhabited-Type l1) + {l1 l2 : Level} {w : UU l1} + where + + frame-inhabited : kripke-frame w l2 → is-inhabited w + frame-inhabited = pr1 + + frame-relation : kripke-frame w l2 → Relation l2 w + frame-relation = type-Relation-Prop ∘ pr2 + +module _ + {l1 : Level} (w : UU l1) (l2 : Level) {l3 : Level} (i : Set l3) (l4 : Level) where - record kripke-model : UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4) where - constructor model - field - model-frame : kripke-frame w l2 - model-valuate : type-Set i → type-Inhabited-Type w → Prop l4 - open kripke-model public - - finite-model : UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4) - finite-model = kripke-model × is-finite (type-Inhabited-Type w) + kripke-model : UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4) + kripke-model = (kripke-frame w l2) × (type-Set i → w → Prop l4) model-class : (l5 : Level) → UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ lsuc l5) model-class l5 = subtype l5 kripke-model @@ -78,11 +78,18 @@ module _ all-models-is-biggest-class _ _ _ = star module _ - {l1 l2 l3 l4 : Level} {w : Inhabited-Type l1} {i : Set l3} + {l1 l2 l3 l4 : Level} {w : UU l1} {i : Set l3} where - model-relation : kripke-model w l2 i l4 → Relation l2 (type-Inhabited-Type w) - model-relation = type-Relation-Prop ∘ frame-relation ∘ model-frame + model-frame : kripke-model w l2 i l4 → kripke-frame w l2 + model-frame = pr1 + + model-valuate : + kripke-model w l2 i l4 → type-Set i → w → Prop l4 + model-valuate = pr2 + + model-relation : kripke-model w l2 i l4 → Relation l2 w + model-relation = frame-relation ∘ model-frame private l = l1 ⊔ l2 ⊔ l4 @@ -92,20 +99,20 @@ module _ infix 5 _⊨M_ infix 5 _⊭M_ - _⊨_ : kripke-model w l2 i l4 × (type-Inhabited-Type w) → formula i → Prop l + _⊨_ : kripke-model w l2 i l4 × w → formula i → Prop l (M , x) ⊨ var n = raise-Prop l (model-valuate M n x) (M , x) ⊨ ⊥ = raise-empty-Prop l (M , x) ⊨ a ⇒ b = implication-Prop ((M , x) ⊨ a) ((M , x) ⊨ b) (M , x) ⊨ □ a = Π-Prop - ( type-Inhabited-Type w) + ( w) ( λ y → function-Prop (model-relation M x y) ((M , y) ⊨ a)) - _⊭_ : kripke-model w l2 i l4 × (type-Inhabited-Type w) → formula i → Prop l + _⊭_ : kripke-model w l2 i l4 × w → formula i → Prop l (M , x) ⊭ a = neg-Prop ((M , x) ⊨ a) _⊨M_ : kripke-model w l2 i l4 → formula i → Prop l - M ⊨M a = Π-Prop (type-Inhabited-Type w) (λ x → (M , x) ⊨ a) + M ⊨M a = Π-Prop w (λ x → (M , x) ⊨ a) _⊭M_ : kripke-model w l2 i l4 → formula i → Prop l M ⊭M a = neg-Prop (M ⊨M a) @@ -114,7 +121,7 @@ module _ decidable-class M = Π-Prop (formula i) (λ a → M ⊨M a) module _ - {l1 l2 l3 l4 l5 : Level} {w : Inhabited-Type l1} {i : Set l3} + {l1 l2 l3 l4 l5 : Level} {w : UU l1} {i : Set l3} where _⊨C_ : @@ -132,7 +139,7 @@ module _ class-modal-logic C a = C ⊨C a module _ - {l1 : Level} (w : Inhabited-Type l1) + {l1 : Level} (w : UU l1) (l2 : Level) {l3 : Level} (i : Set l3) (l4 : Level) @@ -142,33 +149,5 @@ module _ decidable-models M = Π-Prop ( formula i) - ( λ a → - ( Π-Prop - ( type-Inhabited-Type w) - ( λ x → is-decidable-Prop ((M , x) ⊨ a)))) - -module _ - {l1 l2 l3 l4 : Level} {w : Inhabited-Type l1} {i : Set l3} - where - - private - l = l1 ⊔ l2 ⊔ l3 ⊔ l4 - - is-decidable-model-Prop : kripke-model w l2 i l4 → Prop l - is-decidable-model-Prop M = - Π-Prop - ( formula i) - ( λ a → - Π-Prop (type-Inhabited-Type w) (λ x → is-decidable-Prop ((M , x) ⊨ a))) - - is-decidable-model : kripke-model w l2 i l4 → UU l - is-decidable-model = type-Prop ∘ is-decidable-model-Prop - -decidable-model : - {l1 : Level} (w : Inhabited-Type l1) - (l2 : Level) - {l3 : Level} (i : Set l3) - (l4 : Level) → - UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4) -decidable-model w l2 i l4 = Σ (kripke-model w l2 i l4) is-decidable-model + ( λ a → (Π-Prop w (λ x → is-decidable-Prop ((M , x) ⊨ a)))) ``` diff --git a/src/modal-logic/modal-logic-IK.lagda.md b/src/modal-logic/modal-logic-IK.lagda.md index c24aee5138..db30f4e61c 100644 --- a/src/modal-logic/modal-logic-IK.lagda.md +++ b/src/modal-logic/modal-logic-IK.lagda.md @@ -33,6 +33,7 @@ module _ {l : Level} (i : Set l) where +-- TODO: It's not Intuitionistic K modal-logic-IK : formulas l i modal-logic-IK = modal-logic (union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i))) @@ -40,7 +41,7 @@ module _ module _ {l1 l2 : Level} (i : Set l1) - (w : Inhabited-Type l2) + (w : UU l2) (l3 l4 : Level) where diff --git a/src/modal-logic/modal-logic-K.lagda.md b/src/modal-logic/modal-logic-K.lagda.md index eafb8e030f..1b657d2f88 100644 --- a/src/modal-logic/modal-logic-K.lagda.md +++ b/src/modal-logic/modal-logic-K.lagda.md @@ -50,7 +50,7 @@ module _ module _ {l1 l2 l3 : Level} (i : Set l1) (axioms : formulas l2 i) - (w : Inhabited-Type l3) + (w : UU l3) {l4 l5 : Level} where diff --git a/src/modal-logic/soundness.lagda.md b/src/modal-logic/soundness.lagda.md index 0864ba85d6..6fc54c7fd8 100644 --- a/src/modal-logic/soundness.lagda.md +++ b/src/modal-logic/soundness.lagda.md @@ -37,7 +37,7 @@ TODO module _ {l1 l2 l3 l4 l5 l6 : Level} {i : Set l1} (logic : formulas l2 i) - {w : Inhabited-Type l3} (C : model-class w l4 i l5 l6) + {w : UU l3} (C : model-class w l4 i l5 l6) where soundness : UU (l1 ⊔ l2 ⊔ l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6) @@ -46,7 +46,7 @@ module _ module _ {l1 l2 l3 l4 l5 l6 : Level} {i : Set l1} {axioms : formulas l2 i} - {w : Inhabited-Type l3} (C : model-class w l4 i l5 l6) + {w : UU l3} (C : model-class w l4 i l5 l6) where soundness-axioms : @@ -71,7 +71,7 @@ module _ module _ {l1 l2 l3 l4 l5 l6 l7 l8 : Level} {i : Set l1} (ax₁ : formulas l2 i) (ax₂ : formulas l3 i) - {w : Inhabited-Type l4} + {w : UU l4} (C₁ : model-class w l5 i l6 l7) (C₂ : model-class w l5 i l6 l8) (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) where @@ -103,7 +103,7 @@ module _ module _ {l1 l2 l3 l4 l5 l6 l7 : Level} {i : Set l1} (ax₁ : formulas l2 i) (ax₂ : formulas l3 i) - {w : Inhabited-Type l4} + {w : UU l4} where soundness-modal-logic-union-subclass-left-sublevels : @@ -151,7 +151,7 @@ module _ module _ {l1 l2 l3 l4 l5 l6 l7 : Level} {i : Set l1} (ax₁ : formulas l2 i) (ax₂ : formulas l3 i) - {w : Inhabited-Type l4} + {w : UU l4} (C : model-class w l5 i l6 l7) (sound₁ : soundness ax₁ C) (sound₂ : soundness ax₂ C) where From 88cb89f5379a2b4790a06099e62a2e63c476f6a9 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Fri, 27 Oct 2023 12:13:43 +0300 Subject: [PATCH 21/63] Add finite models --- src/modal-logic/kripke-semantics.lagda.md | 76 ++++++++++++++++++++--- src/modal-logic/modal-logic-K.lagda.md | 15 ++++- src/modal-logic/soundness.lagda.md | 15 +++++ 3 files changed, 95 insertions(+), 11 deletions(-) diff --git a/src/modal-logic/kripke-semantics.lagda.md b/src/modal-logic/kripke-semantics.lagda.md index 07e3f9c51d..e341d838d7 100644 --- a/src/modal-logic/kripke-semantics.lagda.md +++ b/src/modal-logic/kripke-semantics.lagda.md @@ -9,7 +9,9 @@ module modal-logic.kripke-semantics where ```agda open import foundation.binary-relations open import foundation.cartesian-product-types +open import foundation.coproduct-types open import foundation.decidable-propositions +open import foundation.decidable-relations open import foundation.decidable-types open import foundation.dependent-pair-types open import foundation.empty-types @@ -26,6 +28,7 @@ open import foundation.universe-levels open import modal-logic.formulas open import modal-logic.logic-syntax +open import univalent-combinatorics.decidable-dependent-function-types open import univalent-combinatorics.finite-types ``` @@ -54,8 +57,11 @@ module _ frame-inhabited : kripke-frame w l2 → is-inhabited w frame-inhabited = pr1 + frame-relation-Prop : kripke-frame w l2 → Relation-Prop l2 w + frame-relation-Prop = pr2 + frame-relation : kripke-frame w l2 → Relation l2 w - frame-relation = type-Relation-Prop ∘ pr2 + frame-relation = type-Relation-Prop ∘ frame-relation-Prop module _ {l1 : Level} (w : UU l1) @@ -65,7 +71,7 @@ module _ where kripke-model : UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4) - kripke-model = (kripke-frame w l2) × (type-Set i → w → Prop l4) + kripke-model = kripke-frame w l2 × (type-Set i → w → Prop l4) model-class : (l5 : Level) → UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ lsuc l5) model-class l5 = subtype l5 kripke-model @@ -73,12 +79,12 @@ module _ all-models : model-class lzero all-models _ = unit-Prop - all-models-is-biggest-class : + all-models-is-largest-class : {l : Level} (C : model-class l) → C ⊆ all-models - all-models-is-biggest-class _ _ _ = star + all-models-is-largest-class _ _ _ = star module _ - {l1 l2 l3 l4 : Level} {w : UU l1} {i : Set l3} + {l1 l2 l3 l4 : Level} {w : UU l1} (i : Set l3) where model-frame : kripke-model w l2 i l4 → kripke-frame w l2 @@ -88,9 +94,16 @@ module _ kripke-model w l2 i l4 → type-Set i → w → Prop l4 model-valuate = pr2 + model-relation-Prop : kripke-model w l2 i l4 → Relation-Prop l2 w + model-relation-Prop = frame-relation-Prop ∘ model-frame + model-relation : kripke-model w l2 i l4 → Relation l2 w model-relation = frame-relation ∘ model-frame +module _ + {l1 l2 l3 l4 : Level} {w : UU l1} {i : Set l3} + where + private l = l1 ⊔ l2 ⊔ l4 @@ -100,13 +113,13 @@ module _ infix 5 _⊭M_ _⊨_ : kripke-model w l2 i l4 × w → formula i → Prop l - (M , x) ⊨ var n = raise-Prop l (model-valuate M n x) + (M , x) ⊨ var n = raise-Prop l (model-valuate i M n x) (M , x) ⊨ ⊥ = raise-empty-Prop l (M , x) ⊨ a ⇒ b = implication-Prop ((M , x) ⊨ a) ((M , x) ⊨ b) (M , x) ⊨ □ a = Π-Prop ( w) - ( λ y → function-Prop (model-relation M x y) ((M , y) ⊨ a)) + ( λ y → function-Prop (model-relation i M x y) ((M , y) ⊨ a)) _⊭_ : kripke-model w l2 i l4 × w → formula i → Prop l (M , x) ⊭ a = neg-Prop ((M , x) ⊨ a) @@ -121,10 +134,11 @@ module _ decidable-class M = Π-Prop (formula i) (λ a → M ⊨M a) module _ - {l1 l2 l3 l4 l5 : Level} {w : UU l1} {i : Set l3} + {l1 l2 l3 l4 : Level} {w : UU l1} {i : Set l3} where _⊨C_ : + {l5 : Level} → model-class w l2 i l4 l5 → formula i → Prop (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ l5) @@ -134,10 +148,21 @@ module _ ( λ M → function-Prop (is-in-subtype C M) (M ⊨M a)) class-modal-logic : + {l5 : Level} → model-class w l2 i l4 l5 → formulas (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ l5) i class-modal-logic C a = C ⊨C a + -- TODO: rename + class-modal-logic-monotic : + {l5 l6 : Level} + (C₁ : model-class w l2 i l4 l5) + (C₂ : model-class w l2 i l4 l6) → + C₁ ⊆ C₂ → + class-modal-logic C₂ ⊆ class-modal-logic C₁ + class-modal-logic-monotic C₁ C₂ sub _ in-modal-logic-C₂ M in-C₁ = + in-modal-logic-C₂ M (sub M in-C₁) + module _ {l1 : Level} (w : UU l1) (l2 : Level) @@ -150,4 +175,39 @@ module _ Π-Prop ( formula i) ( λ a → (Π-Prop w (λ x → is-decidable-Prop ((M , x) ⊨ a)))) + + finite-models : model-class w l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4) + finite-models M = + prod-Prop + ( is-finite-Prop w) + ( prod-Prop + ( Π-Prop + ( w) + ( λ x → + ( Π-Prop ( w) + ( λ y → is-decidable-Prop (model-relation-Prop i M x y))))) + ( Π-Prop + ( w) + ( λ x → + ( Π-Prop + ( type-Set i) + ( λ n → is-decidable-Prop (model-valuate i M n x)))))) + + finite-models-subclass-decidable-models : finite-models ⊆ decidable-models + finite-models-subclass-decidable-models M (_ , _ , dec-v) (var n) x = + is-decidable-raise (l1 ⊔ l2 ⊔ l4) _ (dec-v x n) + finite-models-subclass-decidable-models M _ ⊥ x = + inr map-inv-raise + finite-models-subclass-decidable-models M is-fin (a ⇒ b) x = + is-decidable-function-type + ( finite-models-subclass-decidable-models M is-fin a x) + ( finite-models-subclass-decidable-models M is-fin b x) + finite-models-subclass-decidable-models + M is-fin@(w-is-fin , dec-r , dec-v) (□ a) x = + is-decidable-Π-is-finite + ( w-is-fin) + ( λ y → + ( is-decidable-function-type + ( dec-r x y) + ( finite-models-subclass-decidable-models M is-fin a y))) ``` diff --git a/src/modal-logic/modal-logic-K.lagda.md b/src/modal-logic/modal-logic-K.lagda.md index 1b657d2f88..69338a6124 100644 --- a/src/modal-logic/modal-logic-K.lagda.md +++ b/src/modal-logic/modal-logic-K.lagda.md @@ -38,8 +38,8 @@ module _ modal-logic-K : formulas l i modal-logic-K = modal-logic (union-subtype (modal-logic-IK i) (ax-dn i)) - IK-subtype-K : modal-logic-IK i ⊆ modal-logic-K - IK-subtype-K = + IK-subset-K : modal-logic-IK i ⊆ modal-logic-K + IK-subset-K = transitive-leq-subtype ( modal-logic-IK i) ( union-subtype (modal-logic-IK i) (ax-dn i)) @@ -64,5 +64,14 @@ module _ ( decidable-models w l4 i l5) ( soundness-IK i w l4 l5) ( ax-dn-soundness i w l4 l5) - ( all-models-is-biggest-class w l4 i l5 (decidable-models w l4 i l5)) + ( all-models-is-largest-class w l4 i l5 (decidable-models w l4 i l5)) + + soundness-K-finite : soundness (modal-logic-K i) (finite-models w l4 i l5) + soundness-K-finite = + soundness-subclass + ( modal-logic-K i) + ( decidable-models w l4 i l5) + ( finite-models w l4 i l5) + ( finite-models-subclass-decidable-models w l4 i l5) + ( soundness-K) ``` diff --git a/src/modal-logic/soundness.lagda.md b/src/modal-logic/soundness.lagda.md index 6fc54c7fd8..f3cf248347 100644 --- a/src/modal-logic/soundness.lagda.md +++ b/src/modal-logic/soundness.lagda.md @@ -68,6 +68,21 @@ module _ ## Properties ```agda +module _ + {l1 l2 l3 l4 l5 l6 l7 : Level} + {i : Set l1} (logic : formulas l2 i) + {w : UU l3} + (C₁ : model-class w l4 i l5 l6) (C₂ : model-class w l4 i l5 l7) + where + + soundness-subclass : C₂ ⊆ C₁ → soundness logic C₁ → soundness logic C₂ + soundness-subclass sub = + transitive-leq-subtype + ( logic) + ( class-modal-logic C₁) + ( class-modal-logic C₂) + ( class-modal-logic-monotic C₂ C₁ sub) + module _ {l1 l2 l3 l4 l5 l6 l7 l8 : Level} {i : Set l1} (ax₁ : formulas l2 i) (ax₂ : formulas l3 i) From ef73d85900cb3cd64fdffffc0d48f5ee1e6ed777 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Sat, 28 Oct 2023 18:44:58 +0300 Subject: [PATCH 22/63] Experiments with axioms definitions --- src/modal-logic/axioms.lagda.md | 46 +++++++++++++++++++++++++++++++ src/modal-logic/formulas.lagda.md | 2 +- 2 files changed, 47 insertions(+), 1 deletion(-) diff --git a/src/modal-logic/axioms.lagda.md b/src/modal-logic/axioms.lagda.md index 34c3953b2c..a584e97ddd 100644 --- a/src/modal-logic/axioms.lagda.md +++ b/src/modal-logic/axioms.lagda.md @@ -8,9 +8,11 @@ module modal-logic.axioms where ```agda open import foundation.conjunction +open import foundation.contractible-types open import foundation.coproduct-types open import foundation.dependent-pair-types open import foundation.empty-types +open import foundation.existential-quantification open import foundation.function-types open import foundation.identity-types open import foundation.inhabited-types @@ -67,6 +69,23 @@ module _ {-# CATCHALL #-} ax-dn _ = raise-empty-Prop l + ax-dn' : formulas l i + ax-dn' f = ∃-Prop (formula i) (λ a → f = (~~ a ⇒ a)) + + ax-dn'' : formulas l i + pr1 (ax-dn'' f) = Σ (formula i) (λ a → f = (~~ a ⇒ a)) + pr2 (ax-dn'' (((a₁ ⇒ ⊥) ⇒ ⊥) ⇒ a₁)) (a₁ , refl) = + is-prop-is-contr + ( is-contr-Σ-is-prop + ( a₁) + ( refl) + ( λ _ → is-set-formula i _ _) + ( λ _ → arrow-right-eq)) + ( a₁ , refl) + where + arrow-right-eq : {a b x y : formula i} → x ⇒ a = y ⇒ b → a = b + arrow-right-eq refl = refl + module _ {l1 l2 : Level} (i : Set l1) @@ -152,4 +171,31 @@ module _ ax-dn-soundness (((_ ⇒ ⊥) ⇒ var _) ⇒ _) (map-raise ()) ax-dn-soundness (((_ ⇒ ⊥) ⇒ _ ⇒ _) ⇒ _) (map-raise ()) ax-dn-soundness (((_ ⇒ ⊥) ⇒ □ _) ⇒ _) (map-raise ()) + + ax-dn-soundness' : soundness (ax-dn' i) (decidable-models w l3 i l4) + ax-dn-soundness' a = + map-universal-property-trunc-Prop + ( decidable-models w l3 i l4 ⊨C a) + ( aux) + -- with inductor + -- ( λ { + -- (b , refl) M is-dec x f → + -- ( ind-coprod + -- ( λ _ → type-Prop ((M , x) ⊨ b)) + -- ( id) + -- ( λ fna → raise-ex-falso _ (f (λ fa -> map-raise (fna fa)))) + -- ( is-dec b x))}) + where + aux : + Σ (formula i) (λ b → a = ((b ⇒ ⊥) ⇒ ⊥) ⇒ b) → + type-Prop (decidable-models w l3 i l4 ⊨C a) + aux (b , refl) M is-dec x f with (is-dec b x) + ... | inl fa = fa + ... | inr fna = raise-ex-falso _ (f (λ fa -> map-raise (fna fa))) + + ax-dn-soundness'' : soundness (ax-dn'' i) (decidable-models w l3 i l4) + ax-dn-soundness'' + (((a ⇒ ⊥) ⇒ ⊥) ⇒ a) (a , refl) _ is-dec x f with (is-dec a x) + ... | inl fa = fa + ... | inr fna = raise-ex-falso _ (f (λ fa -> map-raise (fna fa))) ``` diff --git a/src/modal-logic/formulas.lagda.md b/src/modal-logic/formulas.lagda.md index 30c8771c53..21297041e9 100644 --- a/src/modal-logic/formulas.lagda.md +++ b/src/modal-logic/formulas.lagda.md @@ -36,7 +36,7 @@ module _ {l : Level} (i : Set l) where - infixr 6 _⇒_ + infixr 7 _⇒_ infixr 25 □_ data formula : UU l where From 28fc80eec6bf2f4e293fefc423af062795ed3970 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Sun, 29 Oct 2023 17:56:23 +0300 Subject: [PATCH 23/63] Add completeness definition --- src/modal-logic.lagda.md | 1 + src/modal-logic/completeness.lagda.md | 66 +++++++++++++++++++++++++++ src/modal-logic/soundness.lagda.md | 23 ++++------ 3 files changed, 77 insertions(+), 13 deletions(-) create mode 100644 src/modal-logic/completeness.lagda.md diff --git a/src/modal-logic.lagda.md b/src/modal-logic.lagda.md index d5092c6f45..03b7657257 100644 --- a/src/modal-logic.lagda.md +++ b/src/modal-logic.lagda.md @@ -4,6 +4,7 @@ module modal-logic where open import modal-logic.axioms public +open import modal-logic.completeness public open import modal-logic.formulas public open import modal-logic.kripke-semantics public open import modal-logic.logic-syntax public diff --git a/src/modal-logic/completeness.lagda.md b/src/modal-logic/completeness.lagda.md new file mode 100644 index 0000000000..1ac2fdba26 --- /dev/null +++ b/src/modal-logic/completeness.lagda.md @@ -0,0 +1,66 @@ +# Modal logic completeness + +```agda +module modal-logic.completeness where +``` + +
Imports + +```agda +open import foundation.dependent-pair-types +open import foundation.inhabited-types +open import foundation.intersections-subtypes +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.sets +open import foundation.subtypes +open import foundation.transport-along-identifications +open import foundation.unions-subtypes +open import foundation.universe-levels + +open import foundation-core.coproduct-types + +open import modal-logic.formulas +open import modal-logic.kripke-semantics +open import modal-logic.logic-syntax +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l1 l2 l3 l4 l5 l6 : Level} + {i : Set l1} (logic : formulas l2 i) + {w : UU l3} (C : model-class w l4 i l5 l6) + where + + completeness : UU (l1 ⊔ l2 ⊔ l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6) + completeness = class-modal-logic C ⊆ logic +``` + +## Properties + +```agda +module _ + {l1 l2 l3 l4 l5 l6 l7 : Level} + {i : Set l1} (logic : formulas l2 i) + {w : UU l3} + (C₁ : model-class w l4 i l5 l6) (C₂ : model-class w l4 i l5 l7) + where + + completeness-subclass : + C₁ ⊆ C₂ → completeness logic C₁ → completeness logic C₂ + completeness-subclass sub complete = + transitive-leq-subtype + ( class-modal-logic C₂) + ( class-modal-logic C₁) + ( logic) + ( complete) + ( class-modal-logic-monotic C₁ C₂ sub) +``` diff --git a/src/modal-logic/soundness.lagda.md b/src/modal-logic/soundness.lagda.md index f3cf248347..adfbeb89fe 100644 --- a/src/modal-logic/soundness.lagda.md +++ b/src/modal-logic/soundness.lagda.md @@ -41,8 +41,12 @@ module _ where soundness : UU (l1 ⊔ l2 ⊔ l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6) - soundness = logic ⊆ (class-modal-logic C) + soundness = logic ⊆ class-modal-logic C +``` + +## Properties +```agda module _ {l1 l2 l3 l4 l5 l6 : Level} {i : Set l1} {axioms : formulas l2 i} @@ -50,24 +54,17 @@ module _ where soundness-axioms : - ((a : formula i) → is-in-subtype axioms a → type-Prop (C ⊨C a)) → - {a : formula i} → axioms ⊢ a → type-Prop (C ⊨C a) + soundness axioms C → {a : formula i} → axioms ⊢ a → type-Prop (C ⊨C a) soundness-axioms H (ax x) = H _ x soundness-axioms H (mp dab da) M in-class x = soundness-axioms H dab M in-class x (soundness-axioms H da M in-class x) soundness-axioms H (nec d) M in-class _ y _ = soundness-axioms H d M in-class y - soundness-deduction : - ((a : formula i) → is-in-subtype axioms a → type-Prop (C ⊨C a)) → - soundness (modal-logic axioms) C - soundness-deduction H a = + soundness-modal-logic : soundness axioms C → soundness (modal-logic axioms) C + soundness-modal-logic H a = map-universal-property-trunc-Prop (C ⊨C a) (soundness-axioms H) -``` - -## Properties -```agda module _ {l1 l2 l3 l4 l5 l6 l7 : Level} {i : Set l1} (logic : formulas l2 i) @@ -113,7 +110,7 @@ module _ soundness-modal-logic-union : soundness (modal-logic (union-subtype ax₁ ax₂)) (intersection-subtype C₁ C₂) soundness-modal-logic-union = - soundness-deduction (intersection-subtype C₁ C₂) soundness-union + soundness-modal-logic (intersection-subtype C₁ C₂) soundness-union module _ {l1 l2 l3 l4 l5 l6 l7 : Level} @@ -181,5 +178,5 @@ module _ soundness-modal-logic-union-same-class : soundness (modal-logic (union-subtype ax₁ ax₂)) C soundness-modal-logic-union-same-class = - soundness-deduction C soundness-union-same-class + soundness-modal-logic C soundness-union-same-class ``` From 2ddac8ec270bbb134d51233c3c63989925737be6 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Mon, 6 Nov 2023 19:56:54 +0300 Subject: [PATCH 24/63] Refactor axioms --- src/modal-logic/axioms.lagda.md | 212 +++++++++++++----------------- src/modal-logic/formulas.lagda.md | 9 ++ 2 files changed, 100 insertions(+), 121 deletions(-) diff --git a/src/modal-logic/axioms.lagda.md b/src/modal-logic/axioms.lagda.md index a584e97ddd..ee0c527aaf 100644 --- a/src/modal-logic/axioms.lagda.md +++ b/src/modal-logic/axioms.lagda.md @@ -7,11 +7,13 @@ module modal-logic.axioms where
Imports ```agda -open import foundation.conjunction +open import foundation.action-on-identifications-binary-functions +open import foundation.action-on-identifications-functions open import foundation.contractible-types open import foundation.coproduct-types open import foundation.dependent-pair-types open import foundation.empty-types +open import foundation.equality-dependent-pair-types open import foundation.existential-quantification open import foundation.function-types open import foundation.identity-types @@ -44,47 +46,98 @@ module _ where ax-k : formulas l i - ax-k (a₁ ⇒ b ⇒ a₂) = Id-formula-Prop a₁ a₂ - {-# CATCHALL #-} - ax-k _ = raise-empty-Prop l + pr1 (ax-k f) = Σ (formula i) (λ a → Σ (formula i) (λ b → f = a ⇒ b ⇒ a)) + pr2 (ax-k (a ⇒ b ⇒ a)) (_ , _ , refl) = + is-prop-is-contr + ( is-contr-Σ-is-prop + ( a) + ( b , refl) + ( λ x y → + ( is-prop-is-contr + ( is-contr-Σ-is-prop + ( b) + ( ap _ ((eq-implication-left ∘ pr2) y)) + ( λ _ → is-set-formula i _ _) + ( λ _ → eq-implication-left ∘ eq-implication-right)) + ( y))) + ( λ _ → eq-implication-left ∘ pr2)) + ( a , b , refl) ax-s : formulas l i - ax-s ((a₁ ⇒ b₁ ⇒ c₁) ⇒ (a₂ ⇒ b₂) ⇒ a₃ ⇒ c₂) = - Id-formula-Prop a₁ a₂ ∧ - Id-formula-Prop a₂ a₃ ∧ - Id-formula-Prop b₁ b₂ ∧ - Id-formula-Prop c₁ c₂ - {-# CATCHALL #-} - ax-s _ = raise-empty-Prop l + pr1 (ax-s f) = + Σ + ( formula i) + ( λ a → + ( Σ + ( formula i) + ( λ b → Σ (formula i) (λ c → f = (a ⇒ b ⇒ c) ⇒ (a ⇒ b) ⇒ a ⇒ c)))) + pr2 (ax-s ((a ⇒ b ⇒ c) ⇒ (a ⇒ b) ⇒ a ⇒ c)) (_ , _ , _ , refl) = + is-prop-is-contr + ( is-contr-Σ-is-prop + ( a) + ( b , c , refl) + ( λ x y → + ( is-prop-is-contr + ( is-contr-Σ-is-prop + ( pr1 y) + ( pr2 y) + ( λ z e → + ( is-prop-is-contr + ( is-contr-Σ-is-prop + ( c) + ( ap-binary _ + ( eq-implication-left (eq-implication-left (pr2 (pr2 y)))) + ( eq-implication-left + ( eq-implication-right (eq-implication-left (pr2 e))))) + ( λ _ → is-set-formula i _ _) + ( λ _ → + ( eq-implication-right ∘ + eq-implication-right ∘ + eq-implication-right))) + ( e))) + ( λ z e → + _∙_ + ( inv + ( eq-implication-left + ( eq-implication-right + ( eq-implication-left (pr2 (pr2 y)))))) + ( eq-implication-left + ( eq-implication-right + ( eq-implication-left (pr2 e)))))) + ( y))) + ( λ _ → eq-implication-left ∘ eq-implication-left ∘ pr2 ∘ pr2)) + ( a , b , c , refl) ax-n : formulas l i - ax-n (□ (a₁ ⇒ b₁) ⇒ □ a₂ ⇒ □ b₂) = - Id-formula-Prop a₁ a₂ ∧ - Id-formula-Prop b₁ b₂ - {-# CATCHALL #-} - ax-n _ = raise-empty-Prop l + pr1 (ax-n f) = + Σ (formula i) (λ a → Σ (formula i) (λ b → f = □ (a ⇒ b) ⇒ □ a ⇒ □ b)) + pr2 (ax-n (□ (a ⇒ b) ⇒ □ a ⇒ □ b)) (_ , _ , refl) = + is-prop-is-contr + ( is-contr-Σ-is-prop + ( a) + ( b , refl) + ( λ x y → + ( is-prop-is-contr + ( is-contr-Σ-is-prop + ( b) + ( ap _ + ( eq-box (eq-implication-left (eq-implication-right (pr2 y))))) + ( λ _ → is-set-formula i _ _) + ( λ _ → eq-box ∘ eq-implication-right ∘ eq-implication-right)) + ( y))) + ( λ _ → eq-box ∘ eq-implication-left ∘ eq-implication-right ∘ pr2)) + ( a , b , refl) ax-dn : formulas l i - ax-dn (((a₁ ⇒ ⊥) ⇒ ⊥) ⇒ a₂) = Id-formula-Prop a₁ a₂ - {-# CATCHALL #-} - ax-dn _ = raise-empty-Prop l - - ax-dn' : formulas l i - ax-dn' f = ∃-Prop (formula i) (λ a → f = (~~ a ⇒ a)) - - ax-dn'' : formulas l i - pr1 (ax-dn'' f) = Σ (formula i) (λ a → f = (~~ a ⇒ a)) - pr2 (ax-dn'' (((a₁ ⇒ ⊥) ⇒ ⊥) ⇒ a₁)) (a₁ , refl) = + pr1 (ax-dn f) = Σ (formula i) (λ a → f = ~~ a ⇒ a) + pr2 (ax-dn (((a ⇒ ⊥) ⇒ ⊥) ⇒ a)) (_ , refl) = is-prop-is-contr ( is-contr-Σ-is-prop - ( a₁) + ( a) ( refl) ( λ _ → is-set-formula i _ _) - ( λ _ → arrow-right-eq)) - ( a₁ , refl) - where - arrow-right-eq : {a b x y : formula i} → x ⇒ a = y ⇒ b → a = b - arrow-right-eq refl = refl + ( λ _ → eq-implication-right)) + ( a , refl) module _ {l1 l2 : Level} @@ -94,108 +147,25 @@ module _ where ax-k-soundness : soundness (ax-k i) (all-models w l3 i l4) - ax-k-soundness (a₁ ⇒ b ⇒ a₁) refl M in-class x fa _ = fa - ax-k-soundness (var _) (map-raise ()) - ax-k-soundness ⊥ (map-raise ()) - ax-k-soundness (□ _) (map-raise ()) - ax-k-soundness (_ ⇒ var _) (map-raise ()) - ax-k-soundness (_ ⇒ ⊥) (map-raise ()) - ax-k-soundness (_ ⇒ □ _) (map-raise ()) + ax-k-soundness (a ⇒ b ⇒ a) (_ , _ , refl) M _ x fa _ = fa ax-n-soundness : soundness (ax-n i) (all-models w l3 i l4) ax-n-soundness - (□ (a₁ ⇒ b₁) ⇒ □ a₁ ⇒ □ b₁) - (refl , refl) + (□ (a ⇒ b) ⇒ □ a ⇒ □ b) + (_ , _ , refl) M in-class x fab fa y r = fab y r (fa y r) - ax-n-soundness (var _) (map-raise ()) - ax-n-soundness ⊥ (map-raise ()) - ax-n-soundness (□ _) (map-raise ()) - ax-n-soundness (var _ ⇒ _) (map-raise ()) - ax-n-soundness (⊥ ⇒ _) (map-raise ()) - ax-n-soundness ((_ ⇒ _) ⇒ _) (map-raise ()) - ax-n-soundness (□ var _ ⇒ _) (map-raise ()) - ax-n-soundness (□ ⊥ ⇒ _) (map-raise ()) - ax-n-soundness (□ □ _ ⇒ _) (map-raise ()) - ax-n-soundness (□ (_ ⇒ _) ⇒ var _) (map-raise ()) - ax-n-soundness (□ (_ ⇒ _) ⇒ ⊥) (map-raise ()) - ax-n-soundness (□ (_ ⇒ _) ⇒ □ _) (map-raise ()) - ax-n-soundness (□ (_ ⇒ _) ⇒ var _ ⇒ _) (map-raise ()) - ax-n-soundness (□ (_ ⇒ _) ⇒ ⊥ ⇒ _) (map-raise ()) - ax-n-soundness (□ (_ ⇒ _) ⇒ (_ ⇒ _) ⇒ _) (map-raise ()) - ax-n-soundness (□ (_ ⇒ _) ⇒ □ _ ⇒ var _) (map-raise ()) - ax-n-soundness (□ (_ ⇒ _) ⇒ □ _ ⇒ ⊥) (map-raise ()) - ax-n-soundness (□ (_ ⇒ _) ⇒ □ _ ⇒ _ ⇒ _) (map-raise ()) ax-s-soundness : soundness (ax-s i) (all-models w l3 i l4) ax-s-soundness ((a ⇒ b ⇒ c) ⇒ (a ⇒ b) ⇒ a ⇒ c) - (refl , refl , refl , refl) + (_ , _ , _ , refl) M in-class x fabc fab fa = fabc fa (fab fa) - ax-s-soundness (var _) (map-raise ()) - ax-s-soundness ⊥ (map-raise ()) - ax-s-soundness (□ _) (map-raise ()) - ax-s-soundness (var _ ⇒ _) (map-raise ()) - ax-s-soundness (⊥ ⇒ _) (map-raise ()) - ax-s-soundness (□ _ ⇒ _) (map-raise ()) - ax-s-soundness ((_ ⇒ var _) ⇒ _) (map-raise ()) - ax-s-soundness ((_ ⇒ ⊥) ⇒ _) (map-raise ()) - ax-s-soundness ((_ ⇒ □ _) ⇒ _) (map-raise ()) - ax-s-soundness ((_ ⇒ _ ⇒ _) ⇒ var _) (map-raise ()) - ax-s-soundness ((_ ⇒ _ ⇒ _) ⇒ ⊥) (map-raise ()) - ax-s-soundness ((_ ⇒ _ ⇒ _) ⇒ □ _) (map-raise ()) - ax-s-soundness ((_ ⇒ _ ⇒ _) ⇒ var _ ⇒ _) (map-raise ()) - ax-s-soundness ((_ ⇒ _ ⇒ _) ⇒ ⊥ ⇒ _) (map-raise ()) - ax-s-soundness ((_ ⇒ _ ⇒ _) ⇒ □ _ ⇒ _) (map-raise ()) - ax-s-soundness ((_ ⇒ _ ⇒ _) ⇒ (_ ⇒ _) ⇒ var _) (map-raise ()) - ax-s-soundness ((_ ⇒ _ ⇒ _) ⇒ (_ ⇒ _) ⇒ ⊥) (map-raise ()) - ax-s-soundness ((_ ⇒ _ ⇒ _) ⇒ (_ ⇒ _) ⇒ □ _) (map-raise ()) ax-dn-soundness : soundness (ax-dn i) (decidable-models w l3 i l4) - ax-dn-soundness (((a ⇒ ⊥) ⇒ ⊥) ⇒ a) refl M is-dec x f with is-dec a x - ... | inl fa = fa - ... | inr fna = raise-ex-falso _ (f (λ fa → map-raise (fna fa))) - ax-dn-soundness (var _) (map-raise ()) - ax-dn-soundness ⊥ (map-raise ()) - ax-dn-soundness (□ _) (map-raise ()) - ax-dn-soundness (var _ ⇒ _) (map-raise ()) - ax-dn-soundness (⊥ ⇒ _) (map-raise ()) - ax-dn-soundness (□ _ ⇒ _) (map-raise ()) - ax-dn-soundness ((var _ ⇒ _) ⇒ _) (map-raise ()) - ax-dn-soundness ((⊥ ⇒ _) ⇒ _) (map-raise ()) - ax-dn-soundness ((□ _ ⇒ _) ⇒ _) (map-raise ()) - ax-dn-soundness (((_ ⇒ var _) ⇒ _) ⇒ _) (map-raise ()) - ax-dn-soundness (((_ ⇒ _ ⇒ _) ⇒ _) ⇒ _) (map-raise ()) - ax-dn-soundness (((_ ⇒ □ _) ⇒ _) ⇒ _) (map-raise ()) - ax-dn-soundness (((_ ⇒ ⊥) ⇒ var _) ⇒ _) (map-raise ()) - ax-dn-soundness (((_ ⇒ ⊥) ⇒ _ ⇒ _) ⇒ _) (map-raise ()) - ax-dn-soundness (((_ ⇒ ⊥) ⇒ □ _) ⇒ _) (map-raise ()) - - ax-dn-soundness' : soundness (ax-dn' i) (decidable-models w l3 i l4) - ax-dn-soundness' a = - map-universal-property-trunc-Prop - ( decidable-models w l3 i l4 ⊨C a) - ( aux) - -- with inductor - -- ( λ { - -- (b , refl) M is-dec x f → - -- ( ind-coprod - -- ( λ _ → type-Prop ((M , x) ⊨ b)) - -- ( id) - -- ( λ fna → raise-ex-falso _ (f (λ fa -> map-raise (fna fa)))) - -- ( is-dec b x))}) - where - aux : - Σ (formula i) (λ b → a = ((b ⇒ ⊥) ⇒ ⊥) ⇒ b) → - type-Prop (decidable-models w l3 i l4 ⊨C a) - aux (b , refl) M is-dec x f with (is-dec b x) - ... | inl fa = fa - ... | inr fna = raise-ex-falso _ (f (λ fa -> map-raise (fna fa))) - - ax-dn-soundness'' : soundness (ax-dn'' i) (decidable-models w l3 i l4) - ax-dn-soundness'' - (((a ⇒ ⊥) ⇒ ⊥) ⇒ a) (a , refl) _ is-dec x f with (is-dec a x) + ax-dn-soundness (((a ⇒ ⊥) ⇒ ⊥) ⇒ a) (_ , refl) _ is-dec x f + with (is-dec a x) ... | inl fa = fa ... | inr fna = raise-ex-falso _ (f (λ fa -> map-raise (fna fa))) ``` diff --git a/src/modal-logic/formulas.lagda.md b/src/modal-logic/formulas.lagda.md index 21297041e9..d8d41e0d68 100644 --- a/src/modal-logic/formulas.lagda.md +++ b/src/modal-logic/formulas.lagda.md @@ -72,6 +72,15 @@ module _ ◇_ : formula i → formula i ◇ a = ~ □ ~ a + + eq-implication-left : {a b c d : formula i} → a ⇒ b = c ⇒ d → a = c + eq-implication-left refl = refl + + eq-implication-right : {a b c d : formula i} → a ⇒ b = c ⇒ d → b = d + eq-implication-right refl = refl + + eq-box : {a b : formula i} → □ a = □ b → a = b + eq-box refl = refl ``` ## Properties From e17134df3650da0820cc6a4327ee71d1115e633d Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Fri, 10 Nov 2023 14:05:59 +0300 Subject: [PATCH 25/63] Add uncompleted proof of completeness K --- src/foundation/subtypes.lagda.md | 20 + src/foundation/unions-subtypes.lagda.md | 26 +- src/modal-logic.lagda.md | 2 + src/modal-logic/axioms.lagda.md | 12 +- .../classical-completeness-K.lagda.md | 561 ++++++++++++++++++ src/modal-logic/logic-syntax.lagda.md | 28 +- src/modal-logic/weak-deduction.lagda.md | 300 ++++++++++ src/order-theory.lagda.md | 1 + .../maximal-elements-posets.lagda.md | 47 ++ 9 files changed, 985 insertions(+), 12 deletions(-) create mode 100644 src/modal-logic/classical-completeness-K.lagda.md create mode 100644 src/modal-logic/weak-deduction.lagda.md create mode 100644 src/order-theory/maximal-elements-posets.lagda.md diff --git a/src/foundation/subtypes.lagda.md b/src/foundation/subtypes.lagda.md index c91fc2e0d4..b3a5265b3f 100644 --- a/src/foundation/subtypes.lagda.md +++ b/src/foundation/subtypes.lagda.md @@ -15,6 +15,7 @@ open import foundation.equality-dependent-function-types open import foundation.fundamental-theorem-of-identity-types open import foundation.logical-equivalences open import foundation.propositional-extensionality +open import foundation.raising-universe-levels open import foundation.universe-levels open import foundation-core.cartesian-product-types @@ -143,10 +144,16 @@ module _ {l1 : Level} {A : UU l1} where + --- TODO: replace with equiv-fam equiv-subtypes : {l2 l3 : Level} (P : subtype l2 A) (Q : subtype l3 A) → UU (l1 ⊔ l2 ⊔ l3) equiv-subtypes P Q = (x : A) → is-in-subtype P x ≃ is-in-subtype Q x + inv-equiv-subtypes : + {l2 l3 : Level} (P : subtype l2 A) (Q : subtype l3 A) → + equiv-subtypes P Q → equiv-subtypes Q P + inv-equiv-subtypes P Q e x = inv-equiv (e x) + equiv-antisymmetric-leq-subtype : {l2 l3 : Level} (P : subtype l2 A) (Q : subtype l3 A) → P ⊆ Q → Q ⊆ P → equiv-subtypes P Q @@ -178,6 +185,19 @@ pr1 (subtype-Set l2 A) = subtype l2 A pr2 (subtype-Set l2 A) = is-set-subtype ``` +### TODO + +```agda +raise-subtype : + {l1 l2 : Level} {A : UU l1} (l3 : Level) → subtype l2 A → subtype (l2 ⊔ l3) A +raise-subtype l3 P x = raise-Prop l3 (P x) + +compute-raise-subtype : + {l1 l2 : Level} {A : UU l1} (l3 : Level) (S : subtype l2 A) → + equiv-subtypes S (raise-subtype l3 S) +compute-raise-subtype l3 S x = compute-raise l3 (type-Prop (S x)) +``` + ### Characterisation of embeddings into subtypes ```agda diff --git a/src/foundation/unions-subtypes.lagda.md b/src/foundation/unions-subtypes.lagda.md index 18f2c00960..fafc5cc44c 100644 --- a/src/foundation/unions-subtypes.lagda.md +++ b/src/foundation/unions-subtypes.lagda.md @@ -10,13 +10,14 @@ module foundation.unions-subtypes where open import foundation.decidable-subtypes open import foundation.dependent-pair-types open import foundation.disjunction +open import foundation.function-types +open import foundation.identity-types open import foundation.large-locale-of-subtypes open import foundation.powersets open import foundation.propositional-truncations +open import foundation.subtypes open import foundation.universe-levels -open import foundation-core.subtypes - open import order-theory.least-upper-bounds-large-posets ``` @@ -101,4 +102,25 @@ module _ subtype-union-right : Q ⊆ union-subtype P Q subtype-union-right x = inr-disj-Prop (P x) (Q x) + + subtype-union-both : + {l3 : Level} (S : subtype l3 X) → P ⊆ S → Q ⊆ S → union-subtype P Q ⊆ S + subtype-union-both S P-sub-S Q-sub-S x = + elim-disj-Prop (P x) (Q x) (S x) (P-sub-S x , Q-sub-S x) + +module _ + {l1 l2 : Level} {X : UU l1} (P : subtype l2 X) + where + + subtype-union-same : union-subtype P P ⊆ P + subtype-union-same = + subtype-union-both P P P (refl-leq-subtype P) (refl-leq-subtype P) + + eq-union-same : P = union-subtype P P + eq-union-same = + antisymmetric-leq-subtype + ( P) + ( union-subtype P P) + ( subtype-union-left P P) + ( subtype-union-same) ``` diff --git a/src/modal-logic.lagda.md b/src/modal-logic.lagda.md index 03b7657257..57a79b9506 100644 --- a/src/modal-logic.lagda.md +++ b/src/modal-logic.lagda.md @@ -4,6 +4,7 @@ module modal-logic where open import modal-logic.axioms public +open import modal-logic.classical-completeness-K public open import modal-logic.completeness public open import modal-logic.formulas public open import modal-logic.kripke-semantics public @@ -11,4 +12,5 @@ open import modal-logic.logic-syntax public open import modal-logic.modal-logic-IK public open import modal-logic.modal-logic-K public open import modal-logic.soundness public +open import modal-logic.weak-deduction public ``` diff --git a/src/modal-logic/axioms.lagda.md b/src/modal-logic/axioms.lagda.md index ee0c527aaf..94970e46d5 100644 --- a/src/modal-logic/axioms.lagda.md +++ b/src/modal-logic/axioms.lagda.md @@ -96,14 +96,14 @@ module _ eq-implication-right))) ( e))) ( λ z e → - _∙_ - ( inv + ( _∙_ + ( inv + ( eq-implication-left + ( eq-implication-right + ( eq-implication-left (pr2 (pr2 y)))))) ( eq-implication-left ( eq-implication-right - ( eq-implication-left (pr2 (pr2 y)))))) - ( eq-implication-left - ( eq-implication-right - ( eq-implication-left (pr2 e)))))) + ( eq-implication-left (pr2 e))))))) ( y))) ( λ _ → eq-implication-left ∘ eq-implication-left ∘ pr2 ∘ pr2)) ( a , b , c , refl) diff --git a/src/modal-logic/classical-completeness-K.lagda.md b/src/modal-logic/classical-completeness-K.lagda.md new file mode 100644 index 0000000000..4cde9ac3c0 --- /dev/null +++ b/src/modal-logic/classical-completeness-K.lagda.md @@ -0,0 +1,561 @@ +# Classical proof of completeness K + +```agda +module modal-logic.classical-completeness-K where +``` + +
Imports + +```agda +open import foundation.cartesian-product-types +open import foundation.coproduct-types +open import foundation.dependent-pair-types +open import foundation.disjunction +open import foundation.empty-types +open import foundation.existential-quantification +open import foundation.function-types +open import foundation.identity-types +open import foundation.inhabited-types +open import foundation.intersections-subtypes +open import foundation.law-of-excluded-middle +open import foundation.logical-equivalences +open import foundation.negation +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.raising-universe-levels +open import foundation.sets +open import foundation.subtypes +open import foundation.transport-along-identifications +open import foundation.unit-type +open import foundation.unions-subtypes +open import foundation.universe-levels + +open import order-theory.maximal-elements-posets +open import order-theory.posets +open import order-theory.preorders +open import order-theory.subposets +open import order-theory.chains-posets +open import order-theory.top-elements-posets + +open import modal-logic.axioms +open import modal-logic.completeness +open import modal-logic.formulas +open import modal-logic.kripke-semantics +open import modal-logic.logic-syntax +open import modal-logic.soundness +open import modal-logic.weak-deduction +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l1 : Level} (l2 : Level) (A : UU l1) + where + + subtypes-leq-Preorder : Preorder (l1 ⊔ lsuc l2) (l1 ⊔ l2) + pr1 subtypes-leq-Preorder = subtype l2 A + pr1 (pr2 subtypes-leq-Preorder) = leq-prop-subtype + pr1 (pr2 (pr2 subtypes-leq-Preorder)) = refl-leq-subtype + pr2 (pr2 (pr2 subtypes-leq-Preorder)) = transitive-leq-subtype + + subtypes-leq-Poset : Poset (l1 ⊔ lsuc l2) (l1 ⊔ l2) + pr1 subtypes-leq-Poset = subtypes-leq-Preorder + pr2 subtypes-leq-Poset = antisymmetric-leq-subtype + +module _ + {l1 l2 : Level} {i : Set l1} (axioms : formulas l2 i) + (L-is-cons : is-consistent-modal-logic (weak-modal-logic axioms)) + (contains-ax-k : ax-k i ⊆ weak-modal-logic axioms) + (contains-ax-s : ax-s i ⊆ weak-modal-logic axioms) + (contains-ax-dn : ax-dn i ⊆ weak-modal-logic axioms) + where + + is-L-consistent-theory-Prop : formulas (l1 ⊔ l2) i → Prop (l1 ⊔ l2) + is-L-consistent-theory-Prop t = + is-consistent-modal-logic-Prop (weak-modal-logic (union-subtype axioms t)) + + is-L-consistent-theory : formulas (l1 ⊔ l2) i → UU (l1 ⊔ l2) + is-L-consistent-theory = type-Prop ∘ is-L-consistent-theory-Prop + + L-consistent-theory : UU (lsuc l1 ⊔ lsuc l2) + L-consistent-theory = Σ (formulas (l1 ⊔ l2) i) is-L-consistent-theory + + L-consistent-theory-leq-Prop : + L-consistent-theory → L-consistent-theory → Prop (l1 ⊔ l2) + L-consistent-theory-leq-Prop x y = leq-prop-subtype (pr1 x) (pr1 y) + + L-consistent-theory-leq : L-consistent-theory → L-consistent-theory → UU (l1 ⊔ l2) + L-consistent-theory-leq x y = type-Prop (L-consistent-theory-leq-Prop x y) + + theories-Poset : Poset (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) + theories-Poset = subtypes-leq-Poset (l1 ⊔ l2) (formula i) + + L-consistent-theories-Poset : Poset (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) + L-consistent-theories-Poset = + poset-Subposet theories-Poset is-L-consistent-theory-Prop + + is-L-complete-theory-Prop : L-consistent-theory → Prop (lsuc l1 ⊔ lsuc l2) + is-L-complete-theory-Prop = + is-maximal-element-Poset-Prop L-consistent-theories-Poset + + is-L-complete-theory : L-consistent-theory → UU (lsuc l1 ⊔ lsuc l2) + is-L-complete-theory = type-Prop ∘ is-L-complete-theory-Prop + + complete-theory-closed : + (x : L-consistent-theory) → is-L-complete-theory x → + {a : formula i} → pr1 x ▷ a → is-in-subtype (pr1 x) a + complete-theory-closed x is-comp {a} wd = + tr + ( λ y → is-in-subtype (pr1 y) a) + ( is-comp + ( theory-add-formula a (pr1 x) + , λ bot-in-logic → + ( pr2 x + ( subset-weak-modal-logic-if-subset-axioms + ( subtype-union-both + ( axioms) + ( theory-add-formula a (pr1 x)) + ( weak-modal-logic (union-subtype axioms (pr1 x))) + ( transitive-leq-subtype + ( axioms) + ( union-subtype axioms (pr1 x)) + ( weak-modal-logic (union-subtype axioms (pr1 x))) + ( axioms-subset-weak-modal-logic + ( union-subtype axioms (pr1 x))) + ( subtype-union-left axioms (pr1 x))) + ( subtype-union-both + ( Id-formula-Prop a) + ( pr1 x) + ( weak-modal-logic (union-subtype axioms (pr1 x))) + ( λ { .a refl → + ( weak-modal-logic-monotic + ( subtype-union-right axioms (pr1 x)) + ( a) + ( unit-trunc-Prop wd)) }) + ( transitive-leq-subtype + ( pr1 x) + ( union-subtype axioms (pr1 x)) + ( weak-modal-logic (union-subtype axioms (pr1 x))) + ( axioms-subset-weak-modal-logic + (union-subtype axioms (pr1 x))) + ( subtype-union-right axioms (pr1 x))))) + ( ⊥) + ( bot-in-logic)))) + ( subtype-union-right (Id-formula-Prop a) (pr1 x))) + ( unit-trunc-Prop (inl refl)) + + weak-modal-logic-subset-complete-theory : + (x : L-consistent-theory) → is-L-complete-theory x → + weak-modal-logic (pr1 x) ⊆ pr1 x + weak-modal-logic-subset-complete-theory x is-comp a = + map-universal-property-trunc-Prop + ( pr1 x a) + ( complete-theory-closed x is-comp) + + deduction-axioms-L-complete-theory : + (x : L-consistent-theory) → is-L-complete-theory x → + {a : formula i} → axioms ▷ a → is-in-subtype (pr1 x) a + deduction-axioms-L-complete-theory x is-comp {a} wd = + tr + ( λ y → is-in-subtype (pr1 y) a) + ( is-comp + ( theory-add-formula a (pr1 x) + , λ bot-in-logic → + ( pr2 x + ( subset-weak-modal-logic-if-subset-axioms + ( subtype-union-both + ( axioms) + ( theory-add-formula a (pr1 x)) + ( weak-modal-logic (union-subtype axioms (pr1 x))) + ( transitive-leq-subtype + ( axioms) + ( union-subtype axioms (pr1 x)) + ( weak-modal-logic (union-subtype axioms (pr1 x))) + ( axioms-subset-weak-modal-logic + ( union-subtype axioms (pr1 x))) + ( subtype-union-left axioms (pr1 x))) + ( subtype-union-both + ( Id-formula-Prop a) + ( pr1 x) + ( weak-modal-logic (union-subtype axioms (pr1 x))) + ( λ { .a refl → + ( weak-modal-logic-monotic + ( subtype-union-left axioms (pr1 x)) + ( a) + ( unit-trunc-Prop wd)) }) + ( transitive-leq-subtype + ( pr1 x) + ( union-subtype axioms (pr1 x)) + ( weak-modal-logic (union-subtype axioms (pr1 x))) + ( axioms-subset-weak-modal-logic + (union-subtype axioms (pr1 x))) + ( subtype-union-right axioms (pr1 x))))) + ( ⊥) + ( bot-in-logic)))) + ( subtype-union-right (Id-formula-Prop a) (pr1 x))) + ( subtype-union-left (Id-formula-Prop a) (pr1 x) a refl) + + logic-subset-L-complete-theory : + (x : L-consistent-theory) → is-L-complete-theory x → + weak-modal-logic (axioms) ⊆ pr1 x + logic-subset-L-complete-theory x is-comp a = + map-universal-property-trunc-Prop + ( pr1 x a) + ( deduction-axioms-L-complete-theory x is-comp) + + axioms-subset-L-complete-theory : + (x : L-consistent-theory) → is-L-complete-theory x → axioms ⊆ pr1 x + axioms-subset-L-complete-theory x is-comp = + transitive-leq-subtype + ( axioms) + ( weak-modal-logic axioms) + ( pr1 x) + ( logic-subset-L-complete-theory x is-comp) + ( axioms-subset-weak-modal-logic axioms) + + complete-theory-contains-all-formulas : + LEM (l1 ⊔ l2) → + (x : L-consistent-theory) → is-L-complete-theory x → + (a : formula i) → type-disj-Prop ((pr1 x) a) ((pr1 x) (~ a)) + complete-theory-contains-all-formulas lem x is-comp a with lem ((pr1 x) a) + ... | inl a-in-logic = inl-disj-Prop ((pr1 x) a) ((pr1 x) (~ a)) a-in-logic + ... | inr a-not-in-logic = + unit-trunc-Prop + ( inr + ( tr + ( λ y → is-in-subtype (pr1 y) (~ a)) + ( is-comp + ( theory-add-formula (~ a) (pr1 x) + , λ bot-in-logic → + ( a-not-in-logic + ( weak-modal-logic-subset-complete-theory + ( x) + ( is-comp) + ( a) + ( apply-universal-property-trunc-Prop + ( forward-implication + ( deduction-lemma + ( pr1 x) + ( transitive-leq-subtype + ( ax-k i) + ( weak-modal-logic (axioms)) + ( weak-modal-logic (pr1 x)) + ( weak-modal-logic-monotic + ( axioms-subset-L-complete-theory x is-comp)) + ( contains-ax-k)) + ( transitive-leq-subtype + ( ax-s i) + ( weak-modal-logic (axioms)) + ( weak-modal-logic (pr1 x)) + ( weak-modal-logic-monotic + ( axioms-subset-L-complete-theory x is-comp)) + ( contains-ax-s)) + ( ~ a) + ( ⊥)) + ( weak-modal-logic-monotic + { ax₁ = union-subtype axioms (theory-add-formula (~ a) (pr1 x))} + { ax₂ = theory-add-formula (~ a) (pr1 x)} + ( subtype-union-both + ( axioms) + ( theory-add-formula (~ a) (pr1 x)) + ( theory-add-formula (~ a) (pr1 x)) + ( transitive-leq-subtype + ( axioms) + ( pr1 x) + ( theory-add-formula (~ a) (pr1 x)) + ( subtype-union-right + ( Id-formula-Prop (~ a)) + ( pr1 x)) + ( axioms-subset-L-complete-theory x is-comp)) + ( refl-leq-subtype + ( theory-add-formula (~ a) (pr1 x)))) + ( ⊥) + ( bot-in-logic))) + ( weak-modal-logic (pr1 x) a) + ( λ wd-bot → + ( unit-trunc-Prop + ( w-mp + ( w-ax + ( logic-subset-L-complete-theory + ( x) + ( is-comp) + ( ~~ a ⇒ a) + ( contains-ax-dn _ (a , refl)))) + ( wd-bot)))))))) + ( subtype-union-right (Id-formula-Prop (~ a)) (pr1 x))) + ( subtype-union-left (Id-formula-Prop (~ a)) (pr1 x) (~ a) refl))) + -- tr + -- ( λ y → type-disj-Prop ((pr1 y) a) ((pr1 y) (~ a))) + -- ( is-comp + -- ( theory-add-formula (~ a) (pr1 x) + -- , λ bot-in-logic → + -- ( {! !})) + -- ( subtype-union-right (Id-formula-Prop (~ a)) (pr1 x))) + -- ( unit-trunc-Prop + -- ( inr (subtype-union-left (Id-formula-Prop (~ a)) (pr1 x) (~ a) refl))) + + complete-theory-implication : + LEM (l1 ⊔ l2) → + (x : L-consistent-theory) → is-L-complete-theory x → + {a b : formula i} → (is-in-subtype (pr1 x) a → is-in-subtype (pr1 x) b) → + is-in-subtype (pr1 x) (a ⇒ b) + complete-theory-implication lem x is-comp {a} {b} imp = + apply-universal-property-trunc-Prop + ( complete-theory-contains-all-formulas lem x is-comp a) + ( (pr1 x) (a ⇒ b)) + ( λ { (inl a-in-logic) → + ( weak-modal-logic-subset-complete-theory + ( x) + ( is-comp) + ( a ⇒ b) + ( forward-implication + ( deduction-lemma + ( pr1 x) + ( transitive-leq-subtype + ( ax-k i) + ( weak-modal-logic axioms) + ( weak-modal-logic (pr1 x)) + ( transitive-leq-subtype + ( weak-modal-logic axioms) + ( pr1 x) + ( weak-modal-logic (pr1 x)) + ( axioms-subset-weak-modal-logic (pr1 x)) + ( logic-subset-L-complete-theory x is-comp)) + ( contains-ax-k)) + ( transitive-leq-subtype + ( ax-s i) + ( weak-modal-logic axioms) + ( weak-modal-logic (pr1 x)) + ( transitive-leq-subtype + ( weak-modal-logic axioms) + ( pr1 x) + ( weak-modal-logic (pr1 x)) + ( axioms-subset-weak-modal-logic (pr1 x)) + ( logic-subset-L-complete-theory x is-comp)) + ( contains-ax-s)) + ( a) + ( b)) + ( weak-modal-logic-monotic + { ax₁ = pr1 x} -- TODO : make explicit + { ax₂ = union-subtype (Id-formula-Prop a) (pr1 x)} + ( subtype-union-right (Id-formula-Prop a) (pr1 x)) + ( b) + ( unit-trunc-Prop (w-ax (imp a-in-logic)))))) + ; (inr not-a-in-logic) → + ( weak-modal-logic-subset-complete-theory + ( x) + ( is-comp) + ( a ⇒ b) + ( forward-implication + ( deduction-lemma + ( pr1 x) + ( transitive-leq-subtype + ( ax-k i) + ( weak-modal-logic axioms) + ( weak-modal-logic (pr1 x)) + ( transitive-leq-subtype + ( weak-modal-logic axioms) + ( pr1 x) + ( weak-modal-logic (pr1 x)) + ( axioms-subset-weak-modal-logic (pr1 x)) + ( logic-subset-L-complete-theory x is-comp)) + ( contains-ax-k)) + ( transitive-leq-subtype + ( ax-s i) + ( weak-modal-logic axioms) + ( weak-modal-logic (pr1 x)) + ( transitive-leq-subtype + ( weak-modal-logic axioms) + ( pr1 x) + ( weak-modal-logic (pr1 x)) + ( axioms-subset-weak-modal-logic (pr1 x)) + ( logic-subset-L-complete-theory x is-comp)) + ( contains-ax-s)) + ( a) + ( b)) + ( logic-ex-falso + ( theory-add-formula a (pr1 x)) + ( transitive-leq-subtype + ( ax-k i) + ( weak-modal-logic (pr1 x)) + ( weak-modal-logic (theory-add-formula a (pr1 x))) + ( weak-modal-logic-monotic + ( subtype-union-right (Id-formula-Prop a) (pr1 x))) + ( transitive-leq-subtype + ( ax-k i) + ( weak-modal-logic axioms) + ( weak-modal-logic (pr1 x)) + ( transitive-leq-subtype + ( weak-modal-logic axioms) + ( pr1 x) + ( weak-modal-logic (pr1 x)) + ( axioms-subset-weak-modal-logic (pr1 x)) + ( logic-subset-L-complete-theory x is-comp)) + ( contains-ax-k))) + ( transitive-leq-subtype + ( ax-s i) + ( weak-modal-logic (pr1 x)) + ( weak-modal-logic (theory-add-formula a (pr1 x))) + ( weak-modal-logic-monotic + ( subtype-union-right (Id-formula-Prop a) (pr1 x))) + ( transitive-leq-subtype + ( ax-s i) + ( weak-modal-logic axioms) + ( weak-modal-logic (pr1 x)) + ( transitive-leq-subtype + ( weak-modal-logic axioms) + ( pr1 x) + ( weak-modal-logic (pr1 x)) + ( axioms-subset-weak-modal-logic (pr1 x)) + ( logic-subset-L-complete-theory x is-comp)) + ( contains-ax-s))) + ( transitive-leq-subtype + ( ax-dn i) + ( weak-modal-logic (pr1 x)) + ( weak-modal-logic (theory-add-formula a (pr1 x))) + ( weak-modal-logic-monotic + ( subtype-union-right (Id-formula-Prop a) (pr1 x))) + ( transitive-leq-subtype + ( ax-dn i) + ( weak-modal-logic axioms) + ( weak-modal-logic (pr1 x)) + ( transitive-leq-subtype + ( weak-modal-logic axioms) + ( pr1 x) + ( weak-modal-logic (pr1 x)) + ( axioms-subset-weak-modal-logic (pr1 x)) + ( logic-subset-L-complete-theory x is-comp)) + ( contains-ax-dn))) + ( a) + ( b) + ( weak-modal-logic-monotic + { ax₁ = Id-formula-Prop a} + { ax₂ = union-subtype (Id-formula-Prop a) (pr1 x)} + ( subtype-union-left (Id-formula-Prop a) (pr1 x)) + ( a) + ( unit-trunc-Prop (w-ax refl))) + ( weak-modal-logic-monotic + { ax₁ = pr1 x} + { ax₂ = union-subtype (Id-formula-Prop a) (pr1 x)} + ( subtype-union-right (Id-formula-Prop a) (pr1 x)) + ( ~ a) + ( unit-trunc-Prop (w-ax not-a-in-logic)))))) }) + + canonical-worlds : subtype (lsuc l1 ⊔ lsuc l2) (formulas (l1 ⊔ l2) i) + canonical-worlds x = + Σ-Prop + ( is-L-consistent-theory-Prop x) + ( λ is-cons → is-L-complete-theory-Prop (x , is-cons)) + + canonical-kripke-model-world-type : UU (lsuc l1 ⊔ lsuc l2) + canonical-kripke-model-world-type = + Σ (formulas (l1 ⊔ l2) i) (is-in-subtype canonical-worlds) + + postulate + lindenbaum : + (x : L-consistent-theory) → + ∃ ( L-consistent-theory) + ( λ y → L-consistent-theory-leq x y × is-L-complete-theory y) + + is-L-consistent-axioms : is-L-consistent-theory (raise-subtype l1 axioms) + is-L-consistent-axioms = + map-universal-property-trunc-Prop + ( empty-Prop) + ( tr + ( λ x → x ▷ ⊥ → empty) + ( helper) + ( λ x → + ( L-is-cons + ( unit-trunc-Prop + ( deduction-equiv-axioms + ( inv-equiv-subtypes + ( axioms) + ( raise-subtype l1 axioms) + ( compute-raise-subtype l1 axioms)) + ( x)))))) + where + -- TODO: move to unions-subtypes as lemma + helper : + raise-subtype l1 axioms = union-subtype axioms (raise-subtype l1 axioms) + helper = + antisymmetric-leq-subtype + ( raise-subtype l1 axioms) + ( union-subtype axioms (raise-subtype l1 axioms)) + ( subtype-union-right axioms (raise-subtype l1 axioms)) + ( subtype-union-both + ( axioms) + ( raise-subtype l1 axioms) + ( raise-subtype l1 axioms) + ( λ _ → map-raise) + ( λ _ → id)) + + is-inhabited-canonical-kripke-model-world : + is-inhabited canonical-kripke-model-world-type + is-inhabited-canonical-kripke-model-world = + apply-universal-property-trunc-Prop + ( lindenbaum (raise-subtype l1 axioms , is-L-consistent-axioms)) + ( is-inhabited-Prop canonical-kripke-model-world-type) + ( λ x → unit-trunc-Prop ((pr1 (pr1 x)) , ((pr2 (pr1 x)) , (pr2 (pr2 x))))) + + canonical-kripke-model : + kripke-model (canonical-kripke-model-world-type) (l1 ⊔ l2) i (l1 ⊔ l2) + pr1 (pr1 canonical-kripke-model) = + is-inhabited-canonical-kripke-model-world + pr2 (pr1 canonical-kripke-model) x y = + Π-Prop + ( formula i) + ( λ a → function-Prop (is-in-subtype (pr1 x) (□ a)) ((pr1 y) a)) + pr2 canonical-kripke-model n x = pr1 x (var n) + + module _ + (lem : LEM (l1 ⊔ l2)) + where + + canonical-model-theorem : + (a : formula i) + (x : canonical-kripke-model-world-type) → + pr1 x a ⇔ ((canonical-kripke-model , x) ⊨ a) + pr1 (canonical-model-theorem (var n) x) in-logic = + map-raise in-logic + pr1 (canonical-model-theorem ⊥ x) in-logic = + map-raise + ( pr1 + ( pr2 x) + ( unit-trunc-Prop + ( w-ax (subtype-union-right axioms (pr1 x) ⊥ in-logic)))) + pr1 (canonical-model-theorem (a ⇒ b) x) in-logic fa = + pr1 + ( canonical-model-theorem b x) + ( weak-modal-logic-subset-complete-theory + ( pr1 x , pr1 (pr2 x)) + ( pr2 (pr2 x)) + ( b) + ( unit-trunc-Prop + ( w-mp + ( w-ax in-logic) + ( w-ax (backward-implication (canonical-model-theorem a x) fa))))) + pr1 (canonical-model-theorem (□ a) x) in-logic y xRy = + forward-implication (canonical-model-theorem a y) (xRy a in-logic) + pr2 (canonical-model-theorem (var n) x) f = + map-inv-raise f + pr2 (canonical-model-theorem ⊥ x) (map-raise ()) + pr2 (canonical-model-theorem (a ⇒ b) x) f = + complete-theory-implication + ( lem) + ( pr1 x , pr1 (pr2 x)) + ( pr2 (pr2 x)) + ( λ a-in-logic → + ( backward-implication + ( canonical-model-theorem b x) + ( f + ( forward-implication (canonical-model-theorem a x) a-in-logic)))) + pr2 (canonical-model-theorem (□ a) x) f = + {! !} +``` diff --git a/src/modal-logic/logic-syntax.lagda.md b/src/modal-logic/logic-syntax.lagda.md index 35f9733276..c96a37ac5e 100644 --- a/src/modal-logic/logic-syntax.lagda.md +++ b/src/modal-logic/logic-syntax.lagda.md @@ -12,10 +12,12 @@ open import foundation.dependent-pair-types open import foundation.empty-types open import foundation.function-types open import foundation.identity-types +open import foundation.negation open import foundation.propositional-truncations open import foundation.propositions open import foundation.sets open import foundation.subtypes +open import foundation.unions-subtypes open import foundation.unit-type open import foundation.universe-levels @@ -38,6 +40,13 @@ module _ formulas : UU (l1 ⊔ lsuc l2) formulas = subtype l2 (formula i) +module _ + {l1 l2 : Level} {i : Set l1} + where + + theory-add-formula : formula i → formulas l2 i → formulas (l1 ⊔ l2) i + theory-add-formula a = union-subtype (Id-formula-Prop a) + module _ {l1 l2 : Level} {i : Set l1} where @@ -52,6 +61,18 @@ module _ modal-logic : formulas l2 i → formulas (l1 ⊔ l2) i modal-logic axioms a = trunc-Prop (axioms ⊢ a) + is-contradictory-modal-logic-Prop : formulas l2 i → Prop l2 + is-contradictory-modal-logic-Prop logic = logic ⊥ + + is-contradictory-modal-logic : formulas l2 i → UU l2 + is-contradictory-modal-logic = type-Prop ∘ is-contradictory-modal-logic-Prop + + is-consistent-modal-logic-Prop : formulas l2 i → Prop l2 + is-consistent-modal-logic-Prop = neg-Prop ∘ is-contradictory-modal-logic-Prop + + is-consistent-modal-logic : formulas l2 i → UU l2 + is-consistent-modal-logic = type-Prop ∘ is-consistent-modal-logic-Prop + module _ {l1 : Level} {i : Set l1} where @@ -63,7 +84,7 @@ module _ modal-logic-closed : {l2 : Level} {axioms : formulas l2 i} {a : formula i} → modal-logic axioms ⊢ a → - type-Prop (modal-logic axioms a) + is-in-subtype (modal-logic axioms) a modal-logic-closed (ax x) = x modal-logic-closed {axioms = axioms} {a} (mp tdab tda) = apply-twice-universal-property-trunc-Prop @@ -116,10 +137,9 @@ module _ {l1 l2 l3 : Level} {i : Set l1} {ax₁ : formulas l2 i} {ax₂ : formulas l3 i} where - -- TODO: change name - modal-logic-CHOOSE-NAME : + subset-modal-logic-if-subset-axioms : ax₁ ⊆ modal-logic ax₂ → modal-logic ax₁ ⊆ modal-logic ax₂ - modal-logic-CHOOSE-NAME leq = + subset-modal-logic-if-subset-axioms leq = transitive-leq-subtype ( modal-logic ax₁) ( modal-logic (modal-logic ax₂)) diff --git a/src/modal-logic/weak-deduction.lagda.md b/src/modal-logic/weak-deduction.lagda.md new file mode 100644 index 0000000000..1d7b70804d --- /dev/null +++ b/src/modal-logic/weak-deduction.lagda.md @@ -0,0 +1,300 @@ +# Weak deduction + +```agda +module modal-logic.weak-deduction where +``` + +
Imports + +```agda +open import foundation.coproduct-types +open import foundation.dependent-pair-types +open import foundation.disjunction +open import foundation.empty-types +open import foundation.equivalences +open import foundation.function-types +open import foundation.identity-types +open import foundation.logical-equivalences +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.sets +open import foundation.subtypes +open import foundation.unit-type +open import foundation.unions-subtypes +open import foundation.universe-levels + +open import modal-logic.axioms +open import modal-logic.formulas +open import modal-logic.logic-syntax +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l1 l2 : Level} {i : Set l1} + where + + infix 5 _▷_ + + data _▷_ (axioms : formulas l2 i) : formula i → UU (l1 ⊔ l2) where + w-ax : {a : formula i} → is-in-subtype axioms a → axioms ▷ a + w-mp : {a b : formula i} → axioms ▷ a ⇒ b → axioms ▷ a → axioms ▷ b + + weak-modal-logic : formulas l2 i → formulas (l1 ⊔ l2) i + weak-modal-logic axioms a = trunc-Prop (axioms ▷ a) + +module _ + {l1 l2 : Level} {i : Set l1} (axioms : formulas l2 i) + where + + deduction-weak-deduction : {a : formula i} → axioms ▷ a → axioms ⊢ a + deduction-weak-deduction (w-ax x) = ax x + deduction-weak-deduction (w-mp dab da) = + mp (deduction-weak-deduction dab) (deduction-weak-deduction da) + + weak-modal-logic-subset-modal-logic : + weak-modal-logic axioms ⊆ modal-logic axioms + weak-modal-logic-subset-modal-logic a = + map-universal-property-trunc-Prop + ( modal-logic axioms a) + ( unit-trunc-Prop ∘ deduction-weak-deduction) + +module _ + {l1 l2 l3 : Level} {i : Set l1} + {axioms₁ : formulas l2 i} {axioms₂ : formulas l3 i} + (e : equiv-subtypes axioms₁ axioms₂) + where + + deduction-equiv-axioms : + {a : formula i} → axioms₁ ▷ a → axioms₂ ▷ a + deduction-equiv-axioms {a} (w-ax x) = + w-ax (map-equiv (e a) x) + deduction-equiv-axioms (w-mp wdab wda) = + w-mp (deduction-equiv-axioms wdab) (deduction-equiv-axioms wda) + +module _ + {l1 : Level} {i : Set l1} + where + + axioms-subset-weak-modal-logic : + {l2 : Level} (axioms : formulas l2 i) → axioms ⊆ weak-modal-logic axioms + axioms-subset-weak-modal-logic _ a H = unit-trunc-Prop (w-ax H) + + weak-modal-logic-closed : + {l2 : Level} {axioms : formulas l2 i} {a : formula i} → + weak-modal-logic axioms ▷ a → + is-in-subtype (weak-modal-logic axioms) a + weak-modal-logic-closed (w-ax x) = x + weak-modal-logic-closed {axioms = axioms} {a} (w-mp tdab tda) = + apply-twice-universal-property-trunc-Prop + ( weak-modal-logic-closed tdab) + ( weak-modal-logic-closed tda) + ( weak-modal-logic axioms a) + (λ dab da → unit-trunc-Prop (w-mp dab da)) + + subset-double-weak-modal-logic : + {l2 : Level} + (axioms : formulas l2 i) → + weak-modal-logic (weak-modal-logic axioms) ⊆ weak-modal-logic axioms + subset-double-weak-modal-logic axioms a = + map-universal-property-trunc-Prop + ( weak-modal-logic axioms a) + ( weak-modal-logic-closed) + + weak-modal-logic-idempotent : + {l2 : Level} + (axioms : formulas l2 i) → + weak-modal-logic axioms = weak-modal-logic (weak-modal-logic axioms) + weak-modal-logic-idempotent axioms = + antisymmetric-leq-subtype _ _ + ( axioms-subset-weak-modal-logic (weak-modal-logic axioms)) + ( subset-double-weak-modal-logic axioms) + +module _ + {l1 l2 l3 : Level} {i : Set l1} {ax₁ : formulas l2 i} {ax₂ : formulas l3 i} + (leq : ax₁ ⊆ ax₂) + where + + weak-deduction-monotic : {a : formula i} → ax₁ ▷ a → ax₂ ▷ a + weak-deduction-monotic (w-ax x) = w-ax (leq _ x) + weak-deduction-monotic (w-mp wdab wda) = + w-mp (weak-deduction-monotic wdab) (weak-deduction-monotic wda) + + weak-modal-logic-monotic : weak-modal-logic ax₁ ⊆ weak-modal-logic ax₂ + weak-modal-logic-monotic a = + map-universal-property-trunc-Prop + ( weak-modal-logic ax₂ a) + ( unit-trunc-Prop ∘ weak-deduction-monotic) + +module _ + {l1 l2 l3 : Level} {i : Set l1} {ax₁ : formulas l2 i} {ax₂ : formulas l3 i} + where + + subset-weak-modal-logic-if-subset-axioms : + ax₁ ⊆ weak-modal-logic ax₂ → weak-modal-logic ax₁ ⊆ weak-modal-logic ax₂ + subset-weak-modal-logic-if-subset-axioms leq = + transitive-leq-subtype + ( weak-modal-logic ax₁) + ( weak-modal-logic (weak-modal-logic ax₂)) + ( weak-modal-logic ax₂) + ( subset-double-weak-modal-logic ax₂) + ( weak-modal-logic-monotic leq) + +module _ + {l1 l2 : Level} {i : Set l1} (axioms : formulas l2 i) + where + + backward-deduction-lemma : + {a b : formula i} → + axioms ▷ a ⇒ b → + is-in-subtype (weak-modal-logic (theory-add-formula a axioms)) (b) + backward-deduction-lemma {a} wab = + unit-trunc-Prop + ( w-mp + ( weak-deduction-monotic + ( subtype-union-right ((Id-formula-Prop a)) axioms) wab) + ( w-ax (subtype-union-left (Id-formula-Prop a) axioms a refl))) + + module _ + -- TODO: maybe just in axioms? + (contains-ax-k : ax-k i ⊆ weak-modal-logic axioms) + (contains-ax-s : ax-s i ⊆ weak-modal-logic axioms) + where + + deduction-a->a : + (a : formula i) → is-in-subtype (weak-modal-logic axioms) (a ⇒ a) + deduction-a->a a = + apply-three-times-universal-property-trunc-Prop + ( contains-ax-s _ (a , a ⇒ a , a , refl)) + ( contains-ax-k _ (a , a ⇒ a , refl)) + ( contains-ax-k _ (a , a , refl)) + ( (weak-modal-logic axioms) (a ⇒ a)) + ( λ x y z → unit-trunc-Prop (w-mp (w-mp x y) z)) + + forward-deduction-lemma : + (a : formula i) {b : formula i} → + theory-add-formula a axioms ▷ b → + is-in-subtype (weak-modal-logic axioms) (a ⇒ b) + forward-deduction-lemma a {b} (w-ax x) = + elim-disj-Prop + ( Id-formula-Prop a b) + ( axioms b) + ( (weak-modal-logic axioms) (a ⇒ b)) + ( (λ { refl → deduction-a->a a }) + , (λ in-axioms → + ( apply-universal-property-trunc-Prop + ( contains-ax-k _ (b , a , refl)) + ( (weak-modal-logic axioms) (a ⇒ b)) + ( λ x → unit-trunc-Prop (w-mp x (w-ax in-axioms)))))) + ( x) + forward-deduction-lemma a {b} (w-mp {c} twdcb twdc) = + apply-three-times-universal-property-trunc-Prop + ( forward-deduction-lemma a twdcb) + ( forward-deduction-lemma a twdc) + ( contains-ax-s _ (a , c , b , refl)) + ( (weak-modal-logic axioms) (a ⇒ b)) + ( λ wdacb wdac x → + ( unit-trunc-Prop (w-mp (w-mp x wdacb) wdac))) + + deduction-lemma : + (a b : formula i) → + weak-modal-logic (theory-add-formula a axioms) b ⇔ + weak-modal-logic axioms (a ⇒ b) + pr1 (deduction-lemma a b) = + map-universal-property-trunc-Prop + ( (weak-modal-logic axioms) (a ⇒ b)) + ( forward-deduction-lemma a) + pr2 (deduction-lemma a b) = + map-universal-property-trunc-Prop + ( weak-modal-logic (theory-add-formula a axioms) b) + ( backward-deduction-lemma) + + +module _ + {l1 l2 : Level} {i : Set l1} (axioms : formulas l2 i) + (contains-ax-k : ax-k i ⊆ weak-modal-logic axioms) + (contains-ax-s : ax-s i ⊆ weak-modal-logic axioms) + (contains-ax-dn : ax-dn i ⊆ weak-modal-logic axioms) + where + + deduction-ex-falso : + (a b : formula i) → + is-in-subtype (weak-modal-logic axioms) (~ a ⇒ a ⇒ b) + deduction-ex-falso a b = + forward-implication + ( deduction-lemma axioms contains-ax-k contains-ax-s (~ a) (a ⇒ b)) + ( forward-implication + ( deduction-lemma + ( theory-add-formula (~ a) axioms) + ( transitive-leq-subtype + ( ax-k i) + ( weak-modal-logic axioms) + ( weak-modal-logic (theory-add-formula (a ⇒ ⊥) axioms)) + ( weak-modal-logic-monotic + ( subtype-union-right (Id-formula-Prop (~ a)) axioms)) + ( contains-ax-k)) + ( transitive-leq-subtype + ( ax-s i) + ( weak-modal-logic axioms) + ( weak-modal-logic (theory-add-formula (a ⇒ ⊥) axioms)) + ( weak-modal-logic-monotic + ( subtype-union-right (Id-formula-Prop (~ a)) axioms)) + ( contains-ax-s)) + ( a) + ( b)) + ( apply-twice-universal-property-trunc-Prop + ( contains-ax-k _ (⊥ , ~ b , refl)) + ( contains-ax-dn _ ((b , refl))) + ( weak-modal-logic + ( theory-add-formula a (theory-add-formula (~ a) axioms)) b) + ( λ x y → + ( unit-trunc-Prop + ( w-mp {a = ~~ b} + ( weak-deduction-monotic + ( transitive-leq-subtype + ( axioms) + ( theory-add-formula (~ a) axioms) + ( theory-add-formula a + ( theory-add-formula (~ a) axioms)) + ( subtype-union-right + ( Id-formula-Prop a) + ( theory-add-formula (~ a) axioms)) + ( subtype-union-right (Id-formula-Prop (~ a)) axioms)) + ( y)) + ( w-mp {a = ⊥} + ( weak-deduction-monotic + ( transitive-leq-subtype + ( axioms) + ( theory-add-formula (~ a) axioms) + ( theory-add-formula a + ( theory-add-formula (~ a) axioms)) + ( subtype-union-right + ( Id-formula-Prop a) + ( theory-add-formula (~ a) axioms)) + ( subtype-union-right (Id-formula-Prop (~ a)) axioms)) + ( x)) + ( w-mp {a = a} + ( w-ax + ( unit-trunc-Prop (inr (unit-trunc-Prop (inl refl))))) + ( w-ax (unit-trunc-Prop (inl refl)))))))))) + + logic-ex-falso : + (a b : formula i) → + is-in-subtype (weak-modal-logic axioms) a → + is-in-subtype (weak-modal-logic axioms) (~ a) → + is-in-subtype (weak-modal-logic axioms) b + logic-ex-falso a b a-in-logic not-a-in-logic = + apply-three-times-universal-property-trunc-Prop + ( a-in-logic) + ( not-a-in-logic) + ( deduction-ex-falso a b) + ( (weak-modal-logic axioms) b) + ( λ x y z → unit-trunc-Prop (w-mp (w-mp z y) x)) +``` diff --git a/src/order-theory.lagda.md b/src/order-theory.lagda.md index 1ce9b78699..b6ab4908bf 100644 --- a/src/order-theory.lagda.md +++ b/src/order-theory.lagda.md @@ -75,6 +75,7 @@ open import order-theory.lower-sets-large-posets public open import order-theory.lower-types-preorders public open import order-theory.maximal-chains-posets public open import order-theory.maximal-chains-preorders public +open import order-theory.maximal-elements-posets public open import order-theory.meet-semilattices public open import order-theory.meet-suplattices public open import order-theory.nuclei-large-locales public diff --git a/src/order-theory/maximal-elements-posets.lagda.md b/src/order-theory/maximal-elements-posets.lagda.md new file mode 100644 index 0000000000..a60677c230 --- /dev/null +++ b/src/order-theory/maximal-elements-posets.lagda.md @@ -0,0 +1,47 @@ +# Maximal elements in posets + +```agda +module order-theory.maximal-elements-posets where +``` + +
Imports + +```agda +open import foundation.dependent-pair-types +open import foundation.identity-types +open import foundation.propositions +open import foundation.universe-levels + +open import order-theory.posets +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l1 l2 : Level} (X : Poset l1 l2) + where + + is-maximal-element-Poset-Prop : type-Poset X → Prop (l1 ⊔ l2) + is-maximal-element-Poset-Prop x = + Π-Prop + ( type-Poset X) + ( λ y → function-Prop (leq-Poset X x y) (y = x , is-set-type-Poset X y x)) + + is-maximal-element-Poset : type-Poset X → UU (l1 ⊔ l2) + is-maximal-element-Poset x = type-Prop (is-maximal-element-Poset-Prop x) + + is-prop-is-maximal-element-Poset : + (x : type-Poset X) → is-prop (is-maximal-element-Poset x) + is-prop-is-maximal-element-Poset x = + is-prop-type-Prop (is-maximal-element-Poset-Prop x) + + has-maximal-element-Poset : UU (l1 ⊔ l2) + has-maximal-element-Poset = Σ (type-Poset X) is-maximal-element-Poset +``` From c1ae1ffd04d5e427ec639300839b9fc8ee246c28 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Fri, 10 Nov 2023 14:19:20 +0300 Subject: [PATCH 26/63] Run pre-commit --- .../classical-completeness-K.lagda.md | 40 ++++++++----------- src/modal-logic/weak-deduction.lagda.md | 5 +-- 2 files changed, 19 insertions(+), 26 deletions(-) diff --git a/src/modal-logic/classical-completeness-K.lagda.md b/src/modal-logic/classical-completeness-K.lagda.md index 4cde9ac3c0..8cc18c71d2 100644 --- a/src/modal-logic/classical-completeness-K.lagda.md +++ b/src/modal-logic/classical-completeness-K.lagda.md @@ -26,17 +26,10 @@ open import foundation.raising-universe-levels open import foundation.sets open import foundation.subtypes open import foundation.transport-along-identifications -open import foundation.unit-type open import foundation.unions-subtypes +open import foundation.unit-type open import foundation.universe-levels -open import order-theory.maximal-elements-posets -open import order-theory.posets -open import order-theory.preorders -open import order-theory.subposets -open import order-theory.chains-posets -open import order-theory.top-elements-posets - open import modal-logic.axioms open import modal-logic.completeness open import modal-logic.formulas @@ -44,6 +37,13 @@ open import modal-logic.kripke-semantics open import modal-logic.logic-syntax open import modal-logic.soundness open import modal-logic.weak-deduction + +open import order-theory.chains-posets +open import order-theory.maximal-elements-posets +open import order-theory.posets +open import order-theory.preorders +open import order-theory.subposets +open import order-theory.top-elements-posets ```
@@ -91,7 +91,8 @@ module _ L-consistent-theory → L-consistent-theory → Prop (l1 ⊔ l2) L-consistent-theory-leq-Prop x y = leq-prop-subtype (pr1 x) (pr1 y) - L-consistent-theory-leq : L-consistent-theory → L-consistent-theory → UU (l1 ⊔ l2) + L-consistent-theory-leq : + L-consistent-theory → L-consistent-theory → UU (l1 ⊔ l2) L-consistent-theory-leq x y = type-Prop (L-consistent-theory-leq-Prop x y) theories-Poset : Poset (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) @@ -138,7 +139,7 @@ module _ ( weak-modal-logic-monotic ( subtype-union-right axioms (pr1 x)) ( a) - ( unit-trunc-Prop wd)) }) + ( unit-trunc-Prop wd))}) ( transitive-leq-subtype ( pr1 x) ( union-subtype axioms (pr1 x)) @@ -189,7 +190,7 @@ module _ ( weak-modal-logic-monotic ( subtype-union-left axioms (pr1 x)) ( a) - ( unit-trunc-Prop wd)) }) + ( unit-trunc-Prop wd))}) ( transitive-leq-subtype ( pr1 x) ( union-subtype axioms (pr1 x)) @@ -250,7 +251,7 @@ module _ ( weak-modal-logic-monotic ( axioms-subset-L-complete-theory x is-comp)) ( contains-ax-k)) - ( transitive-leq-subtype + ( transitive-leq-subtype ( ax-s i) ( weak-modal-logic (axioms)) ( weak-modal-logic (pr1 x)) @@ -260,7 +261,9 @@ module _ ( ~ a) ( ⊥)) ( weak-modal-logic-monotic - { ax₁ = union-subtype axioms (theory-add-formula (~ a) (pr1 x))} + { ax₁ = + ( union-subtype axioms + ( theory-add-formula (~ a) (pr1 x)))} { ax₂ = theory-add-formula (~ a) (pr1 x)} ( subtype-union-both ( axioms) @@ -291,15 +294,6 @@ module _ ( wd-bot)))))))) ( subtype-union-right (Id-formula-Prop (~ a)) (pr1 x))) ( subtype-union-left (Id-formula-Prop (~ a)) (pr1 x) (~ a) refl))) - -- tr - -- ( λ y → type-disj-Prop ((pr1 y) a) ((pr1 y) (~ a))) - -- ( is-comp - -- ( theory-add-formula (~ a) (pr1 x) - -- , λ bot-in-logic → - -- ( {! !})) - -- ( subtype-union-right (Id-formula-Prop (~ a)) (pr1 x))) - -- ( unit-trunc-Prop - -- ( inr (subtype-union-left (Id-formula-Prop (~ a)) (pr1 x) (~ a) refl))) complete-theory-implication : LEM (l1 ⊔ l2) → @@ -446,7 +440,7 @@ module _ { ax₂ = union-subtype (Id-formula-Prop a) (pr1 x)} ( subtype-union-right (Id-formula-Prop a) (pr1 x)) ( ~ a) - ( unit-trunc-Prop (w-ax not-a-in-logic)))))) }) + ( unit-trunc-Prop (w-ax not-a-in-logic))))))}) canonical-worlds : subtype (lsuc l1 ⊔ lsuc l2) (formulas (l1 ⊔ l2) i) canonical-worlds x = diff --git a/src/modal-logic/weak-deduction.lagda.md b/src/modal-logic/weak-deduction.lagda.md index 1d7b70804d..e786905e04 100644 --- a/src/modal-logic/weak-deduction.lagda.md +++ b/src/modal-logic/weak-deduction.lagda.md @@ -19,8 +19,8 @@ open import foundation.propositional-truncations open import foundation.propositions open import foundation.sets open import foundation.subtypes -open import foundation.unit-type open import foundation.unions-subtypes +open import foundation.unit-type open import foundation.universe-levels open import modal-logic.axioms @@ -187,7 +187,7 @@ module _ ( Id-formula-Prop a b) ( axioms b) ( (weak-modal-logic axioms) (a ⇒ b)) - ( (λ { refl → deduction-a->a a }) + ( (λ { refl → deduction-a->a a}) , (λ in-axioms → ( apply-universal-property-trunc-Prop ( contains-ax-k _ (b , a , refl)) @@ -216,7 +216,6 @@ module _ ( weak-modal-logic (theory-add-formula a axioms) b) ( backward-deduction-lemma) - module _ {l1 l2 : Level} {i : Set l1} (axioms : formulas l2 i) (contains-ax-k : ax-k i ⊆ weak-modal-logic axioms) From b358d85b5facf110e6313664d80a5dbee7e67540 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Mon, 5 Feb 2024 21:16:40 +0300 Subject: [PATCH 27/63] Proof canonical model theorem --- .../law-of-excluded-middle.lagda.md | 22 +- .../propositional-resizing.lagda.md | 49 + src/foundation/unions-subtypes.lagda.md | 133 +- src/lists/lists.lagda.md | 17 + src/lists/reversing-lists.lagda.md | 12 + src/modal-logic.lagda.md | 2 +- src/modal-logic/axioms.lagda.md | 158 ++- .../canonical-model-theorem.lagda.md | 1068 +++++++++++++++++ .../classical-completeness-K.lagda.md | 555 --------- src/modal-logic/logic-syntax.lagda.md | 49 +- src/modal-logic/weak-deduction.lagda.md | 344 +++++- src/order-theory.lagda.md | 2 + src/order-theory/chains-posets.lagda.md | 28 + src/order-theory/chains-union.lagda.md | 121 ++ .../maximal-elements-posets.lagda.md | 80 ++ src/order-theory/subtypes-leq-posets.lagda.md | 39 + 16 files changed, 2020 insertions(+), 659 deletions(-) create mode 100644 src/modal-logic/canonical-model-theorem.lagda.md delete mode 100644 src/modal-logic/classical-completeness-K.lagda.md create mode 100644 src/order-theory/chains-union.lagda.md create mode 100644 src/order-theory/subtypes-leq-posets.lagda.md diff --git a/src/foundation/law-of-excluded-middle.lagda.md b/src/foundation/law-of-excluded-middle.lagda.md index d82d58a594..6a67a25ac9 100644 --- a/src/foundation/law-of-excluded-middle.lagda.md +++ b/src/foundation/law-of-excluded-middle.lagda.md @@ -9,11 +9,12 @@ module foundation.law-of-excluded-middle where ```agda open import foundation.decidable-types open import foundation.dependent-pair-types +open import foundation.propositions +open import foundation.raising-universe-levels open import foundation.universe-levels open import foundation-core.decidable-propositions open import foundation-core.negation -open import foundation-core.propositions open import univalent-combinatorics.2-element-types ``` @@ -29,12 +30,27 @@ The {{#concept "law of excluded middle" Agda=LEM}} asserts that any ## Definition ```agda -LEM : (l : Level) → UU (lsuc l) -LEM l = (P : Prop l) → is-decidable (type-Prop P) +module _ + (l : Level) + where + + LEM-Prop : Prop (lsuc l) + LEM-Prop = Π-Prop (Prop l) is-decidable-Prop + + LEM : UU (lsuc l) + LEM = type-Prop LEM-Prop ``` ## Properties +### TODO + +```agda +lower-LEM : {l1 : Level} (l2 : Level) → LEM (l1 ⊔ l2) → LEM l1 +lower-LEM l2 lem P = + is-decidable-equiv (compute-raise l2 (type-Prop P)) (lem (raise-Prop l2 P)) +``` + ### Given LEM, we obtain a map from the type of propositions to the type of decidable propositions ```agda diff --git a/src/foundation/propositional-resizing.lagda.md b/src/foundation/propositional-resizing.lagda.md index 46d48494aa..c71522f94a 100644 --- a/src/foundation/propositional-resizing.lagda.md +++ b/src/foundation/propositional-resizing.lagda.md @@ -7,9 +7,17 @@ module foundation.propositional-resizing where
Imports ```agda +open import foundation.decidable-types open import foundation.dependent-pair-types +open import foundation.empty-types +open import foundation.law-of-excluded-middle +open import foundation.negation +open import foundation.raising-universe-levels +open import foundation.unit-type open import foundation.universe-levels +open import foundation-core.coproduct-types +open import foundation-core.equivalences open import foundation-core.propositions open import foundation-core.subtypes ``` @@ -35,6 +43,47 @@ propositional-resizing l1 l2 = ( λ Ω → (P : Prop l2) → Σ (pr1 Ω) (λ u → type-equiv-Prop (pr2 Ω u) P)) ``` +## TODO + +```agda +module _ + {l1 l2 : Level} ((Ω , prop-resize) : propositional-resizing l1 l2) + where + + resize : Prop l2 → Prop l1 + resize P = pr2 Ω (pr1 (prop-resize P)) + + is-equiv-resize : (P : Prop l2) → type-equiv-Prop (resize P) P + is-equiv-resize P = pr2 (prop-resize P) + +unit-equiv-true : + {l : Level} (P : Prop l) → type-Prop P → type-equiv-Prop unit-Prop P +pr1 (unit-equiv-true P p) _ = p +pr2 (unit-equiv-true P p) = + is-equiv-is-prop is-prop-unit (is-prop-type-Prop P) (λ _ → star) + +empty-equiv-false : + {l : Level} (P : Prop l) → ¬ (type-Prop P) → type-equiv-Prop empty-Prop P +pr1 (empty-equiv-false P np) = ex-falso +pr2 (empty-equiv-false P np) = + is-equiv-is-prop is-prop-empty (is-prop-type-Prop P) np + +propositional-resizing-LEM : + (l1 : Level) {l2 : Level} → LEM l2 → propositional-resizing l1 l2 +pr1 (pr1 (propositional-resizing-LEM l1 lem)) = raise-unit l1 + raise-unit l1 +pr2 (pr1 (propositional-resizing-LEM l1 lem)) (inl _) = raise-unit-Prop l1 +pr2 (pr1 (propositional-resizing-LEM l1 lem)) (inr _) = raise-empty-Prop l1 +pr2 (propositional-resizing-LEM l1 lem) P with lem P +... | inl p = + pair + ( inl raise-star) + ( unit-equiv-true P p ∘e inv-equiv (compute-raise l1 unit)) +... | inr np = + pair + ( inr raise-star) + ( empty-equiv-false P np ∘e inv-equiv (compute-raise l1 empty)) +``` + ## See also - [The large locale of propositions](foundation.large-locale-of-propositions.md) diff --git a/src/foundation/unions-subtypes.lagda.md b/src/foundation/unions-subtypes.lagda.md index fafc5cc44c..289a680c0e 100644 --- a/src/foundation/unions-subtypes.lagda.md +++ b/src/foundation/unions-subtypes.lagda.md @@ -13,6 +13,7 @@ open import foundation.disjunction open import foundation.function-types open import foundation.identity-types open import foundation.large-locale-of-subtypes +open import foundation.logical-equivalences open import foundation.powersets open import foundation.propositional-truncations open import foundation.subtypes @@ -104,10 +105,140 @@ module _ subtype-union-right x = inr-disj-Prop (P x) (Q x) subtype-union-both : - {l3 : Level} (S : subtype l3 X) → P ⊆ S → Q ⊆ S → union-subtype P Q ⊆ S + {l4 : Level} (S : subtype l4 X) → P ⊆ S → Q ⊆ S → union-subtype P Q ⊆ S subtype-union-both S P-sub-S Q-sub-S x = elim-disj-Prop (P x) (Q x) (S x) (P-sub-S x , Q-sub-S x) +module _ + {l1 l2 l3 : Level} {X : UU l1} (P : subtype l2 X) (Q : subtype l3 X) + where + + subset-union-comm : + union-subtype P Q ⊆ union-subtype Q P + subset-union-comm = + subtype-union-both P Q + ( union-subtype Q P) + ( subtype-union-right Q P) + ( subtype-union-left Q P) + +module _ + {l1 l2 l3 l4 : Level} {X : UU l1} + (P : subtype l2 X) (Q : subtype l3 X) (S : subtype l4 X) + where + + forward-subset-union-assoc : + union-subtype P (union-subtype Q S) ⊆ union-subtype (union-subtype P Q) S + forward-subset-union-assoc = + subtype-union-both + ( P) + ( union-subtype Q S) + ( union-subtype (union-subtype P Q) S) + ( transitive-leq-subtype + ( P) + ( union-subtype P Q) + ( union-subtype (union-subtype P Q) S) + ( subtype-union-left (union-subtype P Q) S) + ( subtype-union-left P Q)) + ( subtype-union-both Q S + ( union-subtype (union-subtype P Q) S) + ( transitive-leq-subtype + ( Q) + ( union-subtype P Q) + ( union-subtype (union-subtype P Q) S) + ( subtype-union-left (union-subtype P Q) S) + ( subtype-union-right P Q)) + ( subtype-union-right (union-subtype P Q) S)) + + backward-subset-union-assoc : + union-subtype (union-subtype P Q) S ⊆ union-subtype P (union-subtype Q S) + backward-subset-union-assoc = + subtype-union-both + ( union-subtype P Q) + ( S) + ( union-subtype P (union-subtype Q S)) + ( subtype-union-both P Q + ( union-subtype P (union-subtype Q S)) + ( subtype-union-left P (union-subtype Q S)) + ( transitive-leq-subtype + ( Q) + ( union-subtype Q S) + ( union-subtype P (union-subtype Q S)) + ( subtype-union-right P (union-subtype Q S)) + ( subtype-union-left Q S))) + ( transitive-leq-subtype + ( S) + ( union-subtype Q S) + ( union-subtype P (union-subtype Q S)) + ( subtype-union-right P (union-subtype Q S)) + ( subtype-union-right Q S)) + +module _ + {l1 : Level} {X : UU l1} + where + + subset-union-subsets : + {l2 l3 l4 l5 : Level} + (P1 : subtype l2 X) (Q1 : subtype l3 X) + (P2 : subtype l4 X) (Q2 : subtype l5 X) → + P1 ⊆ P2 → Q1 ⊆ Q2 → + union-subtype P1 Q1 ⊆ union-subtype P2 Q2 + subset-union-subsets P1 Q1 P2 Q2 P1-sub-P2 Q1-sub-Q2 = + subtype-union-both P1 Q1 (union-subtype P2 Q2) + ( transitive-leq-subtype P1 P2 (union-subtype P2 Q2) + ( subtype-union-left P2 Q2) + ( P1-sub-P2)) + ( transitive-leq-subtype Q1 Q2 (union-subtype P2 Q2) + ( subtype-union-right P2 Q2) + ( Q1-sub-Q2)) + + subset-union-subset-left : + {l2 l3 l4 : Level} + (P1 : subtype l2 X) (P2 : subtype l3 X) (Q : subtype l4 X) → + P1 ⊆ P2 → + union-subtype P1 Q ⊆ union-subtype P2 Q + subset-union-subset-left P1 P2 Q P1-sub-P2 = + subtype-union-both P1 Q (union-subtype P2 Q) + ( transitive-leq-subtype P1 P2 (union-subtype P2 Q) + ( subtype-union-left P2 Q) + ( P1-sub-P2)) + ( subtype-union-right P2 Q) + + subset-union-subset-right : + {l2 l3 l4 : Level} + (P : subtype l2 X) (Q1 : subtype l3 X) (Q2 : subtype l4 X) → + Q1 ⊆ Q2 → + union-subtype P Q1 ⊆ union-subtype P Q2 + subset-union-subset-right P Q1 Q2 Q1-sub-Q2 = + subtype-union-both P Q1 (union-subtype P Q2) + ( subtype-union-left P Q2) + ( transitive-leq-subtype Q1 Q2 (union-subtype P Q2) + ( subtype-union-right P Q2) + ( Q1-sub-Q2)) + +module _ + {l1 l2 l3 l4 : Level} {X : UU l1} + (P : subtype l2 X) (Q : subtype l3 X) (S : subtype l4 X) + where + + union-swap-1-2 : + union-subtype P (union-subtype Q S) ⊆ union-subtype Q (union-subtype P S) + union-swap-1-2 = + transitive-leq-subtype + ( union-subtype P (union-subtype Q S)) + ( union-subtype (union-subtype Q P) S) + ( union-subtype Q (union-subtype P S)) + ( backward-subset-union-assoc Q P S) + ( transitive-leq-subtype + ( union-subtype P (union-subtype Q S)) + ( union-subtype (union-subtype P Q) S) + ( union-subtype (union-subtype Q P) S) + ( subset-union-subset-left + ( union-subtype P Q) + ( union-subtype Q P) + ( S) + ( subset-union-comm P Q)) + ( forward-subset-union-assoc P Q S)) + module _ {l1 l2 : Level} {X : UU l1} (P : subtype l2 X) where diff --git a/src/lists/lists.lagda.md b/src/lists/lists.lagda.md index 3a019086f1..a9f46d3ae1 100644 --- a/src/lists/lists.lagda.md +++ b/src/lists/lists.lagda.md @@ -125,6 +125,23 @@ is-nonnil-is-cons-list l ((a , l') , refl) q = is-nonnil-cons-list a l' q ``` +### TODO: change title + +```agda +last-in-snoc-list : + {l : Level} {A : UU l} (xs : list A) (x : A) → x ∈-list (snoc xs x) +last-in-snoc-list nil x = is-head x nil +last-in-snoc-list (cons y xs) x = + is-in-tail x y (snoc xs x) (last-in-snoc-list xs x) + +rest-in-snoc-list : + {l : Level} {A : UU l} (xs : list A) (x y : A) → + y ∈-list xs → y ∈-list (snoc xs x) +rest-in-snoc-list (cons y xs) x .y (is-head .y .xs) = is-head y (snoc xs x) +rest-in-snoc-list (cons y xs) x z (is-in-tail .z .y .xs in-list) = + is-in-tail z y (snoc xs x) (rest-in-snoc-list xs x z in-list) +``` + ### A list that uses cons is not nil ```agda diff --git a/src/lists/reversing-lists.lagda.md b/src/lists/reversing-lists.lagda.md index 551bed29c5..9de38ce0a7 100644 --- a/src/lists/reversing-lists.lagda.md +++ b/src/lists/reversing-lists.lagda.md @@ -130,4 +130,16 @@ remove-last-element-reverse-list x = ( inv (reverse-reverse-list (remove-last-element-list (reverse-list x)))) ∙ ( ( inv (ap reverse-list (tail-reverse-list (reverse-list x)))) ∙ ( ap (reverse-list ∘ tail-list) (reverse-reverse-list x))) + +forward-contains-reverse-list : + {l1 : Level} {A : UU l1} (x : A) (xs : list A) → + (x ∈-list xs) → (x ∈-list (reverse-list xs)) +forward-contains-reverse-list x (cons .x xs) (is-head .x .xs) = + last-in-snoc-list (reverse-list xs) x +forward-contains-reverse-list x (cons y xs) (is-in-tail .x .y .xs x-in-list) = + rest-in-snoc-list + ( reverse-list xs) + ( y) + ( x) + ( forward-contains-reverse-list x xs x-in-list) ``` diff --git a/src/modal-logic.lagda.md b/src/modal-logic.lagda.md index 57a79b9506..e3e2b011b7 100644 --- a/src/modal-logic.lagda.md +++ b/src/modal-logic.lagda.md @@ -4,7 +4,7 @@ module modal-logic where open import modal-logic.axioms public -open import modal-logic.classical-completeness-K public +open import modal-logic.canonical-model-theorem public open import modal-logic.completeness public open import modal-logic.formulas public open import modal-logic.kripke-semantics public diff --git a/src/modal-logic/axioms.lagda.md b/src/modal-logic/axioms.lagda.md index 94970e46d5..d8e9a8c9d9 100644 --- a/src/modal-logic/axioms.lagda.md +++ b/src/modal-logic/axioms.lagda.md @@ -8,24 +8,30 @@ module modal-logic.axioms where ```agda open import foundation.action-on-identifications-binary-functions +open import foundation.action-on-identifications-dependent-functions open import foundation.action-on-identifications-functions open import foundation.contractible-types -open import foundation.coproduct-types open import foundation.dependent-pair-types open import foundation.empty-types -open import foundation.equality-dependent-pair-types open import foundation.existential-quantification -open import foundation.function-types -open import foundation.identity-types +open import foundation.function-extensionality open import foundation.inhabited-types open import foundation.propositional-truncations -open import foundation.propositions open import foundation.raising-universe-levels -open import foundation.sets -open import foundation.subtypes open import foundation.unit-type open import foundation.universe-levels +open import foundation-core.coproduct-types +open import foundation-core.dependent-identifications +open import foundation-core.equality-dependent-pair-types +open import foundation-core.function-types +open import foundation-core.identity-types +open import foundation-core.injective-maps +open import foundation-core.propositions +open import foundation-core.sets +open import foundation-core.subtypes +open import foundation-core.transport-along-identifications + open import modal-logic.formulas open import modal-logic.kripke-semantics open import modal-logic.logic-syntax @@ -45,99 +51,69 @@ module _ {l : Level} (i : Set l) where - ax-k : formulas l i - pr1 (ax-k f) = Σ (formula i) (λ a → Σ (formula i) (λ b → f = a ⇒ b ⇒ a)) - pr2 (ax-k (a ⇒ b ⇒ a)) (_ , _ , refl) = + ax-1-parameter : (h : formula i → formula i) → is-injective h → formulas l i + pr1 (ax-1-parameter h inj f) = Σ (formula i) (λ a → f = h a) + pr2 (ax-1-parameter h inj f) (a , refl) = + is-prop-is-contr + ( is-contr-Σ-is-prop a refl + ( λ _ → is-set-formula i _ _) + ( λ x → inj)) + ( a , refl) + + ax-2-parameters : + (h : formula i → formula i → formula i) → + ({x x' y y' : formula i} → h x y = h x' y' → x = x') → + ({x x' y y' : formula i} → h x y = h x' y' → y = y') → + formulas l i + pr1 (ax-2-parameters h inj-1 inj-2 f) = + Σ (formula i) (λ a → Σ (formula i) (λ b → f = h a b)) + pr2 (ax-2-parameters h inj-1 inj-2 f) (a , b , refl) = is-prop-is-contr - ( is-contr-Σ-is-prop - ( a) - ( b , refl) - ( λ x y → - ( is-prop-is-contr - ( is-contr-Σ-is-prop - ( b) - ( ap _ ((eq-implication-left ∘ pr2) y)) - ( λ _ → is-set-formula i _ _) - ( λ _ → eq-implication-left ∘ eq-implication-right)) - ( y))) - ( λ _ → eq-implication-left ∘ pr2)) + ( is-contr-Σ-is-prop a (b , refl) + ( λ x → is-prop-type-Prop (ax-1-parameter (h x) inj-2 (h a b))) + ( λ x (y , e) → inj-1 e)) ( a , b , refl) - ax-s : formulas l i - pr1 (ax-s f) = - Σ - ( formula i) - ( λ a → - ( Σ - ( formula i) - ( λ b → Σ (formula i) (λ c → f = (a ⇒ b ⇒ c) ⇒ (a ⇒ b) ⇒ a ⇒ c)))) - pr2 (ax-s ((a ⇒ b ⇒ c) ⇒ (a ⇒ b) ⇒ a ⇒ c)) (_ , _ , _ , refl) = + ax-3-parameters : + (h : formula i → formula i → formula i → formula i) → + ({x x' y y' z z' : formula i} → h x y z = h x' y' z' → x = x') → + ({x x' y y' z z' : formula i} → h x y z = h x' y' z' → y = y') → + ({x x' y y' z z' : formula i} → h x y z = h x' y' z' → z = z') → + formulas l i + pr1 (ax-3-parameters h inj-1 inj-2 inj-3 f) = + Σ ( formula i) + ( λ a → Σ (formula i) (λ b → Σ (formula i) ( λ c → f = h a b c))) + pr2 (ax-3-parameters h inj-1 inj-2 inj-3 f) (a , b , c , refl) = is-prop-is-contr - ( is-contr-Σ-is-prop - ( a) - ( b , c , refl) - ( λ x y → - ( is-prop-is-contr - ( is-contr-Σ-is-prop - ( pr1 y) - ( pr2 y) - ( λ z e → - ( is-prop-is-contr - ( is-contr-Σ-is-prop - ( c) - ( ap-binary _ - ( eq-implication-left (eq-implication-left (pr2 (pr2 y)))) - ( eq-implication-left - ( eq-implication-right (eq-implication-left (pr2 e))))) - ( λ _ → is-set-formula i _ _) - ( λ _ → - ( eq-implication-right ∘ - eq-implication-right ∘ - eq-implication-right))) - ( e))) - ( λ z e → - ( _∙_ - ( inv - ( eq-implication-left - ( eq-implication-right - ( eq-implication-left (pr2 (pr2 y)))))) - ( eq-implication-left - ( eq-implication-right - ( eq-implication-left (pr2 e))))))) - ( y))) - ( λ _ → eq-implication-left ∘ eq-implication-left ∘ pr2 ∘ pr2)) + ( is-contr-Σ-is-prop a (b , c , refl) + ( λ x → is-prop-type-Prop (ax-2-parameters (h x) inj-2 inj-3 (h a b c))) + ( λ x (y , z , e) → inj-1 e)) ( a , b , c , refl) + ax-k : formulas l i + ax-k = + ax-2-parameters + ( λ a b → a ⇒ b ⇒ a) + ( eq-implication-left) + ( eq-implication-left ∘ eq-implication-right) + + ax-s : formulas l i + ax-s = + ax-3-parameters + ( λ a b c → (a ⇒ b ⇒ c) ⇒ (a ⇒ b) ⇒ a ⇒ c) + ( eq-implication-left ∘ eq-implication-left) + ( eq-implication-left ∘ eq-implication-right ∘ eq-implication-left) + ( eq-implication-right ∘ eq-implication-right ∘ eq-implication-left) + ax-n : formulas l i - pr1 (ax-n f) = - Σ (formula i) (λ a → Σ (formula i) (λ b → f = □ (a ⇒ b) ⇒ □ a ⇒ □ b)) - pr2 (ax-n (□ (a ⇒ b) ⇒ □ a ⇒ □ b)) (_ , _ , refl) = - is-prop-is-contr - ( is-contr-Σ-is-prop - ( a) - ( b , refl) - ( λ x y → - ( is-prop-is-contr - ( is-contr-Σ-is-prop - ( b) - ( ap _ - ( eq-box (eq-implication-left (eq-implication-right (pr2 y))))) - ( λ _ → is-set-formula i _ _) - ( λ _ → eq-box ∘ eq-implication-right ∘ eq-implication-right)) - ( y))) - ( λ _ → eq-box ∘ eq-implication-left ∘ eq-implication-right ∘ pr2)) - ( a , b , refl) + ax-n = + ax-2-parameters + ( λ a b → □ (a ⇒ b) ⇒ □ a ⇒ □ b) + ( eq-implication-left ∘ eq-box ∘ eq-implication-left) + ( eq-implication-right ∘ eq-box ∘ eq-implication-left) ax-dn : formulas l i - pr1 (ax-dn f) = Σ (formula i) (λ a → f = ~~ a ⇒ a) - pr2 (ax-dn (((a ⇒ ⊥) ⇒ ⊥) ⇒ a)) (_ , refl) = - is-prop-is-contr - ( is-contr-Σ-is-prop - ( a) - ( refl) - ( λ _ → is-set-formula i _ _) - ( λ _ → eq-implication-right)) - ( a , refl) + ax-dn = ax-1-parameter ( λ a → ~~ a ⇒ a) ( eq-implication-right) module _ {l1 l2 : Level} diff --git a/src/modal-logic/canonical-model-theorem.lagda.md b/src/modal-logic/canonical-model-theorem.lagda.md new file mode 100644 index 0000000000..aedb36fa63 --- /dev/null +++ b/src/modal-logic/canonical-model-theorem.lagda.md @@ -0,0 +1,1068 @@ +# Canonical model theorem + +```agda +module modal-logic.canonical-model-theorem where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.cartesian-product-types +open import foundation.coproduct-types +open import foundation.dependent-pair-types +open import foundation.disjunction +open import foundation.empty-types +open import foundation.existential-quantification +open import foundation.function-types +open import foundation.identity-types +open import foundation.inhabited-types +open import foundation.intersections-subtypes +open import foundation.law-of-excluded-middle +open import foundation.logical-equivalences +open import foundation.negation +open import foundation.propositional-resizing +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.raising-universe-levels +open import foundation.sets +open import foundation.subtypes +open import foundation.transport-along-identifications +open import foundation.unions-subtypes +open import foundation.unit-type +open import foundation.universe-levels + +open import foundation-core.equivalences + +open import lists.lists +open import lists.reversing-lists + +open import modal-logic.axioms +open import modal-logic.completeness +open import modal-logic.formulas +open import modal-logic.kripke-semantics +open import modal-logic.logic-syntax +open import modal-logic.soundness +open import modal-logic.weak-deduction + +open import order-theory.chains-posets +open import order-theory.maximal-elements-posets +open import order-theory.posets +open import order-theory.preorders +open import order-theory.subposets +open import order-theory.subtypes-leq-posets +open import order-theory.top-elements-posets +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l1 l2 : Level} {i : Set l1} (axioms : formulas l2 i) + (zorn : Zorn (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) l2) + (prop-resize : propositional-resizing (l1 ⊔ l2) (lsuc (l1 ⊔ l2))) + (L-is-cons : is-consistent-modal-logic (modal-logic axioms)) + (contains-ax-k : ax-k i ⊆ modal-logic axioms) + (contains-ax-s : ax-s i ⊆ modal-logic axioms) + (contains-ax-dn : ax-dn i ⊆ modal-logic axioms) + (contains-ax-n : ax-n i ⊆ modal-logic axioms) + where + + logic : formulas (l1 ⊔ l2) i + logic = modal-logic axioms + + is-L-consistent-theory-Prop : formulas (l1 ⊔ l2) i → Prop (l1 ⊔ l2) + is-L-consistent-theory-Prop t = + is-consistent-modal-logic-Prop (weak-modal-logic (union-subtype logic t)) + + is-L-consistent-theory : formulas (l1 ⊔ l2) i → UU (l1 ⊔ l2) + is-L-consistent-theory = type-Prop ∘ is-L-consistent-theory-Prop + + L-consistent-theory : UU (lsuc l1 ⊔ lsuc l2) + L-consistent-theory = Σ (formulas (l1 ⊔ l2) i) is-L-consistent-theory + + L-consistent-theory-leq-Prop : + L-consistent-theory → L-consistent-theory → Prop (l1 ⊔ l2) + L-consistent-theory-leq-Prop x y = leq-prop-subtype (pr1 x) (pr1 y) + + L-consistent-theory-leq : + L-consistent-theory → L-consistent-theory → UU (l1 ⊔ l2) + L-consistent-theory-leq x y = type-Prop (L-consistent-theory-leq-Prop x y) + + theories-Poset : Poset (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) + theories-Poset = subtypes-leq-Poset (l1 ⊔ l2) (formula i) + + L-consistent-theories-Poset : Poset (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) + L-consistent-theories-Poset = + poset-Subposet theories-Poset is-L-consistent-theory-Prop + + is-L-complete-theory-Prop : L-consistent-theory → Prop (lsuc l1 ⊔ lsuc l2) + is-L-complete-theory-Prop = + is-maximal-element-Poset-Prop L-consistent-theories-Poset + + is-L-complete-theory : L-consistent-theory → UU (lsuc l1 ⊔ lsuc l2) + is-L-complete-theory = type-Prop ∘ is-L-complete-theory-Prop + + weak-modal-logic-subset-complete-theory : + (x : L-consistent-theory) → is-L-complete-theory x → + weak-modal-logic (pr1 x) ⊆ pr1 x + weak-modal-logic-subset-complete-theory x is-comp a in-logic = + tr + ( λ y → is-in-subtype (pr1 y) a) + ( is-comp + ( pair + ( theory-add-formula a (pr1 x)) + ( λ bot-in-logic → + ( pr2 x + ( subset-weak-modal-logic-if-subset-axioms + ( transitive-leq-subtype + ( union-subtype logic (theory-add-formula a (pr1 x))) + ( theory-add-formula a (union-subtype logic (pr1 x))) + ( weak-modal-logic (union-subtype logic (pr1 x))) + ( subtype-union-both + ( Id-formula-Prop a) + ( union-subtype logic (pr1 x)) + ( weak-modal-logic (union-subtype logic (pr1 x))) + ( λ { .a refl → + ( weak-modal-logic-monotic + ( subtype-union-right logic (pr1 x)) + ( a) + ( in-logic))}) + ( axioms-subset-weak-modal-logic + ( union-subtype logic (pr1 x)))) + ( union-swap-1-2 logic (Id-formula-Prop a) (pr1 x))) + ( ⊥) + ( bot-in-logic))))) + ( subtype-union-right (Id-formula-Prop a) (pr1 x))) + ( unit-trunc-Prop (inl refl)) + + logic-subset-L-complete-theory : + (x : L-consistent-theory) → is-L-complete-theory x → logic ⊆ pr1 x + logic-subset-L-complete-theory x is-comp a in-logic = + tr + ( λ y → is-in-subtype (pr1 y) a) + ( is-comp + ( pair + ( theory-add-formula a (pr1 x)) + ( λ bot-in-logic → + ( pr2 x + ( weak-modal-logic-monotic + ( transitive-leq-subtype + ( union-subtype logic (theory-add-formula a (pr1 x))) + ( union-subtype + ( union-subtype logic (Id-formula-Prop a)) (pr1 x)) + ( union-subtype logic (pr1 x)) + ( subset-union-subset-left + ( union-subtype logic (Id-formula-Prop a)) + ( logic) + ( pr1 x) + ( subtype-union-both + ( logic) + ( Id-formula-Prop a) + ( logic) + ( refl-leq-subtype logic) + ( λ { .a refl → in-logic}))) + ( forward-subset-union-assoc + ( logic) + ( Id-formula-Prop a) + ( pr1 x))) + ( ⊥) + ( bot-in-logic))))) + ( subtype-union-right (Id-formula-Prop a) (pr1 x))) + ( subtype-union-left (Id-formula-Prop a) (pr1 x) a refl) + + complete-theory-contains-all-formulas : + LEM (l1 ⊔ l2) → + (x : L-consistent-theory) → is-L-complete-theory x → + (a : formula i) → type-disj-Prop ((pr1 x) a) ((pr1 x) (~ a)) + complete-theory-contains-all-formulas lem x is-comp a with lem ((pr1 x) a) + ... | inl a-in-logic = inl-disj-Prop ((pr1 x) a) ((pr1 x) (~ a)) a-in-logic + ... | inr a-not-in-logic = + unit-trunc-Prop + ( inr + ( tr + ( λ y → is-in-subtype (pr1 y) (~ a)) + ( is-comp + ( theory-add-formula (~ a) (pr1 x) + , λ bot-in-logic → + ( a-not-in-logic + ( weak-modal-logic-subset-complete-theory + ( x) + ( is-comp) + ( a) + ( apply-universal-property-trunc-Prop + ( forward-implication + ( deduction-lemma + ( pr1 x) + ( transitive-leq-subtype + ( ax-k i) + ( logic) + ( weak-modal-logic (pr1 x)) + ( transitive-leq-subtype + ( logic) + ( pr1 x) + ( weak-modal-logic (pr1 x)) + ( axioms-subset-weak-modal-logic (pr1 x)) + ( logic-subset-L-complete-theory x is-comp)) + ( contains-ax-k)) + ( transitive-leq-subtype + ( ax-s i) + ( logic) + ( weak-modal-logic (pr1 x)) + ( transitive-leq-subtype + ( logic) + ( pr1 x) + ( weak-modal-logic (pr1 x)) + ( axioms-subset-weak-modal-logic (pr1 x)) + ( logic-subset-L-complete-theory x is-comp)) + ( contains-ax-s)) + ( ~ a) + ( ⊥)) + ( weak-modal-logic-monotic + { ax₁ = + ( union-subtype logic + ( theory-add-formula (~ a) (pr1 x)))} + { ax₂ = theory-add-formula (~ a) (pr1 x)} + ( subtype-union-both + ( logic) + ( theory-add-formula (~ a) (pr1 x)) + ( theory-add-formula (~ a) (pr1 x)) + ( transitive-leq-subtype + ( logic) + ( pr1 x) + ( theory-add-formula (~ a) (pr1 x)) + ( subtype-union-right + ( Id-formula-Prop (~ a)) + ( pr1 x)) + ( logic-subset-L-complete-theory x is-comp)) + ( refl-leq-subtype + (theory-add-formula (~ a) (pr1 x)))) + ( ⊥) + ( bot-in-logic))) + ( weak-modal-logic (pr1 x) a) + ( λ wd-bot → + ( unit-trunc-Prop + ( w-mp + ( w-ax + ( logic-subset-L-complete-theory + ( x) + ( is-comp) + ( ~~ a ⇒ a) + ( contains-ax-dn _ (a , refl)))) + ( wd-bot)))))))) + ( subtype-union-right (Id-formula-Prop (~ a)) (pr1 x))) + ( subtype-union-left (Id-formula-Prop (~ a)) (pr1 x) (~ a) refl))) + + complete-theory-implication : + LEM (l1 ⊔ l2) → + (x : L-consistent-theory) → is-L-complete-theory x → + {a b : formula i} → (is-in-subtype (pr1 x) a → is-in-subtype (pr1 x) b) → + is-in-subtype (pr1 x) (a ⇒ b) + complete-theory-implication lem x is-comp {a} {b} imp = + apply-universal-property-trunc-Prop + ( complete-theory-contains-all-formulas lem x is-comp a) + ( pr1 x (a ⇒ b)) + ( λ { (inl a-in-logic) → + ( weak-modal-logic-subset-complete-theory + ( x) + ( is-comp) + ( a ⇒ b) + ( forward-implication + ( deduction-lemma + ( pr1 x) + ( transitive-leq-subtype + ( ax-k i) + ( logic) + ( weak-modal-logic (pr1 x)) + ( transitive-leq-subtype + ( logic) + ( pr1 x) + ( weak-modal-logic (pr1 x)) + ( axioms-subset-weak-modal-logic (pr1 x)) + ( logic-subset-L-complete-theory x is-comp)) + ( contains-ax-k)) + ( transitive-leq-subtype + ( ax-s i) + ( logic) + ( weak-modal-logic (pr1 x)) + ( transitive-leq-subtype + ( logic) + ( pr1 x) + ( weak-modal-logic (pr1 x)) + ( axioms-subset-weak-modal-logic (pr1 x)) + ( logic-subset-L-complete-theory x is-comp)) + ( contains-ax-s)) + ( a) + ( b)) + ( weak-modal-logic-monotic + { ax₁ = pr1 x} -- TODO: make explicit + { ax₂ = union-subtype (Id-formula-Prop a) (pr1 x)} + ( subtype-union-right (Id-formula-Prop a) (pr1 x)) + ( b) + ( unit-trunc-Prop (w-ax (imp a-in-logic)))))) + ; (inr not-a-in-logic) → + ( weak-modal-logic-subset-complete-theory + ( x) + ( is-comp) + ( a ⇒ b) + ( forward-implication + ( deduction-lemma + ( pr1 x) + ( transitive-leq-subtype + ( ax-k i) + ( logic) + ( weak-modal-logic (pr1 x)) + ( transitive-leq-subtype + ( logic) + ( pr1 x) + ( weak-modal-logic (pr1 x)) + ( axioms-subset-weak-modal-logic (pr1 x)) + ( logic-subset-L-complete-theory x is-comp)) + ( contains-ax-k)) + ( transitive-leq-subtype + ( ax-s i) + ( logic) + ( weak-modal-logic (pr1 x)) + ( transitive-leq-subtype + ( logic) + ( pr1 x) + ( weak-modal-logic (pr1 x)) + ( axioms-subset-weak-modal-logic (pr1 x)) + ( logic-subset-L-complete-theory x is-comp)) + ( contains-ax-s)) + ( a) + ( b)) + ( logic-ex-falso + ( theory-add-formula a (pr1 x)) + ( transitive-leq-subtype + ( ax-k i) + ( weak-modal-logic (pr1 x)) + ( weak-modal-logic (theory-add-formula a (pr1 x))) + ( weak-modal-logic-monotic + ( subtype-union-right (Id-formula-Prop a) (pr1 x))) + ( transitive-leq-subtype + ( ax-k i) + ( logic) + ( weak-modal-logic (pr1 x)) + ( transitive-leq-subtype + ( logic) + ( pr1 x) + ( weak-modal-logic (pr1 x)) + ( axioms-subset-weak-modal-logic (pr1 x)) + ( logic-subset-L-complete-theory x is-comp)) + ( contains-ax-k))) + ( transitive-leq-subtype + ( ax-s i) + ( weak-modal-logic (pr1 x)) + ( weak-modal-logic (theory-add-formula a (pr1 x))) + ( weak-modal-logic-monotic + ( subtype-union-right (Id-formula-Prop a) (pr1 x))) + ( transitive-leq-subtype + ( ax-s i) + ( logic) + ( weak-modal-logic (pr1 x)) + ( transitive-leq-subtype + ( logic) + ( pr1 x) + ( weak-modal-logic (pr1 x)) + ( axioms-subset-weak-modal-logic (pr1 x)) + ( logic-subset-L-complete-theory x is-comp)) + ( contains-ax-s))) + ( transitive-leq-subtype + ( ax-dn i) + ( weak-modal-logic (pr1 x)) + ( weak-modal-logic (theory-add-formula a (pr1 x))) + ( weak-modal-logic-monotic + ( subtype-union-right (Id-formula-Prop a) (pr1 x))) + ( transitive-leq-subtype + ( ax-dn i) + ( logic) + ( weak-modal-logic (pr1 x)) + ( transitive-leq-subtype + ( logic) + ( pr1 x) + ( weak-modal-logic (pr1 x)) + ( axioms-subset-weak-modal-logic (pr1 x)) + ( logic-subset-L-complete-theory x is-comp)) + ( contains-ax-dn))) + ( a) + ( b) + ( weak-modal-logic-monotic + { ax₁ = Id-formula-Prop a} + { ax₂ = union-subtype (Id-formula-Prop a) (pr1 x)} + ( subtype-union-left (Id-formula-Prop a) (pr1 x)) + ( a) + ( unit-trunc-Prop (w-ax refl))) + ( weak-modal-logic-monotic + { ax₁ = pr1 x} + { ax₂ = union-subtype (Id-formula-Prop a) (pr1 x)} + ( subtype-union-right (Id-formula-Prop a) (pr1 x)) + ( ~ a) + ( unit-trunc-Prop (w-ax not-a-in-logic))))))}) + + canonical-worlds : subtype (lsuc l1 ⊔ lsuc l2) (formulas (l1 ⊔ l2) i) + canonical-worlds x = + Σ-Prop + ( is-L-consistent-theory-Prop x) + ( λ is-cons → is-L-complete-theory-Prop (x , is-cons)) + + canonical-kripke-model-world-type : UU (lsuc l1 ⊔ lsuc l2) + canonical-kripke-model-world-type = + Σ (formulas (l1 ⊔ l2) i) (is-in-subtype canonical-worlds) + + lindenbaum : + (x : L-consistent-theory) → + ∃ ( L-consistent-theory) + ( λ y → L-consistent-theory-leq x y × is-L-complete-theory y) + lindenbaum x = + apply-universal-property-trunc-Prop + ( zorn + ( L-consistent-big-theories-Poset) + ( unit-trunc-Prop ( x , refl-leq-subtype (pr1 x))) + ( λ C is-inhabited-C → + ( intro-∃ + ( pair + ( pair + ( resized-chain-union C) + ( is-L-consistent-union C is-inhabited-C)) + ( λ a a-in-x → + ( map-equiv + ( inv-equiv (is-equiv-resize prop-resize (chain-union C a))) + ( apply-universal-property-trunc-Prop + ( is-inhabited-C) + ( exists-Prop + ( type-subtype (pr1 C)) + ( λ s → pr1 (pr1 (pr1 s)) a)) + ( λ y → intro-∃ y (pr2 (pr1 y) a a-in-x)))))) + ( λ y a a-in-y → + ( map-equiv + ( inv-equiv (is-equiv-resize prop-resize (chain-union C a))) + ( intro-∃ y a-in-y)))))) + ( ∃-Prop + ( L-consistent-theory) + ( λ y → L-consistent-theory-leq x y × is-L-complete-theory y)) + ( λ ((m , x-leq-m) , is-max) → + ( intro-∃ m + ( x-leq-m , λ y m-leq-y → + ( ap pr1 + ( is-max + ( pair y + ( transitive-leq-Poset + ( L-consistent-theories-Poset) + ( x) + ( m) + ( y) + ( m-leq-y) + ( x-leq-m))) + ( m-leq-y)))))) + where + L-consistent-big-theories-Poset : Poset (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) + L-consistent-big-theories-Poset = + poset-Subposet + ( L-consistent-theories-Poset) + ( L-consistent-theory-leq-Prop x) + + chain-union : + chain-Poset l2 L-consistent-big-theories-Poset → + formulas (lsuc l1 ⊔ lsuc l2) i + chain-union C a = + exists-Prop + ( type-subtype + ( sub-preorder-chain-Poset + ( L-consistent-big-theories-Poset) + ( C))) + ( λ s → pr1 (pr1 (pr1 s)) a) + + resized-chain-union : + chain-Poset l2 L-consistent-big-theories-Poset → type-Poset theories-Poset + resized-chain-union C = resize prop-resize ∘ chain-union C + + in-chain-weak-deduction-list-chain-union : + (C : chain-Poset l2 L-consistent-big-theories-Poset) + (e : type-chain-Poset L-consistent-big-theories-Poset C) + {a : formula i} + {l : list (formula i)} → + in-list l ⊆ union-subtype logic (chain-union C) → + in-list l ▷ a → + ∃ ( type-chain-Poset L-consistent-big-theories-Poset C) + ( λ y → + ( is-in-subtype + ( weak-modal-logic (union-subtype logic (pr1 (pr1 (pr1 y))))) + ( a))) + in-chain-weak-deduction-list-chain-union C e {a} {nil} sub wd = + intro-∃ e (ex-falso (nil-no-deduction wd)) + in-chain-weak-deduction-list-chain-union + C e {a} {cons y l} sub (w-ax is-ax) = + apply-universal-property-trunc-Prop + ( sub a is-ax) + ( ∃-Prop ( type-chain-Poset L-consistent-big-theories-Poset C) + ( λ y → + ( is-in-subtype + ( weak-modal-logic (union-subtype logic (pr1 (pr1 (pr1 y))))) + ( a)))) + ( λ + { (inl in-logic) → + ( intro-∃ e + ( weak-modal-logic-ax + ( subtype-union-left logic (pr1 (pr1 (pr1 e))) a in-logic))) + ; (inr in-union) → + ( apply-universal-property-trunc-Prop + ( in-union) + ( ∃-Prop ( type-chain-Poset L-consistent-big-theories-Poset C) + ( λ y → + ( is-in-subtype + ( weak-modal-logic + ( union-subtype logic (pr1 (pr1 (pr1 y))))) + ( a)))) + ( λ (y , a-in-y) → + ( intro-∃ y + ( weak-modal-logic-ax + (subtype-union-right + ( logic) + ( pr1 (pr1 (pr1 y))) + ( a) + ( a-in-y))))))}) + in-chain-weak-deduction-list-chain-union + C e {a} {l@(cons _ _)} sub (w-mp {b} wdba wdb) = + apply-twice-universal-property-trunc-Prop + ( in-chain-weak-deduction-list-chain-union C e {l = l} sub wdba) + ( in-chain-weak-deduction-list-chain-union C e {l = l} sub wdb) + ( ∃-Prop + ( type-chain-Poset L-consistent-big-theories-Poset C) + ( λ y → + ( is-in-subtype + ( weak-modal-logic (union-subtype logic (pr1 (pr1 (pr1 y))))) + ( a)))) + ( λ (y , ba-in-y) (z , b-in-z) → + ( apply-universal-property-trunc-Prop + ( pr2 C y z) + ( ∃-Prop + ( type-chain-Poset L-consistent-big-theories-Poset C) + ( λ y → + ( is-in-subtype + ( weak-modal-logic (union-subtype logic (pr1 (pr1 (pr1 y))))) + ( a)))) + ( λ + { (inl y-leq-z) → + ( intro-∃ z + ( weak-modal-logic-mp + ( weak-modal-logic-monotic + {ax₁ = union-subtype logic (pr1 (pr1 (pr1 y)))} + {ax₂ = union-subtype logic (pr1 (pr1 (pr1 z)))} + ( subset-union-subset-right + ( logic) + ( pr1 (pr1 (pr1 y))) + ( pr1 (pr1 (pr1 z))) + ( y-leq-z)) + ( b ⇒ a) + ( ba-in-y)) + ( b-in-z))) + ; (inr z-leq-y) → + ( intro-∃ y + ( weak-modal-logic-mp + (ba-in-y) + ( weak-modal-logic-monotic + {ax₁ = union-subtype logic (pr1 (pr1 (pr1 z)))} + {ax₂ = union-subtype logic (pr1 (pr1 (pr1 y)))} + ( subset-union-subset-right + ( logic) + ( pr1 (pr1 (pr1 z))) + ( pr1 (pr1 (pr1 y))) + ( z-leq-y)) + ( b) + ( b-in-z))))}))) + + in-chain-in-chain-union : + (C : chain-Poset l2 L-consistent-big-theories-Poset) → + is-inhabited (type-chain-Poset L-consistent-big-theories-Poset C) → + {a : formula i} → + is-in-subtype + ( weak-modal-logic (union-subtype logic (chain-union C))) a → + ∃ ( type-subtype + ( sub-preorder-chain-Poset L-consistent-big-theories-Poset C)) + ( λ x → + ( is-in-subtype + ( weak-modal-logic (union-subtype logic (pr1 (pr1 (pr1 x))))) + ( a))) + in-chain-in-chain-union C is-inhabited-C {a} in-union = + apply-twice-universal-property-trunc-Prop + ( in-union) + ( is-inhabited-C) + ( ∃-Prop + ( type-subtype + ( sub-preorder-chain-Poset L-consistent-big-theories-Poset C)) + ( λ x → + ( is-in-subtype + ( weak-modal-logic (union-subtype logic (pr1 (pr1 (pr1 x))))) + ( a)))) + ( λ wd c → + ( let (l , sub , wdl) = list-assumptions-weak-deduction wd + in in-chain-weak-deduction-list-chain-union C c sub wdl)) + + is-L-consistent-union : + (C : chain-Poset l2 L-consistent-big-theories-Poset) → + is-inhabited (type-chain-Poset L-consistent-big-theories-Poset C) → + is-L-consistent-theory (resized-chain-union C) + is-L-consistent-union C C-is-inhabited bot-in-union = + apply-universal-property-trunc-Prop + ( in-chain-in-chain-union C C-is-inhabited + ( weak-modal-logic-monotic + {ax₁ = union-subtype logic (resized-chain-union C)} + {ax₂ = union-subtype logic (chain-union C)} + ( subset-union-subset-right + ( logic) + ( resized-chain-union C) + ( chain-union C) + ( λ b → + ( map-equiv (is-equiv-resize prop-resize (chain-union C b))))) + ( ⊥) + ( bot-in-union))) + ( empty-Prop) + ( λ (y , bot-in-y) → pr2 (pr1 (pr1 y)) bot-in-y) + + is-L-consistent-logic : is-L-consistent-theory logic + is-L-consistent-logic bot-in-logic = + L-is-cons + ( transitive-leq-subtype + ( weak-modal-logic (union-subtype logic logic)) + ( modal-logic (union-subtype logic logic)) + ( logic) + ( transitive-leq-subtype + ( modal-logic (union-subtype logic logic)) + ( modal-logic logic) + ( logic) + ( subset-double-modal-logic axioms) + ( modal-logic-monotic + ( subtype-union-both + ( logic) + ( logic) + ( logic) + ( refl-leq-subtype logic) + ( refl-leq-subtype logic)))) + ( weak-modal-logic-subset-modal-logic (union-subtype logic logic)) + ( ⊥) + ( bot-in-logic)) + + is-inhabited-canonical-kripke-model-world : + is-inhabited canonical-kripke-model-world-type + is-inhabited-canonical-kripke-model-world = + apply-universal-property-trunc-Prop + ( lindenbaum (logic , is-L-consistent-logic)) + ( is-inhabited-Prop canonical-kripke-model-world-type) + ( λ x → unit-trunc-Prop ((pr1 (pr1 x)) , ((pr2 (pr1 x)) , (pr2 (pr2 x))))) + + canonical-kripke-model : + kripke-model (canonical-kripke-model-world-type) (l1 ⊔ l2) i (l1 ⊔ l2) + pr1 (pr1 canonical-kripke-model) = + is-inhabited-canonical-kripke-model-world + pr2 (pr1 canonical-kripke-model) x y = + Π-Prop + ( formula i) + ( λ a → hom-Prop (pr1 x (□ a)) (pr1 y a)) + pr2 canonical-kripke-model n x = pr1 x (var n) + + complete-theory-box : + LEM (l1 ⊔ l2) → + (x : L-consistent-theory) + (is-comp : is-L-complete-theory x) + (a : formula i) → + ( (y : formulas (l1 ⊔ l2) i) → + (is-canonical : is-in-subtype canonical-worlds y) → + model-relation i canonical-kripke-model + ( pr1 x , pr2 x , is-comp) + ( y , is-canonical) → + is-in-subtype y a) → + is-in-subtype (pr1 x) (□ a) + complete-theory-box lem x is-comp a h = + apply-universal-property-trunc-Prop + ( complete-theory-contains-all-formulas lem x is-comp (□ a)) + ( pr1 x (□ a)) + ( λ { (inl box-in-logic) → box-in-logic + ; (inr not-box-in-logic) → + ( ex-falso + ( apply-universal-property-trunc-Prop + ( lindenbaum (w not-box-in-logic)) + ( empty-Prop) + ( λ (y , w-leq-y , y-is-comp) → + ( pr2 y + ( weak-modal-logic-monotic + { ax₁ = pr1 y} + { ax₂ = union-subtype logic (pr1 y)} + ( subtype-union-right logic (pr1 y)) + ( ⊥) + ( weak-modal-logic-mp + ( axioms-subset-weak-modal-logic + ( pr1 y) + ( ~ a) + ( w-leq-y + ( ~ a) + ( subtype-union-left + ( Id-formula-Prop (~ a)) + ( λ b → pr1 x (□ b)) + ( ~ a) + ( refl)))) + ( weak-modal-logic-ax + ( h (pr1 y) (pr2 y , y-is-comp) + ( λ b box-in-logic → + ( w-leq-y b + ( subtype-union-right + ( Id-formula-Prop (~ a)) + ( λ b → pr1 x (□ b)) + ( b) + ( box-in-logic))))))))))))}) + where + + list-to-implications : + formula i → + (l : list (formula i)) → + formula i + list-to-implications f nil = f + list-to-implications f (cons g l) = list-to-implications (g ⇒ f) l + + list-to-implications-rev : + formula i → + (l : list (formula i)) → + formula i + list-to-implications-rev f nil = f + list-to-implications-rev f (cons g l) = g ⇒ list-to-implications-rev f l + + list-to-implication-rev-snoc : + (f g : formula i) (l : list (formula i)) → + list-to-implications f (snoc l g) = g ⇒ list-to-implications f l + list-to-implication-rev-snoc f g nil = refl + list-to-implication-rev-snoc f g (cons h l) = + list-to-implication-rev-snoc (h ⇒ f) g l + + reverse-list-to-implications : + (f : formula i) (l : list (formula i)) → + list-to-implications f (reverse-list l) = list-to-implications-rev f l + reverse-list-to-implications f nil = refl + reverse-list-to-implications f (cons g l) = + ( list-to-implication-rev-snoc f g (reverse-list l)) ∙ + ( ap (λ x → g ⇒ x) (reverse-list-to-implications f l)) + + move-right : + {l' : Level} (axioms : formulas l' i) + (f : formula i) (l : list (formula i)) → + ax-k i ⊆ weak-modal-logic axioms → + ax-s i ⊆ weak-modal-logic axioms → + is-in-subtype (weak-modal-logic (union-subtype axioms (in-list l))) f → + is-in-subtype (weak-modal-logic axioms) (list-to-implications f l) + move-right axioms f nil _ _ = + weak-modal-logic-monotic + ( subtype-union-both + ( axioms) + ( in-list nil) + ( axioms) + ( refl-leq-subtype axioms) + ( λ _ → ex-falso ∘ empty-in-list-nil)) + ( f) + move-right axioms f (cons c l) contains-ax-k contains-ax-s wd = + move-right axioms (c ⇒ f) l contains-ax-k contains-ax-s + ( forward-implication + ( deduction-lemma + ( union-subtype axioms (in-list l)) + ( transitive-leq-subtype + ( ax-k i) + ( weak-modal-logic axioms) + ( weak-modal-logic (union-subtype axioms (in-list l))) + ( weak-modal-logic-monotic + ( subtype-union-left axioms (in-list l))) + ( contains-ax-k)) + ( transitive-leq-subtype + ( ax-s i) + ( weak-modal-logic axioms) + ( weak-modal-logic (union-subtype axioms (in-list l))) + ( weak-modal-logic-monotic + ( subtype-union-left axioms (in-list l))) + ( contains-ax-s)) + ( c) + ( f)) + ( weak-modal-logic-monotic + ( transitive-leq-subtype + ( union-subtype axioms (in-list (cons c l))) + ( union-subtype axioms (theory-add-formula c (in-list l))) + ( theory-add-formula c (union-subtype axioms (in-list l))) + ( union-swap-1-2 axioms (Id-formula-Prop c) (in-list l)) + ( subset-union-subset-right + ( axioms) + ( in-list (cons c l)) + ( theory-add-formula c (in-list l)) + ( backward-subset-head-add c l))) + ( f) + ( wd))) + + aux'''' : + (l : list (formula i)) → + in-list l ⊆ (λ b → pr1 x (□ b)) → + is-in-subtype + ( weak-modal-logic (union-subtype logic (pr1 x))) + ( □ (list-to-implications-rev a l)) → + is-in-subtype (weak-modal-logic (union-subtype logic (pr1 x))) (□ a) + aux'''' nil sub in-logic = in-logic + aux'''' (cons c l) sub in-logic = + aux'''' l + ( transitive-leq-subtype + ( in-list l) + ( in-list (cons c l)) + ( λ b → pr1 x (□ b)) + ( sub) + ( subset-in-tail c l)) + ( weak-modal-logic-mp + { a = □ c} + { b = □ (list-to-implications-rev a l)} + ( weak-modal-logic-mp + { a = □ (c ⇒ list-to-implications-rev a l)} + { b = □ c ⇒ □ (list-to-implications-rev a l)} + ( weak-modal-logic-monotic + { ax₁ = logic} + { ax₂ = union-subtype logic (pr1 x)} + ( subtype-union-left logic (pr1 x)) + ( □ (c ⇒ list-to-implications-rev a l) ⇒ + □ c ⇒ □ list-to-implications-rev a l) + ( weak-modal-logic-ax + ( contains-ax-n _ (c , list-to-implications-rev a l , refl)))) + ( in-logic)) + ( weak-modal-logic-ax + ( subtype-union-right logic + ( pr1 x) + ( □ c) + ( sub c (unit-trunc-Prop (is-head c l)))))) + + aux''' : + (l : list (formula i)) → + in-list l ⊆ (λ b → pr1 x (□ b)) → + is-in-subtype (weak-modal-logic (union-subtype logic (in-list l))) (a) → + is-in-subtype (weak-modal-logic (union-subtype logic (pr1 x))) (□ a) + aux''' l sub in-logic = + aux'''' l sub + ( weak-modal-logic-ax + ( subtype-union-left + ( logic) + ( pr1 x) + ( □ (list-to-implications-rev a l)) + ( modal-logic-nec + ( tr + ( is-in-subtype logic) + ( reverse-list-to-implications a l) + ( subset-double-modal-logic + ( axioms) + ( list-to-implications a (reverse-list l)) + ( weak-modal-logic-subset-modal-logic + ( logic) + ( list-to-implications a (reverse-list l)) + ( move-right logic a (reverse-list l) + ( transitive-leq-subtype + ( ax-k i) + ( logic) + ( weak-modal-logic logic) + ( axioms-subset-weak-modal-logic logic) + ( contains-ax-k)) + ( transitive-leq-subtype + ( ax-s i) + ( logic) + ( weak-modal-logic logic) + ( axioms-subset-weak-modal-logic logic) + ( contains-ax-s)) + ( weak-modal-logic-monotic + ( subset-union-subset-right logic + ( in-list l) + ( in-list (reverse-list l)) + ( subset-reversing-list l)) + ( a) + ( in-logic))))))))) + + aux'' : + (l : list (formula i)) → + in-list l ⊆ (λ b → pr1 x (□ b)) → + is-in-subtype + ( weak-modal-logic + ( theory-add-formula (~ a) (union-subtype logic (in-list l)))) + ( ⊥) → + is-in-subtype (weak-modal-logic (union-subtype logic (pr1 x))) (□ a) + aux'' l sub in-logic = + aux''' l sub + ( weak-modal-logic-mp + { a = ~~ a} + { b = a} + ( weak-modal-logic-ax + ( subtype-union-left logic + ( in-list l) + ( _) + ( contains-ax-dn _ (a , refl)))) + ( forward-implication + ( deduction-lemma + ( union-subtype logic (in-list l)) + ( transitive-leq-subtype + ( ax-k i) + ( logic) + ( weak-modal-logic (union-subtype logic (in-list l))) + ( transitive-leq-subtype + ( logic) + ( union-subtype logic (in-list l)) + ( weak-modal-logic (union-subtype logic (in-list l))) + ( axioms-subset-weak-modal-logic + ( union-subtype logic (in-list l))) + ( subtype-union-left logic (in-list l))) + ( contains-ax-k)) + ( transitive-leq-subtype + ( ax-s i) + ( logic) + ( weak-modal-logic (union-subtype logic (in-list l))) + ( transitive-leq-subtype + ( logic) + ( union-subtype logic (in-list l)) + ( weak-modal-logic (union-subtype logic (in-list l))) + ( axioms-subset-weak-modal-logic + ( union-subtype logic (in-list l))) + ( subtype-union-left logic (in-list l))) + ( contains-ax-s)) + ( ~ a) + ( ⊥)) + ( in-logic))) + + aux' : + (l : list (formula i)) → + in-list l ⊆ (λ b → pr1 x (□ b)) → + is-in-subtype + ( weak-modal-logic + ( union-subtype logic (theory-add-formula (~ a) (in-list l)))) + ( ⊥) → + is-in-subtype (weak-modal-logic (union-subtype logic (pr1 x))) (□ a) + aux' l sub in-logic = + aux'' l sub + ( weak-modal-logic-monotic + ( union-swap-1-2 + ( logic) + ( Id-formula-Prop (~ a)) + ( in-list l)) + ( ⊥) + ( in-logic)) + + aux : + (l : list (formula i)) → + in-list l ⊆ + union-subtype logic (theory-add-formula (~ a) (λ b → pr1 x (□ b))) → + is-in-subtype (weak-modal-logic (in-list l)) ⊥ → + is-in-subtype (weak-modal-logic (union-subtype logic (pr1 x))) (□ a) + aux l sub in-logic = + apply-universal-property-trunc-Prop + ( lists-in-union-lists l logic + ( theory-add-formula (~ a) (λ b → pr1 x (□ b))) + ( sub)) + ( weak-modal-logic (union-subtype logic (pr1 x)) (□ a)) + ( λ (l-ax , l-w , l-sub-union , l-ax-sub-axioms , l-w-sub-w) → + ( apply-universal-property-trunc-Prop + ( lists-in-union-lists l-w + ( Id-formula-Prop (~ a)) + ( λ b → pr1 x (□ b)) + ( l-w-sub-w)) + ( weak-modal-logic (union-subtype logic (pr1 x)) (□ a)) + ( λ (l-not-a , l-boxes , l-sub-union' , l-not-a-sub , l-boxes-sub) → + ( aux' + ( l-boxes) + ( l-boxes-sub) + ( weak-modal-logic-monotic + { ax₁ = in-list l} + { ax₂ = + union-subtype logic + ( theory-add-formula (~ a) (in-list l-boxes))} + ( transitive-leq-subtype + ( in-list l) + ( union-subtype (in-list l-ax) (in-list l-w)) + ( union-subtype + ( logic) + ( theory-add-formula (~ a) (in-list l-boxes))) + ( subset-union-subsets + ( in-list l-ax) + ( in-list l-w) + ( logic) + ( theory-add-formula (~ a) (in-list l-boxes)) + ( l-ax-sub-axioms) + ( transitive-leq-subtype + ( in-list l-w) + ( union-subtype (in-list l-not-a) (in-list l-boxes)) + ( theory-add-formula (~ a) (in-list l-boxes)) + ( subset-union-subset-left + ( in-list l-not-a) + ( Id-formula-Prop (~ a)) + ( in-list l-boxes) + ( l-not-a-sub)) + ( l-sub-union'))) + ( l-sub-union)) + ( ⊥) + ( in-logic)))))) + + w : is-in-subtype (pr1 x) (~ □ a) → L-consistent-theory + pr1 (w b) = theory-add-formula (~ a) (λ b → pr1 x (□ b)) + pr2 (w not-box-in-logic) bot-in-logic = + apply-universal-property-trunc-Prop + ( bot-in-logic) + ( empty-Prop) + ( λ wd-bot → + ( let (l , sub , wd) = list-assumptions-weak-deduction wd-bot + in pr2 x + ( weak-modal-logic-mp + ( weak-modal-logic-ax + ( subtype-union-right logic (pr1 x) _ not-box-in-logic)) + ( aux l sub (unit-trunc-Prop wd))))) + + module _ + (lem : LEM (l1 ⊔ l2)) + where + + canonical-model-theorem : + (a : formula i) + (x : canonical-kripke-model-world-type) → + pr1 x a ⇔ ((canonical-kripke-model , x) ⊨ a) + pr1 (canonical-model-theorem (var n) x) in-logic = + map-raise in-logic + pr1 (canonical-model-theorem ⊥ x) in-logic = + map-raise + ( pr1 + ( pr2 x) + ( unit-trunc-Prop + ( w-ax (subtype-union-right logic (pr1 x) ⊥ in-logic)))) + pr1 (canonical-model-theorem (a ⇒ b) x) in-logic fa = + pr1 + ( canonical-model-theorem b x) + ( weak-modal-logic-subset-complete-theory + ( pr1 x , pr1 (pr2 x)) + ( pr2 (pr2 x)) + ( b) + ( unit-trunc-Prop + ( w-mp + ( w-ax in-logic) + ( w-ax (backward-implication (canonical-model-theorem a x) fa))))) + pr1 (canonical-model-theorem (□ a) x) in-logic y xRy = + forward-implication (canonical-model-theorem a y) (xRy a in-logic) + pr2 (canonical-model-theorem (var n) x) f = + map-inv-raise f + pr2 (canonical-model-theorem ⊥ x) (map-raise ()) + pr2 (canonical-model-theorem (a ⇒ b) x) f = + complete-theory-implication + ( lem) + ( pr1 x , pr1 (pr2 x)) + ( pr2 (pr2 x)) + ( λ a-in-logic → + ( backward-implication + ( canonical-model-theorem b x) + ( f + ( forward-implication (canonical-model-theorem a x) a-in-logic)))) + pr2 (canonical-model-theorem (□ a) x) f = + complete-theory-box + ( lem) + ( pr1 x , pr1 (pr2 x)) + ( pr2 (pr2 x)) + ( a) + ( λ y is-canonical access → + ( backward-implication + ( canonical-model-theorem a (y , is-canonical)) + ( f (y , is-canonical) access))) +``` diff --git a/src/modal-logic/classical-completeness-K.lagda.md b/src/modal-logic/classical-completeness-K.lagda.md deleted file mode 100644 index 8cc18c71d2..0000000000 --- a/src/modal-logic/classical-completeness-K.lagda.md +++ /dev/null @@ -1,555 +0,0 @@ -# Classical proof of completeness K - -```agda -module modal-logic.classical-completeness-K where -``` - -
Imports - -```agda -open import foundation.cartesian-product-types -open import foundation.coproduct-types -open import foundation.dependent-pair-types -open import foundation.disjunction -open import foundation.empty-types -open import foundation.existential-quantification -open import foundation.function-types -open import foundation.identity-types -open import foundation.inhabited-types -open import foundation.intersections-subtypes -open import foundation.law-of-excluded-middle -open import foundation.logical-equivalences -open import foundation.negation -open import foundation.propositional-truncations -open import foundation.propositions -open import foundation.raising-universe-levels -open import foundation.sets -open import foundation.subtypes -open import foundation.transport-along-identifications -open import foundation.unions-subtypes -open import foundation.unit-type -open import foundation.universe-levels - -open import modal-logic.axioms -open import modal-logic.completeness -open import modal-logic.formulas -open import modal-logic.kripke-semantics -open import modal-logic.logic-syntax -open import modal-logic.soundness -open import modal-logic.weak-deduction - -open import order-theory.chains-posets -open import order-theory.maximal-elements-posets -open import order-theory.posets -open import order-theory.preorders -open import order-theory.subposets -open import order-theory.top-elements-posets -``` - -
- -## Idea - -TODO - -## Definition - -```agda -module _ - {l1 : Level} (l2 : Level) (A : UU l1) - where - - subtypes-leq-Preorder : Preorder (l1 ⊔ lsuc l2) (l1 ⊔ l2) - pr1 subtypes-leq-Preorder = subtype l2 A - pr1 (pr2 subtypes-leq-Preorder) = leq-prop-subtype - pr1 (pr2 (pr2 subtypes-leq-Preorder)) = refl-leq-subtype - pr2 (pr2 (pr2 subtypes-leq-Preorder)) = transitive-leq-subtype - - subtypes-leq-Poset : Poset (l1 ⊔ lsuc l2) (l1 ⊔ l2) - pr1 subtypes-leq-Poset = subtypes-leq-Preorder - pr2 subtypes-leq-Poset = antisymmetric-leq-subtype - -module _ - {l1 l2 : Level} {i : Set l1} (axioms : formulas l2 i) - (L-is-cons : is-consistent-modal-logic (weak-modal-logic axioms)) - (contains-ax-k : ax-k i ⊆ weak-modal-logic axioms) - (contains-ax-s : ax-s i ⊆ weak-modal-logic axioms) - (contains-ax-dn : ax-dn i ⊆ weak-modal-logic axioms) - where - - is-L-consistent-theory-Prop : formulas (l1 ⊔ l2) i → Prop (l1 ⊔ l2) - is-L-consistent-theory-Prop t = - is-consistent-modal-logic-Prop (weak-modal-logic (union-subtype axioms t)) - - is-L-consistent-theory : formulas (l1 ⊔ l2) i → UU (l1 ⊔ l2) - is-L-consistent-theory = type-Prop ∘ is-L-consistent-theory-Prop - - L-consistent-theory : UU (lsuc l1 ⊔ lsuc l2) - L-consistent-theory = Σ (formulas (l1 ⊔ l2) i) is-L-consistent-theory - - L-consistent-theory-leq-Prop : - L-consistent-theory → L-consistent-theory → Prop (l1 ⊔ l2) - L-consistent-theory-leq-Prop x y = leq-prop-subtype (pr1 x) (pr1 y) - - L-consistent-theory-leq : - L-consistent-theory → L-consistent-theory → UU (l1 ⊔ l2) - L-consistent-theory-leq x y = type-Prop (L-consistent-theory-leq-Prop x y) - - theories-Poset : Poset (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) - theories-Poset = subtypes-leq-Poset (l1 ⊔ l2) (formula i) - - L-consistent-theories-Poset : Poset (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) - L-consistent-theories-Poset = - poset-Subposet theories-Poset is-L-consistent-theory-Prop - - is-L-complete-theory-Prop : L-consistent-theory → Prop (lsuc l1 ⊔ lsuc l2) - is-L-complete-theory-Prop = - is-maximal-element-Poset-Prop L-consistent-theories-Poset - - is-L-complete-theory : L-consistent-theory → UU (lsuc l1 ⊔ lsuc l2) - is-L-complete-theory = type-Prop ∘ is-L-complete-theory-Prop - - complete-theory-closed : - (x : L-consistent-theory) → is-L-complete-theory x → - {a : formula i} → pr1 x ▷ a → is-in-subtype (pr1 x) a - complete-theory-closed x is-comp {a} wd = - tr - ( λ y → is-in-subtype (pr1 y) a) - ( is-comp - ( theory-add-formula a (pr1 x) - , λ bot-in-logic → - ( pr2 x - ( subset-weak-modal-logic-if-subset-axioms - ( subtype-union-both - ( axioms) - ( theory-add-formula a (pr1 x)) - ( weak-modal-logic (union-subtype axioms (pr1 x))) - ( transitive-leq-subtype - ( axioms) - ( union-subtype axioms (pr1 x)) - ( weak-modal-logic (union-subtype axioms (pr1 x))) - ( axioms-subset-weak-modal-logic - ( union-subtype axioms (pr1 x))) - ( subtype-union-left axioms (pr1 x))) - ( subtype-union-both - ( Id-formula-Prop a) - ( pr1 x) - ( weak-modal-logic (union-subtype axioms (pr1 x))) - ( λ { .a refl → - ( weak-modal-logic-monotic - ( subtype-union-right axioms (pr1 x)) - ( a) - ( unit-trunc-Prop wd))}) - ( transitive-leq-subtype - ( pr1 x) - ( union-subtype axioms (pr1 x)) - ( weak-modal-logic (union-subtype axioms (pr1 x))) - ( axioms-subset-weak-modal-logic - (union-subtype axioms (pr1 x))) - ( subtype-union-right axioms (pr1 x))))) - ( ⊥) - ( bot-in-logic)))) - ( subtype-union-right (Id-formula-Prop a) (pr1 x))) - ( unit-trunc-Prop (inl refl)) - - weak-modal-logic-subset-complete-theory : - (x : L-consistent-theory) → is-L-complete-theory x → - weak-modal-logic (pr1 x) ⊆ pr1 x - weak-modal-logic-subset-complete-theory x is-comp a = - map-universal-property-trunc-Prop - ( pr1 x a) - ( complete-theory-closed x is-comp) - - deduction-axioms-L-complete-theory : - (x : L-consistent-theory) → is-L-complete-theory x → - {a : formula i} → axioms ▷ a → is-in-subtype (pr1 x) a - deduction-axioms-L-complete-theory x is-comp {a} wd = - tr - ( λ y → is-in-subtype (pr1 y) a) - ( is-comp - ( theory-add-formula a (pr1 x) - , λ bot-in-logic → - ( pr2 x - ( subset-weak-modal-logic-if-subset-axioms - ( subtype-union-both - ( axioms) - ( theory-add-formula a (pr1 x)) - ( weak-modal-logic (union-subtype axioms (pr1 x))) - ( transitive-leq-subtype - ( axioms) - ( union-subtype axioms (pr1 x)) - ( weak-modal-logic (union-subtype axioms (pr1 x))) - ( axioms-subset-weak-modal-logic - ( union-subtype axioms (pr1 x))) - ( subtype-union-left axioms (pr1 x))) - ( subtype-union-both - ( Id-formula-Prop a) - ( pr1 x) - ( weak-modal-logic (union-subtype axioms (pr1 x))) - ( λ { .a refl → - ( weak-modal-logic-monotic - ( subtype-union-left axioms (pr1 x)) - ( a) - ( unit-trunc-Prop wd))}) - ( transitive-leq-subtype - ( pr1 x) - ( union-subtype axioms (pr1 x)) - ( weak-modal-logic (union-subtype axioms (pr1 x))) - ( axioms-subset-weak-modal-logic - (union-subtype axioms (pr1 x))) - ( subtype-union-right axioms (pr1 x))))) - ( ⊥) - ( bot-in-logic)))) - ( subtype-union-right (Id-formula-Prop a) (pr1 x))) - ( subtype-union-left (Id-formula-Prop a) (pr1 x) a refl) - - logic-subset-L-complete-theory : - (x : L-consistent-theory) → is-L-complete-theory x → - weak-modal-logic (axioms) ⊆ pr1 x - logic-subset-L-complete-theory x is-comp a = - map-universal-property-trunc-Prop - ( pr1 x a) - ( deduction-axioms-L-complete-theory x is-comp) - - axioms-subset-L-complete-theory : - (x : L-consistent-theory) → is-L-complete-theory x → axioms ⊆ pr1 x - axioms-subset-L-complete-theory x is-comp = - transitive-leq-subtype - ( axioms) - ( weak-modal-logic axioms) - ( pr1 x) - ( logic-subset-L-complete-theory x is-comp) - ( axioms-subset-weak-modal-logic axioms) - - complete-theory-contains-all-formulas : - LEM (l1 ⊔ l2) → - (x : L-consistent-theory) → is-L-complete-theory x → - (a : formula i) → type-disj-Prop ((pr1 x) a) ((pr1 x) (~ a)) - complete-theory-contains-all-formulas lem x is-comp a with lem ((pr1 x) a) - ... | inl a-in-logic = inl-disj-Prop ((pr1 x) a) ((pr1 x) (~ a)) a-in-logic - ... | inr a-not-in-logic = - unit-trunc-Prop - ( inr - ( tr - ( λ y → is-in-subtype (pr1 y) (~ a)) - ( is-comp - ( theory-add-formula (~ a) (pr1 x) - , λ bot-in-logic → - ( a-not-in-logic - ( weak-modal-logic-subset-complete-theory - ( x) - ( is-comp) - ( a) - ( apply-universal-property-trunc-Prop - ( forward-implication - ( deduction-lemma - ( pr1 x) - ( transitive-leq-subtype - ( ax-k i) - ( weak-modal-logic (axioms)) - ( weak-modal-logic (pr1 x)) - ( weak-modal-logic-monotic - ( axioms-subset-L-complete-theory x is-comp)) - ( contains-ax-k)) - ( transitive-leq-subtype - ( ax-s i) - ( weak-modal-logic (axioms)) - ( weak-modal-logic (pr1 x)) - ( weak-modal-logic-monotic - ( axioms-subset-L-complete-theory x is-comp)) - ( contains-ax-s)) - ( ~ a) - ( ⊥)) - ( weak-modal-logic-monotic - { ax₁ = - ( union-subtype axioms - ( theory-add-formula (~ a) (pr1 x)))} - { ax₂ = theory-add-formula (~ a) (pr1 x)} - ( subtype-union-both - ( axioms) - ( theory-add-formula (~ a) (pr1 x)) - ( theory-add-formula (~ a) (pr1 x)) - ( transitive-leq-subtype - ( axioms) - ( pr1 x) - ( theory-add-formula (~ a) (pr1 x)) - ( subtype-union-right - ( Id-formula-Prop (~ a)) - ( pr1 x)) - ( axioms-subset-L-complete-theory x is-comp)) - ( refl-leq-subtype - ( theory-add-formula (~ a) (pr1 x)))) - ( ⊥) - ( bot-in-logic))) - ( weak-modal-logic (pr1 x) a) - ( λ wd-bot → - ( unit-trunc-Prop - ( w-mp - ( w-ax - ( logic-subset-L-complete-theory - ( x) - ( is-comp) - ( ~~ a ⇒ a) - ( contains-ax-dn _ (a , refl)))) - ( wd-bot)))))))) - ( subtype-union-right (Id-formula-Prop (~ a)) (pr1 x))) - ( subtype-union-left (Id-formula-Prop (~ a)) (pr1 x) (~ a) refl))) - - complete-theory-implication : - LEM (l1 ⊔ l2) → - (x : L-consistent-theory) → is-L-complete-theory x → - {a b : formula i} → (is-in-subtype (pr1 x) a → is-in-subtype (pr1 x) b) → - is-in-subtype (pr1 x) (a ⇒ b) - complete-theory-implication lem x is-comp {a} {b} imp = - apply-universal-property-trunc-Prop - ( complete-theory-contains-all-formulas lem x is-comp a) - ( (pr1 x) (a ⇒ b)) - ( λ { (inl a-in-logic) → - ( weak-modal-logic-subset-complete-theory - ( x) - ( is-comp) - ( a ⇒ b) - ( forward-implication - ( deduction-lemma - ( pr1 x) - ( transitive-leq-subtype - ( ax-k i) - ( weak-modal-logic axioms) - ( weak-modal-logic (pr1 x)) - ( transitive-leq-subtype - ( weak-modal-logic axioms) - ( pr1 x) - ( weak-modal-logic (pr1 x)) - ( axioms-subset-weak-modal-logic (pr1 x)) - ( logic-subset-L-complete-theory x is-comp)) - ( contains-ax-k)) - ( transitive-leq-subtype - ( ax-s i) - ( weak-modal-logic axioms) - ( weak-modal-logic (pr1 x)) - ( transitive-leq-subtype - ( weak-modal-logic axioms) - ( pr1 x) - ( weak-modal-logic (pr1 x)) - ( axioms-subset-weak-modal-logic (pr1 x)) - ( logic-subset-L-complete-theory x is-comp)) - ( contains-ax-s)) - ( a) - ( b)) - ( weak-modal-logic-monotic - { ax₁ = pr1 x} -- TODO : make explicit - { ax₂ = union-subtype (Id-formula-Prop a) (pr1 x)} - ( subtype-union-right (Id-formula-Prop a) (pr1 x)) - ( b) - ( unit-trunc-Prop (w-ax (imp a-in-logic)))))) - ; (inr not-a-in-logic) → - ( weak-modal-logic-subset-complete-theory - ( x) - ( is-comp) - ( a ⇒ b) - ( forward-implication - ( deduction-lemma - ( pr1 x) - ( transitive-leq-subtype - ( ax-k i) - ( weak-modal-logic axioms) - ( weak-modal-logic (pr1 x)) - ( transitive-leq-subtype - ( weak-modal-logic axioms) - ( pr1 x) - ( weak-modal-logic (pr1 x)) - ( axioms-subset-weak-modal-logic (pr1 x)) - ( logic-subset-L-complete-theory x is-comp)) - ( contains-ax-k)) - ( transitive-leq-subtype - ( ax-s i) - ( weak-modal-logic axioms) - ( weak-modal-logic (pr1 x)) - ( transitive-leq-subtype - ( weak-modal-logic axioms) - ( pr1 x) - ( weak-modal-logic (pr1 x)) - ( axioms-subset-weak-modal-logic (pr1 x)) - ( logic-subset-L-complete-theory x is-comp)) - ( contains-ax-s)) - ( a) - ( b)) - ( logic-ex-falso - ( theory-add-formula a (pr1 x)) - ( transitive-leq-subtype - ( ax-k i) - ( weak-modal-logic (pr1 x)) - ( weak-modal-logic (theory-add-formula a (pr1 x))) - ( weak-modal-logic-monotic - ( subtype-union-right (Id-formula-Prop a) (pr1 x))) - ( transitive-leq-subtype - ( ax-k i) - ( weak-modal-logic axioms) - ( weak-modal-logic (pr1 x)) - ( transitive-leq-subtype - ( weak-modal-logic axioms) - ( pr1 x) - ( weak-modal-logic (pr1 x)) - ( axioms-subset-weak-modal-logic (pr1 x)) - ( logic-subset-L-complete-theory x is-comp)) - ( contains-ax-k))) - ( transitive-leq-subtype - ( ax-s i) - ( weak-modal-logic (pr1 x)) - ( weak-modal-logic (theory-add-formula a (pr1 x))) - ( weak-modal-logic-monotic - ( subtype-union-right (Id-formula-Prop a) (pr1 x))) - ( transitive-leq-subtype - ( ax-s i) - ( weak-modal-logic axioms) - ( weak-modal-logic (pr1 x)) - ( transitive-leq-subtype - ( weak-modal-logic axioms) - ( pr1 x) - ( weak-modal-logic (pr1 x)) - ( axioms-subset-weak-modal-logic (pr1 x)) - ( logic-subset-L-complete-theory x is-comp)) - ( contains-ax-s))) - ( transitive-leq-subtype - ( ax-dn i) - ( weak-modal-logic (pr1 x)) - ( weak-modal-logic (theory-add-formula a (pr1 x))) - ( weak-modal-logic-monotic - ( subtype-union-right (Id-formula-Prop a) (pr1 x))) - ( transitive-leq-subtype - ( ax-dn i) - ( weak-modal-logic axioms) - ( weak-modal-logic (pr1 x)) - ( transitive-leq-subtype - ( weak-modal-logic axioms) - ( pr1 x) - ( weak-modal-logic (pr1 x)) - ( axioms-subset-weak-modal-logic (pr1 x)) - ( logic-subset-L-complete-theory x is-comp)) - ( contains-ax-dn))) - ( a) - ( b) - ( weak-modal-logic-monotic - { ax₁ = Id-formula-Prop a} - { ax₂ = union-subtype (Id-formula-Prop a) (pr1 x)} - ( subtype-union-left (Id-formula-Prop a) (pr1 x)) - ( a) - ( unit-trunc-Prop (w-ax refl))) - ( weak-modal-logic-monotic - { ax₁ = pr1 x} - { ax₂ = union-subtype (Id-formula-Prop a) (pr1 x)} - ( subtype-union-right (Id-formula-Prop a) (pr1 x)) - ( ~ a) - ( unit-trunc-Prop (w-ax not-a-in-logic))))))}) - - canonical-worlds : subtype (lsuc l1 ⊔ lsuc l2) (formulas (l1 ⊔ l2) i) - canonical-worlds x = - Σ-Prop - ( is-L-consistent-theory-Prop x) - ( λ is-cons → is-L-complete-theory-Prop (x , is-cons)) - - canonical-kripke-model-world-type : UU (lsuc l1 ⊔ lsuc l2) - canonical-kripke-model-world-type = - Σ (formulas (l1 ⊔ l2) i) (is-in-subtype canonical-worlds) - - postulate - lindenbaum : - (x : L-consistent-theory) → - ∃ ( L-consistent-theory) - ( λ y → L-consistent-theory-leq x y × is-L-complete-theory y) - - is-L-consistent-axioms : is-L-consistent-theory (raise-subtype l1 axioms) - is-L-consistent-axioms = - map-universal-property-trunc-Prop - ( empty-Prop) - ( tr - ( λ x → x ▷ ⊥ → empty) - ( helper) - ( λ x → - ( L-is-cons - ( unit-trunc-Prop - ( deduction-equiv-axioms - ( inv-equiv-subtypes - ( axioms) - ( raise-subtype l1 axioms) - ( compute-raise-subtype l1 axioms)) - ( x)))))) - where - -- TODO: move to unions-subtypes as lemma - helper : - raise-subtype l1 axioms = union-subtype axioms (raise-subtype l1 axioms) - helper = - antisymmetric-leq-subtype - ( raise-subtype l1 axioms) - ( union-subtype axioms (raise-subtype l1 axioms)) - ( subtype-union-right axioms (raise-subtype l1 axioms)) - ( subtype-union-both - ( axioms) - ( raise-subtype l1 axioms) - ( raise-subtype l1 axioms) - ( λ _ → map-raise) - ( λ _ → id)) - - is-inhabited-canonical-kripke-model-world : - is-inhabited canonical-kripke-model-world-type - is-inhabited-canonical-kripke-model-world = - apply-universal-property-trunc-Prop - ( lindenbaum (raise-subtype l1 axioms , is-L-consistent-axioms)) - ( is-inhabited-Prop canonical-kripke-model-world-type) - ( λ x → unit-trunc-Prop ((pr1 (pr1 x)) , ((pr2 (pr1 x)) , (pr2 (pr2 x))))) - - canonical-kripke-model : - kripke-model (canonical-kripke-model-world-type) (l1 ⊔ l2) i (l1 ⊔ l2) - pr1 (pr1 canonical-kripke-model) = - is-inhabited-canonical-kripke-model-world - pr2 (pr1 canonical-kripke-model) x y = - Π-Prop - ( formula i) - ( λ a → function-Prop (is-in-subtype (pr1 x) (□ a)) ((pr1 y) a)) - pr2 canonical-kripke-model n x = pr1 x (var n) - - module _ - (lem : LEM (l1 ⊔ l2)) - where - - canonical-model-theorem : - (a : formula i) - (x : canonical-kripke-model-world-type) → - pr1 x a ⇔ ((canonical-kripke-model , x) ⊨ a) - pr1 (canonical-model-theorem (var n) x) in-logic = - map-raise in-logic - pr1 (canonical-model-theorem ⊥ x) in-logic = - map-raise - ( pr1 - ( pr2 x) - ( unit-trunc-Prop - ( w-ax (subtype-union-right axioms (pr1 x) ⊥ in-logic)))) - pr1 (canonical-model-theorem (a ⇒ b) x) in-logic fa = - pr1 - ( canonical-model-theorem b x) - ( weak-modal-logic-subset-complete-theory - ( pr1 x , pr1 (pr2 x)) - ( pr2 (pr2 x)) - ( b) - ( unit-trunc-Prop - ( w-mp - ( w-ax in-logic) - ( w-ax (backward-implication (canonical-model-theorem a x) fa))))) - pr1 (canonical-model-theorem (□ a) x) in-logic y xRy = - forward-implication (canonical-model-theorem a y) (xRy a in-logic) - pr2 (canonical-model-theorem (var n) x) f = - map-inv-raise f - pr2 (canonical-model-theorem ⊥ x) (map-raise ()) - pr2 (canonical-model-theorem (a ⇒ b) x) f = - complete-theory-implication - ( lem) - ( pr1 x , pr1 (pr2 x)) - ( pr2 (pr2 x)) - ( λ a-in-logic → - ( backward-implication - ( canonical-model-theorem b x) - ( f - ( forward-implication (canonical-model-theorem a x) a-in-logic)))) - pr2 (canonical-model-theorem (□ a) x) f = - {! !} -``` diff --git a/src/modal-logic/logic-syntax.lagda.md b/src/modal-logic/logic-syntax.lagda.md index c96a37ac5e..e8f90fea29 100644 --- a/src/modal-logic/logic-syntax.lagda.md +++ b/src/modal-logic/logic-syntax.lagda.md @@ -40,13 +40,6 @@ module _ formulas : UU (l1 ⊔ lsuc l2) formulas = subtype l2 (formula i) -module _ - {l1 l2 : Level} {i : Set l1} - where - - theory-add-formula : formula i → formulas l2 i → formulas (l1 ⊔ l2) i - theory-add-formula a = union-subtype (Id-formula-Prop a) - module _ {l1 l2 : Level} {i : Set l1} where @@ -73,6 +66,35 @@ module _ is-consistent-modal-logic : formulas l2 i → UU l2 is-consistent-modal-logic = type-Prop ∘ is-consistent-modal-logic-Prop +module _ + {l1 l2 : Level} {i : Set l1} {axioms : formulas l2 i} + where + + modal-logic-ax : + {a : formula i} → + is-in-subtype axioms a → + is-in-subtype (modal-logic axioms) a + modal-logic-ax = unit-trunc-Prop ∘ ax + + modal-logic-mp : + {a b : formula i} → + is-in-subtype (modal-logic axioms) (a ⇒ b) → + is-in-subtype (modal-logic axioms) a → + is-in-subtype (modal-logic axioms) b + modal-logic-mp {a} {b} tdab tda = + apply-twice-universal-property-trunc-Prop tdab tda + ( modal-logic axioms b) + ( λ dab da → unit-trunc-Prop (mp dab da)) + + modal-logic-nec : + {a : formula i} → + is-in-subtype (modal-logic axioms) a → + is-in-subtype (modal-logic axioms) (□ a) + modal-logic-nec {a} = + map-universal-property-trunc-Prop + ( modal-logic axioms (□ a)) + ( λ da → unit-trunc-Prop (nec da)) + module _ {l1 : Level} {i : Set l1} where @@ -146,4 +168,17 @@ module _ ( modal-logic ax₂) ( subset-double-modal-logic ax₂) ( modal-logic-monotic leq) + +module _ + {l1 l2 : Level} {i : Set l1} (a : formula i) (logic : formulas l2 i) + where + + theory-add-formula : formulas (l1 ⊔ l2) i + theory-add-formula = union-subtype (Id-formula-Prop a) logic + + formula-in-add-formula : is-in-subtype theory-add-formula a + formula-in-add-formula = subtype-union-left (Id-formula-Prop a) logic a refl + + subset-add-formula : logic ⊆ theory-add-formula + subset-add-formula = subtype-union-right (Id-formula-Prop a) logic ``` diff --git a/src/modal-logic/weak-deduction.lagda.md b/src/modal-logic/weak-deduction.lagda.md index e786905e04..4bb40a8de0 100644 --- a/src/modal-logic/weak-deduction.lagda.md +++ b/src/modal-logic/weak-deduction.lagda.md @@ -7,11 +7,13 @@ module modal-logic.weak-deduction where
Imports ```agda +open import foundation.cartesian-product-types open import foundation.coproduct-types open import foundation.dependent-pair-types open import foundation.disjunction open import foundation.empty-types open import foundation.equivalences +open import foundation.existential-quantification open import foundation.function-types open import foundation.identity-types open import foundation.logical-equivalences @@ -23,6 +25,12 @@ open import foundation.unions-subtypes open import foundation.unit-type open import foundation.universe-levels +open import foundation-core.negation + +open import lists.concatenation-lists +open import lists.lists +open import lists.reversing-lists + open import modal-logic.axioms open import modal-logic.formulas open import modal-logic.logic-syntax @@ -50,6 +58,26 @@ module _ weak-modal-logic : formulas l2 i → formulas (l1 ⊔ l2) i weak-modal-logic axioms a = trunc-Prop (axioms ▷ a) +module _ + {l1 l2 : Level} {i : Set l1} {axioms : formulas l2 i} + where + + weak-modal-logic-ax : + {a : formula i} → + is-in-subtype axioms a → + is-in-subtype (weak-modal-logic axioms) a + weak-modal-logic-ax = unit-trunc-Prop ∘ w-ax + + weak-modal-logic-mp : + {a b : formula i} → + is-in-subtype (weak-modal-logic axioms) (a ⇒ b) → + is-in-subtype (weak-modal-logic axioms) a → + is-in-subtype (weak-modal-logic axioms) b + weak-modal-logic-mp {a} {b} twdab twda = + apply-twice-universal-property-trunc-Prop twdab twda + ( weak-modal-logic axioms b) + ( λ wdab wda → unit-trunc-Prop (w-mp wdab wda)) + module _ {l1 l2 : Level} {i : Set l1} (axioms : formulas l2 i) where @@ -147,6 +175,320 @@ module _ ( subset-double-weak-modal-logic ax₂) ( weak-modal-logic-monotic leq) +-- TODO: move to lists + +in-list : {l : Level} {A : UU l} → list A → A → Prop l +in-list l a = trunc-Prop (a ∈-list l) + +subset-in-tail : + {l : Level} {A : UU l} (a : A) (l : list A) → + in-list l ⊆ in-list (cons a l) +subset-in-tail a l x = + map-universal-property-trunc-Prop + ( in-list (cons a l) x) + ( unit-trunc-Prop ∘ (is-in-tail x a l)) + +in-concat-list : + {l : Level} {A : UU l} (l1 l2 : list A) → + (a : A) → a ∈-list (concat-list l1 l2) → (a ∈-list l1) + (a ∈-list l2) +in-concat-list nil l2 a in-list = inr in-list +in-concat-list (cons x l1) l2 a (is-head a _) = inl (is-head a l1) +in-concat-list (cons x l1) l2 a (is-in-tail _ _ _ in-tail) + with in-concat-list l1 l2 a in-tail +... | inl in-l1 = inl (is-in-tail a x l1 in-l1) +... | inr in-l2 = inr in-l2 + +subset-reversing-list : + {l : Level} {A : UU l} (l : list A) → in-list l ⊆ in-list (reverse-list l) +subset-reversing-list l x = + map-universal-property-trunc-Prop + ( in-list (reverse-list l) x) + ( unit-trunc-Prop ∘ forward-contains-reverse-list x l) + +in-sum-concat-list-Prop : + {l : Level} {A : UU l} (l1 l2 : list A) → + (a : A) → a ∈-list (concat-list l1 l2) → + type-Prop (in-list l1 a) + type-Prop (in-list l2 a) +in-sum-concat-list-Prop l1 l2 a in-concat with in-concat-list l1 l2 a in-concat +... | inl in-l1 = inl (unit-trunc-Prop in-l1) +... | inr in-l2 = inr (unit-trunc-Prop in-l2) + +union-lists-concat : + {l : Level} {A : UU l} (l1 l2 : list A) → + in-list (concat-list l1 l2) ⊆ union-subtype (in-list l1) (in-list l2) +union-lists-concat l1 l2 a = + map-universal-property-trunc-Prop + ( union-subtype (in-list l1) (in-list l2) a) + ( unit-trunc-Prop ∘ (in-sum-concat-list-Prop l1 l2 a)) + +in-concat-left : + {l : Level} {A : UU l} (l1 l2 : list A) + {a : A} → a ∈-list l1 → a ∈-list (concat-list l1 l2) +in-concat-left _ _ (is-head a _) = + is-head a _ +in-concat-left _ l2 (is-in-tail a x l1 in-l1) = + is-in-tail a x (concat-list l1 l2) (in-concat-left l1 l2 in-l1) + +in-concat-right : + {l : Level} {A : UU l} (l1 l2 : list A) + {a : A} → a ∈-list l2 → a ∈-list (concat-list l1 l2) +in-concat-right nil l2 in-l2 = in-l2 +in-concat-right (cons x l1) l2 in-l2 = + is-in-tail _ x (concat-list l1 l2) (in-concat-right l1 l2 in-l2) + +in-concat-list-sum-Prop : + {l : Level} {A : UU l} (l1 l2 : list A) → + (a : A) → type-Prop (in-list l1 a) + type-Prop (in-list l2 a) → + type-Prop (in-list (concat-list l1 l2) a) +in-concat-list-sum-Prop l1 l2 a (inl in-l1) = + apply-universal-property-trunc-Prop + ( in-l1) + ( in-list (concat-list l1 l2) a) + ( unit-trunc-Prop ∘ (in-concat-left l1 l2)) +in-concat-list-sum-Prop l1 l2 a (inr in-l2) = + apply-universal-property-trunc-Prop + ( in-l2) + ( in-list (concat-list l1 l2) a) + ( unit-trunc-Prop ∘ (in-concat-right l1 l2)) + +in-concat-lists-union : + {l : Level} {A : UU l} (l1 l2 : list A) → + union-subtype (in-list l1) (in-list l2) ⊆ in-list (concat-list l1 l2) +in-concat-lists-union l1 l2 a = + map-universal-property-trunc-Prop + ( in-list (concat-list l1 l2) a) + ( in-concat-list-sum-Prop l1 l2 a) + +empty-in-list-nil : + {l : Level} {A : UU l} {x : A} → is-in-subtype (in-list nil) x → empty +empty-in-list-nil = map-universal-property-trunc-Prop empty-Prop ( λ ()) + +equiv-subset-head-tail : + {l1 l2 : Level} {A : UU l1} (x : A) (xs : list A) (a : subtype l2 A) → + (leq-prop-subtype (in-list (cons x xs)) a) ⇔ + prod-Prop (a x) (leq-prop-subtype (in-list xs) a) +pr1 (equiv-subset-head-tail x xs a) sub = + pair + ( sub x (unit-trunc-Prop (is-head x xs))) + ( transitive-leq-subtype + ( in-list xs) + ( in-list (cons x xs)) + ( a) + ( sub) + ( subset-in-tail x xs)) +pr2 (equiv-subset-head-tail x xs a) (x-in-a , sub) y = + map-universal-property-trunc-Prop + ( a y) + ( λ + { (is-head .x .xs) → x-in-a + ; (is-in-tail .y .x .xs y-in-list) → sub y (unit-trunc-Prop y-in-list)}) + +lists-in-union-lists : + {l1 l2 l3 : Level} {A : UU l1} + (xs : list A) (a : subtype l2 A) (b : subtype l3 A) → + in-list xs ⊆ union-subtype a b → + ∃ + ( list A) + ( λ xsl → + ( Σ + ( list A) + ( λ xsr → + ( in-list xs ⊆ union-subtype (in-list xsl) (in-list xsr)) × + ( in-list xsl ⊆ a) × + ( in-list xsr ⊆ b)))) +lists-in-union-lists nil a b sub = + unit-trunc-Prop + ( triple nil nil + ( triple + ( λ _ → ex-falso ∘ empty-in-list-nil) + ( λ _ → ex-falso ∘ empty-in-list-nil) + ( λ _ → ex-falso ∘ empty-in-list-nil))) +lists-in-union-lists (cons x xs) a b sub = + apply-twice-universal-property-trunc-Prop + ( lists-in-union-lists xs a b + ( transitive-leq-subtype + ( in-list xs) + ( in-list (cons x xs)) + ( union-subtype a b) + ( sub) + ( subset-in-tail x xs))) + ( sub x (unit-trunc-Prop (is-head x xs))) + ( ∃-Prop _ (λ _ → Σ _ _)) + ( λ (xsl , xsr , sub-lists , sub-xsl , sub-xsr) → + ( ind-coprod + ( λ _ → _) + ( λ x-in-a → + ( unit-trunc-Prop + ( triple (cons x xsl) xsr + ( triple + ( λ y → + ( map-universal-property-trunc-Prop + ( union-subtype (in-list (cons x xsl)) (in-list xsr) y) + ( λ + { (is-head .x .xs) → + ( subtype-union-left + ( in-list (cons x xsl)) + ( in-list xsr) + ( y) + ( unit-trunc-Prop (is-head y xsl))) + ; (is-in-tail .y .x .xs y-in-xs) → + ( subset-union-subset-left + ( in-list xsl) + ( in-list (cons x xsl)) + ( in-list xsr) + ( subset-in-tail x xsl) + ( y) + ( sub-lists y (unit-trunc-Prop y-in-xs)))}))) + ( backward-implication + ( equiv-subset-head-tail x xsl a) + ( x-in-a , sub-xsl)) + ( sub-xsr))))) + ( λ x-in-b → + ( unit-trunc-Prop + ( triple xsl (cons x xsr) + ( triple + ( λ y → + ( map-universal-property-trunc-Prop + ( union-subtype (in-list xsl) (in-list (cons x xsr)) y) + ( λ + { (is-head .x .xs) → + ( subtype-union-right + ( in-list xsl) + ( in-list (cons x xsr)) + ( y) + ( unit-trunc-Prop (is-head y xsr))) + ; (is-in-tail .y .x .xs y-in-xs) → + ( subset-union-subset-right + ( in-list xsl) + ( in-list xsr) + ( in-list (cons x xsr)) + ( subset-in-tail x xsr) + ( y) + ( sub-lists y (unit-trunc-Prop y-in-xs)))}))) + ( sub-xsl) + ( backward-implication + ( equiv-subset-head-tail x xsr b) + ( x-in-b , sub-xsr)))))))) + +module _ + {l1 : Level} {i : Set l1} + where + + list-assumptions-weak-deduction : + {l2 : Level} {logic : formulas l2 i} {a : formula i} → + logic ▷ a → + Σ (list (formula i)) (λ l → (in-list l ⊆ logic) × (in-list l ▷ a)) + pr1 (list-assumptions-weak-deduction {a = a} (w-ax x)) = + unit-list a + pr1 (pr2 (list-assumptions-weak-deduction {logic = logic} (w-ax x))) b = + map-universal-property-trunc-Prop ( logic b) ( λ { (is-head _ _) → x}) + pr2 (pr2 (list-assumptions-weak-deduction {a = a} (w-ax x))) = + w-ax (unit-trunc-Prop (is-head a nil)) + pr1 (list-assumptions-weak-deduction (w-mp wdcb wdc)) = + concat-list + ( pr1 (list-assumptions-weak-deduction wdcb)) + ( pr1 (list-assumptions-weak-deduction wdc)) + pr1 (pr2 (list-assumptions-weak-deduction {logic = logic} (w-mp wdcb wdc))) + b b-in-list = + subtype-union-both + ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) + ( in-list (pr1 (list-assumptions-weak-deduction wdc))) + ( logic) + ( pr1 (pr2 (list-assumptions-weak-deduction wdcb))) + ( pr1 (pr2 (list-assumptions-weak-deduction wdc))) + ( b) + ( union-lists-concat + ( pr1 (list-assumptions-weak-deduction wdcb)) + ( pr1 (list-assumptions-weak-deduction wdc)) + ( b) + ( b-in-list)) + pr2 (pr2 (list-assumptions-weak-deduction (w-mp wdcb wdc))) = + w-mp + ( weak-deduction-monotic + ( transitive-leq-subtype + ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) + ( union-subtype + ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) + ( in-list (pr1 (list-assumptions-weak-deduction wdc)))) + ( in-list + ( concat-list + ( pr1 (list-assumptions-weak-deduction wdcb)) + ( pr1 (list-assumptions-weak-deduction wdc)))) + ( in-concat-lists-union + ( pr1 (list-assumptions-weak-deduction wdcb)) + ( pr1 (list-assumptions-weak-deduction wdc))) + ( subtype-union-left + ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) + ( in-list (pr1 (list-assumptions-weak-deduction wdc))))) + ( pr2 (pr2 (list-assumptions-weak-deduction wdcb)))) + ( weak-deduction-monotic + ( transitive-leq-subtype + ( in-list (pr1 (list-assumptions-weak-deduction wdc))) + ( union-subtype + ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) + ( in-list (pr1 (list-assumptions-weak-deduction wdc)))) + ( in-list + ( concat-list + ( pr1 (list-assumptions-weak-deduction wdcb)) + ( pr1 (list-assumptions-weak-deduction wdc)))) + ( in-concat-lists-union + ( pr1 (list-assumptions-weak-deduction wdcb)) + ( pr1 (list-assumptions-weak-deduction wdc))) + ( subtype-union-right + ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) + ( in-list (pr1 (list-assumptions-weak-deduction wdc))))) + ( pr2 (pr2 (list-assumptions-weak-deduction wdc)))) + + nil-no-deduction : {a : formula i} → ¬ (in-list nil ▷ a) + nil-no-deduction (w-ax x) = + apply-universal-property-trunc-Prop x empty-Prop (λ ()) + nil-no-deduction (w-mp wd _) = nil-no-deduction wd + + -- list-assumptions-weak-modal-logic : + -- {l2 : Level} {logic : formulas l2 i} {a : formula i} → + -- is-in-subtype (weak-modal-logic logic) a → + -- ∃ ( list (formula i)) + -- ( λ l → + -- ( in-list l ⊆ logic) × + -- ( is-in-subtype (weak-modal-logic (in-list l)) a)) + -- list-assumptions-weak-modal-logic {l2} {logic} {a} = + -- map-universal-property-trunc-Prop + -- ( ∃-Prop + -- ( list (formula i)) + -- ( λ l → + -- ( in-list l ⊆ logic) × + -- ( is-in-subtype (weak-modal-logic (in-list l)) a)) ) + -- ( λ wda → + -- ( let (l , sub , comp) = list-assumptions-weak-deduction wda + -- in unit-trunc-Prop (l , sub , (unit-trunc-Prop comp)))) + + forward-subset-head-add : + (a : formula i) (l : list (formula i)) → + theory-add-formula a (in-list l) ⊆ in-list (cons a l) + forward-subset-head-add a l = + subtype-union-both + ( Id-formula-Prop a) + ( in-list l) + ( in-list (cons a l)) + ( λ { .a refl → unit-trunc-Prop (is-head a l)}) + ( subset-in-tail a l) + + backward-subset-head-add : + (a : formula i) (l : list (formula i)) → + in-list (cons a l) ⊆ theory-add-formula a (in-list l) + backward-subset-head-add a l x = + map-universal-property-trunc-Prop + ( theory-add-formula a (in-list l) x) + ( λ + { (is-head .a .l) → + ( subtype-union-left (Id-formula-Prop a) (in-list l) a refl) + ; (is-in-tail .x .a .l t) → + ( subtype-union-right + ( Id-formula-Prop a) + ( in-list l) + ( x) + ( unit-trunc-Prop t))}) + module _ {l1 l2 : Level} {i : Set l1} (axioms : formulas l2 i) where @@ -154,7 +496,7 @@ module _ backward-deduction-lemma : {a b : formula i} → axioms ▷ a ⇒ b → - is-in-subtype (weak-modal-logic (theory-add-formula a axioms)) (b) + is-in-subtype (weak-modal-logic (theory-add-formula a axioms)) b backward-deduction-lemma {a} wab = unit-trunc-Prop ( w-mp diff --git a/src/order-theory.lagda.md b/src/order-theory.lagda.md index b6ab4908bf..2c49831616 100644 --- a/src/order-theory.lagda.md +++ b/src/order-theory.lagda.md @@ -10,6 +10,7 @@ open import order-theory.bottom-elements-posets public open import order-theory.bottom-elements-preorders public open import order-theory.chains-posets public open import order-theory.chains-preorders public +open import order-theory.chains-union public open import order-theory.closure-operators-large-locales public open import order-theory.closure-operators-large-posets public open import order-theory.commuting-squares-of-galois-connections-large-posets public @@ -101,6 +102,7 @@ open import order-theory.similarity-of-order-preserving-maps-large-posets public open import order-theory.similarity-of-order-preserving-maps-large-preorders public open import order-theory.subposets public open import order-theory.subpreorders public +open import order-theory.subtypes-leq-posets public open import order-theory.suplattices public open import order-theory.top-elements-large-posets public open import order-theory.top-elements-posets public diff --git a/src/order-theory/chains-posets.lagda.md b/src/order-theory/chains-posets.lagda.md index 609bdfb03e..8e3f7b7748 100644 --- a/src/order-theory/chains-posets.lagda.md +++ b/src/order-theory/chains-posets.lagda.md @@ -7,7 +7,11 @@ module order-theory.chains-posets where
Imports ```agda +open import foundation.dependent-pair-types +open import foundation.existential-quantification +open import foundation.function-types open import foundation.propositions +open import foundation.subtypes open import foundation.universe-levels open import order-theory.chains-preorders @@ -74,4 +78,28 @@ module _ is-prop (inclusion-chain-Poset C D) is-prop-inclusion-chain-Poset = is-prop-inclusion-chain-Preorder (preorder-Poset X) + +-- TODO: move to separate file +module _ + {l1 l2 l3 : Level} (X : Poset l1 l2) (C : chain-Poset l3 X) + where + + is-chain-upper-bound-Prop : (x : type-Poset X) → Prop (l1 ⊔ l2 ⊔ l3) + is-chain-upper-bound-Prop x = + Π-Prop + ( type-chain-Poset X C) + ( λ y → + ( leq-Poset-Prop X + ( inclusion-subtype (sub-preorder-chain-Poset X C) y) + ( x))) + + is-chain-upper-bound : (x : type-Poset X) → UU (l1 ⊔ l2 ⊔ l3) + is-chain-upper-bound = type-Prop ∘ is-chain-upper-bound-Prop + + has-chain-upper-bound-Prop : Prop (l1 ⊔ l2 ⊔ l3) + has-chain-upper-bound-Prop = + exists-Prop (type-Poset X) is-chain-upper-bound-Prop + + has-chain-upper-bound : UU (l1 ⊔ l2 ⊔ l3) + has-chain-upper-bound = type-Prop has-chain-upper-bound-Prop ``` diff --git a/src/order-theory/chains-union.lagda.md b/src/order-theory/chains-union.lagda.md new file mode 100644 index 0000000000..3acee548c3 --- /dev/null +++ b/src/order-theory/chains-union.lagda.md @@ -0,0 +1,121 @@ +# Chain union + +```agda +module order-theory.chains-union where +``` + +
Imports + +```agda +open import foundation.dependent-pair-types +open import foundation.equivalences +open import foundation.existential-quantification +open import foundation.function-types +open import foundation.propositional-resizing +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.subtypes +open import foundation.unit-type +open import foundation.universe-levels + +open import order-theory.chains-posets +open import order-theory.maximal-elements-posets +open import order-theory.posets +open import order-theory.subposets +open import order-theory.subtypes-leq-posets +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l1 l2 l3 : Level} (A : UU l1) (C : chain-Poset l3 (subtypes-leq-Poset l2 A)) + where + + in-chain-union : subtype (l1 ⊔ lsuc l2 ⊔ l3) A + in-chain-union a = + exists-Prop + ( type-chain-Poset (subtypes-leq-Poset l2 A) C) + ( λ x → + ( inclusion-subtype + ( sub-preorder-chain-Poset (subtypes-leq-Poset l2 A) C) + ( x) + ( a))) + + is-chain-union : + {l4 : Level} (u : subtype l4 A) → UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ l4) + is-chain-union = equiv-subtypes in-chain-union + + chain-union-is-upper-bound : + (u : subtype l2 A) → + is-chain-union u → + is-chain-upper-bound (subtypes-leq-Poset l2 A) C u + chain-union-is-upper-bound u u-is-union s a a-in-s = + map-equiv (u-is-union a) (intro-∃ s a-in-s) + + module _ + (prop-resize : propositional-resizing l2 (l1 ⊔ lsuc l2 ⊔ l3)) + where + + chain-union : Σ (subtype l2 A) is-chain-union + pr1 chain-union = resize prop-resize ∘ in-chain-union + pr2 chain-union = inv-equiv ∘ is-equiv-resize prop-resize ∘ in-chain-union + +module _ + {l1 l2 l3 l4 : Level} (A : UU l1) + (S : subtype l3 (subtype l2 A)) + (C : chain-Poset l4 (poset-Subposet (subtypes-leq-Poset l2 A) S)) + where + + private + poset : Poset (l1 ⊔ lsuc l2 ⊔ l3) (l1 ⊔ l2) + poset = poset-Subposet (subtypes-leq-Poset l2 A) S + + in-sub-chain-union : subtype (l1 ⊔ lsuc l2 ⊔ l3 ⊔ l4) A + in-sub-chain-union a = + exists-Prop + ( type-chain-Poset poset C) + ( λ x → + ( inclusion-subtype S + ( inclusion-subtype (sub-preorder-chain-Poset poset C) x) a)) + + is-sub-chain-union : (u : type-Poset poset) → UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ l4) + is-sub-chain-union u = + equiv-subtypes in-sub-chain-union (inclusion-subtype S u) + + sub-chain-union-is-upper-bound : + (u : type-Poset poset) → + is-sub-chain-union u → + is-chain-upper-bound poset C u + sub-chain-union-is-upper-bound u u-is-union s a a-in-s = + map-equiv (u-is-union a) (intro-∃ s a-in-s) + + module _ + (x : type-chain-Poset poset C) + (subtype-monotic : + (y : subtype l2 A) → + inclusion-subtype S + ( inclusion-subtype (sub-preorder-chain-Poset poset C) x) ⊆ y → + is-in-subtype S y) + (prop-resize : propositional-resizing l2 (l1 ⊔ lsuc l2 ⊔ l3 ⊔ l4)) + where + + sub-chain-union : Σ (type-Poset poset) is-sub-chain-union + pr1 (pr1 sub-chain-union) = + resize prop-resize ∘ in-sub-chain-union + pr2 (pr1 sub-chain-union) = + subtype-monotic + ( resize prop-resize ∘ in-sub-chain-union) + ( λ a a-in-x → + ( map-equiv + ( inv-equiv (is-equiv-resize prop-resize (in-sub-chain-union a))) + ( intro-∃ x a-in-x))) + pr2 sub-chain-union = + inv-equiv ∘ is-equiv-resize prop-resize ∘ in-sub-chain-union +``` diff --git a/src/order-theory/maximal-elements-posets.lagda.md b/src/order-theory/maximal-elements-posets.lagda.md index a60677c230..c52a961f44 100644 --- a/src/order-theory/maximal-elements-posets.lagda.md +++ b/src/order-theory/maximal-elements-posets.lagda.md @@ -7,12 +7,26 @@ module order-theory.maximal-elements-posets where
Imports ```agda +open import foundation.action-on-identifications-functions open import foundation.dependent-pair-types +open import foundation.empty-types +open import foundation.existential-quantification open import foundation.identity-types +open import foundation.inhabited-types +open import foundation.law-of-excluded-middle +open import foundation.logical-equivalences +open import foundation.propositional-truncations open import foundation.propositions +open import foundation.subtypes +open import foundation.unit-type open import foundation.universe-levels +open import foundation-core.coproduct-types + +open import order-theory.chains-posets open import order-theory.posets +open import order-theory.subposets +open import order-theory.upper-bounds-posets ```
@@ -44,4 +58,70 @@ module _ has-maximal-element-Poset : UU (l1 ⊔ l2) has-maximal-element-Poset = Σ (type-Poset X) is-maximal-element-Poset + +module _ + (l1 l2 l3 : Level) + where + + Zorn-Prop : Prop (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3) + Zorn-Prop = + Π-Prop + ( Poset l1 l2) + ( λ X → + ( function-Prop + ( is-inhabited (type-Poset X)) + ( function-Prop + ( (C : chain-Poset l3 X) → + is-inhabited (type-chain-Poset X C) → + has-chain-upper-bound X C) + ( ∃-Prop (type-Poset X) (is-maximal-element-Poset X))))) + + Zorn : UU (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3) + Zorn = type-Prop Zorn-Prop + + Zorn'-Prop : Prop (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3) + Zorn'-Prop = + Π-Prop + ( Poset l1 l2) + ( λ X → + ( function-Prop + ( (C : chain-Poset l3 X) → has-chain-upper-bound X C) + ( ∃-Prop (type-Poset X) (is-maximal-element-Poset X)))) + + Zorn' : UU (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3) + Zorn' = type-Prop Zorn'-Prop + + Zorn'-Zorn : Zorn → Zorn' + Zorn'-Zorn zorn X H = + zorn X + ( apply-universal-property-trunc-Prop + ( H + ( pair + (λ _ → raise-empty-Prop l3) + (λ (_ , f) → raise-ex-falso l3 f))) + ( is-inhabited-Prop (type-Poset X)) + ( λ (x , _) → unit-trunc-Prop x)) + ( λ C _ → H C) + + module _ + (lem : LEM (l1 ⊔ l3)) + where + + Zorn-Zorn' : Zorn' → Zorn + Zorn-Zorn' zorn' X is-inhabited-X H = zorn' X chain-upper-bound + where + chain-upper-bound : (C : chain-Poset l3 X) → has-chain-upper-bound X C + chain-upper-bound C with lem (is-inhabited-Prop (type-chain-Poset X C)) + ... | inl is-inhabited-C = H C is-inhabited-C + ... | inr is-empty-C = + apply-universal-property-trunc-Prop + ( is-inhabited-X) + ( has-chain-upper-bound-Prop X C) + ( λ x → + ( intro-∃ x + ( λ (y , y-in-C) → ex-falso (is-empty-C (intro-∃ y y-in-C))))) + + Zorn-iff-Zorn' : Zorn-Prop ⇔ Zorn'-Prop + pr1 (Zorn-iff-Zorn') = Zorn'-Zorn + pr2 (Zorn-iff-Zorn') = Zorn-Zorn' ``` diff --git a/src/order-theory/subtypes-leq-posets.lagda.md b/src/order-theory/subtypes-leq-posets.lagda.md new file mode 100644 index 0000000000..e0149381c3 --- /dev/null +++ b/src/order-theory/subtypes-leq-posets.lagda.md @@ -0,0 +1,39 @@ +# Subtypes leq Posets +```agda +module order-theory.subtypes-leq-posets where +``` + +
Imports + +```agda +open import foundation.dependent-pair-types +open import foundation.subtypes +open import foundation.universe-levels + +open import order-theory.posets +open import order-theory.preorders +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l1 : Level} (l2 : Level) (A : UU l1) + where + + subtypes-leq-Preorder : Preorder (l1 ⊔ lsuc l2) (l1 ⊔ l2) + pr1 subtypes-leq-Preorder = subtype l2 A + pr1 (pr2 subtypes-leq-Preorder) = leq-prop-subtype + pr1 (pr2 (pr2 subtypes-leq-Preorder)) = refl-leq-subtype + pr2 (pr2 (pr2 subtypes-leq-Preorder)) = transitive-leq-subtype + + subtypes-leq-Poset : Poset (l1 ⊔ lsuc l2) (l1 ⊔ l2) + pr1 subtypes-leq-Poset = subtypes-leq-Preorder + pr2 subtypes-leq-Poset = antisymmetric-leq-subtype +``` From 9e3bb018f3136d3e192911b6ffa5e4456b029b74 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Mon, 5 Feb 2024 21:22:54 +0300 Subject: [PATCH 28/63] fix pre-commit --- src/order-theory/subtypes-leq-posets.lagda.md | 1 + 1 file changed, 1 insertion(+) diff --git a/src/order-theory/subtypes-leq-posets.lagda.md b/src/order-theory/subtypes-leq-posets.lagda.md index e0149381c3..9a96cc9b6b 100644 --- a/src/order-theory/subtypes-leq-posets.lagda.md +++ b/src/order-theory/subtypes-leq-posets.lagda.md @@ -1,4 +1,5 @@ # Subtypes leq Posets + ```agda module order-theory.subtypes-leq-posets where ``` From fd1cc318f899e82cd1287a03b7628be3a8a2221d Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Tue, 6 Feb 2024 18:07:26 +0300 Subject: [PATCH 29/63] Prove the second point of canonical model theorem --- .../canonical-model-theorem.lagda.md | 81 +++++++++++++++++++ 1 file changed, 81 insertions(+) diff --git a/src/modal-logic/canonical-model-theorem.lagda.md b/src/modal-logic/canonical-model-theorem.lagda.md index aedb36fa63..2f450f9eb1 100644 --- a/src/modal-logic/canonical-model-theorem.lagda.md +++ b/src/modal-logic/canonical-model-theorem.lagda.md @@ -1065,4 +1065,85 @@ module _ ( backward-implication ( canonical-model-theorem a (y , is-canonical)) ( f (y , is-canonical) access))) + + canonical-model-theorem' : + (a : formula i) → + neg-Prop (logic a) ⇔ neg-Prop (canonical-kripke-model ⊨M a) + pr1 (canonical-model-theorem' a) nf in-logic = + apply-universal-property-trunc-Prop + ( lindenbaum x) + ( empty-Prop) + ( λ w → + ( pr2 (pr1 w) + ( weak-modal-logic-mp + { a = a} + ( weak-modal-logic-ax + ( transitive-leq-subtype + ( theory-add-formula (~ a) logic) + ( pr1 (pr1 w)) + ( union-subtype logic (pr1 (pr1 w))) + ( subtype-union-right logic (pr1 (pr1 w))) + ( pr1 (pr2 w)) + ( ~ a) + ( formula-in-add-formula (~ a) logic))) + ( weak-modal-logic-ax + ( subtype-union-right + ( logic) + ( pr1 (pr1 w)) + ( a) + ( backward-implication + ( canonical-model-theorem a + ( pr1 (pr1 w) , pr2 (pr1 w) , pr2 (pr2 w))) + ( in-logic + ( pr1 (pr1 w) , pr2 (pr1 w) , pr2 (pr2 w))))))))) + where + x : L-consistent-theory + pr1 x = theory-add-formula (~ a) logic + pr2 x bot-in-logic = + nf + ( modal-logic-mp + { a = ~~ a} + { b = a} + ( contains-ax-dn _ (a , refl)) + ( subset-double-modal-logic + ( axioms) + ( ~~ a) + ( weak-modal-logic-subset-modal-logic + ( logic) + ( ~~ a) + ( forward-implication + ( deduction-lemma logic + ( transitive-leq-subtype + ( ax-k i) + ( logic) + ( weak-modal-logic logic) + ( axioms-subset-weak-modal-logic logic) + ( contains-ax-k)) + ( transitive-leq-subtype + ( ax-s i) + ( logic) + ( weak-modal-logic logic) + ( axioms-subset-weak-modal-logic logic) + ( contains-ax-s)) + ( ~ a) + ( ⊥)) + ( weak-modal-logic-monotic + ( subtype-union-both + ( logic) + ( theory-add-formula (~ a) logic) + ( theory-add-formula (~ a) logic) + ( subtype-union-right (Id-formula-Prop (~ a)) logic) + ( refl-leq-subtype (theory-add-formula (~ a) logic))) + ( ⊥) + ( bot-in-logic)))))) + pr2 (canonical-model-theorem' a) = + map-neg + ( λ in-logic x → + ( forward-implication + ( canonical-model-theorem a x) + ( logic-subset-L-complete-theory + ( pr1 x , pr1 (pr2 x)) + ( pr2 (pr2 x)) + ( a) + ( in-logic)))) ``` From 39bccd5202607a24b486bd43b48bb2a6919963f5 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Wed, 7 Feb 2024 00:02:15 +0300 Subject: [PATCH 30/63] Prove completeness K --- src/modal-logic.lagda.md | 1 + .../canonical-model-theorem.lagda.md | 44 +++++++- src/modal-logic/completeness-K.lagda.md | 84 +++++++++++++++ src/modal-logic/completeness.lagda.md | 51 +++++---- src/modal-logic/modal-logic-IK.lagda.md | 41 +++++++ src/modal-logic/modal-logic-K.lagda.md | 102 +++++++++++++++--- 6 files changed, 282 insertions(+), 41 deletions(-) create mode 100644 src/modal-logic/completeness-K.lagda.md diff --git a/src/modal-logic.lagda.md b/src/modal-logic.lagda.md index e3e2b011b7..d6859a669d 100644 --- a/src/modal-logic.lagda.md +++ b/src/modal-logic.lagda.md @@ -6,6 +6,7 @@ module modal-logic where open import modal-logic.axioms public open import modal-logic.canonical-model-theorem public open import modal-logic.completeness public +open import modal-logic.completeness-K public open import modal-logic.formulas public open import modal-logic.kripke-semantics public open import modal-logic.logic-syntax public diff --git a/src/modal-logic/canonical-model-theorem.lagda.md b/src/modal-logic/canonical-model-theorem.lagda.md index 2f450f9eb1..fbbfe28757 100644 --- a/src/modal-logic/canonical-model-theorem.lagda.md +++ b/src/modal-logic/canonical-model-theorem.lagda.md @@ -42,6 +42,7 @@ open import modal-logic.completeness open import modal-logic.formulas open import modal-logic.kripke-semantics open import modal-logic.logic-syntax +open import modal-logic.modal-logic-K open import modal-logic.soundness open import modal-logic.weak-deduction @@ -68,15 +69,48 @@ module _ (zorn : Zorn (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) l2) (prop-resize : propositional-resizing (l1 ⊔ l2) (lsuc (l1 ⊔ l2))) (L-is-cons : is-consistent-modal-logic (modal-logic axioms)) - (contains-ax-k : ax-k i ⊆ modal-logic axioms) - (contains-ax-s : ax-s i ⊆ modal-logic axioms) - (contains-ax-dn : ax-dn i ⊆ modal-logic axioms) - (contains-ax-n : ax-n i ⊆ modal-logic axioms) + (contains-K : modal-logic-K i ⊆ modal-logic axioms) where logic : formulas (l1 ⊔ l2) i logic = modal-logic axioms + contains-ax-k : ax-k i ⊆ logic + contains-ax-k = + transitive-leq-subtype + ( ax-k i) + ( modal-logic-K i) + ( logic) + ( contains-K) + ( K-contains-ax-k i) + + contains-ax-s : ax-s i ⊆ logic + contains-ax-s = + transitive-leq-subtype + ( ax-s i) + ( modal-logic-K i) + ( logic) + ( contains-K) + ( K-contains-ax-s i) + + contains-ax-n : ax-n i ⊆ logic + contains-ax-n = + transitive-leq-subtype + ( ax-n i) + ( modal-logic-K i) + ( logic) + ( contains-K) + ( K-contains-ax-n i) + + contains-ax-dn : ax-dn i ⊆ logic + contains-ax-dn = + transitive-leq-subtype + ( ax-dn i) + ( modal-logic-K i) + ( logic) + ( contains-K) + ( K-contains-ax-dn i) + is-L-consistent-theory-Prop : formulas (l1 ⊔ l2) i → Prop (l1 ⊔ l2) is-L-consistent-theory-Prop t = is-consistent-modal-logic-Prop (weak-modal-logic (union-subtype logic t)) @@ -180,7 +214,7 @@ module _ complete-theory-contains-all-formulas : LEM (l1 ⊔ l2) → (x : L-consistent-theory) → is-L-complete-theory x → - (a : formula i) → type-disj-Prop ((pr1 x) a) ((pr1 x) (~ a)) + (a : formula i) → type-disj-Prop (pr1 x a) (pr1 x (~ a)) complete-theory-contains-all-formulas lem x is-comp a with lem ((pr1 x) a) ... | inl a-in-logic = inl-disj-Prop ((pr1 x) a) ((pr1 x) (~ a)) a-in-logic ... | inr a-not-in-logic = diff --git a/src/modal-logic/completeness-K.lagda.md b/src/modal-logic/completeness-K.lagda.md new file mode 100644 index 0000000000..ec1100e427 --- /dev/null +++ b/src/modal-logic/completeness-K.lagda.md @@ -0,0 +1,84 @@ +# Completeness of K + +```agda +module modal-logic.completeness-K where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.cartesian-product-types +open import foundation.coproduct-types +open import foundation.dependent-pair-types +open import foundation.disjunction +open import foundation.empty-types +open import foundation.existential-quantification +open import foundation.function-types +open import foundation.identity-types +open import foundation.inhabited-types +open import foundation.intersections-subtypes +open import foundation.law-of-excluded-middle +open import foundation.logical-equivalences +open import foundation.negation +open import foundation.propositional-resizing +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.raising-universe-levels +open import foundation.sets +open import foundation.subtypes +open import foundation.transport-along-identifications +open import foundation.unions-subtypes +open import foundation.unit-type +open import foundation.universe-levels + +open import modal-logic.axioms +open import modal-logic.canonical-model-theorem +open import modal-logic.completeness +open import modal-logic.formulas +open import modal-logic.kripke-semantics +open import modal-logic.logic-syntax +open import modal-logic.modal-logic-IK +open import modal-logic.modal-logic-K +open import modal-logic.soundness +open import modal-logic.weak-deduction + +open import order-theory.maximal-elements-posets +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +-- TODO: levels +module _ + {l1 : Level} + (i : Set l1) + (lem : LEM l1) + (zorn : Zorn (lsuc l1) l1 l1) + (prop-resize : propositional-resizing l1 (lsuc l1)) + where + + completeness-K : + completeness (modal-logic-K i) (lsuc l1) (λ w → all-models w l1 i l1) + completeness-K a in-kripke-logic with lem ((modal-logic-K i) a) + ... | inl in-logic = in-logic + ... | inr not-in-logic = + ex-falso + ( forward-implication + ( canonical-model-theorem' + ( union-subtype (modal-logic-IK i) (ax-dn i)) + ( zorn) + ( prop-resize) + ( is-consistent-K i) + ( refl-leq-subtype (modal-logic-K i)) + ( lem) + ( a)) + ( not-in-logic) + ( in-kripke-logic _ _ star)) +``` diff --git a/src/modal-logic/completeness.lagda.md b/src/modal-logic/completeness.lagda.md index 1ac2fdba26..6af35273d9 100644 --- a/src/modal-logic/completeness.lagda.md +++ b/src/modal-logic/completeness.lagda.md @@ -34,33 +34,46 @@ TODO ## Definition ```agda +-- module _ +-- {l1 l2 l3 l4 l5 l6 : Level} +-- {i : Set l1} (logic : formulas l2 i) +-- {w : UU l3} (C : model-class w l4 i l5 l6) +-- where + +-- completeness : UU (l1 ⊔ l2 ⊔ l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6) +-- completeness = class-modal-logic C ⊆ logic + module _ - {l1 l2 l3 l4 l5 l6 : Level} + {l1 l2 l3 l4 l5 : Level} {i : Set l1} (logic : formulas l2 i) - {w : UU l3} (C : model-class w l4 i l5 l6) + (l6 : Level) (C : (w : UU l6) → model-class w l3 i l4 l5) where - completeness : UU (l1 ⊔ l2 ⊔ l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6) - completeness = class-modal-logic C ⊆ logic + completeness : UU (l1 ⊔ l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ l5 ⊔ lsuc l6) + -- completeness = (w : UU l6) → class-modal-logic (C w) ⊆ logic + completeness = + (a : formula i) → + ( (w : UU l6) → is-in-subtype (class-modal-logic (C w)) a) → + is-in-subtype logic a ``` ## Properties ```agda -module _ - {l1 l2 l3 l4 l5 l6 l7 : Level} - {i : Set l1} (logic : formulas l2 i) - {w : UU l3} - (C₁ : model-class w l4 i l5 l6) (C₂ : model-class w l4 i l5 l7) - where +-- module _ +-- {l1 l2 l3 l4 l5 l6 l7 : Level} +-- {i : Set l1} (logic : formulas l2 i) +-- {w : UU l3} +-- (C₁ : model-class w l4 i l5 l6) (C₂ : model-class w l4 i l5 l7) +-- where - completeness-subclass : - C₁ ⊆ C₂ → completeness logic C₁ → completeness logic C₂ - completeness-subclass sub complete = - transitive-leq-subtype - ( class-modal-logic C₂) - ( class-modal-logic C₁) - ( logic) - ( complete) - ( class-modal-logic-monotic C₁ C₂ sub) +-- completeness-subclass : +-- C₁ ⊆ C₂ → completeness logic C₁ → completeness logic C₂ +-- completeness-subclass sub complete = +-- transitive-leq-subtype +-- ( class-modal-logic C₂) +-- ( class-modal-logic C₁) +-- ( logic) +-- ( complete) +-- ( class-modal-logic-monotic C₁ C₂ sub) ``` diff --git a/src/modal-logic/modal-logic-IK.lagda.md b/src/modal-logic/modal-logic-IK.lagda.md index db30f4e61c..32c2781baf 100644 --- a/src/modal-logic/modal-logic-IK.lagda.md +++ b/src/modal-logic/modal-logic-IK.lagda.md @@ -10,6 +10,7 @@ module modal-logic.modal-logic-IK where open import foundation.inhabited-types open import foundation.intersections-subtypes open import foundation.sets +open import foundation.subtypes open import foundation.unions-subtypes open import foundation.universe-levels @@ -38,6 +39,46 @@ module _ modal-logic-IK = modal-logic (union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i))) + IK-contains-ax-k : ax-k i ⊆ modal-logic-IK + IK-contains-ax-k = + transitive-leq-subtype + ( ax-k i) + ( union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i))) + ( modal-logic (union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i)))) + ( axioms-subset-modal-logic + ( union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i)))) + ( subtype-union-left (ax-k i) (union-subtype (ax-s i) (ax-n i))) + + IK-contains-ax-s : ax-s i ⊆ modal-logic-IK + IK-contains-ax-s = + transitive-leq-subtype + ( ax-s i) + ( union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i))) + ( modal-logic (union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i)))) + ( axioms-subset-modal-logic + ( union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i)))) + ( transitive-leq-subtype + ( ax-s i) + ( union-subtype (ax-s i) (ax-n i)) + ( union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i))) + ( subtype-union-right (ax-k i) (union-subtype (ax-s i) (ax-n i))) + ( subtype-union-left (ax-s i) (ax-n i))) + + IK-contains-ax-n : ax-n i ⊆ modal-logic-IK + IK-contains-ax-n = + transitive-leq-subtype + ( ax-n i) + ( union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i))) + ( modal-logic (union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i)))) + ( axioms-subset-modal-logic + ( union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i)))) + ( transitive-leq-subtype + ( ax-n i) + ( union-subtype (ax-s i) (ax-n i)) + ( union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i))) + ( subtype-union-right (ax-k i) (union-subtype (ax-s i) (ax-n i))) + ( subtype-union-right (ax-s i) (ax-n i))) + module _ {l1 l2 : Level} (i : Set l1) diff --git a/src/modal-logic/modal-logic-K.lagda.md b/src/modal-logic/modal-logic-K.lagda.md index 69338a6124..dc3a2fb171 100644 --- a/src/modal-logic/modal-logic-K.lagda.md +++ b/src/modal-logic/modal-logic-K.lagda.md @@ -7,19 +7,30 @@ module modal-logic.modal-logic-K where
Imports ```agda +open import foundation.decidable-types +open import foundation.dependent-pair-types +open import foundation.empty-types open import foundation.inhabited-types open import foundation.intersections-subtypes +open import foundation.negation +open import foundation.propositional-truncations +open import foundation.raising-universe-levels open import foundation.sets open import foundation.subtypes open import foundation.unions-subtypes +open import foundation.unit-type open import foundation.universe-levels +open import foundation-core.coproduct-types + open import modal-logic.axioms open import modal-logic.formulas open import modal-logic.kripke-semantics open import modal-logic.logic-syntax open import modal-logic.modal-logic-IK open import modal-logic.soundness + +open import univalent-combinatorics.finite-types ```
@@ -36,7 +47,8 @@ module _ where modal-logic-K : formulas l i - modal-logic-K = modal-logic (union-subtype (modal-logic-IK i) (ax-dn i)) + modal-logic-K = + modal-logic (union-subtype (modal-logic-IK i) (ax-dn i)) IK-subset-K : modal-logic-IK i ⊆ modal-logic-K IK-subset-K = @@ -47,31 +59,87 @@ module _ ( axioms-subset-modal-logic _) ( subtype-union-left (modal-logic-IK i) (ax-dn i)) + K-contains-ax-k : ax-k i ⊆ modal-logic-K + K-contains-ax-k = + transitive-leq-subtype + ( ax-k i) + ( modal-logic-IK i) + ( modal-logic-K) + ( IK-subset-K) + ( IK-contains-ax-k i) + + K-contains-ax-s : ax-s i ⊆ modal-logic-K + K-contains-ax-s = + transitive-leq-subtype + ( ax-s i) + ( modal-logic-IK i) + ( modal-logic-K) + ( IK-subset-K) + ( IK-contains-ax-s i) + + K-contains-ax-n : ax-n i ⊆ modal-logic-K + K-contains-ax-n = + transitive-leq-subtype + ( ax-n i) + ( modal-logic-IK i) + ( modal-logic-K) + ( IK-subset-K) + ( IK-contains-ax-n i) + + K-contains-ax-dn : ax-dn i ⊆ modal-logic-K + K-contains-ax-dn = + transitive-leq-subtype + ( ax-dn i) + ( union-subtype (modal-logic-IK i) (ax-dn i)) + ( modal-logic (union-subtype (modal-logic-IK i) (ax-dn i))) + ( axioms-subset-modal-logic + ( union-subtype (modal-logic-IK i) (ax-dn i))) + ( subtype-union-right (modal-logic-IK i) (ax-dn i)) + module _ - {l1 l2 l3 : Level} - (i : Set l1) (axioms : formulas l2 i) - (w : UU l3) - {l4 l5 : Level} + {l1 l2 l3 l4 : Level} + (i : Set l1) (w : UU l2) where - soundness-K : soundness (modal-logic-K i) (decidable-models w l4 i l5) + soundness-K : soundness (modal-logic-K i) (decidable-models w l3 i l4) soundness-K = soundness-modal-logic-union-subclass-right-sublevels ( modal-logic-IK i) ( ax-dn i) - ( l1 ⊔ l3 ⊔ l4 ⊔ l5) - ( all-models w l4 i l5) - ( decidable-models w l4 i l5) - ( soundness-IK i w l4 l5) - ( ax-dn-soundness i w l4 l5) - ( all-models-is-largest-class w l4 i l5 (decidable-models w l4 i l5)) - - soundness-K-finite : soundness (modal-logic-K i) (finite-models w l4 i l5) + ( l1 ⊔ l2 ⊔ l3 ⊔ l4) + ( all-models w l3 i l4) + ( decidable-models w l3 i l4) + ( soundness-IK i w l3 l4) + ( ax-dn-soundness i w l3 l4) + ( all-models-is-largest-class w l3 i l4 (decidable-models w l3 i l4)) + + soundness-K-finite : soundness (modal-logic-K i) (finite-models w l3 i l4) soundness-K-finite = soundness-subclass ( modal-logic-K i) - ( decidable-models w l4 i l5) - ( finite-models w l4 i l5) - ( finite-models-subclass-decidable-models w l4 i l5) + ( decidable-models w l3 i l4) + ( finite-models w l3 i l4) + ( finite-models-subclass-decidable-models w l3 i l4) ( soundness-K) + +module _ + {l1 : Level} (i : Set l1) + where + + is-consistent-K : is-consistent-modal-logic (modal-logic-K i) + is-consistent-K bot-in-logic = + map-inv-raise + ( soundness-K-finite + ( i) + ( unit) + ( ⊥) + ( bot-in-logic) + ( pair + ( (unit-trunc-Prop star) , (λ _ _ → empty-Prop)) + ( λ _ _ → empty-Prop)) + ( triple + ( is-finite-unit) + ( λ _ _ → inr (λ r → r)) + ( λ _ _ → is-decidable-empty)) + ( star)) ``` From 200901aabb24c7bc20436ef06ac66a68a467a09f Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Wed, 14 Feb 2024 21:07:49 +0300 Subject: [PATCH 31/63] fix renamed functions --- src/foundation/unions-subtypes.lagda.md | 6 +- src/modal-logic/axioms.lagda.md | 16 ++--- .../canonical-model-theorem.lagda.md | 40 ++++++------ src/modal-logic/formulas.lagda.md | 62 +++++++++---------- src/modal-logic/kripke-semantics.lagda.md | 8 +-- src/modal-logic/logic-syntax.lagda.md | 4 +- src/modal-logic/weak-deduction.lagda.md | 40 ++++++------ 7 files changed, 88 insertions(+), 88 deletions(-) diff --git a/src/foundation/unions-subtypes.lagda.md b/src/foundation/unions-subtypes.lagda.md index 289a680c0e..ce86a03ba0 100644 --- a/src/foundation/unions-subtypes.lagda.md +++ b/src/foundation/unions-subtypes.lagda.md @@ -99,15 +99,15 @@ module _ where subtype-union-left : P ⊆ union-subtype P Q - subtype-union-left x = inl-disj-Prop (P x) (Q x) + subtype-union-left x = inl-disjunction-Prop (P x) (Q x) subtype-union-right : Q ⊆ union-subtype P Q - subtype-union-right x = inr-disj-Prop (P x) (Q x) + subtype-union-right x = inr-disjunction-Prop (P x) (Q x) subtype-union-both : {l4 : Level} (S : subtype l4 X) → P ⊆ S → Q ⊆ S → union-subtype P Q ⊆ S subtype-union-both S P-sub-S Q-sub-S x = - elim-disj-Prop (P x) (Q x) (S x) (P-sub-S x , Q-sub-S x) + elim-disjunction-Prop (P x) (Q x) (S x) (P-sub-S x , Q-sub-S x) module _ {l1 l2 l3 : Level} {X : UU l1} (P : subtype l2 X) (Q : subtype l3 X) diff --git a/src/modal-logic/axioms.lagda.md b/src/modal-logic/axioms.lagda.md index d8e9a8c9d9..d7f94b33a9 100644 --- a/src/modal-logic/axioms.lagda.md +++ b/src/modal-logic/axioms.lagda.md @@ -93,14 +93,14 @@ module _ ax-k : formulas l i ax-k = ax-2-parameters - ( λ a b → a ⇒ b ⇒ a) + ( λ a b → a →ₘ b →ₘ a) ( eq-implication-left) ( eq-implication-left ∘ eq-implication-right) ax-s : formulas l i ax-s = ax-3-parameters - ( λ a b c → (a ⇒ b ⇒ c) ⇒ (a ⇒ b) ⇒ a ⇒ c) + ( λ a b c → (a →ₘ b →ₘ c) →ₘ (a →ₘ b) →ₘ a →ₘ c) ( eq-implication-left ∘ eq-implication-left) ( eq-implication-left ∘ eq-implication-right ∘ eq-implication-left) ( eq-implication-right ∘ eq-implication-right ∘ eq-implication-left) @@ -108,12 +108,12 @@ module _ ax-n : formulas l i ax-n = ax-2-parameters - ( λ a b → □ (a ⇒ b) ⇒ □ a ⇒ □ b) + ( λ a b → □ (a →ₘ b) →ₘ □ a →ₘ □ b) ( eq-implication-left ∘ eq-box ∘ eq-implication-left) ( eq-implication-right ∘ eq-box ∘ eq-implication-left) ax-dn : formulas l i - ax-dn = ax-1-parameter ( λ a → ~~ a ⇒ a) ( eq-implication-right) + ax-dn = ax-1-parameter ( λ a → ~~ a →ₘ a) ( eq-implication-right) module _ {l1 l2 : Level} @@ -123,24 +123,24 @@ module _ where ax-k-soundness : soundness (ax-k i) (all-models w l3 i l4) - ax-k-soundness (a ⇒ b ⇒ a) (_ , _ , refl) M _ x fa _ = fa + ax-k-soundness (a →ₘ b →ₘ a) (_ , _ , refl) M _ x fa _ = fa ax-n-soundness : soundness (ax-n i) (all-models w l3 i l4) ax-n-soundness - (□ (a ⇒ b) ⇒ □ a ⇒ □ b) + (□ (a →ₘ b) →ₘ □ a →ₘ □ b) (_ , _ , refl) M in-class x fab fa y r = fab y r (fa y r) ax-s-soundness : soundness (ax-s i) (all-models w l3 i l4) ax-s-soundness - ((a ⇒ b ⇒ c) ⇒ (a ⇒ b) ⇒ a ⇒ c) + ((a →ₘ b →ₘ c) →ₘ (a →ₘ b) →ₘ a →ₘ c) (_ , _ , _ , refl) M in-class x fabc fab fa = fabc fa (fab fa) ax-dn-soundness : soundness (ax-dn i) (decidable-models w l3 i l4) - ax-dn-soundness (((a ⇒ ⊥) ⇒ ⊥) ⇒ a) (_ , refl) _ is-dec x f + ax-dn-soundness (((a →ₘ ⊥) →ₘ ⊥) →ₘ a) (_ , refl) _ is-dec x f with (is-dec a x) ... | inl fa = fa ... | inr fna = raise-ex-falso _ (f (λ fa -> map-raise (fna fa))) diff --git a/src/modal-logic/canonical-model-theorem.lagda.md b/src/modal-logic/canonical-model-theorem.lagda.md index fbbfe28757..d2b3f522dd 100644 --- a/src/modal-logic/canonical-model-theorem.lagda.md +++ b/src/modal-logic/canonical-model-theorem.lagda.md @@ -214,9 +214,9 @@ module _ complete-theory-contains-all-formulas : LEM (l1 ⊔ l2) → (x : L-consistent-theory) → is-L-complete-theory x → - (a : formula i) → type-disj-Prop (pr1 x a) (pr1 x (~ a)) + (a : formula i) → type-disjunction-Prop (pr1 x a) (pr1 x (~ a)) complete-theory-contains-all-formulas lem x is-comp a with lem ((pr1 x) a) - ... | inl a-in-logic = inl-disj-Prop ((pr1 x) a) ((pr1 x) (~ a)) a-in-logic + ... | inl a-in-logic = inl-disjunction-Prop ((pr1 x) a) ((pr1 x) (~ a)) a-in-logic ... | inr a-not-in-logic = unit-trunc-Prop ( inr @@ -287,7 +287,7 @@ module _ ( logic-subset-L-complete-theory ( x) ( is-comp) - ( ~~ a ⇒ a) + ( ~~ a →ₘ a) ( contains-ax-dn _ (a , refl)))) ( wd-bot)))))))) ( subtype-union-right (Id-formula-Prop (~ a)) (pr1 x))) @@ -297,16 +297,16 @@ module _ LEM (l1 ⊔ l2) → (x : L-consistent-theory) → is-L-complete-theory x → {a b : formula i} → (is-in-subtype (pr1 x) a → is-in-subtype (pr1 x) b) → - is-in-subtype (pr1 x) (a ⇒ b) + is-in-subtype (pr1 x) (a →ₘ b) complete-theory-implication lem x is-comp {a} {b} imp = apply-universal-property-trunc-Prop ( complete-theory-contains-all-formulas lem x is-comp a) - ( pr1 x (a ⇒ b)) + ( pr1 x (a →ₘ b)) ( λ { (inl a-in-logic) → ( weak-modal-logic-subset-complete-theory ( x) ( is-comp) - ( a ⇒ b) + ( a →ₘ b) ( forward-implication ( deduction-lemma ( pr1 x) @@ -344,7 +344,7 @@ module _ ( weak-modal-logic-subset-complete-theory ( x) ( is-comp) - ( a ⇒ b) + ( a →ₘ b) ( forward-implication ( deduction-lemma ( pr1 x) @@ -594,7 +594,7 @@ module _ ( pr1 (pr1 (pr1 y))) ( pr1 (pr1 (pr1 z))) ( y-leq-z)) - ( b ⇒ a) + ( b →ₘ a) ( ba-in-y)) ( b-in-z))) ; (inr z-leq-y) → @@ -757,21 +757,21 @@ module _ (l : list (formula i)) → formula i list-to-implications f nil = f - list-to-implications f (cons g l) = list-to-implications (g ⇒ f) l + list-to-implications f (cons g l) = list-to-implications (g →ₘ f) l list-to-implications-rev : formula i → (l : list (formula i)) → formula i list-to-implications-rev f nil = f - list-to-implications-rev f (cons g l) = g ⇒ list-to-implications-rev f l + list-to-implications-rev f (cons g l) = g →ₘ list-to-implications-rev f l list-to-implication-rev-snoc : (f g : formula i) (l : list (formula i)) → - list-to-implications f (snoc l g) = g ⇒ list-to-implications f l + list-to-implications f (snoc l g) = g →ₘ list-to-implications f l list-to-implication-rev-snoc f g nil = refl list-to-implication-rev-snoc f g (cons h l) = - list-to-implication-rev-snoc (h ⇒ f) g l + list-to-implication-rev-snoc (h →ₘ f) g l reverse-list-to-implications : (f : formula i) (l : list (formula i)) → @@ -779,7 +779,7 @@ module _ reverse-list-to-implications f nil = refl reverse-list-to-implications f (cons g l) = ( list-to-implication-rev-snoc f g (reverse-list l)) ∙ - ( ap (λ x → g ⇒ x) (reverse-list-to-implications f l)) + ( ap (λ x → g →ₘ x) (reverse-list-to-implications f l)) move-right : {l' : Level} (axioms : formulas l' i) @@ -798,7 +798,7 @@ module _ ( λ _ → ex-falso ∘ empty-in-list-nil)) ( f) move-right axioms f (cons c l) contains-ax-k contains-ax-s wd = - move-right axioms (c ⇒ f) l contains-ax-k contains-ax-s + move-right axioms (c →ₘ f) l contains-ax-k contains-ax-s ( forward-implication ( deduction-lemma ( union-subtype axioms (in-list l)) @@ -852,14 +852,14 @@ module _ { a = □ c} { b = □ (list-to-implications-rev a l)} ( weak-modal-logic-mp - { a = □ (c ⇒ list-to-implications-rev a l)} - { b = □ c ⇒ □ (list-to-implications-rev a l)} + { a = □ (c →ₘ list-to-implications-rev a l)} + { b = □ c →ₘ □ (list-to-implications-rev a l)} ( weak-modal-logic-monotic { ax₁ = logic} { ax₂ = union-subtype logic (pr1 x)} ( subtype-union-left logic (pr1 x)) - ( □ (c ⇒ list-to-implications-rev a l) ⇒ - □ c ⇒ □ list-to-implications-rev a l) + ( □ (c →ₘ list-to-implications-rev a l) →ₘ + □ c →ₘ □ list-to-implications-rev a l) ( weak-modal-logic-ax ( contains-ax-n _ (c , list-to-implications-rev a l , refl)))) ( in-logic)) @@ -1063,7 +1063,7 @@ module _ ( pr2 x) ( unit-trunc-Prop ( w-ax (subtype-union-right logic (pr1 x) ⊥ in-logic)))) - pr1 (canonical-model-theorem (a ⇒ b) x) in-logic fa = + pr1 (canonical-model-theorem (a →ₘ b) x) in-logic fa = pr1 ( canonical-model-theorem b x) ( weak-modal-logic-subset-complete-theory @@ -1079,7 +1079,7 @@ module _ pr2 (canonical-model-theorem (var n) x) f = map-inv-raise f pr2 (canonical-model-theorem ⊥ x) (map-raise ()) - pr2 (canonical-model-theorem (a ⇒ b) x) f = + pr2 (canonical-model-theorem (a →ₘ b) x) f = complete-theory-implication ( lem) ( pr1 x , pr1 (pr2 x)) diff --git a/src/modal-logic/formulas.lagda.md b/src/modal-logic/formulas.lagda.md index d8d41e0d68..8812a6f39e 100644 --- a/src/modal-logic/formulas.lagda.md +++ b/src/modal-logic/formulas.lagda.md @@ -36,13 +36,13 @@ module _ {l : Level} (i : Set l) where - infixr 7 _⇒_ + infixr 7 _→ₘ_ infixr 25 □_ data formula : UU l where var : type-Set i → formula ⊥ : formula - _⇒_ : formula → formula → formula + _→ₘ_ : formula → formula → formula □_ : formula → formula module _ @@ -56,16 +56,16 @@ module _ infixr 25 ◇_ ~_ : formula i → formula i - ~ a = a ⇒ ⊥ + ~ a = a →ₘ ⊥ ~~_ : formula i → formula i ~~ a = ~ ~ a -- _∨_ : formula i → formula i → formula i - -- a ∨ b = ~ a ⇒ b + -- a ∨ b = ~ a →ₘ b -- _∧_ : formula i → formula i → formula i - -- a ∧ b = ~ (a ⇒ ~ b) + -- a ∧ b = ~ (a →ₘ ~ b) ⊤ : formula i ⊤ = ~ ⊥ @@ -73,10 +73,10 @@ module _ ◇_ : formula i → formula i ◇ a = ~ □ ~ a - eq-implication-left : {a b c d : formula i} → a ⇒ b = c ⇒ d → a = c + eq-implication-left : {a b c d : formula i} → a →ₘ b = c →ₘ d → a = c eq-implication-left refl = refl - eq-implication-right : {a b c d : formula i} → a ⇒ b = c ⇒ d → b = d + eq-implication-right : {a b c d : formula i} → a →ₘ b = c →ₘ d → b = d eq-implication-right refl = refl eq-box : {a b : formula i} → □ a = □ b → a = b @@ -95,25 +95,25 @@ module _ Eq-formula : formula i → formula i → UU l Eq-formula (var n) (var m) = n = m Eq-formula (var _) ⊥ = raise-empty l - Eq-formula (var _) (_ ⇒ _) = raise-empty l + Eq-formula (var _) (_ →ₘ _) = raise-empty l Eq-formula (var -) (□ _) = raise-empty l Eq-formula ⊥ (var _) = raise-empty l Eq-formula ⊥ ⊥ = raise-unit l - Eq-formula ⊥ (_ ⇒ _) = raise-empty l + Eq-formula ⊥ (_ →ₘ _) = raise-empty l Eq-formula ⊥ (□ _) = raise-empty l - Eq-formula (_ ⇒ _) (var _) = raise-empty l - Eq-formula (_ ⇒ _) ⊥ = raise-empty l - Eq-formula (a ⇒ b) (c ⇒ d) = (Eq-formula a c) × (Eq-formula b d) - Eq-formula (_ ⇒ _) (□ _) = raise-empty l + Eq-formula (_ →ₘ _) (var _) = raise-empty l + Eq-formula (_ →ₘ _) ⊥ = raise-empty l + Eq-formula (a →ₘ b) (c →ₘ d) = (Eq-formula a c) × (Eq-formula b d) + Eq-formula (_ →ₘ _) (□ _) = raise-empty l Eq-formula (□ _) (var _) = raise-empty l Eq-formula (□ _) ⊥ = raise-empty l - Eq-formula (□ _) (_ ⇒ _) = raise-empty l + Eq-formula (□ _) (_ →ₘ _) = raise-empty l Eq-formula (□ a) (□ c) = Eq-formula a c refl-Eq-formula : (a : formula i) → Eq-formula a a refl-Eq-formula (var n) = refl refl-Eq-formula ⊥ = raise-star - refl-Eq-formula (a ⇒ b) = (refl-Eq-formula a) , (refl-Eq-formula b) + refl-Eq-formula (a →ₘ b) = (refl-Eq-formula a) , (refl-Eq-formula b) refl-Eq-formula (□ a) = refl-Eq-formula a Eq-eq-formula : {a b : formula i} → a = b → Eq-formula a b @@ -122,39 +122,39 @@ module _ eq-Eq-formula : {a b : formula i} → Eq-formula a b → a = b eq-Eq-formula {var _} {var _} refl = refl eq-Eq-formula {var _} {⊥} (map-raise ()) - eq-Eq-formula {var _} {_ ⇒ _} (map-raise ()) + eq-Eq-formula {var _} {_ →ₘ _} (map-raise ()) eq-Eq-formula {var _} {□ _} (map-raise ()) eq-Eq-formula {⊥} {var _} (map-raise ()) eq-Eq-formula {⊥} {⊥} _ = refl - eq-Eq-formula {⊥} {_ ⇒ _} (map-raise ()) + eq-Eq-formula {⊥} {_ →ₘ _} (map-raise ()) eq-Eq-formula {⊥} {□ _} (map-raise ()) - eq-Eq-formula {_ ⇒ _} {var _} (map-raise ()) - eq-Eq-formula {_ ⇒ _} {⊥} (map-raise ()) - eq-Eq-formula {a ⇒ b} {c ⇒ d} (eq1 , eq2) = - ap (λ (x , y) → x ⇒ y) (eq-pair (eq-Eq-formula eq1) (eq-Eq-formula eq2)) - eq-Eq-formula {_ ⇒ _} {□ _} (map-raise ()) + eq-Eq-formula {_ →ₘ _} {var _} (map-raise ()) + eq-Eq-formula {_ →ₘ _} {⊥} (map-raise ()) + eq-Eq-formula {a →ₘ b} {c →ₘ d} (eq1 , eq2) = + ap (λ (x , y) → x →ₘ y) (eq-pair (eq-Eq-formula eq1) (eq-Eq-formula eq2)) + eq-Eq-formula {_ →ₘ _} {□ _} (map-raise ()) eq-Eq-formula {□ _} {var _} (map-raise ()) eq-Eq-formula {□ _} {⊥} (map-raise ()) - eq-Eq-formula {□ _} {_ ⇒ _} (map-raise ()) + eq-Eq-formula {□ _} {_ →ₘ _} (map-raise ()) eq-Eq-formula {□ _} {□ _} eq = ap □_ (eq-Eq-formula eq) is-prop-Eq-formula : (a b : formula i) → is-prop (Eq-formula a b) is-prop-Eq-formula (var n) (var m) = is-prop-type-Prop (Id-Prop i n m) is-prop-Eq-formula (var _) ⊥ = is-prop-raise-empty - is-prop-Eq-formula (var _) (_ ⇒ _) = is-prop-raise-empty + is-prop-Eq-formula (var _) (_ →ₘ _) = is-prop-raise-empty is-prop-Eq-formula (var -) (□ _) = is-prop-raise-empty is-prop-Eq-formula ⊥ (var _) = is-prop-raise-empty is-prop-Eq-formula ⊥ ⊥ = is-prop-raise-unit - is-prop-Eq-formula ⊥ (_ ⇒ _) = is-prop-raise-empty + is-prop-Eq-formula ⊥ (_ →ₘ _) = is-prop-raise-empty is-prop-Eq-formula ⊥ (□ _) = is-prop-raise-empty - is-prop-Eq-formula (_ ⇒ _) (var _) = is-prop-raise-empty - is-prop-Eq-formula (_ ⇒ _) ⊥ = is-prop-raise-empty - is-prop-Eq-formula (a ⇒ b) (c ⇒ d) = - is-prop-prod (is-prop-Eq-formula a c) (is-prop-Eq-formula b d) - is-prop-Eq-formula (_ ⇒ _) (□ _) = is-prop-raise-empty + is-prop-Eq-formula (_ →ₘ _) (var _) = is-prop-raise-empty + is-prop-Eq-formula (_ →ₘ _) ⊥ = is-prop-raise-empty + is-prop-Eq-formula (a →ₘ b) (c →ₘ d) = + is-prop-product (is-prop-Eq-formula a c) (is-prop-Eq-formula b d) + is-prop-Eq-formula (_ →ₘ _) (□ _) = is-prop-raise-empty is-prop-Eq-formula (□ _) (var _) = is-prop-raise-empty is-prop-Eq-formula (□ _) ⊥ = is-prop-raise-empty - is-prop-Eq-formula (□ _) (_ ⇒ _) = is-prop-raise-empty + is-prop-Eq-formula (□ _) (_ →ₘ _) = is-prop-raise-empty is-prop-Eq-formula (□ a) (□ c) = is-prop-Eq-formula a c is-set-formula : is-set (formula i) diff --git a/src/modal-logic/kripke-semantics.lagda.md b/src/modal-logic/kripke-semantics.lagda.md index e341d838d7..4ffde22489 100644 --- a/src/modal-logic/kripke-semantics.lagda.md +++ b/src/modal-logic/kripke-semantics.lagda.md @@ -115,7 +115,7 @@ module _ _⊨_ : kripke-model w l2 i l4 × w → formula i → Prop l (M , x) ⊨ var n = raise-Prop l (model-valuate i M n x) (M , x) ⊨ ⊥ = raise-empty-Prop l - (M , x) ⊨ a ⇒ b = implication-Prop ((M , x) ⊨ a) ((M , x) ⊨ b) + (M , x) ⊨ a →ₘ b = implication-Prop ((M , x) ⊨ a) ((M , x) ⊨ b) (M , x) ⊨ □ a = Π-Prop ( w) @@ -178,9 +178,9 @@ module _ finite-models : model-class w l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4) finite-models M = - prod-Prop + product-Prop ( is-finite-Prop w) - ( prod-Prop + ( product-Prop ( Π-Prop ( w) ( λ x → @@ -198,7 +198,7 @@ module _ is-decidable-raise (l1 ⊔ l2 ⊔ l4) _ (dec-v x n) finite-models-subclass-decidable-models M _ ⊥ x = inr map-inv-raise - finite-models-subclass-decidable-models M is-fin (a ⇒ b) x = + finite-models-subclass-decidable-models M is-fin (a →ₘ b) x = is-decidable-function-type ( finite-models-subclass-decidable-models M is-fin a x) ( finite-models-subclass-decidable-models M is-fin b x) diff --git a/src/modal-logic/logic-syntax.lagda.md b/src/modal-logic/logic-syntax.lagda.md index e8f90fea29..1f34f40ca3 100644 --- a/src/modal-logic/logic-syntax.lagda.md +++ b/src/modal-logic/logic-syntax.lagda.md @@ -48,7 +48,7 @@ module _ data _⊢_ (axioms : formulas l2 i) : formula i → UU (l1 ⊔ l2) where ax : {a : formula i} → is-in-subtype axioms a → axioms ⊢ a - mp : {a b : formula i} → axioms ⊢ a ⇒ b → axioms ⊢ a → axioms ⊢ b + mp : {a b : formula i} → axioms ⊢ a →ₘ b → axioms ⊢ a → axioms ⊢ b nec : {a : formula i} → axioms ⊢ a → axioms ⊢ □ a modal-logic : formulas l2 i → formulas (l1 ⊔ l2) i @@ -78,7 +78,7 @@ module _ modal-logic-mp : {a b : formula i} → - is-in-subtype (modal-logic axioms) (a ⇒ b) → + is-in-subtype (modal-logic axioms) (a →ₘ b) → is-in-subtype (modal-logic axioms) a → is-in-subtype (modal-logic axioms) b modal-logic-mp {a} {b} tdab tda = diff --git a/src/modal-logic/weak-deduction.lagda.md b/src/modal-logic/weak-deduction.lagda.md index 4bb40a8de0..34307928bd 100644 --- a/src/modal-logic/weak-deduction.lagda.md +++ b/src/modal-logic/weak-deduction.lagda.md @@ -53,7 +53,7 @@ module _ data _▷_ (axioms : formulas l2 i) : formula i → UU (l1 ⊔ l2) where w-ax : {a : formula i} → is-in-subtype axioms a → axioms ▷ a - w-mp : {a b : formula i} → axioms ▷ a ⇒ b → axioms ▷ a → axioms ▷ b + w-mp : {a b : formula i} → axioms ▷ a →ₘ b → axioms ▷ a → axioms ▷ b weak-modal-logic : formulas l2 i → formulas (l1 ⊔ l2) i weak-modal-logic axioms a = trunc-Prop (axioms ▷ a) @@ -70,7 +70,7 @@ module _ weak-modal-logic-mp : {a b : formula i} → - is-in-subtype (weak-modal-logic axioms) (a ⇒ b) → + is-in-subtype (weak-modal-logic axioms) (a →ₘ b) → is-in-subtype (weak-modal-logic axioms) a → is-in-subtype (weak-modal-logic axioms) b weak-modal-logic-mp {a} {b} twdab twda = @@ -266,7 +266,7 @@ empty-in-list-nil = map-universal-property-trunc-Prop empty-Prop ( λ ()) equiv-subset-head-tail : {l1 l2 : Level} {A : UU l1} (x : A) (xs : list A) (a : subtype l2 A) → (leq-prop-subtype (in-list (cons x xs)) a) ⇔ - prod-Prop (a x) (leq-prop-subtype (in-list xs) a) + product-Prop (a x) (leq-prop-subtype (in-list xs) a) pr1 (equiv-subset-head-tail x xs a) sub = pair ( sub x (unit-trunc-Prop (is-head x xs))) @@ -315,7 +315,7 @@ lists-in-union-lists (cons x xs) a b sub = ( sub x (unit-trunc-Prop (is-head x xs))) ( ∃-Prop _ (λ _ → Σ _ _)) ( λ (xsl , xsr , sub-lists , sub-xsl , sub-xsr) → - ( ind-coprod + ( ind-coproduct ( λ _ → _) ( λ x-in-a → ( unit-trunc-Prop @@ -495,7 +495,7 @@ module _ backward-deduction-lemma : {a b : formula i} → - axioms ▷ a ⇒ b → + axioms ▷ a →ₘ b → is-in-subtype (weak-modal-logic (theory-add-formula a axioms)) b backward-deduction-lemma {a} wab = unit-trunc-Prop @@ -511,29 +511,29 @@ module _ where deduction-a->a : - (a : formula i) → is-in-subtype (weak-modal-logic axioms) (a ⇒ a) + (a : formula i) → is-in-subtype (weak-modal-logic axioms) (a →ₘ a) deduction-a->a a = apply-three-times-universal-property-trunc-Prop - ( contains-ax-s _ (a , a ⇒ a , a , refl)) - ( contains-ax-k _ (a , a ⇒ a , refl)) + ( contains-ax-s _ (a , a →ₘ a , a , refl)) + ( contains-ax-k _ (a , a →ₘ a , refl)) ( contains-ax-k _ (a , a , refl)) - ( (weak-modal-logic axioms) (a ⇒ a)) + ( (weak-modal-logic axioms) (a →ₘ a)) ( λ x y z → unit-trunc-Prop (w-mp (w-mp x y) z)) forward-deduction-lemma : (a : formula i) {b : formula i} → theory-add-formula a axioms ▷ b → - is-in-subtype (weak-modal-logic axioms) (a ⇒ b) + is-in-subtype (weak-modal-logic axioms) (a →ₘ b) forward-deduction-lemma a {b} (w-ax x) = - elim-disj-Prop + elim-disjunction-Prop ( Id-formula-Prop a b) ( axioms b) - ( (weak-modal-logic axioms) (a ⇒ b)) + ( (weak-modal-logic axioms) (a →ₘ b)) ( (λ { refl → deduction-a->a a}) , (λ in-axioms → ( apply-universal-property-trunc-Prop ( contains-ax-k _ (b , a , refl)) - ( (weak-modal-logic axioms) (a ⇒ b)) + ( (weak-modal-logic axioms) (a →ₘ b)) ( λ x → unit-trunc-Prop (w-mp x (w-ax in-axioms)))))) ( x) forward-deduction-lemma a {b} (w-mp {c} twdcb twdc) = @@ -541,17 +541,17 @@ module _ ( forward-deduction-lemma a twdcb) ( forward-deduction-lemma a twdc) ( contains-ax-s _ (a , c , b , refl)) - ( (weak-modal-logic axioms) (a ⇒ b)) + ( (weak-modal-logic axioms) (a →ₘ b)) ( λ wdacb wdac x → ( unit-trunc-Prop (w-mp (w-mp x wdacb) wdac))) deduction-lemma : (a b : formula i) → weak-modal-logic (theory-add-formula a axioms) b ⇔ - weak-modal-logic axioms (a ⇒ b) + weak-modal-logic axioms (a →ₘ b) pr1 (deduction-lemma a b) = map-universal-property-trunc-Prop - ( (weak-modal-logic axioms) (a ⇒ b)) + ( (weak-modal-logic axioms) (a →ₘ b)) ( forward-deduction-lemma a) pr2 (deduction-lemma a b) = map-universal-property-trunc-Prop @@ -567,24 +567,24 @@ module _ deduction-ex-falso : (a b : formula i) → - is-in-subtype (weak-modal-logic axioms) (~ a ⇒ a ⇒ b) + is-in-subtype (weak-modal-logic axioms) (~ a →ₘ a →ₘ b) deduction-ex-falso a b = forward-implication - ( deduction-lemma axioms contains-ax-k contains-ax-s (~ a) (a ⇒ b)) + ( deduction-lemma axioms contains-ax-k contains-ax-s (~ a) (a →ₘ b)) ( forward-implication ( deduction-lemma ( theory-add-formula (~ a) axioms) ( transitive-leq-subtype ( ax-k i) ( weak-modal-logic axioms) - ( weak-modal-logic (theory-add-formula (a ⇒ ⊥) axioms)) + ( weak-modal-logic (theory-add-formula (a →ₘ ⊥) axioms)) ( weak-modal-logic-monotic ( subtype-union-right (Id-formula-Prop (~ a)) axioms)) ( contains-ax-k)) ( transitive-leq-subtype ( ax-s i) ( weak-modal-logic axioms) - ( weak-modal-logic (theory-add-formula (a ⇒ ⊥) axioms)) + ( weak-modal-logic (theory-add-formula (a →ₘ ⊥) axioms)) ( weak-modal-logic-monotic ( subtype-union-right (Id-formula-Prop (~ a)) axioms)) ( contains-ax-s)) From 77f1901c243604f573b8b862a4ed6362e82b7d21 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Wed, 14 Feb 2024 21:20:07 +0300 Subject: [PATCH 32/63] fix pre-commit --- src/modal-logic/canonical-model-theorem.lagda.md | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/src/modal-logic/canonical-model-theorem.lagda.md b/src/modal-logic/canonical-model-theorem.lagda.md index d2b3f522dd..e256347e34 100644 --- a/src/modal-logic/canonical-model-theorem.lagda.md +++ b/src/modal-logic/canonical-model-theorem.lagda.md @@ -216,7 +216,8 @@ module _ (x : L-consistent-theory) → is-L-complete-theory x → (a : formula i) → type-disjunction-Prop (pr1 x a) (pr1 x (~ a)) complete-theory-contains-all-formulas lem x is-comp a with lem ((pr1 x) a) - ... | inl a-in-logic = inl-disjunction-Prop ((pr1 x) a) ((pr1 x) (~ a)) a-in-logic + ... | inl a-in-logic = + inl-disjunction-Prop ((pr1 x) a) ((pr1 x) (~ a)) a-in-logic ... | inr a-not-in-logic = unit-trunc-Prop ( inr From f31d02dcc19fff72abc483d0bafa67d6d233c88d Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Tue, 26 Mar 2024 01:17:24 +0300 Subject: [PATCH 33/63] Small refactor and add canonical axioms --- src/foundation/unions-subtypes.lagda.md | 4 + src/modal-logic.lagda.md | 1 - src/modal-logic/axioms.lagda.md | 35 +++++-- src/modal-logic/completeness-K.lagda.md | 8 +- src/modal-logic/completeness.lagda.md | 12 +-- src/modal-logic/formulas.lagda.md | 3 + src/modal-logic/kripke-semantics.lagda.md | 100 ++++++++++-------- src/modal-logic/modal-logic-IK.lagda.md | 102 ------------------ src/modal-logic/modal-logic-K.lagda.md | 121 +++++++++++++++------- src/modal-logic/soundness.lagda.md | 47 ++++----- 10 files changed, 200 insertions(+), 233 deletions(-) delete mode 100644 src/modal-logic/modal-logic-IK.lagda.md diff --git a/src/foundation/unions-subtypes.lagda.md b/src/foundation/unions-subtypes.lagda.md index ce86a03ba0..c078ad7c3e 100644 --- a/src/foundation/unions-subtypes.lagda.md +++ b/src/foundation/unions-subtypes.lagda.md @@ -40,6 +40,10 @@ module _ union-subtype : subtype l1 X → subtype l2 X → subtype (l1 ⊔ l2) X union-subtype P Q x = (P x) ∨ (Q x) + union-subtype P Q x = disjunction-Prop (P x) (Q x) + + infixl 5 _∪_ + _∪_ = union-subtype ``` ### Unions of decidable subtypes diff --git a/src/modal-logic.lagda.md b/src/modal-logic.lagda.md index d6859a669d..3a3f9e7e88 100644 --- a/src/modal-logic.lagda.md +++ b/src/modal-logic.lagda.md @@ -10,7 +10,6 @@ open import modal-logic.completeness-K public open import modal-logic.formulas public open import modal-logic.kripke-semantics public open import modal-logic.logic-syntax public -open import modal-logic.modal-logic-IK public open import modal-logic.modal-logic-K public open import modal-logic.soundness public open import modal-logic.weak-deduction public diff --git a/src/modal-logic/axioms.lagda.md b/src/modal-logic/axioms.lagda.md index d7f94b33a9..461baaaedf 100644 --- a/src/modal-logic/axioms.lagda.md +++ b/src/modal-logic/axioms.lagda.md @@ -113,34 +113,49 @@ module _ ( eq-implication-right ∘ eq-box ∘ eq-implication-left) ax-dn : formulas l i - ax-dn = ax-1-parameter ( λ a → ~~ a →ₘ a) ( eq-implication-right) + ax-dn = ax-1-parameter ( λ a → ~~ a →ₘ a) eq-implication-right + + ax-m : formulas l i + ax-m = ax-1-parameter (λ a → □ a →ₘ a) eq-implication-right + + ax-b : formulas l i + ax-b = ax-1-parameter (λ a → a →ₘ □ ◇ a) eq-implication-left + + ax-d : formulas l i + ax-d = ax-1-parameter (λ a → □ a →ₘ ◇ a) (eq-box ∘ eq-implication-left) + + ax-4 : formulas l i + ax-4 = ax-1-parameter (λ a → □ a →ₘ □ □ a) (eq-box ∘ eq-implication-left) + + ax-5 : formulas l i + ax-5 = + ax-1-parameter ( λ a → ◇ a →ₘ □ ◇ a) ( eq-diamond ∘ eq-implication-left) module _ {l1 l2 : Level} (i : Set l1) - (w : UU l2) (l3 l4 : Level) where - ax-k-soundness : soundness (ax-k i) (all-models w l3 i l4) - ax-k-soundness (a →ₘ b →ₘ a) (_ , _ , refl) M _ x fa _ = fa + ax-k-soundness : soundness (ax-k i) (all-models l2 l3 i l4) + ax-k-soundness (a →ₘ b →ₘ a) (_ , _ , refl) w M _ x fa _ = fa - ax-n-soundness : soundness (ax-n i) (all-models w l3 i l4) + ax-n-soundness : soundness (ax-n i) (all-models l2 l3 i l4) ax-n-soundness (□ (a →ₘ b) →ₘ □ a →ₘ □ b) (_ , _ , refl) - M in-class x fab fa y r = + w M in-class x fab fa y r = fab y r (fa y r) - ax-s-soundness : soundness (ax-s i) (all-models w l3 i l4) + ax-s-soundness : soundness (ax-s i) (all-models l2 l3 i l4) ax-s-soundness ((a →ₘ b →ₘ c) →ₘ (a →ₘ b) →ₘ a →ₘ c) (_ , _ , _ , refl) - M in-class x fabc fab fa = + w M in-class x fabc fab fa = fabc fa (fab fa) - ax-dn-soundness : soundness (ax-dn i) (decidable-models w l3 i l4) - ax-dn-soundness (((a →ₘ ⊥) →ₘ ⊥) →ₘ a) (_ , refl) _ is-dec x f + ax-dn-soundness : soundness (ax-dn i) (decidable-models l2 l3 i l4) + ax-dn-soundness .(~~ a →ₘ a) (a , refl) w M is-dec x f with (is-dec a x) ... | inl fa = fa ... | inr fna = raise-ex-falso _ (f (λ fa -> map-raise (fna fa))) diff --git a/src/modal-logic/completeness-K.lagda.md b/src/modal-logic/completeness-K.lagda.md index ec1100e427..5e25b064f2 100644 --- a/src/modal-logic/completeness-K.lagda.md +++ b/src/modal-logic/completeness-K.lagda.md @@ -38,7 +38,6 @@ open import modal-logic.completeness open import modal-logic.formulas open import modal-logic.kripke-semantics open import modal-logic.logic-syntax -open import modal-logic.modal-logic-IK open import modal-logic.modal-logic-K open import modal-logic.soundness open import modal-logic.weak-deduction @@ -64,15 +63,14 @@ module _ (prop-resize : propositional-resizing l1 (lsuc l1)) where - completeness-K : - completeness (modal-logic-K i) (lsuc l1) (λ w → all-models w l1 i l1) - completeness-K a in-kripke-logic with lem ((modal-logic-K i) a) + completeness-K : completeness (modal-logic-K i) (all-models (lsuc l1) l1 i l1) + completeness-K a in-kripke-logic with lem (modal-logic-K i a) ... | inl in-logic = in-logic ... | inr not-in-logic = ex-falso ( forward-implication ( canonical-model-theorem' - ( union-subtype (modal-logic-IK i) (ax-dn i)) + ( modal-logic-K-axioms i) ( zorn) ( prop-resize) ( is-consistent-K i) diff --git a/src/modal-logic/completeness.lagda.md b/src/modal-logic/completeness.lagda.md index 6af35273d9..ec5370684e 100644 --- a/src/modal-logic/completeness.lagda.md +++ b/src/modal-logic/completeness.lagda.md @@ -44,17 +44,13 @@ TODO -- completeness = class-modal-logic C ⊆ logic module _ - {l1 l2 l3 l4 l5 : Level} + {l1 l2 l3 l4 l5 l6 : Level} {i : Set l1} (logic : formulas l2 i) - (l6 : Level) (C : (w : UU l6) → model-class w l3 i l4 l5) + (C : model-class l3 l4 i l5 l6) where - completeness : UU (l1 ⊔ l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ l5 ⊔ lsuc l6) - -- completeness = (w : UU l6) → class-modal-logic (C w) ⊆ logic - completeness = - (a : formula i) → - ( (w : UU l6) → is-in-subtype (class-modal-logic (C w)) a) → - is-in-subtype logic a + completeness : UU (l1 ⊔ l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6) + completeness = class-modal-logic C ⊆ logic ``` ## Properties diff --git a/src/modal-logic/formulas.lagda.md b/src/modal-logic/formulas.lagda.md index 8812a6f39e..d1172e441d 100644 --- a/src/modal-logic/formulas.lagda.md +++ b/src/modal-logic/formulas.lagda.md @@ -81,6 +81,9 @@ module _ eq-box : {a b : formula i} → □ a = □ b → a = b eq-box refl = refl + + eq-diamond : {a b : formula i} → ◇ a = ◇ b → a = b + eq-diamond refl = refl ``` ## Properties diff --git a/src/modal-logic/kripke-semantics.lagda.md b/src/modal-logic/kripke-semantics.lagda.md index 4ffde22489..e7656274e4 100644 --- a/src/modal-logic/kripke-semantics.lagda.md +++ b/src/modal-logic/kripke-semantics.lagda.md @@ -73,8 +73,24 @@ module _ kripke-model : UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4) kripke-model = kripke-frame w l2 × (type-Set i → w → Prop l4) - model-class : (l5 : Level) → UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ lsuc l5) - model-class l5 = subtype l5 kripke-model +module _ + (l1 l2 : Level) + {l3 : Level} (i : Set l3) + (l4 : Level) + where + + model-with-world : UU (lsuc l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4) + model-with-world = Σ (UU l1) (λ w → kripke-model w l2 i l4) + + world-model-with-world : model-with-world → UU l1 + world-model-with-world = pr1 + + model-with-world-model : + (M : model-with-world) → kripke-model (world-model-with-world M) l2 i l4 + model-with-world-model = pr2 + + model-class : (l5 : Level) → UU (lsuc l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ lsuc l5) + model-class l5 = subtype l5 model-with-world all-models : model-class lzero all-models _ = unit-Prop @@ -90,8 +106,7 @@ module _ model-frame : kripke-model w l2 i l4 → kripke-frame w l2 model-frame = pr1 - model-valuate : - kripke-model w l2 i l4 → type-Set i → w → Prop l4 + model-valuate : kripke-model w l2 i l4 → type-Set i → w → Prop l4 model-valuate = pr2 model-relation-Prop : kripke-model w l2 i l4 → Relation-Prop l2 w @@ -130,62 +145,63 @@ module _ _⊭M_ : kripke-model w l2 i l4 → formula i → Prop l M ⊭M a = neg-Prop (M ⊨M a) - decidable-class : model-class w l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4) - decidable-class M = Π-Prop (formula i) (λ a → M ⊨M a) - module _ - {l1 l2 l3 l4 : Level} {w : UU l1} {i : Set l3} + {l1 l2 l3 l4 : Level} {i : Set l3} where _⊨C_ : {l5 : Level} → - model-class w l2 i l4 l5 → + model-class l1 l2 i l4 l5 → formula i → - Prop (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ l5) + Prop (lsuc l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ l5) C ⊨C a = Π-Prop - ( kripke-model w l2 i l4) - ( λ M → function-Prop (is-in-subtype C M) (M ⊨M a)) + ( UU l1) + ( λ w → + ( Π-Prop + ( kripke-model w l2 i l4)) + ( λ M → function-Prop (is-in-subtype C (w , M)) (M ⊨M a))) class-modal-logic : {l5 : Level} → - model-class w l2 i l4 l5 → - formulas (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ l5) i + model-class l1 l2 i l4 l5 → + formulas (lsuc l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ l5) i class-modal-logic C a = C ⊨C a -- TODO: rename class-modal-logic-monotic : {l5 l6 : Level} - (C₁ : model-class w l2 i l4 l5) - (C₂ : model-class w l2 i l4 l6) → + (C₁ : model-class l1 l2 i l4 l5) + (C₂ : model-class l1 l2 i l4 l6) → C₁ ⊆ C₂ → class-modal-logic C₂ ⊆ class-modal-logic C₁ - class-modal-logic-monotic C₁ C₂ sub _ in-modal-logic-C₂ M in-C₁ = - in-modal-logic-C₂ M (sub M in-C₁) + class-modal-logic-monotic C₁ C₂ sub _ in-modal-logic-C₂ w M in-C₁ = + in-modal-logic-C₂ w M (sub (w , M) in-C₁) module _ - {l1 : Level} (w : UU l1) - (l2 : Level) + (l1 l2 : Level) {l3 : Level} (i : Set l3) (l4 : Level) where - decidable-models : model-class w l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4) - decidable-models M = + decidable-models : model-class l1 l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4) + decidable-models (w , M) = Π-Prop ( formula i) - ( λ a → (Π-Prop w (λ x → is-decidable-Prop ((M , x) ⊨ a)))) + (λ a → Π-Prop w (λ x → is-decidable-Prop ((M , x) ⊨ a))) + -- (λ a → is-decidable-Prop (M ⊨M a)) - finite-models : model-class w l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4) - finite-models M = + finite-models : model-class l1 l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4) + finite-models (w , M) = product-Prop ( is-finite-Prop w) ( product-Prop ( Π-Prop ( w) ( λ x → - ( Π-Prop ( w) - ( λ y → is-decidable-Prop (model-relation-Prop i M x y))))) + ( Π-Prop + ( w) + ( λ y → is-decidable-Prop (model-relation-Prop i M x y))))) ( Π-Prop ( w) ( λ x → @@ -193,21 +209,21 @@ module _ ( type-Set i) ( λ n → is-decidable-Prop (model-valuate i M n x)))))) - finite-models-subclass-decidable-models : finite-models ⊆ decidable-models - finite-models-subclass-decidable-models M (_ , _ , dec-v) (var n) x = - is-decidable-raise (l1 ⊔ l2 ⊔ l4) _ (dec-v x n) - finite-models-subclass-decidable-models M _ ⊥ x = - inr map-inv-raise - finite-models-subclass-decidable-models M is-fin (a →ₘ b) x = - is-decidable-function-type - ( finite-models-subclass-decidable-models M is-fin a x) - ( finite-models-subclass-decidable-models M is-fin b x) - finite-models-subclass-decidable-models - M is-fin@(w-is-fin , dec-r , dec-v) (□ a) x = + finite-models-subclass-decidable-models : + finite-models ⊆ decidable-models + finite-models-subclass-decidable-models (w , M) (w-is-fin , dec-r , dec-v) = + lemma + where + lemma : + (a : formula i) (x : w) → is-decidable (type-Prop ((M , x) ⊨ a)) + lemma (var n) x = + is-decidable-raise (l1 ⊔ l2 ⊔ l4) _ (dec-v x n) + lemma ⊥ x = + inr map-inv-raise + lemma (a →ₘ b) x = + is-decidable-function-type (lemma a x) (lemma b x) + lemma (□ a) x = is-decidable-Π-is-finite ( w-is-fin) - ( λ y → - ( is-decidable-function-type - ( dec-r x y) - ( finite-models-subclass-decidable-models M is-fin a y))) + ( λ y → is-decidable-function-type (dec-r x y) (lemma a y)) ``` diff --git a/src/modal-logic/modal-logic-IK.lagda.md b/src/modal-logic/modal-logic-IK.lagda.md deleted file mode 100644 index 32c2781baf..0000000000 --- a/src/modal-logic/modal-logic-IK.lagda.md +++ /dev/null @@ -1,102 +0,0 @@ -# Modal logic IK - -```agda -module modal-logic.modal-logic-IK where -``` - -
Imports - -```agda -open import foundation.inhabited-types -open import foundation.intersections-subtypes -open import foundation.sets -open import foundation.subtypes -open import foundation.unions-subtypes -open import foundation.universe-levels - -open import modal-logic.axioms -open import modal-logic.formulas -open import modal-logic.kripke-semantics -open import modal-logic.logic-syntax -open import modal-logic.soundness -``` - -
- -## Idea - -TODO - -## Definition - -```agda -module _ - {l : Level} (i : Set l) - where - --- TODO: It's not Intuitionistic K - modal-logic-IK : formulas l i - modal-logic-IK = - modal-logic (union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i))) - - IK-contains-ax-k : ax-k i ⊆ modal-logic-IK - IK-contains-ax-k = - transitive-leq-subtype - ( ax-k i) - ( union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i))) - ( modal-logic (union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i)))) - ( axioms-subset-modal-logic - ( union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i)))) - ( subtype-union-left (ax-k i) (union-subtype (ax-s i) (ax-n i))) - - IK-contains-ax-s : ax-s i ⊆ modal-logic-IK - IK-contains-ax-s = - transitive-leq-subtype - ( ax-s i) - ( union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i))) - ( modal-logic (union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i)))) - ( axioms-subset-modal-logic - ( union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i)))) - ( transitive-leq-subtype - ( ax-s i) - ( union-subtype (ax-s i) (ax-n i)) - ( union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i))) - ( subtype-union-right (ax-k i) (union-subtype (ax-s i) (ax-n i))) - ( subtype-union-left (ax-s i) (ax-n i))) - - IK-contains-ax-n : ax-n i ⊆ modal-logic-IK - IK-contains-ax-n = - transitive-leq-subtype - ( ax-n i) - ( union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i))) - ( modal-logic (union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i)))) - ( axioms-subset-modal-logic - ( union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i)))) - ( transitive-leq-subtype - ( ax-n i) - ( union-subtype (ax-s i) (ax-n i)) - ( union-subtype (ax-k i) (union-subtype (ax-s i) (ax-n i))) - ( subtype-union-right (ax-k i) (union-subtype (ax-s i) (ax-n i))) - ( subtype-union-right (ax-s i) (ax-n i))) - -module _ - {l1 l2 : Level} - (i : Set l1) - (w : UU l2) - (l3 l4 : Level) - where - - soundness-IK : soundness (modal-logic-IK i) (all-models w l3 i l4) - soundness-IK = - soundness-modal-logic-union-same-class - ( ax-k i) - ( union-subtype (ax-s i) (ax-n i)) - ( all-models w l3 i l4) - ( ax-k-soundness i w l3 l4) - ( soundness-union-same-class - ( ax-s i) - ( ax-n i) - ( all-models w l3 i l4) - ( ax-s-soundness i w l3 l4) - ( ax-n-soundness i w l3 l4)) -``` diff --git a/src/modal-logic/modal-logic-K.lagda.md b/src/modal-logic/modal-logic-K.lagda.md index dc3a2fb171..ffec7968c3 100644 --- a/src/modal-logic/modal-logic-K.lagda.md +++ b/src/modal-logic/modal-logic-K.lagda.md @@ -27,7 +27,6 @@ open import modal-logic.axioms open import modal-logic.formulas open import modal-logic.kripke-semantics open import modal-logic.logic-syntax -open import modal-logic.modal-logic-IK open import modal-logic.soundness open import univalent-combinatorics.finite-types @@ -46,80 +45,122 @@ module _ {l : Level} (i : Set l) where + modal-logic-K-axioms : formulas l i + modal-logic-K-axioms = ax-k i ∪ ax-s i ∪ ax-n i ∪ ax-dn i + modal-logic-K : formulas l i - modal-logic-K = - modal-logic (union-subtype (modal-logic-IK i) (ax-dn i)) + modal-logic-K = modal-logic modal-logic-K-axioms - IK-subset-K : modal-logic-IK i ⊆ modal-logic-K - IK-subset-K = + K-axioms-contains-ax-k : ax-k i ⊆ modal-logic-K-axioms + K-axioms-contains-ax-k = transitive-leq-subtype - ( modal-logic-IK i) - ( union-subtype (modal-logic-IK i) (ax-dn i)) - ( modal-logic-K) - ( axioms-subset-modal-logic _) - ( subtype-union-left (modal-logic-IK i) (ax-dn i)) + ( ax-k i) + ( ax-k i ∪ ax-s i) + ( modal-logic-K-axioms) + ( transitive-leq-subtype + ( ax-k i ∪ ax-s i) + ( ax-k i ∪ ax-s i ∪ ax-n i) + ( modal-logic-K-axioms) + ( subtype-union-left (ax-k i ∪ ax-s i ∪ ax-n i) (ax-dn i)) + ( subtype-union-left (ax-k i ∪ ax-s i) (ax-n i))) + ( subtype-union-left (ax-k i) (ax-s i)) + + K-axioms-contains-ax-s : ax-s i ⊆ modal-logic-K-axioms + K-axioms-contains-ax-s = + transitive-leq-subtype + ( ax-s i) + ( ax-k i ∪ ax-s i) + ( modal-logic-K-axioms) + ( transitive-leq-subtype + ( ax-k i ∪ ax-s i) + ( ax-k i ∪ ax-s i ∪ ax-n i) + ( modal-logic-K-axioms) + ( subtype-union-left (ax-k i ∪ ax-s i ∪ ax-n i) (ax-dn i)) + ( subtype-union-left (ax-k i ∪ ax-s i) (ax-n i))) + ( subtype-union-right (ax-k i) (ax-s i)) + + K-axioms-contains-ax-n : ax-n i ⊆ modal-logic-K-axioms + K-axioms-contains-ax-n = + transitive-leq-subtype + ( ax-n i) + ( ax-k i ∪ ax-s i ∪ ax-n i) + ( modal-logic-K-axioms) + ( subtype-union-left (ax-k i ∪ ax-s i ∪ ax-n i) (ax-dn i)) + ( subtype-union-right (ax-k i ∪ ax-s i) (ax-n i)) + + K-axioms-contains-ax-dn : ax-dn i ⊆ modal-logic-K-axioms + K-axioms-contains-ax-dn = + subtype-union-right (ax-k i ∪ ax-s i ∪ ax-n i) (ax-dn i) K-contains-ax-k : ax-k i ⊆ modal-logic-K K-contains-ax-k = transitive-leq-subtype ( ax-k i) - ( modal-logic-IK i) + ( modal-logic-K-axioms) ( modal-logic-K) - ( IK-subset-K) - ( IK-contains-ax-k i) + ( axioms-subset-modal-logic modal-logic-K-axioms) + ( K-axioms-contains-ax-k) K-contains-ax-s : ax-s i ⊆ modal-logic-K K-contains-ax-s = transitive-leq-subtype ( ax-s i) - ( modal-logic-IK i) + ( modal-logic-K-axioms) ( modal-logic-K) - ( IK-subset-K) - ( IK-contains-ax-s i) + ( axioms-subset-modal-logic modal-logic-K-axioms) + ( K-axioms-contains-ax-s) K-contains-ax-n : ax-n i ⊆ modal-logic-K K-contains-ax-n = transitive-leq-subtype ( ax-n i) - ( modal-logic-IK i) + ( modal-logic-K-axioms) ( modal-logic-K) - ( IK-subset-K) - ( IK-contains-ax-n i) + ( axioms-subset-modal-logic modal-logic-K-axioms) + ( K-axioms-contains-ax-n) K-contains-ax-dn : ax-dn i ⊆ modal-logic-K K-contains-ax-dn = transitive-leq-subtype ( ax-dn i) - ( union-subtype (modal-logic-IK i) (ax-dn i)) - ( modal-logic (union-subtype (modal-logic-IK i) (ax-dn i))) - ( axioms-subset-modal-logic - ( union-subtype (modal-logic-IK i) (ax-dn i))) - ( subtype-union-right (modal-logic-IK i) (ax-dn i)) + ( modal-logic-K-axioms) + ( modal-logic-K) + ( axioms-subset-modal-logic modal-logic-K-axioms) + ( K-axioms-contains-ax-dn) module _ - {l1 l2 l3 l4 : Level} - (i : Set l1) (w : UU l2) + {l1 l2 l3 l4 : Level} (i : Set l1) where - soundness-K : soundness (modal-logic-K i) (decidable-models w l3 i l4) + soundness-K : soundness (modal-logic-K i) (decidable-models l2 l3 i l4) soundness-K = soundness-modal-logic-union-subclass-right-sublevels - ( modal-logic-IK i) + ( ax-k i ∪ ax-s i ∪ ax-n i) ( ax-dn i) ( l1 ⊔ l2 ⊔ l3 ⊔ l4) - ( all-models w l3 i l4) - ( decidable-models w l3 i l4) - ( soundness-IK i w l3 l4) - ( ax-dn-soundness i w l3 l4) - ( all-models-is-largest-class w l3 i l4 (decidable-models w l3 i l4)) - - soundness-K-finite : soundness (modal-logic-K i) (finite-models w l3 i l4) + ( all-models l2 l3 i l4) + ( decidable-models l2 l3 i l4) + ( soundness-union-same-class + ( ax-k i ∪ ax-s i) + ( ax-n i) + ( all-models l2 l3 i l4) + ( soundness-union-same-class + ( ax-k i) + ( ax-s i) + ( all-models l2 l3 i l4) + ( ax-k-soundness i l3 l4) + ( ax-s-soundness i l3 l4)) + ( ax-n-soundness i l3 l4)) + ( ax-dn-soundness i l3 l4) + ( all-models-is-largest-class l2 l3 i l4 (decidable-models l2 l3 i l4)) + + soundness-K-finite : soundness (modal-logic-K i) (finite-models l2 l3 i l4) soundness-K-finite = soundness-subclass ( modal-logic-K i) - ( decidable-models w l3 i l4) - ( finite-models w l3 i l4) - ( finite-models-subclass-decidable-models w l3 i l4) + ( decidable-models l2 l3 i l4) + ( finite-models l2 l3 i l4) + ( finite-models-subclass-decidable-models l2 l3 i l4) ( soundness-K) module _ @@ -131,11 +172,11 @@ module _ map-inv-raise ( soundness-K-finite ( i) - ( unit) ( ⊥) ( bot-in-logic) + ( unit) ( pair - ( (unit-trunc-Prop star) , (λ _ _ → empty-Prop)) + ( pair (unit-trunc-Prop star) (λ _ _ → empty-Prop)) ( λ _ _ → empty-Prop)) ( triple ( is-finite-unit) diff --git a/src/modal-logic/soundness.lagda.md b/src/modal-logic/soundness.lagda.md index adfbeb89fe..00f5f9123c 100644 --- a/src/modal-logic/soundness.lagda.md +++ b/src/modal-logic/soundness.lagda.md @@ -37,10 +37,10 @@ TODO module _ {l1 l2 l3 l4 l5 l6 : Level} {i : Set l1} (logic : formulas l2 i) - {w : UU l3} (C : model-class w l4 i l5 l6) + (C : model-class l3 l4 i l5 l6) where - soundness : UU (l1 ⊔ l2 ⊔ l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6) + soundness : UU (l1 ⊔ l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6) soundness = logic ⊆ class-modal-logic C ``` @@ -50,16 +50,16 @@ module _ module _ {l1 l2 l3 l4 l5 l6 : Level} {i : Set l1} {axioms : formulas l2 i} - {w : UU l3} (C : model-class w l4 i l5 l6) + (C : model-class l3 l4 i l5 l6) where soundness-axioms : soundness axioms C → {a : formula i} → axioms ⊢ a → type-Prop (C ⊨C a) soundness-axioms H (ax x) = H _ x - soundness-axioms H (mp dab da) M in-class x = - soundness-axioms H dab M in-class x (soundness-axioms H da M in-class x) - soundness-axioms H (nec d) M in-class _ y _ = - soundness-axioms H d M in-class y + soundness-axioms H (mp dab da) w M in-class x = + soundness-axioms H dab w M in-class x (soundness-axioms H da w M in-class x) + soundness-axioms H (nec d) w M in-class _ y _ = + soundness-axioms H d w M in-class y soundness-modal-logic : soundness axioms C → soundness (modal-logic axioms) C soundness-modal-logic H a = @@ -68,8 +68,7 @@ module _ module _ {l1 l2 l3 l4 l5 l6 l7 : Level} {i : Set l1} (logic : formulas l2 i) - {w : UU l3} - (C₁ : model-class w l4 i l5 l6) (C₂ : model-class w l4 i l5 l7) + (C₁ : model-class l3 l4 i l5 l6) (C₂ : model-class l3 l4 i l5 l7) where soundness-subclass : C₂ ⊆ C₁ → soundness logic C₁ → soundness logic C₂ @@ -83,29 +82,29 @@ module _ module _ {l1 l2 l3 l4 l5 l6 l7 l8 : Level} {i : Set l1} (ax₁ : formulas l2 i) (ax₂ : formulas l3 i) - {w : UU l4} - (C₁ : model-class w l5 i l6 l7) (C₂ : model-class w l5 i l6 l8) + (C₁ : model-class l4 l5 i l6 l7) (C₂ : model-class l4 l5 i l6 l8) (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) where forces-in-intersection : + (w : UU l4) (M : kripke-model w l5 i l6) → - is-in-subtype (intersection-subtype C₁ C₂) M → + is-in-subtype (intersection-subtype C₁ C₂) (w , M) → (a : formula i) → is-in-subtype ax₁ a + is-in-subtype ax₂ a → type-Prop (M ⊨M a) - forces-in-intersection M in-class a (inl d) = - sound₁ a d M (subtype-intersection-left C₁ C₂ M in-class) - forces-in-intersection M in-class a (inr d) = - sound₂ a d M (subtype-intersection-right C₁ C₂ M in-class) + forces-in-intersection w M in-class a (inl d) = + sound₁ a d w M (subtype-intersection-left C₁ C₂ (w , M) in-class) + forces-in-intersection w M in-class a (inr d) = + sound₂ a d w M (subtype-intersection-right C₁ C₂ (w , M) in-class) soundness-union : soundness (union-subtype ax₁ ax₂) (intersection-subtype C₁ C₂) - soundness-union a is-ax M in-class = + soundness-union a is-ax w M in-class = apply-universal-property-trunc-Prop ( is-ax) ( M ⊨M a) - ( forces-in-intersection M in-class a) + ( forces-in-intersection w M in-class a) soundness-modal-logic-union : soundness (modal-logic (union-subtype ax₁ ax₂)) (intersection-subtype C₁ C₂) @@ -115,12 +114,11 @@ module _ module _ {l1 l2 l3 l4 l5 l6 l7 : Level} {i : Set l1} (ax₁ : formulas l2 i) (ax₂ : formulas l3 i) - {w : UU l4} where soundness-modal-logic-union-subclass-left-sublevels : (l8 : Level) - (C₁ : model-class w l5 i l6 (l7 ⊔ l8)) (C₂ : model-class w l5 i l6 l7) + (C₁ : model-class l4 l5 i l6 (l7 ⊔ l8)) (C₂ : model-class l4 l5 i l6 l7) (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) → C₁ ⊆ C₂ → soundness (modal-logic (union-subtype ax₁ ax₂)) C₁ @@ -133,7 +131,7 @@ module _ soundness-modal-logic-union-subclass-right-sublevels : (l8 : Level) - (C₁ : model-class w l5 i l6 l7) (C₂ : model-class w l5 i l6 (l7 ⊔ l8)) + (C₁ : model-class l4 l5 i l6 l7) (C₂ : model-class l4 l5 i l6 (l7 ⊔ l8)) (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) → C₂ ⊆ C₁ → soundness (modal-logic (union-subtype ax₁ ax₂)) C₂ @@ -145,7 +143,7 @@ module _ ( soundness-modal-logic-union ax₁ ax₂ C₁ C₂ sound₁ sound₂) soundness-modal-logic-union-subclass-left : - (C₁ : model-class w l5 i l6 l7) (C₂ : model-class w l5 i l6 l7) + (C₁ : model-class l4 l5 i l6 l7) (C₂ : model-class l4 l5 i l6 l7) (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) → C₁ ⊆ C₂ → soundness (modal-logic (union-subtype ax₁ ax₂)) C₁ @@ -153,7 +151,7 @@ module _ soundness-modal-logic-union-subclass-left-sublevels lzero soundness-modal-logic-union-subclass-right : - (C₁ : model-class w l5 i l6 l7) (C₂ : model-class w l5 i l6 l7) + (C₁ : model-class l4 l5 i l6 l7) (C₂ : model-class l4 l5 i l6 l7) (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) → C₂ ⊆ C₁ → soundness (modal-logic (union-subtype ax₁ ax₂)) C₂ @@ -163,8 +161,7 @@ module _ module _ {l1 l2 l3 l4 l5 l6 l7 : Level} {i : Set l1} (ax₁ : formulas l2 i) (ax₂ : formulas l3 i) - {w : UU l4} - (C : model-class w l5 i l6 l7) + (C : model-class l4 l5 i l6 l7) (sound₁ : soundness ax₁ C) (sound₂ : soundness ax₂ C) where From eb312914592884af45a651b4442b560ca9175a93 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Wed, 27 Mar 2024 19:47:59 +0300 Subject: [PATCH 34/63] Add S5 soundness and refactor kripke-frame --- src/modal-logic.lagda.md | 1 + src/modal-logic/axioms.lagda.md | 31 ++- .../canonical-model-theorem.lagda.md | 7 +- src/modal-logic/completeness-K.lagda.md | 2 +- src/modal-logic/kripke-semantics.lagda.md | 260 +++++++++++++----- src/modal-logic/modal-logic-K.lagda.md | 16 +- src/modal-logic/modal-logic-S5.lagda.md | 96 +++++++ src/modal-logic/soundness.lagda.md | 51 +++- 8 files changed, 363 insertions(+), 101 deletions(-) create mode 100644 src/modal-logic/modal-logic-S5.lagda.md diff --git a/src/modal-logic.lagda.md b/src/modal-logic.lagda.md index 3a3f9e7e88..9e4a865ab3 100644 --- a/src/modal-logic.lagda.md +++ b/src/modal-logic.lagda.md @@ -11,6 +11,7 @@ open import modal-logic.formulas public open import modal-logic.kripke-semantics public open import modal-logic.logic-syntax public open import modal-logic.modal-logic-K public +open import modal-logic.modal-logic-S5 public open import modal-logic.soundness public open import modal-logic.weak-deduction public ``` diff --git a/src/modal-logic/axioms.lagda.md b/src/modal-logic/axioms.lagda.md index 461baaaedf..57630935f1 100644 --- a/src/modal-logic/axioms.lagda.md +++ b/src/modal-logic/axioms.lagda.md @@ -138,25 +138,48 @@ module _ where ax-k-soundness : soundness (ax-k i) (all-models l2 l3 i l4) - ax-k-soundness (a →ₘ b →ₘ a) (_ , _ , refl) w M _ x fa _ = fa + ax-k-soundness (a →ₘ b →ₘ a) (_ , _ , refl) M _ x fa _ = fa ax-n-soundness : soundness (ax-n i) (all-models l2 l3 i l4) ax-n-soundness (□ (a →ₘ b) →ₘ □ a →ₘ □ b) (_ , _ , refl) - w M in-class x fab fa y r = + M in-class x fab fa y r = fab y r (fa y r) ax-s-soundness : soundness (ax-s i) (all-models l2 l3 i l4) ax-s-soundness ((a →ₘ b →ₘ c) →ₘ (a →ₘ b) →ₘ a →ₘ c) (_ , _ , _ , refl) - w M in-class x fabc fab fa = + M in-class x fabc fab fa = fabc fa (fab fa) ax-dn-soundness : soundness (ax-dn i) (decidable-models l2 l3 i l4) - ax-dn-soundness .(~~ a →ₘ a) (a , refl) w M is-dec x f + ax-dn-soundness .(~~ a →ₘ a) (a , refl) M is-dec x f with (is-dec a x) ... | inl fa = fa ... | inr fna = raise-ex-falso _ (f (λ fa -> map-raise (fna fa))) + + ax-m-soundness : soundness (ax-m i) (reflexive-class l2 l3 i l4) + ax-m-soundness (□ a →ₘ a) (_ , refl) M is-refl x fa = fa x (is-refl x) + + ax-b-soundness : soundness (ax-b i) (symmetry-class l2 l3 i l4) + ax-b-soundness .(a →ₘ □ ◇ a) (a , refl) M is-sym x fa y r contra = + contra x (is-sym x y r) fa + + ax-d-soundness : soundness (ax-d i) (serial-class l2 l3 i l4) + ax-d-soundness .(□ a →ₘ ◇ a) (a , refl) M is-serial x fa contra = + map-raise + ( apply-universal-property-trunc-Prop + ( is-serial x) + ( empty-Prop) + ( λ (y , r) → map-inv-raise (contra y r (fa y r)))) + + ax-4-soundness : soundness (ax-4 i) (transitivity-class l2 l3 i l4) + ax-4-soundness .(□ a →ₘ □ □ a) (a , refl) M is-trans x fa y r-xy z r-yz = + fa z (is-trans x y z r-yz r-xy) + + ax-5-sooundness : soundness (ax-5 i) (euclidean-class l2 l3 i l4) + ax-5-sooundness .(◇ a →ₘ □ ◇ a) (a , refl) M is-eucl x fa y r-xy contra = + fa (λ z r-xz → contra z (is-eucl x y z r-xy r-xz)) ``` diff --git a/src/modal-logic/canonical-model-theorem.lagda.md b/src/modal-logic/canonical-model-theorem.lagda.md index e256347e34..0dbc1cd756 100644 --- a/src/modal-logic/canonical-model-theorem.lagda.md +++ b/src/modal-logic/canonical-model-theorem.lagda.md @@ -693,8 +693,9 @@ module _ ( λ x → unit-trunc-Prop ((pr1 (pr1 x)) , ((pr2 (pr1 x)) , (pr2 (pr2 x))))) canonical-kripke-model : - kripke-model (canonical-kripke-model-world-type) (l1 ⊔ l2) i (l1 ⊔ l2) - pr1 (pr1 canonical-kripke-model) = + kripke-model (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) i (l1 ⊔ l2) + pr1 (pr1 (pr1 canonical-kripke-model)) = canonical-kripke-model-world-type + pr2 (pr1 (pr1 canonical-kripke-model)) = is-inhabited-canonical-kripke-model-world pr2 (pr1 canonical-kripke-model) x y = Π-Prop @@ -709,7 +710,7 @@ module _ (a : formula i) → ( (y : formulas (l1 ⊔ l2) i) → (is-canonical : is-in-subtype canonical-worlds y) → - model-relation i canonical-kripke-model + relation-kripke-model i canonical-kripke-model ( pr1 x , pr2 x , is-comp) ( y , is-canonical) → is-in-subtype y a) → diff --git a/src/modal-logic/completeness-K.lagda.md b/src/modal-logic/completeness-K.lagda.md index 5e25b064f2..3d9516a2bc 100644 --- a/src/modal-logic/completeness-K.lagda.md +++ b/src/modal-logic/completeness-K.lagda.md @@ -78,5 +78,5 @@ module _ ( lem) ( a)) ( not-in-logic) - ( in-kripke-logic _ _ star)) + ( in-kripke-logic _ star)) ``` diff --git a/src/modal-logic/kripke-semantics.lagda.md b/src/modal-logic/kripke-semantics.lagda.md index e7656274e4..45d51c1fb9 100644 --- a/src/modal-logic/kripke-semantics.lagda.md +++ b/src/modal-logic/kripke-semantics.lagda.md @@ -15,9 +15,14 @@ open import foundation.decidable-relations open import foundation.decidable-types open import foundation.dependent-pair-types open import foundation.empty-types +open import foundation.existential-quantification +open import foundation.function-extensionality open import foundation.function-types open import foundation.inhabited-types +open import foundation.intersections-subtypes +open import foundation.iterated-dependent-product-types open import foundation.negation +open import foundation.propositional-extensionality open import foundation.propositions open import foundation.raising-universe-levels open import foundation.sets @@ -25,6 +30,9 @@ open import foundation.subtypes open import foundation.unit-type open import foundation.universe-levels +open import foundation-core.equivalence-relations +open import foundation-core.identity-types + open import modal-logic.formulas open import modal-logic.logic-syntax @@ -44,111 +52,203 @@ TODO ```agda module _ - {l1 : Level} (w : UU l1) (l2 : Level) + (l1 l2 : Level) where - kripke-frame : UU (l1 ⊔ lsuc l2) - kripke-frame = is-inhabited w × Relation-Prop l2 w + kripke-frame : UU (lsuc l1 ⊔ lsuc l2) + kripke-frame = Σ (Inhabited-Type l1) (Relation-Prop l2 ∘ type-Inhabited-Type) module _ - {l1 l2 : Level} {w : UU l1} + {l1 l2 : Level} where - frame-inhabited : kripke-frame w l2 → is-inhabited w - frame-inhabited = pr1 + type-kripke-frame : kripke-frame l1 l2 → UU l1 + type-kripke-frame = type-Inhabited-Type ∘ pr1 + + is-inhabited-type-kripke-frame : + (F : kripke-frame l1 l2) → is-inhabited (type-kripke-frame F) + is-inhabited-type-kripke-frame = is-inhabited-type-Inhabited-Type ∘ pr1 - frame-relation-Prop : kripke-frame w l2 → Relation-Prop l2 w - frame-relation-Prop = pr2 + relation-Prop-kripke-frame : + (F : kripke-frame l1 l2) → Relation-Prop l2 (type-kripke-frame F) + relation-Prop-kripke-frame = pr2 - frame-relation : kripke-frame w l2 → Relation l2 w - frame-relation = type-Relation-Prop ∘ frame-relation-Prop + relation-kripke-frame : + (F : kripke-frame l1 l2) → Relation l2 (type-kripke-frame F) + relation-kripke-frame = type-Relation-Prop ∘ relation-Prop-kripke-frame module _ - {l1 : Level} (w : UU l1) - (l2 : Level) - {l3 : Level} (i : Set l3) - (l4 : Level) + (l1 l2 : Level) {l3 : Level} (i : Set l3) (l4 : Level) where - kripke-model : UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4) - kripke-model = kripke-frame w l2 × (type-Set i → w → Prop l4) + kripke-model : UU (lsuc l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4) + kripke-model = + Σ (kripke-frame l1 l2) (λ F → type-Set i → type-kripke-frame F → Prop l4) module _ - (l1 l2 : Level) - {l3 : Level} (i : Set l3) - (l4 : Level) + {l1 l2 l3 l4 : Level} (i : Set l3) where - model-with-world : UU (lsuc l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4) - model-with-world = Σ (UU l1) (λ w → kripke-model w l2 i l4) + kripke-frame-kripke-model : kripke-model l1 l2 i l4 → kripke-frame l1 l2 + kripke-frame-kripke-model = pr1 + + type-kripke-model : kripke-model l1 l2 i l4 → UU l1 + type-kripke-model = type-kripke-frame ∘ kripke-frame-kripke-model - world-model-with-world : model-with-world → UU l1 - world-model-with-world = pr1 + valuate-kripke-model : + (M : kripke-model l1 l2 i l4) → type-Set i → type-kripke-model M → Prop l4 + valuate-kripke-model = pr2 - model-with-world-model : - (M : model-with-world) → kripke-model (world-model-with-world M) l2 i l4 - model-with-world-model = pr2 + relation-Prop-kripke-model : + (M : kripke-model l1 l2 i l4) → Relation-Prop l2 (type-kripke-model M) + relation-Prop-kripke-model = + relation-Prop-kripke-frame ∘ kripke-frame-kripke-model + + relation-kripke-model : + (M : kripke-model l1 l2 i l4) → Relation l2 (type-kripke-model M) + relation-kripke-model = + relation-kripke-frame ∘ kripke-frame-kripke-model + +module _ + (l1 l2 : Level) {l3 : Level} (i : Set l3) (l4 : Level) + where model-class : (l5 : Level) → UU (lsuc l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ lsuc l5) - model-class l5 = subtype l5 model-with-world + model-class l5 = subtype l5 (kripke-model l1 l2 i l4) all-models : model-class lzero all-models _ = unit-Prop +module _ + {l1 l2 l3 l4 : Level} (i : Set l3) + where + all-models-is-largest-class : - {l : Level} (C : model-class l) → C ⊆ all-models + {l5 : Level} (C : model-class l1 l2 i l4 l5) → C ⊆ all-models l1 l2 i l4 all-models-is-largest-class _ _ _ = star +-- TODO: move to binary relations +module _ + {l1 l2 : Level} {A : UU l1} (R : Relation l2 A) + where + + is-serial : UU (l1 ⊔ l2) + is-serial = (x : A) → ∃ A (λ y → R x y) + + is-euclidean : UU (l1 ⊔ l2) + is-euclidean = (x y z : A) → R x y → R x z → R y z + module _ - {l1 l2 l3 l4 : Level} {w : UU l1} (i : Set l3) + {l1 l2 : Level} {A : UU l1} (R : Relation-Prop l2 A) where - model-frame : kripke-model w l2 i l4 → kripke-frame w l2 - model-frame = pr1 + is-serial-Relation-Prop : UU (l1 ⊔ l2) + is-serial-Relation-Prop = is-serial (type-Relation-Prop R) - model-valuate : kripke-model w l2 i l4 → type-Set i → w → Prop l4 - model-valuate = pr2 + is-prop-is-serial-Relation-Prop : is-prop is-serial-Relation-Prop + is-prop-is-serial-Relation-Prop = + is-prop-Π (λ x → is-prop-∃ A _) - model-relation-Prop : kripke-model w l2 i l4 → Relation-Prop l2 w - model-relation-Prop = frame-relation-Prop ∘ model-frame + is-euclidean-Relation-Prop : UU (l1 ⊔ l2) + is-euclidean-Relation-Prop = is-euclidean (type-Relation-Prop R) - model-relation : kripke-model w l2 i l4 → Relation l2 w - model-relation = frame-relation ∘ model-frame + is-prop-is-euclidean-Relation-Prop : is-prop is-euclidean-Relation-Prop + is-prop-is-euclidean-Relation-Prop = + is-prop-iterated-Π 3 + ( λ x y z → + ( is-prop-function-type + ( is-prop-function-type (is-prop-type-Relation-Prop R y z)))) module _ - {l1 l2 l3 l4 : Level} {w : UU l1} {i : Set l3} + (l1 l2 : Level) {l3 : Level} (i : Set l3) (l4 : Level) where - private - l = l1 ⊔ l2 ⊔ l4 + relation-property-class : + {l5 : Level} → + ({A : UU l1} → subtype l5 (Relation-Prop l2 A)) → + model-class l1 l2 i l4 l5 + relation-property-class property M = + property (relation-Prop-kripke-model i M) + + -- intersection-relation-property-class : + -- {l5 l6 : Level} → + -- (R₁ : {A : UU l1} → subtype l5 (Relation-Prop l2 A)) → + -- (R₂ : {A : UU l1} → subtype l5 (Relation-Prop l2 A)) → + -- Id + -- ( intersection-subtype + -- ( relation-property-class R₁) + -- ( relation-property-class R₂)) + -- ( relation-property-class (intersection-subtype R₁ R₂)) + -- intersection-relation-property-class R₁ R₂ = refl + + reflexive-class : model-class l1 l2 i l4 (l1 ⊔ l2) + reflexive-class = + relation-property-class + ( λ x → + ( is-reflexive-Relation-Prop x , is-prop-is-reflexive-Relation-Prop x)) + + symmetry-class : model-class l1 l2 i l4 (l1 ⊔ l2) + symmetry-class = + relation-property-class + ( λ x → + ( is-symmetric-Relation-Prop x , is-prop-is-symmetric-Relation-Prop x)) + + transitivity-class : model-class l1 l2 i l4 (l1 ⊔ l2) + transitivity-class = + relation-property-class + ( λ x → + ( pair + ( is-transitive-Relation-Prop x) + ( is-prop-is-transitive-Relation-Prop x))) + + serial-class : model-class l1 l2 i l4 (l1 ⊔ l2) + serial-class = + relation-property-class + ( λ x → + ( is-serial-Relation-Prop x , is-prop-is-serial-Relation-Prop x)) + + euclidean-class : model-class l1 l2 i l4 (l1 ⊔ l2) + euclidean-class = + relation-property-class + ( λ x → + ( is-euclidean-Relation-Prop x , is-prop-is-euclidean-Relation-Prop x)) + + equivalence-class : model-class l1 l2 i l4 (l1 ⊔ l2) + equivalence-class = relation-property-class is-equivalence-relation-Prop + +module _ + {l1 l2 l3 l4 : Level} {i : Set l3} + where infix 5 _⊨_ infix 5 _⊭_ infix 5 _⊨M_ infix 5 _⊭M_ - _⊨_ : kripke-model w l2 i l4 × w → formula i → Prop l - (M , x) ⊨ var n = raise-Prop l (model-valuate i M n x) - (M , x) ⊨ ⊥ = raise-empty-Prop l + _⊨_ : + Σ (kripke-model l1 l2 i l4) (type-kripke-model i) → + formula i → + Prop (l1 ⊔ l2 ⊔ l4) + (M , x) ⊨ var n = raise-Prop (l1 ⊔ l2) (valuate-kripke-model i M n x) + (M , x) ⊨ ⊥ = raise-empty-Prop (l1 ⊔ l2 ⊔ l4) (M , x) ⊨ a →ₘ b = implication-Prop ((M , x) ⊨ a) ((M , x) ⊨ b) (M , x) ⊨ □ a = Π-Prop - ( w) - ( λ y → function-Prop (model-relation i M x y) ((M , y) ⊨ a)) + ( type-kripke-model i M) + ( λ y → function-Prop (relation-kripke-model i M x y) ((M , y) ⊨ a)) - _⊭_ : kripke-model w l2 i l4 × w → formula i → Prop l + _⊭_ : + Σ (kripke-model l1 l2 i l4) (type-kripke-model i) → + formula i → + Prop (l1 ⊔ l2 ⊔ l4) (M , x) ⊭ a = neg-Prop ((M , x) ⊨ a) - _⊨M_ : kripke-model w l2 i l4 → formula i → Prop l - M ⊨M a = Π-Prop w (λ x → (M , x) ⊨ a) + _⊨M_ : kripke-model l1 l2 i l4 → formula i → Prop (l1 ⊔ l2 ⊔ l4) + M ⊨M a = Π-Prop (type-kripke-model i M) (λ x → (M , x) ⊨ a) - _⊭M_ : kripke-model w l2 i l4 → formula i → Prop l + _⊭M_ : kripke-model l1 l2 i l4 → formula i → Prop (l1 ⊔ l2 ⊔ l4) M ⊭M a = neg-Prop (M ⊨M a) -module _ - {l1 l2 l3 l4 : Level} {i : Set l3} - where - _⊨C_ : {l5 : Level} → model-class l1 l2 i l4 l5 → @@ -156,11 +256,8 @@ module _ Prop (lsuc l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ l5) C ⊨C a = Π-Prop - ( UU l1) - ( λ w → - ( Π-Prop - ( kripke-model w l2 i l4)) - ( λ M → function-Prop (is-in-subtype C (w , M)) (M ⊨M a))) + ( kripke-model l1 l2 i l4) + ( λ M → function-Prop (is-in-subtype C M) (M ⊨M a)) class-modal-logic : {l5 : Level} → @@ -175,8 +272,8 @@ module _ (C₂ : model-class l1 l2 i l4 l6) → C₁ ⊆ C₂ → class-modal-logic C₂ ⊆ class-modal-logic C₁ - class-modal-logic-monotic C₁ C₂ sub _ in-modal-logic-C₂ w M in-C₁ = - in-modal-logic-C₂ w M (sub (w , M) in-C₁) + class-modal-logic-monotic C₁ C₂ sub _ in-modal-logic-C₂ M in-C₁ = + in-modal-logic-C₂ M (sub M in-C₁) module _ (l1 l2 : Level) @@ -185,39 +282,44 @@ module _ where decidable-models : model-class l1 l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4) - decidable-models (w , M) = + decidable-models M = Π-Prop ( formula i) - (λ a → Π-Prop w (λ x → is-decidable-Prop ((M , x) ⊨ a))) - -- (λ a → is-decidable-Prop (M ⊨M a)) + ( λ a → + ( Π-Prop + ( type-kripke-model i M) + ( λ x → is-decidable-Prop ((M , x) ⊨ a)))) + + -- decidable-subclass : + -- {l5 : Level} → finite-models : model-class l1 l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4) - finite-models (w , M) = + finite-models M = product-Prop - ( is-finite-Prop w) + ( is-finite-Prop (type-kripke-model i M)) ( product-Prop ( Π-Prop - ( w) + ( type-kripke-model i M) ( λ x → ( Π-Prop - ( w) - ( λ y → is-decidable-Prop (model-relation-Prop i M x y))))) + ( type-kripke-model i M) + ( λ y → is-decidable-Prop (relation-Prop-kripke-model i M x y))))) ( Π-Prop - ( w) + ( type-kripke-model i M) ( λ x → ( Π-Prop ( type-Set i) - ( λ n → is-decidable-Prop (model-valuate i M n x)))))) + ( λ n → is-decidable-Prop (valuate-kripke-model i M n x)))))) finite-models-subclass-decidable-models : finite-models ⊆ decidable-models - finite-models-subclass-decidable-models (w , M) (w-is-fin , dec-r , dec-v) = - lemma + finite-models-subclass-decidable-models M (w-is-fin , dec-r , dec-v) = lemma where lemma : - (a : formula i) (x : w) → is-decidable (type-Prop ((M , x) ⊨ a)) + (a : formula i) (x : type-kripke-model i M) → + is-decidable (type-Prop ((M , x) ⊨ a)) lemma (var n) x = - is-decidable-raise (l1 ⊔ l2 ⊔ l4) _ (dec-v x n) + is-decidable-raise (l1 ⊔ l2) _ (dec-v x n) lemma ⊥ x = inr map-inv-raise lemma (a →ₘ b) x = @@ -226,4 +328,14 @@ module _ is-decidable-Π-is-finite ( w-is-fin) ( λ y → is-decidable-function-type (dec-r x y) (lemma a y)) + +module _ + {l1 l2 l3 l4 : Level} (i : Set l3) + where + + decidable-subclass : + {l5 : Level} → + model-class l1 l2 i l4 l5 → + model-class l1 l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) + decidable-subclass C = intersection-subtype (decidable-models l1 l2 i l4) C ``` diff --git a/src/modal-logic/modal-logic-K.lagda.md b/src/modal-logic/modal-logic-K.lagda.md index ffec7968c3..29ae9f51a1 100644 --- a/src/modal-logic/modal-logic-K.lagda.md +++ b/src/modal-logic/modal-logic-K.lagda.md @@ -132,9 +132,10 @@ module _ {l1 l2 l3 l4 : Level} (i : Set l1) where - soundness-K : soundness (modal-logic-K i) (decidable-models l2 l3 i l4) - soundness-K = - soundness-modal-logic-union-subclass-right-sublevels + soundness-K-axioms : + soundness (modal-logic-K-axioms i) (decidable-models l2 l3 i l4) + soundness-K-axioms = + soundness-union-subclass-right-sublevels ( ax-k i ∪ ax-s i ∪ ax-n i) ( ax-dn i) ( l1 ⊔ l2 ⊔ l3 ⊔ l4) @@ -152,7 +153,11 @@ module _ ( ax-s-soundness i l3 l4)) ( ax-n-soundness i l3 l4)) ( ax-dn-soundness i l3 l4) - ( all-models-is-largest-class l2 l3 i l4 (decidable-models l2 l3 i l4)) + ( all-models-is-largest-class i (decidable-models l2 l3 i l4)) + + soundness-K : soundness (modal-logic-K i) (decidable-models l2 l3 i l4) + soundness-K = + soundness-modal-logic (decidable-models l2 l3 i l4) soundness-K-axioms soundness-K-finite : soundness (modal-logic-K i) (finite-models l2 l3 i l4) soundness-K-finite = @@ -174,9 +179,8 @@ module _ ( i) ( ⊥) ( bot-in-logic) - ( unit) ( pair - ( pair (unit-trunc-Prop star) (λ _ _ → empty-Prop)) + ( pair (unit , unit-trunc-Prop star) (λ _ _ → empty-Prop)) ( λ _ _ → empty-Prop)) ( triple ( is-finite-unit) diff --git a/src/modal-logic/modal-logic-S5.lagda.md b/src/modal-logic/modal-logic-S5.lagda.md new file mode 100644 index 0000000000..5da5367c2f --- /dev/null +++ b/src/modal-logic/modal-logic-S5.lagda.md @@ -0,0 +1,96 @@ +# Modal logic S5 + +```agda +module modal-logic.modal-logic-S5 where +``` + +
Imports + +```agda +open import foundation.intersections-subtypes +open import foundation.sets +open import foundation.subtypes +open import foundation.unions-subtypes +open import foundation.universe-levels + +open import foundation-core.coproduct-types + +open import modal-logic.axioms +open import modal-logic.formulas +open import modal-logic.kripke-semantics +open import modal-logic.logic-syntax +open import modal-logic.modal-logic-K +open import modal-logic.soundness +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l : Level} (i : Set l) + where + + modal-logic-S5-axioms : formulas l i + modal-logic-S5-axioms = modal-logic-K-axioms i ∪ (ax-m i ∪ (ax-b i ∪ ax-4 i)) + + modal-logic-S5 : formulas l i + modal-logic-S5 = modal-logic modal-logic-S5-axioms + + modal-logic-K-sub-S5 : modal-logic-K i ⊆ modal-logic-S5 + modal-logic-K-sub-S5 = + modal-logic-monotic + ( subtype-union-left + ( modal-logic-K-axioms i) + ( ax-m i ∪ (ax-b i ∪ ax-4 i))) + +module _ + {l1 l2 l3 l4 : Level} (i : Set l1) + where + + soundness-S5-additional-axioms : + soundness (ax-m i ∪ (ax-b i ∪ ax-4 i)) (equivalence-class l2 l3 i l4) + soundness-S5-additional-axioms = + soundness-union + ( ax-m i) + ( ax-b i ∪ ax-4 i) + ( reflexive-class l2 l3 i l4) + ( intersection-subtype + ( symmetry-class l2 l3 i l4) + ( transitivity-class l2 l3 i l4)) + ( ax-m-soundness i l3 l4) + ( soundness-union + ( ax-b i) + ( ax-4 i) + ( symmetry-class l2 l3 i l4) + ( transitivity-class l2 l3 i l4) + ( ax-b-soundness i l3 l4) + ( ax-4-soundness i l3 l4)) + + soundness-S5-axioms : + soundness + ( modal-logic-S5-axioms i) + ( decidable-subclass i (equivalence-class l2 l3 i l4)) + soundness-S5-axioms = + soundness-union + ( modal-logic-K-axioms i) + ( ax-m i ∪ (ax-b i ∪ ax-4 i)) + ( decidable-models l2 l3 i l4) + ( equivalence-class l2 l3 i l4) + ( soundness-K-axioms i) + ( soundness-S5-additional-axioms) + + soundness-S5 : + soundness + ( modal-logic-S5 i) + ( decidable-subclass i (equivalence-class l2 l3 i l4)) + soundness-S5 = + soundness-modal-logic + ( decidable-subclass i (equivalence-class l2 l3 i l4)) + ( soundness-S5-axioms) +``` diff --git a/src/modal-logic/soundness.lagda.md b/src/modal-logic/soundness.lagda.md index 00f5f9123c..69265fe09e 100644 --- a/src/modal-logic/soundness.lagda.md +++ b/src/modal-logic/soundness.lagda.md @@ -56,10 +56,10 @@ module _ soundness-axioms : soundness axioms C → {a : formula i} → axioms ⊢ a → type-Prop (C ⊨C a) soundness-axioms H (ax x) = H _ x - soundness-axioms H (mp dab da) w M in-class x = - soundness-axioms H dab w M in-class x (soundness-axioms H da w M in-class x) - soundness-axioms H (nec d) w M in-class _ y _ = - soundness-axioms H d w M in-class y + soundness-axioms H (mp dab da) M in-class x = + soundness-axioms H dab M in-class x (soundness-axioms H da M in-class x) + soundness-axioms H (nec d) M in-class _ y _ = + soundness-axioms H d M in-class y soundness-modal-logic : soundness axioms C → soundness (modal-logic axioms) C soundness-modal-logic H a = @@ -87,24 +87,23 @@ module _ where forces-in-intersection : - (w : UU l4) - (M : kripke-model w l5 i l6) → - is-in-subtype (intersection-subtype C₁ C₂) (w , M) → + (M : kripke-model l4 l5 i l6) → + is-in-subtype (intersection-subtype C₁ C₂) M → (a : formula i) → is-in-subtype ax₁ a + is-in-subtype ax₂ a → type-Prop (M ⊨M a) - forces-in-intersection w M in-class a (inl d) = - sound₁ a d w M (subtype-intersection-left C₁ C₂ (w , M) in-class) - forces-in-intersection w M in-class a (inr d) = - sound₂ a d w M (subtype-intersection-right C₁ C₂ (w , M) in-class) + forces-in-intersection M in-class a (inl d) = + sound₁ a d M (subtype-intersection-left C₁ C₂ M in-class) + forces-in-intersection M in-class a (inr d) = + sound₂ a d M (subtype-intersection-right C₁ C₂ M in-class) soundness-union : soundness (union-subtype ax₁ ax₂) (intersection-subtype C₁ C₂) - soundness-union a is-ax w M in-class = + soundness-union a is-ax M in-class = apply-universal-property-trunc-Prop ( is-ax) ( M ⊨M a) - ( forces-in-intersection w M in-class a) + ( forces-in-intersection M in-class a) soundness-modal-logic-union : soundness (modal-logic (union-subtype ax₁ ax₂)) (intersection-subtype C₁ C₂) @@ -116,6 +115,32 @@ module _ {i : Set l1} (ax₁ : formulas l2 i) (ax₂ : formulas l3 i) where + soundness-union-subclass-left-sublevels : + (l8 : Level) + (C₁ : model-class l4 l5 i l6 (l7 ⊔ l8)) (C₂ : model-class l4 l5 i l6 l7) + (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) → + C₁ ⊆ C₂ → + soundness (union-subtype ax₁ ax₂) C₁ + soundness-union-subclass-left-sublevels + l8 C₁ C₂ sound₁ sound₂ C₁-sub-C₂ = + tr + ( soundness (union-subtype ax₁ ax₂)) + ( intersection-subtype-left-sublevels l8 C₁ C₂ C₁-sub-C₂) + ( soundness-union ax₁ ax₂ C₁ C₂ sound₁ sound₂) + + soundness-union-subclass-right-sublevels : + (l8 : Level) + (C₁ : model-class l4 l5 i l6 l7) (C₂ : model-class l4 l5 i l6 (l7 ⊔ l8)) + (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) → + C₂ ⊆ C₁ → + soundness (union-subtype ax₁ ax₂) C₂ + soundness-union-subclass-right-sublevels + l8 C₁ C₂ sound₁ sound₂ C₂-sub-C₁ = + tr + ( soundness (union-subtype ax₁ ax₂)) + ( intersection-subtype-right-sublevels l8 C₁ C₂ C₂-sub-C₁) + ( soundness-union ax₁ ax₂ C₁ C₂ sound₁ sound₂) + soundness-modal-logic-union-subclass-left-sublevels : (l8 : Level) (C₁ : model-class l4 l5 i l6 (l7 ⊔ l8)) (C₂ : model-class l4 l5 i l6 l7) From 0e1fd0e448fb80f7e99e37eb992c28c722de9c78 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Fri, 29 Mar 2024 23:22:02 +0300 Subject: [PATCH 35/63] Proof completeness of S5 --- src/modal-logic.lagda.md | 3 + src/modal-logic/axioms.lagda.md | 22 +- .../canonical-model-theorem.lagda.md | 139 ++++++++---- src/modal-logic/canonical-theories.lagda.md | 165 ++++++++++++++ src/modal-logic/completeness-S5.lagda.md | 111 +++++++++ src/modal-logic/formulas-deduction.lagda.md | 210 ++++++++++++++++++ src/modal-logic/kripke-semantics.lagda.md | 59 +++-- src/modal-logic/logic-syntax.lagda.md | 3 + src/modal-logic/modal-logic-S5.lagda.md | 151 ++++++++++++- 9 files changed, 789 insertions(+), 74 deletions(-) create mode 100644 src/modal-logic/canonical-theories.lagda.md create mode 100644 src/modal-logic/completeness-S5.lagda.md create mode 100644 src/modal-logic/formulas-deduction.lagda.md diff --git a/src/modal-logic.lagda.md b/src/modal-logic.lagda.md index 9e4a865ab3..31f5501140 100644 --- a/src/modal-logic.lagda.md +++ b/src/modal-logic.lagda.md @@ -5,9 +5,12 @@ module modal-logic where open import modal-logic.axioms public open import modal-logic.canonical-model-theorem public +open import modal-logic.canonical-theories public open import modal-logic.completeness public open import modal-logic.completeness-K public +open import modal-logic.completeness-S5 public open import modal-logic.formulas public +open import modal-logic.formulas-deduction public open import modal-logic.kripke-semantics public open import modal-logic.logic-syntax public open import modal-logic.modal-logic-K public diff --git a/src/modal-logic/axioms.lagda.md b/src/modal-logic/axioms.lagda.md index 57630935f1..0f95536c79 100644 --- a/src/modal-logic/axioms.lagda.md +++ b/src/modal-logic/axioms.lagda.md @@ -138,19 +138,19 @@ module _ where ax-k-soundness : soundness (ax-k i) (all-models l2 l3 i l4) - ax-k-soundness (a →ₘ b →ₘ a) (_ , _ , refl) M _ x fa _ = fa + ax-k-soundness .(a →ₘ b →ₘ a) (a , b , refl) M _ x fa _ = fa ax-n-soundness : soundness (ax-n i) (all-models l2 l3 i l4) ax-n-soundness - (□ (a →ₘ b) →ₘ □ a →ₘ □ b) - (_ , _ , refl) + .(□ (a →ₘ b) →ₘ □ a →ₘ □ b) + (a , b , refl) M in-class x fab fa y r = fab y r (fa y r) ax-s-soundness : soundness (ax-s i) (all-models l2 l3 i l4) ax-s-soundness - ((a →ₘ b →ₘ c) →ₘ (a →ₘ b) →ₘ a →ₘ c) - (_ , _ , _ , refl) + .((a →ₘ b →ₘ c) →ₘ (a →ₘ b) →ₘ a →ₘ c) + (a , b , c , refl) M in-class x fabc fab fa = fabc fa (fab fa) @@ -160,14 +160,14 @@ module _ ... | inl fa = fa ... | inr fna = raise-ex-falso _ (f (λ fa -> map-raise (fna fa))) - ax-m-soundness : soundness (ax-m i) (reflexive-class l2 l3 i l4) - ax-m-soundness (□ a →ₘ a) (_ , refl) M is-refl x fa = fa x (is-refl x) + ax-m-soundness : soundness (ax-m i) (reflexive-kripke-class l2 l3 i l4) + ax-m-soundness .(□ a →ₘ a) (a , refl) M is-refl x fa = fa x (is-refl x) - ax-b-soundness : soundness (ax-b i) (symmetry-class l2 l3 i l4) + ax-b-soundness : soundness (ax-b i) (symmetry-kripke-class l2 l3 i l4) ax-b-soundness .(a →ₘ □ ◇ a) (a , refl) M is-sym x fa y r contra = contra x (is-sym x y r) fa - ax-d-soundness : soundness (ax-d i) (serial-class l2 l3 i l4) + ax-d-soundness : soundness (ax-d i) (serial-kripke-class l2 l3 i l4) ax-d-soundness .(□ a →ₘ ◇ a) (a , refl) M is-serial x fa contra = map-raise ( apply-universal-property-trunc-Prop @@ -175,11 +175,11 @@ module _ ( empty-Prop) ( λ (y , r) → map-inv-raise (contra y r (fa y r)))) - ax-4-soundness : soundness (ax-4 i) (transitivity-class l2 l3 i l4) + ax-4-soundness : soundness (ax-4 i) (transitivity-kripke-class l2 l3 i l4) ax-4-soundness .(□ a →ₘ □ □ a) (a , refl) M is-trans x fa y r-xy z r-yz = fa z (is-trans x y z r-yz r-xy) - ax-5-sooundness : soundness (ax-5 i) (euclidean-class l2 l3 i l4) + ax-5-sooundness : soundness (ax-5 i) (euclidean-kripke-class l2 l3 i l4) ax-5-sooundness .(◇ a →ₘ □ ◇ a) (a , refl) M is-eucl x fa y r-xy contra = fa (λ z r-xz → contra z (is-eucl x y z r-xy r-xz)) ``` diff --git a/src/modal-logic/canonical-model-theorem.lagda.md b/src/modal-logic/canonical-model-theorem.lagda.md index 0dbc1cd756..e4b231cd70 100644 --- a/src/modal-logic/canonical-model-theorem.lagda.md +++ b/src/modal-logic/canonical-model-theorem.lagda.md @@ -40,6 +40,7 @@ open import lists.reversing-lists open import modal-logic.axioms open import modal-logic.completeness open import modal-logic.formulas +open import modal-logic.formulas-deduction open import modal-logic.kripke-semantics open import modal-logic.logic-syntax open import modal-logic.modal-logic-K @@ -72,44 +73,45 @@ module _ (contains-K : modal-logic-K i ⊆ modal-logic axioms) where - logic : formulas (l1 ⊔ l2) i - logic = modal-logic axioms + private + logic : formulas (l1 ⊔ l2) i + logic = modal-logic axioms - contains-ax-k : ax-k i ⊆ logic - contains-ax-k = - transitive-leq-subtype - ( ax-k i) - ( modal-logic-K i) - ( logic) - ( contains-K) - ( K-contains-ax-k i) + contains-ax-k : ax-k i ⊆ logic + contains-ax-k = + transitive-leq-subtype + ( ax-k i) + ( modal-logic-K i) + ( logic) + ( contains-K) + ( K-contains-ax-k i) - contains-ax-s : ax-s i ⊆ logic - contains-ax-s = - transitive-leq-subtype - ( ax-s i) - ( modal-logic-K i) - ( logic) - ( contains-K) - ( K-contains-ax-s i) + contains-ax-s : ax-s i ⊆ logic + contains-ax-s = + transitive-leq-subtype + ( ax-s i) + ( modal-logic-K i) + ( logic) + ( contains-K) + ( K-contains-ax-s i) - contains-ax-n : ax-n i ⊆ logic - contains-ax-n = - transitive-leq-subtype - ( ax-n i) - ( modal-logic-K i) - ( logic) - ( contains-K) - ( K-contains-ax-n i) + contains-ax-n : ax-n i ⊆ logic + contains-ax-n = + transitive-leq-subtype + ( ax-n i) + ( modal-logic-K i) + ( logic) + ( contains-K) + ( K-contains-ax-n i) - contains-ax-dn : ax-dn i ⊆ logic - contains-ax-dn = - transitive-leq-subtype - ( ax-dn i) - ( modal-logic-K i) - ( logic) - ( contains-K) - ( K-contains-ax-dn i) + contains-ax-dn : ax-dn i ⊆ logic + contains-ax-dn = + transitive-leq-subtype + ( ax-dn i) + ( modal-logic-K i) + ( logic) + ( contains-K) + ( K-contains-ax-dn i) is-L-consistent-theory-Prop : formulas (l1 ⊔ l2) i → Prop (l1 ⊔ l2) is-L-consistent-theory-Prop t = @@ -294,6 +296,61 @@ module _ ( subtype-union-right (Id-formula-Prop (~ a)) (pr1 x))) ( subtype-union-left (Id-formula-Prop (~ a)) (pr1 x) (~ a) refl))) + -- postulate + -- helper : (a : formula i) → ◇ (~ a) → ~ □ a + + -- TODO: rename + lemma-box-diamond : + LEM (l1 ⊔ l2) → + (x : L-consistent-theory) → is-L-complete-theory x → + (y : L-consistent-theory) → is-L-complete-theory y → + (λ a → + ( exists-Prop + ( formula i) + ( λ b → product-Prop (Id-formula-Prop a (◇ b)) ( pr1 y b)))) ⊆ pr1 x → + (λ a → pr1 x (□ a)) ⊆ pr1 y + lemma-box-diamond lem x x-is-comp y y-is-comp sub a box-a-in-x = + apply-universal-property-trunc-Prop + ( complete-theory-contains-all-formulas lem y y-is-comp a) + ( pr1 y a) + ( λ + { (inl a-in-y) → a-in-y + ; (inr not-a-in-y) → + ( ex-falso + ( pr2 x + ( weak-modal-logic-mp + ( weak-modal-logic-diamond-negate + ( i) + ( modal-logic axioms ∪ pr1 x) + ( transitive-leq-subtype + ( modal-logic-K i) + ( modal-logic axioms ∪ pr1 x) + ( weak-modal-logic (modal-logic axioms ∪ pr1 x)) + ( axioms-subset-weak-modal-logic + ( modal-logic axioms ∪ pr1 x)) + ( transitive-leq-subtype + ( modal-logic-K i) + ( modal-logic axioms) + ( modal-logic axioms ∪ pr1 x) + ( subtype-union-left (modal-logic axioms) (pr1 x)) + ( contains-K))) + ( axioms-subset-weak-modal-logic + ( modal-logic axioms ∪ pr1 x) + ( ◇ ~ a) + ( subtype-union-right + ( modal-logic axioms) + ( pr1 x) + ( ◇ ~ a) + ( sub (◇ ~ a) (intro-∃ (~ a) (refl , not-a-in-y)))))) + ( axioms-subset-weak-modal-logic + ( modal-logic axioms ∪ pr1 x) + ( □ a) + ( subtype-union-right + ( modal-logic axioms) + ( pr1 x) + ( □ a) + ( box-a-in-x))))))}) + complete-theory-implication : LEM (l1 ⊔ l2) → (x : L-consistent-theory) → is-L-complete-theory x → @@ -692,15 +749,17 @@ module _ ( is-inhabited-Prop canonical-kripke-model-world-type) ( λ x → unit-trunc-Prop ((pr1 (pr1 x)) , ((pr2 (pr1 x)) , (pr2 (pr2 x))))) - canonical-kripke-model : - kripke-model (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) i (l1 ⊔ l2) - pr1 (pr1 (pr1 canonical-kripke-model)) = canonical-kripke-model-world-type - pr2 (pr1 (pr1 canonical-kripke-model)) = - is-inhabited-canonical-kripke-model-world - pr2 (pr1 canonical-kripke-model) x y = + canonical-kripke-frame : kripke-frame (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) + pr1 (pr1 canonical-kripke-frame) = canonical-kripke-model-world-type + pr2 (pr1 canonical-kripke-frame) = is-inhabited-canonical-kripke-model-world + pr2 canonical-kripke-frame x y = Π-Prop ( formula i) ( λ a → hom-Prop (pr1 x (□ a)) (pr1 y a)) + + canonical-kripke-model : + kripke-model (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) i (l1 ⊔ l2) + pr1 canonical-kripke-model = canonical-kripke-frame pr2 canonical-kripke-model n x = pr1 x (var n) complete-theory-box : diff --git a/src/modal-logic/canonical-theories.lagda.md b/src/modal-logic/canonical-theories.lagda.md new file mode 100644 index 0000000000..781384f385 --- /dev/null +++ b/src/modal-logic/canonical-theories.lagda.md @@ -0,0 +1,165 @@ +# Canonical threories + +```agda +module modal-logic.canonical-theories where +``` + +
Imports + +```agda +open import foundation.dependent-pair-types +open import foundation.empty-types +open import foundation.identity-types +open import foundation.law-of-excluded-middle +open import foundation.logical-equivalences +open import foundation.propositional-resizing +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.sets +open import foundation.subtypes +open import foundation.universe-levels + +open import foundation-core.coproduct-types + +open import modal-logic.axioms +open import modal-logic.canonical-model-theorem +open import modal-logic.completeness +open import modal-logic.formulas +open import modal-logic.kripke-semantics +open import modal-logic.logic-syntax +open import modal-logic.modal-logic-K +open import modal-logic.modal-logic-S5 +open import modal-logic.soundness +open import modal-logic.weak-deduction + +open import order-theory.maximal-elements-posets +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +-- TODO: levels +module _ + {l1 : Level} + (i : Set l1) + (lem : LEM l1) + (zorn : Zorn (lsuc l1) l1 l1) + (prop-resize : propositional-resizing l1 (lsuc l1)) + (axioms : modal-theory l1 i) + (is-cons : is-consistent-modal-logic (modal-logic axioms)) + (contains-K : modal-logic-K i ⊆ modal-logic axioms) + where + + is-canonical-ax-m : + ax-m i ⊆ modal-logic axioms → + is-in-subtype + ( reflexive-kripke-class (lsuc l1) l1 i l1) + ( canonical-kripke-model axioms zorn prop-resize is-cons contains-K) + is-canonical-ax-m ax-m-subset x a box-a-in-x = + weak-modal-logic-subset-complete-theory + ( axioms) + ( zorn) + ( prop-resize) + ( is-cons) + ( contains-K) + ( pr1 x , pr1 (pr2 x)) + ( pr2 (pr2 x)) + ( a) + ( weak-modal-logic-mp + ( weak-modal-logic-ax + ( logic-subset-L-complete-theory + ( axioms) + ( zorn) + ( prop-resize) + ( is-cons) + ( contains-K) + ( pr1 x , pr1 (pr2 x)) + ( pr2 (pr2 x)) + ( □ a →ₘ a) + ( ax-m-subset _ (a , refl)))) + ( weak-modal-logic-ax box-a-in-x)) + + is-canonical-ax-b : + ax-b i ⊆ modal-logic axioms → + is-in-subtype + ( symmetry-kripke-class (lsuc l1) l1 i l1) + ( canonical-kripke-model axioms zorn prop-resize is-cons contains-K) + is-canonical-ax-b ax-b-subset x y r-xy a box-a-in-y = + lemma-box-diamond + ( axioms) + ( zorn) + ( prop-resize) + ( is-cons) + ( contains-K) + ( lem) + ( pr1 y , pr1 (pr2 y)) + ( pr2 (pr2 y)) + ( pr1 x , pr1 (pr2 x)) + ( pr2 (pr2 x)) + ( λ b → + ( map-universal-property-trunc-Prop + ( pr1 y b) + ( λ { (c , refl , c-in-x) → + ( r-xy (◇ c) + ( weak-modal-logic-subset-complete-theory + ( axioms) + ( zorn) + ( prop-resize) + ( is-cons) + ( contains-K) + ( pr1 x , pr1 (pr2 x)) + ( pr2 (pr2 x)) + ( □ ◇ c) + ( weak-modal-logic-mp + ( weak-modal-logic-ax + ( logic-subset-L-complete-theory + ( axioms) + ( zorn) + ( prop-resize) + ( is-cons) + ( contains-K) + ( pr1 x , pr1 (pr2 x)) + ( pr2 (pr2 x)) + ( c →ₘ □ ◇ c) + ( ax-b-subset _ (c , refl)))) + ( weak-modal-logic-ax c-in-x))))}))) + ( a) + ( box-a-in-y) + + is-canonical-ax-4 : + ax-4 i ⊆ modal-logic axioms → + is-in-subtype + ( transitivity-kripke-class (lsuc l1) l1 i l1) + ( canonical-kripke-model axioms zorn prop-resize is-cons contains-K) + is-canonical-ax-4 ax-4-subset x y z r-yz r-xy a box-a-in-x = + r-yz a + ( r-xy (□ a) + ( weak-modal-logic-subset-complete-theory + ( axioms) + ( zorn) + ( prop-resize) + ( is-cons) + ( contains-K) + ( pr1 x , pr1 (pr2 x)) + ( pr2 (pr2 x)) + ( □ □ a) + ( weak-modal-logic-mp + ( weak-modal-logic-ax + ( logic-subset-L-complete-theory + ( axioms) + ( zorn) + ( prop-resize) + ( is-cons) + ( contains-K) + ( pr1 x , pr1 (pr2 x)) + ( pr2 (pr2 x)) + ( □ a →ₘ □ □ a) + ( ax-4-subset _ (a , refl)))) + ( weak-modal-logic-ax box-a-in-x)))) +``` diff --git a/src/modal-logic/completeness-S5.lagda.md b/src/modal-logic/completeness-S5.lagda.md new file mode 100644 index 0000000000..602a3d9c0d --- /dev/null +++ b/src/modal-logic/completeness-S5.lagda.md @@ -0,0 +1,111 @@ +# Completeness of S5 + +```agda +module modal-logic.completeness-S5 where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.cartesian-product-types +open import foundation.coproduct-types +open import foundation.dependent-pair-types +open import foundation.disjunction +open import foundation.empty-types +open import foundation.existential-quantification +open import foundation.function-types +open import foundation.identity-types +open import foundation.inhabited-types +open import foundation.intersections-subtypes +open import foundation.law-of-excluded-middle +open import foundation.logical-equivalences +open import foundation.negation +open import foundation.propositional-resizing +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.raising-universe-levels +open import foundation.sets +open import foundation.subtypes +open import foundation.transport-along-identifications +open import foundation.unions-subtypes +open import foundation.unit-type +open import foundation.universe-levels + +open import modal-logic.axioms +open import modal-logic.canonical-model-theorem +open import modal-logic.canonical-theories +open import modal-logic.completeness +open import modal-logic.formulas +open import modal-logic.kripke-semantics +open import modal-logic.logic-syntax +open import modal-logic.modal-logic-S5 +open import modal-logic.soundness +open import modal-logic.weak-deduction + +open import order-theory.maximal-elements-posets +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +-- TODO: levels +module _ + {l1 : Level} + (i : Set l1) + (lem : LEM l1) + (zorn : Zorn (lsuc l1) l1 l1) + (prop-resize : propositional-resizing l1 (lsuc l1)) + where + + S5-canonical-model-is-equivalence : + is-in-subtype + ( equivalence-kripke-class (lsuc l1) l1 i l1) + ( canonical-kripke-model + ( modal-logic-S5-axioms i) + ( zorn) + ( prop-resize) + ( is-consistent-S5 i) + ( modal-logic-K-sub-S5 i)) + S5-canonical-model-is-equivalence = + triple + ( is-canonical-ax-m i lem zorn prop-resize + ( modal-logic-S5-axioms i) + ( is-consistent-S5 i) + ( modal-logic-K-sub-S5 i) + ( modal-logic-S5-contains-ax-m i)) + ( is-canonical-ax-b i lem zorn prop-resize + ( modal-logic-S5-axioms i) + ( is-consistent-S5 i) + ( modal-logic-K-sub-S5 i) + ( modal-logic-S5-contains-ax-b i)) + ( is-canonical-ax-4 i lem zorn prop-resize + ( modal-logic-S5-axioms i) + ( is-consistent-S5 i) + ( modal-logic-K-sub-S5 i) + ( modal-logic-S5-contains-ax-4 i)) + + completeness-S5 : + completeness (modal-logic-S5 i) (equivalence-kripke-class (lsuc l1) l1 i l1) + completeness-S5 a in-kripke-logic with lem (modal-logic-S5 i a) + ... | inl in-logic = in-logic + ... | inr not-in-logic = + ex-falso + ( forward-implication + ( canonical-model-theorem' + ( modal-logic-S5-axioms i) + ( zorn) + ( prop-resize) + ( is-consistent-S5 i) + ( modal-logic-K-sub-S5 i) + ( lem) + ( a)) + ( not-in-logic) + ( in-kripke-logic _ S5-canonical-model-is-equivalence)) +``` diff --git a/src/modal-logic/formulas-deduction.lagda.md b/src/modal-logic/formulas-deduction.lagda.md new file mode 100644 index 0000000000..09b8f2bea4 --- /dev/null +++ b/src/modal-logic/formulas-deduction.lagda.md @@ -0,0 +1,210 @@ +# Formulas deduction + +```agda +module modal-logic.formulas-deduction where +``` + +
Imports + +```agda +open import foundation.dependent-pair-types +open import foundation.logical-equivalences +open import foundation.sets +open import foundation.subtypes +open import foundation.unions-subtypes +open import foundation.universe-levels + +open import foundation-core.identity-types + +open import modal-logic.axioms +open import modal-logic.formulas +open import modal-logic.logic-syntax +open import modal-logic.modal-logic-K +open import modal-logic.weak-deduction +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l1 l2 : Level} (i : Set l1) + (axioms : modal-theory l2 i) + (is-normal : modal-logic-K i ⊆ modal-logic axioms) + where + + private + contains-ax-k : ax-k i ⊆ modal-logic axioms + contains-ax-k = + transitive-leq-subtype + ( ax-k i) + ( modal-logic-K i) + ( modal-logic axioms) + is-normal + ( K-contains-ax-k i) + + contains-ax-s : ax-s i ⊆ modal-logic axioms + contains-ax-s = + transitive-leq-subtype + ( ax-s i) + ( modal-logic-K i) + ( modal-logic axioms) + is-normal + ( K-contains-ax-s i) + + contains-ax-n : ax-n i ⊆ modal-logic axioms + contains-ax-n = + transitive-leq-subtype + ( ax-n i) + ( modal-logic-K i) + ( modal-logic axioms) + is-normal + ( K-contains-ax-n i) + + contains-ax-dn : ax-dn i ⊆ modal-logic axioms + contains-ax-dn = + transitive-leq-subtype + ( ax-dn i) + ( modal-logic-K i) + ( modal-logic axioms) + is-normal + ( K-contains-ax-dn i) + + modal-logic-const : + (a : formula i) {b : formula i} → + is-in-subtype (modal-logic axioms) b → + is-in-subtype (modal-logic axioms) (a →ₘ b) + modal-logic-const a {b} b-in-logic = + modal-logic-mp (contains-ax-k _ (b , a , refl)) (b-in-logic) + + modal-logic-transitivity : + {a b c : formula i} → + is-in-subtype (modal-logic axioms) (b →ₘ c) → + is-in-subtype (modal-logic axioms) (a →ₘ b) → + is-in-subtype (modal-logic axioms) (a →ₘ c) + modal-logic-transitivity {a} {b} {c} bc ab = + modal-logic-mp + ( modal-logic-mp + (contains-ax-s _ (a , b , c , refl)) + (modal-logic-const a bc)) + ( ab) + + modal-logic-implication-box' : + {a b : formula i} → + is-in-subtype (modal-logic axioms) (a →ₘ b) → + is-in-subtype (modal-logic axioms) (□ a →ₘ □ b) + modal-logic-implication-box' {a} {b} ab = + modal-logic-mp (contains-ax-n _ (a , b , refl)) (modal-logic-nec ab) + + modal-logic-implication-box : + {a b : formula i} → + is-in-subtype (modal-logic axioms) (a →ₘ b) → + is-in-subtype (modal-logic axioms) (□ a) → + is-in-subtype (modal-logic axioms) (□ b) + modal-logic-implication-box {a} {b} ab boxa = + modal-logic-mp + ( modal-logic-mp (contains-ax-n _ (a , b , refl)) (modal-logic-nec ab)) + ( boxa) + + modal-logic-implication-dn : + (a : formula i) → is-in-subtype (modal-logic axioms) (a →ₘ ~~ a) + modal-logic-implication-dn a = + subset-double-modal-logic axioms + ( a →ₘ (a →ₘ ⊥) →ₘ ⊥) + ( weak-modal-logic-subset-modal-logic + ( modal-logic axioms) + ( a →ₘ (a →ₘ ⊥) →ₘ ⊥) + ( forward-implication + ( deduction-lemma + ( modal-logic axioms) + ( transitive-leq-subtype + ( ax-k i) + ( modal-logic axioms) + ( weak-modal-logic (modal-logic axioms)) + ( axioms-subset-weak-modal-logic (modal-logic axioms)) + ( contains-ax-k)) + ( transitive-leq-subtype + ( ax-s i) + ( modal-logic axioms) + ( weak-modal-logic (modal-logic axioms)) + ( axioms-subset-weak-modal-logic (modal-logic axioms)) + ( contains-ax-s)) + ( a) + ( (a →ₘ ⊥) →ₘ ⊥)) + ( forward-implication + ( deduction-lemma + ( theory-add-formula a (modal-logic axioms)) + ( transitive-leq-subtype + ( ax-k i) + ( theory-add-formula a (modal-logic axioms)) + ( weak-modal-logic (theory-add-formula a (modal-logic axioms))) + ( axioms-subset-weak-modal-logic + ( theory-add-formula a (modal-logic axioms))) + ( transitive-leq-subtype + ( ax-k i) + ( modal-logic axioms) + ( theory-add-formula a (modal-logic axioms)) + ( subtype-union-right + ( Id-formula-Prop a) + ( modal-logic axioms)) + ( contains-ax-k))) + ( transitive-leq-subtype + ( ax-s i) + ( theory-add-formula a (modal-logic axioms)) + ( weak-modal-logic (theory-add-formula a (modal-logic axioms))) + ( axioms-subset-weak-modal-logic + ( theory-add-formula a (modal-logic axioms))) + ( transitive-leq-subtype + ( ax-s i) + ( modal-logic axioms) + ( theory-add-formula a (modal-logic axioms)) + ( subtype-union-right + ( Id-formula-Prop a) + ( modal-logic axioms)) + ( contains-ax-s))) + ( a →ₘ ⊥) + ( ⊥)) + ( weak-modal-logic-mp + ( weak-modal-logic-ax + ( subtype-union-left + ( Id-formula-Prop (~ a)) + ( theory-add-formula a (modal-logic axioms)) + ( ~ a) + ( refl))) + ( weak-modal-logic-ax + ( subtype-union-right + ( Id-formula-Prop (~ a)) + ( theory-add-formula a (modal-logic axioms)) + ( a) + ( subtype-union-left + ( Id-formula-Prop a) + ( modal-logic axioms) + ( a) + ( refl)))))))) + + modal-logic-diamond-negate : + {a : formula i} → + is-in-subtype (modal-logic axioms) (◇ ~ a) → + is-in-subtype (modal-logic axioms) (~ □ a) + modal-logic-diamond-negate {a} dia-a = + modal-logic-transitivity + ( dia-a) + ( modal-logic-implication-box' (modal-logic-implication-dn a)) + +module _ + {l1 l2 : Level} (i : Set l1) + (axioms : modal-theory l2 i) + (is-normal : modal-logic-K i ⊆ weak-modal-logic axioms) + where + + postulate + weak-modal-logic-diamond-negate : + {a : formula i} → + is-in-subtype (weak-modal-logic axioms) (◇ ~ a) → + is-in-subtype (weak-modal-logic axioms) (~ □ a) +``` diff --git a/src/modal-logic/kripke-semantics.lagda.md b/src/modal-logic/kripke-semantics.lagda.md index 45d51c1fb9..083992938f 100644 --- a/src/modal-logic/kripke-semantics.lagda.md +++ b/src/modal-logic/kripke-semantics.lagda.md @@ -62,6 +62,10 @@ module _ {l1 l2 : Level} where + -- TODO: replace with theorem + postulate + is-set-kripke-frame : is-set (kripke-frame l1 l2) + type-kripke-frame : kripke-frame l1 l2 → UU l1 type-kripke-frame = type-Inhabited-Type ∘ pr1 @@ -181,40 +185,41 @@ module _ -- ( relation-property-class (intersection-subtype R₁ R₂)) -- intersection-relation-property-class R₁ R₂ = refl - reflexive-class : model-class l1 l2 i l4 (l1 ⊔ l2) - reflexive-class = + reflexive-kripke-class : model-class l1 l2 i l4 (l1 ⊔ l2) + reflexive-kripke-class = relation-property-class ( λ x → ( is-reflexive-Relation-Prop x , is-prop-is-reflexive-Relation-Prop x)) - symmetry-class : model-class l1 l2 i l4 (l1 ⊔ l2) - symmetry-class = + symmetry-kripke-class : model-class l1 l2 i l4 (l1 ⊔ l2) + symmetry-kripke-class = relation-property-class ( λ x → ( is-symmetric-Relation-Prop x , is-prop-is-symmetric-Relation-Prop x)) - transitivity-class : model-class l1 l2 i l4 (l1 ⊔ l2) - transitivity-class = + transitivity-kripke-class : model-class l1 l2 i l4 (l1 ⊔ l2) + transitivity-kripke-class = relation-property-class ( λ x → ( pair ( is-transitive-Relation-Prop x) ( is-prop-is-transitive-Relation-Prop x))) - serial-class : model-class l1 l2 i l4 (l1 ⊔ l2) - serial-class = + serial-kripke-class : model-class l1 l2 i l4 (l1 ⊔ l2) + serial-kripke-class = relation-property-class ( λ x → ( is-serial-Relation-Prop x , is-prop-is-serial-Relation-Prop x)) - euclidean-class : model-class l1 l2 i l4 (l1 ⊔ l2) - euclidean-class = + euclidean-kripke-class : model-class l1 l2 i l4 (l1 ⊔ l2) + euclidean-kripke-class = relation-property-class ( λ x → ( is-euclidean-Relation-Prop x , is-prop-is-euclidean-Relation-Prop x)) - equivalence-class : model-class l1 l2 i l4 (l1 ⊔ l2) - equivalence-class = relation-property-class is-equivalence-relation-Prop + equivalence-kripke-class : model-class l1 l2 i l4 (l1 ⊔ l2) + equivalence-kripke-class = + relation-property-class is-equivalence-relation-Prop module _ {l1 l2 l3 l4 : Level} {i : Set l3} @@ -263,7 +268,7 @@ module _ {l5 : Level} → model-class l1 l2 i l4 l5 → formulas (lsuc l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ l5) i - class-modal-logic C a = C ⊨C a + class-modal-logic = _⊨C_ -- TODO: rename class-modal-logic-monotic : @@ -275,6 +280,28 @@ module _ class-modal-logic-monotic C₁ C₂ sub _ in-modal-logic-C₂ M in-C₁ = in-modal-logic-C₂ M (sub M in-C₁) +module _ + {l1 l2 : Level} + {l3 : Level} (i : Set l3) + (l4 : Level) + where + + kripke-frame-model-class : + kripke-frame l1 l2 → + model-class l1 l2 i l4 (lsuc l1 ⊔ lsuc l2) + pr1 (kripke-frame-model-class F M) = + kripke-frame-kripke-model i M = F + pr2 (kripke-frame-model-class F M) = + is-set-kripke-frame (kripke-frame-kripke-model i M) F + + -- frame-modal-logic : + -- kripke-frame l1 l2 → + -- formulas (l1 ⊔ l2 ⊔ l3 ⊔ lsuc l4) i + -- frame-modal-logic F a = + -- Π-Prop + -- ( type-Set i → type-kripke-frame F → Prop l4) + -- ( λ v → (F , v) ⊨M a) + module _ (l1 l2 : Level) {l3 : Level} (i : Set l3) @@ -338,4 +365,10 @@ module _ model-class l1 l2 i l4 l5 → model-class l1 l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) decidable-subclass C = intersection-subtype (decidable-models l1 l2 i l4) C + + finite-subclass : + {l5 : Level} → + model-class l1 l2 i l4 l5 → + model-class l1 l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) + finite-subclass C = intersection-subtype (finite-models l1 l2 i l4) C ``` diff --git a/src/modal-logic/logic-syntax.lagda.md b/src/modal-logic/logic-syntax.lagda.md index 1f34f40ca3..b2a10cc495 100644 --- a/src/modal-logic/logic-syntax.lagda.md +++ b/src/modal-logic/logic-syntax.lagda.md @@ -40,6 +40,9 @@ module _ formulas : UU (l1 ⊔ lsuc l2) formulas = subtype l2 (formula i) + modal-theory : UU (l1 ⊔ lsuc l2) + modal-theory = formulas + module _ {l1 l2 : Level} {i : Set l1} where diff --git a/src/modal-logic/modal-logic-S5.lagda.md b/src/modal-logic/modal-logic-S5.lagda.md index 5da5367c2f..4c1d217121 100644 --- a/src/modal-logic/modal-logic-S5.lagda.md +++ b/src/modal-logic/modal-logic-S5.lagda.md @@ -7,10 +7,16 @@ module modal-logic.modal-logic-S5 where
Imports ```agda +open import foundation.decidable-types +open import foundation.dependent-pair-types +open import foundation.empty-types open import foundation.intersections-subtypes +open import foundation.propositional-truncations +open import foundation.raising-universe-levels open import foundation.sets open import foundation.subtypes open import foundation.unions-subtypes +open import foundation.unit-type open import foundation.universe-levels open import foundation-core.coproduct-types @@ -21,6 +27,8 @@ open import modal-logic.kripke-semantics open import modal-logic.logic-syntax open import modal-logic.modal-logic-K open import modal-logic.soundness + +open import univalent-combinatorics.finite-types ```
@@ -49,48 +57,171 @@ module _ ( modal-logic-K-axioms i) ( ax-m i ∪ (ax-b i ∪ ax-4 i))) + modal-logic-S5-axioms-contains-ax-m : ax-m i ⊆ modal-logic-S5-axioms + modal-logic-S5-axioms-contains-ax-m = + transitive-leq-subtype + ( ax-m i) + ( ax-m i ∪ (ax-b i ∪ ax-4 i)) + ( modal-logic-K-axioms i ∪ (ax-m i ∪ (ax-b i ∪ ax-4 i))) + ( subtype-union-right + ( modal-logic-K-axioms i) + ( ax-m i ∪ (ax-b i ∪ ax-4 i))) + ( subtype-union-left + ( ax-m i) + ( ax-b i ∪ ax-4 i)) + + modal-logic-S5-axioms-contains-ax-b : ax-b i ⊆ modal-logic-S5-axioms + modal-logic-S5-axioms-contains-ax-b = + transitive-leq-subtype + ( ax-b i) + ( ax-m i ∪ (ax-b i ∪ ax-4 i)) + ( modal-logic-K-axioms i ∪ (ax-m i ∪ (ax-b i ∪ ax-4 i))) + ( subtype-union-right + ( modal-logic-K-axioms i) + ( ax-m i ∪ (ax-b i ∪ ax-4 i))) + ( transitive-leq-subtype + ( ax-b i) + ( ax-b i ∪ ax-4 i) + ( ax-m i ∪ (ax-b i ∪ ax-4 i)) + ( subtype-union-right (ax-m i) (ax-b i ∪ ax-4 i)) + ( subtype-union-left (ax-b i) (ax-4 i))) + + modal-logic-S5-axioms-contains-ax-4 : ax-4 i ⊆ modal-logic-S5-axioms + modal-logic-S5-axioms-contains-ax-4 = + transitive-leq-subtype + ( ax-4 i) + ( ax-m i ∪ (ax-b i ∪ ax-4 i)) + ( modal-logic-K-axioms i ∪ (ax-m i ∪ (ax-b i ∪ ax-4 i))) + ( subtype-union-right + ( modal-logic-K-axioms i) + ( ax-m i ∪ (ax-b i ∪ ax-4 i))) + ( transitive-leq-subtype + ( ax-4 i) + ( ax-b i ∪ ax-4 i) + ( ax-m i ∪ (ax-b i ∪ ax-4 i)) + ( subtype-union-right (ax-m i) (ax-b i ∪ ax-4 i)) + ( subtype-union-right (ax-b i) (ax-4 i))) + + modal-logic-S5-contains-ax-m : ax-m i ⊆ modal-logic-S5 + modal-logic-S5-contains-ax-m = + transitive-leq-subtype + ( ax-m i) + ( modal-logic-S5-axioms) + ( modal-logic-S5) + ( axioms-subset-modal-logic modal-logic-S5-axioms) + ( modal-logic-S5-axioms-contains-ax-m) + + modal-logic-S5-contains-ax-b : ax-b i ⊆ modal-logic-S5 + modal-logic-S5-contains-ax-b = + transitive-leq-subtype + ( ax-b i) + ( modal-logic-S5-axioms) + ( modal-logic-S5) + ( axioms-subset-modal-logic modal-logic-S5-axioms) + ( modal-logic-S5-axioms-contains-ax-b) + + modal-logic-S5-contains-ax-4 : ax-4 i ⊆ modal-logic-S5 + modal-logic-S5-contains-ax-4 = + transitive-leq-subtype + ( ax-4 i) + ( modal-logic-S5-axioms) + ( modal-logic-S5) + ( axioms-subset-modal-logic modal-logic-S5-axioms) + ( modal-logic-S5-axioms-contains-ax-4) + module _ {l1 l2 l3 l4 : Level} (i : Set l1) where soundness-S5-additional-axioms : - soundness (ax-m i ∪ (ax-b i ∪ ax-4 i)) (equivalence-class l2 l3 i l4) + soundness (ax-m i ∪ (ax-b i ∪ ax-4 i)) (equivalence-kripke-class l2 l3 i l4) soundness-S5-additional-axioms = soundness-union ( ax-m i) ( ax-b i ∪ ax-4 i) - ( reflexive-class l2 l3 i l4) + ( reflexive-kripke-class l2 l3 i l4) ( intersection-subtype - ( symmetry-class l2 l3 i l4) - ( transitivity-class l2 l3 i l4)) + ( symmetry-kripke-class l2 l3 i l4) + ( transitivity-kripke-class l2 l3 i l4)) ( ax-m-soundness i l3 l4) ( soundness-union ( ax-b i) ( ax-4 i) - ( symmetry-class l2 l3 i l4) - ( transitivity-class l2 l3 i l4) + ( symmetry-kripke-class l2 l3 i l4) + ( transitivity-kripke-class l2 l3 i l4) ( ax-b-soundness i l3 l4) ( ax-4-soundness i l3 l4)) soundness-S5-axioms : soundness ( modal-logic-S5-axioms i) - ( decidable-subclass i (equivalence-class l2 l3 i l4)) + ( decidable-subclass i (equivalence-kripke-class l2 l3 i l4)) soundness-S5-axioms = soundness-union ( modal-logic-K-axioms i) ( ax-m i ∪ (ax-b i ∪ ax-4 i)) ( decidable-models l2 l3 i l4) - ( equivalence-class l2 l3 i l4) + ( equivalence-kripke-class l2 l3 i l4) ( soundness-K-axioms i) ( soundness-S5-additional-axioms) soundness-S5 : soundness ( modal-logic-S5 i) - ( decidable-subclass i (equivalence-class l2 l3 i l4)) + ( decidable-subclass i (equivalence-kripke-class l2 l3 i l4)) soundness-S5 = soundness-modal-logic - ( decidable-subclass i (equivalence-class l2 l3 i l4)) + ( decidable-subclass i (equivalence-kripke-class l2 l3 i l4)) ( soundness-S5-axioms) + + soundness-S5-finite : + soundness + ( modal-logic-S5 i) + ( finite-subclass i (equivalence-kripke-class l2 l3 i l4)) + soundness-S5-finite = + soundness-subclass + ( modal-logic-S5 i) + ( decidable-subclass i (equivalence-kripke-class l2 l3 i l4)) + ( finite-subclass i (equivalence-kripke-class l2 l3 i l4)) + ( subtype-both-intersection + ( decidable-models l2 l3 i l4) + ( equivalence-kripke-class l2 l3 i l4) + ( finite-subclass i (equivalence-kripke-class l2 l3 i l4)) + ( transitive-leq-subtype + ( finite-subclass i (equivalence-kripke-class l2 l3 i l4)) + ( finite-models l2 l3 i l4) + ( decidable-models l2 l3 i l4) + ( finite-models-subclass-decidable-models l2 l3 i l4) + ( subtype-intersection-left + ( finite-models l2 l3 i l4) + ( equivalence-kripke-class l2 l3 i l4))) + ( subtype-intersection-right + ( finite-models l2 l3 i l4) + ( equivalence-kripke-class l2 l3 i l4))) + ( soundness-S5) + +module _ + {l1 : Level} (i : Set l1) + where + + is-consistent-S5 : is-consistent-modal-logic (modal-logic-S5 i) + is-consistent-S5 bot-in-logic = + map-inv-raise + ( soundness-S5-finite + ( i) + ( ⊥) + ( bot-in-logic) + ( pair + ( pair (unit , unit-trunc-Prop star) (λ _ _ → unit-Prop)) + ( λ _ _ → unit-Prop)) + ( pair + ( triple + ( is-finite-unit) + ( λ _ _ → inl star) + ( λ _ _ → is-decidable-unit)) + ( triple + ( λ x → star) + ( λ x y _ → star) + ( λ x y z _ _ → star))) + ( star)) ``` From b019f5bbfca31bd356f0b959573ba4dece2bb92f Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Sat, 6 Apr 2024 01:10:19 +0300 Subject: [PATCH 36/63] Add filtrations and finite approximability --- .../law-of-excluded-middle.lagda.md | 30 +- src/foundation/surjective-maps.lagda.md | 8 + src/lists/lists.lagda.md | 27 + src/modal-logic.lagda.md | 4 + .../canonical-model-theorem.lagda.md | 1 - .../finite-approximability.lagda.md | 196 +++++ ...kripke-models-filtrations-theorem.lagda.md | 190 ++++ .../kripke-models-filtrations.lagda.md | 829 ++++++++++++++++++ src/modal-logic/kripke-semantics.lagda.md | 45 +- src/modal-logic/modal-logic-K.lagda.md | 2 +- src/modal-logic/modal-logic-S5.lagda.md | 2 +- src/modal-logic/modal-logic-decision.lagda.md | 720 +++++++++++++++ src/modal-logic/weak-deduction.lagda.md | 22 + .../kuratowsky-finite-sets.lagda.md | 42 + 14 files changed, 2102 insertions(+), 16 deletions(-) create mode 100644 src/modal-logic/finite-approximability.lagda.md create mode 100644 src/modal-logic/kripke-models-filtrations-theorem.lagda.md create mode 100644 src/modal-logic/kripke-models-filtrations.lagda.md create mode 100644 src/modal-logic/modal-logic-decision.lagda.md diff --git a/src/foundation/law-of-excluded-middle.lagda.md b/src/foundation/law-of-excluded-middle.lagda.md index 6a67a25ac9..97258ff4c8 100644 --- a/src/foundation/law-of-excluded-middle.lagda.md +++ b/src/foundation/law-of-excluded-middle.lagda.md @@ -7,16 +7,21 @@ module foundation.law-of-excluded-middle where
Imports ```agda +open import foundation.booleans +open import foundation.decidable-propositions open import foundation.decidable-types open import foundation.dependent-pair-types open import foundation.propositions open import foundation.raising-universe-levels open import foundation.universe-levels -open import foundation-core.decidable-propositions +open import foundation-core.equality-dependent-pair-types +open import foundation-core.equivalences +open import foundation-core.identity-types open import foundation-core.negation open import univalent-combinatorics.2-element-types +open import univalent-combinatorics.finite-types ```
@@ -61,6 +66,29 @@ pr1 (pr2 (decidable-prop-Prop lem P)) = is-prop-type-Prop P pr2 (pr2 (decidable-prop-Prop lem P)) = lem P ``` +### TODO: Given LEM, Prop equiv Decidable-Prop + +```agda +prop-equiv-decidable-prop : + {l : Level} → LEM l → Prop l ≃ Decidable-Prop l +prop-equiv-decidable-prop lem = + pair + ( decidable-prop-Prop lem) + ( is-equiv-is-invertible + ( prop-Decidable-Prop) + ( λ P → + ( eq-pair-Σ + ( refl) + ( eq-is-prop (is-prop-is-decidable-prop (type-Decidable-Prop P))))) + ( λ P → eq-pair-Σ refl (eq-is-prop (is-prop-is-prop (type-Prop P))))) + +is-finite-Prop-LEM : {l : Level} → LEM l → is-finite (Prop l) +is-finite-Prop-LEM lem = + is-finite-equiv' + ( equiv-bool-Decidable-Prop ∘e prop-equiv-decidable-prop lem) + ( is-finite-bool) +``` + ### The unrestricted law of excluded middle does not hold ```agda diff --git a/src/foundation/surjective-maps.lagda.md b/src/foundation/surjective-maps.lagda.md index a22a750dac..808d18f658 100644 --- a/src/foundation/surjective-maps.lagda.md +++ b/src/foundation/surjective-maps.lagda.md @@ -463,6 +463,14 @@ module _ is-surjective g → is-surjective h → is-surjective (g ∘ h) is-surjective-comp {g} {h} = is-surjective-left-map-triangle (g ∘ h) g h refl-htpy + + surjection-comp : B ↠ X → A ↠ B → A ↠ X + surjection-comp g h = + pair + (map-surjection g ∘ map-surjection h) + (is-surjective-comp + ( is-surjective-map-surjection g) + ( is-surjective-map-surjection h)) ``` ### Functoriality of products preserves being surjective diff --git a/src/lists/lists.lagda.md b/src/lists/lists.lagda.md index a9f46d3ae1..59175a778c 100644 --- a/src/lists/lists.lagda.md +++ b/src/lists/lists.lagda.md @@ -31,6 +31,8 @@ open import foundation.truncated-types open import foundation.truncation-levels open import foundation.unit-type open import foundation.universe-levels + +open import univalent-combinatorics.standard-finite-types ```
@@ -142,6 +144,31 @@ rest-in-snoc-list (cons y xs) x z (is-in-tail .z .y .xs in-list) = is-in-tail z y (snoc xs x) (rest-in-snoc-list xs x z in-list) ``` +### TODO: change title + +```agda +module _ + {l : Level} {A : UU l} + where + + component-list : (l : list A) → Fin (length-list l) → A + component-list (cons x l) (inl k) = component-list l k + component-list (cons x l) (inr k) = x + + index-in-list : (a : A) (l : list A) → a ∈-list l → Fin (length-list l) + index-in-list a (cons .a l) (is-head .a .l) = + inr star + index-in-list a (cons x l) (is-in-tail .a .x .l in-list) = + inl (index-in-list a l in-list) + + eq-component-list-index-in-list : + (a : A) (l : list A) (in-list : a ∈-list l) → + a = component-list l (index-in-list a l in-list) + eq-component-list-index-in-list a .(cons a l) (is-head .a l) = refl + eq-component-list-index-in-list a .(cons x l) (is-in-tail .a x l in-list) = + eq-component-list-index-in-list a l in-list +``` + ### A list that uses cons is not nil ```agda diff --git a/src/modal-logic.lagda.md b/src/modal-logic.lagda.md index 31f5501140..e7c40e8ee0 100644 --- a/src/modal-logic.lagda.md +++ b/src/modal-logic.lagda.md @@ -9,12 +9,16 @@ open import modal-logic.canonical-theories public open import modal-logic.completeness public open import modal-logic.completeness-K public open import modal-logic.completeness-S5 public +open import modal-logic.finite-approximability public open import modal-logic.formulas public open import modal-logic.formulas-deduction public +open import modal-logic.kripke-models-filtrations public +open import modal-logic.kripke-models-filtrations-theorem public open import modal-logic.kripke-semantics public open import modal-logic.logic-syntax public open import modal-logic.modal-logic-K public open import modal-logic.modal-logic-S5 public +open import modal-logic.modal-logic-decision public open import modal-logic.soundness public open import modal-logic.weak-deduction public ``` diff --git a/src/modal-logic/canonical-model-theorem.lagda.md b/src/modal-logic/canonical-model-theorem.lagda.md index e4b231cd70..8eb907d55f 100644 --- a/src/modal-logic/canonical-model-theorem.lagda.md +++ b/src/modal-logic/canonical-model-theorem.lagda.md @@ -812,7 +812,6 @@ module _ ( b) ( box-in-logic))))))))))))}) where - list-to-implications : formula i → (l : list (formula i)) → diff --git a/src/modal-logic/finite-approximability.lagda.md b/src/modal-logic/finite-approximability.lagda.md new file mode 100644 index 0000000000..f51c202428 --- /dev/null +++ b/src/modal-logic/finite-approximability.lagda.md @@ -0,0 +1,196 @@ +# Finite approximability + +```agda +module modal-logic.finite-approximability where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.binary-relations +open import foundation.coproduct-types +open import foundation.decidable-types +open import foundation.dependent-pair-types +open import foundation.equivalence-classes +open import foundation.equivalences +open import foundation.existential-quantification +open import foundation.function-types +open import foundation.identity-types +open import foundation.inhabited-types +open import foundation.intersections-subtypes +open import foundation.law-of-excluded-middle +open import foundation.logical-equivalences +open import foundation.propositional-resizing +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.raising-universe-levels +open import foundation.sets +open import foundation.subtypes +open import foundation.transport-along-identifications +open import foundation.unions-subtypes +open import foundation.unit-type +open import foundation.universe-levels + +open import foundation-core.cartesian-product-types +open import foundation-core.embeddings +open import foundation-core.equivalence-relations +open import foundation-core.invertible-maps + +open import modal-logic.completeness +open import modal-logic.completeness-K +open import modal-logic.formulas +open import modal-logic.kripke-models-filtrations +open import modal-logic.kripke-semantics +open import modal-logic.logic-syntax +open import modal-logic.modal-logic-K +open import modal-logic.modal-logic-decision +open import modal-logic.soundness +open import modal-logic.weak-deduction + +open import order-theory.maximal-elements-posets + +open import univalent-combinatorics.finite-types +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l1 : Level} (i : Set l1) + where + + is-finitely-approximable-Prop : + {l2 : Level} (l3 l4 l5 l6 : Level) → + modal-theory l2 i → + Prop (l1 ⊔ l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ lsuc l6) + is-finitely-approximable-Prop l3 l4 l5 l6 logic = + ∃-Prop + ( model-class l3 l4 i l5 l6) + ( λ C → + ( product + ( soundness logic (finite-subclass i C)) + ( completeness logic (finite-subclass i C)))) + + is-finitely-approximable : + {l2 : Level} (l3 l4 l5 l6 : Level) → + modal-theory l2 i → + UU (l1 ⊔ l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ lsuc l6) + is-finitely-approximable l3 l4 l5 l6 logic = + type-Prop (is-finitely-approximable-Prop l3 l4 l5 l6 logic) + + module _ + (lem : LEM (lsuc (lsuc l1))) + (zorn : Zorn (lsuc l1) l1 l1) + (prop-resize : propositional-resizing l1 (lsuc l1)) + where + + K-finite-class : + model-class (lsuc (lsuc l1)) (lsuc l1) i (lsuc l1) (lsuc (lsuc (lsuc l1))) + K-finite-class = + filtrate-class i + ( all-models (lsuc l1) l1 i l1) + ( minimal-kripke-model-filtration i) + + K-finite-class-sub-filtration-models : + K-finite-class ⊆ + filtration-models + (lsuc (lsuc l1)) (lsuc l1) i (lsuc l1) l1 (lsuc l1) l1 l1 + K-finite-class-sub-filtration-models M* = + map-universal-property-trunc-Prop + ( filtration-models _ _ i _ _ _ _ _ M*) + ( λ ((a , M , _) , p) → + ( tr (is-in-subtype (filtration-models _ _ i _ _ _ _ _)) (inv p) + ( intro-∃ + ( in-list (subformulas-list i a) , M) + ( pair + (is-finite-subformulas-list + ( i) + ( lower-LEM (lsuc (lsuc l1)) lem) + ( a)) + ( is-kripke-model-filtration-minimal-kripke-model-filtration + ( i) + ( in-list (subformulas-list i a)) + ( M) + ( is-modal-theory-closed-under-subformulas-subformulas-list + ( i) + ( a))))))) + + completeness-K-filtration : + completeness (modal-logic-K i) (K-finite-class) + completeness-K-filtration = + filtrate-completeness i + ( all-models (lsuc l1) l1 i l1) + ( minimal-kripke-model-filtration i) + ( λ (M , _) theory theory-is-closed → + ( is-kripke-model-filtration-minimal-kripke-model-filtration i + ( theory) + ( M) + ( theory-is-closed))) + ( modal-logic-K i) + ( completeness-K i (lower-LEM (lsuc (lsuc l1)) lem) zorn prop-resize) + + soundness-K-filtration : + soundness (modal-logic-K i) (K-finite-class) + soundness-K-filtration = + filtrate-soundness i + ( all-models (lsuc l1) l1 i l1) + ( minimal-kripke-model-filtration i) + ( modal-logic-K i) + ( all-models (lsuc (lsuc l1)) (lsuc l1) i (lsuc l1)) + ( λ _ _ → star) + ( λ a in-logic M _ x → + ( soundness-K i a in-logic M (λ b y → lem ((M , y) ⊨ b)) x)) + + is-finitely-approximable-K : + is-finitely-approximable + ( lsuc (lsuc l1)) + ( lsuc l1) + ( lsuc l1) + ( lsuc (lsuc (lsuc l1))) + ( modal-logic-K i) + is-finitely-approximable-K = + intro-∃ + ( K-finite-class) + ( pair + ( soundness-subclass + ( modal-logic-K i) + ( K-finite-class) + ( finite-subclass i K-finite-class) + ( subtype-intersection-right + ( finite-models (lsuc (lsuc l1)) (lsuc l1) i (lsuc l1)) + ( K-finite-class)) + ( soundness-K-filtration)) + ( transitive-leq-subtype + ( class-modal-logic (finite-subclass i K-finite-class)) + ( class-modal-logic K-finite-class) + ( modal-logic-K i) + ( completeness-K-filtration) + ( λ a a-in-logic M M-in-class → + ( a-in-logic M + ( pair + ( transitive-leq-subtype + ( K-finite-class) + ( filtration-models + ( lsuc (lsuc l1)) + ( lsuc l1) + ( i) + ( lsuc l1) + ( l1) + ( lsuc l1) + ( l1) + ( l1)) + ( finite-models (lsuc (lsuc l1)) (lsuc l1) i (lsuc l1)) + ( filtration-models-subset-finite-models i l1 + (lsuc l1) l1 l1 lem) + ( K-finite-class-sub-filtration-models) + ( M) + ( M-in-class)) + ( M-in-class)))))) +``` diff --git a/src/modal-logic/kripke-models-filtrations-theorem.lagda.md b/src/modal-logic/kripke-models-filtrations-theorem.lagda.md new file mode 100644 index 0000000000..01026a8dbc --- /dev/null +++ b/src/modal-logic/kripke-models-filtrations-theorem.lagda.md @@ -0,0 +1,190 @@ +# Kripke models filtrations theorem + +```agda +module modal-logic.kripke-models-filtrations-theorem where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.binary-relations +open import foundation.coproduct-types +open import foundation.decidable-types +open import foundation.dependent-pair-types +open import foundation.equivalence-classes +open import foundation.equivalences +open import foundation.existential-quantification +open import foundation.function-types +open import foundation.identity-types +open import foundation.inhabited-types +open import foundation.law-of-excluded-middle +open import foundation.logical-equivalences +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.raising-universe-levels +open import foundation.sets +open import foundation.subtypes +open import foundation.transport-along-identifications +open import foundation.unit-type +open import foundation.universe-levels + +open import foundation-core.cartesian-product-types +open import foundation-core.embeddings +open import foundation-core.equivalence-relations +open import foundation-core.invertible-maps + +open import modal-logic.completeness +open import modal-logic.formulas +open import modal-logic.kripke-models-filtrations +open import modal-logic.kripke-semantics +open import modal-logic.logic-syntax +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l1 l2 l3 l4 l5 l6 l7 l8 : Level} (i : Set l3) + (theory : modal-theory l5 i) + (theory-is-closed : is-modal-theory-closed-under-subformulas i theory) + (M : kripke-model l1 l2 i l4) (M* : kripke-model l6 l7 i l8) + where + + kripke-models-filtrations-theorem' : + (is-filtration : is-kripke-model-filtration-Σ i theory M M*) + (a : formula i) → + is-in-subtype theory a → + (x : type-kripke-model i M) → + type-iff-Prop + ( (M , x) ⊨ a) + ( pair + ( M*) + ( map-equiv-is-kripke-model-filtration i theory M M* is-filtration + ( class (Φ-equivalence i theory M) x)) + ⊨ a) + pr1 (kripke-models-filtrations-theorem' is-filtration (var n) in-theory x) f = + map-raise + ( forward-implication + ( is-filtration-valuate-is-kripke-model-filtration + ( i) + ( theory) + ( M) + ( M*) + ( is-filtration) + ( n) + ( x)) + ( in-theory , map-inv-raise f)) + pr1 (kripke-models-filtrations-theorem' is-filtration ⊥ in-theory x) = + map-raise ∘ map-inv-raise + pr1 (kripke-models-filtrations-theorem' is-filtration (a →ₘ b) in-theory x) + fab fa = + forward-implication + ( kripke-models-filtrations-theorem' is-filtration b + ( pr2 (theory-is-closed in-theory)) + ( x)) + ( fab + ( backward-implication + ( kripke-models-filtrations-theorem' is-filtration a + ( pr1 (theory-is-closed in-theory)) + ( x)) + ( fa))) + pr1 (kripke-models-filtrations-theorem' is-filtration (□ a) in-theory x) + f y* r-xy = + apply-universal-property-trunc-Prop + ( is-inhabited-subtype-equivalence-class + ( Φ-equivalence i theory M) + ( map-inv-equiv-is-kripke-model-filtration i theory M M* + ( is-filtration) + ( y*))) + ( (M* , y*) ⊨ a) + ( λ (y , y-in-class) → + ( let y*-id-class + = concat + ( ap + ( map-equiv-is-kripke-model-filtration i theory M M* + ( is-filtration)) + ( eq-class-equivalence-class + ( Φ-equivalence i theory M) + ( map-inv-equiv-is-kripke-model-filtration i theory M + ( M*) + ( is-filtration) + ( y*)) + ( y-in-class))) + ( y*) + ( is-retraction-map-retraction-map-equiv + ( inv-equiv + ( equiv-is-kripke-model-filtration i theory M M* + ( is-filtration))) + ( y*)) + in tr + ( λ z* → type-Prop ((M* , z*) ⊨ a)) + ( y*-id-class) + ( forward-implication + ( kripke-models-filtrations-theorem' is-filtration a + ( theory-is-closed in-theory) + ( y)) + ( filtration-relation-upper-bound-is-kripke-model-filtration + ( i) + ( theory) + ( M) + ( M*) + ( is-filtration) + ( a) + ( in-theory) + ( x) + ( y) + ( tr + ( relation-kripke-model i M* + ( map-equiv-is-kripke-model-filtration i theory M M* + ( is-filtration) + ( class (Φ-equivalence i theory M) x))) + ( inv y*-id-class) + ( r-xy)) + (f))))) + pr2 (kripke-models-filtrations-theorem' is-filtration (var n) in-theory x) f = + map-raise + ( pr2 + ( backward-implication + ( is-filtration-valuate-is-kripke-model-filtration i theory M M* + ( is-filtration) + ( n) + ( x)) + ( map-inv-raise f))) + pr2 (kripke-models-filtrations-theorem' is-filtration ⊥ in-theory x) = + map-raise ∘ map-inv-raise + pr2 (kripke-models-filtrations-theorem' is-filtration (a →ₘ b) in-theory x) + fab fa = + backward-implication + ( kripke-models-filtrations-theorem' is-filtration b + ( pr2 (theory-is-closed in-theory)) + ( x)) + ( fab + ( forward-implication + ( kripke-models-filtrations-theorem' is-filtration a + ( pr1 (theory-is-closed in-theory)) + ( x)) + ( fa))) + pr2 (kripke-models-filtrations-theorem' is-filtration (□ a) in-theory x) + f y r-xy = + backward-implication + ( kripke-models-filtrations-theorem' is-filtration a + ( theory-is-closed in-theory) + ( y)) + ( f + ( map-equiv-is-kripke-model-filtration i theory M M* is-filtration + ( class (Φ-equivalence i theory M) y)) + ( filtration-relation-lower-bound-is-kripke-model-filtration i theory + ( M) + ( M*) + ( is-filtration) + ( x) + ( y) + ( r-xy))) +``` diff --git a/src/modal-logic/kripke-models-filtrations.lagda.md b/src/modal-logic/kripke-models-filtrations.lagda.md new file mode 100644 index 0000000000..1aef6d6dc5 --- /dev/null +++ b/src/modal-logic/kripke-models-filtrations.lagda.md @@ -0,0 +1,829 @@ +# Kripke models filtrations + +```agda +module modal-logic.kripke-models-filtrations where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.binary-relations +open import foundation.coproduct-types +open import foundation.decidable-types +open import foundation.dependent-pair-types +open import foundation.equivalence-classes +open import foundation.equivalences +open import foundation.existential-quantification +open import foundation.function-extensionality +open import foundation.function-types +open import foundation.identity-types +open import foundation.inhabited-types +open import foundation.logical-equivalences +open import foundation.negation +open import foundation.propositional-extensionality +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.raising-universe-levels +open import foundation.sets +open import foundation.subtypes +open import foundation.unit-type +open import foundation.universe-levels + +open import foundation-core.cartesian-product-types +open import foundation-core.embeddings +open import foundation-core.equality-dependent-pair-types +open import foundation-core.equivalence-relations +open import foundation-core.injective-maps +open import foundation-core.transport-along-identifications + +open import modal-logic.axioms +open import modal-logic.formulas +open import modal-logic.kripke-semantics +open import modal-logic.logic-syntax + +open import univalent-combinatorics.finite-types +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +-- TODO: move to new file +module _ + {l1 l2 : Level} {A : UU l1} + where + + data transitive-closure (R : Relation l2 A) : A → A → UU (l1 ⊔ l2) + where + base* : {x y : A} → R x y → transitive-closure R x y + step* : + {x y z : A} → + R x y → + transitive-closure R y z → + transitive-closure R x z + + transitive-closure-Prop : + (R : Relation l2 A) → Relation-Prop (l1 ⊔ l2) A + transitive-closure-Prop R x y = + trunc-Prop (transitive-closure R x y) + + is-transitive-transitive-closure : + (R : Relation l2 A) → is-transitive (transitive-closure R) + is-transitive-transitive-closure R x y z c-yz (base* r-xy) = + step* r-xy c-yz + is-transitive-transitive-closure + R x y z c-yz (step* {y = u} r-xu c-uy) = + step* r-xu (is-transitive-transitive-closure R u y z c-yz c-uy) + + is-transitive-transitive-closure-Prop : + (R : Relation l2 A) → + is-transitive-Relation-Prop (transitive-closure-Prop R) + is-transitive-transitive-closure-Prop R x y z c-yz c-xy = + apply-twice-universal-property-trunc-Prop + ( c-yz) + ( c-xy) + ( transitive-closure-Prop R x z) + ( λ r-yz r-xy → + ( unit-trunc-Prop (is-transitive-transitive-closure R x y z r-yz r-xy))) + + transitive-closure-preserves-reflexivity : + (R : Relation l2 A) → + is-reflexive R → + is-reflexive (transitive-closure R) + transitive-closure-preserves-reflexivity R is-refl x = base* (is-refl x) + + transitive-closure-preserves-symmetry : + (R : Relation l2 A) → + is-symmetric R → + is-symmetric (transitive-closure R) + transitive-closure-preserves-symmetry R is-sym x y (base* r) = + base* (is-sym x y r) + transitive-closure-preserves-symmetry R is-sym x y (step* {y = u} r-xu c-uy) = + is-transitive-transitive-closure R y u x + ( base* (is-sym x u r-xu)) + ( transitive-closure-preserves-symmetry R is-sym u y c-uy) + + transitive-closure-Prop-preserves-reflexivity : + (R : Relation l2 A) → + is-reflexive R → + is-reflexive-Relation-Prop (transitive-closure-Prop R) + transitive-closure-Prop-preserves-reflexivity R is-refl x = + unit-trunc-Prop (transitive-closure-preserves-reflexivity R is-refl x) + + transitive-closure-Prop-preserves-symmetry : + (R : Relation l2 A) → + is-symmetric R → + is-symmetric-Relation-Prop (transitive-closure-Prop R) + transitive-closure-Prop-preserves-symmetry R is-sym x y = + map-universal-property-trunc-Prop + ( transitive-closure-Prop R y x) + ( unit-trunc-Prop ∘ transitive-closure-preserves-symmetry R is-sym x y) + +module _ + {l1 : Level} (i : Set l1) + where + + module _ + {l2 : Level} (theory : modal-theory l2 i) + where + + is-modal-theory-has-subformulas-Prop : formula i → Prop l2 + is-modal-theory-has-subformulas-Prop (var _) = raise-unit-Prop l2 + is-modal-theory-has-subformulas-Prop ⊥ = raise-unit-Prop l2 + is-modal-theory-has-subformulas-Prop (a →ₘ b) = + product-Prop (theory a) (theory b) + is-modal-theory-has-subformulas-Prop (□ a) = theory a + + is-modal-theory-has-subformulas : formula i → UU l2 + is-modal-theory-has-subformulas = + type-Prop ∘ is-modal-theory-has-subformulas-Prop + + is-modal-theory-closed-under-subformulas-Prop : Prop (l1 ⊔ l2) + is-modal-theory-closed-under-subformulas-Prop = + implicit-Π-Prop + ( formula i) + ( λ a → + ( function-Prop + ( is-in-subtype theory a) + ( is-modal-theory-has-subformulas-Prop a))) + + is-modal-theory-closed-under-subformulas : UU (l1 ⊔ l2) + is-modal-theory-closed-under-subformulas = + type-Prop (is-modal-theory-closed-under-subformulas-Prop) + + is-modal-theory-closed-under-subformulas-condition : + ( {a b : formula i} → + is-in-subtype theory (a →ₘ b) → + is-in-subtype theory a × is-in-subtype theory b) → + ( {a : formula i} → is-in-subtype theory (□ a) → is-in-subtype theory a) → + is-modal-theory-closed-under-subformulas + is-modal-theory-closed-under-subformulas-condition + h-impl h-box {var n} _ = raise-star + is-modal-theory-closed-under-subformulas-condition + h-impl h-box {⊥} _ = raise-star + is-modal-theory-closed-under-subformulas-condition + h-impl h-box {a →ₘ b} = h-impl + is-modal-theory-closed-under-subformulas-condition + h-impl h-box {□ a} = h-box + +module _ + {l1 l2 l3 l4 l5 : Level} (i : Set l3) + (theory : modal-theory l5 i) + (M : kripke-model l1 l2 i l4) + where + + Φ-equivalence : + equivalence-relation (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) (type-kripke-model i M) + pr1 Φ-equivalence x y = + Π-Prop + ( formula i) + ( λ a → + ( function-Prop + ( is-in-subtype theory a) + ( iff-Prop ((M , x) ⊨ a) ((M , y) ⊨ a)))) + pr1 (pr2 Φ-equivalence) x a in-theory = id , id + pr1 (pr2 (pr2 Φ-equivalence)) x y r a in-theory = + inv-iff (r a in-theory) + pr2 (pr2 (pr2 Φ-equivalence)) x y z r-xy r-yz a in-theory = + r-xy a in-theory ∘iff r-yz a in-theory + + valuate-function-equivalence-class : + equivalence-class Φ-equivalence → + UU (lsuc l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ l5) + valuate-function-equivalence-class class = + Σ ( type-subtype theory → Prop (l1 ⊔ l2 ⊔ l4)) + ( λ f → + ( (a , in-theory) : type-subtype theory) + ( (x , _) : + type-subtype (subtype-equivalence-class Φ-equivalence class)) → + f (a , in-theory) ⇔ ((M , x) ⊨ a)) + + is-prop-valuate-function-equivalence-class : + (class : equivalence-class Φ-equivalence) → + is-prop (valuate-function-equivalence-class class) + is-prop-valuate-function-equivalence-class class = + apply-universal-property-trunc-Prop + ( is-inhabited-subtype-equivalence-class Φ-equivalence class) + ( is-prop _ , is-prop-is-prop _) + ( λ x → + ( is-prop-all-elements-equal + ( λ (f , f-val) (g , g-val) → + ( eq-pair-Σ + ( eq-htpy + ( λ a → + ( eq-iff + ( λ fa → + ( backward-implication + ( g-val a x) + ( forward-implication (f-val a x) fa))) + ( λ ga → + ( backward-implication + ( f-val a x) + ( forward-implication (g-val a x) ga)))))) + ( eq-is-prop + ( is-prop-Π + ( λ (a , in-theory) → + ( is-prop-Π + ( λ (x , _) → + ( is-prop-iff-Prop + ( g (a , in-theory)) + ( (M , x) ⊨ a))))))))))) + + function-equivalence-class : + (class : equivalence-class Φ-equivalence) → + valuate-function-equivalence-class class + function-equivalence-class class = + apply-universal-property-trunc-Prop + ( is-inhabited-subtype-equivalence-class Φ-equivalence class) + ( pair + ( valuate-function-equivalence-class class) + ( is-prop-valuate-function-equivalence-class class)) + ( λ (x , x-in-class) → + ( pair + ( λ (a , _) → (M , x) ⊨ a) + ( λ (a , in-theory) (y , y-in-class) → + ( apply-effectiveness-class Φ-equivalence {x} {y} + ( concat + ( eq-effective-quotient' Φ-equivalence x class x-in-class) + ( _) + ( inv + ( eq-effective-quotient' Φ-equivalence y class y-in-class))) + ( a) + ( in-theory))))) + + map-valuate-function-equivalence-class : + (class : equivalence-class Φ-equivalence) → + valuate-function-equivalence-class class → + type-subtype theory → Prop (l1 ⊔ l2 ⊔ l4) + map-valuate-function-equivalence-class _ = pr1 + + map-function-equivalence-class : + equivalence-class Φ-equivalence → + type-subtype theory → Prop (l1 ⊔ l2 ⊔ l4) + map-function-equivalence-class class = + map-valuate-function-equivalence-class + ( class) + ( function-equivalence-class class) + + is-injective-map-function-equivalence-class : + is-injective map-function-equivalence-class + is-injective-map-function-equivalence-class {x-class} {y-class} p = + let (f , f-val) = function-equivalence-class x-class + (g , g-val) = function-equivalence-class y-class + in apply-twice-universal-property-trunc-Prop + ( is-inhabited-subtype-equivalence-class Φ-equivalence x-class) + ( is-inhabited-subtype-equivalence-class Φ-equivalence y-class) + ( pair + ( x-class = y-class) + ( is-set-equivalence-class Φ-equivalence x-class y-class)) + ( λ (x , x-in-class) (y , y-in-class) → + ( equational-reasoning + x-class = class Φ-equivalence x by + ( inv + ( eq-class-equivalence-class + ( Φ-equivalence) + ( x-class) + ( x-in-class))) + = class Φ-equivalence y by + ( apply-effectiveness-class' + ( Φ-equivalence) + ( λ a a-in-theory → + ( g-val (a , a-in-theory) (y , y-in-class) + ∘iff iff-eq (htpy-eq p (a , a-in-theory)) + ∘iff inv-iff + ( f-val + ( a , a-in-theory) + ( x , x-in-class))))) + = y-class by + ( eq-class-equivalence-class + ( Φ-equivalence) + ( y-class) + ( y-in-class)))) + + injection-map-function-equivalence-class : + injection + ( equivalence-class Φ-equivalence) + ( type-subtype theory → Prop (l1 ⊔ l2 ⊔ l4)) + pr1 injection-map-function-equivalence-class = + map-function-equivalence-class + pr2 injection-map-function-equivalence-class = + is-injective-map-function-equivalence-class + + is-bounded-valuate-Prop : + {l6 l7 : Level} {A : UU l6} → + (type-Set i → A → Prop l7) → + Prop (l3 ⊔ l5 ⊔ l6 ⊔ l7) + is-bounded-valuate-Prop {A = A} V = + Π-Prop (type-Set i) + ( λ n → + ( function-Prop + ( ¬ (is-in-subtype theory (var n))) + ( Π-Prop A + ( λ x → + ( neg-Prop (V n x)))))) + + -- is-bounded-valuate : + -- {l6 l7 : Level} {A : UU l6} → + -- (type-Set i → A → Prop l7) → + -- UU (l3 ⊔ l5 ⊔ l6 ⊔ l7) + -- is-bounded-valuate = type-Prop ∘ is-bounded-valuate-Prop + + module _ + {l6 l7 l8 : Level} (M* : kripke-model l6 l7 i l8) + where + + is-filtration-valuate-Prop : + equivalence-class Φ-equivalence ≃ type-kripke-model i M* → + Prop (l1 ⊔ l3 ⊔ l4 ⊔ l5 ⊔ l8) + is-filtration-valuate-Prop e = + Π-Prop (type-Set i) + ( λ n → + ( Π-Prop (type-kripke-model i M) + ( λ x → iff-Prop + ( product-Prop (theory (var n)) (valuate-kripke-model i M n x)) + ( valuate-kripke-model i M* n + ( map-equiv e (class Φ-equivalence x)))))) + + is-filtration-valuate : + equivalence-class Φ-equivalence ≃ type-kripke-model i M* → + UU (l1 ⊔ l3 ⊔ l4 ⊔ l5 ⊔ l8) + is-filtration-valuate = type-Prop ∘ is-filtration-valuate-Prop + + filtration-relation-lower-bound-Prop : + equivalence-class Φ-equivalence ≃ type-kripke-model i M* → + Prop (l1 ⊔ l2 ⊔ l7) + filtration-relation-lower-bound-Prop e = + Π-Prop (type-kripke-model i M) + ( λ x → + ( Π-Prop (type-kripke-model i M) + ( λ y → + ( function-Prop (relation-kripke-model i M x y) + ( relation-Prop-kripke-model i M* + ( map-equiv e (class Φ-equivalence x)) + ( map-equiv e (class Φ-equivalence y))))))) + + filtration-relation-lower-bound : + equivalence-class Φ-equivalence ≃ type-kripke-model i M* → + UU (l1 ⊔ l2 ⊔ l7) + filtration-relation-lower-bound = + type-Prop ∘ filtration-relation-lower-bound-Prop + + filtration-relation-upper-bound-Prop : + equivalence-class Φ-equivalence ≃ type-kripke-model i M* → + Prop (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5 ⊔ l7) + filtration-relation-upper-bound-Prop e = + Π-Prop (formula i) + ( λ a → + ( function-Prop (is-in-subtype theory (□ a)) + ( Π-Prop (type-kripke-model i M) + ( λ x → + ( Π-Prop (type-kripke-model i M) + ( λ y → + ( function-Prop + ( relation-kripke-model i M* + ( map-equiv e (class Φ-equivalence x)) + ( map-equiv e (class Φ-equivalence y))) + ( hom-Prop ((M , x) ⊨ □ a) + ( (M , y) ⊨ a))))))))) + + filtration-relation-upper-bound : + equivalence-class Φ-equivalence ≃ type-kripke-model i M* → + UU (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5 ⊔ l7) + filtration-relation-upper-bound = + type-Prop ∘ filtration-relation-upper-bound-Prop + + -- is-alpha-Σ : + -- UU (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6 ⊔ l8) + -- is-alpha-Σ = + -- product + -- ( equivalence-class Φ-equivalence ≃ type-kripke-model i M*) + -- ( is-bounded-valuate (valuate-kripke-model i M*)) + + -- is-alpha-Prop : + -- Prop (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6 ⊔ l8) + -- is-alpha-Prop = trunc-Prop is-alpha-Σ + + is-kripke-model-filtration-Σ : + UU (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6 ⊔ l7 ⊔ l8) + is-kripke-model-filtration-Σ = + Σ ( equivalence-class Φ-equivalence ≃ type-kripke-model i M*) + ( λ e → + ( product + ( is-filtration-valuate e) + ( product + ( filtration-relation-lower-bound e) + ( filtration-relation-upper-bound e)))) + + is-kripke-model-filtration-Prop : + Prop (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6 ⊔ l7 ⊔ l8) + is-kripke-model-filtration-Prop = trunc-Prop is-kripke-model-filtration-Σ + + is-kripke-model-filtration : + UU (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6 ⊔ l7 ⊔ l8) + is-kripke-model-filtration = type-Prop is-kripke-model-filtration-Prop + + equiv-is-kripke-model-filtration : + is-kripke-model-filtration-Σ → + equivalence-class Φ-equivalence ≃ type-kripke-model i M* + equiv-is-kripke-model-filtration = pr1 + + map-equiv-is-kripke-model-filtration : + is-kripke-model-filtration-Σ → + equivalence-class Φ-equivalence → type-kripke-model i M* + map-equiv-is-kripke-model-filtration = + map-equiv ∘ equiv-is-kripke-model-filtration + + map-inv-equiv-is-kripke-model-filtration : + is-kripke-model-filtration-Σ → + type-kripke-model i M* → equivalence-class Φ-equivalence + map-inv-equiv-is-kripke-model-filtration = + map-inv-equiv ∘ equiv-is-kripke-model-filtration + + is-filtration-valuate-is-kripke-model-filtration : + (e : is-kripke-model-filtration-Σ) → + is-filtration-valuate (equiv-is-kripke-model-filtration e) + is-filtration-valuate-is-kripke-model-filtration = pr1 ∘ pr2 + + filtration-relation-lower-bound-is-kripke-model-filtration : + (e : is-kripke-model-filtration-Σ) → + filtration-relation-lower-bound (equiv-is-kripke-model-filtration e) + filtration-relation-lower-bound-is-kripke-model-filtration = + pr1 ∘ pr2 ∘ pr2 + + filtration-relation-upper-bound-is-kripke-model-filtration : + (e : is-kripke-model-filtration-Σ) → + filtration-relation-upper-bound (equiv-is-kripke-model-filtration e) + filtration-relation-upper-bound-is-kripke-model-filtration = + pr2 ∘ pr2 ∘ pr2 + + class-x-eq-x*' : + (e : equivalence-class Φ-equivalence ≃ type-kripke-model i M*) → + (x : type-kripke-model i M) + (x* : type-kripke-model i M*) → + is-in-equivalence-class Φ-equivalence (map-inv-equiv e x*) x → + map-equiv e (class Φ-equivalence x) = x* + class-x-eq-x*' e x x* x-in-x* = + concat + ( ap + ( map-equiv e) + ( eq-class-equivalence-class Φ-equivalence + ( map-inv-equiv e x*) + ( x-in-x*))) + ( x*) + ( is-section-map-section-map-equiv e x*) + + class-x-eq-x* : + (is-filt : is-kripke-model-filtration-Σ) + (x : type-kripke-model i M) + (x* : type-kripke-model i M*) → + is-in-equivalence-class Φ-equivalence + ( map-inv-equiv-is-kripke-model-filtration is-filt x*) x → + map-equiv-is-kripke-model-filtration is-filt (class Φ-equivalence x) = x* + class-x-eq-x* = class-x-eq-x*' ∘ equiv-is-kripke-model-filtration + + filtration-relation-preserves-reflexivity : + (e : equivalence-class Φ-equivalence ≃ type-kripke-model i M*) → + type-Prop (filtration-relation-lower-bound-Prop e) → + is-in-subtype (reflexive-kripke-class l1 l2 i l4) M → + is-in-subtype (reflexive-kripke-class l6 l7 i l8) M* + filtration-relation-preserves-reflexivity e bound is-refl x* = + apply-universal-property-trunc-Prop + ( is-inhabited-subtype-equivalence-class Φ-equivalence + ( map-inv-equiv e x*)) + ( relation-Prop-kripke-model i M* x* x*) + ( λ (x , x-in-x*) → + ( tr + ( λ y → relation-kripke-model i M* y y) + ( class-x-eq-x*' e x x* x-in-x*) + ( bound x x (is-refl x)))) + + filtration-preserves-reflexivity : + is-kripke-model-filtration → + is-in-subtype (reflexive-kripke-class l1 l2 i l4) M → + is-in-subtype (reflexive-kripke-class l6 l7 i l8) M* + filtration-preserves-reflexivity t-is-filt is-refl class = + apply-universal-property-trunc-Prop + ( t-is-filt) + ( relation-Prop-kripke-model i M* class class) + ( λ is-filt → + ( apply-universal-property-trunc-Prop + ( is-inhabited-subtype-equivalence-class Φ-equivalence + ( map-inv-equiv-is-kripke-model-filtration is-filt class)) + ( relation-Prop-kripke-model i M* class class) + ( λ (x , in-class) → + ( tr + ( λ y → relation-kripke-model i M* y y) + ( class-x-eq-x* is-filt x class in-class) + ( filtration-relation-lower-bound-is-kripke-model-filtration + ( is-filt) + ( x) + ( x) + ( is-refl x)))))) + + is-inhabited-equivalence-classes : + is-inhabited (equivalence-class Φ-equivalence) + is-inhabited-equivalence-classes = + map-is-inhabited + ( class Φ-equivalence) + ( is-inhabited-type-kripke-model i M) + + minimal-kripke-model-filtration-relation : + equivalence-class Φ-equivalence → + equivalence-class Φ-equivalence → + Prop (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) + minimal-kripke-model-filtration-relation x* y* = + ∃-Prop + ( type-kripke-model i M × type-kripke-model i M) + ( λ (x , y) → + ( product + ( relation-kripke-model i M x y) + ( product + ( is-in-equivalence-class Φ-equivalence x* x) + ( is-in-equivalence-class Φ-equivalence y* y)))) + + minimal-kripke-model-filtration-valuate : + type-Set i → + equivalence-class Φ-equivalence → + Prop (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) + minimal-kripke-model-filtration-valuate n x* = + product-Prop + ( theory (var n)) + ( Π-Prop + ( type-kripke-model i M) + ( λ x → + ( function-Prop + ( is-in-equivalence-class Φ-equivalence x* x) + ( valuate-kripke-model i M n x)))) + + minimal-kripke-model-filtration : + kripke-model + ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) + ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) + ( i) + ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) + minimal-kripke-model-filtration = + pair + ( pair + ( equivalence-class Φ-equivalence , is-inhabited-equivalence-classes) + ( minimal-kripke-model-filtration-relation)) + ( minimal-kripke-model-filtration-valuate) + + minimal-transitive-kripke-model-filtration : + kripke-model + ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) + ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) + ( i) + ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) + minimal-transitive-kripke-model-filtration = + pair + ( pair + ( equivalence-class Φ-equivalence , is-inhabited-equivalence-classes) + ( transitive-closure-Prop + ( type-Relation-Prop minimal-kripke-model-filtration-relation))) + ( minimal-kripke-model-filtration-valuate) + + module _ + (theory-is-closed : is-modal-theory-closed-under-subformulas i theory) + where + + proof-upper-bound : + (a : formula i) → + is-in-subtype theory (□ a) → + (x y : type-kripke-model i M) → + relation-kripke-model i minimal-kripke-model-filtration + ( class Φ-equivalence x) + ( class Φ-equivalence y) → + type-Prop ((M , x) ⊨ □ a) → + type-Prop ((M , y) ⊨ a) + proof-upper-bound a box-in-theory x y r-xy x-forces-box = + apply-universal-property-trunc-Prop + ( r-xy) + ( (M , y) ⊨ a) + ( λ ((x' , y') , r-xy' , iff-x , iff-y) → + ( backward-implication + ( iff-y a (theory-is-closed box-in-theory)) + ( forward-implication + ( iff-x (□ a) box-in-theory) + ( x-forces-box) + ( y') + ( r-xy')))) + + proof-lower-bound : + (x y : type-kripke-model i M) → + relation-kripke-model i M x y → + type-Prop + ( minimal-kripke-model-filtration-relation + ( class Φ-equivalence x) + ( class Φ-equivalence y)) + proof-lower-bound x y r = + intro-∃ (x , y) (r , (λ _ _ → id , id) , (λ _ _ → id , id)) + + is-kripke-model-filtration-minimal-kripke-model-filtration : + is-kripke-model-filtration minimal-kripke-model-filtration + is-kripke-model-filtration-minimal-kripke-model-filtration = + intro-∃ + ( id-equiv) + ( triple + ( λ n x → + ( pair + ( λ (in-theory , val-n) → + ( pair in-theory + ( λ y eq-xy → + ( map-inv-raise + ( forward-implication + ( eq-xy (var n) in-theory) + ( map-raise val-n))))))) + ( λ (in-theory , val-n) → in-theory , val-n x (λ _ _ → id , id))) + ( proof-lower-bound) + ( proof-upper-bound)) + + helper : + is-in-subtype (transitivity-kripke-class l1 l2 i l4) M → + (a : formula i) → + is-in-subtype theory (□ a) → + (x y : type-kripke-model i M) → + transitive-closure + ( relation-kripke-model i minimal-kripke-model-filtration) + ( class Φ-equivalence x) + ( class Φ-equivalence y) → + type-Prop ((M , x) ⊨ □ a) → + type-Prop ((M , y) ⊨ a) + helper M-is-trans a box-in-theory x y (base* r-xy) x-forces-box = + proof-upper-bound a box-in-theory x y r-xy x-forces-box + helper M-is-trans a box-in-theory x y (step* {y = z*} r-xz c-zy) + x-forces-box = + apply-universal-property-trunc-Prop + ( is-inhabited-subtype-equivalence-class Φ-equivalence z*) + ( (M , y) ⊨ a) + ( λ (z , z-in-z*) → + aux z (eq-class-equivalence-class Φ-equivalence z* z-in-z*)) + where + aux : + (z : type-kripke-model i M) → + class Φ-equivalence z = z* → + type-Prop ((M , y) ⊨ a) + aux z refl = + apply-universal-property-trunc-Prop + ( r-xz) + ( (M , y) ⊨ a) + ( λ ((x' , z') , r-xz' , iff-x , iff-z) → + ( helper M-is-trans a box-in-theory z y c-zy + ( backward-implication + ( iff-z (□ a) box-in-theory) + ( ax-4-soundness i l2 l4 (□ a →ₘ □ □ a) (a , refl) M M-is-trans + ( x') + ( forward-implication + ( iff-x (□ a) box-in-theory) + ( x-forces-box)) + ( z') + ( r-xz'))))) + + filtration-relation-upper-bound-Prop-minimal-transitive-kripke-model-filtration : + is-in-subtype (transitivity-kripke-class l1 l2 i l4) M → + filtration-relation-upper-bound + minimal-transitive-kripke-model-filtration id-equiv + filtration-relation-upper-bound-Prop-minimal-transitive-kripke-model-filtration + M-is-trans a box-in-theory x y r-xy x-forces-box = + apply-universal-property-trunc-Prop + ( r-xy) + ( (M , y) ⊨ a) + ( λ r → helper M-is-trans a box-in-theory x y r x-forces-box) + + is-kripke-model-filtration-minimal-transitive-kripke-model-filtration : + is-in-subtype (transitivity-kripke-class l1 l2 i l4) M → + is-kripke-model-filtration minimal-transitive-kripke-model-filtration + is-kripke-model-filtration-minimal-transitive-kripke-model-filtration + M-is-trans = + intro-∃ + ( id-equiv) + ( triple + ( λ n x → + ( pair + ( λ (in-theory , val-n) → + ( pair in-theory + ( λ y eq-xy → + ( map-inv-raise + ( forward-implication + ( eq-xy (var n) in-theory) + ( map-raise val-n))))))) + ( λ (in-theory , val-n) → in-theory , val-n x (λ _ _ → id , id))) + ( λ x y r → unit-trunc-Prop (base* (proof-lower-bound x y r))) + ( filtration-relation-upper-bound-Prop-minimal-transitive-kripke-model-filtration + ( M-is-trans))) + + minimal-filtration-preserves-reflexivity : + is-in-subtype (reflexive-kripke-class l1 l2 i l4) M → + is-in-subtype + ( reflexive-kripke-class + ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) + ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) + ( i) + ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( minimal-kripke-model-filtration) + minimal-filtration-preserves-reflexivity = + filtration-preserves-reflexivity + ( minimal-kripke-model-filtration) + ( is-kripke-model-filtration-minimal-kripke-model-filtration) + + minimal-kripke-model-filtration-relation-preserves-symmetry : + is-in-subtype (symmetry-kripke-class l1 l2 i l4) M → + is-symmetric-Relation-Prop minimal-kripke-model-filtration-relation + minimal-kripke-model-filtration-relation-preserves-symmetry is-sym x* y* = + map-universal-property-trunc-Prop + ( relation-Prop-kripke-model i minimal-kripke-model-filtration y* x*) + ( λ ((x , y) , r-xy , x-in-x* , y-in-y*) → + ( intro-∃ (y , x) (is-sym x y r-xy , y-in-y* , x-in-x*))) + + minimal-filtration-preserves-symmetry : + is-in-subtype (symmetry-kripke-class l1 l2 i l4) M → + is-in-subtype + ( symmetry-kripke-class + ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) + ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) + ( i) + ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( minimal-kripke-model-filtration) + minimal-filtration-preserves-symmetry = + minimal-kripke-model-filtration-relation-preserves-symmetry + + minimal-transitive-filtration-preserves-reflexivity : + is-in-subtype (reflexive-kripke-class l1 l2 i l4) M → + is-in-subtype + ( reflexive-kripke-class + ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) + ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) + ( i) + ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( minimal-transitive-kripke-model-filtration) + minimal-transitive-filtration-preserves-reflexivity is-refl = + transitive-closure-Prop-preserves-reflexivity + ( type-Relation-Prop minimal-kripke-model-filtration-relation) + ( minimal-filtration-preserves-reflexivity is-refl) + + minimal-transitive-kripke-model-filtration-preserves-symmetry : + is-in-subtype (symmetry-kripke-class l1 l2 i l4) M → + is-in-subtype + ( symmetry-kripke-class + ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) + ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) + ( i) + ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( minimal-transitive-kripke-model-filtration) + minimal-transitive-kripke-model-filtration-preserves-symmetry is-sym = + transitive-closure-Prop-preserves-symmetry + ( type-Relation-Prop minimal-kripke-model-filtration-relation) + ( minimal-filtration-preserves-symmetry is-sym) + + minimal-transitive-kripke-model-filtration-is-transitive : + is-in-subtype + ( transitivity-kripke-class + ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) + ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) + ( i) + ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( minimal-transitive-kripke-model-filtration) + minimal-transitive-kripke-model-filtration-is-transitive = + is-transitive-transitive-closure-Prop + ( type-Relation-Prop minimal-kripke-model-filtration-relation) + + minimal-transitive-filtration-preserves-equivalence : + is-in-subtype (equivalence-kripke-class l1 l2 i l4) M → + is-in-subtype + ( equivalence-kripke-class + ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) + ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) + ( i) + ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( minimal-transitive-kripke-model-filtration) + minimal-transitive-filtration-preserves-equivalence + (is-refl , is-sym , _) = + triple + ( minimal-transitive-filtration-preserves-reflexivity is-refl) + ( minimal-transitive-kripke-model-filtration-preserves-symmetry + ( is-sym)) + ( minimal-transitive-kripke-model-filtration-is-transitive) + +module _ + (l1 l2 : Level) + {l3 : Level} (i : Set l3) + (l4 : Level) + (l5 l6 l7 l8 : Level) + where + + filtration-models : + model-class l1 l2 i l4 + (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4 ⊔ lsuc l5 ⊔ lsuc l6 ⊔ lsuc l7 ⊔ lsuc l8) + filtration-models M* = + ∃-Prop + ( modal-theory l5 i × kripke-model l6 l7 i l8) + ( λ (theory , M) → + ( product + ( is-finite (type-subtype theory)) + ( is-kripke-model-filtration i theory M M*))) +``` diff --git a/src/modal-logic/kripke-semantics.lagda.md b/src/modal-logic/kripke-semantics.lagda.md index 083992938f..74dd41d71d 100644 --- a/src/modal-logic/kripke-semantics.lagda.md +++ b/src/modal-logic/kripke-semantics.lagda.md @@ -66,12 +66,16 @@ module _ postulate is-set-kripke-frame : is-set (kripke-frame l1 l2) + Inhabited-Type-kripke-frame : kripke-frame l1 l2 → Inhabited-Type l1 + Inhabited-Type-kripke-frame = pr1 + type-kripke-frame : kripke-frame l1 l2 → UU l1 - type-kripke-frame = type-Inhabited-Type ∘ pr1 + type-kripke-frame = type-Inhabited-Type ∘ Inhabited-Type-kripke-frame is-inhabited-type-kripke-frame : (F : kripke-frame l1 l2) → is-inhabited (type-kripke-frame F) - is-inhabited-type-kripke-frame = is-inhabited-type-Inhabited-Type ∘ pr1 + is-inhabited-type-kripke-frame = + is-inhabited-type-Inhabited-Type ∘ Inhabited-Type-kripke-frame relation-Prop-kripke-frame : (F : kripke-frame l1 l2) → Relation-Prop l2 (type-kripke-frame F) @@ -96,12 +100,17 @@ module _ kripke-frame-kripke-model : kripke-model l1 l2 i l4 → kripke-frame l1 l2 kripke-frame-kripke-model = pr1 + Inhabited-Type-kripke-model : kripke-model l1 l2 i l4 → Inhabited-Type l1 + Inhabited-Type-kripke-model = + Inhabited-Type-kripke-frame ∘ kripke-frame-kripke-model + type-kripke-model : kripke-model l1 l2 i l4 → UU l1 type-kripke-model = type-kripke-frame ∘ kripke-frame-kripke-model - valuate-kripke-model : - (M : kripke-model l1 l2 i l4) → type-Set i → type-kripke-model M → Prop l4 - valuate-kripke-model = pr2 + is-inhabited-type-kripke-model : + (M : kripke-model l1 l2 i l4) → is-inhabited (type-kripke-model M) + is-inhabited-type-kripke-model = + is-inhabited-type-kripke-frame ∘ kripke-frame-kripke-model relation-Prop-kripke-model : (M : kripke-model l1 l2 i l4) → Relation-Prop l2 (type-kripke-model M) @@ -113,6 +122,10 @@ module _ relation-kripke-model = relation-kripke-frame ∘ kripke-frame-kripke-model + valuate-kripke-model : + (M : kripke-model l1 l2 i l4) → type-Set i → type-kripke-model M → Prop l4 + valuate-kripke-model = pr2 + module _ (l1 l2 : Level) {l3 : Level} (i : Set l3) (l4 : Level) where @@ -317,9 +330,7 @@ module _ ( type-kripke-model i M) ( λ x → is-decidable-Prop ((M , x) ⊨ a)))) - -- decidable-subclass : - -- {l5 : Level} → - + -- TODO: maybe dicidable-finite? finite-models : model-class l1 l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4) finite-models M = product-Prop @@ -338,8 +349,12 @@ module _ ( type-Set i) ( λ n → is-decidable-Prop (valuate-kripke-model i M n x)))))) +module _ + {l1 l2 l3 l4 : Level} (i : Set l3) + where + finite-models-subclass-decidable-models : - finite-models ⊆ decidable-models + finite-models l1 l2 i l4 ⊆ decidable-models l1 l2 i l4 finite-models-subclass-decidable-models M (w-is-fin , dec-r , dec-v) = lemma where lemma : @@ -356,9 +371,15 @@ module _ ( w-is-fin) ( λ y → is-decidable-function-type (dec-r x y) (lemma a y)) -module _ - {l1 l2 l3 l4 : Level} (i : Set l3) - where + is-finite-model-valuate-decidable-models : + (M : kripke-model l1 l2 i l4) → + is-in-subtype (finite-models l1 l2 i l4) M → + (a : formula i) → + is-decidable (type-Prop (M ⊨M a)) + is-finite-model-valuate-decidable-models M sub-fin a = + is-decidable-Π-is-finite + ( pr1 (sub-fin)) + ( finite-models-subclass-decidable-models M sub-fin a) decidable-subclass : {l5 : Level} → diff --git a/src/modal-logic/modal-logic-K.lagda.md b/src/modal-logic/modal-logic-K.lagda.md index 29ae9f51a1..d176c1e385 100644 --- a/src/modal-logic/modal-logic-K.lagda.md +++ b/src/modal-logic/modal-logic-K.lagda.md @@ -165,7 +165,7 @@ module _ ( modal-logic-K i) ( decidable-models l2 l3 i l4) ( finite-models l2 l3 i l4) - ( finite-models-subclass-decidable-models l2 l3 i l4) + ( finite-models-subclass-decidable-models i) ( soundness-K) module _ diff --git a/src/modal-logic/modal-logic-S5.lagda.md b/src/modal-logic/modal-logic-S5.lagda.md index 4c1d217121..be025827a3 100644 --- a/src/modal-logic/modal-logic-S5.lagda.md +++ b/src/modal-logic/modal-logic-S5.lagda.md @@ -191,7 +191,7 @@ module _ ( finite-subclass i (equivalence-kripke-class l2 l3 i l4)) ( finite-models l2 l3 i l4) ( decidable-models l2 l3 i l4) - ( finite-models-subclass-decidable-models l2 l3 i l4) + ( finite-models-subclass-decidable-models i) ( subtype-intersection-left ( finite-models l2 l3 i l4) ( equivalence-kripke-class l2 l3 i l4))) diff --git a/src/modal-logic/modal-logic-decision.lagda.md b/src/modal-logic/modal-logic-decision.lagda.md new file mode 100644 index 0000000000..f21c46f372 --- /dev/null +++ b/src/modal-logic/modal-logic-decision.lagda.md @@ -0,0 +1,720 @@ +# Modal logic decision + +```agda +module modal-logic.modal-logic-decision where +``` + +
Imports + +```agda +open import elementary-number-theory.natural-numbers + +open import foundation.action-on-identifications-binary-functions +open import foundation.action-on-identifications-dependent-functions +open import foundation.action-on-identifications-functions +open import foundation.booleans +open import foundation.contractible-types +open import foundation.decidable-types +open import foundation.dependent-pair-types +open import foundation.embeddings +open import foundation.empty-types +open import foundation.equivalence-classes +open import foundation.equivalences +open import foundation.existential-quantification +open import foundation.function-extensionality +open import foundation.inhabited-types +open import foundation.law-of-excluded-middle +open import foundation.logical-equivalences +open import foundation.propositional-extensionality +open import foundation.propositional-truncations +open import foundation.raising-universe-levels +open import foundation.sets +open import foundation.subtypes +open import foundation.surjective-maps +open import foundation.transport-along-identifications +open import foundation.unions-subtypes +open import foundation.unit-type +open import foundation.universe-levels + +open import foundation-core.cartesian-product-types +open import foundation-core.coproduct-types +open import foundation-core.dependent-identifications +open import foundation-core.equality-dependent-pair-types +open import foundation-core.fibers-of-maps +open import foundation-core.function-types +open import foundation-core.identity-types +open import foundation-core.injective-maps +open import foundation-core.propositions + +open import lists.arrays +open import lists.concatenation-lists +open import lists.lists +open import lists.reversing-lists + +open import modal-logic.completeness +open import modal-logic.formulas +open import modal-logic.kripke-models-filtrations +open import modal-logic.kripke-models-filtrations-theorem +open import modal-logic.kripke-semantics +open import modal-logic.logic-syntax +open import modal-logic.soundness +open import modal-logic.weak-deduction + +open import univalent-combinatorics.counting +open import univalent-combinatorics.decidable-dependent-function-types +open import univalent-combinatorics.finite-types +open import univalent-combinatorics.function-types +open import univalent-combinatorics.kuratowsky-finite-sets +open import univalent-combinatorics.small-types +open import univalent-combinatorics.standard-finite-types +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l1 l2 l3 l4 l5 : Level} + (i : Set l3) + (theory : modal-theory l2 i) + (C : model-class l1 l2 i l4 l5) + (C-sub-fin : C ⊆ finite-models l1 l2 i l4) + (C-is-fin : is-finite (type-subtype C)) + where + + decision-procedure : + (a : formula i) → + is-decidable + ( (M : type-subtype C) → type-Prop (inclusion-subtype C M ⊨M a)) + decision-procedure a = + is-decidable-Π-is-finite + ( C-is-fin) + ( λ (M , M-in-C) → + ( is-finite-model-valuate-decidable-models i M (C-sub-fin M M-in-C) a)) + + decision-procedure' : (a : formula i) → bool + decision-procedure' a with decision-procedure a + ... | inl _ = true + ... | inr _ = false + + decision-procedure-correctness : + soundness theory C → + completeness theory C → + (a : formula i) → + theory a ⇔ prop-bool (decision-procedure' a) + pr1 (decision-procedure-correctness sound complete a) in-theory + with decision-procedure a + ... | inl _ = star + ... | inr not-valid-in-C = + not-valid-in-C (λ (M , M-in-C) x → sound a in-theory M M-in-C x) + pr2 (decision-procedure-correctness sound complete a) _ + with decision-procedure a + ... | inl valid-in-C = complete a (λ M M-in-C → valid-in-C (M , M-in-C)) + +module _ + {l : Level} {i : Set l} + where + + collect-vars : formula i → list (type-Set i) + collect-vars (var x) = cons x nil + collect-vars ⊥ = nil + collect-vars (a →ₘ b) = concat-list (collect-vars a) (collect-vars b) + collect-vars (□ a) = collect-vars a + +module _ + {l1 l2 l3 l4 l5 : Level} (i : Set l3) + (F : kripke-frame l1 l2) + where + + valuates-equals-on-vars : + {l6 : Level} → + subtype l6 (type-Set i) → + (type-Set i → type-kripke-frame F → Prop l4) → + (type-Set i → type-kripke-frame F → Prop l5) → + UU (l1 ⊔ l3 ⊔ l4 ⊔ l5 ⊔ l6) + valuates-equals-on-vars vars V V' = + ((a , _) : type-subtype vars) → (x : type-kripke-frame F) → V a x ⇔ V' a x + + valuates-equals-on-vars-subtype : + {l6 l7 : Level} + (vars₁ : subtype l6 (type-Set i)) + (vars₂ : subtype l7 (type-Set i)) → + vars₂ ⊆ vars₁ → + (V : type-Set i → type-kripke-frame F → Prop l4) + (V' : type-Set i → type-kripke-frame F → Prop l5) → + valuates-equals-on-vars (vars₁) V V' → + valuates-equals-on-vars (vars₂) V V' + valuates-equals-on-vars-subtype vars₁ vars₂ sub V V' eq (a , in-vars) = + eq (a , sub a in-vars) + + affect-only-occured-vars : + (a : formula i) → + (V : type-Set i → type-kripke-frame F → Prop l4) → + (V' : type-Set i → type-kripke-frame F → Prop l5) → + valuates-equals-on-vars (in-list (collect-vars a)) V V' → + (x : type-kripke-frame F) → + (((F , V) , x) ⊨ a) ⇔ (((F , V') , x) ⊨ a) + affect-only-occured-vars (var n) V V' eq x = + pair + ( λ v → + ( map-raise + ( forward-implication + ( eq (n , unit-trunc-Prop (is-head n nil)) x) + ( map-inv-raise v)))) + ( λ v → + ( map-raise + ( backward-implication + ( eq (n , unit-trunc-Prop (is-head n nil)) x) + ( map-inv-raise v)))) + affect-only-occured-vars ⊥ V V' eq x = + (map-raise ∘ map-inv-raise) , (map-raise ∘ map-inv-raise) + affect-only-occured-vars (a →ₘ b) V V' eq x = + pair + ( λ fab fa → + ( forward-implication + ( affect-only-occured-vars b V V' + ( valuates-equals-on-vars-subtype + ( in-list (collect-vars (a →ₘ b))) + ( in-list (collect-vars b)) + ( subset-in-concat-right (collect-vars a) (collect-vars b)) + ( V) + ( V') + ( eq)) + ( x)) + ( fab + ( backward-implication + ( affect-only-occured-vars a V V' + ( valuates-equals-on-vars-subtype + ( in-list (collect-vars (a →ₘ b))) + ( in-list (collect-vars a)) + ( subset-in-concat-left (collect-vars a) (collect-vars b)) + ( V) + ( V') + ( eq)) + ( x)) + ( fa))))) + ( λ fab fa → + ( backward-implication + ( affect-only-occured-vars b V V' + ( valuates-equals-on-vars-subtype + ( in-list (collect-vars (a →ₘ b))) + ( in-list (collect-vars b)) + ( subset-in-concat-right (collect-vars a) (collect-vars b)) + ( V) + ( V') + ( eq)) + ( x)) + ( fab + ( forward-implication + ( affect-only-occured-vars a V V' + ( valuates-equals-on-vars-subtype + ( in-list (collect-vars (a →ₘ b))) + ( in-list (collect-vars a)) + ( subset-in-concat-left (collect-vars a) (collect-vars b)) + ( V) + ( V') + ( eq)) + ( x)) + ( fa))))) + affect-only-occured-vars (□ a) V V' eq x = + pair + ( λ f y r → + ( forward-implication (affect-only-occured-vars a V V' eq y) (f y r))) + ( λ f y r → + ( backward-implication (affect-only-occured-vars a V V' eq y) (f y r))) + +is-kuratowsky-finite-set-Prop' : {l : Level} → Set l → Prop l +is-kuratowsky-finite-set-Prop' X = + ∃-Prop (list (type-Set X)) (λ l → (x : type-Set X) → type-Prop (in-list l x)) + +is-kuratowsky-finite-set' : {l : Level} → Set l → UU l +is-kuratowsky-finite-set' X = type-Prop (is-kuratowsky-finite-set-Prop' X) + +is-kuratowsky-finite-set-is-kuratowsky-finite-set' : + {l : Level} (X : Set l) → + is-kuratowsky-finite-set' X → is-kuratowsky-finite-set X +is-kuratowsky-finite-set-is-kuratowsky-finite-set' X = + map-universal-property-trunc-Prop + ( is-kuratowsky-finite-set-Prop X) + ( λ (l , all-in-list) → + ( intro-∃ + ( length-list l) + ( pair + ( component-list l) + ( λ x → + ( apply-universal-property-trunc-Prop + ( all-in-list x) + ( trunc-Prop (fiber _ x)) + ( λ x-in-list → + ( unit-trunc-Prop + ( pair + ( index-in-list x l x-in-list) + ( inv + ( eq-component-list-index-in-list x l x-in-list)))))))))) + +map-list : + {l1 l2 : Level} {A : UU l1} {B : UU l2} + (l : list A) (f : Σ A (λ a → a ∈-list l) → B) → + list B +map-list nil f = nil +map-list {A = A} {B} (cons x l) f = cons (f (x , is-head x l)) (map-list l f') + where + f' : Σ A (λ a → a ∈-list l) → B + f' (a , in-list) = f (a , is-in-tail a x l in-list) + +in-map-list : + {l1 l2 : Level} {A : UU l1} {B : UU l2} + {l : list A} (f : Σ A (λ a → a ∈-list l) → B) + {a : A} (a-in-l : a ∈-list l) → + f (a , a-in-l) ∈-list map-list l f +in-map-list f (is-head _ l) = is-head _ _ +in-map-list {A = A} {B} f (is-in-tail _ x l a-in-l) = + is-in-tail _ _ _ (in-map-list _ a-in-l) + +module _ + {l : Level} (i : Set l) + where + + subformulas-list : formula i → list (formula i) + subformulas-list a = cons a (rest a) + where + rest : formula i → list (formula i) + rest (var x) = nil + rest ⊥ = nil + rest (a →ₘ b) = concat-list (subformulas-list a) (subformulas-list b) + rest (□ a) = subformulas-list a + + subformulas-list-has-subimpl : + (a : formula i) {x y : formula i} → + (x →ₘ y) ∈-list subformulas-list a → + (x ∈-list subformulas-list a) × (y ∈-list subformulas-list a) + subformulas-list-has-subimpl .(x →ₘ y) {x} {y} (is-head .(x →ₘ y) _) = + pair + ( is-in-tail x (x →ₘ y) _ + ( in-concat-left + ( subformulas-list x) + ( subformulas-list y) + ( is-head x _))) + ( is-in-tail y (x →ₘ y) _ + ( in-concat-right + ( subformulas-list x) + ( subformulas-list y) + ( is-head y _))) + subformulas-list-has-subimpl + (a →ₘ b) {x} {y} (is-in-tail .(x →ₘ y) .(a →ₘ b) _ xy-in-list) + with + in-concat-list + ( subformulas-list a) + ( subformulas-list b) + ( x →ₘ y) + ( xy-in-list) + ... | inl xy-in-left = + let (x-in-tail , y-in-tail) = subformulas-list-has-subimpl a xy-in-left + in pair + ( is-in-tail x (a →ₘ b) _ + ( in-concat-left (subformulas-list a) (subformulas-list b) x-in-tail)) + ( is-in-tail y (a →ₘ b) _ + ( in-concat-left (subformulas-list a) (subformulas-list b) y-in-tail)) + ... | inr xy-in-right = + let (x-in-tail , y-in-tail) = subformulas-list-has-subimpl b xy-in-right + in pair + ( is-in-tail x (a →ₘ b) _ + ( in-concat-right (subformulas-list a) (subformulas-list b) x-in-tail)) + ( is-in-tail y (a →ₘ b) _ + ( in-concat-right (subformulas-list a) (subformulas-list b) y-in-tail)) + subformulas-list-has-subimpl + (□ a) {x} {y} (is-in-tail .(x →ₘ y) .(□ a) _ xy-in-list) = + let (x-in-tail , y-in-tail) = subformulas-list-has-subimpl a xy-in-list + in (is-in-tail x (□ a) _ x-in-tail) , (is-in-tail y (□ a) _ y-in-tail) + + subformulas-list-has-subbox : + (a : formula i) {x : formula i} → + □ x ∈-list subformulas-list a → + x ∈-list subformulas-list a + subformulas-list-has-subbox .(□ x) {x} (is-head .(□ x) _) = + is-in-tail x (□ x) _ (is-head x _) + subformulas-list-has-subbox + (a →ₘ b) {x} (is-in-tail .(□ x) .(a →ₘ b) _ x-in-list) + with + in-concat-list (subformulas-list a) (subformulas-list b) (□ x) x-in-list + ... | inl x-in-left = + is-in-tail x (a →ₘ b) _ + ( in-concat-left (subformulas-list a) (subformulas-list b) + ( subformulas-list-has-subbox a x-in-left)) + ... | inr x-in-right = + is-in-tail x (a →ₘ b) _ + ( in-concat-right (subformulas-list a) (subformulas-list b) + ( subformulas-list-has-subbox b x-in-right)) + subformulas-list-has-subbox (□ a) {x} (is-in-tail .(□ x) .(□ a) _ x-in-list) = + is-in-tail x (□ a) _ (subformulas-list-has-subbox a x-in-list) + + is-modal-theory-closed-under-subformulas-subformulas-list : + (a : formula i) → + is-modal-theory-closed-under-subformulas i (in-list (subformulas-list a)) + is-modal-theory-closed-under-subformulas-subformulas-list a = + is-modal-theory-closed-under-subformulas-condition + ( i) + ( in-list (subformulas-list a)) + ( map-universal-property-trunc-Prop + ( product-Prop + ( in-list (subformulas-list a) _) + ( in-list (subformulas-list a) _)) + ( λ impl-in-list → + ( let (a-in-list , b-in-list) = + subformulas-list-has-subimpl a impl-in-list + in (unit-trunc-Prop a-in-list) , (unit-trunc-Prop b-in-list)))) + ( map-universal-property-trunc-Prop + ( in-list (subformulas-list a) _) + ( unit-trunc-Prop ∘ subformulas-list-has-subbox a)) + + subformulas-Set : formula i → Set l + subformulas-Set a = set-subset (formula-Set i) (in-list (subformulas-list a)) + + subformulas-Set-list : (a : formula i) → list (type-Set (subformulas-Set a)) + subformulas-Set-list a = + map-list + ( subformulas-list a) + ( λ (x , in-list) → x , unit-trunc-Prop in-list) + + is-kuratowsky-finite'-subformulas-list : + (a : formula i) → is-kuratowsky-finite-set' (subformulas-Set a) + is-kuratowsky-finite'-subformulas-list a = + intro-∃ + ( subformulas-Set-list a) + ( λ (b , trunc-b-in-list) → + ( apply-universal-property-trunc-Prop + ( trunc-b-in-list) + ( in-list (subformulas-Set-list a) (b , trunc-b-in-list)) + ( λ b-in-list → + ( unit-trunc-Prop + ( tr + ( λ i → (b , i) ∈-list subformulas-Set-list a) + ( eq-is-prop is-prop-type-trunc-Prop) + ( in-map-list + (λ (x , in-list) → x , unit-trunc-Prop in-list) + ( b-in-list))))))) + + is-kuratowsky-finite-subformulas-list : + (a : formula i) → is-kuratowsky-finite-set (subformulas-Set a) + is-kuratowsky-finite-subformulas-list a = + is-kuratowsky-finite-set-is-kuratowsky-finite-set' + ( subformulas-Set a) + ( is-kuratowsky-finite'-subformulas-list a) + + is-finite-subformulas-list : + LEM l → + (a : formula i) → + is-finite (type-subtype (in-list (subformulas-list a))) + is-finite-subformulas-list lem a = + is-finite-is-kuratowsky-finite-set + (subformulas-Set a) lem (is-kuratowsky-finite-subformulas-list a) + + is-finite-subtypes-subformulas-list : + {l2 : Level} → + LEM l → + LEM l2 → + (a : formula i) → + is-finite (type-subtype (in-list (subformulas-list a)) → Prop l2) + is-finite-subtypes-subformulas-list lem lem2 a = + is-finite-function-type + ( is-finite-subformulas-list lem a) + ( is-finite-Prop-LEM lem2) + +module _ + {l1 l2 : Level} (A : Set l1) (B : Set l2) + (lem : LEM l1) + (lem2 : LEM (l1 ⊔ l2)) + where + + surjection-from-injection : + type-Set A → + injection (type-Set A) (type-Set B) → + type-Set B → type-Set A + surjection-from-injection a (f , is-inj) b + with + lem2 + ( pair + ( Σ (type-Set A) (λ a → f a = b)) + ( is-prop-all-elements-equal + ( λ x y → + ( eq-pair-Σ + ( is-inj (pr2 x ∙ inv (pr2 y))) + ( eq-is-prop (is-set-type-Set B _ _)))))) + ... | inl x = pr1 x + ... | inr x = a + + is-invertable-surjection-from-injection : + (a0 : type-Set A) → + (inj@(f , _) : injection (type-Set A) (type-Set B)) → + (a : type-Set A) → + surjection-from-injection a0 inj (f a) = a + is-invertable-surjection-from-injection a0 (f , is-inj) a + with + lem2 + ( pair + ( Σ (type-Set A) (λ a' → f a' = f a)) + ( is-prop-all-elements-equal + ( λ x y → + ( eq-pair-Σ + ( is-inj (pr2 x ∙ inv (pr2 y))) + ( eq-is-prop (is-set-type-Set B _ _)))))) + ... | inl x = is-inj (pr2 x) + ... | inr x = ex-falso (x (a , refl)) + + is-surjective-surjection-from-injection : + (a : type-Set A) → + (inj : injection (type-Set A) (type-Set B)) → + is-surjective (surjection-from-injection a inj) + is-surjective-surjection-from-injection a0 (f , is-inj) a = + unit-trunc-Prop + ( f a , is-invertable-surjection-from-injection a0 (f , is-inj) a) + + kuratowsky-finite-set-injection : + injection (type-Set A) (type-Set B) → + is-kuratowsky-finite-set B → + is-kuratowsky-finite-set A + kuratowsky-finite-set-injection inj b-is-fin + with lem (is-inhabited-Prop (type-Set A)) + ... | inl x = + apply-twice-universal-property-trunc-Prop + ( x) + ( b-is-fin) + ( is-kuratowsky-finite-set-Prop A) + ( λ a (n , e) → + ( unit-trunc-Prop + ( triple + ( n) + ( surjection-from-injection a inj ∘ pr1 e) + ( is-surjective-comp + ( is-surjective-surjection-from-injection a inj) + ( pr2 e))))) + ... | inr x = + is-kuratowsky-finite-set-is-finite A + ( is-finite-is-empty (x ∘ unit-trunc-Prop)) + + is-finite-injection : + injection (type-Set A) (type-Set B) → + is-finite (type-Set B) → + is-finite (type-Set A) + is-finite-injection inj b-is-fin = + is-finite-is-kuratowsky-finite-set A + lem (kuratowsky-finite-set-injection inj + ( is-kuratowsky-finite-set-is-finite B b-is-fin)) + +module _ + {l1 l2 l3 l4 l5 : Level} (i : Set l3) + (M : kripke-model l1 l2 i l4) + (lem : LEM (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5)) + (theory : modal-theory l5 i) + (is-fin : is-finite (type-subtype theory)) + where + + is-finite-equivalence-classes-filtration : + is-finite (equivalence-class (Φ-equivalence i theory M)) + is-finite-equivalence-classes-filtration = + is-finite-injection + ( equivalence-class-Set (Φ-equivalence i theory M)) + ( function-Set (type-subtype theory) (Prop-Set (l1 ⊔ l2 ⊔ l4))) + ( lem) + ( lem) + ( injection-map-function-equivalence-class i theory M) + ( is-finite-function-type + ( is-fin) + ( is-finite-Prop-LEM + ( lower-LEM (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) lem))) + + is-small-equivalence-classes-filtration : + (l : Level) → is-small l (equivalence-class (Φ-equivalence i theory M)) + is-small-equivalence-classes-filtration l = + is-small-is-finite l + ( pair + ( equivalence-class (Φ-equivalence i theory M)) + ( is-finite-equivalence-classes-filtration)) + +module _ + {l1 l2 l3 l4 : Level} (i : Set l3) + -- (theory : modal-theory l5 i) + -- (is-fin : is-finite (type-subtype theory)) + (l5 l6 l7 l8 : Level) + (lem : LEM (l2 ⊔ lsuc l3 ⊔ l4 ⊔ lsuc l5 ⊔ lsuc l6 ⊔ lsuc l7 ⊔ lsuc l8)) + where + + filtration-models-subset-finite-models : + filtration-models l1 l2 i l4 l5 l6 l7 l8 ⊆ finite-models l1 l2 i l4 + filtration-models-subset-finite-models M* = + map-universal-property-trunc-Prop + ( finite-models l1 l2 i l4 M*) + ( λ ((theory , M) , is-fin , t-is-filt) → + ( apply-universal-property-trunc-Prop + ( t-is-filt) + ( finite-models l1 l2 i l4 M*) + ( λ is-filt → + ( triple + ( is-finite-equiv + ( equiv-is-kripke-model-filtration i theory M M* is-filt) + ( is-finite-equivalence-classes-filtration i M + ( lower-LEM (l2 ⊔ l4) lem) + ( theory) + ( is-fin))) + ( λ x y → + ( lower-LEM + ( lsuc l3 ⊔ l4 ⊔ lsuc l5 ⊔ lsuc l6 ⊔ lsuc l7 ⊔ lsuc l8) + ( lem) + ( relation-Prop-kripke-model i M* x y))) + ( λ x n → + ( lower-LEM + ( l2 ⊔ lsuc l3 ⊔ lsuc l5 ⊔ lsuc l6 ⊔ lsuc l7 ⊔ lsuc l8) + ( lem) + ( valuate-kripke-model i M* n x))))))) + + -- module _ + -- (l6 l7 l8 : Level) + -- where + + -- private + -- w' = type-is-small (is-small-equivalence-classes-filtration l6) + + -- extend : + -- (type-subtype (λ n → theory (var n)) → w' → Prop l8) → + -- type-Set i → w' → Prop l8 + -- extend f n x + -- with + -- lower-LEM + -- ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) + -- ( lem) + -- ( theory (var n)) + -- ... | inl in-theory = f (n , in-theory) x + -- ... | inr _ = raise-empty-Prop l8 + + -- is-bounded-valuate-extend : + -- (f : type-subtype (λ n → theory (var n)) → w' → Prop l8) → + -- is-bounded-valuate i theory M (extend f) + -- is-bounded-valuate-extend f n not-in-theory x val + -- with + -- lower-LEM + -- ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) + -- ( lem) + -- ( theory (var n)) + -- ... | inl in-theory = not-in-theory in-theory + -- ... | inr not-in-theory' = map-inv-raise val + + -- last : + -- (w' → w' → Prop l7) + -- × (type-subtype (λ n → theory (var n)) → w' → Prop l8) → + -- type-subtype (is-alpha-Prop i theory M {l6} {l7} {l8}) + -- last (r , v) = + -- pair + -- ( pair + -- ( pair + -- ( pair + -- ( w') + -- ( map-is-inhabited + -- ( λ x → + -- ( map-equiv-is-small + -- ( is-small-equivalence-classes-filtration l6) + -- ( class (Φ-equivalence i theory M) x))) + -- ( is-inhabited-type-kripke-model i M))) + -- ( r)) + -- ( extend v)) + -- ( unit-trunc-Prop + -- ( pair + -- ( equiv-is-small (is-small-equivalence-classes-filtration l6)) + -- ( is-bounded-valuate-extend v))) + + -- is-surjective-last : is-surjective last + -- is-surjective-last (r , v) = + -- unit-trunc-Prop + -- (pair + -- ( pair + -- ( λ x y → + -- ( relation-Prop-kripke-model i r + -- ( {! !}) + -- ( {! !}))) + -- ( {! !})) + -- ( {! !})) + +module _ + {l1 l2 l3 l4 l5 l6 l7 l8 : Level} (i : Set l3) + (C : model-class l1 l2 i l4 l5) + (filtration : modal-theory l3 i → + kripke-model l1 l2 i l4 → + kripke-model l6 l7 i l8) + where + + filtrate-class : + model-class l6 l7 i l8 + ( lsuc l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ l5 ⊔ lsuc l6 ⊔ lsuc l7 ⊔ lsuc l8) + filtrate-class M* = + ∃-Prop (formula i × type-subtype C) + ( λ (a , (M , _)) → + ( M* = filtration (in-list (subformulas-list i a)) M)) + + module _ + (filtration-is-filtration : + ((M , _) : type-subtype C) + (theory : modal-theory l3 i) → + is-modal-theory-closed-under-subformulas i theory → + is-kripke-model-filtration i theory M (filtration theory M)) + where + filtrate-completeness : + {l9 : Level} (logic : modal-theory l9 i) → + completeness logic C → + completeness logic filtrate-class + filtrate-completeness logic complete a in-logic = + complete a + ( λ M M-in-class x → + apply-universal-property-trunc-Prop + ( filtration-is-filtration + ( M , M-in-class) + ( in-list (subformulas-list i a)) + ( is-modal-theory-closed-under-subformulas-subformulas-list i a)) + ( (M , x) ⊨ a) + ( λ is-filt → + ( backward-implication + ( kripke-models-filtrations-theorem' i + ( in-list (subformulas-list i a)) + ( is-modal-theory-closed-under-subformulas-subformulas-list + ( i) + ( a)) + ( M) + ( filtration (in-list (subformulas-list i a)) M) + ( is-filt) + ( a) + ( unit-trunc-Prop (is-head a _)) + ( x)) + ( in-logic + ( filtration (in-list (cons a _)) M) + ( intro-∃ (a , (M , M-in-class)) refl) + ( map-equiv-is-kripke-model-filtration + ( i) + ( in-list (subformulas-list i a)) + ( M) + ( filtration (in-list (cons a _)) M) + ( is-filt) + ( class + ( Φ-equivalence i + ( in-list (subformulas-list i a)) + ( M)) + ( x))))))) + + filtrate-soundness : + {l9 l10 : Level} (logic : modal-theory l9 i) → + (C₂ : model-class l6 l7 i l8 l10) → + ( ((M , _) : type-subtype C) → + (a : formula i) → + is-in-subtype C₂ (filtration (in-list (subformulas-list i a)) M)) → + soundness logic C₂ → + soundness logic filtrate-class + filtrate-soundness logic C₂ H sound a in-logic M* in-class = + apply-universal-property-trunc-Prop + ( in-class) + ( M* ⊨M a) + ( λ ((b , (M , in-C)) , p) → + ( sound a in-logic M* + ( tr (is-in-subtype C₂) (inv p) (H (M , in-C) b)))) +``` diff --git a/src/modal-logic/weak-deduction.lagda.md b/src/modal-logic/weak-deduction.lagda.md index 34307928bd..43bfea902a 100644 --- a/src/modal-logic/weak-deduction.lagda.md +++ b/src/modal-logic/weak-deduction.lagda.md @@ -259,6 +259,28 @@ in-concat-lists-union l1 l2 a = ( in-list (concat-list l1 l2) a) ( in-concat-list-sum-Prop l1 l2 a) +subset-in-concat-left : + {l : Level} {A : UU l} (l1 l2 : list A) → + in-list l1 ⊆ in-list (concat-list l1 l2) +subset-in-concat-left l1 l2 = + transitive-leq-subtype + ( in-list l1) + ( union-subtype (in-list l1) (in-list l2)) + ( in-list (concat-list l1 l2)) + ( in-concat-lists-union l1 l2) + ( subtype-union-left (in-list l1) (in-list l2)) + +subset-in-concat-right : + {l : Level} {A : UU l} (l1 l2 : list A) → + in-list l2 ⊆ in-list (concat-list l1 l2) +subset-in-concat-right l1 l2 = + transitive-leq-subtype + ( in-list l2) + ( union-subtype (in-list l1) (in-list l2)) + ( in-list (concat-list l1 l2)) + ( in-concat-lists-union l1 l2) + ( subtype-union-right (in-list l1) (in-list l2)) + empty-in-list-nil : {l : Level} {A : UU l} {x : A} → is-in-subtype (in-list nil) x → empty empty-in-list-nil = map-universal-property-trunc-Prop empty-Prop ( λ ()) diff --git a/src/univalent-combinatorics/kuratowsky-finite-sets.lagda.md b/src/univalent-combinatorics/kuratowsky-finite-sets.lagda.md index adc3eb2069..78a8b2a4c6 100644 --- a/src/univalent-combinatorics/kuratowsky-finite-sets.lagda.md +++ b/src/univalent-combinatorics/kuratowsky-finite-sets.lagda.md @@ -12,12 +12,18 @@ open import elementary-number-theory.natural-numbers open import foundation.decidable-equality open import foundation.dependent-pair-types open import foundation.existential-quantification +open import foundation.law-of-excluded-middle +open import foundation.logical-equivalences open import foundation.propositional-truncations open import foundation.propositions open import foundation.sets open import foundation.surjective-maps open import foundation.universe-levels +open import foundation-core.equivalences +open import foundation-core.function-types +open import foundation-core.identity-types + open import univalent-combinatorics.equality-finite-types open import univalent-combinatorics.finite-types open import univalent-combinatorics.image-of-maps @@ -85,3 +91,39 @@ has-decidable-equality-is-finite-type-𝔽-Kuratowsky : has-decidable-equality-is-finite-type-𝔽-Kuratowsky X H = has-decidable-equality-is-finite H ``` + +### TODO: change title + +```agda +is-kuratowsky-finite-set-surjection : + {l1 l2 : Level} (X : Set l1) (Y : Set l2) → + type-Set X ↠ type-Set Y → + is-kuratowsky-finite-set X → + is-kuratowsky-finite-set Y +is-kuratowsky-finite-set-surjection X Y f = + map-universal-property-trunc-Prop + ( is-kuratowsky-finite-set-Prop Y) + ( λ (n , g) → (intro-∃ n (surjection-comp f g))) + +is-kuratowsky-finite-set-is-finite : + {l : Level} (X : Set l) → + is-finite (type-Set X) → + is-kuratowsky-finite-set X +is-kuratowsky-finite-set-is-finite X = + map-universal-property-trunc-Prop + ( is-kuratowsky-finite-set-Prop X) + ( λ (n , e) → intro-∃ n (map-equiv e , is-surjective-map-equiv e)) +``` + +### Classical facts + +```agda +is-finite-is-kuratowsky-finite-set : + {l : Level} (X : Set l) → + LEM l → + is-kuratowsky-finite-set X → is-finite (type-Set X) +is-finite-is-kuratowsky-finite-set X lem is-fin = + is-finite-has-decidable-equality-type-𝔽-Kuratowsky + ( X , is-fin) + ( λ x y → lem (Id-Prop X x y)) +``` From f41eaff07f1272def8139fc156f84113bfa97c65 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Sat, 6 Apr 2024 16:19:40 +0300 Subject: [PATCH 37/63] Rename map-list to dependent-map-list --- src/modal-logic/modal-logic-decision.lagda.md | 21 ++++++++++--------- 1 file changed, 11 insertions(+), 10 deletions(-) diff --git a/src/modal-logic/modal-logic-decision.lagda.md b/src/modal-logic/modal-logic-decision.lagda.md index f21c46f372..b9c3eb49d8 100644 --- a/src/modal-logic/modal-logic-decision.lagda.md +++ b/src/modal-logic/modal-logic-decision.lagda.md @@ -257,24 +257,25 @@ is-kuratowsky-finite-set-is-kuratowsky-finite-set' X = ( inv ( eq-component-list-index-in-list x l x-in-list)))))))))) -map-list : +dependent-map-list : {l1 l2 : Level} {A : UU l1} {B : UU l2} (l : list A) (f : Σ A (λ a → a ∈-list l) → B) → list B -map-list nil f = nil -map-list {A = A} {B} (cons x l) f = cons (f (x , is-head x l)) (map-list l f') +dependent-map-list nil f = nil +dependent-map-list {A = A} {B} (cons x l) f = + cons (f (x , is-head x l)) (dependent-map-list l f') where f' : Σ A (λ a → a ∈-list l) → B f' (a , in-list) = f (a , is-in-tail a x l in-list) -in-map-list : +in-dependent-map-list : {l1 l2 : Level} {A : UU l1} {B : UU l2} {l : list A} (f : Σ A (λ a → a ∈-list l) → B) {a : A} (a-in-l : a ∈-list l) → - f (a , a-in-l) ∈-list map-list l f -in-map-list f (is-head _ l) = is-head _ _ -in-map-list {A = A} {B} f (is-in-tail _ x l a-in-l) = - is-in-tail _ _ _ (in-map-list _ a-in-l) + f (a , a-in-l) ∈-list dependent-map-list l f +in-dependent-map-list f (is-head _ l) = is-head _ _ +in-dependent-map-list {A = A} {B} f (is-in-tail _ x l a-in-l) = + is-in-tail _ _ _ (in-dependent-map-list _ a-in-l) module _ {l : Level} (i : Set l) @@ -377,7 +378,7 @@ module _ subformulas-Set-list : (a : formula i) → list (type-Set (subformulas-Set a)) subformulas-Set-list a = - map-list + dependent-map-list ( subformulas-list a) ( λ (x , in-list) → x , unit-trunc-Prop in-list) @@ -395,7 +396,7 @@ module _ ( tr ( λ i → (b , i) ∈-list subformulas-Set-list a) ( eq-is-prop is-prop-type-trunc-Prop) - ( in-map-list + ( in-dependent-map-list (λ (x , in-list) → x , unit-trunc-Prop in-list) ( b-in-list))))))) From 53eb3bf2e9acfe19de4a694f4b2779c30a9f80c8 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Wed, 1 May 2024 00:41:07 +0300 Subject: [PATCH 38/63] Small fixes --- .../decidable-propositions.lagda.md | 34 ++++++++- .../intersections-subtypes.lagda.md | 3 + .../law-of-excluded-middle.lagda.md | 37 ++++++++- src/foundation/unions-subtypes.lagda.md | 2 +- src/modal-logic/axioms.lagda.md | 16 ++-- .../canonical-model-theorem.lagda.md | 68 ++++++++++++----- src/modal-logic/canonical-theories.lagda.md | 1 - src/modal-logic/completeness-K.lagda.md | 25 +++---- src/modal-logic/completeness-S5.lagda.md | 25 +++---- src/modal-logic/completeness.lagda.md | 11 +-- src/modal-logic/formulas.lagda.md | 24 +++--- .../kripke-models-filtrations.lagda.md | 75 ++++++------------- src/modal-logic/kripke-semantics.lagda.md | 44 +---------- src/modal-logic/logic-syntax.lagda.md | 38 ++++++---- src/modal-logic/modal-logic-decision.lagda.md | 44 +++++------ src/modal-logic/soundness.lagda.md | 33 ++++---- 16 files changed, 261 insertions(+), 219 deletions(-) diff --git a/src/foundation-core/decidable-propositions.lagda.md b/src/foundation-core/decidable-propositions.lagda.md index 7134f32b4b..fbb712b15a 100644 --- a/src/foundation-core/decidable-propositions.lagda.md +++ b/src/foundation-core/decidable-propositions.lagda.md @@ -11,13 +11,16 @@ open import foundation.coproduct-types open import foundation.decidable-types open import foundation.dependent-pair-types open import foundation.double-negation +open import foundation.empty-types +open import foundation.equivalences +open import foundation.logical-equivalences open import foundation.negation open import foundation.propositional-truncations +open import foundation.small-types open import foundation.unit-type open import foundation.universe-levels open import foundation-core.cartesian-product-types -open import foundation-core.empty-types open import foundation-core.function-types open import foundation-core.functoriality-dependent-pair-types open import foundation-core.propositions @@ -189,3 +192,32 @@ is-merely-decidable-is-decidable-trunc-Prop A (inl x) = is-merely-decidable-is-decidable-trunc-Prop A (inr f) = unit-trunc-Prop (inr (f ∘ unit-trunc-Prop)) ``` + +### TODO + +```agda +-- TODO: not only for decidable +equiv-empty-is-decidable-prop : + {l : Level} {A : UU l} → is-decidable-prop A → ¬ A → A ≃ empty +equiv-empty-is-decidable-prop {l} {A} (is-p , _) contra = + equiv-iff (A , is-p) empty-Prop contra ex-falso + +-- TODO: not only for decidable +equiv-unit-is-decidable-prop : + {l : Level} {A : UU l} → is-decidable-prop A → A → A ≃ unit +equiv-unit-is-decidable-prop {l} {A} (is-p , _) a = + equiv-iff (A , is-p) unit-Prop (λ _ → star) (λ _ → a) + +equiv-empty-or-unit-is-decidable-prop : + {l : Level} {A : UU l} → is-decidable-prop A → (A ≃ unit) + (A ≃ empty) +equiv-empty-or-unit-is-decidable-prop {l} {A} H@(_ , is-d) with is-d +... | inl contra = inl (equiv-unit-is-decidable-prop H contra) +... | inr a = inr (equiv-empty-is-decidable-prop H a) + +is-small-prop-is-decidable-prop : + {l1 : Level} (l2 : Level) (A : UU l1) → is-decidable-prop A → is-small l2 A +is-small-prop-is-decidable-prop l2 A H + with equiv-empty-or-unit-is-decidable-prop H +... | inl e = is-small-equiv unit e (raise-unit l2 , compute-raise-unit l2) +... | inr e = is-small-equiv empty e (raise-empty l2 , compute-raise-empty l2) +``` diff --git a/src/foundation/intersections-subtypes.lagda.md b/src/foundation/intersections-subtypes.lagda.md index 00891864f6..31fdec6e97 100644 --- a/src/foundation/intersections-subtypes.lagda.md +++ b/src/foundation/intersections-subtypes.lagda.md @@ -61,6 +61,9 @@ module _ pr1 (p x r) pr2 (pr2 (is-greatest-binary-lower-bound-intersection-subtype R) p) x r = pr2 (p x r) + + infixl 15 _∩_ + _∩_ = intersection-subtype ``` ### The intersection of two decidable subtypes diff --git a/src/foundation/law-of-excluded-middle.lagda.md b/src/foundation/law-of-excluded-middle.lagda.md index 97258ff4c8..1a8a6ed07e 100644 --- a/src/foundation/law-of-excluded-middle.lagda.md +++ b/src/foundation/law-of-excluded-middle.lagda.md @@ -13,6 +13,8 @@ open import foundation.decidable-types open import foundation.dependent-pair-types open import foundation.propositions open import foundation.raising-universe-levels +open import foundation.small-types +open import foundation.subtypes open import foundation.universe-levels open import foundation-core.equality-dependent-pair-types @@ -66,7 +68,7 @@ pr1 (pr2 (decidable-prop-Prop lem P)) = is-prop-type-Prop P pr2 (pr2 (decidable-prop-Prop lem P)) = lem P ``` -### TODO: Given LEM, Prop equiv Decidable-Prop +### Given LEM, Prop equiv Decidable-Prop ```agda prop-equiv-decidable-prop : @@ -87,6 +89,39 @@ is-finite-Prop-LEM lem = is-finite-equiv' ( equiv-bool-Decidable-Prop ∘e prop-equiv-decidable-prop lem) ( is-finite-bool) + +is-small-Prop-LEM : {l1 : Level} (l2 : Level) → LEM l1 → is-small l2 (Prop l1) +is-small-Prop-LEM {l1} l2 lem = + is-small-equiv + ( Decidable-Prop l1) + ( prop-equiv-decidable-prop lem) + ( is-small-Decidable-Prop l1 l2) + +is-small-type-Prop-LEM : + {l1 : Level} (l2 : Level) → LEM l1 → (P : Prop l1) → is-small l2 (type-Prop P) +is-small-type-Prop-LEM l2 lem P = + is-small-prop-is-decidable-prop l2 (type-Prop P) (is-prop-type-Prop P , lem P) + +-- is-small-type-Prop-LEM : +-- {l1 : Level} (l2 : Level) → LEM l1 → is-small l2 (type-Prop l1) +-- is-small-type-Prop-LEM {l1} {l2} lem = +-- ? +-- -- is-small-Π +-- -- (is-small-Prop-LEM l2 lem) +-- -- (λ P → is-small-Π (is-small-Prop-LEM l2 lem) (λ _ → is-small-Prop-LEM l2 lem)) +``` + +### Given LEM, type subtype is small + +```agda +is-small-type-subtype-LEM : + {l1 l2 : Level} {A : UU l1} (P : subtype l2 A) → + LEM l2 → + is-small l1 (type-subtype P) +is-small-type-subtype-LEM {l1} {l2} {A} P lem = + is-small-Σ {l1} {l2} {l1} {l1} + (is-small' {l1} {A}) + (λ x → is-small-type-Prop-LEM l1 lem (P x)) ``` ### The unrestricted law of excluded middle does not hold diff --git a/src/foundation/unions-subtypes.lagda.md b/src/foundation/unions-subtypes.lagda.md index c078ad7c3e..7643d606d7 100644 --- a/src/foundation/unions-subtypes.lagda.md +++ b/src/foundation/unions-subtypes.lagda.md @@ -42,7 +42,7 @@ module _ union-subtype P Q x = (P x) ∨ (Q x) union-subtype P Q x = disjunction-Prop (P x) (Q x) - infixl 5 _∪_ + infixl 10 _∪_ _∪_ = union-subtype ``` diff --git a/src/modal-logic/axioms.lagda.md b/src/modal-logic/axioms.lagda.md index 0f95536c79..015787c87f 100644 --- a/src/modal-logic/axioms.lagda.md +++ b/src/modal-logic/axioms.lagda.md @@ -113,7 +113,7 @@ module _ ( eq-implication-right ∘ eq-box ∘ eq-implication-left) ax-dn : formulas l i - ax-dn = ax-1-parameter ( λ a → ~~ a →ₘ a) eq-implication-right + ax-dn = ax-1-parameter (λ a → ~~ a →ₘ a) eq-implication-right ax-m : formulas l i ax-m = ax-1-parameter (λ a → □ a →ₘ a) eq-implication-right @@ -140,13 +140,6 @@ module _ ax-k-soundness : soundness (ax-k i) (all-models l2 l3 i l4) ax-k-soundness .(a →ₘ b →ₘ a) (a , b , refl) M _ x fa _ = fa - ax-n-soundness : soundness (ax-n i) (all-models l2 l3 i l4) - ax-n-soundness - .(□ (a →ₘ b) →ₘ □ a →ₘ □ b) - (a , b , refl) - M in-class x fab fa y r = - fab y r (fa y r) - ax-s-soundness : soundness (ax-s i) (all-models l2 l3 i l4) ax-s-soundness .((a →ₘ b →ₘ c) →ₘ (a →ₘ b) →ₘ a →ₘ c) @@ -154,6 +147,13 @@ module _ M in-class x fabc fab fa = fabc fa (fab fa) + ax-n-soundness : soundness (ax-n i) (all-models l2 l3 i l4) + ax-n-soundness + .(□ (a →ₘ b) →ₘ □ a →ₘ □ b) + (a , b , refl) + M in-class x fab fa y r = + fab y r (fa y r) + ax-dn-soundness : soundness (ax-dn i) (decidable-models l2 l3 i l4) ax-dn-soundness .(~~ a →ₘ a) (a , refl) M is-dec x f with (is-dec a x) diff --git a/src/modal-logic/canonical-model-theorem.lagda.md b/src/modal-logic/canonical-model-theorem.lagda.md index 8eb907d55f..e53e8cf6de 100644 --- a/src/modal-logic/canonical-model-theorem.lagda.md +++ b/src/modal-logic/canonical-model-theorem.lagda.md @@ -54,6 +54,11 @@ open import order-theory.preorders open import order-theory.subposets open import order-theory.subtypes-leq-posets open import order-theory.top-elements-posets + +open import foundation.images +open import foundation.replacement +open import foundation.locally-small-types +open import foundation-core.small-types ```
@@ -66,7 +71,7 @@ TODO ```agda module _ - {l1 l2 : Level} {i : Set l1} (axioms : formulas l2 i) + {l1 l2 : Level} {i : Set l1} (axioms : modal-theory l2 i) (zorn : Zorn (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) l2) (prop-resize : propositional-resizing (l1 ⊔ l2) (lsuc (l1 ⊔ l2))) (L-is-cons : is-consistent-modal-logic (modal-logic axioms)) @@ -74,7 +79,7 @@ module _ where private - logic : formulas (l1 ⊔ l2) i + logic : modal-theory (l1 ⊔ l2) i logic = modal-logic axioms contains-ax-k : ax-k i ⊆ logic @@ -113,15 +118,15 @@ module _ ( contains-K) ( K-contains-ax-dn i) - is-L-consistent-theory-Prop : formulas (l1 ⊔ l2) i → Prop (l1 ⊔ l2) + is-L-consistent-theory-Prop : modal-theory (l1 ⊔ l2) i → Prop (l1 ⊔ l2) is-L-consistent-theory-Prop t = - is-consistent-modal-logic-Prop (weak-modal-logic (union-subtype logic t)) + is-consistent-modal-logic-Prop (weak-modal-logic (logic ∪ t)) - is-L-consistent-theory : formulas (l1 ⊔ l2) i → UU (l1 ⊔ l2) + is-L-consistent-theory : modal-theory (l1 ⊔ l2) i → UU (l1 ⊔ l2) is-L-consistent-theory = type-Prop ∘ is-L-consistent-theory-Prop L-consistent-theory : UU (lsuc l1 ⊔ lsuc l2) - L-consistent-theory = Σ (formulas (l1 ⊔ l2) i) is-L-consistent-theory + L-consistent-theory = type-subtype is-L-consistent-theory-Prop L-consistent-theory-leq-Prop : L-consistent-theory → L-consistent-theory → Prop (l1 ⊔ l2) @@ -145,6 +150,9 @@ module _ is-L-complete-theory : L-consistent-theory → UU (lsuc l1 ⊔ lsuc l2) is-L-complete-theory = type-Prop ∘ is-L-complete-theory-Prop + L-complete-theory : UU (lsuc l1 ⊔ lsuc l2) + L-complete-theory = type-subtype is-L-complete-theory-Prop + weak-modal-logic-subset-complete-theory : (x : L-consistent-theory) → is-L-complete-theory x → weak-modal-logic (pr1 x) ⊆ pr1 x @@ -213,11 +221,11 @@ module _ ( subtype-union-right (Id-formula-Prop a) (pr1 x))) ( subtype-union-left (Id-formula-Prop a) (pr1 x) a refl) - complete-theory-contains-all-formulas : + disjuctivity-L-complete-theory : LEM (l1 ⊔ l2) → (x : L-consistent-theory) → is-L-complete-theory x → - (a : formula i) → type-disjunction-Prop (pr1 x a) (pr1 x (~ a)) - complete-theory-contains-all-formulas lem x is-comp a with lem ((pr1 x) a) + (a : formula i) → pr1 x a ∨ pr1 x (~ a) + disjuctivity-L-complete-theory lem x is-comp a with lem (pr1 x a) ... | inl a-in-logic = inl-disjunction-Prop ((pr1 x) a) ((pr1 x) (~ a)) a-in-logic ... | inr a-not-in-logic = @@ -311,7 +319,7 @@ module _ (λ a → pr1 x (□ a)) ⊆ pr1 y lemma-box-diamond lem x x-is-comp y y-is-comp sub a box-a-in-x = apply-universal-property-trunc-Prop - ( complete-theory-contains-all-formulas lem y y-is-comp a) + ( disjuctivity-L-complete-theory lem y y-is-comp a) ( pr1 y a) ( λ { (inl a-in-y) → a-in-y @@ -358,7 +366,7 @@ module _ is-in-subtype (pr1 x) (a →ₘ b) complete-theory-implication lem x is-comp {a} {b} imp = apply-universal-property-trunc-Prop - ( complete-theory-contains-all-formulas lem x is-comp a) + ( disjuctivity-L-complete-theory lem x is-comp a) ( pr1 x (a →ₘ b)) ( λ { (inl a-in-logic) → ( weak-modal-logic-subset-complete-theory @@ -498,7 +506,7 @@ module _ ( ~ a) ( unit-trunc-Prop (w-ax not-a-in-logic))))))}) - canonical-worlds : subtype (lsuc l1 ⊔ lsuc l2) (formulas (l1 ⊔ l2) i) + canonical-worlds : subtype (lsuc l1 ⊔ lsuc l2) (modal-theory (l1 ⊔ l2) i) canonical-worlds x = Σ-Prop ( is-L-consistent-theory-Prop x) @@ -506,7 +514,7 @@ module _ canonical-kripke-model-world-type : UU (lsuc l1 ⊔ lsuc l2) canonical-kripke-model-world-type = - Σ (formulas (l1 ⊔ l2) i) (is-in-subtype canonical-worlds) + Σ (modal-theory (l1 ⊔ l2) i) (is-in-subtype canonical-worlds) lindenbaum : (x : L-consistent-theory) → @@ -562,7 +570,7 @@ module _ chain-union : chain-Poset l2 L-consistent-big-theories-Poset → - formulas (lsuc l1 ⊔ lsuc l2) i + modal-theory (lsuc l1 ⊔ lsuc l2) i chain-union C a = exists-Prop ( type-subtype @@ -767,7 +775,7 @@ module _ (x : L-consistent-theory) (is-comp : is-L-complete-theory x) (a : formula i) → - ( (y : formulas (l1 ⊔ l2) i) → + ( (y : modal-theory (l1 ⊔ l2) i) → (is-canonical : is-in-subtype canonical-worlds y) → relation-kripke-model i canonical-kripke-model ( pr1 x , pr2 x , is-comp) @@ -776,7 +784,7 @@ module _ is-in-subtype (pr1 x) (□ a) complete-theory-box lem x is-comp a h = apply-universal-property-trunc-Prop - ( complete-theory-contains-all-formulas lem x is-comp (□ a)) + ( disjuctivity-L-complete-theory lem x is-comp (□ a)) ( pr1 x (□ a)) ( λ { (inl box-in-logic) → box-in-logic ; (inr not-box-in-logic) → @@ -842,7 +850,7 @@ module _ ( ap (λ x → g →ₘ x) (reverse-list-to-implications f l)) move-right : - {l' : Level} (axioms : formulas l' i) + {l' : Level} (axioms : modal-theory l' i) (f : formula i) (l : list (formula i)) → ax-k i ⊆ weak-modal-logic axioms → ax-s i ⊆ weak-modal-logic axioms → @@ -1240,4 +1248,30 @@ module _ ( pr2 (pr2 x)) ( a) ( in-logic)))) + + canonical-model-theorem'' : + (a : formula i) → logic a ⇔ (canonical-kripke-model ⊨M a) + pr1 (canonical-model-theorem'' a) in-logic x = + forward-implication + ( canonical-model-theorem a x) + ( logic-subset-L-complete-theory + ( pr1 x , pr1 (pr2 x)) + ( pr2 (pr2 x)) + ( a) + ( in-logic)) + pr2 (canonical-model-theorem'' a) in-class-logic with lem (logic a) + ... | inl x = x + ... | inr x = + ex-falso + ( forward-implication (canonical-model-theorem' a) x in-class-logic) + + canonical-model-completness : + {l3 : Level} + (C : model-class (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) i (l1 ⊔ l2) l3) → + is-in-subtype C canonical-kripke-model → + completeness logic C + canonical-model-completness C model-in-class a a-in-class-logic = + backward-implication + ( canonical-model-theorem'' a) + ( a-in-class-logic canonical-kripke-model model-in-class) ``` diff --git a/src/modal-logic/canonical-theories.lagda.md b/src/modal-logic/canonical-theories.lagda.md index 781384f385..722b1564f6 100644 --- a/src/modal-logic/canonical-theories.lagda.md +++ b/src/modal-logic/canonical-theories.lagda.md @@ -44,7 +44,6 @@ TODO ## Definition ```agda --- TODO: levels module _ {l1 : Level} (i : Set l1) diff --git a/src/modal-logic/completeness-K.lagda.md b/src/modal-logic/completeness-K.lagda.md index 3d9516a2bc..a75a6584a2 100644 --- a/src/modal-logic/completeness-K.lagda.md +++ b/src/modal-logic/completeness-K.lagda.md @@ -64,19 +64,14 @@ module _ where completeness-K : completeness (modal-logic-K i) (all-models (lsuc l1) l1 i l1) - completeness-K a in-kripke-logic with lem (modal-logic-K i a) - ... | inl in-logic = in-logic - ... | inr not-in-logic = - ex-falso - ( forward-implication - ( canonical-model-theorem' - ( modal-logic-K-axioms i) - ( zorn) - ( prop-resize) - ( is-consistent-K i) - ( refl-leq-subtype (modal-logic-K i)) - ( lem) - ( a)) - ( not-in-logic) - ( in-kripke-logic _ star)) + completeness-K = + canonical-model-completness + ( modal-logic-K-axioms i) + ( zorn) + ( prop-resize) + ( is-consistent-K i) + ( refl-leq-subtype (modal-logic-K i)) + ( lem) + ( all-models (lsuc l1) l1 i l1) + ( star) ``` diff --git a/src/modal-logic/completeness-S5.lagda.md b/src/modal-logic/completeness-S5.lagda.md index 602a3d9c0d..fac72921a4 100644 --- a/src/modal-logic/completeness-S5.lagda.md +++ b/src/modal-logic/completeness-S5.lagda.md @@ -93,19 +93,14 @@ module _ completeness-S5 : completeness (modal-logic-S5 i) (equivalence-kripke-class (lsuc l1) l1 i l1) - completeness-S5 a in-kripke-logic with lem (modal-logic-S5 i a) - ... | inl in-logic = in-logic - ... | inr not-in-logic = - ex-falso - ( forward-implication - ( canonical-model-theorem' - ( modal-logic-S5-axioms i) - ( zorn) - ( prop-resize) - ( is-consistent-S5 i) - ( modal-logic-K-sub-S5 i) - ( lem) - ( a)) - ( not-in-logic) - ( in-kripke-logic _ S5-canonical-model-is-equivalence)) + completeness-S5 = + canonical-model-completness + ( modal-logic-S5-axioms i) + ( zorn) + ( prop-resize) + ( is-consistent-S5 i) + ( modal-logic-K-sub-S5 i) + ( lem) + ( equivalence-kripke-class (lsuc l1) l1 i l1) + ( S5-canonical-model-is-equivalence) ``` diff --git a/src/modal-logic/completeness.lagda.md b/src/modal-logic/completeness.lagda.md index ec5370684e..d03a0646e5 100644 --- a/src/modal-logic/completeness.lagda.md +++ b/src/modal-logic/completeness.lagda.md @@ -44,13 +44,14 @@ TODO -- completeness = class-modal-logic C ⊆ logic module _ - {l1 l2 l3 l4 l5 l6 : Level} - {i : Set l1} (logic : formulas l2 i) - (C : model-class l3 l4 i l5 l6) + {l1 l2 l3 l4 l5 l6 : Level} {i : Set l1} where - completeness : UU (l1 ⊔ l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6) - completeness = class-modal-logic C ⊆ logic + completeness : + modal-theory l2 i → + model-class l3 l4 i l5 l6 → + UU (l1 ⊔ l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6) + completeness logic C = class-modal-logic C ⊆ logic ``` ## Properties diff --git a/src/modal-logic/formulas.lagda.md b/src/modal-logic/formulas.lagda.md index d1172e441d..0c72d6bbea 100644 --- a/src/modal-logic/formulas.lagda.md +++ b/src/modal-logic/formulas.lagda.md @@ -33,17 +33,17 @@ TODO ```agda module _ - {l : Level} (i : Set l) + {l : Level} where infixr 7 _→ₘ_ infixr 25 □_ - data formula : UU l where - var : type-Set i → formula - ⊥ : formula - _→ₘ_ : formula → formula → formula - □_ : formula → formula + data formula (i : Set l) : UU l where + var : type-Set i → formula i + ⊥ : formula i + _→ₘ_ : formula i → formula i → formula i + □_ : formula i → formula i module _ {l : Level} {i : Set l} @@ -51,8 +51,8 @@ module _ infixr 25 ~_ infixr 25 ~~_ - -- infixl 10 _∨_ - -- infixl 15 _∧_ + infixl 10 _∨ₘ_ + infixl 15 _∧ₘ_ infixr 25 ◇_ ~_ : formula i → formula i @@ -61,11 +61,11 @@ module _ ~~_ : formula i → formula i ~~ a = ~ ~ a - -- _∨_ : formula i → formula i → formula i - -- a ∨ b = ~ a →ₘ b + _∨ₘ_ : formula i → formula i → formula i + a ∨ₘ b = ~ a →ₘ b - -- _∧_ : formula i → formula i → formula i - -- a ∧ b = ~ (a →ₘ ~ b) + _∧ₘ_ : formula i → formula i → formula i + a ∧ₘ b = ~ (a →ₘ ~ b) ⊤ : formula i ⊤ = ~ ⊥ diff --git a/src/modal-logic/kripke-models-filtrations.lagda.md b/src/modal-logic/kripke-models-filtrations.lagda.md index 1aef6d6dc5..93d47aa1e0 100644 --- a/src/modal-logic/kripke-models-filtrations.lagda.md +++ b/src/modal-logic/kripke-models-filtrations.lagda.md @@ -257,19 +257,10 @@ module _ ( a) ( in-theory))))) - map-valuate-function-equivalence-class : - (class : equivalence-class Φ-equivalence) → - valuate-function-equivalence-class class → - type-subtype theory → Prop (l1 ⊔ l2 ⊔ l4) - map-valuate-function-equivalence-class _ = pr1 - map-function-equivalence-class : equivalence-class Φ-equivalence → type-subtype theory → Prop (l1 ⊔ l2 ⊔ l4) - map-function-equivalence-class class = - map-valuate-function-equivalence-class - ( class) - ( function-equivalence-class class) + map-function-equivalence-class = pr1 ∘ function-equivalence-class is-injective-map-function-equivalence-class : is-injective map-function-equivalence-class @@ -315,25 +306,6 @@ module _ pr2 injection-map-function-equivalence-class = is-injective-map-function-equivalence-class - is-bounded-valuate-Prop : - {l6 l7 : Level} {A : UU l6} → - (type-Set i → A → Prop l7) → - Prop (l3 ⊔ l5 ⊔ l6 ⊔ l7) - is-bounded-valuate-Prop {A = A} V = - Π-Prop (type-Set i) - ( λ n → - ( function-Prop - ( ¬ (is-in-subtype theory (var n))) - ( Π-Prop A - ( λ x → - ( neg-Prop (V n x)))))) - - -- is-bounded-valuate : - -- {l6 l7 : Level} {A : UU l6} → - -- (type-Set i → A → Prop l7) → - -- UU (l3 ⊔ l5 ⊔ l6 ⊔ l7) - -- is-bounded-valuate = type-Prop ∘ is-bounded-valuate-Prop - module _ {l6 l7 l8 : Level} (M* : kripke-model l6 l7 i l8) where @@ -534,9 +506,7 @@ module _ ( is-inhabited-type-kripke-model i M) minimal-kripke-model-filtration-relation : - equivalence-class Φ-equivalence → - equivalence-class Φ-equivalence → - Prop (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) + Relation-Prop (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) (equivalence-class Φ-equivalence) minimal-kripke-model-filtration-relation x* y* = ∃-Prop ( type-kripke-model i M × type-kripke-model i M) @@ -563,16 +533,18 @@ module _ minimal-kripke-model-filtration : kripke-model - ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) ( i) ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) - minimal-kripke-model-filtration = - pair - ( pair - ( equivalence-class Φ-equivalence , is-inhabited-equivalence-classes) - ( minimal-kripke-model-filtration-relation)) - ( minimal-kripke-model-filtration-valuate) + pr1 (pr1 (pr1 minimal-kripke-model-filtration)) = + equivalence-class Φ-equivalence + pr2 (pr1 (pr1 minimal-kripke-model-filtration)) = + is-inhabited-equivalence-classes + pr2 (pr1 minimal-kripke-model-filtration) = + minimal-kripke-model-filtration-relation + pr2 minimal-kripke-model-filtration = + minimal-kripke-model-filtration-valuate minimal-transitive-kripke-model-filtration : kripke-model @@ -580,13 +552,15 @@ module _ ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) ( i) ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) - minimal-transitive-kripke-model-filtration = - pair - ( pair - ( equivalence-class Φ-equivalence , is-inhabited-equivalence-classes) - ( transitive-closure-Prop - ( type-Relation-Prop minimal-kripke-model-filtration-relation))) - ( minimal-kripke-model-filtration-valuate) + pr1 (pr1 (pr1 minimal-transitive-kripke-model-filtration)) = + equivalence-class Φ-equivalence + pr2 (pr1 (pr1 minimal-transitive-kripke-model-filtration)) = + is-inhabited-equivalence-classes + pr2 (pr1 minimal-transitive-kripke-model-filtration) = + transitive-closure-Prop + ( type-Relation-Prop minimal-kripke-model-filtration-relation) + pr2 minimal-transitive-kripke-model-filtration = + minimal-kripke-model-filtration-valuate module _ (theory-is-closed : is-modal-theory-closed-under-subformulas i theory) @@ -783,8 +757,8 @@ module _ minimal-transitive-kripke-model-filtration-is-transitive : is-in-subtype ( transitivity-kripke-class - ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) - ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) ( i) ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) ( minimal-transitive-kripke-model-filtration) @@ -796,8 +770,8 @@ module _ is-in-subtype (equivalence-kripke-class l1 l2 i l4) M → is-in-subtype ( equivalence-kripke-class - ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) - ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) ( i) ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) ( minimal-transitive-kripke-model-filtration) @@ -815,7 +789,6 @@ module _ (l4 : Level) (l5 l6 l7 l8 : Level) where - filtration-models : model-class l1 l2 i l4 (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4 ⊔ lsuc l5 ⊔ lsuc l6 ⊔ lsuc l7 ⊔ lsuc l8) diff --git a/src/modal-logic/kripke-semantics.lagda.md b/src/modal-logic/kripke-semantics.lagda.md index 74dd41d71d..d8754eed50 100644 --- a/src/modal-logic/kripke-semantics.lagda.md +++ b/src/modal-logic/kripke-semantics.lagda.md @@ -62,10 +62,6 @@ module _ {l1 l2 : Level} where - -- TODO: replace with theorem - postulate - is-set-kripke-frame : is-set (kripke-frame l1 l2) - Inhabited-Type-kripke-frame : kripke-frame l1 l2 → Inhabited-Type l1 Inhabited-Type-kripke-frame = pr1 @@ -239,9 +235,8 @@ module _ where infix 5 _⊨_ - infix 5 _⊭_ infix 5 _⊨M_ - infix 5 _⊭M_ + infix 5 _⊨C_ _⊨_ : Σ (kripke-model l1 l2 i l4) (type-kripke-model i) → @@ -255,18 +250,9 @@ module _ ( type-kripke-model i M) ( λ y → function-Prop (relation-kripke-model i M x y) ((M , y) ⊨ a)) - _⊭_ : - Σ (kripke-model l1 l2 i l4) (type-kripke-model i) → - formula i → - Prop (l1 ⊔ l2 ⊔ l4) - (M , x) ⊭ a = neg-Prop ((M , x) ⊨ a) - _⊨M_ : kripke-model l1 l2 i l4 → formula i → Prop (l1 ⊔ l2 ⊔ l4) M ⊨M a = Π-Prop (type-kripke-model i M) (λ x → (M , x) ⊨ a) - _⊭M_ : kripke-model l1 l2 i l4 → formula i → Prop (l1 ⊔ l2 ⊔ l4) - M ⊭M a = neg-Prop (M ⊨M a) - _⊨C_ : {l5 : Level} → model-class l1 l2 i l4 l5 → @@ -280,7 +266,7 @@ module _ class-modal-logic : {l5 : Level} → model-class l1 l2 i l4 l5 → - formulas (lsuc l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ l5) i + modal-theory (lsuc l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ l5) i class-modal-logic = _⊨C_ -- TODO: rename @@ -293,28 +279,6 @@ module _ class-modal-logic-monotic C₁ C₂ sub _ in-modal-logic-C₂ M in-C₁ = in-modal-logic-C₂ M (sub M in-C₁) -module _ - {l1 l2 : Level} - {l3 : Level} (i : Set l3) - (l4 : Level) - where - - kripke-frame-model-class : - kripke-frame l1 l2 → - model-class l1 l2 i l4 (lsuc l1 ⊔ lsuc l2) - pr1 (kripke-frame-model-class F M) = - kripke-frame-kripke-model i M = F - pr2 (kripke-frame-model-class F M) = - is-set-kripke-frame (kripke-frame-kripke-model i M) F - - -- frame-modal-logic : - -- kripke-frame l1 l2 → - -- formulas (l1 ⊔ l2 ⊔ l3 ⊔ lsuc l4) i - -- frame-modal-logic F a = - -- Π-Prop - -- ( type-Set i → type-kripke-frame F → Prop l4) - -- ( λ v → (F , v) ⊨M a) - module _ (l1 l2 : Level) {l3 : Level} (i : Set l3) @@ -385,11 +349,11 @@ module _ {l5 : Level} → model-class l1 l2 i l4 l5 → model-class l1 l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) - decidable-subclass C = intersection-subtype (decidable-models l1 l2 i l4) C + decidable-subclass C = (decidable-models l1 l2 i l4) ∩ C finite-subclass : {l5 : Level} → model-class l1 l2 i l4 l5 → model-class l1 l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) - finite-subclass C = intersection-subtype (finite-models l1 l2 i l4) C + finite-subclass C = (finite-models l1 l2 i l4) ∩ C ``` diff --git a/src/modal-logic/logic-syntax.lagda.md b/src/modal-logic/logic-syntax.lagda.md index b2a10cc495..2921dbe176 100644 --- a/src/modal-logic/logic-syntax.lagda.md +++ b/src/modal-logic/logic-syntax.lagda.md @@ -37,11 +37,12 @@ module _ {l1 : Level} (l2 : Level) (i : Set l1) where - formulas : UU (l1 ⊔ lsuc l2) - formulas = subtype l2 (formula i) - modal-theory : UU (l1 ⊔ lsuc l2) - modal-theory = formulas + modal-theory = subtype l2 (formula i) + + -- TODO: remove + formulas : UU (l1 ⊔ lsuc l2) + formulas = modal-theory module _ {l1 l2 : Level} {i : Set l1} @@ -49,24 +50,35 @@ module _ infix 5 _⊢_ - data _⊢_ (axioms : formulas l2 i) : formula i → UU (l1 ⊔ l2) where + data _⊢_ (axioms : modal-theory l2 i) : formula i → UU (l1 ⊔ l2) where ax : {a : formula i} → is-in-subtype axioms a → axioms ⊢ a mp : {a b : formula i} → axioms ⊢ a →ₘ b → axioms ⊢ a → axioms ⊢ b nec : {a : formula i} → axioms ⊢ a → axioms ⊢ □ a - modal-logic : formulas l2 i → formulas (l1 ⊔ l2) i + is-modal-logic-Prop : modal-theory l2 i → Prop (l1 ⊔ l2) + is-modal-logic-Prop theory = + implicit-Π-Prop (formula i) (λ a → function-Prop (theory ⊢ a) (theory a)) + + is-modal-logic : modal-theory l2 i → UU (l1 ⊔ l2) + is-modal-logic = type-Prop ∘ is-modal-logic-Prop + + T-modal-logic : UU (l1 ⊔ lsuc l2) + T-modal-logic = Σ (modal-theory l2 i) is-modal-logic + + -- TODO: rename to modal-logic-closure + modal-logic : modal-theory l2 i → modal-theory (l1 ⊔ l2) i modal-logic axioms a = trunc-Prop (axioms ⊢ a) - is-contradictory-modal-logic-Prop : formulas l2 i → Prop l2 + is-contradictory-modal-logic-Prop : modal-theory l2 i → Prop l2 is-contradictory-modal-logic-Prop logic = logic ⊥ - is-contradictory-modal-logic : formulas l2 i → UU l2 + is-contradictory-modal-logic : modal-theory l2 i → UU l2 is-contradictory-modal-logic = type-Prop ∘ is-contradictory-modal-logic-Prop - is-consistent-modal-logic-Prop : formulas l2 i → Prop l2 + is-consistent-modal-logic-Prop : modal-theory l2 i → Prop l2 is-consistent-modal-logic-Prop = neg-Prop ∘ is-contradictory-modal-logic-Prop - is-consistent-modal-logic : formulas l2 i → UU l2 + is-consistent-modal-logic : modal-theory l2 i → UU l2 is-consistent-modal-logic = type-Prop ∘ is-consistent-modal-logic-Prop module _ @@ -107,16 +119,14 @@ module _ axioms-subset-modal-logic _ a H = unit-trunc-Prop (ax H) modal-logic-closed : - {l2 : Level} {axioms : formulas l2 i} {a : formula i} → - modal-logic axioms ⊢ a → - is-in-subtype (modal-logic axioms) a + {l2 : Level} {axioms : formulas l2 i} → is-modal-logic (modal-logic axioms) modal-logic-closed (ax x) = x modal-logic-closed {axioms = axioms} {a} (mp tdab tda) = apply-twice-universal-property-trunc-Prop ( modal-logic-closed tdab) ( modal-logic-closed tda) ( modal-logic axioms a) - (λ dab da → unit-trunc-Prop (mp dab da)) + ( λ dab da → unit-trunc-Prop (mp dab da)) modal-logic-closed {axioms = axioms} {a} (nec d) = apply-universal-property-trunc-Prop ( modal-logic-closed d) diff --git a/src/modal-logic/modal-logic-decision.lagda.md b/src/modal-logic/modal-logic-decision.lagda.md index b9c3eb49d8..ab395be441 100644 --- a/src/modal-logic/modal-logic-decision.lagda.md +++ b/src/modal-logic/modal-logic-decision.lagda.md @@ -87,18 +87,18 @@ module _ (C-is-fin : is-finite (type-subtype C)) where - decision-procedure : + decision-procedure' : (a : formula i) → is-decidable ( (M : type-subtype C) → type-Prop (inclusion-subtype C M ⊨M a)) - decision-procedure a = + decision-procedure' a = is-decidable-Π-is-finite ( C-is-fin) ( λ (M , M-in-C) → ( is-finite-model-valuate-decidable-models i M (C-sub-fin M M-in-C) a)) - decision-procedure' : (a : formula i) → bool - decision-procedure' a with decision-procedure a + decision-procedure : (a : formula i) → bool + decision-procedure a with decision-procedure' a ... | inl _ = true ... | inr _ = false @@ -106,14 +106,14 @@ module _ soundness theory C → completeness theory C → (a : formula i) → - theory a ⇔ prop-bool (decision-procedure' a) + theory a ⇔ prop-bool (decision-procedure a) pr1 (decision-procedure-correctness sound complete a) in-theory - with decision-procedure a + with decision-procedure' a ... | inl _ = star ... | inr not-valid-in-C = not-valid-in-C (λ (M , M-in-C) x → sound a in-theory M M-in-C x) pr2 (decision-procedure-correctness sound complete a) _ - with decision-procedure a + with decision-procedure' a ... | inl valid-in-C = complete a (λ M M-in-C → valid-in-C (M , M-in-C)) module _ @@ -290,6 +290,9 @@ module _ rest (a →ₘ b) = concat-list (subformulas-list a) (subformulas-list b) rest (□ a) = subformulas-list a + subformulas : formula i → modal-theory l i + subformulas a = in-list (subformulas-list a) + subformulas-list-has-subimpl : (a : formula i) {x y : formula i} → (x →ₘ y) ∈-list subformulas-list a → @@ -354,10 +357,10 @@ module _ subformulas-list-has-subbox (□ a) {x} (is-in-tail .(□ x) .(□ a) _ x-in-list) = is-in-tail x (□ a) _ (subformulas-list-has-subbox a x-in-list) - is-modal-theory-closed-under-subformulas-subformulas-list : + is-modal-theory-closed-under-subformulas-subformulas : (a : formula i) → - is-modal-theory-closed-under-subformulas i (in-list (subformulas-list a)) - is-modal-theory-closed-under-subformulas-subformulas-list a = + is-modal-theory-closed-under-subformulas i (subformulas a) + is-modal-theory-closed-under-subformulas-subformulas a = is-modal-theory-closed-under-subformulas-condition ( i) ( in-list (subformulas-list a)) @@ -648,12 +651,10 @@ module _ where filtrate-class : - model-class l6 l7 i l8 - ( lsuc l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ l5 ⊔ lsuc l6 ⊔ lsuc l7 ⊔ lsuc l8) + model-class l6 l7 i l8 ( l3 ⊔ l5 ⊔ lsuc (l1 ⊔ l2 ⊔ l4 ⊔ l6 ⊔ l7 ⊔ l8)) filtrate-class M* = ∃-Prop (formula i × type-subtype C) - ( λ (a , (M , _)) → - ( M* = filtration (in-list (subformulas-list i a)) M)) + ( λ (a , (M , _)) → M* = filtration (subformulas i a) M) module _ (filtration-is-filtration : @@ -662,6 +663,7 @@ module _ is-modal-theory-closed-under-subformulas i theory → is-kripke-model-filtration i theory M (filtration theory M)) where + filtrate-completeness : {l9 : Level} (logic : modal-theory l9 i) → completeness logic C → @@ -672,18 +674,18 @@ module _ apply-universal-property-trunc-Prop ( filtration-is-filtration ( M , M-in-class) - ( in-list (subformulas-list i a)) - ( is-modal-theory-closed-under-subformulas-subformulas-list i a)) + ( subformulas i a) + ( is-modal-theory-closed-under-subformulas-subformulas i a)) ( (M , x) ⊨ a) ( λ is-filt → ( backward-implication ( kripke-models-filtrations-theorem' i - ( in-list (subformulas-list i a)) - ( is-modal-theory-closed-under-subformulas-subformulas-list + ( subformulas i a) + ( is-modal-theory-closed-under-subformulas-subformulas ( i) ( a)) ( M) - ( filtration (in-list (subformulas-list i a)) M) + ( filtration (subformulas i a) M) ( is-filt) ( a) ( unit-trunc-Prop (is-head a _)) @@ -693,13 +695,13 @@ module _ ( intro-∃ (a , (M , M-in-class)) refl) ( map-equiv-is-kripke-model-filtration ( i) - ( in-list (subformulas-list i a)) + ( subformulas i a) ( M) ( filtration (in-list (cons a _)) M) ( is-filt) ( class ( Φ-equivalence i - ( in-list (subformulas-list i a)) + ( subformulas i a) ( M)) ( x))))))) diff --git a/src/modal-logic/soundness.lagda.md b/src/modal-logic/soundness.lagda.md index 69265fe09e..06b7e07b89 100644 --- a/src/modal-logic/soundness.lagda.md +++ b/src/modal-logic/soundness.lagda.md @@ -34,14 +34,12 @@ TODO ## Definition ```agda -module _ - {l1 l2 l3 l4 l5 l6 : Level} - {i : Set l1} (logic : formulas l2 i) - (C : model-class l3 l4 i l5 l6) - where - - soundness : UU (l1 ⊔ l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6) - soundness = logic ⊆ class-modal-logic C +soundness : + {l1 l2 l3 l4 l5 l6 : Level} {i : Set l1} → + modal-theory l2 i → + model-class l3 l4 i l5 l6 → + UU (l1 ⊔ l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6) +soundness logic C = logic ⊆ class-modal-logic C ``` ## Properties @@ -81,32 +79,33 @@ module _ module _ {l1 l2 l3 l4 l5 l6 l7 l8 : Level} - {i : Set l1} (ax₁ : formulas l2 i) (ax₂ : formulas l3 i) + {i : Set l1} (theory₁ : formulas l2 i) (theory₂ : formulas l3 i) (C₁ : model-class l4 l5 i l6 l7) (C₂ : model-class l4 l5 i l6 l8) - (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) + (sound₁ : soundness theory₁ C₁) (sound₂ : soundness theory₂ C₂) where forces-in-intersection : (M : kripke-model l4 l5 i l6) → - is-in-subtype (intersection-subtype C₁ C₂) M → + is-in-subtype (C₁ ∩ C₂) M → (a : formula i) → - is-in-subtype ax₁ a + is-in-subtype ax₂ a → + is-in-subtype theory₁ a + is-in-subtype theory₂ a → type-Prop (M ⊨M a) forces-in-intersection M in-class a (inl d) = sound₁ a d M (subtype-intersection-left C₁ C₂ M in-class) forces-in-intersection M in-class a (inr d) = sound₂ a d M (subtype-intersection-right C₁ C₂ M in-class) - soundness-union : - soundness (union-subtype ax₁ ax₂) (intersection-subtype C₁ C₂) - soundness-union a is-ax M in-class = + soundness-union : soundness (theory₁ ∪ theory₂) (C₁ ∩ C₂) + soundness-union a is-theory M in-class = apply-universal-property-trunc-Prop - ( is-ax) + ( is-theory) ( M ⊨M a) ( forces-in-intersection M in-class a) soundness-modal-logic-union : - soundness (modal-logic (union-subtype ax₁ ax₂)) (intersection-subtype C₁ C₂) + soundness + (modal-logic (union-subtype theory₁ theory₂)) + (intersection-subtype C₁ C₂) soundness-modal-logic-union = soundness-modal-logic (intersection-subtype C₁ C₂) soundness-union From e51080abdecbfe29cfbf529a78c60830e92691f1 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Thu, 2 May 2024 04:36:05 +0300 Subject: [PATCH 39/63] Refactor weak-deduction and L-complete theories --- src/foundation/unions-subtypes.lagda.md | 18 + src/modal-logic/L-complete-theories.lagda.md | 339 ++++ .../canonical-model-theorem.lagda.md | 5 - src/modal-logic/logic-syntax.lagda.md | 72 +- src/modal-logic/weak-deduction.lagda.md | 1451 ++++++++++------- 5 files changed, 1326 insertions(+), 559 deletions(-) create mode 100644 src/modal-logic/L-complete-theories.lagda.md diff --git a/src/foundation/unions-subtypes.lagda.md b/src/foundation/unions-subtypes.lagda.md index 7643d606d7..795acdc60e 100644 --- a/src/foundation/unions-subtypes.lagda.md +++ b/src/foundation/unions-subtypes.lagda.md @@ -15,6 +15,7 @@ open import foundation.identity-types open import foundation.large-locale-of-subtypes open import foundation.logical-equivalences open import foundation.powersets +open import foundation.propositions open import foundation.propositional-truncations open import foundation.subtypes open import foundation.universe-levels @@ -108,6 +109,23 @@ module _ subtype-union-right : Q ⊆ union-subtype P Q subtype-union-right x = inr-disjunction-Prop (P x) (Q x) + elim-union-subtype : + {l4 : Level} (f : X → Prop l4) → + ((x : X) → is-in-subtype P x → type-Prop (f x)) → + ((x : X) → is-in-subtype Q x → type-Prop (f x)) → + (x : X) → is-in-subtype (P ∪ Q) x → type-Prop (f x) + elim-union-subtype f H-P H-Q x = + elim-disjunction-Prop (P x) (Q x) (f x) (H-P x , H-Q x) + + elim-union-subtype' : + {l4 : Level} (R : Prop l4) → + ((x : X) → is-in-subtype P x → type-Prop R) → + ((x : X) → is-in-subtype Q x → type-Prop R) → + (x : X) → is-in-subtype (P ∪ Q) x → type-Prop R + elim-union-subtype' R H-P H-Q x = + elim-disjunction-Prop (P x) (Q x) R (H-P x , H-Q x) + -- elim-disjunction-Prop (P x) (Q x) (f x) (H-P x , H-Q x) p + subtype-union-both : {l4 : Level} (S : subtype l4 X) → P ⊆ S → Q ⊆ S → union-subtype P Q ⊆ S subtype-union-both S P-sub-S Q-sub-S x = diff --git a/src/modal-logic/L-complete-theories.lagda.md b/src/modal-logic/L-complete-theories.lagda.md new file mode 100644 index 0000000000..c296f2da2a --- /dev/null +++ b/src/modal-logic/L-complete-theories.lagda.md @@ -0,0 +1,339 @@ +# L-complete theories + +```agda +module modal-logic.L-complete-theories where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +-- open import foundation.cartesian-product-types +open import foundation.binary-relations +open import foundation.coproduct-types +open import foundation.dependent-pair-types +open import foundation.disjunction +-- open import foundation.empty-types +-- open import foundation.existential-quantification +open import foundation.function-types +open import foundation.identity-types +-- open import foundation.inhabited-types +-- open import foundation.intersections-subtypes +open import foundation.law-of-excluded-middle +open import foundation.logical-equivalences +open import foundation.negation +-- open import foundation.propositional-resizing +open import foundation.propositional-truncations +open import foundation.propositions +-- open import foundation.raising-universe-levels +open import foundation.sets +open import foundation.subtypes +open import foundation.transport-along-identifications +open import foundation.unions-subtypes +-- open import foundation.unit-type +open import foundation.universe-levels + +-- open import foundation-core.equivalences + +-- open import lists.lists +-- open import lists.reversing-lists + +open import modal-logic.axioms +-- open import modal-logic.completeness +open import modal-logic.formulas +-- open import modal-logic.formulas-deduction +-- open import modal-logic.kripke-semantics +open import modal-logic.logic-syntax +-- open import modal-logic.modal-logic-K +-- open import modal-logic.soundness +open import modal-logic.weak-deduction + +-- open import order-theory.chains-posets +open import order-theory.maximal-elements-posets +open import order-theory.posets +-- open import order-theory.preorders +open import order-theory.subposets +open import order-theory.subtypes-leq-posets +-- open import order-theory.top-elements-posets +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l1 l2 : Level} {i : Set l1} + (logic : modal-theory l2 i) + (logic-is-modal-logic : is-modal-logic logic) + where + + is-L-consistent-theory-Prop : + {l3 : Level} → modal-theory l3 i → Prop (l1 ⊔ l2 ⊔ l3) + is-L-consistent-theory-Prop theory = + is-consistent-modal-logic-Prop (weak-modal-logic-closure (logic ∪ theory)) + + is-L-consistent-theory : + {l3 : Level} → modal-theory l3 i → UU (l1 ⊔ l2 ⊔ l3) + is-L-consistent-theory = type-Prop ∘ is-L-consistent-theory-Prop + + L-consistent-theory : (l3 : Level) → UU (l1 ⊔ l2 ⊔ lsuc l3) + L-consistent-theory l3 = type-subtype (is-L-consistent-theory-Prop {l3}) + + modal-theory-L-consistent-theory : + {l3 : Level} → L-consistent-theory l3 → modal-theory l3 i + modal-theory-L-consistent-theory = + inclusion-subtype is-L-consistent-theory-Prop + + is-L-consistent-theory-modal-theory-L-consistent-theory : + {l3 : Level} (theory : L-consistent-theory l3) → + is-L-consistent-theory (modal-theory-L-consistent-theory theory) + is-L-consistent-theory-modal-theory-L-consistent-theory = + is-in-subtype-inclusion-subtype is-L-consistent-theory-Prop + + is-L-consistent-antimonotic : + {l3 l4 : Level} + (theory₁ : modal-theory l3 i) → + (theory₂ : modal-theory l4 i) → + theory₁ ⊆ theory₂ → + is-L-consistent-theory theory₂ → + is-L-consistent-theory theory₁ + is-L-consistent-antimonotic theory₁ theory₂ leq is-cons = + is-consistent-modal-logic-antimonotic + ( weak-modal-logic-closure (logic ∪ theory₁)) + ( weak-modal-logic-closure (logic ∪ theory₂)) + ( weak-modal-logic-closure-monotic + ( subset-union-subset-right logic theory₁ theory₂ leq)) + ( is-cons) + + L-consistent-theory-leq-Prop : + {l3 : Level} → Relation-Prop (l1 ⊔ l3) (L-consistent-theory l3) + -- {l3 : Level} → L-consistent-theory l3 → L-consistent-theory → Prop (l1 ⊔ l2) + L-consistent-theory-leq-Prop x y = + leq-prop-subtype + ( modal-theory-L-consistent-theory x) + ( modal-theory-L-consistent-theory y) + + L-consistent-theory-leq : + {l3 : Level} → Relation (l1 ⊔ l3) (L-consistent-theory l3) + L-consistent-theory-leq = type-Relation-Prop L-consistent-theory-leq-Prop + + theories-Poset : (l3 : Level) → Poset (l1 ⊔ lsuc l3) (l1 ⊔ l3) + theories-Poset l3 = subtypes-leq-Poset l3 (formula i) + + L-consistent-theories-Poset : + (l3 : Level) → Poset (l1 ⊔ l2 ⊔ lsuc l3) (l1 ⊔ l3) + L-consistent-theories-Poset l3 = + poset-Subposet (theories-Poset l3) (is-L-consistent-theory-Prop) + + is-L-complete-theory-Prop : + {l3 : Level} → L-consistent-theory l3 → Prop (l1 ⊔ l2 ⊔ lsuc l3) + is-L-complete-theory-Prop {l3} = + is-maximal-element-Poset-Prop (L-consistent-theories-Poset l3) + + is-L-complete-theory : + {l3 : Level} → L-consistent-theory l3 → UU (l1 ⊔ l2 ⊔ lsuc l3) + is-L-complete-theory = type-Prop ∘ is-L-complete-theory-Prop + + L-complete-theory : (l3 : Level) → UU (l1 ⊔ l2 ⊔ lsuc l3) + L-complete-theory l3 = type-subtype (is-L-complete-theory-Prop {l3}) + + L-consistent-theory-L-complete-theory : + {l3 : Level} → L-complete-theory l3 → L-consistent-theory l3 + L-consistent-theory-L-complete-theory = inclusion-subtype is-L-complete-theory-Prop + + is-L-complete-theory-L-consistent-theory : + {l3 : Level} (theory : L-complete-theory l3) → + is-L-complete-theory (L-consistent-theory-L-complete-theory theory) + is-L-complete-theory-L-consistent-theory = + is-in-subtype-inclusion-subtype is-L-complete-theory-Prop + + modal-theory-L-complete-theory : + {l3 : Level} → L-complete-theory l3 → modal-theory l3 i + modal-theory-L-complete-theory = + modal-theory-L-consistent-theory ∘ L-consistent-theory-L-complete-theory + + eq-is-L-consistent-union-L-complete : + {l3 l4 : Level} → + (((theory , _) , _) : L-complete-theory (l1 ⊔ l2 ⊔ l3 ⊔ l4)) → + (theory' : modal-theory l4 i) → + is-L-consistent-theory (theory' ∪ theory) → + theory' ∪ theory = theory + eq-is-L-consistent-union-L-complete + ((theory , is-cons) , is-comp) theory' is-L-cons = + ap + ( modal-theory-L-consistent-theory) + ( is-comp + ( theory' ∪ theory , is-L-cons) + ( subtype-union-right theory' theory)) + + union-L-consistent : + {l3 : Level} → + L-consistent-theory l3 → + L-consistent-theory (l1 ⊔ l2 ⊔ l3) + pr1 (union-L-consistent (theory , is-cons)) = + weak-modal-logic-closure (logic ∪ theory) + pr2 (union-L-consistent (theory , is-cons)) bot-in-logic = + is-cons + ( transitive-leq-subtype + ( weak-modal-logic-closure + ( logic ∪ weak-modal-logic-closure (logic ∪ theory))) + ( weak-modal-logic-closure (weak-modal-logic-closure (logic ∪ theory))) + ( weak-modal-logic-closure (logic ∪ theory)) + ( is-weak-modal-logic-weak-modal-logic-closure) + ( weak-modal-logic-closure-monotic + ( subtype-union-both + ( logic) + ( weak-modal-logic-closure (logic ∪ theory)) + ( weak-modal-logic-closure (logic ∪ theory)) + ( transitive-leq-subtype + ( logic) + ( logic ∪ theory) + ( weak-modal-logic-closure (logic ∪ theory)) + ( subset-axioms-weak-modal-logic) + ( subtype-union-left logic theory)) + ( refl-leq-subtype (weak-modal-logic-closure (logic ∪ theory))))) + ( ⊥) + ( bot-in-logic)) + + module _ + {l3 : Level} + (t@((theory , is-cons) , is-comp) : L-complete-theory (l1 ⊔ l2 ⊔ l3)) + where + + eq-union-L-consistent : + weak-modal-logic-closure (logic ∪ theory) = theory + eq-union-L-consistent = + ap modal-theory-L-consistent-theory + ( is-comp + ( union-L-consistent (theory , is-cons)) + ( transitive-leq-subtype + ( theory) + ( logic ∪ theory) + ( weak-modal-logic-closure (logic ∪ theory)) + ( subset-axioms-weak-modal-logic) + ( subtype-union-right logic theory))) + + subset-union-L-consistent : + weak-modal-logic-closure (logic ∪ theory) ⊆ theory + subset-union-L-consistent = + inv-tr + ( _⊆ theory) + ( eq-union-L-consistent) + ( refl-leq-subtype theory) + + subset-union-logic-L-complete-theory : + logic ∪ theory ⊆ theory + subset-union-logic-L-complete-theory = + transitive-leq-subtype + ( logic ∪ theory) + ( modal-theory-L-consistent-theory + ( union-L-consistent (theory , is-cons))) + ( theory) + ( subset-union-L-consistent) + ( subset-axioms-weak-modal-logic) + + subset-logic-L-complete-theory : + logic ⊆ theory + subset-logic-L-complete-theory = + transitive-leq-subtype + ( logic) + ( logic ∪ theory) + ( theory) + ( subset-union-logic-L-complete-theory) + ( subtype-union-left logic theory) + + is-weak-modal-logic-L-complete-theory : + is-weak-modal-logic theory + is-weak-modal-logic-L-complete-theory = + transitive-leq-subtype + ( weak-modal-logic-closure theory) + ( modal-theory-L-consistent-theory + ( union-L-consistent (theory , is-cons))) + ( theory) + ( subset-union-L-consistent) + ( weak-modal-logic-closure-monotic (subtype-union-right logic theory)) + + module _ + (contains-ax-k : ax-k i ⊆ logic) + (contains-ax-s : ax-s i ⊆ logic) + (contains-ax-dn : ax-dn i ⊆ logic) + where + + private + contains-ax-k' : ax-k i ⊆ theory + contains-ax-k' = + transitive-leq-subtype (ax-k i) logic theory + ( subset-logic-L-complete-theory) + ( contains-ax-k) + + contains-ax-s' : ax-s i ⊆ theory + contains-ax-s' = + transitive-leq-subtype (ax-s i) logic theory + ( subset-logic-L-complete-theory) + ( contains-ax-s) + + contains-ax-dn' : ax-dn i ⊆ theory + contains-ax-dn' = + transitive-leq-subtype (ax-dn i) logic theory + ( subset-logic-L-complete-theory) + ( contains-ax-dn) + + is-L-consistent-add-formula-not-in-logic : + {a : formula i} → + ¬ (is-in-subtype theory a) → + is-L-consistent-theory (theory-add-formula (~ a) theory) + is-L-consistent-add-formula-not-in-logic {a} not-in-logic bot-in-logic = + not-in-logic + ( weak-modal-logic-mp theory + ( is-weak-modal-logic-L-complete-theory) + ( contains-ax-dn' (~~ a →ₘ a) (a , refl)) + ( is-weak-modal-logic-L-complete-theory (~~ a) + ( forward-implication + ( deduction-lemma + ( theory) + ( contains-ax-k') + ( contains-ax-s') + ( ~ a) + ( ⊥)) + ( weak-modal-logic-closure-monotic + ( subtype-union-both logic (theory-add-formula (~ a) theory) + ( theory-add-formula (~ a) theory) + ( transitive-leq-subtype + ( logic) + ( theory) + ( theory-add-formula (~ a) theory) + ( subset-add-formula (~ a) theory) + ( subset-logic-L-complete-theory)) + ( refl-leq-subtype (theory-add-formula (~ a) theory))) + ( ⊥) + ( bot-in-logic))))) + + contains-negation-not-contains-formula-L-complete-theory : + {a : formula i} → + ¬ (is-in-subtype theory a) → + is-in-subtype theory (~ a) + contains-negation-not-contains-formula-L-complete-theory {a} not-in-logic = + tr + ( λ t → is-in-subtype t (~ a)) + ( eq-is-L-consistent-union-L-complete + {l3 = l3} {l4 = l1} + ( t) + ( Id-formula-Prop (~ a)) + ( is-L-consistent-add-formula-not-in-logic not-in-logic)) + ( formula-in-add-formula (~ a) theory) + + is-disjuctive-L-complete-theory : + LEM (l1 ⊔ l2 ⊔ l3) → + is-disjuctive-modal-theory theory + is-disjuctive-L-complete-theory lem a with lem (theory a) + ... | inl a-in-logic = inl a-in-logic + ... | inr a-not-in-logic = + inr + ( contains-negation-not-contains-formula-L-complete-theory + ( a-not-in-logic)) +``` diff --git a/src/modal-logic/canonical-model-theorem.lagda.md b/src/modal-logic/canonical-model-theorem.lagda.md index e53e8cf6de..f362423a75 100644 --- a/src/modal-logic/canonical-model-theorem.lagda.md +++ b/src/modal-logic/canonical-model-theorem.lagda.md @@ -54,11 +54,6 @@ open import order-theory.preorders open import order-theory.subposets open import order-theory.subtypes-leq-posets open import order-theory.top-elements-posets - -open import foundation.images -open import foundation.replacement -open import foundation.locally-small-types -open import foundation-core.small-types ```
diff --git a/src/modal-logic/logic-syntax.lagda.md b/src/modal-logic/logic-syntax.lagda.md index 2921dbe176..77775d4574 100644 --- a/src/modal-logic/logic-syntax.lagda.md +++ b/src/modal-logic/logic-syntax.lagda.md @@ -8,6 +8,7 @@ module modal-logic.logic-syntax where ```agda open import foundation.conjunction +open import foundation.coproduct-types open import foundation.dependent-pair-types open import foundation.empty-types open import foundation.function-types @@ -81,6 +82,25 @@ module _ is-consistent-modal-logic : modal-theory l2 i → UU l2 is-consistent-modal-logic = type-Prop ∘ is-consistent-modal-logic-Prop +module _ + {l1 : Level} {i : Set l1} + where + + is-contradictory-modal-logic-monotic : + {l2 l3 : Level} (ax₁ : modal-theory l2 i) (ax₂ : modal-theory l3 i) → + ax₁ ⊆ ax₂ → + is-contradictory-modal-logic ax₁ → + is-contradictory-modal-logic ax₂ + is-contradictory-modal-logic-monotic ax₁ ax₂ leq = leq ⊥ + + is-consistent-modal-logic-antimonotic : + {l2 l3 : Level} (ax₁ : modal-theory l2 i) (ax₂ : modal-theory l3 i) → + ax₁ ⊆ ax₂ → + is-consistent-modal-logic ax₂ → + is-consistent-modal-logic ax₁ + is-consistent-modal-logic-antimonotic ax₁ ax₂ leq is-cons = + is-cons ∘ is-contradictory-modal-logic-monotic ax₁ ax₂ leq + module _ {l1 l2 : Level} {i : Set l1} {axioms : formulas l2 i} where @@ -183,15 +203,57 @@ module _ ( modal-logic-monotic leq) module _ - {l1 l2 : Level} {i : Set l1} (a : formula i) (logic : formulas l2 i) + {l1 l2 : Level} {i : Set l1} (a : formula i) (theory : modal-theory l2 i) where + -- TODO: make Id-formula to be a function for 1 element modal-theory theory-add-formula : formulas (l1 ⊔ l2) i - theory-add-formula = union-subtype (Id-formula-Prop a) logic + theory-add-formula = (Id-formula-Prop a) ∪ theory formula-in-add-formula : is-in-subtype theory-add-formula a - formula-in-add-formula = subtype-union-left (Id-formula-Prop a) logic a refl + formula-in-add-formula = subtype-union-left (Id-formula-Prop a) theory a refl + + subset-add-formula : theory ⊆ theory-add-formula + subset-add-formula = subtype-union-right (Id-formula-Prop a) theory + + elim-theory-add-formula : + {l3 : Level} (P : formula i → Prop l3) → + type-Prop (P a) → + ((x : formula i) → is-in-subtype theory x → type-Prop (P x)) → + (x : formula i) → is-in-subtype theory-add-formula x → type-Prop (P x) + elim-theory-add-formula P H-a H-rest = + elim-union-subtype (Id-formula-Prop a) theory P + ( λ { .a refl → H-a }) + ( H-rest) + + subset-theory-add-formula : + {l3 : Level} (theory' : formulas l3 i) → + is-in-subtype theory' a → + theory ⊆ theory' → + theory-add-formula ⊆ theory' + subset-theory-add-formula theory' a-in = + subtype-union-both + ( Id-formula-Prop a) + ( theory) + ( theory') + ( λ { .a refl → a-in }) + +module _ + {l1 : Level} {i : Set l1} + where - subset-add-formula : logic ⊆ theory-add-formula - subset-add-formula = subtype-union-right (Id-formula-Prop a) logic + is-disjuctive-modal-theory : + {l2 : Level} → modal-theory l2 i → UU (l1 ⊔ l2) + is-disjuctive-modal-theory theory = + (a : formula i) → is-in-subtype theory a + is-in-subtype theory (~ a) + + theory-add-formula-union-right : + (a : formula i) + {l2 l3 : Level} + (theory : formulas l2 i) + (theory' : formulas l3 i) → + theory ∪ theory-add-formula a theory' ⊆ + theory-add-formula a (theory ∪ theory') + theory-add-formula-union-right a theory theory' = + union-swap-1-2 theory (Id-formula-Prop a) theory' ``` diff --git a/src/modal-logic/weak-deduction.lagda.md b/src/modal-logic/weak-deduction.lagda.md index 43bfea902a..823f8e3a3a 100644 --- a/src/modal-logic/weak-deduction.lagda.md +++ b/src/modal-logic/weak-deduction.lagda.md @@ -9,6 +9,7 @@ module modal-logic.weak-deduction where ```agda open import foundation.cartesian-product-types open import foundation.coproduct-types +open import foundation.conjunction open import foundation.dependent-pair-types open import foundation.disjunction open import foundation.empty-types @@ -49,615 +50,967 @@ module _ {l1 l2 : Level} {i : Set l1} where - infix 5 _▷_ + is-weak-deduction-Prop : + {axioms : modal-theory l2 i} {a : formula i} → axioms ⊢ a → Prop lzero + is-weak-deduction-Prop (ax x) = unit-Prop + is-weak-deduction-Prop (mp dab da) = + conjunction-Prop (is-weak-deduction-Prop dab) (is-weak-deduction-Prop da) + is-weak-deduction-Prop (nec d) = empty-Prop - data _▷_ (axioms : formulas l2 i) : formula i → UU (l1 ⊔ l2) where - w-ax : {a : formula i} → is-in-subtype axioms a → axioms ▷ a - w-mp : {a b : formula i} → axioms ▷ a →ₘ b → axioms ▷ a → axioms ▷ b + is-weak-deduction : + {axioms : modal-theory l2 i} {a : formula i} → axioms ⊢ a → UU lzero + is-weak-deduction = type-Prop ∘ is-weak-deduction-Prop - weak-modal-logic : formulas l2 i → formulas (l1 ⊔ l2) i - weak-modal-logic axioms a = trunc-Prop (axioms ▷ a) + infix 5 _⊢w_ + + _⊢w_ : modal-theory l2 i → formula i → UU (l1 ⊔ l2) + axioms ⊢w a = type-subtype (is-weak-deduction-Prop {axioms} {a}) + + deduction-weak-deduction : + {axioms : modal-theory l2 i} {a : formula i} → axioms ⊢w a → axioms ⊢ a + deduction-weak-deduction = inclusion-subtype is-weak-deduction-Prop + + is-weak-deduction-deduction-weak-deduction : + {axioms : modal-theory l2 i} {a : formula i} (d : axioms ⊢w a) → + is-weak-deduction (deduction-weak-deduction d) + is-weak-deduction-deduction-weak-deduction = + is-in-subtype-inclusion-subtype is-weak-deduction-Prop + + weak-deduction-ax : + {axioms : modal-theory l2 i} {a : formula i} → + is-in-subtype axioms a → + axioms ⊢w a + weak-deduction-ax in-axioms = ax in-axioms , star + + weak-deduction-mp : + {axioms : modal-theory l2 i} {a b : formula i} → + axioms ⊢w (a →ₘ b) → + axioms ⊢w a → + axioms ⊢w b + weak-deduction-mp (dab , is-w-dab) (da , is-w-da) = + mp dab da , is-w-dab , is-w-da + + -- ind-weak-deduction : + -- {l : Level} + -- {P : {axioms : modal-theory l2 i} {a : formula i} → axioms ⊢w a → UU l} → + -- ( {axioms : modal-theory l2 i} {a : formula i} + -- (in-axioms : is-in-subtype axioms a) → + -- P {axioms} (weak-deduction-ax in-axioms)) → + -- ( {axioms : modal-theory l2 i} {a b : formula i} + -- (dab : axioms ⊢w a →ₘ b) (da : axioms ⊢w a) → + -- P dab → P da → P (weak-deduction-mp dab da)) → + -- {axioms : modal-theory l2 i} {a : formula i} (wd : axioms ⊢w a) → P wd + -- ind-weak-deduction H-ax H-mp (ax in-axioms , _) = + -- H-ax in-axioms + -- ind-weak-deduction {P = P} H-ax H-mp (mp dba db , is-w-dba , is-w-db) = + -- H-mp + -- ( dba , is-w-dba) + -- ( db , is-w-db) + -- ( ind-weak-deduction {P = P} H-ax H-mp (dba , is-w-dba)) + -- ( ind-weak-deduction {P = P} H-ax H-mp (db , is-w-db)) + + -- rec-weak-deduction : + -- {l : Level} + -- {P : UU l} → + -- ( {axioms : modal-theory l2 i} {a : formula i} → + -- is-in-subtype axioms a → P) → + -- ( {axioms : modal-theory l2 i} {a b : formula i} + -- (dab : axioms ⊢w a →ₘ b) (da : axioms ⊢w a) → + -- P → P → P) → + -- {axioms : modal-theory l2 i} {a : formula i} → + -- axioms ⊢w a → P + -- rec-weak-deduction = ind-weak-deduction + + ind-weak-deduction : + {l : Level} {axioms : modal-theory l2 i} + {P : {a : formula i} → axioms ⊢w a → UU l} → + ( {a : formula i} (in-axioms : is-in-subtype axioms a) → + P (weak-deduction-ax in-axioms)) → + ( {a b : formula i} {dab : axioms ⊢w a →ₘ b} {da : axioms ⊢w a} → + P dab → P da → P (weak-deduction-mp dab da)) → + {a : formula i} (wd : axioms ⊢w a) → P wd + ind-weak-deduction H-ax H-mp (ax in-axioms , _) = + H-ax in-axioms + ind-weak-deduction {P = P} H-ax H-mp (mp dba db , is-w-dba , is-w-db) = + H-mp + ( ind-weak-deduction {P = P} H-ax H-mp (dba , is-w-dba)) + ( ind-weak-deduction {P = P} H-ax H-mp (db , is-w-db)) + + weak-modal-logic-closure : modal-theory l2 i → modal-theory (l1 ⊔ l2) i + weak-modal-logic-closure axioms a = trunc-Prop (axioms ⊢w a) + + is-weak-modal-logic-Prop : modal-theory l2 i → Prop (l1 ⊔ l2) + is-weak-modal-logic-Prop theory = + leq-prop-subtype (weak-modal-logic-closure theory) theory + -- implicit-Π-Prop (formula i) (λ a → function-Prop (theory ⊢w a) (theory a)) + + is-weak-modal-logic : modal-theory l2 i → UU (l1 ⊔ l2) + is-weak-modal-logic = type-Prop ∘ is-weak-modal-logic-Prop + + is-in-weak-deduction-closure-weak-deduction : + {axioms : modal-theory l2 i} {a : formula i} → + axioms ⊢w a → is-in-subtype (weak-modal-logic-closure axioms) a + is-in-weak-deduction-closure-weak-deduction = unit-trunc-Prop module _ - {l1 l2 : Level} {i : Set l1} {axioms : formulas l2 i} + {l1 l2 : Level} {i : Set l1} (axioms : modal-theory l2 i) where - weak-modal-logic-ax : + weak-modal-logic-closure-ax : {a : formula i} → is-in-subtype axioms a → - is-in-subtype (weak-modal-logic axioms) a - weak-modal-logic-ax = unit-trunc-Prop ∘ w-ax + is-in-subtype (weak-modal-logic-closure axioms) a + weak-modal-logic-closure-ax = + is-in-weak-deduction-closure-weak-deduction ∘ weak-deduction-ax - weak-modal-logic-mp : + weak-modal-logic-closure-mp : {a b : formula i} → - is-in-subtype (weak-modal-logic axioms) (a →ₘ b) → - is-in-subtype (weak-modal-logic axioms) a → - is-in-subtype (weak-modal-logic axioms) b - weak-modal-logic-mp {a} {b} twdab twda = + is-in-subtype (weak-modal-logic-closure axioms) (a →ₘ b) → + is-in-subtype (weak-modal-logic-closure axioms) a → + is-in-subtype (weak-modal-logic-closure axioms) b + weak-modal-logic-closure-mp {a} {b} twdab twda = apply-twice-universal-property-trunc-Prop twdab twda - ( weak-modal-logic axioms b) - ( λ wdab wda → unit-trunc-Prop (w-mp wdab wda)) + ( weak-modal-logic-closure axioms b) + ( λ wdab wda → + ( is-in-weak-deduction-closure-weak-deduction + ( weak-deduction-mp wdab wda))) module _ - {l1 l2 : Level} {i : Set l1} (axioms : formulas l2 i) - where - - deduction-weak-deduction : {a : formula i} → axioms ▷ a → axioms ⊢ a - deduction-weak-deduction (w-ax x) = ax x - deduction-weak-deduction (w-mp dab da) = - mp (deduction-weak-deduction dab) (deduction-weak-deduction da) - - weak-modal-logic-subset-modal-logic : - weak-modal-logic axioms ⊆ modal-logic axioms - weak-modal-logic-subset-modal-logic a = - map-universal-property-trunc-Prop - ( modal-logic axioms a) - ( unit-trunc-Prop ∘ deduction-weak-deduction) - -module _ - {l1 l2 l3 : Level} {i : Set l1} - {axioms₁ : formulas l2 i} {axioms₂ : formulas l3 i} - (e : equiv-subtypes axioms₁ axioms₂) + {l1 l2 : Level} {i : Set l1} + (theory : modal-theory l2 i) (is-w : is-weak-modal-logic theory) where - deduction-equiv-axioms : - {a : formula i} → axioms₁ ▷ a → axioms₂ ▷ a - deduction-equiv-axioms {a} (w-ax x) = - w-ax (map-equiv (e a) x) - deduction-equiv-axioms (w-mp wdab wda) = - w-mp (deduction-equiv-axioms wdab) (deduction-equiv-axioms wda) + weak-modal-logic-mp : + {a b : formula i} → + is-in-subtype theory (a →ₘ b) → + is-in-subtype theory a → + is-in-subtype theory b + weak-modal-logic-mp {a} {b} dab da = + is-w b + ( weak-modal-logic-closure-mp theory + ( weak-modal-logic-closure-ax theory dab) + ( weak-modal-logic-closure-ax theory da)) module _ - {l1 : Level} {i : Set l1} + {l1 l2 : Level} {i : Set l1} {axioms : modal-theory l2 i} where - axioms-subset-weak-modal-logic : - {l2 : Level} (axioms : formulas l2 i) → axioms ⊆ weak-modal-logic axioms - axioms-subset-weak-modal-logic _ a H = unit-trunc-Prop (w-ax H) - weak-modal-logic-closed : - {l2 : Level} {axioms : formulas l2 i} {a : formula i} → - weak-modal-logic axioms ▷ a → - is-in-subtype (weak-modal-logic axioms) a - weak-modal-logic-closed (w-ax x) = x - weak-modal-logic-closed {axioms = axioms} {a} (w-mp tdab tda) = - apply-twice-universal-property-trunc-Prop - ( weak-modal-logic-closed tdab) - ( weak-modal-logic-closed tda) - ( weak-modal-logic axioms a) - (λ dab da → unit-trunc-Prop (w-mp dab da)) - - subset-double-weak-modal-logic : - {l2 : Level} - (axioms : formulas l2 i) → - weak-modal-logic (weak-modal-logic axioms) ⊆ weak-modal-logic axioms - subset-double-weak-modal-logic axioms a = + {a : formula i} → weak-modal-logic-closure axioms ⊢w a → + is-in-subtype (weak-modal-logic-closure axioms) a + weak-modal-logic-closed = + ind-weak-deduction + ( λ in-axioms → in-axioms) + ( λ in-logic-ab in-logic-a → + ( weak-modal-logic-closure-mp axioms in-logic-ab in-logic-a)) + + is-weak-modal-logic-weak-modal-logic-closure : + is-weak-modal-logic (weak-modal-logic-closure axioms) + is-weak-modal-logic-weak-modal-logic-closure a = map-universal-property-trunc-Prop - ( weak-modal-logic axioms a) + ( weak-modal-logic-closure axioms a) ( weak-modal-logic-closed) - weak-modal-logic-idempotent : - {l2 : Level} - (axioms : formulas l2 i) → - weak-modal-logic axioms = weak-modal-logic (weak-modal-logic axioms) - weak-modal-logic-idempotent axioms = - antisymmetric-leq-subtype _ _ - ( axioms-subset-weak-modal-logic (weak-modal-logic axioms)) - ( subset-double-weak-modal-logic axioms) + -- weak-modal-logic-closure-closed : + -- is-weak-modal-logic axioms → + -- weak-modal-logic-closure axioms ⊆ axioms + -- weak-modal-logic-closure-closed is-logic a = + -- map-universal-property-trunc-Prop (axioms a) is-logic + +-- module _ +-- {l1 l2 : Level} {i : Set l1} {axioms : modal-theory l2 i} +-- where + + subset-axioms-weak-modal-logic : axioms ⊆ weak-modal-logic-closure axioms + subset-axioms-weak-modal-logic a = weak-modal-logic-closure-ax axioms + +-- subset-double-weak-modal-logic-closure : +-- weak-modal-logic-closure (weak-modal-logic-closure axioms) ⊆ +-- weak-modal-logic-closure axioms +-- subset-double-weak-modal-logic-closure = +-- weak-modal-logic-closure-closed is-weak-modal-logic-weak-modal-logic-closure module _ - {l1 l2 l3 : Level} {i : Set l1} {ax₁ : formulas l2 i} {ax₂ : formulas l3 i} + {l1 l2 l3 : Level} {i : Set l1} + {ax₁ : modal-theory l2 i} + {ax₂ : modal-theory l3 i} (leq : ax₁ ⊆ ax₂) where - weak-deduction-monotic : {a : formula i} → ax₁ ▷ a → ax₂ ▷ a - weak-deduction-monotic (w-ax x) = w-ax (leq _ x) - weak-deduction-monotic (w-mp wdab wda) = - w-mp (weak-deduction-monotic wdab) (weak-deduction-monotic wda) + weak-deduction-monotic : {a : formula i} → ax₁ ⊢w a → ax₂ ⊢w a + weak-deduction-monotic = + ind-weak-deduction + ( λ {a} in-axioms → weak-deduction-ax (leq a in-axioms)) + ( λ dab da → weak-deduction-mp dab da) - weak-modal-logic-monotic : weak-modal-logic ax₁ ⊆ weak-modal-logic ax₂ - weak-modal-logic-monotic a = + weak-modal-logic-closure-monotic : + weak-modal-logic-closure ax₁ ⊆ weak-modal-logic-closure ax₂ + weak-modal-logic-closure-monotic a = map-universal-property-trunc-Prop - ( weak-modal-logic ax₂ a) + ( weak-modal-logic-closure ax₂ a) ( unit-trunc-Prop ∘ weak-deduction-monotic) module _ - {l1 l2 l3 : Level} {i : Set l1} {ax₁ : formulas l2 i} {ax₂ : formulas l3 i} + {l1 l2 l3 : Level} {i : Set l1} + {ax₁ : modal-theory l2 i} + {ax₂ : modal-theory l3 i} where - subset-weak-modal-logic-if-subset-axioms : - ax₁ ⊆ weak-modal-logic ax₂ → weak-modal-logic ax₁ ⊆ weak-modal-logic ax₂ - subset-weak-modal-logic-if-subset-axioms leq = + subset-weak-modal-logic-subset-axioms : + ax₁ ⊆ weak-modal-logic-closure ax₂ → + weak-modal-logic-closure ax₁ ⊆ weak-modal-logic-closure ax₂ + subset-weak-modal-logic-subset-axioms leq = transitive-leq-subtype - ( weak-modal-logic ax₁) - ( weak-modal-logic (weak-modal-logic ax₂)) - ( weak-modal-logic ax₂) - ( subset-double-weak-modal-logic ax₂) - ( weak-modal-logic-monotic leq) - --- TODO: move to lists - -in-list : {l : Level} {A : UU l} → list A → A → Prop l -in-list l a = trunc-Prop (a ∈-list l) - -subset-in-tail : - {l : Level} {A : UU l} (a : A) (l : list A) → - in-list l ⊆ in-list (cons a l) -subset-in-tail a l x = - map-universal-property-trunc-Prop - ( in-list (cons a l) x) - ( unit-trunc-Prop ∘ (is-in-tail x a l)) - -in-concat-list : - {l : Level} {A : UU l} (l1 l2 : list A) → - (a : A) → a ∈-list (concat-list l1 l2) → (a ∈-list l1) + (a ∈-list l2) -in-concat-list nil l2 a in-list = inr in-list -in-concat-list (cons x l1) l2 a (is-head a _) = inl (is-head a l1) -in-concat-list (cons x l1) l2 a (is-in-tail _ _ _ in-tail) - with in-concat-list l1 l2 a in-tail -... | inl in-l1 = inl (is-in-tail a x l1 in-l1) -... | inr in-l2 = inr in-l2 - -subset-reversing-list : - {l : Level} {A : UU l} (l : list A) → in-list l ⊆ in-list (reverse-list l) -subset-reversing-list l x = - map-universal-property-trunc-Prop - ( in-list (reverse-list l) x) - ( unit-trunc-Prop ∘ forward-contains-reverse-list x l) - -in-sum-concat-list-Prop : - {l : Level} {A : UU l} (l1 l2 : list A) → - (a : A) → a ∈-list (concat-list l1 l2) → - type-Prop (in-list l1 a) + type-Prop (in-list l2 a) -in-sum-concat-list-Prop l1 l2 a in-concat with in-concat-list l1 l2 a in-concat -... | inl in-l1 = inl (unit-trunc-Prop in-l1) -... | inr in-l2 = inr (unit-trunc-Prop in-l2) - -union-lists-concat : - {l : Level} {A : UU l} (l1 l2 : list A) → - in-list (concat-list l1 l2) ⊆ union-subtype (in-list l1) (in-list l2) -union-lists-concat l1 l2 a = - map-universal-property-trunc-Prop - ( union-subtype (in-list l1) (in-list l2) a) - ( unit-trunc-Prop ∘ (in-sum-concat-list-Prop l1 l2 a)) - -in-concat-left : - {l : Level} {A : UU l} (l1 l2 : list A) - {a : A} → a ∈-list l1 → a ∈-list (concat-list l1 l2) -in-concat-left _ _ (is-head a _) = - is-head a _ -in-concat-left _ l2 (is-in-tail a x l1 in-l1) = - is-in-tail a x (concat-list l1 l2) (in-concat-left l1 l2 in-l1) - -in-concat-right : - {l : Level} {A : UU l} (l1 l2 : list A) - {a : A} → a ∈-list l2 → a ∈-list (concat-list l1 l2) -in-concat-right nil l2 in-l2 = in-l2 -in-concat-right (cons x l1) l2 in-l2 = - is-in-tail _ x (concat-list l1 l2) (in-concat-right l1 l2 in-l2) - -in-concat-list-sum-Prop : - {l : Level} {A : UU l} (l1 l2 : list A) → - (a : A) → type-Prop (in-list l1 a) + type-Prop (in-list l2 a) → - type-Prop (in-list (concat-list l1 l2) a) -in-concat-list-sum-Prop l1 l2 a (inl in-l1) = - apply-universal-property-trunc-Prop - ( in-l1) - ( in-list (concat-list l1 l2) a) - ( unit-trunc-Prop ∘ (in-concat-left l1 l2)) -in-concat-list-sum-Prop l1 l2 a (inr in-l2) = - apply-universal-property-trunc-Prop - ( in-l2) - ( in-list (concat-list l1 l2) a) - ( unit-trunc-Prop ∘ (in-concat-right l1 l2)) - -in-concat-lists-union : - {l : Level} {A : UU l} (l1 l2 : list A) → - union-subtype (in-list l1) (in-list l2) ⊆ in-list (concat-list l1 l2) -in-concat-lists-union l1 l2 a = - map-universal-property-trunc-Prop - ( in-list (concat-list l1 l2) a) - ( in-concat-list-sum-Prop l1 l2 a) - -subset-in-concat-left : - {l : Level} {A : UU l} (l1 l2 : list A) → - in-list l1 ⊆ in-list (concat-list l1 l2) -subset-in-concat-left l1 l2 = - transitive-leq-subtype - ( in-list l1) - ( union-subtype (in-list l1) (in-list l2)) - ( in-list (concat-list l1 l2)) - ( in-concat-lists-union l1 l2) - ( subtype-union-left (in-list l1) (in-list l2)) - -subset-in-concat-right : - {l : Level} {A : UU l} (l1 l2 : list A) → - in-list l2 ⊆ in-list (concat-list l1 l2) -subset-in-concat-right l1 l2 = - transitive-leq-subtype - ( in-list l2) - ( union-subtype (in-list l1) (in-list l2)) - ( in-list (concat-list l1 l2)) - ( in-concat-lists-union l1 l2) - ( subtype-union-right (in-list l1) (in-list l2)) - -empty-in-list-nil : - {l : Level} {A : UU l} {x : A} → is-in-subtype (in-list nil) x → empty -empty-in-list-nil = map-universal-property-trunc-Prop empty-Prop ( λ ()) - -equiv-subset-head-tail : - {l1 l2 : Level} {A : UU l1} (x : A) (xs : list A) (a : subtype l2 A) → - (leq-prop-subtype (in-list (cons x xs)) a) ⇔ - product-Prop (a x) (leq-prop-subtype (in-list xs) a) -pr1 (equiv-subset-head-tail x xs a) sub = - pair - ( sub x (unit-trunc-Prop (is-head x xs))) - ( transitive-leq-subtype - ( in-list xs) - ( in-list (cons x xs)) - ( a) - ( sub) - ( subset-in-tail x xs)) -pr2 (equiv-subset-head-tail x xs a) (x-in-a , sub) y = - map-universal-property-trunc-Prop - ( a y) - ( λ - { (is-head .x .xs) → x-in-a - ; (is-in-tail .y .x .xs y-in-list) → sub y (unit-trunc-Prop y-in-list)}) - -lists-in-union-lists : - {l1 l2 l3 : Level} {A : UU l1} - (xs : list A) (a : subtype l2 A) (b : subtype l3 A) → - in-list xs ⊆ union-subtype a b → - ∃ - ( list A) - ( λ xsl → - ( Σ - ( list A) - ( λ xsr → - ( in-list xs ⊆ union-subtype (in-list xsl) (in-list xsr)) × - ( in-list xsl ⊆ a) × - ( in-list xsr ⊆ b)))) -lists-in-union-lists nil a b sub = - unit-trunc-Prop - ( triple nil nil - ( triple - ( λ _ → ex-falso ∘ empty-in-list-nil) - ( λ _ → ex-falso ∘ empty-in-list-nil) - ( λ _ → ex-falso ∘ empty-in-list-nil))) -lists-in-union-lists (cons x xs) a b sub = - apply-twice-universal-property-trunc-Prop - ( lists-in-union-lists xs a b - ( transitive-leq-subtype - ( in-list xs) - ( in-list (cons x xs)) - ( union-subtype a b) - ( sub) - ( subset-in-tail x xs))) - ( sub x (unit-trunc-Prop (is-head x xs))) - ( ∃-Prop _ (λ _ → Σ _ _)) - ( λ (xsl , xsr , sub-lists , sub-xsl , sub-xsr) → - ( ind-coproduct - ( λ _ → _) - ( λ x-in-a → - ( unit-trunc-Prop - ( triple (cons x xsl) xsr - ( triple - ( λ y → - ( map-universal-property-trunc-Prop - ( union-subtype (in-list (cons x xsl)) (in-list xsr) y) - ( λ - { (is-head .x .xs) → - ( subtype-union-left - ( in-list (cons x xsl)) - ( in-list xsr) - ( y) - ( unit-trunc-Prop (is-head y xsl))) - ; (is-in-tail .y .x .xs y-in-xs) → - ( subset-union-subset-left - ( in-list xsl) - ( in-list (cons x xsl)) - ( in-list xsr) - ( subset-in-tail x xsl) - ( y) - ( sub-lists y (unit-trunc-Prop y-in-xs)))}))) - ( backward-implication - ( equiv-subset-head-tail x xsl a) - ( x-in-a , sub-xsl)) - ( sub-xsr))))) - ( λ x-in-b → - ( unit-trunc-Prop - ( triple xsl (cons x xsr) - ( triple - ( λ y → - ( map-universal-property-trunc-Prop - ( union-subtype (in-list xsl) (in-list (cons x xsr)) y) - ( λ - { (is-head .x .xs) → - ( subtype-union-right - ( in-list xsl) - ( in-list (cons x xsr)) - ( y) - ( unit-trunc-Prop (is-head y xsr))) - ; (is-in-tail .y .x .xs y-in-xs) → - ( subset-union-subset-right - ( in-list xsl) - ( in-list xsr) - ( in-list (cons x xsr)) - ( subset-in-tail x xsr) - ( y) - ( sub-lists y (unit-trunc-Prop y-in-xs)))}))) - ( sub-xsl) - ( backward-implication - ( equiv-subset-head-tail x xsr b) - ( x-in-b , sub-xsr)))))))) - -module _ - {l1 : Level} {i : Set l1} - where - - list-assumptions-weak-deduction : - {l2 : Level} {logic : formulas l2 i} {a : formula i} → - logic ▷ a → - Σ (list (formula i)) (λ l → (in-list l ⊆ logic) × (in-list l ▷ a)) - pr1 (list-assumptions-weak-deduction {a = a} (w-ax x)) = - unit-list a - pr1 (pr2 (list-assumptions-weak-deduction {logic = logic} (w-ax x))) b = - map-universal-property-trunc-Prop ( logic b) ( λ { (is-head _ _) → x}) - pr2 (pr2 (list-assumptions-weak-deduction {a = a} (w-ax x))) = - w-ax (unit-trunc-Prop (is-head a nil)) - pr1 (list-assumptions-weak-deduction (w-mp wdcb wdc)) = - concat-list - ( pr1 (list-assumptions-weak-deduction wdcb)) - ( pr1 (list-assumptions-weak-deduction wdc)) - pr1 (pr2 (list-assumptions-weak-deduction {logic = logic} (w-mp wdcb wdc))) - b b-in-list = - subtype-union-both - ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) - ( in-list (pr1 (list-assumptions-weak-deduction wdc))) - ( logic) - ( pr1 (pr2 (list-assumptions-weak-deduction wdcb))) - ( pr1 (pr2 (list-assumptions-weak-deduction wdc))) - ( b) - ( union-lists-concat - ( pr1 (list-assumptions-weak-deduction wdcb)) - ( pr1 (list-assumptions-weak-deduction wdc)) - ( b) - ( b-in-list)) - pr2 (pr2 (list-assumptions-weak-deduction (w-mp wdcb wdc))) = - w-mp - ( weak-deduction-monotic - ( transitive-leq-subtype - ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) - ( union-subtype - ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) - ( in-list (pr1 (list-assumptions-weak-deduction wdc)))) - ( in-list - ( concat-list - ( pr1 (list-assumptions-weak-deduction wdcb)) - ( pr1 (list-assumptions-weak-deduction wdc)))) - ( in-concat-lists-union - ( pr1 (list-assumptions-weak-deduction wdcb)) - ( pr1 (list-assumptions-weak-deduction wdc))) - ( subtype-union-left - ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) - ( in-list (pr1 (list-assumptions-weak-deduction wdc))))) - ( pr2 (pr2 (list-assumptions-weak-deduction wdcb)))) - ( weak-deduction-monotic - ( transitive-leq-subtype - ( in-list (pr1 (list-assumptions-weak-deduction wdc))) - ( union-subtype - ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) - ( in-list (pr1 (list-assumptions-weak-deduction wdc)))) - ( in-list - ( concat-list - ( pr1 (list-assumptions-weak-deduction wdcb)) - ( pr1 (list-assumptions-weak-deduction wdc)))) - ( in-concat-lists-union - ( pr1 (list-assumptions-weak-deduction wdcb)) - ( pr1 (list-assumptions-weak-deduction wdc))) - ( subtype-union-right - ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) - ( in-list (pr1 (list-assumptions-weak-deduction wdc))))) - ( pr2 (pr2 (list-assumptions-weak-deduction wdc)))) - - nil-no-deduction : {a : formula i} → ¬ (in-list nil ▷ a) - nil-no-deduction (w-ax x) = - apply-universal-property-trunc-Prop x empty-Prop (λ ()) - nil-no-deduction (w-mp wd _) = nil-no-deduction wd - - -- list-assumptions-weak-modal-logic : - -- {l2 : Level} {logic : formulas l2 i} {a : formula i} → - -- is-in-subtype (weak-modal-logic logic) a → - -- ∃ ( list (formula i)) - -- ( λ l → - -- ( in-list l ⊆ logic) × - -- ( is-in-subtype (weak-modal-logic (in-list l)) a)) - -- list-assumptions-weak-modal-logic {l2} {logic} {a} = - -- map-universal-property-trunc-Prop - -- ( ∃-Prop - -- ( list (formula i)) - -- ( λ l → - -- ( in-list l ⊆ logic) × - -- ( is-in-subtype (weak-modal-logic (in-list l)) a)) ) - -- ( λ wda → - -- ( let (l , sub , comp) = list-assumptions-weak-deduction wda - -- in unit-trunc-Prop (l , sub , (unit-trunc-Prop comp)))) - - forward-subset-head-add : - (a : formula i) (l : list (formula i)) → - theory-add-formula a (in-list l) ⊆ in-list (cons a l) - forward-subset-head-add a l = - subtype-union-both - ( Id-formula-Prop a) - ( in-list l) - ( in-list (cons a l)) - ( λ { .a refl → unit-trunc-Prop (is-head a l)}) - ( subset-in-tail a l) - - backward-subset-head-add : - (a : formula i) (l : list (formula i)) → - in-list (cons a l) ⊆ theory-add-formula a (in-list l) - backward-subset-head-add a l x = - map-universal-property-trunc-Prop - ( theory-add-formula a (in-list l) x) - ( λ - { (is-head .a .l) → - ( subtype-union-left (Id-formula-Prop a) (in-list l) a refl) - ; (is-in-tail .x .a .l t) → - ( subtype-union-right - ( Id-formula-Prop a) - ( in-list l) - ( x) - ( unit-trunc-Prop t))}) + ( weak-modal-logic-closure ax₁) + ( weak-modal-logic-closure (weak-modal-logic-closure ax₂)) + ( weak-modal-logic-closure ax₂) + ( is-weak-modal-logic-weak-modal-logic-closure) + ( weak-modal-logic-closure-monotic leq) module _ - {l1 l2 : Level} {i : Set l1} (axioms : formulas l2 i) + {l1 l2 : Level} {i : Set l1} + (axioms : modal-theory l2 i) + -- (is-w : is-weak-modal-logic logic) where backward-deduction-lemma : {a b : formula i} → - axioms ▷ a →ₘ b → - is-in-subtype (weak-modal-logic (theory-add-formula a axioms)) b + axioms ⊢w a →ₘ b → + theory-add-formula a axioms ⊢w b backward-deduction-lemma {a} wab = - unit-trunc-Prop - ( w-mp - ( weak-deduction-monotic - ( subtype-union-right ((Id-formula-Prop a)) axioms) wab) - ( w-ax (subtype-union-left (Id-formula-Prop a) axioms a refl))) + weak-deduction-mp + ( weak-deduction-monotic + ( subset-add-formula a axioms) + ( wab)) + ( weak-deduction-ax (formula-in-add-formula a axioms)) module _ - -- TODO: maybe just in axioms? - (contains-ax-k : ax-k i ⊆ weak-modal-logic axioms) - (contains-ax-s : ax-s i ⊆ weak-modal-logic axioms) + (contains-ax-k : ax-k i ⊆ axioms) + (contains-ax-s : ax-s i ⊆ axioms) where + -- TODO: move to file with deduction deduction-a->a : - (a : formula i) → is-in-subtype (weak-modal-logic axioms) (a →ₘ a) + (a : formula i) → axioms ⊢w a →ₘ a deduction-a->a a = - apply-three-times-universal-property-trunc-Prop - ( contains-ax-s _ (a , a →ₘ a , a , refl)) - ( contains-ax-k _ (a , a →ₘ a , refl)) - ( contains-ax-k _ (a , a , refl)) - ( (weak-modal-logic axioms) (a →ₘ a)) - ( λ x y z → unit-trunc-Prop (w-mp (w-mp x y) z)) + weak-deduction-mp + ( weak-deduction-mp + ( weak-deduction-ax (contains-ax-s _ (a , a →ₘ a , a , refl))) + ( weak-deduction-ax (contains-ax-k _ (a , a →ₘ a , refl)))) + ( weak-deduction-ax (contains-ax-k _ (a , a , refl))) forward-deduction-lemma : (a : formula i) {b : formula i} → - theory-add-formula a axioms ▷ b → - is-in-subtype (weak-modal-logic axioms) (a →ₘ b) - forward-deduction-lemma a {b} (w-ax x) = - elim-disjunction-Prop - ( Id-formula-Prop a b) - ( axioms b) - ( (weak-modal-logic axioms) (a →ₘ b)) - ( (λ { refl → deduction-a->a a}) - , (λ in-axioms → - ( apply-universal-property-trunc-Prop - ( contains-ax-k _ (b , a , refl)) - ( (weak-modal-logic axioms) (a →ₘ b)) - ( λ x → unit-trunc-Prop (w-mp x (w-ax in-axioms)))))) - ( x) - forward-deduction-lemma a {b} (w-mp {c} twdcb twdc) = - apply-three-times-universal-property-trunc-Prop - ( forward-deduction-lemma a twdcb) - ( forward-deduction-lemma a twdc) - ( contains-ax-s _ (a , c , b , refl)) - ( (weak-modal-logic axioms) (a →ₘ b)) - ( λ wdacb wdac x → - ( unit-trunc-Prop (w-mp (w-mp x wdacb) wdac))) + theory-add-formula a axioms ⊢w b → + is-in-subtype (weak-modal-logic-closure axioms) (a →ₘ b) + forward-deduction-lemma a = + ind-weak-deduction + ( λ {b} b-in-axioms → + ( elim-theory-add-formula a axioms + ( λ x → weak-modal-logic-closure axioms (a →ₘ x)) + ( is-in-weak-deduction-closure-weak-deduction (deduction-a->a a)) + ( λ x x-in-axioms → + ( weak-modal-logic-closure-mp axioms + ( weak-modal-logic-closure-ax axioms + ( contains-ax-k (x →ₘ a →ₘ x) (x , a , refl))) + ( weak-modal-logic-closure-ax axioms x-in-axioms))) + ( b) + ( b-in-axioms))) + ( λ {b} {c} dabc dab → + ( weak-modal-logic-closure-mp axioms + ( weak-modal-logic-closure-mp axioms + ( weak-modal-logic-closure-ax axioms + ( contains-ax-s _ (a , b , c , refl))) + ( dabc)) + ( dab))) deduction-lemma : (a b : formula i) → - weak-modal-logic (theory-add-formula a axioms) b ⇔ - weak-modal-logic axioms (a →ₘ b) + weak-modal-logic-closure (theory-add-formula a axioms) b ⇔ + weak-modal-logic-closure axioms (a →ₘ b) pr1 (deduction-lemma a b) = map-universal-property-trunc-Prop - ( (weak-modal-logic axioms) (a →ₘ b)) + ( weak-modal-logic-closure axioms (a →ₘ b)) ( forward-deduction-lemma a) pr2 (deduction-lemma a b) = map-universal-property-trunc-Prop - ( weak-modal-logic (theory-add-formula a axioms) b) - ( backward-deduction-lemma) + ( weak-modal-logic-closure (theory-add-formula a axioms) b) + ( is-in-weak-deduction-closure-weak-deduction ∘ + backward-deduction-lemma) + +-- module _ +-- -- TODO: maybe just in axioms? +-- (contains-ax-k : ax-k i ⊆ weak-modal-logic axioms) +-- (contains-ax-s : ax-s i ⊆ weak-modal-logic axioms) +-- where + +-- forward-deduction-lemma : +-- (a : formula i) {b : formula i} → +-- theory-add-formula a axioms ▷ b → +-- is-in-subtype (weak-modal-logic axioms) (a →ₘ b) +-- forward-deduction-lemma a {b} (w-ax x) = +-- elim-disjunction-Prop +-- ( Id-formula-Prop a b) +-- ( axioms b) +-- ( (weak-modal-logic axioms) (a →ₘ b)) +-- ( (λ { refl → deduction-a->a a}) +-- , (λ in-axioms → +-- ( apply-universal-property-trunc-Prop +-- ( contains-ax-k _ (b , a , refl)) +-- ( (weak-modal-logic axioms) (a →ₘ b)) +-- ( λ x → unit-trunc-Prop (w-mp x (w-ax in-axioms)))))) +-- ( x) +-- forward-deduction-lemma a {b} (w-mp {c} twdcb twdc) = +-- apply-three-times-universal-property-trunc-Prop +-- ( forward-deduction-lemma a twdcb) +-- ( forward-deduction-lemma a twdc) +-- ( contains-ax-s _ (a , c , b , refl)) +-- ( (weak-modal-logic axioms) (a →ₘ b)) +-- ( λ wdacb wdac x → +-- ( unit-trunc-Prop (w-mp (w-mp x wdacb) wdac))) + +-- deduction-lemma : +-- (a b : formula i) → +-- weak-modal-logic (theory-add-formula a axioms) b ⇔ +-- weak-modal-logic axioms (a →ₘ b) +-- pr1 (deduction-lemma a b) = +-- map-universal-property-trunc-Prop +-- ( (weak-modal-logic axioms) (a →ₘ b)) +-- ( forward-deduction-lemma a) +-- pr2 (deduction-lemma a b) = +-- map-universal-property-trunc-Prop +-- ( weak-modal-logic (theory-add-formula a axioms) b) +-- ( backward-deduction-lemma) +``` -module _ - {l1 l2 : Level} {i : Set l1} (axioms : formulas l2 i) - (contains-ax-k : ax-k i ⊆ weak-modal-logic axioms) - (contains-ax-s : ax-s i ⊆ weak-modal-logic axioms) - (contains-ax-dn : ax-dn i ⊆ weak-modal-logic axioms) - where +### Old definition - deduction-ex-falso : - (a b : formula i) → - is-in-subtype (weak-modal-logic axioms) (~ a →ₘ a →ₘ b) - deduction-ex-falso a b = - forward-implication - ( deduction-lemma axioms contains-ax-k contains-ax-s (~ a) (a →ₘ b)) - ( forward-implication - ( deduction-lemma - ( theory-add-formula (~ a) axioms) - ( transitive-leq-subtype - ( ax-k i) - ( weak-modal-logic axioms) - ( weak-modal-logic (theory-add-formula (a →ₘ ⊥) axioms)) - ( weak-modal-logic-monotic - ( subtype-union-right (Id-formula-Prop (~ a)) axioms)) - ( contains-ax-k)) - ( transitive-leq-subtype - ( ax-s i) - ( weak-modal-logic axioms) - ( weak-modal-logic (theory-add-formula (a →ₘ ⊥) axioms)) - ( weak-modal-logic-monotic - ( subtype-union-right (Id-formula-Prop (~ a)) axioms)) - ( contains-ax-s)) - ( a) - ( b)) - ( apply-twice-universal-property-trunc-Prop - ( contains-ax-k _ (⊥ , ~ b , refl)) - ( contains-ax-dn _ ((b , refl))) - ( weak-modal-logic - ( theory-add-formula a (theory-add-formula (~ a) axioms)) b) - ( λ x y → - ( unit-trunc-Prop - ( w-mp {a = ~~ b} - ( weak-deduction-monotic - ( transitive-leq-subtype - ( axioms) - ( theory-add-formula (~ a) axioms) - ( theory-add-formula a - ( theory-add-formula (~ a) axioms)) - ( subtype-union-right - ( Id-formula-Prop a) - ( theory-add-formula (~ a) axioms)) - ( subtype-union-right (Id-formula-Prop (~ a)) axioms)) - ( y)) - ( w-mp {a = ⊥} - ( weak-deduction-monotic - ( transitive-leq-subtype - ( axioms) - ( theory-add-formula (~ a) axioms) - ( theory-add-formula a - ( theory-add-formula (~ a) axioms)) - ( subtype-union-right - ( Id-formula-Prop a) - ( theory-add-formula (~ a) axioms)) - ( subtype-union-right (Id-formula-Prop (~ a)) axioms)) - ( x)) - ( w-mp {a = a} - ( w-ax - ( unit-trunc-Prop (inr (unit-trunc-Prop (inl refl))))) - ( w-ax (unit-trunc-Prop (inl refl)))))))))) - - logic-ex-falso : - (a b : formula i) → - is-in-subtype (weak-modal-logic axioms) a → - is-in-subtype (weak-modal-logic axioms) (~ a) → - is-in-subtype (weak-modal-logic axioms) b - logic-ex-falso a b a-in-logic not-a-in-logic = - apply-three-times-universal-property-trunc-Prop - ( a-in-logic) - ( not-a-in-logic) - ( deduction-ex-falso a b) - ( (weak-modal-logic axioms) b) - ( λ x y z → unit-trunc-Prop (w-mp (w-mp z y) x)) +```agda +-- infix 5 _▷_ + +-- data _▷_ (axioms : modal-theory l2 i) : formula i → UU (l1 ⊔ l2) where +-- w-ax : {a : formula i} → is-in-subtype axioms a → axioms ▷ a +-- w-mp : {a b : formula i} → axioms ▷ a →ₘ b → axioms ▷ a → axioms ▷ b + +-- weak-modal-logic : modal-theory l2 i → modal-theory (l1 ⊔ l2) i +-- weak-modal-logic axioms a = trunc-Prop (axioms ▷ a) + +-- is-weak-modal-logic-Prop : modal-theory l2 i → Prop (l1 ⊔ l2) +-- is-weak-modal-logic-Prop theory = +-- implicit-Π-Prop (formula i) (λ a → function-Prop (theory ▷ a) (theory a)) + +-- is-weak-modal-logic : modal-theory l2 i → UU (l1 ⊔ l2) +-- is-weak-modal-logic = type-Prop ∘ is-weak-modal-logic-Prop + +-- module _ +-- {l1 l2 : Level} {i : Set l1} {axioms : modal-theory l2 i} +-- where + +-- weak-modal-logic-ax : +-- {a : formula i} → +-- is-in-subtype axioms a → +-- is-in-subtype (weak-modal-logic axioms) a +-- weak-modal-logic-ax = unit-trunc-Prop ∘ w-ax + +-- weak-modal-logic-mp : +-- {a b : formula i} → +-- is-in-subtype (weak-modal-logic axioms) (a →ₘ b) → +-- is-in-subtype (weak-modal-logic axioms) a → +-- is-in-subtype (weak-modal-logic axioms) b +-- weak-modal-logic-mp {a} {b} twdab twda = +-- apply-twice-universal-property-trunc-Prop twdab twda +-- ( weak-modal-logic axioms b) +-- ( λ wdab wda → unit-trunc-Prop (w-mp wdab wda)) + +-- module _ +-- {l1 l2 : Level} {i : Set l1} (axioms : modal-theory l2 i) +-- where + +-- deduction-weak-deduction : {a : formula i} → axioms ▷ a → axioms ⊢ a +-- deduction-weak-deduction (w-ax x) = ax x +-- deduction-weak-deduction (w-mp dab da) = +-- mp (deduction-weak-deduction dab) (deduction-weak-deduction da) + +-- weak-modal-logic-subset-modal-logic : +-- weak-modal-logic axioms ⊆ modal-logic axioms +-- weak-modal-logic-subset-modal-logic a = +-- map-universal-property-trunc-Prop +-- ( modal-logic axioms a) +-- ( unit-trunc-Prop ∘ deduction-weak-deduction) + +-- module _ +-- {l1 l2 l3 : Level} {i : Set l1} +-- {axioms₁ : modal-theory l2 i} {axioms₂ : modal-theory l3 i} +-- (e : equiv-subtypes axioms₁ axioms₂) +-- where + +-- deduction-equiv-axioms : +-- {a : formula i} → axioms₁ ▷ a → axioms₂ ▷ a +-- deduction-equiv-axioms {a} (w-ax x) = +-- w-ax (map-equiv (e a) x) +-- deduction-equiv-axioms (w-mp wdab wda) = +-- w-mp (deduction-equiv-axioms wdab) (deduction-equiv-axioms wda) + +-- module _ +-- {l1 : Level} {i : Set l1} +-- where + +-- axioms-subset-weak-modal-logic : +-- {l2 : Level} (axioms : modal-theory l2 i) → axioms ⊆ weak-modal-logic axioms +-- axioms-subset-weak-modal-logic _ a H = unit-trunc-Prop (w-ax H) + +-- -- TODO: rename to is-weak-modal-logic-weak-modal-logic-closure +-- weak-modal-logic-closed : +-- {l2 : Level} {axioms : modal-theory l2 i} → +-- is-weak-modal-logic (weak-modal-logic axioms) +-- weak-modal-logic-closed (w-ax x) = x +-- weak-modal-logic-closed {axioms = axioms} {a} (w-mp tdab tda) = +-- apply-twice-universal-property-trunc-Prop +-- ( weak-modal-logic-closed tdab) +-- ( weak-modal-logic-closed tda) +-- ( weak-modal-logic axioms a) +-- (λ dab da → unit-trunc-Prop (w-mp dab da)) + +-- weak-modal-logic-closed' : +-- {l2 : Level} {theory : modal-theory l2 i} → +-- is-weak-modal-logic theory → +-- weak-modal-logic theory ⊆ theory +-- weak-modal-logic-closed' {theory = theory} is-logic a = +-- map-universal-property-trunc-Prop (theory a) is-logic + +-- subset-double-weak-modal-logic : +-- {l2 : Level} +-- (axioms : modal-theory l2 i) → +-- weak-modal-logic (weak-modal-logic axioms) ⊆ weak-modal-logic axioms +-- subset-double-weak-modal-logic axioms = +-- weak-modal-logic-closed' weak-modal-logic-closed + +-- weak-modal-logic-idempotent : +-- {l2 : Level} +-- (axioms : modal-theory l2 i) → +-- weak-modal-logic axioms = weak-modal-logic (weak-modal-logic axioms) +-- weak-modal-logic-idempotent axioms = +-- antisymmetric-leq-subtype _ _ +-- ( axioms-subset-weak-modal-logic (weak-modal-logic axioms)) +-- ( subset-double-weak-modal-logic axioms) + +-- module _ +-- {l1 l2 l3 : Level} {i : Set l1} +-- {ax₁ : modal-theory l2 i} +-- {ax₂ : modal-theory l3 i} +-- (leq : ax₁ ⊆ ax₂) +-- where + +-- weak-deduction-monotic : {a : formula i} → ax₁ ▷ a → ax₂ ▷ a +-- weak-deduction-monotic (w-ax x) = w-ax (leq _ x) +-- weak-deduction-monotic (w-mp wdab wda) = +-- w-mp (weak-deduction-monotic wdab) (weak-deduction-monotic wda) + +-- weak-modal-logic-monotic : weak-modal-logic ax₁ ⊆ weak-modal-logic ax₂ +-- weak-modal-logic-monotic a = +-- map-universal-property-trunc-Prop +-- ( weak-modal-logic ax₂ a) +-- ( unit-trunc-Prop ∘ weak-deduction-monotic) + +-- module _ +-- {l1 l2 l3 : Level} {i : Set l1} +-- {ax₁ : modal-theory l2 i} +-- {ax₂ : modal-theory l3 i} +-- where + +-- subset-weak-modal-logic-if-subset-axioms : +-- ax₁ ⊆ weak-modal-logic ax₂ → weak-modal-logic ax₁ ⊆ weak-modal-logic ax₂ +-- subset-weak-modal-logic-if-subset-axioms leq = +-- transitive-leq-subtype +-- ( weak-modal-logic ax₁) +-- ( weak-modal-logic (weak-modal-logic ax₂)) +-- ( weak-modal-logic ax₂) +-- ( subset-double-weak-modal-logic ax₂) +-- ( weak-modal-logic-monotic leq) + +-- -- TODO: move to lists + +-- in-list : {l : Level} {A : UU l} → list A → A → Prop l +-- in-list l a = trunc-Prop (a ∈-list l) + +-- subset-in-tail : +-- {l : Level} {A : UU l} (a : A) (l : list A) → +-- in-list l ⊆ in-list (cons a l) +-- subset-in-tail a l x = +-- map-universal-property-trunc-Prop +-- ( in-list (cons a l) x) +-- ( unit-trunc-Prop ∘ (is-in-tail x a l)) + +-- in-concat-list : +-- {l : Level} {A : UU l} (l1 l2 : list A) → +-- (a : A) → a ∈-list (concat-list l1 l2) → (a ∈-list l1) + (a ∈-list l2) +-- in-concat-list nil l2 a in-list = inr in-list +-- in-concat-list (cons x l1) l2 a (is-head a _) = inl (is-head a l1) +-- in-concat-list (cons x l1) l2 a (is-in-tail _ _ _ in-tail) +-- with in-concat-list l1 l2 a in-tail +-- ... | inl in-l1 = inl (is-in-tail a x l1 in-l1) +-- ... | inr in-l2 = inr in-l2 + +-- subset-reversing-list : +-- {l : Level} {A : UU l} (l : list A) → in-list l ⊆ in-list (reverse-list l) +-- subset-reversing-list l x = +-- map-universal-property-trunc-Prop +-- ( in-list (reverse-list l) x) +-- ( unit-trunc-Prop ∘ forward-contains-reverse-list x l) + +-- in-sum-concat-list-Prop : +-- {l : Level} {A : UU l} (l1 l2 : list A) → +-- (a : A) → a ∈-list (concat-list l1 l2) → +-- type-Prop (in-list l1 a) + type-Prop (in-list l2 a) +-- in-sum-concat-list-Prop l1 l2 a in-concat with in-concat-list l1 l2 a in-concat +-- ... | inl in-l1 = inl (unit-trunc-Prop in-l1) +-- ... | inr in-l2 = inr (unit-trunc-Prop in-l2) + +-- union-lists-concat : +-- {l : Level} {A : UU l} (l1 l2 : list A) → +-- in-list (concat-list l1 l2) ⊆ union-subtype (in-list l1) (in-list l2) +-- union-lists-concat l1 l2 a = +-- map-universal-property-trunc-Prop +-- ( union-subtype (in-list l1) (in-list l2) a) +-- ( unit-trunc-Prop ∘ (in-sum-concat-list-Prop l1 l2 a)) + +-- in-concat-left : +-- {l : Level} {A : UU l} (l1 l2 : list A) +-- {a : A} → a ∈-list l1 → a ∈-list (concat-list l1 l2) +-- in-concat-left _ _ (is-head a _) = +-- is-head a _ +-- in-concat-left _ l2 (is-in-tail a x l1 in-l1) = +-- is-in-tail a x (concat-list l1 l2) (in-concat-left l1 l2 in-l1) + +-- in-concat-right : +-- {l : Level} {A : UU l} (l1 l2 : list A) +-- {a : A} → a ∈-list l2 → a ∈-list (concat-list l1 l2) +-- in-concat-right nil l2 in-l2 = in-l2 +-- in-concat-right (cons x l1) l2 in-l2 = +-- is-in-tail _ x (concat-list l1 l2) (in-concat-right l1 l2 in-l2) + +-- in-concat-list-sum-Prop : +-- {l : Level} {A : UU l} (l1 l2 : list A) → +-- (a : A) → type-Prop (in-list l1 a) + type-Prop (in-list l2 a) → +-- type-Prop (in-list (concat-list l1 l2) a) +-- in-concat-list-sum-Prop l1 l2 a (inl in-l1) = +-- apply-universal-property-trunc-Prop +-- ( in-l1) +-- ( in-list (concat-list l1 l2) a) +-- ( unit-trunc-Prop ∘ (in-concat-left l1 l2)) +-- in-concat-list-sum-Prop l1 l2 a (inr in-l2) = +-- apply-universal-property-trunc-Prop +-- ( in-l2) +-- ( in-list (concat-list l1 l2) a) +-- ( unit-trunc-Prop ∘ (in-concat-right l1 l2)) + +-- in-concat-lists-union : +-- {l : Level} {A : UU l} (l1 l2 : list A) → +-- union-subtype (in-list l1) (in-list l2) ⊆ in-list (concat-list l1 l2) +-- in-concat-lists-union l1 l2 a = +-- map-universal-property-trunc-Prop +-- ( in-list (concat-list l1 l2) a) +-- ( in-concat-list-sum-Prop l1 l2 a) + +-- subset-in-concat-left : +-- {l : Level} {A : UU l} (l1 l2 : list A) → +-- in-list l1 ⊆ in-list (concat-list l1 l2) +-- subset-in-concat-left l1 l2 = +-- transitive-leq-subtype +-- ( in-list l1) +-- ( union-subtype (in-list l1) (in-list l2)) +-- ( in-list (concat-list l1 l2)) +-- ( in-concat-lists-union l1 l2) +-- ( subtype-union-left (in-list l1) (in-list l2)) + +-- subset-in-concat-right : +-- {l : Level} {A : UU l} (l1 l2 : list A) → +-- in-list l2 ⊆ in-list (concat-list l1 l2) +-- subset-in-concat-right l1 l2 = +-- transitive-leq-subtype +-- ( in-list l2) +-- ( union-subtype (in-list l1) (in-list l2)) +-- ( in-list (concat-list l1 l2)) +-- ( in-concat-lists-union l1 l2) +-- ( subtype-union-right (in-list l1) (in-list l2)) + +-- empty-in-list-nil : +-- {l : Level} {A : UU l} {x : A} → is-in-subtype (in-list nil) x → empty +-- empty-in-list-nil = map-universal-property-trunc-Prop empty-Prop ( λ ()) + +-- equiv-subset-head-tail : +-- {l1 l2 : Level} {A : UU l1} (x : A) (xs : list A) (a : subtype l2 A) → +-- (leq-prop-subtype (in-list (cons x xs)) a) ⇔ +-- product-Prop (a x) (leq-prop-subtype (in-list xs) a) +-- pr1 (equiv-subset-head-tail x xs a) sub = +-- pair +-- ( sub x (unit-trunc-Prop (is-head x xs))) +-- ( transitive-leq-subtype +-- ( in-list xs) +-- ( in-list (cons x xs)) +-- ( a) +-- ( sub) +-- ( subset-in-tail x xs)) +-- pr2 (equiv-subset-head-tail x xs a) (x-in-a , sub) y = +-- map-universal-property-trunc-Prop +-- ( a y) +-- ( λ +-- { (is-head .x .xs) → x-in-a +-- ; (is-in-tail .y .x .xs y-in-list) → sub y (unit-trunc-Prop y-in-list)}) + +-- lists-in-union-lists : +-- {l1 l2 l3 : Level} {A : UU l1} +-- (xs : list A) (a : subtype l2 A) (b : subtype l3 A) → +-- in-list xs ⊆ union-subtype a b → +-- ∃ +-- ( list A) +-- ( λ xsl → +-- ( Σ +-- ( list A) +-- ( λ xsr → +-- ( in-list xs ⊆ union-subtype (in-list xsl) (in-list xsr)) × +-- ( in-list xsl ⊆ a) × +-- ( in-list xsr ⊆ b)))) +-- lists-in-union-lists nil a b sub = +-- unit-trunc-Prop +-- ( triple nil nil +-- ( triple +-- ( λ _ → ex-falso ∘ empty-in-list-nil) +-- ( λ _ → ex-falso ∘ empty-in-list-nil) +-- ( λ _ → ex-falso ∘ empty-in-list-nil))) +-- lists-in-union-lists (cons x xs) a b sub = +-- apply-twice-universal-property-trunc-Prop +-- ( lists-in-union-lists xs a b +-- ( transitive-leq-subtype +-- ( in-list xs) +-- ( in-list (cons x xs)) +-- ( union-subtype a b) +-- ( sub) +-- ( subset-in-tail x xs))) +-- ( sub x (unit-trunc-Prop (is-head x xs))) +-- ( ∃-Prop _ (λ _ → Σ _ _)) +-- ( λ (xsl , xsr , sub-lists , sub-xsl , sub-xsr) → +-- ( ind-coproduct +-- ( λ _ → _) +-- ( λ x-in-a → +-- ( unit-trunc-Prop +-- ( triple (cons x xsl) xsr +-- ( triple +-- ( λ y → +-- ( map-universal-property-trunc-Prop +-- ( union-subtype (in-list (cons x xsl)) (in-list xsr) y) +-- ( λ +-- { (is-head .x .xs) → +-- ( subtype-union-left +-- ( in-list (cons x xsl)) +-- ( in-list xsr) +-- ( y) +-- ( unit-trunc-Prop (is-head y xsl))) +-- ; (is-in-tail .y .x .xs y-in-xs) → +-- ( subset-union-subset-left +-- ( in-list xsl) +-- ( in-list (cons x xsl)) +-- ( in-list xsr) +-- ( subset-in-tail x xsl) +-- ( y) +-- ( sub-lists y (unit-trunc-Prop y-in-xs)))}))) +-- ( backward-implication +-- ( equiv-subset-head-tail x xsl a) +-- ( x-in-a , sub-xsl)) +-- ( sub-xsr))))) +-- ( λ x-in-b → +-- ( unit-trunc-Prop +-- ( triple xsl (cons x xsr) +-- ( triple +-- ( λ y → +-- ( map-universal-property-trunc-Prop +-- ( union-subtype (in-list xsl) (in-list (cons x xsr)) y) +-- ( λ +-- { (is-head .x .xs) → +-- ( subtype-union-right +-- ( in-list xsl) +-- ( in-list (cons x xsr)) +-- ( y) +-- ( unit-trunc-Prop (is-head y xsr))) +-- ; (is-in-tail .y .x .xs y-in-xs) → +-- ( subset-union-subset-right +-- ( in-list xsl) +-- ( in-list xsr) +-- ( in-list (cons x xsr)) +-- ( subset-in-tail x xsr) +-- ( y) +-- ( sub-lists y (unit-trunc-Prop y-in-xs)))}))) +-- ( sub-xsl) +-- ( backward-implication +-- ( equiv-subset-head-tail x xsr b) +-- ( x-in-b , sub-xsr)))))))) + +-- module _ +-- {l1 : Level} {i : Set l1} +-- where + +-- list-assumptions-weak-deduction : +-- {l2 : Level} {logic : modal-theory l2 i} {a : formula i} → +-- logic ▷ a → +-- Σ (list (formula i)) (λ l → (in-list l ⊆ logic) × (in-list l ▷ a)) +-- pr1 (list-assumptions-weak-deduction {a = a} (w-ax x)) = +-- unit-list a +-- pr1 (pr2 (list-assumptions-weak-deduction {logic = logic} (w-ax x))) b = +-- map-universal-property-trunc-Prop ( logic b) ( λ { (is-head _ _) → x}) +-- pr2 (pr2 (list-assumptions-weak-deduction {a = a} (w-ax x))) = +-- w-ax (unit-trunc-Prop (is-head a nil)) +-- pr1 (list-assumptions-weak-deduction (w-mp wdcb wdc)) = +-- concat-list +-- ( pr1 (list-assumptions-weak-deduction wdcb)) +-- ( pr1 (list-assumptions-weak-deduction wdc)) +-- pr1 (pr2 (list-assumptions-weak-deduction {logic = logic} (w-mp wdcb wdc))) +-- b b-in-list = +-- subtype-union-both +-- ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) +-- ( in-list (pr1 (list-assumptions-weak-deduction wdc))) +-- ( logic) +-- ( pr1 (pr2 (list-assumptions-weak-deduction wdcb))) +-- ( pr1 (pr2 (list-assumptions-weak-deduction wdc))) +-- ( b) +-- ( union-lists-concat +-- ( pr1 (list-assumptions-weak-deduction wdcb)) +-- ( pr1 (list-assumptions-weak-deduction wdc)) +-- ( b) +-- ( b-in-list)) +-- pr2 (pr2 (list-assumptions-weak-deduction (w-mp wdcb wdc))) = +-- w-mp +-- ( weak-deduction-monotic +-- ( transitive-leq-subtype +-- ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) +-- ( union-subtype +-- ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) +-- ( in-list (pr1 (list-assumptions-weak-deduction wdc)))) +-- ( in-list +-- ( concat-list +-- ( pr1 (list-assumptions-weak-deduction wdcb)) +-- ( pr1 (list-assumptions-weak-deduction wdc)))) +-- ( in-concat-lists-union +-- ( pr1 (list-assumptions-weak-deduction wdcb)) +-- ( pr1 (list-assumptions-weak-deduction wdc))) +-- ( subtype-union-left +-- ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) +-- ( in-list (pr1 (list-assumptions-weak-deduction wdc))))) +-- ( pr2 (pr2 (list-assumptions-weak-deduction wdcb)))) +-- ( weak-deduction-monotic +-- ( transitive-leq-subtype +-- ( in-list (pr1 (list-assumptions-weak-deduction wdc))) +-- ( union-subtype +-- ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) +-- ( in-list (pr1 (list-assumptions-weak-deduction wdc)))) +-- ( in-list +-- ( concat-list +-- ( pr1 (list-assumptions-weak-deduction wdcb)) +-- ( pr1 (list-assumptions-weak-deduction wdc)))) +-- ( in-concat-lists-union +-- ( pr1 (list-assumptions-weak-deduction wdcb)) +-- ( pr1 (list-assumptions-weak-deduction wdc))) +-- ( subtype-union-right +-- ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) +-- ( in-list (pr1 (list-assumptions-weak-deduction wdc))))) +-- ( pr2 (pr2 (list-assumptions-weak-deduction wdc)))) + +-- nil-no-deduction : {a : formula i} → ¬ (in-list nil ▷ a) +-- nil-no-deduction (w-ax x) = +-- apply-universal-property-trunc-Prop x empty-Prop (λ ()) +-- nil-no-deduction (w-mp wd _) = nil-no-deduction wd + +-- -- list-assumptions-weak-modal-logic : +-- -- {l2 : Level} {logic : modal-theory l2 i} {a : formula i} → +-- -- is-in-subtype (weak-modal-logic logic) a → +-- -- ∃ ( list (formula i)) +-- -- ( λ l → +-- -- ( in-list l ⊆ logic) × +-- -- ( is-in-subtype (weak-modal-logic (in-list l)) a)) +-- -- list-assumptions-weak-modal-logic {l2} {logic} {a} = +-- -- map-universal-property-trunc-Prop +-- -- ( ∃-Prop +-- -- ( list (formula i)) +-- -- ( λ l → +-- -- ( in-list l ⊆ logic) × +-- -- ( is-in-subtype (weak-modal-logic (in-list l)) a)) ) +-- -- ( λ wda → +-- -- ( let (l , sub , comp) = list-assumptions-weak-deduction wda +-- -- in unit-trunc-Prop (l , sub , (unit-trunc-Prop comp)))) + +-- forward-subset-head-add : +-- (a : formula i) (l : list (formula i)) → +-- theory-add-formula a (in-list l) ⊆ in-list (cons a l) +-- forward-subset-head-add a l = +-- subtype-union-both +-- ( Id-formula-Prop a) +-- ( in-list l) +-- ( in-list (cons a l)) +-- ( λ { .a refl → unit-trunc-Prop (is-head a l)}) +-- ( subset-in-tail a l) + +-- backward-subset-head-add : +-- (a : formula i) (l : list (formula i)) → +-- in-list (cons a l) ⊆ theory-add-formula a (in-list l) +-- backward-subset-head-add a l x = +-- map-universal-property-trunc-Prop +-- ( theory-add-formula a (in-list l) x) +-- ( λ +-- { (is-head .a .l) → +-- ( subtype-union-left (Id-formula-Prop a) (in-list l) a refl) +-- ; (is-in-tail .x .a .l t) → +-- ( subtype-union-right +-- ( Id-formula-Prop a) +-- ( in-list l) +-- ( x) +-- ( unit-trunc-Prop t))}) + +-- module _ +-- {l1 l2 : Level} {i : Set l1} (axioms : modal-theory l2 i) +-- where + +-- backward-deduction-lemma : +-- {a b : formula i} → +-- axioms ▷ a →ₘ b → +-- is-in-subtype (weak-modal-logic (theory-add-formula a axioms)) b +-- backward-deduction-lemma {a} wab = +-- unit-trunc-Prop +-- ( w-mp +-- ( weak-deduction-monotic +-- ( subtype-union-right ((Id-formula-Prop a)) axioms) wab) +-- ( w-ax (subtype-union-left (Id-formula-Prop a) axioms a refl))) + +-- module _ +-- -- TODO: maybe just in axioms? +-- (contains-ax-k : ax-k i ⊆ weak-modal-logic axioms) +-- (contains-ax-s : ax-s i ⊆ weak-modal-logic axioms) +-- where + +-- deduction-a->a : +-- (a : formula i) → is-in-subtype (weak-modal-logic axioms) (a →ₘ a) +-- deduction-a->a a = +-- apply-three-times-universal-property-trunc-Prop +-- ( contains-ax-s _ (a , a →ₘ a , a , refl)) +-- ( contains-ax-k _ (a , a →ₘ a , refl)) +-- ( contains-ax-k _ (a , a , refl)) +-- ( (weak-modal-logic axioms) (a →ₘ a)) +-- ( λ x y z → unit-trunc-Prop (w-mp (w-mp x y) z)) + +-- forward-deduction-lemma : +-- (a : formula i) {b : formula i} → +-- theory-add-formula a axioms ▷ b → +-- is-in-subtype (weak-modal-logic axioms) (a →ₘ b) +-- forward-deduction-lemma a {b} (w-ax x) = +-- elim-disjunction-Prop +-- ( Id-formula-Prop a b) +-- ( axioms b) +-- ( (weak-modal-logic axioms) (a →ₘ b)) +-- ( (λ { refl → deduction-a->a a}) +-- , (λ in-axioms → +-- ( apply-universal-property-trunc-Prop +-- ( contains-ax-k _ (b , a , refl)) +-- ( (weak-modal-logic axioms) (a →ₘ b)) +-- ( λ x → unit-trunc-Prop (w-mp x (w-ax in-axioms)))))) +-- ( x) +-- forward-deduction-lemma a {b} (w-mp {c} twdcb twdc) = +-- apply-three-times-universal-property-trunc-Prop +-- ( forward-deduction-lemma a twdcb) +-- ( forward-deduction-lemma a twdc) +-- ( contains-ax-s _ (a , c , b , refl)) +-- ( (weak-modal-logic axioms) (a →ₘ b)) +-- ( λ wdacb wdac x → +-- ( unit-trunc-Prop (w-mp (w-mp x wdacb) wdac))) + +-- deduction-lemma : +-- (a b : formula i) → +-- weak-modal-logic (theory-add-formula a axioms) b ⇔ +-- weak-modal-logic axioms (a →ₘ b) +-- pr1 (deduction-lemma a b) = +-- map-universal-property-trunc-Prop +-- ( (weak-modal-logic axioms) (a →ₘ b)) +-- ( forward-deduction-lemma a) +-- pr2 (deduction-lemma a b) = +-- map-universal-property-trunc-Prop +-- ( weak-modal-logic (theory-add-formula a axioms) b) +-- ( backward-deduction-lemma) + +-- module _ +-- {l1 l2 : Level} {i : Set l1} (axioms : modal-theory l2 i) +-- (contains-ax-k : ax-k i ⊆ weak-modal-logic axioms) +-- (contains-ax-s : ax-s i ⊆ weak-modal-logic axioms) +-- (contains-ax-dn : ax-dn i ⊆ weak-modal-logic axioms) +-- where + +-- deduction-ex-falso : +-- (a b : formula i) → +-- is-in-subtype (weak-modal-logic axioms) (~ a →ₘ a →ₘ b) +-- deduction-ex-falso a b = +-- forward-implication +-- ( deduction-lemma axioms contains-ax-k contains-ax-s (~ a) (a →ₘ b)) +-- ( forward-implication +-- ( deduction-lemma +-- ( theory-add-formula (~ a) axioms) +-- ( transitive-leq-subtype +-- ( ax-k i) +-- ( weak-modal-logic axioms) +-- ( weak-modal-logic (theory-add-formula (a →ₘ ⊥) axioms)) +-- ( weak-modal-logic-monotic +-- ( subtype-union-right (Id-formula-Prop (~ a)) axioms)) +-- ( contains-ax-k)) +-- ( transitive-leq-subtype +-- ( ax-s i) +-- ( weak-modal-logic axioms) +-- ( weak-modal-logic (theory-add-formula (a →ₘ ⊥) axioms)) +-- ( weak-modal-logic-monotic +-- ( subtype-union-right (Id-formula-Prop (~ a)) axioms)) +-- ( contains-ax-s)) +-- ( a) +-- ( b)) +-- ( apply-twice-universal-property-trunc-Prop +-- ( contains-ax-k _ (⊥ , ~ b , refl)) +-- ( contains-ax-dn _ ((b , refl))) +-- ( weak-modal-logic +-- ( theory-add-formula a (theory-add-formula (~ a) axioms)) b) +-- ( λ x y → +-- ( unit-trunc-Prop +-- ( w-mp {a = ~~ b} +-- ( weak-deduction-monotic +-- ( transitive-leq-subtype +-- ( axioms) +-- ( theory-add-formula (~ a) axioms) +-- ( theory-add-formula a +-- ( theory-add-formula (~ a) axioms)) +-- ( subtype-union-right +-- ( Id-formula-Prop a) +-- ( theory-add-formula (~ a) axioms)) +-- ( subtype-union-right (Id-formula-Prop (~ a)) axioms)) +-- ( y)) +-- ( w-mp {a = ⊥} +-- ( weak-deduction-monotic +-- ( transitive-leq-subtype +-- ( axioms) +-- ( theory-add-formula (~ a) axioms) +-- ( theory-add-formula a +-- ( theory-add-formula (~ a) axioms)) +-- ( subtype-union-right +-- ( Id-formula-Prop a) +-- ( theory-add-formula (~ a) axioms)) +-- ( subtype-union-right (Id-formula-Prop (~ a)) axioms)) +-- ( x)) +-- ( w-mp {a = a} +-- ( w-ax +-- ( unit-trunc-Prop (inr (unit-trunc-Prop (inl refl))))) +-- ( w-ax (unit-trunc-Prop (inl refl)))))))))) + +-- logic-ex-falso : +-- (a b : formula i) → +-- is-in-subtype (weak-modal-logic axioms) a → +-- is-in-subtype (weak-modal-logic axioms) (~ a) → +-- is-in-subtype (weak-modal-logic axioms) b +-- logic-ex-falso a b a-in-logic not-a-in-logic = +-- apply-three-times-universal-property-trunc-Prop +-- ( a-in-logic) +-- ( not-a-in-logic) +-- ( deduction-ex-falso a b) +-- ( (weak-modal-logic axioms) b) +-- ( λ x y z → unit-trunc-Prop (w-mp (w-mp z y) x)) ``` From 317fd295c6bc86dd21b9c2ab028918f6ba7d5d6a Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Fri, 3 May 2024 17:41:35 +0300 Subject: [PATCH 40/63] Refactor Lindanbaum's lemma --- src/foundation/subtypes.lagda.md | 31 +- src/foundation/unions-subtypes.lagda.md | 20 + src/lists/concatenation-lists.lagda.md | 19 + src/lists/lists-subtypes.lagda.md | 223 ++++++ src/modal-logic/L-complete-theories.lagda.md | 672 +++++++++++++----- .../l-consistent-theories.lagda.md | 181 +++++ src/modal-logic/lindenbaum.lagda.md | 135 ++++ src/modal-logic/logic-syntax.lagda.md | 42 +- src/modal-logic/weak-deduction.lagda.md | 203 +++--- src/order-theory/chains-posets.lagda.md | 12 +- src/order-theory/chains-preorders.lagda.md | 8 + .../maximal-elements-posets.lagda.md | 89 +-- src/order-theory/zorn.lagda.md | 101 +++ 13 files changed, 1359 insertions(+), 377 deletions(-) create mode 100644 src/lists/lists-subtypes.lagda.md create mode 100644 src/modal-logic/l-consistent-theories.lagda.md create mode 100644 src/modal-logic/lindenbaum.lagda.md create mode 100644 src/order-theory/zorn.lagda.md diff --git a/src/foundation/subtypes.lagda.md b/src/foundation/subtypes.lagda.md index b3a5265b3f..b7dc626d1b 100644 --- a/src/foundation/subtypes.lagda.md +++ b/src/foundation/subtypes.lagda.md @@ -139,6 +139,19 @@ module _ ### The containment relation is antisymmetric +```agda +module _ + {l1 : Level} {A : UU l1} + where + + antisymmetric-leq-subtype : + {l2 : Level} (P Q : subtype l2 A) → P ⊆ Q → Q ⊆ P → P = Q + antisymmetric-leq-subtype P Q H K = + eq-has-same-elements-subtype P Q (λ x → (H x , K x)) +``` + +### TODO: equiv subtypes + ```agda module _ {l1 : Level} {A : UU l1} @@ -155,8 +168,8 @@ module _ inv-equiv-subtypes P Q e x = inv-equiv (e x) equiv-antisymmetric-leq-subtype : - {l2 l3 : Level} (P : subtype l2 A) (Q : subtype l3 A) → P ⊆ Q → Q ⊆ P → - equiv-subtypes P Q + {l2 l3 : Level} (P : subtype l2 A) (Q : subtype l3 A) → + P ⊆ Q → Q ⊆ P → equiv-subtypes P Q equiv-antisymmetric-leq-subtype P Q H K x = equiv-iff-is-prop ( is-prop-is-in-subtype P x) @@ -164,10 +177,16 @@ module _ ( H x) ( K x) - antisymmetric-leq-subtype : - {l2 : Level} (P Q : subtype l2 A) → P ⊆ Q → Q ⊆ P → P = Q - antisymmetric-leq-subtype P Q H K = - eq-has-same-elements-subtype P Q (λ x → (H x , K x)) + subset-equiv-subtypes : + {l2 l3 : Level} (P : subtype l2 A) (Q : subtype l3 A) → + equiv-subtypes P Q → P ⊆ Q + subset-equiv-subtypes P Q e x = map-equiv (e x) + + inv-subset-equiv-subtypes : + {l2 l3 : Level} (P : subtype l2 A) (Q : subtype l3 A) → + equiv-subtypes P Q → Q ⊆ P + inv-subset-equiv-subtypes P Q = + subset-equiv-subtypes Q P ∘ inv-equiv-subtypes P Q ``` ### The type of all subtypes of a type is a set diff --git a/src/foundation/unions-subtypes.lagda.md b/src/foundation/unions-subtypes.lagda.md index 795acdc60e..d24bceb5e0 100644 --- a/src/foundation/unions-subtypes.lagda.md +++ b/src/foundation/unions-subtypes.lagda.md @@ -276,4 +276,24 @@ module _ ( union-subtype P P) ( subtype-union-left P P) ( subtype-union-same) + +module _ + {l1 l2 : Level} {X : UU l1} (P : subtype l2 X) (Q : subtype l2 X) + where + + eq-union-subset-left : P ⊆ Q → P ∪ Q = Q + eq-union-subset-left P-sub-Q = + antisymmetric-leq-subtype + ( P ∪ Q) + ( Q) + ( subtype-union-both P Q Q P-sub-Q (refl-leq-subtype Q)) + ( subtype-union-right P Q) + + eq-union-subset-right : Q ⊆ P → P ∪ Q = P + eq-union-subset-right Q-sub-P = + antisymmetric-leq-subtype + ( P ∪ Q) + ( P) + ( subtype-union-both P Q P (refl-leq-subtype P) Q-sub-P) + ( subtype-union-left P Q) ``` diff --git a/src/lists/concatenation-lists.lagda.md b/src/lists/concatenation-lists.lagda.md index 3607cc2fa2..a8e362c224 100644 --- a/src/lists/concatenation-lists.lagda.md +++ b/src/lists/concatenation-lists.lagda.md @@ -172,3 +172,22 @@ tail-concat-list' (cons a nil) y = refl tail-concat-list' (cons a (cons b x)) y = ap (cons b) (tail-concat-list' (cons b x) y) ``` + +### TODO: Concats + +```agda +in-concat-left : + {l : Level} {A : UU l} (l1 l2 : list A) + {a : A} → a ∈-list l1 → a ∈-list (concat-list l1 l2) +in-concat-left _ _ (is-head a _) = + is-head a _ +in-concat-left _ l2 (is-in-tail a x l1 in-l1) = + is-in-tail a x (concat-list l1 l2) (in-concat-left l1 l2 in-l1) + +in-concat-right : + {l : Level} {A : UU l} (l1 l2 : list A) + {a : A} → a ∈-list l2 → a ∈-list (concat-list l1 l2) +in-concat-right nil l2 in-l2 = in-l2 +in-concat-right (cons x l1) l2 in-l2 = + is-in-tail _ x (concat-list l1 l2) (in-concat-right l1 l2 in-l2) +``` diff --git a/src/lists/lists-subtypes.lagda.md b/src/lists/lists-subtypes.lagda.md new file mode 100644 index 0000000000..b3f6d88463 --- /dev/null +++ b/src/lists/lists-subtypes.lagda.md @@ -0,0 +1,223 @@ +# Lists + +```agda +module lists.lists-subtypes where +``` + +
Imports + +```agda +open import foundation.dependent-pair-types +open import foundation.disjunction +open import foundation-core.empty-types +open import foundation.existential-quantification +open import foundation.logical-equivalences +open import foundation.intersections-subtypes +open import foundation.unions-subtypes +open import foundation.propositional-truncations +open import foundation.universe-levels + +open import foundation-core.cartesian-product-types +open import foundation-core.function-types +open import foundation-core.negation +open import foundation-core.propositions +open import foundation-core.subtypes + +open import lists.lists +open import lists.concatenation-lists +open import lists.reversing-lists +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l : Level} {A : UU l} + where + + list-subtype : list A → subtype l A + list-subtype l a = trunc-Prop (a ∈-list l) + + not-in-list-nil : {a : A} → ¬ (is-in-subtype (list-subtype nil) a) + not-in-list-nil = map-universal-property-trunc-Prop empty-Prop (λ ()) + + subset-list-subtype-nil : + {l2 : Level} (S : subtype l2 A) → list-subtype nil ⊆ S + subset-list-subtype-nil S _ = ex-falso ∘ not-in-list-nil + + in-list-subtype-in-list : + {a : A} {l : list A} → a ∈-list l → is-in-subtype (list-subtype l) a + in-list-subtype-in-list = unit-trunc-Prop + + subset-tail-list-subtype : + {a : A} {l : list A} → list-subtype l ⊆ list-subtype (cons a l) + subset-tail-list-subtype {a} {l} x = + map-universal-property-trunc-Prop + ( list-subtype (cons a l) x) + ( in-list-subtype-in-list ∘ is-in-tail x a l) + + head-in-list-subtype : + {a : A} {l : list A} → is-in-subtype (list-subtype (cons a l)) a + head-in-list-subtype {a} {l} = in-list-subtype-in-list (is-head a l) + + subset-list-subtype-cons : + {a : A} {l : list A} → + {l2 : Level} (S : subtype l2 A) → + is-in-subtype S a → + list-subtype l ⊆ S → + list-subtype (cons a l) ⊆ S + subset-list-subtype-cons {a} {l} S a-in-S l-sub-S x = + map-universal-property-trunc-Prop + ( S x) + ( λ { (is-head .x .l) → a-in-S + ; (is-in-tail .x .a .l t) → l-sub-S x (in-list-subtype-in-list t) }) + + subset-list-subtype-reverse-list : + (l : list A) → list-subtype l ⊆ list-subtype (reverse-list l) + subset-list-subtype-reverse-list l x = + map-universal-property-trunc-Prop + ( list-subtype (reverse-list l) x) + ( in-list-subtype-in-list ∘ forward-contains-reverse-list x l) + + subset-list-subtype-concat-union : + {l1 l2 : list A} → + list-subtype (concat-list l1 l2) ⊆ list-subtype l1 ∪ list-subtype l2 + subset-list-subtype-concat-union {nil} {l2} = + subtype-union-right (list-subtype nil) (list-subtype l2) + subset-list-subtype-concat-union {cons x l1} {l2} = + subset-list-subtype-cons + ( list-subtype (cons x l1) ∪ list-subtype l2) + ( subtype-union-left (list-subtype (cons x l1)) (list-subtype l2) x + ( head-in-list-subtype)) + ( transitive-leq-subtype + ( list-subtype (concat-list l1 l2)) + ( list-subtype l1 ∪ list-subtype l2) + ( list-subtype (cons x l1) ∪ list-subtype l2) + ( subset-union-subset-left + ( list-subtype l1) + ( list-subtype (cons x l1)) + ( list-subtype l2) + ( subset-tail-list-subtype)) + ( subset-list-subtype-concat-union)) + + subset-list-subset-concat-left : + (l1 l2 : list A) → + list-subtype l1 ⊆ list-subtype (concat-list l1 l2) + subset-list-subset-concat-left l1 l2 x = + map-universal-property-trunc-Prop + ( list-subtype (concat-list l1 l2) x) + ( in-list-subtype-in-list ∘ in-concat-left l1 l2) + + subset-list-subset-concat-right : + (l1 l2 : list A) → + list-subtype l2 ⊆ list-subtype (concat-list l1 l2) + subset-list-subset-concat-right l1 l2 x = + map-universal-property-trunc-Prop + ( list-subtype (concat-list l1 l2) x) + ( in-list-subtype-in-list ∘ in-concat-right l1 l2) + + subset-list-subtype-union-concat : + {l1 l2 : list A} → + list-subtype l1 ∪ list-subtype l2 ⊆ list-subtype (concat-list l1 l2) + subset-list-subtype-union-concat {l1} {l2} = + subtype-union-both + ( list-subtype l1) + ( list-subtype l2) + ( list-subtype (concat-list l1 l2)) + ( subset-list-subset-concat-left l1 l2) + ( subset-list-subset-concat-right l1 l2) + + iff-subset-head-tail : + {l2 : Level} (x : A) (l : list A) (a : subtype l2 A) → + (list-subtype (cons x l) ⊆ a) ↔ is-in-subtype a x × (list-subtype l ⊆ a) + pr1 (pr1 (iff-subset-head-tail x l a) leq) = + leq x head-in-list-subtype + pr2 (pr1 (iff-subset-head-tail x l a) leq) = + transitive-leq-subtype (list-subtype l) (list-subtype (cons x l)) a + ( leq) + ( subset-tail-list-subtype) + pr2 (iff-subset-head-tail x xs a) (x-in-a , leq) = + subset-list-subtype-cons a x-in-a leq + + lists-in-union-lists : + {l2 l3 : Level} + (xs : list A) (a : subtype l2 A) (b : subtype l3 A) → + list-subtype xs ⊆ union-subtype a b → + ∃ ( list A × list A) + ( λ (xsl , xsr) → + ( list-subtype xs ⊆ list-subtype xsl ∪ list-subtype xsr) × + ( list-subtype xsl ⊆ a) × + ( list-subtype xsr ⊆ b)) + lists-in-union-lists nil a b sub = + intro-∃ (nil , nil) + ( triple + ( subset-list-subtype-nil (list-subtype nil ∪ list-subtype nil)) + ( subset-list-subtype-nil a) + ( subset-list-subtype-nil b)) + lists-in-union-lists (cons x xs) a b leq = + apply-universal-property-trunc-Prop + ( lists-in-union-lists xs a b + ( transitive-leq-subtype + ( list-subtype xs) + ( list-subtype (cons x xs)) + ( union-subtype a b) + ( leq) + ( subset-tail-list-subtype))) + ( ∃-Prop _ _) + ( λ ((xsl , xsr) , leq-lists , leq-xsl , leq-xsr) → + ( elim-disjunction-Prop (a x) (b x) (∃-Prop _ _) + ( pair + ( λ x-in-a → + ( intro-∃ (cons x xsl , xsr) + ( triple + ( subset-list-subtype-cons + ( list-subtype (cons x xsl) ∪ list-subtype xsr) + ( subtype-union-left + ( list-subtype (cons x xsl)) + ( list-subtype xsr) + ( x) + ( head-in-list-subtype)) + ( transitive-leq-subtype + ( list-subtype xs) + ( list-subtype xsl ∪ list-subtype xsr) + ( list-subtype (cons x xsl) ∪ list-subtype xsr) + ( subset-union-subset-left + ( list-subtype xsl) + ( list-subtype (cons x xsl)) + ( list-subtype xsr) + ( subset-tail-list-subtype)) + ( leq-lists))) + ( backward-implication (iff-subset-head-tail x xsl a) + ( x-in-a , leq-xsl)) + ( leq-xsr)))) + ( λ x-in-b → + ( intro-∃ (xsl , cons x xsr) + ( triple + ( subset-list-subtype-cons + ( list-subtype xsl ∪ list-subtype (cons x xsr)) + ( subtype-union-right + ( list-subtype xsl) + ( list-subtype (cons x xsr)) + ( x) + ( head-in-list-subtype)) + ( transitive-leq-subtype + ( list-subtype xs) + ( list-subtype xsl ∪ list-subtype xsr) + ( list-subtype xsl ∪ list-subtype (cons x xsr)) + ( subset-union-subset-right + ( list-subtype xsl) + ( list-subtype xsr) + ( list-subtype (cons x xsr)) + ( subset-tail-list-subtype)) + ( leq-lists))) + ( leq-xsl) + ( backward-implication (iff-subset-head-tail x xsr b) + ( x-in-b , leq-xsr)))))) + ( leq x head-in-list-subtype))) +``` diff --git a/src/modal-logic/L-complete-theories.lagda.md b/src/modal-logic/L-complete-theories.lagda.md index c296f2da2a..cad804e3bb 100644 --- a/src/modal-logic/L-complete-theories.lagda.md +++ b/src/modal-logic/L-complete-theories.lagda.md @@ -1,60 +1,52 @@ # L-complete theories ```agda -module modal-logic.L-complete-theories where +module modal-logic.l-complete-theories where ```
Imports ```agda open import foundation.action-on-identifications-functions --- open import foundation.cartesian-product-types open import foundation.binary-relations -open import foundation.coproduct-types open import foundation.dependent-pair-types open import foundation.disjunction --- open import foundation.empty-types --- open import foundation.existential-quantification -open import foundation.function-types -open import foundation.identity-types --- open import foundation.inhabited-types --- open import foundation.intersections-subtypes +open import foundation.empty-types +open import foundation.existential-quantification +open import foundation.inhabited-types open import foundation.law-of-excluded-middle open import foundation.logical-equivalences -open import foundation.negation --- open import foundation.propositional-resizing -open import foundation.propositional-truncations -open import foundation.propositions --- open import foundation.raising-universe-levels -open import foundation.sets open import foundation.subtypes +open import foundation.propositional-truncations +open import foundation.propositional-resizing open import foundation.transport-along-identifications open import foundation.unions-subtypes --- open import foundation.unit-type open import foundation.universe-levels --- open import foundation-core.equivalences +open import foundation-core.coproduct-types +open import foundation-core.equality-dependent-pair-types +open import foundation-core.function-types +open import foundation-core.identity-types +open import foundation-core.negation +open import foundation-core.propositions +open import foundation-core.sets +open import foundation-core.subtypes --- open import lists.lists --- open import lists.reversing-lists +open import lists.lists +open import lists.lists-subtypes open import modal-logic.axioms --- open import modal-logic.completeness open import modal-logic.formulas --- open import modal-logic.formulas-deduction --- open import modal-logic.kripke-semantics open import modal-logic.logic-syntax --- open import modal-logic.modal-logic-K --- open import modal-logic.soundness +open import modal-logic.l-consistent-theories open import modal-logic.weak-deduction --- open import order-theory.chains-posets open import order-theory.maximal-elements-posets +open import order-theory.chains-posets open import order-theory.posets --- open import order-theory.preorders open import order-theory.subposets open import order-theory.subtypes-leq-posets --- open import order-theory.top-elements-posets +open import order-theory.zorn ```
@@ -69,82 +61,24 @@ TODO module _ {l1 l2 : Level} {i : Set l1} (logic : modal-theory l2 i) - (logic-is-modal-logic : is-modal-logic logic) where - is-L-consistent-theory-Prop : - {l3 : Level} → modal-theory l3 i → Prop (l1 ⊔ l2 ⊔ l3) - is-L-consistent-theory-Prop theory = - is-consistent-modal-logic-Prop (weak-modal-logic-closure (logic ∪ theory)) - - is-L-consistent-theory : - {l3 : Level} → modal-theory l3 i → UU (l1 ⊔ l2 ⊔ l3) - is-L-consistent-theory = type-Prop ∘ is-L-consistent-theory-Prop - - L-consistent-theory : (l3 : Level) → UU (l1 ⊔ l2 ⊔ lsuc l3) - L-consistent-theory l3 = type-subtype (is-L-consistent-theory-Prop {l3}) - - modal-theory-L-consistent-theory : - {l3 : Level} → L-consistent-theory l3 → modal-theory l3 i - modal-theory-L-consistent-theory = - inclusion-subtype is-L-consistent-theory-Prop - - is-L-consistent-theory-modal-theory-L-consistent-theory : - {l3 : Level} (theory : L-consistent-theory l3) → - is-L-consistent-theory (modal-theory-L-consistent-theory theory) - is-L-consistent-theory-modal-theory-L-consistent-theory = - is-in-subtype-inclusion-subtype is-L-consistent-theory-Prop - - is-L-consistent-antimonotic : - {l3 l4 : Level} - (theory₁ : modal-theory l3 i) → - (theory₂ : modal-theory l4 i) → - theory₁ ⊆ theory₂ → - is-L-consistent-theory theory₂ → - is-L-consistent-theory theory₁ - is-L-consistent-antimonotic theory₁ theory₂ leq is-cons = - is-consistent-modal-logic-antimonotic - ( weak-modal-logic-closure (logic ∪ theory₁)) - ( weak-modal-logic-closure (logic ∪ theory₂)) - ( weak-modal-logic-closure-monotic - ( subset-union-subset-right logic theory₁ theory₂ leq)) - ( is-cons) - - L-consistent-theory-leq-Prop : - {l3 : Level} → Relation-Prop (l1 ⊔ l3) (L-consistent-theory l3) - -- {l3 : Level} → L-consistent-theory l3 → L-consistent-theory → Prop (l1 ⊔ l2) - L-consistent-theory-leq-Prop x y = - leq-prop-subtype - ( modal-theory-L-consistent-theory x) - ( modal-theory-L-consistent-theory y) - - L-consistent-theory-leq : - {l3 : Level} → Relation (l1 ⊔ l3) (L-consistent-theory l3) - L-consistent-theory-leq = type-Relation-Prop L-consistent-theory-leq-Prop - - theories-Poset : (l3 : Level) → Poset (l1 ⊔ lsuc l3) (l1 ⊔ l3) - theories-Poset l3 = subtypes-leq-Poset l3 (formula i) - - L-consistent-theories-Poset : - (l3 : Level) → Poset (l1 ⊔ l2 ⊔ lsuc l3) (l1 ⊔ l3) - L-consistent-theories-Poset l3 = - poset-Subposet (theories-Poset l3) (is-L-consistent-theory-Prop) - is-L-complete-theory-Prop : - {l3 : Level} → L-consistent-theory l3 → Prop (l1 ⊔ l2 ⊔ lsuc l3) + {l3 : Level} → L-consistent-theory logic l3 → Prop (l1 ⊔ l2 ⊔ lsuc l3) is-L-complete-theory-Prop {l3} = - is-maximal-element-Poset-Prop (L-consistent-theories-Poset l3) + is-maximal-element-Poset-Prop (L-consistent-theories-Poset logic l3) is-L-complete-theory : - {l3 : Level} → L-consistent-theory l3 → UU (l1 ⊔ l2 ⊔ lsuc l3) + {l3 : Level} → L-consistent-theory logic l3 → UU (l1 ⊔ l2 ⊔ lsuc l3) is-L-complete-theory = type-Prop ∘ is-L-complete-theory-Prop L-complete-theory : (l3 : Level) → UU (l1 ⊔ l2 ⊔ lsuc l3) L-complete-theory l3 = type-subtype (is-L-complete-theory-Prop {l3}) L-consistent-theory-L-complete-theory : - {l3 : Level} → L-complete-theory l3 → L-consistent-theory l3 - L-consistent-theory-L-complete-theory = inclusion-subtype is-L-complete-theory-Prop + {l3 : Level} → L-complete-theory l3 → L-consistent-theory logic l3 + L-consistent-theory-L-complete-theory = + inclusion-subtype is-L-complete-theory-Prop is-L-complete-theory-L-consistent-theory : {l3 : Level} (theory : L-complete-theory l3) → @@ -155,26 +89,36 @@ module _ modal-theory-L-complete-theory : {l3 : Level} → L-complete-theory l3 → modal-theory l3 i modal-theory-L-complete-theory = - modal-theory-L-consistent-theory ∘ L-consistent-theory-L-complete-theory - - eq-is-L-consistent-union-L-complete : - {l3 l4 : Level} → - (((theory , _) , _) : L-complete-theory (l1 ⊔ l2 ⊔ l3 ⊔ l4)) → - (theory' : modal-theory l4 i) → - is-L-consistent-theory (theory' ∪ theory) → - theory' ∪ theory = theory - eq-is-L-consistent-union-L-complete - ((theory , is-cons) , is-comp) theory' is-L-cons = + modal-theory-L-consistent-theory logic ∘ + L-consistent-theory-L-complete-theory + + is-L-consistent-theory-modal-theory-L-complete-theory : + {l3 : Level} (theory : L-complete-theory l3) → + is-L-consistent-theory logic (modal-theory-L-complete-theory theory) + is-L-consistent-theory-modal-theory-L-complete-theory = + is-L-consistent-theory-modal-theory-L-consistent-theory logic ∘ + L-consistent-theory-L-complete-theory + + module _ + {l3 : Level} + (((theory , is-cons) , is-comp) : L-complete-theory (l1 ⊔ l2 ⊔ l3)) + (theory' : modal-theory l3 i) + where + + eq-is-L-consistent-union-L-complete : + is-L-consistent-theory logic (theory' ∪ theory) → + theory' ∪ theory = theory + eq-is-L-consistent-union-L-complete is-L-cons = ap - ( modal-theory-L-consistent-theory) + ( modal-theory-L-consistent-theory logic) ( is-comp ( theory' ∪ theory , is-L-cons) ( subtype-union-right theory' theory)) union-L-consistent : {l3 : Level} → - L-consistent-theory l3 → - L-consistent-theory (l1 ⊔ l2 ⊔ l3) + L-consistent-theory logic l3 → + L-consistent-theory logic (l1 ⊔ l2 ⊔ l3) pr1 (union-L-consistent (theory , is-cons)) = weak-modal-logic-closure (logic ∪ theory) pr2 (union-L-consistent (theory , is-cons)) bot-in-logic = @@ -201,14 +145,15 @@ module _ ( bot-in-logic)) module _ - {l3 : Level} + (l3 : Level) (t@((theory , is-cons) , is-comp) : L-complete-theory (l1 ⊔ l2 ⊔ l3)) where eq-union-L-consistent : weak-modal-logic-closure (logic ∪ theory) = theory eq-union-L-consistent = - ap modal-theory-L-consistent-theory + ap + ( modal-theory-L-consistent-theory logic) ( is-comp ( union-L-consistent (theory , is-cons)) ( transitive-leq-subtype @@ -231,14 +176,13 @@ module _ subset-union-logic-L-complete-theory = transitive-leq-subtype ( logic ∪ theory) - ( modal-theory-L-consistent-theory + ( modal-theory-L-consistent-theory logic ( union-L-consistent (theory , is-cons))) ( theory) ( subset-union-L-consistent) ( subset-axioms-weak-modal-logic) - subset-logic-L-complete-theory : - logic ⊆ theory + subset-logic-L-complete-theory : logic ⊆ theory subset-logic-L-complete-theory = transitive-leq-subtype ( logic) @@ -247,93 +191,465 @@ module _ ( subset-union-logic-L-complete-theory) ( subtype-union-left logic theory) - is-weak-modal-logic-L-complete-theory : - is-weak-modal-logic theory + is-weak-modal-logic-L-complete-theory : is-weak-modal-logic theory is-weak-modal-logic-L-complete-theory = transitive-leq-subtype ( weak-modal-logic-closure theory) - ( modal-theory-L-consistent-theory + ( modal-theory-L-consistent-theory logic ( union-L-consistent (theory , is-cons))) ( theory) ( subset-union-L-consistent) ( weak-modal-logic-closure-monotic (subtype-union-right logic theory)) + module _ + {l3 : Level} + (t@((theory , is-cons) , is-comp) : L-complete-theory (l1 ⊔ l2 ⊔ l3)) + (theory' : modal-theory l3 i) + where + + eq-is-consistent-union-L-complete : + is-consistent-modal-logic (weak-modal-logic-closure (theory' ∪ theory)) → + theory' ∪ theory = theory + eq-is-consistent-union-L-complete is-cons' = + eq-is-L-consistent-union-L-complete t theory' + ( is-consistent-modal-logic-antimonotic + ( weak-modal-logic-closure (logic ∪ (theory' ∪ theory))) + ( weak-modal-logic-closure (theory' ∪ theory)) + ( weak-modal-logic-closure-monotic + ( subtype-union-both logic (theory' ∪ theory) (theory' ∪ theory) + ( transitive-leq-subtype + ( logic) + ( theory) + ( theory' ∪ theory) + ( subtype-union-right theory' theory) + ( subset-logic-L-complete-theory l3 t)) + ( refl-leq-subtype (theory' ∪ theory)))) + ( is-cons')) + + module _ + (t@((theory , is-cons) , is-comp) : L-complete-theory (l1 ⊔ l2)) + (contains-ax-k : ax-k i ⊆ logic) + (contains-ax-s : ax-s i ⊆ logic) + (contains-ax-dn : ax-dn i ⊆ logic) + where + + private + contains-ax-k' : ax-k i ⊆ theory + contains-ax-k' = + transitive-leq-subtype (ax-k i) logic theory + ( subset-logic-L-complete-theory lzero t) + ( contains-ax-k) + + contains-ax-s' : ax-s i ⊆ theory + contains-ax-s' = + transitive-leq-subtype (ax-s i) logic theory + ( subset-logic-L-complete-theory lzero t) + ( contains-ax-s) + + contains-ax-dn' : ax-dn i ⊆ theory + contains-ax-dn' = + transitive-leq-subtype (ax-dn i) logic theory + ( subset-logic-L-complete-theory lzero t) + ( contains-ax-dn) + + is-L-consistent-add-formula-not-in-logic : + {a : formula i} → + ¬ (is-in-subtype theory a) → + is-L-consistent-theory logic (theory-add-formula (~ a) theory) + is-L-consistent-add-formula-not-in-logic {a} not-in-logic bot-in-logic = + not-in-logic + ( weak-modal-logic-mp + ( is-weak-modal-logic-L-complete-theory lzero t) + ( contains-ax-dn' (~~ a →ₘ a) (a , refl)) + ( is-weak-modal-logic-L-complete-theory lzero t (~~ a) + ( forward-implication + ( deduction-lemma + ( theory) + ( contains-ax-k') + ( contains-ax-s') + ( ~ a) + ( ⊥)) + ( weak-modal-logic-closure-monotic + ( subtype-union-both logic (theory-add-formula (~ a) theory) + ( theory-add-formula (~ a) theory) + ( transitive-leq-subtype + ( logic) + ( theory) + ( theory-add-formula (~ a) theory) + ( subset-add-formula (~ a) theory) + ( subset-logic-L-complete-theory lzero t)) + ( refl-leq-subtype (theory-add-formula (~ a) theory))) + ( ⊥) + ( bot-in-logic))))) + + contains-negation-not-contains-formula-L-complete-theory : + {a : formula i} → + ¬ (is-in-subtype theory a) → + is-in-subtype theory (~ a) + contains-negation-not-contains-formula-L-complete-theory {a} not-in-logic = + tr + ( λ t → is-in-subtype t (~ a)) + ( eq-is-L-consistent-union-L-complete t + ( Id-formula-Prop (~ a)) + ( is-L-consistent-add-formula-not-in-logic not-in-logic)) + ( formula-in-add-formula (~ a) theory) + + is-disjuctive-L-complete-theory : + LEM (l1 ⊔ l2) → + is-disjuctive-modal-theory theory + is-disjuctive-L-complete-theory lem a with lem (theory a) + ... | inl a-in-logic = inl a-in-logic + ... | inr a-not-in-logic = + inr + ( contains-negation-not-contains-formula-L-complete-theory + ( a-not-in-logic)) + + is-inhabited-L-complete-exists-complete-L-consistent-theory : + {l3 : Level} → + ∃ (L-consistent-theory logic l3) is-L-complete-theory → + is-inhabited (L-complete-theory l3) + is-inhabited-L-complete-exists-complete-L-consistent-theory {l3} = + elim-exists-Prop + ( is-L-complete-theory-Prop) + ( is-inhabited-Prop (L-complete-theory l3)) + ( λ theory is-comp → unit-trunc-Prop (theory , is-comp)) + + module _ + {l3 l4 : Level} + (C : chain-Poset l4 (L-consistent-theories-Poset logic l3)) + where + + private + P : Poset (l1 ⊔ l2 ⊔ lsuc l3) (l1 ⊔ l3) + P = L-consistent-theories-Poset logic l3 + + modal-theory-chain-element : + type-chain-Poset P C → modal-theory l3 i + modal-theory-chain-element = + modal-theory-L-consistent-theory logic ∘ + type-Poset-type-chain-Poset P C + + L-union : type-chain-Poset P C → modal-theory (l1 ⊔ l2 ⊔ l3) i + L-union x = + weak-modal-logic-closure (logic ∪ modal-theory-chain-element x) + + theory-subset-L-union : + (x : type-chain-Poset P C) → modal-theory-chain-element x ⊆ L-union x + theory-subset-L-union x = + transitive-leq-subtype + ( modal-theory-chain-element x) + ( logic ∪ modal-theory-chain-element x) + ( L-union x) + ( subset-axioms-weak-modal-logic) + ( subtype-union-right logic (modal-theory-chain-element x)) + + leq-L-union : + (x y : type-chain-Poset P C) → + modal-theory-chain-element x ⊆ modal-theory-chain-element y → + L-union x ⊆ L-union y + leq-L-union x y leq = + weak-modal-logic-closure-monotic + ( subset-union-subset-right logic + ( modal-theory-chain-element x) + ( modal-theory-chain-element y) + ( leq)) + + chain-union-modal-theory : + modal-theory (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4) i + chain-union-modal-theory a = + ∃-Prop (type-chain-Poset P C) + ( λ x → is-in-subtype (modal-theory-chain-element x) a) + + exists-chain-element-with-formula-Prop : + (a : formula i) → Prop (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4) + exists-chain-element-with-formula-Prop a = + ∃-Prop (type-chain-Poset P C) + ( λ x → + ( is-in-subtype + ( weak-modal-logic-closure (logic ∪ modal-theory-chain-element x)) + ( a))) + + exists-chain-element-with-formula : + (a : formula i) → UU (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4) + exists-chain-element-with-formula = + type-Prop ∘ exists-chain-element-with-formula-Prop + module _ (contains-ax-k : ax-k i ⊆ logic) (contains-ax-s : ax-s i ⊆ logic) - (contains-ax-dn : ax-dn i ⊆ logic) where private - contains-ax-k' : ax-k i ⊆ theory - contains-ax-k' = - transitive-leq-subtype (ax-k i) logic theory - ( subset-logic-L-complete-theory) + contains-ax-k' : + {l5 : Level} (theory : modal-theory l5 i) → ax-k i ⊆ logic ∪ theory + contains-ax-k' theory = + transitive-leq-subtype (ax-k i) logic (logic ∪ theory) + ( subtype-union-left logic theory) ( contains-ax-k) - contains-ax-s' : ax-s i ⊆ theory - contains-ax-s' = - transitive-leq-subtype (ax-s i) logic theory - ( subset-logic-L-complete-theory) + contains-ax-s' : + {l5 : Level} (theory : modal-theory l5 i) → ax-s i ⊆ logic ∪ theory + contains-ax-s' theory = + transitive-leq-subtype (ax-s i) logic (logic ∪ theory) + ( subtype-union-left logic theory) ( contains-ax-s) - contains-ax-dn' : ax-dn i ⊆ theory - contains-ax-dn' = - transitive-leq-subtype (ax-dn i) logic theory - ( subset-logic-L-complete-theory) - ( contains-ax-dn) - - is-L-consistent-add-formula-not-in-logic : + L-union-deduction-lemma : + (l : list (formula i)) → + (h a : formula i) → + is-in-subtype + ( weak-modal-logic-closure (logic ∪ list-subtype (cons h l))) a → + is-in-subtype + ( weak-modal-logic-closure (logic ∪ list-subtype l)) (h →ₘ a) + L-union-deduction-lemma l h a in-logic = + forward-implication + ( deduction-lemma + ( logic ∪ list-subtype l) + ( contains-ax-k' (list-subtype l)) + ( contains-ax-s' (list-subtype l)) + ( h) + ( a)) + ( weak-modal-logic-closure-monotic + ( transitive-leq-subtype + ( logic ∪ list-subtype (cons h l)) + ( logic ∪ theory-add-formula h (list-subtype l)) + ( theory-add-formula h (logic ∪ list-subtype l)) + ( theory-add-formula-union-right h logic (list-subtype l)) + ( subset-union-subset-right + ( logic) + ( list-subtype (cons h l)) + ( theory-add-formula h (list-subtype l)) + ( backward-subset-head-add h l))) + ( a) + ( in-logic)) + + in-chain-in-chain-union-assumptions : + is-inhabited (type-chain-Poset P C) → + (l : list (formula i)) → + list-subtype l ⊆ chain-union-modal-theory → {a : formula i} → - ¬ (is-in-subtype theory a) → - is-L-consistent-theory (theory-add-formula (~ a) theory) - is-L-consistent-add-formula-not-in-logic {a} not-in-logic bot-in-logic = - not-in-logic - ( weak-modal-logic-mp theory - ( is-weak-modal-logic-L-complete-theory) - ( contains-ax-dn' (~~ a →ₘ a) (a , refl)) - ( is-weak-modal-logic-L-complete-theory (~~ a) - ( forward-implication - ( deduction-lemma - ( theory) - ( contains-ax-k') - ( contains-ax-s') - ( ~ a) - ( ⊥)) - ( weak-modal-logic-closure-monotic - ( subtype-union-both logic (theory-add-formula (~ a) theory) - ( theory-add-formula (~ a) theory) - ( transitive-leq-subtype - ( logic) - ( theory) - ( theory-add-formula (~ a) theory) - ( subset-add-formula (~ a) theory) - ( subset-logic-L-complete-theory)) - ( refl-leq-subtype (theory-add-formula (~ a) theory))) - ( ⊥) - ( bot-in-logic))))) - - contains-negation-not-contains-formula-L-complete-theory : + is-in-subtype (weak-modal-logic-closure (logic ∪ list-subtype l)) a → + exists-chain-element-with-formula a + in-chain-in-chain-union-assumptions is-inh nil leq {a} in-logic = + apply-universal-property-trunc-Prop + ( is-inh) + ( exists-chain-element-with-formula-Prop a) + ( λ x → + ( intro-∃ x + ( weak-modal-logic-closure-monotic + ( subtype-union-both logic (list-subtype nil) + ( logic ∪ modal-theory-chain-element x) + ( subtype-union-left logic (modal-theory-chain-element x)) + ( subset-list-subtype-nil + ( logic ∪ modal-theory-chain-element x))) + ( a) + ( in-logic)))) + in-chain-in-chain-union-assumptions is-inh (cons h l) leq {a} in-logic = + apply-twice-universal-property-trunc-Prop + ( leq h (head-in-list-subtype)) + ( in-chain-in-chain-union-assumptions is-inh l + ( transitive-leq-subtype + ( list-subtype l) + ( list-subtype (cons h l)) + ( chain-union-modal-theory) + ( leq) + ( subset-tail-list-subtype)) + ( L-union-deduction-lemma l h a in-logic)) + ( exists-chain-element-with-formula-Prop a) + ( λ (x , h-in-x) (y , ha-in-y) → + ( elim-disjunction-Prop + ( leq-Poset-Prop P + ( type-Poset-type-chain-Poset P C x) + ( type-Poset-type-chain-Poset P C y)) + ( leq-Poset-Prop P + ( type-Poset-type-chain-Poset P C y) + ( type-Poset-type-chain-Poset P C x)) + ( exists-chain-element-with-formula-Prop a) + ( pair + ( λ x-leq-y → + ( intro-∃ y + ( weak-modal-logic-closure-mp + ( ha-in-y) + ( leq-L-union x y x-leq-y h + ( theory-subset-L-union x h h-in-x))))) + ( λ y-leq-x → + ( intro-∃ x + ( weak-modal-logic-closure-mp + ( leq-L-union y x y-leq-x (h →ₘ a) ha-in-y) + ( theory-subset-L-union x h h-in-x))))) + ( is-chain-Subposet-chain-Poset P C x y))) + + in-chain-in-chain-union : + is-inhabited (type-chain-Poset P C) → {a : formula i} → - ¬ (is-in-subtype theory a) → - is-in-subtype theory (~ a) - contains-negation-not-contains-formula-L-complete-theory {a} not-in-logic = - tr - ( λ t → is-in-subtype t (~ a)) - ( eq-is-L-consistent-union-L-complete - {l3 = l3} {l4 = l1} - ( t) - ( Id-formula-Prop (~ a)) - ( is-L-consistent-add-formula-not-in-logic not-in-logic)) - ( formula-in-add-formula (~ a) theory) - - is-disjuctive-L-complete-theory : - LEM (l1 ⊔ l2 ⊔ l3) → - is-disjuctive-modal-theory theory - is-disjuctive-L-complete-theory lem a with lem (theory a) - ... | inl a-in-logic = inl a-in-logic - ... | inr a-not-in-logic = - inr - ( contains-negation-not-contains-formula-L-complete-theory - ( a-not-in-logic)) + is-in-subtype + ( weak-modal-logic-closure (logic ∪ chain-union-modal-theory)) + ( a) → + exists-chain-element-with-formula a + in-chain-in-chain-union is-inh {a} = + map-universal-property-trunc-Prop + ( exists-chain-element-with-formula-Prop a) + ( λ d → + ( apply-universal-property-trunc-Prop + ( lists-in-union-lists + ( list-assumptions-weak-deduction d) + ( logic) + ( chain-union-modal-theory) + ( subset-theory-list-assumptions-weak-deduction d)) + ( exists-chain-element-with-formula-Prop a) + ( λ ((logic-l , theory-l) , leq-lists , leq-logic , leq-theory) → + ( in-chain-in-chain-union-assumptions is-inh theory-l leq-theory + ( weak-modal-logic-closure-monotic + {ax₁ = list-subtype logic-l ∪ list-subtype theory-l} + ( subset-union-subset-left + ( list-subtype logic-l) + ( logic) + ( list-subtype theory-l) + ( leq-logic)) + ( a) + ( weak-modal-logic-closure-monotic leq-lists a + ( is-in-weak-deduction-closure-weak-deduction + ( is-assumptions-list-assumptions-weak-deduction + ( d))))))))) + + is-L-consistent-theory-chain-union-modal-theory : + is-inhabited (type-chain-Poset P C) → + is-L-consistent-theory logic chain-union-modal-theory + is-L-consistent-theory-chain-union-modal-theory is-inh in-logic = + apply-universal-property-trunc-Prop + ( in-chain-in-chain-union is-inh in-logic) + ( empty-Prop) + ( λ (x , in-logic') → + ( is-L-consistent-theory-modal-theory-L-consistent-theory + ( logic) + ( type-Poset-type-chain-Poset P C x) + ( in-logic'))) + + module _ + {l3 l4 : Level} + (prop-resize : propositional-resizing l3 (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4)) + where + + private + P : Poset (l1 ⊔ l2 ⊔ lsuc l3) (l1 ⊔ l3) + P = L-consistent-theories-Poset logic l3 + + resized-chain-union-modal-theory : + chain-Poset l4 P → + modal-theory l3 i + resized-chain-union-modal-theory C = + resize prop-resize ∘ chain-union-modal-theory C + + equiv-resized-chain-union-modal-theory : + (C : chain-Poset l4 P) → + equiv-subtypes + (resized-chain-union-modal-theory C) + (chain-union-modal-theory C) + equiv-resized-chain-union-modal-theory C a = + is-equiv-resize prop-resize (chain-union-modal-theory C a) + + module _ + (contains-ax-k : ax-k i ⊆ logic) + (contains-ax-s : ax-s i ⊆ logic) + where + + is-L-consistent-resized-chain-union-modal-theory : + (C : chain-Poset l4 P) → + is-inhabited (type-chain-Poset P C) → + is-L-consistent-theory logic (resized-chain-union-modal-theory C) + is-L-consistent-resized-chain-union-modal-theory C is-inh = + is-L-consistent-antimonotic logic + ( resized-chain-union-modal-theory C) + ( chain-union-modal-theory C) + ( subset-equiv-subtypes + ( resized-chain-union-modal-theory C) + ( chain-union-modal-theory C) + ( equiv-resized-chain-union-modal-theory C)) + ( is-L-consistent-theory-chain-union-modal-theory C + ( contains-ax-k) + ( contains-ax-s) + ( is-inh)) + + resized-chain-union-L-consistent-theory : + (C : chain-Poset l4 P) → + is-inhabited (type-chain-Poset P C) → + L-consistent-theory logic l3 + pr1 (resized-chain-union-L-consistent-theory C is-inh) = + resized-chain-union-modal-theory C + pr2 (resized-chain-union-L-consistent-theory C is-inh) = + is-L-consistent-resized-chain-union-modal-theory C is-inh + + union-is-chain-upper-bound : + (C : chain-Poset l4 P) → + (is-inh : is-inhabited (type-chain-Poset P C)) → + is-chain-upper-bound P C + ( resized-chain-union-L-consistent-theory C is-inh) + union-is-chain-upper-bound C is-inh x = + transitive-leq-subtype + ( modal-theory-L-consistent-theory logic + ( type-Poset-type-chain-Poset + ( L-consistent-theories-Poset logic l3) + ( C) + ( x))) + ( chain-union-modal-theory C) + ( resized-chain-union-modal-theory C) + ( inv-subset-equiv-subtypes + ( resized-chain-union-modal-theory C) + ( chain-union-modal-theory C) + ( equiv-resized-chain-union-modal-theory C)) + ( λ a in-theory → intro-∃ x in-theory) + + extend-L-consistent-theory : + (zorn : Zorn-non-empty (l1 ⊔ l2 ⊔ lsuc l3) (l1 ⊔ l3) l4) → + is-inhabited (L-consistent-theory logic l3) → + is-inhabited (L-complete-theory l3) + extend-L-consistent-theory zorn is-inh = + ( is-inhabited-L-complete-exists-complete-L-consistent-theory + ( zorn + ( L-consistent-theories-Poset logic l3) + ( is-inh) + ( λ C C-is-inh → + ( intro-∃ + ( resized-chain-union-L-consistent-theory C C-is-inh) + ( union-is-chain-upper-bound C C-is-inh))))) + +module _ + {l1 : Level} {i : Set l1} + {l2 l3 : Level} + (logic₁ : modal-theory l2 i) + (logic₂ : modal-theory l3 i) + (theory : modal-theory (l1 ⊔ l2 ⊔ l3) i) + (is-cons₁ : is-L-consistent-theory logic₁ theory) + (is-cons₂ : is-L-consistent-theory logic₂ theory) + where + + universal-L-complete-theory : + is-L-complete-theory logic₁ (theory , is-cons₁) → + is-L-complete-theory logic₂ (theory , is-cons₂) + universal-L-complete-theory is-comp (theory' , is-cons') leq = + eq-pair-Σ + ( equational-proof) + ( eq-is-prop + ( is-prop-type-Prop (is-L-consistent-theory-Prop logic₂ theory))) + where + complete-theory : L-complete-theory logic₁ (l1 ⊔ l2 ⊔ l3) + complete-theory = (theory , is-cons₁) , is-comp + + equational-proof : theory' = theory + equational-proof = + equational-reasoning + theory' + = theory' ∪ theory + by inv (eq-union-subset-right theory' theory leq) + = theory + by eq-is-consistent-union-L-complete logic₁ complete-theory theory' + ( is-consistent-modal-logic-antimonotic + ( weak-modal-logic-closure (theory' ∪ theory)) + ( weak-modal-logic-closure theory') + ( weak-modal-logic-closure-monotic + ( subtype-union-both theory' theory theory' + ( refl-leq-subtype theory') + ( leq))) + ( is-consistent-modal-theory-L-consistent-theory logic₂ + ( theory' , is-cons'))) ``` diff --git a/src/modal-logic/l-consistent-theories.lagda.md b/src/modal-logic/l-consistent-theories.lagda.md new file mode 100644 index 0000000000..e648feaafd --- /dev/null +++ b/src/modal-logic/l-consistent-theories.lagda.md @@ -0,0 +1,181 @@ +# L-consistent theories + +```agda +module modal-logic.l-consistent-theories where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.binary-relations +open import foundation.dependent-pair-types +open import foundation.disjunction +open import foundation.inhabited-types +open import foundation.law-of-excluded-middle +open import foundation.logical-equivalences +open import foundation.propositional-truncations +open import foundation.transport-along-identifications +open import foundation.unions-subtypes +open import foundation.universe-levels + +open import foundation-core.coproduct-types +open import foundation-core.function-types +open import foundation-core.identity-types +open import foundation-core.negation +open import foundation-core.propositions +open import foundation-core.sets +open import foundation-core.subtypes + +open import modal-logic.axioms +open import modal-logic.formulas +open import modal-logic.logic-syntax +open import modal-logic.weak-deduction + +open import order-theory.maximal-elements-posets +open import order-theory.posets +open import order-theory.subposets +open import order-theory.subtypes-leq-posets +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l1 l2 : Level} {i : Set l1} + (logic : modal-theory l2 i) + where + + is-L-consistent-theory-Prop : + {l3 : Level} → modal-theory l3 i → Prop (l1 ⊔ l2 ⊔ l3) + is-L-consistent-theory-Prop theory = + is-consistent-modal-logic-Prop (weak-modal-logic-closure (logic ∪ theory)) + + is-L-consistent-theory : + {l3 : Level} → modal-theory l3 i → UU (l1 ⊔ l2 ⊔ l3) + is-L-consistent-theory = type-Prop ∘ is-L-consistent-theory-Prop + + L-consistent-theory : (l3 : Level) → UU (l1 ⊔ l2 ⊔ lsuc l3) + L-consistent-theory l3 = type-subtype (is-L-consistent-theory-Prop {l3}) + + modal-theory-L-consistent-theory : + {l3 : Level} → L-consistent-theory l3 → modal-theory l3 i + modal-theory-L-consistent-theory = + inclusion-subtype is-L-consistent-theory-Prop + + is-L-consistent-theory-modal-theory-L-consistent-theory : + {l3 : Level} (theory : L-consistent-theory l3) → + is-L-consistent-theory (modal-theory-L-consistent-theory theory) + is-L-consistent-theory-modal-theory-L-consistent-theory = + is-in-subtype-inclusion-subtype is-L-consistent-theory-Prop + + is-consistent-modal-theory-L-consistent-theory : + {l3 : Level} (theory : L-consistent-theory l3) → + is-consistent-modal-logic + ( weak-modal-logic-closure (modal-theory-L-consistent-theory theory)) + is-consistent-modal-theory-L-consistent-theory theory = + is-consistent-modal-logic-antimonotic + ( weak-modal-logic-closure (modal-theory-L-consistent-theory theory)) + ( weak-modal-logic-closure + (logic ∪ modal-theory-L-consistent-theory theory)) + ( weak-modal-logic-closure-monotic + ( subtype-union-right logic (modal-theory-L-consistent-theory theory))) + ( is-L-consistent-theory-modal-theory-L-consistent-theory theory) + -- is-consistent-modal-logic-antimonotic + -- ( modal-theory-L-consistent-theory theory) + -- ( weak-modal-logic-closure + -- ( logic ∪ modal-theory-L-consistent-theory theory)) + -- ( transitive-leq-subtype + -- ( modal-theory-L-consistent-theory theory) + -- ( logic ∪ modal-theory-L-consistent-theory theory) + -- ( weak-modal-logic-closure + -- ( logic ∪ modal-theory-L-consistent-theory theory)) + -- ( subset-axioms-weak-modal-logic) + -- ( subtype-union-right logic (modal-theory-L-consistent-theory theory))) + -- ( is-L-consistent-theory-modal-theory-L-consistent-theory theory) + + is-L-consistent-antimonotic : + {l3 l4 : Level} + (theory₁ : modal-theory l3 i) → + (theory₂ : modal-theory l4 i) → + theory₁ ⊆ theory₂ → + is-L-consistent-theory theory₂ → + is-L-consistent-theory theory₁ + is-L-consistent-antimonotic theory₁ theory₂ leq = + is-consistent-modal-logic-antimonotic + ( weak-modal-logic-closure (logic ∪ theory₁)) + ( weak-modal-logic-closure (logic ∪ theory₂)) + ( weak-modal-logic-closure-monotic + ( subset-union-subset-right logic theory₁ theory₂ leq)) + + L-consistent-theory-leq-Prop : + {l3 : Level} → Relation-Prop (l1 ⊔ l3) (L-consistent-theory l3) + -- {l3 : Level} → L-consistent-theory l3 → L-consistent-theory → Prop (l1 ⊔ l2) + L-consistent-theory-leq-Prop x y = + leq-prop-subtype + ( modal-theory-L-consistent-theory x) + ( modal-theory-L-consistent-theory y) + + L-consistent-theory-leq : + {l3 : Level} → Relation (l1 ⊔ l3) (L-consistent-theory l3) + L-consistent-theory-leq = type-Relation-Prop L-consistent-theory-leq-Prop + + theories-Poset : (l3 : Level) → Poset (l1 ⊔ lsuc l3) (l1 ⊔ l3) + theories-Poset l3 = subtypes-leq-Poset l3 (formula i) + + L-consistent-theories-Poset : + (l3 : Level) → Poset (l1 ⊔ l2 ⊔ lsuc l3) (l1 ⊔ l3) + L-consistent-theories-Poset l3 = + poset-Subposet (theories-Poset l3) (is-L-consistent-theory-Prop) + + module _ + (is-logic : is-weak-modal-logic logic) + (is-cons : is-consistent-modal-logic logic) + where + + is-L-consistent-theory-logic : + is-L-consistent-theory logic + is-L-consistent-theory-logic = + is-consistent-modal-logic-antimonotic + ( weak-modal-logic-closure (logic ∪ logic)) + ( logic) + ( transitive-leq-subtype + ( weak-modal-logic-closure (logic ∪ logic)) + ( weak-modal-logic-closure logic) + ( logic) + ( is-logic) + ( weak-modal-logic-closure-monotic + ( subtype-union-both logic logic logic + ( refl-leq-subtype logic) + ( refl-leq-subtype logic)))) + (is-cons) + + is-inhabited-L-consistent-theory : is-inhabited (L-consistent-theory l2) + is-inhabited-L-consistent-theory = + unit-trunc-Prop (logic , is-L-consistent-theory-logic) + +module _ + {l1 : Level} {i : Set l1} + where + + is-L-consistent-antimonotic-logic : + {l2 l3 l4 : Level} + (logic₁ : modal-theory l2 i) → + (logic₂ : modal-theory l3 i) → + (theory : modal-theory l4 i) → + logic₁ ⊆ logic₂ → + is-L-consistent-theory logic₂ theory → + is-L-consistent-theory logic₁ theory + is-L-consistent-antimonotic-logic logic₁ logic₂ theory leq = + is-consistent-modal-logic-antimonotic + ( weak-modal-logic-closure (logic₁ ∪ theory)) + ( weak-modal-logic-closure (logic₂ ∪ theory)) + ( weak-modal-logic-closure-monotic + ( subset-union-subset-left logic₁ logic₂ theory leq)) +``` diff --git a/src/modal-logic/lindenbaum.lagda.md b/src/modal-logic/lindenbaum.lagda.md new file mode 100644 index 0000000000..7c7e475cbd --- /dev/null +++ b/src/modal-logic/lindenbaum.lagda.md @@ -0,0 +1,135 @@ +# Lindenbaum's lemma + +```agda +module modal-logic.lindenbaum where +``` + +
Imports + +```agda +open import foundation.dependent-pair-types +open import foundation.existential-quantification +open import foundation.propositional-resizing +open import foundation.propositional-truncations +open import foundation.unions-subtypes +open import foundation.universe-levels + +open import foundation-core.function-types +open import foundation-core.propositions +open import foundation-core.sets +open import foundation-core.subtypes + +open import modal-logic.axioms +open import modal-logic.logic-syntax +open import modal-logic.l-consistent-theories +open import modal-logic.l-complete-theories +open import modal-logic.weak-deduction + +open import order-theory.zorn +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l1 l2 l3 : Level} {i : Set l1} + (logic : modal-theory l2 i) + (contains-ax-k : ax-k i ⊆ logic) + (contains-ax-s : ax-s i ⊆ logic) + (zorn : Zorn-non-empty (lsuc (l1 ⊔ l2 ⊔ l3)) (l1 ⊔ l2 ⊔ l3) l3) + (prop-resize : propositional-resizing (l1 ⊔ l2 ⊔ l3) (lsuc (l1 ⊔ l2 ⊔ l3))) + (x@(theory , is-cons) : L-consistent-theory logic l3) + where + + lindenbaum : + ∃ ( L-complete-theory logic (l1 ⊔ l2 ⊔ l3)) + ( λ y → + ( modal-theory-L-consistent-theory logic x ⊆ + modal-theory-L-complete-theory logic y)) + lindenbaum = + apply-universal-property-trunc-Prop + ( extend-L-consistent-theory L prop-resize contains-ax-k' contains-ax-s' + ( zorn) + ( is-inhabited-L-consistent-theory L + ( is-weak-modal-logic-weak-modal-logic-closure) + ( is-cons))) + ( ∃-Prop _ _) + ( λ y → + ( intro-∃ + ( result-L-complete y) + ( subset-theory-transofrm-L-complete y))) + where + L : modal-theory (l1 ⊔ l2 ⊔ l3) i + L = weak-modal-logic-closure (logic ∪ theory) + + subset-logic-L : logic ⊆ L + subset-logic-L = + transitive-leq-subtype + ( logic) + ( logic ∪ theory) + ( weak-modal-logic-closure (logic ∪ theory)) + ( subset-axioms-weak-modal-logic) + ( subtype-union-left logic theory) + + contains-ax-k' : ax-k i ⊆ L + contains-ax-k' = + transitive-leq-subtype (ax-k i) logic L subset-logic-L contains-ax-k + + contains-ax-s' : ax-s i ⊆ L + contains-ax-s' = + transitive-leq-subtype (ax-s i) logic L subset-logic-L contains-ax-s + + is-L-consistent-result : + (y : L-complete-theory L (l1 ⊔ l2 ⊔ l3)) → + is-L-consistent-theory logic (modal-theory-L-complete-theory L y) + is-L-consistent-result y = + is-L-consistent-antimonotic-logic + ( logic) + ( weak-modal-logic-closure (logic ∪ theory)) + ( modal-theory-L-complete-theory L y) + ( subset-logic-L) + ( is-L-consistent-theory-modal-theory-L-complete-theory L y) + + result-L-consistent : + L-complete-theory L (l1 ⊔ l2 ⊔ l3) → + L-consistent-theory logic (l1 ⊔ l2 ⊔ l3) + pr1 (result-L-consistent y) = modal-theory-L-complete-theory L y + pr2 (result-L-consistent y) = is-L-consistent-result y + + is-L-complete-result : + (y : L-complete-theory L (l1 ⊔ l2 ⊔ l3)) → + is-L-complete-theory logic (result-L-consistent y) + is-L-complete-result y = + universal-L-complete-theory L logic + ( modal-theory-L-complete-theory L y) + ( is-L-consistent-theory-modal-theory-L-complete-theory L y) + ( is-L-consistent-result y) + ( is-L-complete-theory-L-consistent-theory L y) + + result-L-complete : + L-complete-theory L (l1 ⊔ l2 ⊔ l3) → + L-complete-theory logic (l1 ⊔ l2 ⊔ l3) + pr1 (pr1 (result-L-complete y)) = modal-theory-L-complete-theory L y + pr2 (pr1 (result-L-complete y)) = is-L-consistent-result y + pr2 (result-L-complete y) = is-L-complete-result y + + subset-theory-transofrm-L-complete : + (y : L-complete-theory L (l1 ⊔ l2 ⊔ l3)) → + theory ⊆ modal-theory-L-complete-theory logic (result-L-complete y) + subset-theory-transofrm-L-complete y = + ( transitive-leq-subtype theory L + ( modal-theory-L-complete-theory L y) + ( subset-logic-L-complete-theory L l3 y) + ( transitive-leq-subtype + ( theory) + ( logic ∪ theory) + ( L) + ( subset-axioms-weak-modal-logic) + ( subtype-union-right logic theory))) +``` diff --git a/src/modal-logic/logic-syntax.lagda.md b/src/modal-logic/logic-syntax.lagda.md index 77775d4574..104269081e 100644 --- a/src/modal-logic/logic-syntax.lagda.md +++ b/src/modal-logic/logic-syntax.lagda.md @@ -56,20 +56,17 @@ module _ mp : {a b : formula i} → axioms ⊢ a →ₘ b → axioms ⊢ a → axioms ⊢ b nec : {a : formula i} → axioms ⊢ a → axioms ⊢ □ a + -- TODO: rename to modal-logic-closure + modal-logic : modal-theory l2 i → modal-theory (l1 ⊔ l2) i + modal-logic axioms a = trunc-Prop (axioms ⊢ a) + is-modal-logic-Prop : modal-theory l2 i → Prop (l1 ⊔ l2) is-modal-logic-Prop theory = - implicit-Π-Prop (formula i) (λ a → function-Prop (theory ⊢ a) (theory a)) + leq-prop-subtype (modal-logic theory) theory is-modal-logic : modal-theory l2 i → UU (l1 ⊔ l2) is-modal-logic = type-Prop ∘ is-modal-logic-Prop - T-modal-logic : UU (l1 ⊔ lsuc l2) - T-modal-logic = Σ (modal-theory l2 i) is-modal-logic - - -- TODO: rename to modal-logic-closure - modal-logic : modal-theory l2 i → modal-theory (l1 ⊔ l2) i - modal-logic axioms a = trunc-Prop (axioms ⊢ a) - is-contradictory-modal-logic-Prop : modal-theory l2 i → Prop l2 is-contradictory-modal-logic-Prop logic = logic ⊥ @@ -139,19 +136,14 @@ module _ axioms-subset-modal-logic _ a H = unit-trunc-Prop (ax H) modal-logic-closed : - {l2 : Level} {axioms : formulas l2 i} → is-modal-logic (modal-logic axioms) + {l2 : Level} {axioms : formulas l2 i} {a : formula i} → + modal-logic axioms ⊢ a → + is-in-subtype (modal-logic axioms) a modal-logic-closed (ax x) = x - modal-logic-closed {axioms = axioms} {a} (mp tdab tda) = - apply-twice-universal-property-trunc-Prop - ( modal-logic-closed tdab) - ( modal-logic-closed tda) - ( modal-logic axioms a) - ( λ dab da → unit-trunc-Prop (mp dab da)) - modal-logic-closed {axioms = axioms} {a} (nec d) = - apply-universal-property-trunc-Prop - ( modal-logic-closed d) - ( modal-logic axioms a) - ( unit-trunc-Prop ∘ nec) + modal-logic-closed (mp dab da) = + modal-logic-mp (modal-logic-closed dab) (modal-logic-closed da) + modal-logic-closed (nec d) = + modal-logic-nec (modal-logic-closed d) subset-double-modal-logic : {l2 : Level} @@ -256,4 +248,14 @@ module _ theory-add-formula a (theory ∪ theory') theory-add-formula-union-right a theory theory' = union-swap-1-2 theory (Id-formula-Prop a) theory' + + inv-theory-add-formula-union-right : + (a : formula i) + {l2 l3 : Level} + (theory : formulas l2 i) + (theory' : formulas l3 i) → + theory-add-formula a (theory ∪ theory') ⊆ + theory ∪ theory-add-formula a theory' + inv-theory-add-formula-union-right a theory theory' = + union-swap-1-2 (Id-formula-Prop a) theory theory' ``` diff --git a/src/modal-logic/weak-deduction.lagda.md b/src/modal-logic/weak-deduction.lagda.md index 823f8e3a3a..a1ff5951d8 100644 --- a/src/modal-logic/weak-deduction.lagda.md +++ b/src/modal-logic/weak-deduction.lagda.md @@ -30,6 +30,7 @@ open import foundation-core.negation open import lists.concatenation-lists open import lists.lists +open import lists.lists-subtypes open import lists.reversing-lists open import modal-logic.axioms @@ -45,7 +46,7 @@ TODO ## Definition -```agda +````agda module _ {l1 l2 : Level} {i : Set l1} where @@ -123,18 +124,26 @@ module _ ind-weak-deduction : {l : Level} {axioms : modal-theory l2 i} - {P : {a : formula i} → axioms ⊢w a → UU l} → + (P : {a : formula i} → axioms ⊢w a → UU l) → ( {a : formula i} (in-axioms : is-in-subtype axioms a) → P (weak-deduction-ax in-axioms)) → - ( {a b : formula i} {dab : axioms ⊢w a →ₘ b} {da : axioms ⊢w a} → + ( {a b : formula i} (dab : axioms ⊢w a →ₘ b) (da : axioms ⊢w a) → P dab → P da → P (weak-deduction-mp dab da)) → {a : formula i} (wd : axioms ⊢w a) → P wd - ind-weak-deduction H-ax H-mp (ax in-axioms , _) = + ind-weak-deduction P H-ax H-mp (ax in-axioms , _) = H-ax in-axioms - ind-weak-deduction {P = P} H-ax H-mp (mp dba db , is-w-dba , is-w-db) = - H-mp - ( ind-weak-deduction {P = P} H-ax H-mp (dba , is-w-dba)) - ( ind-weak-deduction {P = P} H-ax H-mp (db , is-w-db)) + ind-weak-deduction P H-ax H-mp (mp dba db , is-w-dba , is-w-db) = + H-mp _ _ + ( ind-weak-deduction P H-ax H-mp (dba , is-w-dba)) + ( ind-weak-deduction P H-ax H-mp (db , is-w-db)) + + rec-weak-deduction : + {l : Level} {axioms : modal-theory l2 i} {P : UU l} → + ( {a : formula i} (in-axioms : is-in-subtype axioms a) → P) → + ( {a b : formula i} (dab : axioms ⊢w a →ₘ b) (da : axioms ⊢w a) → + P → P → P) → + {a : formula i} (wd : axioms ⊢w a) → P + rec-weak-deduction {P = P} = ind-weak-deduction (λ _ → P) weak-modal-logic-closure : modal-theory l2 i → modal-theory (l1 ⊔ l2) i weak-modal-logic-closure axioms a = trunc-Prop (axioms ⊢w a) @@ -153,7 +162,7 @@ module _ is-in-weak-deduction-closure-weak-deduction = unit-trunc-Prop module _ - {l1 l2 : Level} {i : Set l1} (axioms : modal-theory l2 i) + {l1 l2 : Level} {i : Set l1} {axioms : modal-theory l2 i} where weak-modal-logic-closure-ax : @@ -177,7 +186,8 @@ module _ module _ {l1 l2 : Level} {i : Set l1} - (theory : modal-theory l2 i) (is-w : is-weak-modal-logic theory) + {theory : modal-theory l2 i} + (is-w : is-weak-modal-logic theory) where weak-modal-logic-mp : @@ -187,9 +197,9 @@ module _ is-in-subtype theory b weak-modal-logic-mp {a} {b} dab da = is-w b - ( weak-modal-logic-closure-mp theory - ( weak-modal-logic-closure-ax theory dab) - ( weak-modal-logic-closure-ax theory da)) + ( weak-modal-logic-closure-mp + ( weak-modal-logic-closure-ax dab) + ( weak-modal-logic-closure-ax da)) module _ {l1 l2 : Level} {i : Set l1} {axioms : modal-theory l2 i} @@ -199,10 +209,10 @@ module _ {a : formula i} → weak-modal-logic-closure axioms ⊢w a → is-in-subtype (weak-modal-logic-closure axioms) a weak-modal-logic-closed = - ind-weak-deduction + ind-weak-deduction _ ( λ in-axioms → in-axioms) - ( λ in-logic-ab in-logic-a → - ( weak-modal-logic-closure-mp axioms in-logic-ab in-logic-a)) + ( λ _ _ in-logic-ab in-logic-a → + ( weak-modal-logic-closure-mp in-logic-ab in-logic-a)) is-weak-modal-logic-weak-modal-logic-closure : is-weak-modal-logic (weak-modal-logic-closure axioms) @@ -211,24 +221,8 @@ module _ ( weak-modal-logic-closure axioms a) ( weak-modal-logic-closed) - -- weak-modal-logic-closure-closed : - -- is-weak-modal-logic axioms → - -- weak-modal-logic-closure axioms ⊆ axioms - -- weak-modal-logic-closure-closed is-logic a = - -- map-universal-property-trunc-Prop (axioms a) is-logic - --- module _ --- {l1 l2 : Level} {i : Set l1} {axioms : modal-theory l2 i} --- where - subset-axioms-weak-modal-logic : axioms ⊆ weak-modal-logic-closure axioms - subset-axioms-weak-modal-logic a = weak-modal-logic-closure-ax axioms - --- subset-double-weak-modal-logic-closure : --- weak-modal-logic-closure (weak-modal-logic-closure axioms) ⊆ --- weak-modal-logic-closure axioms --- subset-double-weak-modal-logic-closure = --- weak-modal-logic-closure-closed is-weak-modal-logic-weak-modal-logic-closure + subset-axioms-weak-modal-logic a = weak-modal-logic-closure-ax module _ {l1 l2 l3 : Level} {i : Set l1} @@ -239,16 +233,16 @@ module _ weak-deduction-monotic : {a : formula i} → ax₁ ⊢w a → ax₂ ⊢w a weak-deduction-monotic = - ind-weak-deduction + ind-weak-deduction _ ( λ {a} in-axioms → weak-deduction-ax (leq a in-axioms)) - ( λ dab da → weak-deduction-mp dab da) + ( λ _ _ dab da → weak-deduction-mp dab da) weak-modal-logic-closure-monotic : weak-modal-logic-closure ax₁ ⊆ weak-modal-logic-closure ax₂ weak-modal-logic-closure-monotic a = map-universal-property-trunc-Prop ( weak-modal-logic-closure ax₂ a) - ( unit-trunc-Prop ∘ weak-deduction-monotic) + ( is-in-weak-deduction-closure-weak-deduction ∘ weak-deduction-monotic) module _ {l1 l2 l3 : Level} {i : Set l1} @@ -267,10 +261,22 @@ module _ ( is-weak-modal-logic-weak-modal-logic-closure) ( weak-modal-logic-closure-monotic leq) +module _ + {l1 : Level} {i : Set l1} + where + + backward-subset-head-add : + (a : formula i) (l : list (formula i)) → + list-subtype (cons a l) ⊆ theory-add-formula a (list-subtype l) + backward-subset-head-add a l = + subset-list-subtype-cons + ( theory-add-formula a (list-subtype l)) + ( formula-in-add-formula a (list-subtype l)) + ( subset-add-formula a (list-subtype l)) + module _ {l1 l2 : Level} {i : Set l1} (axioms : modal-theory l2 i) - -- (is-w : is-weak-modal-logic logic) where backward-deduction-lemma : @@ -304,22 +310,22 @@ module _ theory-add-formula a axioms ⊢w b → is-in-subtype (weak-modal-logic-closure axioms) (a →ₘ b) forward-deduction-lemma a = - ind-weak-deduction + ind-weak-deduction _ ( λ {b} b-in-axioms → ( elim-theory-add-formula a axioms ( λ x → weak-modal-logic-closure axioms (a →ₘ x)) ( is-in-weak-deduction-closure-weak-deduction (deduction-a->a a)) ( λ x x-in-axioms → - ( weak-modal-logic-closure-mp axioms - ( weak-modal-logic-closure-ax axioms + ( weak-modal-logic-closure-mp + ( weak-modal-logic-closure-ax ( contains-ax-k (x →ₘ a →ₘ x) (x , a , refl))) - ( weak-modal-logic-closure-ax axioms x-in-axioms))) + ( weak-modal-logic-closure-ax x-in-axioms))) ( b) ( b-in-axioms))) - ( λ {b} {c} dabc dab → - ( weak-modal-logic-closure-mp axioms - ( weak-modal-logic-closure-mp axioms - ( weak-modal-logic-closure-ax axioms + ( λ {b} {c} _ _ dabc dab → + ( weak-modal-logic-closure-mp + ( weak-modal-logic-closure-mp + ( weak-modal-logic-closure-ax ( contains-ax-s _ (a , b , c , refl))) ( dabc)) ( dab))) @@ -337,50 +343,74 @@ module _ ( weak-modal-logic-closure (theory-add-formula a axioms) b) ( is-in-weak-deduction-closure-weak-deduction ∘ backward-deduction-lemma) +``` --- module _ --- -- TODO: maybe just in axioms? --- (contains-ax-k : ax-k i ⊆ weak-modal-logic axioms) --- (contains-ax-s : ax-s i ⊆ weak-modal-logic axioms) --- where +### TODO: List of assumptions --- forward-deduction-lemma : --- (a : formula i) {b : formula i} → --- theory-add-formula a axioms ▷ b → --- is-in-subtype (weak-modal-logic axioms) (a →ₘ b) --- forward-deduction-lemma a {b} (w-ax x) = --- elim-disjunction-Prop --- ( Id-formula-Prop a b) --- ( axioms b) --- ( (weak-modal-logic axioms) (a →ₘ b)) --- ( (λ { refl → deduction-a->a a}) --- , (λ in-axioms → --- ( apply-universal-property-trunc-Prop --- ( contains-ax-k _ (b , a , refl)) --- ( (weak-modal-logic axioms) (a →ₘ b)) --- ( λ x → unit-trunc-Prop (w-mp x (w-ax in-axioms)))))) --- ( x) --- forward-deduction-lemma a {b} (w-mp {c} twdcb twdc) = --- apply-three-times-universal-property-trunc-Prop --- ( forward-deduction-lemma a twdcb) --- ( forward-deduction-lemma a twdc) --- ( contains-ax-s _ (a , c , b , refl)) --- ( (weak-modal-logic axioms) (a →ₘ b)) --- ( λ wdacb wdac x → --- ( unit-trunc-Prop (w-mp (w-mp x wdacb) wdac))) +```agda +module _ + {l1 : Level} {i : Set l1} + where --- deduction-lemma : --- (a b : formula i) → --- weak-modal-logic (theory-add-formula a axioms) b ⇔ --- weak-modal-logic axioms (a →ₘ b) --- pr1 (deduction-lemma a b) = --- map-universal-property-trunc-Prop --- ( (weak-modal-logic axioms) (a →ₘ b)) --- ( forward-deduction-lemma a) --- pr2 (deduction-lemma a b) = --- map-universal-property-trunc-Prop --- ( weak-modal-logic (theory-add-formula a axioms) b) --- ( backward-deduction-lemma) + list-assumptions-weak-deduction : + {l2 : Level} {theory : modal-theory l2 i} {a : formula i} → + theory ⊢w a → list (formula i) + list-assumptions-weak-deduction = + rec-weak-deduction + ( λ {a} _ → cons a nil) + ( λ _ _ l1 l2 → concat-list l1 l2) + + subset-theory-list-assumptions-weak-deduction : + {l2 : Level} {theory : modal-theory l2 i} {a : formula i} → + (d : theory ⊢w a) → + list-subtype (list-assumptions-weak-deduction d) ⊆ theory + subset-theory-list-assumptions-weak-deduction {theory = theory} = + ind-weak-deduction + ( λ d → list-subtype (list-assumptions-weak-deduction d) ⊆ theory) + ( λ {a} in-axioms → + ( subset-list-subtype-cons theory in-axioms + ( subset-list-subtype-nil theory))) + ( λ dab da sub1 sub2 → + ( transitive-leq-subtype + ( list-subtype + ( concat-list + ( list-assumptions-weak-deduction dab) + ( list-assumptions-weak-deduction da))) + ( list-subtype (list-assumptions-weak-deduction dab) ∪ + list-subtype (list-assumptions-weak-deduction da)) + ( theory) + ( subtype-union-both + ( list-subtype (list-assumptions-weak-deduction dab)) + ( list-subtype (list-assumptions-weak-deduction da)) + ( theory) + ( sub1) + ( sub2)) + ( subset-list-subtype-concat-union))) + + is-assumptions-list-assumptions-weak-deduction : + {l2 : Level} {theory : modal-theory l2 i} {a : formula i} → + (d : theory ⊢w a) → + list-subtype (list-assumptions-weak-deduction d) ⊢w a + is-assumptions-list-assumptions-weak-deduction {theory = theory} = + ind-weak-deduction + ( λ {a} d → list-subtype (list-assumptions-weak-deduction d) ⊢w a) + ( λ _ → weak-deduction-ax head-in-list-subtype) + ( λ dab da rab ra → + ( weak-deduction-monotic + {ax₁ = list-subtype (list-assumptions-weak-deduction dab) ∪ + list-subtype (list-assumptions-weak-deduction da)} + ( subset-list-subtype-union-concat) + ( weak-deduction-mp + ( weak-deduction-monotic + ( subtype-union-left + ( list-subtype (list-assumptions-weak-deduction dab)) + ( list-subtype (list-assumptions-weak-deduction da))) + ( rab)) + ( weak-deduction-monotic + ( subtype-union-right + ( list-subtype (list-assumptions-weak-deduction dab)) + ( list-subtype (list-assumptions-weak-deduction da))) + ( ra))))) ``` ### Old definition @@ -1014,3 +1044,4 @@ module _ -- ( (weak-modal-logic axioms) b) -- ( λ x y z → unit-trunc-Prop (w-mp (w-mp z y) x)) ``` +```` diff --git a/src/order-theory/chains-posets.lagda.md b/src/order-theory/chains-posets.lagda.md index 8e3f7b7748..72ce9a41dc 100644 --- a/src/order-theory/chains-posets.lagda.md +++ b/src/order-theory/chains-posets.lagda.md @@ -57,9 +57,18 @@ module _ sub-preorder-chain-Poset = sub-preorder-chain-Preorder (preorder-Poset X) C + is-chain-Subposet-chain-Poset : + is-chain-Subposet X sub-preorder-chain-Poset + is-chain-Subposet-chain-Poset = + is-chain-Subpreorder-chain-Preorder (preorder-Poset X) C + type-chain-Poset : UU (l1 ⊔ l3) type-chain-Poset = type-chain-Preorder (preorder-Poset X) C + type-Poset-type-chain-Poset : type-chain-Poset → type-Poset X + type-Poset-type-chain-Poset = + type-Preorder-type-chain-Preorder (preorder-Poset X) C + module _ {l1 l2 : Level} (X : Poset l1 l2) where @@ -97,8 +106,7 @@ module _ is-chain-upper-bound = type-Prop ∘ is-chain-upper-bound-Prop has-chain-upper-bound-Prop : Prop (l1 ⊔ l2 ⊔ l3) - has-chain-upper-bound-Prop = - exists-Prop (type-Poset X) is-chain-upper-bound-Prop + has-chain-upper-bound-Prop = ∃-Prop (type-Poset X) is-chain-upper-bound has-chain-upper-bound : UU (l1 ⊔ l2 ⊔ l3) has-chain-upper-bound = type-Prop has-chain-upper-bound-Prop diff --git a/src/order-theory/chains-preorders.lagda.md b/src/order-theory/chains-preorders.lagda.md index edcc516f32..c53a1bb8f1 100644 --- a/src/order-theory/chains-preorders.lagda.md +++ b/src/order-theory/chains-preorders.lagda.md @@ -58,9 +58,17 @@ module _ sub-preorder-chain-Preorder : type-Preorder X → Prop l3 sub-preorder-chain-Preorder = pr1 C + is-chain-Subpreorder-chain-Preorder : + is-chain-Subpreorder X sub-preorder-chain-Preorder + is-chain-Subpreorder-chain-Preorder = pr2 C + type-chain-Preorder : UU (l1 ⊔ l3) type-chain-Preorder = type-subtype sub-preorder-chain-Preorder + type-Preorder-type-chain-Preorder : type-chain-Preorder → type-Preorder X + type-Preorder-type-chain-Preorder = + inclusion-subtype sub-preorder-chain-Preorder + module _ {l1 l2 : Level} (X : Preorder l1 l2) where diff --git a/src/order-theory/maximal-elements-posets.lagda.md b/src/order-theory/maximal-elements-posets.lagda.md index c52a961f44..b807d310ab 100644 --- a/src/order-theory/maximal-elements-posets.lagda.md +++ b/src/order-theory/maximal-elements-posets.lagda.md @@ -7,26 +7,13 @@ module order-theory.maximal-elements-posets where
Imports ```agda -open import foundation.action-on-identifications-functions open import foundation.dependent-pair-types -open import foundation.empty-types -open import foundation.existential-quantification -open import foundation.identity-types -open import foundation.inhabited-types -open import foundation.law-of-excluded-middle -open import foundation.logical-equivalences -open import foundation.propositional-truncations -open import foundation.propositions -open import foundation.subtypes -open import foundation.unit-type open import foundation.universe-levels -open import foundation-core.coproduct-types +open import foundation-core.identity-types +open import foundation-core.propositions -open import order-theory.chains-posets open import order-theory.posets -open import order-theory.subposets -open import order-theory.upper-bounds-posets ```
@@ -46,7 +33,8 @@ module _ is-maximal-element-Poset-Prop x = Π-Prop ( type-Poset X) - ( λ y → function-Prop (leq-Poset X x y) (y = x , is-set-type-Poset X y x)) + ( λ y → + ( function-Prop (leq-Poset X x y) (y = x , is-set-type-Poset X y x))) is-maximal-element-Poset : type-Poset X → UU (l1 ⊔ l2) is-maximal-element-Poset x = type-Prop (is-maximal-element-Poset-Prop x) @@ -55,73 +43,4 @@ module _ (x : type-Poset X) → is-prop (is-maximal-element-Poset x) is-prop-is-maximal-element-Poset x = is-prop-type-Prop (is-maximal-element-Poset-Prop x) - - has-maximal-element-Poset : UU (l1 ⊔ l2) - has-maximal-element-Poset = Σ (type-Poset X) is-maximal-element-Poset - -module _ - (l1 l2 l3 : Level) - where - - Zorn-Prop : Prop (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3) - Zorn-Prop = - Π-Prop - ( Poset l1 l2) - ( λ X → - ( function-Prop - ( is-inhabited (type-Poset X)) - ( function-Prop - ( (C : chain-Poset l3 X) → - is-inhabited (type-chain-Poset X C) → - has-chain-upper-bound X C) - ( ∃-Prop (type-Poset X) (is-maximal-element-Poset X))))) - - Zorn : UU (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3) - Zorn = type-Prop Zorn-Prop - - Zorn'-Prop : Prop (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3) - Zorn'-Prop = - Π-Prop - ( Poset l1 l2) - ( λ X → - ( function-Prop - ( (C : chain-Poset l3 X) → has-chain-upper-bound X C) - ( ∃-Prop (type-Poset X) (is-maximal-element-Poset X)))) - - Zorn' : UU (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3) - Zorn' = type-Prop Zorn'-Prop - - Zorn'-Zorn : Zorn → Zorn' - Zorn'-Zorn zorn X H = - zorn X - ( apply-universal-property-trunc-Prop - ( H - ( pair - (λ _ → raise-empty-Prop l3) - (λ (_ , f) → raise-ex-falso l3 f))) - ( is-inhabited-Prop (type-Poset X)) - ( λ (x , _) → unit-trunc-Prop x)) - ( λ C _ → H C) - - module _ - (lem : LEM (l1 ⊔ l3)) - where - - Zorn-Zorn' : Zorn' → Zorn - Zorn-Zorn' zorn' X is-inhabited-X H = zorn' X chain-upper-bound - where - chain-upper-bound : (C : chain-Poset l3 X) → has-chain-upper-bound X C - chain-upper-bound C with lem (is-inhabited-Prop (type-chain-Poset X C)) - ... | inl is-inhabited-C = H C is-inhabited-C - ... | inr is-empty-C = - apply-universal-property-trunc-Prop - ( is-inhabited-X) - ( has-chain-upper-bound-Prop X C) - ( λ x → - ( intro-∃ x - ( λ (y , y-in-C) → ex-falso (is-empty-C (intro-∃ y y-in-C))))) - - Zorn-iff-Zorn' : Zorn-Prop ⇔ Zorn'-Prop - pr1 (Zorn-iff-Zorn') = Zorn'-Zorn - pr2 (Zorn-iff-Zorn') = Zorn-Zorn' ``` diff --git a/src/order-theory/zorn.lagda.md b/src/order-theory/zorn.lagda.md new file mode 100644 index 0000000000..0077eedb2c --- /dev/null +++ b/src/order-theory/zorn.lagda.md @@ -0,0 +1,101 @@ +# Zorn's lemma + +```agda +module order-theory.zorn where +``` + +
Imports + +```agda +open import foundation.dependent-pair-types +open import foundation.empty-types +open import foundation.existential-quantification +open import foundation.inhabited-types +open import foundation.law-of-excluded-middle +open import foundation.logical-equivalences +open import foundation.propositional-truncations +open import foundation.universe-levels + +open import foundation-core.propositions +open import foundation-core.coproduct-types + +open import order-theory.chains-posets +open import order-theory.posets +open import order-theory.maximal-elements-posets +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + (l1 l2 l3 : Level) + where + + Zorn-Prop : Prop (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3) + Zorn-Prop = + Π-Prop + ( Poset l1 l2) + ( λ X → + ( function-Prop + ( (C : chain-Poset l3 X) → has-chain-upper-bound X C) + ( ∃-Prop (type-Poset X) (is-maximal-element-Poset X)))) + + Zorn : UU (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3) + Zorn = type-Prop Zorn-Prop + + Zorn-non-empty-Prop : Prop (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3) + Zorn-non-empty-Prop = + Π-Prop + ( Poset l1 l2) + ( λ X → + ( function-Prop + ( is-inhabited (type-Poset X)) + ( function-Prop + ( (C : chain-Poset l3 X) → + is-inhabited (type-chain-Poset X C) → + has-chain-upper-bound X C) + ( ∃-Prop (type-Poset X) (is-maximal-element-Poset X))))) + + Zorn-non-empty : UU (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3) + Zorn-non-empty = type-Prop Zorn-non-empty-Prop + + Zorn-Zorn-non-empty : Zorn-non-empty → Zorn + Zorn-Zorn-non-empty zorn X H = + zorn X + ( apply-universal-property-trunc-Prop + ( H + ( pair + ( λ _ → raise-empty-Prop l3) + ( λ (_ , f) → raise-ex-falso l3 f))) + ( is-inhabited-Prop (type-Poset X)) + ( λ (x , _) → unit-trunc-Prop x)) + ( λ C _ → H C) + + module _ + (lem : LEM (l1 ⊔ l3)) + where + + Zorn-non-empty-Zorn : Zorn → Zorn-non-empty + Zorn-non-empty-Zorn zorn X is-inhabited-X H = zorn X chain-upper-bound + where + chain-upper-bound : (C : chain-Poset l3 X) → has-chain-upper-bound X C + chain-upper-bound C with lem (is-inhabited-Prop (type-chain-Poset X C)) + ... | inl is-inhabited-C = H C is-inhabited-C + ... | inr is-empty-C = + apply-universal-property-trunc-Prop + ( is-inhabited-X) + ( has-chain-upper-bound-Prop X C) + ( λ x → + ( intro-∃ x + ( λ (y , y-in-C) → ex-falso (is-empty-C (intro-∃ y y-in-C))))) + + iff-Zorn-non-empty-Zorn : Zorn-non-empty-Prop ⇔ Zorn-Prop + pr1 (iff-Zorn-non-empty-Zorn) = Zorn-Zorn-non-empty + pr2 (iff-Zorn-non-empty-Zorn) = Zorn-non-empty-Zorn +``` From 998bd15b19af0fc4fb77b58e5863e11f17be53a9 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Fri, 3 May 2024 17:43:48 +0300 Subject: [PATCH 41/63] Remove old code from weak-deduction --- src/modal-logic/weak-deduction.lagda.md | 635 +----------------------- 1 file changed, 1 insertion(+), 634 deletions(-) diff --git a/src/modal-logic/weak-deduction.lagda.md b/src/modal-logic/weak-deduction.lagda.md index a1ff5951d8..a8d4553912 100644 --- a/src/modal-logic/weak-deduction.lagda.md +++ b/src/modal-logic/weak-deduction.lagda.md @@ -46,7 +46,7 @@ TODO ## Definition -````agda +```agda module _ {l1 l2 : Level} {i : Set l1} where @@ -412,636 +412,3 @@ module _ ( list-subtype (list-assumptions-weak-deduction da))) ( ra))))) ``` - -### Old definition - -```agda --- infix 5 _▷_ - --- data _▷_ (axioms : modal-theory l2 i) : formula i → UU (l1 ⊔ l2) where --- w-ax : {a : formula i} → is-in-subtype axioms a → axioms ▷ a --- w-mp : {a b : formula i} → axioms ▷ a →ₘ b → axioms ▷ a → axioms ▷ b - --- weak-modal-logic : modal-theory l2 i → modal-theory (l1 ⊔ l2) i --- weak-modal-logic axioms a = trunc-Prop (axioms ▷ a) - --- is-weak-modal-logic-Prop : modal-theory l2 i → Prop (l1 ⊔ l2) --- is-weak-modal-logic-Prop theory = --- implicit-Π-Prop (formula i) (λ a → function-Prop (theory ▷ a) (theory a)) - --- is-weak-modal-logic : modal-theory l2 i → UU (l1 ⊔ l2) --- is-weak-modal-logic = type-Prop ∘ is-weak-modal-logic-Prop - --- module _ --- {l1 l2 : Level} {i : Set l1} {axioms : modal-theory l2 i} --- where - --- weak-modal-logic-ax : --- {a : formula i} → --- is-in-subtype axioms a → --- is-in-subtype (weak-modal-logic axioms) a --- weak-modal-logic-ax = unit-trunc-Prop ∘ w-ax - --- weak-modal-logic-mp : --- {a b : formula i} → --- is-in-subtype (weak-modal-logic axioms) (a →ₘ b) → --- is-in-subtype (weak-modal-logic axioms) a → --- is-in-subtype (weak-modal-logic axioms) b --- weak-modal-logic-mp {a} {b} twdab twda = --- apply-twice-universal-property-trunc-Prop twdab twda --- ( weak-modal-logic axioms b) --- ( λ wdab wda → unit-trunc-Prop (w-mp wdab wda)) - --- module _ --- {l1 l2 : Level} {i : Set l1} (axioms : modal-theory l2 i) --- where - --- deduction-weak-deduction : {a : formula i} → axioms ▷ a → axioms ⊢ a --- deduction-weak-deduction (w-ax x) = ax x --- deduction-weak-deduction (w-mp dab da) = --- mp (deduction-weak-deduction dab) (deduction-weak-deduction da) - --- weak-modal-logic-subset-modal-logic : --- weak-modal-logic axioms ⊆ modal-logic axioms --- weak-modal-logic-subset-modal-logic a = --- map-universal-property-trunc-Prop --- ( modal-logic axioms a) --- ( unit-trunc-Prop ∘ deduction-weak-deduction) - --- module _ --- {l1 l2 l3 : Level} {i : Set l1} --- {axioms₁ : modal-theory l2 i} {axioms₂ : modal-theory l3 i} --- (e : equiv-subtypes axioms₁ axioms₂) --- where - --- deduction-equiv-axioms : --- {a : formula i} → axioms₁ ▷ a → axioms₂ ▷ a --- deduction-equiv-axioms {a} (w-ax x) = --- w-ax (map-equiv (e a) x) --- deduction-equiv-axioms (w-mp wdab wda) = --- w-mp (deduction-equiv-axioms wdab) (deduction-equiv-axioms wda) - --- module _ --- {l1 : Level} {i : Set l1} --- where - --- axioms-subset-weak-modal-logic : --- {l2 : Level} (axioms : modal-theory l2 i) → axioms ⊆ weak-modal-logic axioms --- axioms-subset-weak-modal-logic _ a H = unit-trunc-Prop (w-ax H) - --- -- TODO: rename to is-weak-modal-logic-weak-modal-logic-closure --- weak-modal-logic-closed : --- {l2 : Level} {axioms : modal-theory l2 i} → --- is-weak-modal-logic (weak-modal-logic axioms) --- weak-modal-logic-closed (w-ax x) = x --- weak-modal-logic-closed {axioms = axioms} {a} (w-mp tdab tda) = --- apply-twice-universal-property-trunc-Prop --- ( weak-modal-logic-closed tdab) --- ( weak-modal-logic-closed tda) --- ( weak-modal-logic axioms a) --- (λ dab da → unit-trunc-Prop (w-mp dab da)) - --- weak-modal-logic-closed' : --- {l2 : Level} {theory : modal-theory l2 i} → --- is-weak-modal-logic theory → --- weak-modal-logic theory ⊆ theory --- weak-modal-logic-closed' {theory = theory} is-logic a = --- map-universal-property-trunc-Prop (theory a) is-logic - --- subset-double-weak-modal-logic : --- {l2 : Level} --- (axioms : modal-theory l2 i) → --- weak-modal-logic (weak-modal-logic axioms) ⊆ weak-modal-logic axioms --- subset-double-weak-modal-logic axioms = --- weak-modal-logic-closed' weak-modal-logic-closed - --- weak-modal-logic-idempotent : --- {l2 : Level} --- (axioms : modal-theory l2 i) → --- weak-modal-logic axioms = weak-modal-logic (weak-modal-logic axioms) --- weak-modal-logic-idempotent axioms = --- antisymmetric-leq-subtype _ _ --- ( axioms-subset-weak-modal-logic (weak-modal-logic axioms)) --- ( subset-double-weak-modal-logic axioms) - --- module _ --- {l1 l2 l3 : Level} {i : Set l1} --- {ax₁ : modal-theory l2 i} --- {ax₂ : modal-theory l3 i} --- (leq : ax₁ ⊆ ax₂) --- where - --- weak-deduction-monotic : {a : formula i} → ax₁ ▷ a → ax₂ ▷ a --- weak-deduction-monotic (w-ax x) = w-ax (leq _ x) --- weak-deduction-monotic (w-mp wdab wda) = --- w-mp (weak-deduction-monotic wdab) (weak-deduction-monotic wda) - --- weak-modal-logic-monotic : weak-modal-logic ax₁ ⊆ weak-modal-logic ax₂ --- weak-modal-logic-monotic a = --- map-universal-property-trunc-Prop --- ( weak-modal-logic ax₂ a) --- ( unit-trunc-Prop ∘ weak-deduction-monotic) - --- module _ --- {l1 l2 l3 : Level} {i : Set l1} --- {ax₁ : modal-theory l2 i} --- {ax₂ : modal-theory l3 i} --- where - --- subset-weak-modal-logic-if-subset-axioms : --- ax₁ ⊆ weak-modal-logic ax₂ → weak-modal-logic ax₁ ⊆ weak-modal-logic ax₂ --- subset-weak-modal-logic-if-subset-axioms leq = --- transitive-leq-subtype --- ( weak-modal-logic ax₁) --- ( weak-modal-logic (weak-modal-logic ax₂)) --- ( weak-modal-logic ax₂) --- ( subset-double-weak-modal-logic ax₂) --- ( weak-modal-logic-monotic leq) - --- -- TODO: move to lists - --- in-list : {l : Level} {A : UU l} → list A → A → Prop l --- in-list l a = trunc-Prop (a ∈-list l) - --- subset-in-tail : --- {l : Level} {A : UU l} (a : A) (l : list A) → --- in-list l ⊆ in-list (cons a l) --- subset-in-tail a l x = --- map-universal-property-trunc-Prop --- ( in-list (cons a l) x) --- ( unit-trunc-Prop ∘ (is-in-tail x a l)) - --- in-concat-list : --- {l : Level} {A : UU l} (l1 l2 : list A) → --- (a : A) → a ∈-list (concat-list l1 l2) → (a ∈-list l1) + (a ∈-list l2) --- in-concat-list nil l2 a in-list = inr in-list --- in-concat-list (cons x l1) l2 a (is-head a _) = inl (is-head a l1) --- in-concat-list (cons x l1) l2 a (is-in-tail _ _ _ in-tail) --- with in-concat-list l1 l2 a in-tail --- ... | inl in-l1 = inl (is-in-tail a x l1 in-l1) --- ... | inr in-l2 = inr in-l2 - --- subset-reversing-list : --- {l : Level} {A : UU l} (l : list A) → in-list l ⊆ in-list (reverse-list l) --- subset-reversing-list l x = --- map-universal-property-trunc-Prop --- ( in-list (reverse-list l) x) --- ( unit-trunc-Prop ∘ forward-contains-reverse-list x l) - --- in-sum-concat-list-Prop : --- {l : Level} {A : UU l} (l1 l2 : list A) → --- (a : A) → a ∈-list (concat-list l1 l2) → --- type-Prop (in-list l1 a) + type-Prop (in-list l2 a) --- in-sum-concat-list-Prop l1 l2 a in-concat with in-concat-list l1 l2 a in-concat --- ... | inl in-l1 = inl (unit-trunc-Prop in-l1) --- ... | inr in-l2 = inr (unit-trunc-Prop in-l2) - --- union-lists-concat : --- {l : Level} {A : UU l} (l1 l2 : list A) → --- in-list (concat-list l1 l2) ⊆ union-subtype (in-list l1) (in-list l2) --- union-lists-concat l1 l2 a = --- map-universal-property-trunc-Prop --- ( union-subtype (in-list l1) (in-list l2) a) --- ( unit-trunc-Prop ∘ (in-sum-concat-list-Prop l1 l2 a)) - --- in-concat-left : --- {l : Level} {A : UU l} (l1 l2 : list A) --- {a : A} → a ∈-list l1 → a ∈-list (concat-list l1 l2) --- in-concat-left _ _ (is-head a _) = --- is-head a _ --- in-concat-left _ l2 (is-in-tail a x l1 in-l1) = --- is-in-tail a x (concat-list l1 l2) (in-concat-left l1 l2 in-l1) - --- in-concat-right : --- {l : Level} {A : UU l} (l1 l2 : list A) --- {a : A} → a ∈-list l2 → a ∈-list (concat-list l1 l2) --- in-concat-right nil l2 in-l2 = in-l2 --- in-concat-right (cons x l1) l2 in-l2 = --- is-in-tail _ x (concat-list l1 l2) (in-concat-right l1 l2 in-l2) - --- in-concat-list-sum-Prop : --- {l : Level} {A : UU l} (l1 l2 : list A) → --- (a : A) → type-Prop (in-list l1 a) + type-Prop (in-list l2 a) → --- type-Prop (in-list (concat-list l1 l2) a) --- in-concat-list-sum-Prop l1 l2 a (inl in-l1) = --- apply-universal-property-trunc-Prop --- ( in-l1) --- ( in-list (concat-list l1 l2) a) --- ( unit-trunc-Prop ∘ (in-concat-left l1 l2)) --- in-concat-list-sum-Prop l1 l2 a (inr in-l2) = --- apply-universal-property-trunc-Prop --- ( in-l2) --- ( in-list (concat-list l1 l2) a) --- ( unit-trunc-Prop ∘ (in-concat-right l1 l2)) - --- in-concat-lists-union : --- {l : Level} {A : UU l} (l1 l2 : list A) → --- union-subtype (in-list l1) (in-list l2) ⊆ in-list (concat-list l1 l2) --- in-concat-lists-union l1 l2 a = --- map-universal-property-trunc-Prop --- ( in-list (concat-list l1 l2) a) --- ( in-concat-list-sum-Prop l1 l2 a) - --- subset-in-concat-left : --- {l : Level} {A : UU l} (l1 l2 : list A) → --- in-list l1 ⊆ in-list (concat-list l1 l2) --- subset-in-concat-left l1 l2 = --- transitive-leq-subtype --- ( in-list l1) --- ( union-subtype (in-list l1) (in-list l2)) --- ( in-list (concat-list l1 l2)) --- ( in-concat-lists-union l1 l2) --- ( subtype-union-left (in-list l1) (in-list l2)) - --- subset-in-concat-right : --- {l : Level} {A : UU l} (l1 l2 : list A) → --- in-list l2 ⊆ in-list (concat-list l1 l2) --- subset-in-concat-right l1 l2 = --- transitive-leq-subtype --- ( in-list l2) --- ( union-subtype (in-list l1) (in-list l2)) --- ( in-list (concat-list l1 l2)) --- ( in-concat-lists-union l1 l2) --- ( subtype-union-right (in-list l1) (in-list l2)) - --- empty-in-list-nil : --- {l : Level} {A : UU l} {x : A} → is-in-subtype (in-list nil) x → empty --- empty-in-list-nil = map-universal-property-trunc-Prop empty-Prop ( λ ()) - --- equiv-subset-head-tail : --- {l1 l2 : Level} {A : UU l1} (x : A) (xs : list A) (a : subtype l2 A) → --- (leq-prop-subtype (in-list (cons x xs)) a) ⇔ --- product-Prop (a x) (leq-prop-subtype (in-list xs) a) --- pr1 (equiv-subset-head-tail x xs a) sub = --- pair --- ( sub x (unit-trunc-Prop (is-head x xs))) --- ( transitive-leq-subtype --- ( in-list xs) --- ( in-list (cons x xs)) --- ( a) --- ( sub) --- ( subset-in-tail x xs)) --- pr2 (equiv-subset-head-tail x xs a) (x-in-a , sub) y = --- map-universal-property-trunc-Prop --- ( a y) --- ( λ --- { (is-head .x .xs) → x-in-a --- ; (is-in-tail .y .x .xs y-in-list) → sub y (unit-trunc-Prop y-in-list)}) - --- lists-in-union-lists : --- {l1 l2 l3 : Level} {A : UU l1} --- (xs : list A) (a : subtype l2 A) (b : subtype l3 A) → --- in-list xs ⊆ union-subtype a b → --- ∃ --- ( list A) --- ( λ xsl → --- ( Σ --- ( list A) --- ( λ xsr → --- ( in-list xs ⊆ union-subtype (in-list xsl) (in-list xsr)) × --- ( in-list xsl ⊆ a) × --- ( in-list xsr ⊆ b)))) --- lists-in-union-lists nil a b sub = --- unit-trunc-Prop --- ( triple nil nil --- ( triple --- ( λ _ → ex-falso ∘ empty-in-list-nil) --- ( λ _ → ex-falso ∘ empty-in-list-nil) --- ( λ _ → ex-falso ∘ empty-in-list-nil))) --- lists-in-union-lists (cons x xs) a b sub = --- apply-twice-universal-property-trunc-Prop --- ( lists-in-union-lists xs a b --- ( transitive-leq-subtype --- ( in-list xs) --- ( in-list (cons x xs)) --- ( union-subtype a b) --- ( sub) --- ( subset-in-tail x xs))) --- ( sub x (unit-trunc-Prop (is-head x xs))) --- ( ∃-Prop _ (λ _ → Σ _ _)) --- ( λ (xsl , xsr , sub-lists , sub-xsl , sub-xsr) → --- ( ind-coproduct --- ( λ _ → _) --- ( λ x-in-a → --- ( unit-trunc-Prop --- ( triple (cons x xsl) xsr --- ( triple --- ( λ y → --- ( map-universal-property-trunc-Prop --- ( union-subtype (in-list (cons x xsl)) (in-list xsr) y) --- ( λ --- { (is-head .x .xs) → --- ( subtype-union-left --- ( in-list (cons x xsl)) --- ( in-list xsr) --- ( y) --- ( unit-trunc-Prop (is-head y xsl))) --- ; (is-in-tail .y .x .xs y-in-xs) → --- ( subset-union-subset-left --- ( in-list xsl) --- ( in-list (cons x xsl)) --- ( in-list xsr) --- ( subset-in-tail x xsl) --- ( y) --- ( sub-lists y (unit-trunc-Prop y-in-xs)))}))) --- ( backward-implication --- ( equiv-subset-head-tail x xsl a) --- ( x-in-a , sub-xsl)) --- ( sub-xsr))))) --- ( λ x-in-b → --- ( unit-trunc-Prop --- ( triple xsl (cons x xsr) --- ( triple --- ( λ y → --- ( map-universal-property-trunc-Prop --- ( union-subtype (in-list xsl) (in-list (cons x xsr)) y) --- ( λ --- { (is-head .x .xs) → --- ( subtype-union-right --- ( in-list xsl) --- ( in-list (cons x xsr)) --- ( y) --- ( unit-trunc-Prop (is-head y xsr))) --- ; (is-in-tail .y .x .xs y-in-xs) → --- ( subset-union-subset-right --- ( in-list xsl) --- ( in-list xsr) --- ( in-list (cons x xsr)) --- ( subset-in-tail x xsr) --- ( y) --- ( sub-lists y (unit-trunc-Prop y-in-xs)))}))) --- ( sub-xsl) --- ( backward-implication --- ( equiv-subset-head-tail x xsr b) --- ( x-in-b , sub-xsr)))))))) - --- module _ --- {l1 : Level} {i : Set l1} --- where - --- list-assumptions-weak-deduction : --- {l2 : Level} {logic : modal-theory l2 i} {a : formula i} → --- logic ▷ a → --- Σ (list (formula i)) (λ l → (in-list l ⊆ logic) × (in-list l ▷ a)) --- pr1 (list-assumptions-weak-deduction {a = a} (w-ax x)) = --- unit-list a --- pr1 (pr2 (list-assumptions-weak-deduction {logic = logic} (w-ax x))) b = --- map-universal-property-trunc-Prop ( logic b) ( λ { (is-head _ _) → x}) --- pr2 (pr2 (list-assumptions-weak-deduction {a = a} (w-ax x))) = --- w-ax (unit-trunc-Prop (is-head a nil)) --- pr1 (list-assumptions-weak-deduction (w-mp wdcb wdc)) = --- concat-list --- ( pr1 (list-assumptions-weak-deduction wdcb)) --- ( pr1 (list-assumptions-weak-deduction wdc)) --- pr1 (pr2 (list-assumptions-weak-deduction {logic = logic} (w-mp wdcb wdc))) --- b b-in-list = --- subtype-union-both --- ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) --- ( in-list (pr1 (list-assumptions-weak-deduction wdc))) --- ( logic) --- ( pr1 (pr2 (list-assumptions-weak-deduction wdcb))) --- ( pr1 (pr2 (list-assumptions-weak-deduction wdc))) --- ( b) --- ( union-lists-concat --- ( pr1 (list-assumptions-weak-deduction wdcb)) --- ( pr1 (list-assumptions-weak-deduction wdc)) --- ( b) --- ( b-in-list)) --- pr2 (pr2 (list-assumptions-weak-deduction (w-mp wdcb wdc))) = --- w-mp --- ( weak-deduction-monotic --- ( transitive-leq-subtype --- ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) --- ( union-subtype --- ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) --- ( in-list (pr1 (list-assumptions-weak-deduction wdc)))) --- ( in-list --- ( concat-list --- ( pr1 (list-assumptions-weak-deduction wdcb)) --- ( pr1 (list-assumptions-weak-deduction wdc)))) --- ( in-concat-lists-union --- ( pr1 (list-assumptions-weak-deduction wdcb)) --- ( pr1 (list-assumptions-weak-deduction wdc))) --- ( subtype-union-left --- ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) --- ( in-list (pr1 (list-assumptions-weak-deduction wdc))))) --- ( pr2 (pr2 (list-assumptions-weak-deduction wdcb)))) --- ( weak-deduction-monotic --- ( transitive-leq-subtype --- ( in-list (pr1 (list-assumptions-weak-deduction wdc))) --- ( union-subtype --- ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) --- ( in-list (pr1 (list-assumptions-weak-deduction wdc)))) --- ( in-list --- ( concat-list --- ( pr1 (list-assumptions-weak-deduction wdcb)) --- ( pr1 (list-assumptions-weak-deduction wdc)))) --- ( in-concat-lists-union --- ( pr1 (list-assumptions-weak-deduction wdcb)) --- ( pr1 (list-assumptions-weak-deduction wdc))) --- ( subtype-union-right --- ( in-list (pr1 (list-assumptions-weak-deduction wdcb))) --- ( in-list (pr1 (list-assumptions-weak-deduction wdc))))) --- ( pr2 (pr2 (list-assumptions-weak-deduction wdc)))) - --- nil-no-deduction : {a : formula i} → ¬ (in-list nil ▷ a) --- nil-no-deduction (w-ax x) = --- apply-universal-property-trunc-Prop x empty-Prop (λ ()) --- nil-no-deduction (w-mp wd _) = nil-no-deduction wd - --- -- list-assumptions-weak-modal-logic : --- -- {l2 : Level} {logic : modal-theory l2 i} {a : formula i} → --- -- is-in-subtype (weak-modal-logic logic) a → --- -- ∃ ( list (formula i)) --- -- ( λ l → --- -- ( in-list l ⊆ logic) × --- -- ( is-in-subtype (weak-modal-logic (in-list l)) a)) --- -- list-assumptions-weak-modal-logic {l2} {logic} {a} = --- -- map-universal-property-trunc-Prop --- -- ( ∃-Prop --- -- ( list (formula i)) --- -- ( λ l → --- -- ( in-list l ⊆ logic) × --- -- ( is-in-subtype (weak-modal-logic (in-list l)) a)) ) --- -- ( λ wda → --- -- ( let (l , sub , comp) = list-assumptions-weak-deduction wda --- -- in unit-trunc-Prop (l , sub , (unit-trunc-Prop comp)))) - --- forward-subset-head-add : --- (a : formula i) (l : list (formula i)) → --- theory-add-formula a (in-list l) ⊆ in-list (cons a l) --- forward-subset-head-add a l = --- subtype-union-both --- ( Id-formula-Prop a) --- ( in-list l) --- ( in-list (cons a l)) --- ( λ { .a refl → unit-trunc-Prop (is-head a l)}) --- ( subset-in-tail a l) - --- backward-subset-head-add : --- (a : formula i) (l : list (formula i)) → --- in-list (cons a l) ⊆ theory-add-formula a (in-list l) --- backward-subset-head-add a l x = --- map-universal-property-trunc-Prop --- ( theory-add-formula a (in-list l) x) --- ( λ --- { (is-head .a .l) → --- ( subtype-union-left (Id-formula-Prop a) (in-list l) a refl) --- ; (is-in-tail .x .a .l t) → --- ( subtype-union-right --- ( Id-formula-Prop a) --- ( in-list l) --- ( x) --- ( unit-trunc-Prop t))}) - --- module _ --- {l1 l2 : Level} {i : Set l1} (axioms : modal-theory l2 i) --- where - --- backward-deduction-lemma : --- {a b : formula i} → --- axioms ▷ a →ₘ b → --- is-in-subtype (weak-modal-logic (theory-add-formula a axioms)) b --- backward-deduction-lemma {a} wab = --- unit-trunc-Prop --- ( w-mp --- ( weak-deduction-monotic --- ( subtype-union-right ((Id-formula-Prop a)) axioms) wab) --- ( w-ax (subtype-union-left (Id-formula-Prop a) axioms a refl))) - --- module _ --- -- TODO: maybe just in axioms? --- (contains-ax-k : ax-k i ⊆ weak-modal-logic axioms) --- (contains-ax-s : ax-s i ⊆ weak-modal-logic axioms) --- where - --- deduction-a->a : --- (a : formula i) → is-in-subtype (weak-modal-logic axioms) (a →ₘ a) --- deduction-a->a a = --- apply-three-times-universal-property-trunc-Prop --- ( contains-ax-s _ (a , a →ₘ a , a , refl)) --- ( contains-ax-k _ (a , a →ₘ a , refl)) --- ( contains-ax-k _ (a , a , refl)) --- ( (weak-modal-logic axioms) (a →ₘ a)) --- ( λ x y z → unit-trunc-Prop (w-mp (w-mp x y) z)) - --- forward-deduction-lemma : --- (a : formula i) {b : formula i} → --- theory-add-formula a axioms ▷ b → --- is-in-subtype (weak-modal-logic axioms) (a →ₘ b) --- forward-deduction-lemma a {b} (w-ax x) = --- elim-disjunction-Prop --- ( Id-formula-Prop a b) --- ( axioms b) --- ( (weak-modal-logic axioms) (a →ₘ b)) --- ( (λ { refl → deduction-a->a a}) --- , (λ in-axioms → --- ( apply-universal-property-trunc-Prop --- ( contains-ax-k _ (b , a , refl)) --- ( (weak-modal-logic axioms) (a →ₘ b)) --- ( λ x → unit-trunc-Prop (w-mp x (w-ax in-axioms)))))) --- ( x) --- forward-deduction-lemma a {b} (w-mp {c} twdcb twdc) = --- apply-three-times-universal-property-trunc-Prop --- ( forward-deduction-lemma a twdcb) --- ( forward-deduction-lemma a twdc) --- ( contains-ax-s _ (a , c , b , refl)) --- ( (weak-modal-logic axioms) (a →ₘ b)) --- ( λ wdacb wdac x → --- ( unit-trunc-Prop (w-mp (w-mp x wdacb) wdac))) - --- deduction-lemma : --- (a b : formula i) → --- weak-modal-logic (theory-add-formula a axioms) b ⇔ --- weak-modal-logic axioms (a →ₘ b) --- pr1 (deduction-lemma a b) = --- map-universal-property-trunc-Prop --- ( (weak-modal-logic axioms) (a →ₘ b)) --- ( forward-deduction-lemma a) --- pr2 (deduction-lemma a b) = --- map-universal-property-trunc-Prop --- ( weak-modal-logic (theory-add-formula a axioms) b) --- ( backward-deduction-lemma) - --- module _ --- {l1 l2 : Level} {i : Set l1} (axioms : modal-theory l2 i) --- (contains-ax-k : ax-k i ⊆ weak-modal-logic axioms) --- (contains-ax-s : ax-s i ⊆ weak-modal-logic axioms) --- (contains-ax-dn : ax-dn i ⊆ weak-modal-logic axioms) --- where - --- deduction-ex-falso : --- (a b : formula i) → --- is-in-subtype (weak-modal-logic axioms) (~ a →ₘ a →ₘ b) --- deduction-ex-falso a b = --- forward-implication --- ( deduction-lemma axioms contains-ax-k contains-ax-s (~ a) (a →ₘ b)) --- ( forward-implication --- ( deduction-lemma --- ( theory-add-formula (~ a) axioms) --- ( transitive-leq-subtype --- ( ax-k i) --- ( weak-modal-logic axioms) --- ( weak-modal-logic (theory-add-formula (a →ₘ ⊥) axioms)) --- ( weak-modal-logic-monotic --- ( subtype-union-right (Id-formula-Prop (~ a)) axioms)) --- ( contains-ax-k)) --- ( transitive-leq-subtype --- ( ax-s i) --- ( weak-modal-logic axioms) --- ( weak-modal-logic (theory-add-formula (a →ₘ ⊥) axioms)) --- ( weak-modal-logic-monotic --- ( subtype-union-right (Id-formula-Prop (~ a)) axioms)) --- ( contains-ax-s)) --- ( a) --- ( b)) --- ( apply-twice-universal-property-trunc-Prop --- ( contains-ax-k _ (⊥ , ~ b , refl)) --- ( contains-ax-dn _ ((b , refl))) --- ( weak-modal-logic --- ( theory-add-formula a (theory-add-formula (~ a) axioms)) b) --- ( λ x y → --- ( unit-trunc-Prop --- ( w-mp {a = ~~ b} --- ( weak-deduction-monotic --- ( transitive-leq-subtype --- ( axioms) --- ( theory-add-formula (~ a) axioms) --- ( theory-add-formula a --- ( theory-add-formula (~ a) axioms)) --- ( subtype-union-right --- ( Id-formula-Prop a) --- ( theory-add-formula (~ a) axioms)) --- ( subtype-union-right (Id-formula-Prop (~ a)) axioms)) --- ( y)) --- ( w-mp {a = ⊥} --- ( weak-deduction-monotic --- ( transitive-leq-subtype --- ( axioms) --- ( theory-add-formula (~ a) axioms) --- ( theory-add-formula a --- ( theory-add-formula (~ a) axioms)) --- ( subtype-union-right --- ( Id-formula-Prop a) --- ( theory-add-formula (~ a) axioms)) --- ( subtype-union-right (Id-formula-Prop (~ a)) axioms)) --- ( x)) --- ( w-mp {a = a} --- ( w-ax --- ( unit-trunc-Prop (inr (unit-trunc-Prop (inl refl))))) --- ( w-ax (unit-trunc-Prop (inl refl)))))))))) - --- logic-ex-falso : --- (a b : formula i) → --- is-in-subtype (weak-modal-logic axioms) a → --- is-in-subtype (weak-modal-logic axioms) (~ a) → --- is-in-subtype (weak-modal-logic axioms) b --- logic-ex-falso a b a-in-logic not-a-in-logic = --- apply-three-times-universal-property-trunc-Prop --- ( a-in-logic) --- ( not-a-in-logic) --- ( deduction-ex-falso a b) --- ( (weak-modal-logic axioms) b) --- ( λ x y z → unit-trunc-Prop (w-mp (w-mp z y) x)) -``` -```` From 8d29db398ec607849bec7d35d6bf2075bbc93ff4 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Fri, 3 May 2024 20:32:48 +0300 Subject: [PATCH 42/63] Refactor canonical kripke model --- src/foundation/subtypes.lagda.md | 3 +- src/modal-logic/canonical-model.lagda.md | 72 +++++++++++++++++++ ....lagda.md => l-complete-theories.lagda.md} | 28 +++++++- .../l-consistent-theories.lagda.md | 13 ---- 4 files changed, 99 insertions(+), 17 deletions(-) create mode 100644 src/modal-logic/canonical-model.lagda.md rename src/modal-logic/{L-complete-theories.lagda.md => l-complete-theories.lagda.md} (95%) diff --git a/src/foundation/subtypes.lagda.md b/src/foundation/subtypes.lagda.md index b7dc626d1b..5405df24ea 100644 --- a/src/foundation/subtypes.lagda.md +++ b/src/foundation/subtypes.lagda.md @@ -157,7 +157,6 @@ module _ {l1 : Level} {A : UU l1} where - --- TODO: replace with equiv-fam equiv-subtypes : {l2 l3 : Level} (P : subtype l2 A) (Q : subtype l3 A) → UU (l1 ⊔ l2 ⊔ l3) equiv-subtypes P Q = (x : A) → is-in-subtype P x ≃ is-in-subtype Q x @@ -204,7 +203,7 @@ pr1 (subtype-Set l2 A) = subtype l2 A pr2 (subtype-Set l2 A) = is-set-subtype ``` -### TODO +### TODO: raise subtype ```agda raise-subtype : diff --git a/src/modal-logic/canonical-model.lagda.md b/src/modal-logic/canonical-model.lagda.md new file mode 100644 index 0000000000..bad06098c2 --- /dev/null +++ b/src/modal-logic/canonical-model.lagda.md @@ -0,0 +1,72 @@ +# Canonical model theorem + +```agda +module modal-logic.canonical-model where +``` + +
Imports + +```agda +open import foundation.dependent-pair-types +open import foundation.propositional-resizing +open import foundation.propositional-truncations +open import foundation.sets +open import foundation.universe-levels + +open import foundation-core.propositions +open import foundation-core.subtypes + +open import modal-logic.axioms +open import modal-logic.formulas +open import modal-logic.kripke-semantics +open import modal-logic.logic-syntax +open import modal-logic.l-complete-theories +open import modal-logic.l-consistent-theories +open import modal-logic.weak-deduction + +open import order-theory.zorn +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l1 l2 : Level} {i : Set l1} + (logic : modal-theory l2 i) + (is-logic : is-weak-modal-logic logic) + (is-cons : is-consistent-modal-logic logic) + (contains-ax-k : ax-k i ⊆ logic) + (contains-ax-s : ax-s i ⊆ logic) + (zorn : Zorn-non-empty (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) l2) + (prop-resize : propositional-resizing (l1 ⊔ l2) (lsuc (l1 ⊔ l2))) + where + + canonical-kripke-model : + kripke-model (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) i (l1 ⊔ l2) + pr1 (pr1 (pr1 canonical-kripke-model)) = + L-complete-theory logic (l1 ⊔ l2) + pr2 (pr1 (pr1 canonical-kripke-model)) = + is-inhabited-L-complete-theory + ( logic) + ( prop-resize) + ( zorn) + ( is-logic) + ( is-cons) + ( contains-ax-k) + ( contains-ax-s) + pr2 (pr1 canonical-kripke-model) x y = + Π-Prop + ( formula i) + ( λ a → + ( hom-Prop + ( modal-theory-L-complete-theory logic x (□ a)) + ( modal-theory-L-complete-theory logic y a))) + pr2 canonical-kripke-model n x = + modal-theory-L-complete-theory logic x (var n) +``` diff --git a/src/modal-logic/L-complete-theories.lagda.md b/src/modal-logic/l-complete-theories.lagda.md similarity index 95% rename from src/modal-logic/L-complete-theories.lagda.md rename to src/modal-logic/l-complete-theories.lagda.md index cad804e3bb..5dfa78384a 100644 --- a/src/modal-logic/L-complete-theories.lagda.md +++ b/src/modal-logic/l-complete-theories.lagda.md @@ -604,14 +604,38 @@ module _ is-inhabited (L-consistent-theory logic l3) → is-inhabited (L-complete-theory l3) extend-L-consistent-theory zorn is-inh = - ( is-inhabited-L-complete-exists-complete-L-consistent-theory + is-inhabited-L-complete-exists-complete-L-consistent-theory ( zorn ( L-consistent-theories-Poset logic l3) ( is-inh) ( λ C C-is-inh → ( intro-∃ ( resized-chain-union-L-consistent-theory C C-is-inh) - ( union-is-chain-upper-bound C C-is-inh))))) + ( union-is-chain-upper-bound C C-is-inh)))) + + module _ + {l3 : Level} + (prop-resize : propositional-resizing (l1 ⊔ l2) (lsuc l1 ⊔ lsuc l2 ⊔ l3)) + (zorn : Zorn-non-empty (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) l3) + (is-logic : is-weak-modal-logic logic) + (is-cons : is-consistent-modal-logic logic) + (contains-ax-k : ax-k i ⊆ logic) + (contains-ax-s : ax-s i ⊆ logic) + where + + is-inhabited-L-complete-theory : + is-inhabited (L-complete-theory (l1 ⊔ l2)) + is-inhabited-L-complete-theory = + extend-L-consistent-theory prop-resize contains-ax-k contains-ax-s zorn + ( map-is-inhabited + ( λ (t , t-is-cons) → + ( pair + ( raise-subtype l1 t) + ( is-L-consistent-antimonotic logic (raise-subtype l1 t) t + ( inv-subset-equiv-subtypes t (raise-subtype l1 t) + ( compute-raise-subtype l1 t)) + ( t-is-cons)))) + ( is-inhabited-L-consistent-theory logic is-logic is-cons)) module _ {l1 : Level} {i : Set l1} diff --git a/src/modal-logic/l-consistent-theories.lagda.md b/src/modal-logic/l-consistent-theories.lagda.md index e648feaafd..3d07b7beaf 100644 --- a/src/modal-logic/l-consistent-theories.lagda.md +++ b/src/modal-logic/l-consistent-theories.lagda.md @@ -87,18 +87,6 @@ module _ ( weak-modal-logic-closure-monotic ( subtype-union-right logic (modal-theory-L-consistent-theory theory))) ( is-L-consistent-theory-modal-theory-L-consistent-theory theory) - -- is-consistent-modal-logic-antimonotic - -- ( modal-theory-L-consistent-theory theory) - -- ( weak-modal-logic-closure - -- ( logic ∪ modal-theory-L-consistent-theory theory)) - -- ( transitive-leq-subtype - -- ( modal-theory-L-consistent-theory theory) - -- ( logic ∪ modal-theory-L-consistent-theory theory) - -- ( weak-modal-logic-closure - -- ( logic ∪ modal-theory-L-consistent-theory theory)) - -- ( subset-axioms-weak-modal-logic) - -- ( subtype-union-right logic (modal-theory-L-consistent-theory theory))) - -- ( is-L-consistent-theory-modal-theory-L-consistent-theory theory) is-L-consistent-antimonotic : {l3 l4 : Level} @@ -116,7 +104,6 @@ module _ L-consistent-theory-leq-Prop : {l3 : Level} → Relation-Prop (l1 ⊔ l3) (L-consistent-theory l3) - -- {l3 : Level} → L-consistent-theory l3 → L-consistent-theory → Prop (l1 ⊔ l2) L-consistent-theory-leq-Prop x y = leq-prop-subtype ( modal-theory-L-consistent-theory x) From 02b75b66e5cf74c686fe1a1785dc9434ac066aac Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Sun, 5 May 2024 05:16:43 +0300 Subject: [PATCH 43/63] Refactor canonical model theorem --- src/foundation/decidable-types.lagda.md | 12 + .../law-of-excluded-middle.lagda.md | 13 + src/foundation/unions-subtypes.lagda.md | 2 +- src/lists.lagda.md | 1 + src/lists/lists-subtypes.lagda.md | 12 +- src/modal-logic.lagda.md | 7 +- src/modal-logic/axioms.lagda.md | 25 +- .../canonical-model-theorem.lagda.md | 1663 +++++------------ src/modal-logic/canonical-model.lagda.md | 72 - src/modal-logic/canonical-theories.lagda.md | 185 +- src/modal-logic/completeness-K.lagda.md | 27 +- src/modal-logic/completeness-S5.lagda.md | 43 +- src/modal-logic/completeness.lagda.md | 30 - .../finite-approximability.lagda.md | 4 +- src/modal-logic/formulas-deduction.lagda.md | 278 +-- src/modal-logic/l-complete-theories.lagda.md | 76 +- .../l-consistent-theories.lagda.md | 14 +- src/modal-logic/lindenbaum.lagda.md | 10 +- src/modal-logic/logic-syntax.lagda.md | 155 +- src/modal-logic/modal-logic-K.lagda.md | 19 +- src/modal-logic/modal-logic-S5.lagda.md | 17 +- src/modal-logic/soundness.lagda.md | 29 +- src/modal-logic/weak-deduction.lagda.md | 208 ++- src/order-theory.lagda.md | 1 + src/order-theory/zorn.lagda.md | 4 +- 25 files changed, 1192 insertions(+), 1715 deletions(-) delete mode 100644 src/modal-logic/canonical-model.lagda.md diff --git a/src/foundation/decidable-types.lagda.md b/src/foundation/decidable-types.lagda.md index 3bed2fac28..9ef913eedd 100644 --- a/src/foundation/decidable-types.lagda.md +++ b/src/foundation/decidable-types.lagda.md @@ -270,3 +270,15 @@ module _ is-decidable-raise (inl p) = inl (map-raise p) is-decidable-raise (inr np) = inr (λ p' → np (map-inv-raise p')) ``` + +### For decidable type `Q` `¬ Q → ¬ P` implies `P → Q` + +```agda +contraposition-is-decidable : + {l1 l2 : Level} {P : UU l1} {Q : UU l2} → + is-decidable Q → + (¬ Q → ¬ P) → + P → Q +contraposition-is-decidable {Q = Q} (inl q) f p = q +contraposition-is-decidable {Q = Q} (inr nq) f p = ex-falso (f nq p) +``` diff --git a/src/foundation/law-of-excluded-middle.lagda.md b/src/foundation/law-of-excluded-middle.lagda.md index 1a8a6ed07e..b0b59d0a22 100644 --- a/src/foundation/law-of-excluded-middle.lagda.md +++ b/src/foundation/law-of-excluded-middle.lagda.md @@ -17,6 +17,7 @@ open import foundation.small-types open import foundation.subtypes open import foundation.universe-levels +open import foundation-core.coproduct-types open import foundation-core.equality-dependent-pair-types open import foundation-core.equivalences open import foundation-core.identity-types @@ -50,6 +51,18 @@ module _ ## Properties +-- implies contraposition + +### The law of excluded middle implies the contraposition of a proposition + +```agda +contraposition : + {l1 l2 : Level} → LEM l2 → {P : UU l1} (Q : Prop l2) → + (¬ (type-Prop Q) → ¬ P) → + P → type-Prop Q +contraposition lem Q = contraposition-is-decidable (lem Q) +``` + ### TODO ```agda diff --git a/src/foundation/unions-subtypes.lagda.md b/src/foundation/unions-subtypes.lagda.md index d24bceb5e0..058326ea9a 100644 --- a/src/foundation/unions-subtypes.lagda.md +++ b/src/foundation/unions-subtypes.lagda.md @@ -15,8 +15,8 @@ open import foundation.identity-types open import foundation.large-locale-of-subtypes open import foundation.logical-equivalences open import foundation.powersets -open import foundation.propositions open import foundation.propositional-truncations +open import foundation.propositions open import foundation.subtypes open import foundation.universe-levels diff --git a/src/lists.lagda.md b/src/lists.lagda.md index 18b0ea48f3..380000e74e 100644 --- a/src/lists.lagda.md +++ b/src/lists.lagda.md @@ -11,6 +11,7 @@ open import lists.flattening-lists public open import lists.functoriality-lists public open import lists.lists public open import lists.lists-discrete-types public +open import lists.lists-subtypes public open import lists.permutation-lists public open import lists.permutation-vectors public open import lists.predicates-on-lists public diff --git a/src/lists/lists-subtypes.lagda.md b/src/lists/lists-subtypes.lagda.md index b3f6d88463..e2ef1b4afc 100644 --- a/src/lists/lists-subtypes.lagda.md +++ b/src/lists/lists-subtypes.lagda.md @@ -1,4 +1,4 @@ -# Lists +# Lists subtypes ```agda module lists.lists-subtypes where @@ -9,22 +9,22 @@ module lists.lists-subtypes where ```agda open import foundation.dependent-pair-types open import foundation.disjunction -open import foundation-core.empty-types open import foundation.existential-quantification -open import foundation.logical-equivalences open import foundation.intersections-subtypes -open import foundation.unions-subtypes +open import foundation.logical-equivalences open import foundation.propositional-truncations +open import foundation.unions-subtypes open import foundation.universe-levels open import foundation-core.cartesian-product-types +open import foundation-core.empty-types open import foundation-core.function-types open import foundation-core.negation open import foundation-core.propositions open import foundation-core.subtypes -open import lists.lists open import lists.concatenation-lists +open import lists.lists open import lists.reversing-lists ``` @@ -76,7 +76,7 @@ module _ map-universal-property-trunc-Prop ( S x) ( λ { (is-head .x .l) → a-in-S - ; (is-in-tail .x .a .l t) → l-sub-S x (in-list-subtype-in-list t) }) + ; (is-in-tail .x .a .l t) → l-sub-S x (in-list-subtype-in-list t)}) subset-list-subtype-reverse-list : (l : list A) → list-subtype l ⊆ list-subtype (reverse-list l) diff --git a/src/modal-logic.lagda.md b/src/modal-logic.lagda.md index e7c40e8ee0..fe11bc1e13 100644 --- a/src/modal-logic.lagda.md +++ b/src/modal-logic.lagda.md @@ -7,17 +7,16 @@ open import modal-logic.axioms public open import modal-logic.canonical-model-theorem public open import modal-logic.canonical-theories public open import modal-logic.completeness public -open import modal-logic.completeness-K public -open import modal-logic.completeness-S5 public open import modal-logic.finite-approximability public open import modal-logic.formulas public open import modal-logic.formulas-deduction public open import modal-logic.kripke-models-filtrations public open import modal-logic.kripke-models-filtrations-theorem public open import modal-logic.kripke-semantics public +open import modal-logic.l-complete-theories public +open import modal-logic.l-consistent-theories public +open import modal-logic.lindenbaum public open import modal-logic.logic-syntax public -open import modal-logic.modal-logic-K public -open import modal-logic.modal-logic-S5 public open import modal-logic.modal-logic-decision public open import modal-logic.soundness public open import modal-logic.weak-deduction public diff --git a/src/modal-logic/axioms.lagda.md b/src/modal-logic/axioms.lagda.md index 015787c87f..58cac7d685 100644 --- a/src/modal-logic/axioms.lagda.md +++ b/src/modal-logic/axioms.lagda.md @@ -51,7 +51,8 @@ module _ {l : Level} (i : Set l) where - ax-1-parameter : (h : formula i → formula i) → is-injective h → formulas l i + ax-1-parameter : + (h : formula i → formula i) → is-injective h → modal-theory l i pr1 (ax-1-parameter h inj f) = Σ (formula i) (λ a → f = h a) pr2 (ax-1-parameter h inj f) (a , refl) = is-prop-is-contr @@ -64,7 +65,7 @@ module _ (h : formula i → formula i → formula i) → ({x x' y y' : formula i} → h x y = h x' y' → x = x') → ({x x' y y' : formula i} → h x y = h x' y' → y = y') → - formulas l i + modal-theory l i pr1 (ax-2-parameters h inj-1 inj-2 f) = Σ (formula i) (λ a → Σ (formula i) (λ b → f = h a b)) pr2 (ax-2-parameters h inj-1 inj-2 f) (a , b , refl) = @@ -79,7 +80,7 @@ module _ ({x x' y y' z z' : formula i} → h x y z = h x' y' z' → x = x') → ({x x' y y' z z' : formula i} → h x y z = h x' y' z' → y = y') → ({x x' y y' z z' : formula i} → h x y z = h x' y' z' → z = z') → - formulas l i + modal-theory l i pr1 (ax-3-parameters h inj-1 inj-2 inj-3 f) = Σ ( formula i) ( λ a → Σ (formula i) (λ b → Σ (formula i) ( λ c → f = h a b c))) @@ -90,14 +91,14 @@ module _ ( λ x (y , z , e) → inj-1 e)) ( a , b , c , refl) - ax-k : formulas l i + ax-k : modal-theory l i ax-k = ax-2-parameters ( λ a b → a →ₘ b →ₘ a) ( eq-implication-left) ( eq-implication-left ∘ eq-implication-right) - ax-s : formulas l i + ax-s : modal-theory l i ax-s = ax-3-parameters ( λ a b c → (a →ₘ b →ₘ c) →ₘ (a →ₘ b) →ₘ a →ₘ c) @@ -105,29 +106,29 @@ module _ ( eq-implication-left ∘ eq-implication-right ∘ eq-implication-left) ( eq-implication-right ∘ eq-implication-right ∘ eq-implication-left) - ax-n : formulas l i + ax-n : modal-theory l i ax-n = ax-2-parameters ( λ a b → □ (a →ₘ b) →ₘ □ a →ₘ □ b) ( eq-implication-left ∘ eq-box ∘ eq-implication-left) ( eq-implication-right ∘ eq-box ∘ eq-implication-left) - ax-dn : formulas l i + ax-dn : modal-theory l i ax-dn = ax-1-parameter (λ a → ~~ a →ₘ a) eq-implication-right - ax-m : formulas l i + ax-m : modal-theory l i ax-m = ax-1-parameter (λ a → □ a →ₘ a) eq-implication-right - ax-b : formulas l i + ax-b : modal-theory l i ax-b = ax-1-parameter (λ a → a →ₘ □ ◇ a) eq-implication-left - ax-d : formulas l i + ax-d : modal-theory l i ax-d = ax-1-parameter (λ a → □ a →ₘ ◇ a) (eq-box ∘ eq-implication-left) - ax-4 : formulas l i + ax-4 : modal-theory l i ax-4 = ax-1-parameter (λ a → □ a →ₘ □ □ a) (eq-box ∘ eq-implication-left) - ax-5 : formulas l i + ax-5 : modal-theory l i ax-5 = ax-1-parameter ( λ a → ◇ a →ₘ □ ◇ a) ( eq-diamond ∘ eq-implication-left) diff --git a/src/modal-logic/canonical-model-theorem.lagda.md b/src/modal-logic/canonical-model-theorem.lagda.md index f362423a75..b27845e7f8 100644 --- a/src/modal-logic/canonical-model-theorem.lagda.md +++ b/src/modal-logic/canonical-model-theorem.lagda.md @@ -8,52 +8,41 @@ module modal-logic.canonical-model-theorem where ```agda open import foundation.action-on-identifications-functions -open import foundation.cartesian-product-types -open import foundation.coproduct-types open import foundation.dependent-pair-types -open import foundation.disjunction open import foundation.empty-types -open import foundation.existential-quantification open import foundation.function-types open import foundation.identity-types -open import foundation.inhabited-types -open import foundation.intersections-subtypes open import foundation.law-of-excluded-middle open import foundation.logical-equivalences open import foundation.negation open import foundation.propositional-resizing open import foundation.propositional-truncations -open import foundation.propositions open import foundation.raising-universe-levels open import foundation.sets open import foundation.subtypes open import foundation.transport-along-identifications open import foundation.unions-subtypes -open import foundation.unit-type open import foundation.universe-levels -open import foundation-core.equivalences +open import foundation-core.coproduct-types +open import foundation-core.propositions open import lists.lists +open import lists.lists-subtypes open import lists.reversing-lists open import modal-logic.axioms open import modal-logic.completeness open import modal-logic.formulas -open import modal-logic.formulas-deduction open import modal-logic.kripke-semantics +open import modal-logic.l-complete-theories +open import modal-logic.l-consistent-theories +open import modal-logic.lindenbaum open import modal-logic.logic-syntax -open import modal-logic.modal-logic-K -open import modal-logic.soundness +open import modal-logic.modal-logic-k open import modal-logic.weak-deduction -open import order-theory.chains-posets -open import order-theory.maximal-elements-posets -open import order-theory.posets -open import order-theory.preorders -open import order-theory.subposets -open import order-theory.subtypes-leq-posets -open import order-theory.top-elements-posets +open import order-theory.zorn ```
@@ -66,1207 +55,509 @@ TODO ```agda module _ - {l1 l2 : Level} {i : Set l1} (axioms : modal-theory l2 i) - (zorn : Zorn (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) l2) + {l1 l2 : Level} {i : Set l1} + (logic : modal-theory l2 i) + (is-logic : is-modal-logic logic) + (is-cons : is-consistent-modal-logic logic) + (is-normal : is-normal-modal-logic logic) + (lem : LEM (l1 ⊔ l2)) + (zorn : Zorn-non-empty (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) l2) (prop-resize : propositional-resizing (l1 ⊔ l2) (lsuc (l1 ⊔ l2))) - (L-is-cons : is-consistent-modal-logic (modal-logic axioms)) - (contains-K : modal-logic-K i ⊆ modal-logic axioms) where private - logic : modal-theory (l1 ⊔ l2) i - logic = modal-logic axioms + is-weak-logic : is-weak-modal-logic logic + is-weak-logic = is-weak-modal-logic-is-modal-logic is-logic contains-ax-k : ax-k i ⊆ logic contains-ax-k = - transitive-leq-subtype - ( ax-k i) - ( modal-logic-K i) - ( logic) - ( contains-K) + transitive-leq-subtype (ax-k i) (modal-logic-K i) logic + ( is-normal) ( K-contains-ax-k i) contains-ax-s : ax-s i ⊆ logic contains-ax-s = - transitive-leq-subtype - ( ax-s i) - ( modal-logic-K i) - ( logic) - ( contains-K) + transitive-leq-subtype (ax-s i) (modal-logic-K i) logic + ( is-normal) ( K-contains-ax-s i) contains-ax-n : ax-n i ⊆ logic contains-ax-n = - transitive-leq-subtype - ( ax-n i) - ( modal-logic-K i) - ( logic) - ( contains-K) + transitive-leq-subtype (ax-n i) (modal-logic-K i) logic + ( is-normal) ( K-contains-ax-n i) contains-ax-dn : ax-dn i ⊆ logic contains-ax-dn = - transitive-leq-subtype - ( ax-dn i) - ( modal-logic-K i) - ( logic) - ( contains-K) + transitive-leq-subtype (ax-dn i) (modal-logic-K i) logic + ( is-normal) ( K-contains-ax-dn i) - is-L-consistent-theory-Prop : modal-theory (l1 ⊔ l2) i → Prop (l1 ⊔ l2) - is-L-consistent-theory-Prop t = - is-consistent-modal-logic-Prop (weak-modal-logic (logic ∪ t)) - - is-L-consistent-theory : modal-theory (l1 ⊔ l2) i → UU (l1 ⊔ l2) - is-L-consistent-theory = type-Prop ∘ is-L-consistent-theory-Prop - - L-consistent-theory : UU (lsuc l1 ⊔ lsuc l2) - L-consistent-theory = type-subtype is-L-consistent-theory-Prop - - L-consistent-theory-leq-Prop : - L-consistent-theory → L-consistent-theory → Prop (l1 ⊔ l2) - L-consistent-theory-leq-Prop x y = leq-prop-subtype (pr1 x) (pr1 y) - - L-consistent-theory-leq : - L-consistent-theory → L-consistent-theory → UU (l1 ⊔ l2) - L-consistent-theory-leq x y = type-Prop (L-consistent-theory-leq-Prop x y) - - theories-Poset : Poset (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) - theories-Poset = subtypes-leq-Poset (l1 ⊔ l2) (formula i) - - L-consistent-theories-Poset : Poset (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) - L-consistent-theories-Poset = - poset-Subposet theories-Poset is-L-consistent-theory-Prop - - is-L-complete-theory-Prop : L-consistent-theory → Prop (lsuc l1 ⊔ lsuc l2) - is-L-complete-theory-Prop = - is-maximal-element-Poset-Prop L-consistent-theories-Poset - - is-L-complete-theory : L-consistent-theory → UU (lsuc l1 ⊔ lsuc l2) - is-L-complete-theory = type-Prop ∘ is-L-complete-theory-Prop - - L-complete-theory : UU (lsuc l1 ⊔ lsuc l2) - L-complete-theory = type-subtype is-L-complete-theory-Prop - - weak-modal-logic-subset-complete-theory : - (x : L-consistent-theory) → is-L-complete-theory x → - weak-modal-logic (pr1 x) ⊆ pr1 x - weak-modal-logic-subset-complete-theory x is-comp a in-logic = - tr - ( λ y → is-in-subtype (pr1 y) a) - ( is-comp - ( pair - ( theory-add-formula a (pr1 x)) - ( λ bot-in-logic → - ( pr2 x - ( subset-weak-modal-logic-if-subset-axioms - ( transitive-leq-subtype - ( union-subtype logic (theory-add-formula a (pr1 x))) - ( theory-add-formula a (union-subtype logic (pr1 x))) - ( weak-modal-logic (union-subtype logic (pr1 x))) - ( subtype-union-both - ( Id-formula-Prop a) - ( union-subtype logic (pr1 x)) - ( weak-modal-logic (union-subtype logic (pr1 x))) - ( λ { .a refl → - ( weak-modal-logic-monotic - ( subtype-union-right logic (pr1 x)) - ( a) - ( in-logic))}) - ( axioms-subset-weak-modal-logic - ( union-subtype logic (pr1 x)))) - ( union-swap-1-2 logic (Id-formula-Prop a) (pr1 x))) - ( ⊥) - ( bot-in-logic))))) - ( subtype-union-right (Id-formula-Prop a) (pr1 x))) - ( unit-trunc-Prop (inl refl)) - - logic-subset-L-complete-theory : - (x : L-consistent-theory) → is-L-complete-theory x → logic ⊆ pr1 x - logic-subset-L-complete-theory x is-comp a in-logic = - tr - ( λ y → is-in-subtype (pr1 y) a) - ( is-comp - ( pair - ( theory-add-formula a (pr1 x)) - ( λ bot-in-logic → - ( pr2 x - ( weak-modal-logic-monotic - ( transitive-leq-subtype - ( union-subtype logic (theory-add-formula a (pr1 x))) - ( union-subtype - ( union-subtype logic (Id-formula-Prop a)) (pr1 x)) - ( union-subtype logic (pr1 x)) - ( subset-union-subset-left - ( union-subtype logic (Id-formula-Prop a)) - ( logic) - ( pr1 x) - ( subtype-union-both - ( logic) - ( Id-formula-Prop a) - ( logic) - ( refl-leq-subtype logic) - ( λ { .a refl → in-logic}))) - ( forward-subset-union-assoc - ( logic) - ( Id-formula-Prop a) - ( pr1 x))) - ( ⊥) - ( bot-in-logic))))) - ( subtype-union-right (Id-formula-Prop a) (pr1 x))) - ( subtype-union-left (Id-formula-Prop a) (pr1 x) a refl) - - disjuctivity-L-complete-theory : - LEM (l1 ⊔ l2) → - (x : L-consistent-theory) → is-L-complete-theory x → - (a : formula i) → pr1 x a ∨ pr1 x (~ a) - disjuctivity-L-complete-theory lem x is-comp a with lem (pr1 x a) - ... | inl a-in-logic = - inl-disjunction-Prop ((pr1 x) a) ((pr1 x) (~ a)) a-in-logic - ... | inr a-not-in-logic = - unit-trunc-Prop - ( inr - ( tr - ( λ y → is-in-subtype (pr1 y) (~ a)) - ( is-comp - ( theory-add-formula (~ a) (pr1 x) - , λ bot-in-logic → - ( a-not-in-logic - ( weak-modal-logic-subset-complete-theory - ( x) - ( is-comp) - ( a) - ( apply-universal-property-trunc-Prop - ( forward-implication - ( deduction-lemma - ( pr1 x) - ( transitive-leq-subtype - ( ax-k i) - ( logic) - ( weak-modal-logic (pr1 x)) - ( transitive-leq-subtype - ( logic) - ( pr1 x) - ( weak-modal-logic (pr1 x)) - ( axioms-subset-weak-modal-logic (pr1 x)) - ( logic-subset-L-complete-theory x is-comp)) - ( contains-ax-k)) - ( transitive-leq-subtype - ( ax-s i) - ( logic) - ( weak-modal-logic (pr1 x)) - ( transitive-leq-subtype - ( logic) - ( pr1 x) - ( weak-modal-logic (pr1 x)) - ( axioms-subset-weak-modal-logic (pr1 x)) - ( logic-subset-L-complete-theory x is-comp)) - ( contains-ax-s)) - ( ~ a) - ( ⊥)) - ( weak-modal-logic-monotic - { ax₁ = - ( union-subtype logic - ( theory-add-formula (~ a) (pr1 x)))} - { ax₂ = theory-add-formula (~ a) (pr1 x)} - ( subtype-union-both - ( logic) - ( theory-add-formula (~ a) (pr1 x)) - ( theory-add-formula (~ a) (pr1 x)) - ( transitive-leq-subtype - ( logic) - ( pr1 x) - ( theory-add-formula (~ a) (pr1 x)) - ( subtype-union-right - ( Id-formula-Prop (~ a)) - ( pr1 x)) - ( logic-subset-L-complete-theory x is-comp)) - ( refl-leq-subtype - (theory-add-formula (~ a) (pr1 x)))) - ( ⊥) - ( bot-in-logic))) - ( weak-modal-logic (pr1 x) a) - ( λ wd-bot → - ( unit-trunc-Prop - ( w-mp - ( w-ax - ( logic-subset-L-complete-theory - ( x) - ( is-comp) - ( ~~ a →ₘ a) - ( contains-ax-dn _ (a , refl)))) - ( wd-bot)))))))) - ( subtype-union-right (Id-formula-Prop (~ a)) (pr1 x))) - ( subtype-union-left (Id-formula-Prop (~ a)) (pr1 x) (~ a) refl))) - - -- postulate - -- helper : (a : formula i) → ◇ (~ a) → ~ □ a - - -- TODO: rename - lemma-box-diamond : - LEM (l1 ⊔ l2) → - (x : L-consistent-theory) → is-L-complete-theory x → - (y : L-consistent-theory) → is-L-complete-theory y → - (λ a → - ( exists-Prop - ( formula i) - ( λ b → product-Prop (Id-formula-Prop a (◇ b)) ( pr1 y b)))) ⊆ pr1 x → - (λ a → pr1 x (□ a)) ⊆ pr1 y - lemma-box-diamond lem x x-is-comp y y-is-comp sub a box-a-in-x = - apply-universal-property-trunc-Prop - ( disjuctivity-L-complete-theory lem y y-is-comp a) - ( pr1 y a) - ( λ - { (inl a-in-y) → a-in-y - ; (inr not-a-in-y) → - ( ex-falso - ( pr2 x - ( weak-modal-logic-mp - ( weak-modal-logic-diamond-negate - ( i) - ( modal-logic axioms ∪ pr1 x) - ( transitive-leq-subtype - ( modal-logic-K i) - ( modal-logic axioms ∪ pr1 x) - ( weak-modal-logic (modal-logic axioms ∪ pr1 x)) - ( axioms-subset-weak-modal-logic - ( modal-logic axioms ∪ pr1 x)) - ( transitive-leq-subtype - ( modal-logic-K i) - ( modal-logic axioms) - ( modal-logic axioms ∪ pr1 x) - ( subtype-union-left (modal-logic axioms) (pr1 x)) - ( contains-K))) - ( axioms-subset-weak-modal-logic - ( modal-logic axioms ∪ pr1 x) - ( ◇ ~ a) - ( subtype-union-right - ( modal-logic axioms) - ( pr1 x) - ( ◇ ~ a) - ( sub (◇ ~ a) (intro-∃ (~ a) (refl , not-a-in-y)))))) - ( axioms-subset-weak-modal-logic - ( modal-logic axioms ∪ pr1 x) - ( □ a) - ( subtype-union-right - ( modal-logic axioms) - ( pr1 x) - ( □ a) - ( box-a-in-x))))))}) - - complete-theory-implication : - LEM (l1 ⊔ l2) → - (x : L-consistent-theory) → is-L-complete-theory x → - {a b : formula i} → (is-in-subtype (pr1 x) a → is-in-subtype (pr1 x) b) → - is-in-subtype (pr1 x) (a →ₘ b) - complete-theory-implication lem x is-comp {a} {b} imp = - apply-universal-property-trunc-Prop - ( disjuctivity-L-complete-theory lem x is-comp a) - ( pr1 x (a →ₘ b)) - ( λ { (inl a-in-logic) → - ( weak-modal-logic-subset-complete-theory - ( x) - ( is-comp) - ( a →ₘ b) - ( forward-implication - ( deduction-lemma - ( pr1 x) - ( transitive-leq-subtype - ( ax-k i) - ( logic) - ( weak-modal-logic (pr1 x)) - ( transitive-leq-subtype - ( logic) - ( pr1 x) - ( weak-modal-logic (pr1 x)) - ( axioms-subset-weak-modal-logic (pr1 x)) - ( logic-subset-L-complete-theory x is-comp)) - ( contains-ax-k)) - ( transitive-leq-subtype - ( ax-s i) - ( logic) - ( weak-modal-logic (pr1 x)) - ( transitive-leq-subtype - ( logic) - ( pr1 x) - ( weak-modal-logic (pr1 x)) - ( axioms-subset-weak-modal-logic (pr1 x)) - ( logic-subset-L-complete-theory x is-comp)) - ( contains-ax-s)) - ( a) - ( b)) - ( weak-modal-logic-monotic - { ax₁ = pr1 x} -- TODO: make explicit - { ax₂ = union-subtype (Id-formula-Prop a) (pr1 x)} - ( subtype-union-right (Id-formula-Prop a) (pr1 x)) - ( b) - ( unit-trunc-Prop (w-ax (imp a-in-logic)))))) - ; (inr not-a-in-logic) → - ( weak-modal-logic-subset-complete-theory - ( x) - ( is-comp) - ( a →ₘ b) - ( forward-implication - ( deduction-lemma - ( pr1 x) - ( transitive-leq-subtype - ( ax-k i) - ( logic) - ( weak-modal-logic (pr1 x)) - ( transitive-leq-subtype - ( logic) - ( pr1 x) - ( weak-modal-logic (pr1 x)) - ( axioms-subset-weak-modal-logic (pr1 x)) - ( logic-subset-L-complete-theory x is-comp)) - ( contains-ax-k)) - ( transitive-leq-subtype - ( ax-s i) - ( logic) - ( weak-modal-logic (pr1 x)) - ( transitive-leq-subtype - ( logic) - ( pr1 x) - ( weak-modal-logic (pr1 x)) - ( axioms-subset-weak-modal-logic (pr1 x)) - ( logic-subset-L-complete-theory x is-comp)) - ( contains-ax-s)) - ( a) - ( b)) - ( logic-ex-falso - ( theory-add-formula a (pr1 x)) - ( transitive-leq-subtype - ( ax-k i) - ( weak-modal-logic (pr1 x)) - ( weak-modal-logic (theory-add-formula a (pr1 x))) - ( weak-modal-logic-monotic - ( subtype-union-right (Id-formula-Prop a) (pr1 x))) - ( transitive-leq-subtype - ( ax-k i) - ( logic) - ( weak-modal-logic (pr1 x)) - ( transitive-leq-subtype - ( logic) - ( pr1 x) - ( weak-modal-logic (pr1 x)) - ( axioms-subset-weak-modal-logic (pr1 x)) - ( logic-subset-L-complete-theory x is-comp)) - ( contains-ax-k))) - ( transitive-leq-subtype - ( ax-s i) - ( weak-modal-logic (pr1 x)) - ( weak-modal-logic (theory-add-formula a (pr1 x))) - ( weak-modal-logic-monotic - ( subtype-union-right (Id-formula-Prop a) (pr1 x))) - ( transitive-leq-subtype - ( ax-s i) - ( logic) - ( weak-modal-logic (pr1 x)) - ( transitive-leq-subtype - ( logic) - ( pr1 x) - ( weak-modal-logic (pr1 x)) - ( axioms-subset-weak-modal-logic (pr1 x)) - ( logic-subset-L-complete-theory x is-comp)) - ( contains-ax-s))) - ( transitive-leq-subtype - ( ax-dn i) - ( weak-modal-logic (pr1 x)) - ( weak-modal-logic (theory-add-formula a (pr1 x))) - ( weak-modal-logic-monotic - ( subtype-union-right (Id-formula-Prop a) (pr1 x))) - ( transitive-leq-subtype - ( ax-dn i) - ( logic) - ( weak-modal-logic (pr1 x)) - ( transitive-leq-subtype - ( logic) - ( pr1 x) - ( weak-modal-logic (pr1 x)) - ( axioms-subset-weak-modal-logic (pr1 x)) - ( logic-subset-L-complete-theory x is-comp)) - ( contains-ax-dn))) - ( a) - ( b) - ( weak-modal-logic-monotic - { ax₁ = Id-formula-Prop a} - { ax₂ = union-subtype (Id-formula-Prop a) (pr1 x)} - ( subtype-union-left (Id-formula-Prop a) (pr1 x)) - ( a) - ( unit-trunc-Prop (w-ax refl))) - ( weak-modal-logic-monotic - { ax₁ = pr1 x} - { ax₂ = union-subtype (Id-formula-Prop a) (pr1 x)} - ( subtype-union-right (Id-formula-Prop a) (pr1 x)) - ( ~ a) - ( unit-trunc-Prop (w-ax not-a-in-logic))))))}) - - canonical-worlds : subtype (lsuc l1 ⊔ lsuc l2) (modal-theory (l1 ⊔ l2) i) - canonical-worlds x = - Σ-Prop - ( is-L-consistent-theory-Prop x) - ( λ is-cons → is-L-complete-theory-Prop (x , is-cons)) - - canonical-kripke-model-world-type : UU (lsuc l1 ⊔ lsuc l2) - canonical-kripke-model-world-type = - Σ (modal-theory (l1 ⊔ l2) i) (is-in-subtype canonical-worlds) - - lindenbaum : - (x : L-consistent-theory) → - ∃ ( L-consistent-theory) - ( λ y → L-consistent-theory-leq x y × is-L-complete-theory y) - lindenbaum x = - apply-universal-property-trunc-Prop - ( zorn - ( L-consistent-big-theories-Poset) - ( unit-trunc-Prop ( x , refl-leq-subtype (pr1 x))) - ( λ C is-inhabited-C → - ( intro-∃ - ( pair - ( pair - ( resized-chain-union C) - ( is-L-consistent-union C is-inhabited-C)) - ( λ a a-in-x → - ( map-equiv - ( inv-equiv (is-equiv-resize prop-resize (chain-union C a))) - ( apply-universal-property-trunc-Prop - ( is-inhabited-C) - ( exists-Prop - ( type-subtype (pr1 C)) - ( λ s → pr1 (pr1 (pr1 s)) a)) - ( λ y → intro-∃ y (pr2 (pr1 y) a a-in-x)))))) - ( λ y a a-in-y → - ( map-equiv - ( inv-equiv (is-equiv-resize prop-resize (chain-union C a))) - ( intro-∃ y a-in-y)))))) - ( ∃-Prop - ( L-consistent-theory) - ( λ y → L-consistent-theory-leq x y × is-L-complete-theory y)) - ( λ ((m , x-leq-m) , is-max) → - ( intro-∃ m - ( x-leq-m , λ y m-leq-y → - ( ap pr1 - ( is-max - ( pair y - ( transitive-leq-Poset - ( L-consistent-theories-Poset) - ( x) - ( m) - ( y) - ( m-leq-y) - ( x-leq-m))) - ( m-leq-y)))))) - where - L-consistent-big-theories-Poset : Poset (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) - L-consistent-big-theories-Poset = - poset-Subposet - ( L-consistent-theories-Poset) - ( L-consistent-theory-leq-Prop x) - - chain-union : - chain-Poset l2 L-consistent-big-theories-Poset → - modal-theory (lsuc l1 ⊔ lsuc l2) i - chain-union C a = - exists-Prop - ( type-subtype - ( sub-preorder-chain-Poset - ( L-consistent-big-theories-Poset) - ( C))) - ( λ s → pr1 (pr1 (pr1 s)) a) - - resized-chain-union : - chain-Poset l2 L-consistent-big-theories-Poset → type-Poset theories-Poset - resized-chain-union C = resize prop-resize ∘ chain-union C - - in-chain-weak-deduction-list-chain-union : - (C : chain-Poset l2 L-consistent-big-theories-Poset) - (e : type-chain-Poset L-consistent-big-theories-Poset C) - {a : formula i} - {l : list (formula i)} → - in-list l ⊆ union-subtype logic (chain-union C) → - in-list l ▷ a → - ∃ ( type-chain-Poset L-consistent-big-theories-Poset C) - ( λ y → - ( is-in-subtype - ( weak-modal-logic (union-subtype logic (pr1 (pr1 (pr1 y))))) - ( a))) - in-chain-weak-deduction-list-chain-union C e {a} {nil} sub wd = - intro-∃ e (ex-falso (nil-no-deduction wd)) - in-chain-weak-deduction-list-chain-union - C e {a} {cons y l} sub (w-ax is-ax) = - apply-universal-property-trunc-Prop - ( sub a is-ax) - ( ∃-Prop ( type-chain-Poset L-consistent-big-theories-Poset C) - ( λ y → - ( is-in-subtype - ( weak-modal-logic (union-subtype logic (pr1 (pr1 (pr1 y))))) - ( a)))) - ( λ - { (inl in-logic) → - ( intro-∃ e - ( weak-modal-logic-ax - ( subtype-union-left logic (pr1 (pr1 (pr1 e))) a in-logic))) - ; (inr in-union) → - ( apply-universal-property-trunc-Prop - ( in-union) - ( ∃-Prop ( type-chain-Poset L-consistent-big-theories-Poset C) - ( λ y → - ( is-in-subtype - ( weak-modal-logic - ( union-subtype logic (pr1 (pr1 (pr1 y))))) - ( a)))) - ( λ (y , a-in-y) → - ( intro-∃ y - ( weak-modal-logic-ax - (subtype-union-right - ( logic) - ( pr1 (pr1 (pr1 y))) - ( a) - ( a-in-y))))))}) - in-chain-weak-deduction-list-chain-union - C e {a} {l@(cons _ _)} sub (w-mp {b} wdba wdb) = - apply-twice-universal-property-trunc-Prop - ( in-chain-weak-deduction-list-chain-union C e {l = l} sub wdba) - ( in-chain-weak-deduction-list-chain-union C e {l = l} sub wdb) - ( ∃-Prop - ( type-chain-Poset L-consistent-big-theories-Poset C) - ( λ y → - ( is-in-subtype - ( weak-modal-logic (union-subtype logic (pr1 (pr1 (pr1 y))))) - ( a)))) - ( λ (y , ba-in-y) (z , b-in-z) → - ( apply-universal-property-trunc-Prop - ( pr2 C y z) - ( ∃-Prop - ( type-chain-Poset L-consistent-big-theories-Poset C) - ( λ y → - ( is-in-subtype - ( weak-modal-logic (union-subtype logic (pr1 (pr1 (pr1 y))))) - ( a)))) - ( λ - { (inl y-leq-z) → - ( intro-∃ z - ( weak-modal-logic-mp - ( weak-modal-logic-monotic - {ax₁ = union-subtype logic (pr1 (pr1 (pr1 y)))} - {ax₂ = union-subtype logic (pr1 (pr1 (pr1 z)))} - ( subset-union-subset-right - ( logic) - ( pr1 (pr1 (pr1 y))) - ( pr1 (pr1 (pr1 z))) - ( y-leq-z)) - ( b →ₘ a) - ( ba-in-y)) - ( b-in-z))) - ; (inr z-leq-y) → - ( intro-∃ y - ( weak-modal-logic-mp - (ba-in-y) - ( weak-modal-logic-monotic - {ax₁ = union-subtype logic (pr1 (pr1 (pr1 z)))} - {ax₂ = union-subtype logic (pr1 (pr1 (pr1 y)))} - ( subset-union-subset-right - ( logic) - ( pr1 (pr1 (pr1 z))) - ( pr1 (pr1 (pr1 y))) - ( z-leq-y)) - ( b) - ( b-in-z))))}))) - - in-chain-in-chain-union : - (C : chain-Poset l2 L-consistent-big-theories-Poset) → - is-inhabited (type-chain-Poset L-consistent-big-theories-Poset C) → - {a : formula i} → - is-in-subtype - ( weak-modal-logic (union-subtype logic (chain-union C))) a → - ∃ ( type-subtype - ( sub-preorder-chain-Poset L-consistent-big-theories-Poset C)) - ( λ x → - ( is-in-subtype - ( weak-modal-logic (union-subtype logic (pr1 (pr1 (pr1 x))))) - ( a))) - in-chain-in-chain-union C is-inhabited-C {a} in-union = - apply-twice-universal-property-trunc-Prop - ( in-union) - ( is-inhabited-C) - ( ∃-Prop - ( type-subtype - ( sub-preorder-chain-Poset L-consistent-big-theories-Poset C)) - ( λ x → - ( is-in-subtype - ( weak-modal-logic (union-subtype logic (pr1 (pr1 (pr1 x))))) - ( a)))) - ( λ wd c → - ( let (l , sub , wdl) = list-assumptions-weak-deduction wd - in in-chain-weak-deduction-list-chain-union C c sub wdl)) - - is-L-consistent-union : - (C : chain-Poset l2 L-consistent-big-theories-Poset) → - is-inhabited (type-chain-Poset L-consistent-big-theories-Poset C) → - is-L-consistent-theory (resized-chain-union C) - is-L-consistent-union C C-is-inhabited bot-in-union = - apply-universal-property-trunc-Prop - ( in-chain-in-chain-union C C-is-inhabited - ( weak-modal-logic-monotic - {ax₁ = union-subtype logic (resized-chain-union C)} - {ax₂ = union-subtype logic (chain-union C)} - ( subset-union-subset-right - ( logic) - ( resized-chain-union C) - ( chain-union C) - ( λ b → - ( map-equiv (is-equiv-resize prop-resize (chain-union C b))))) - ( ⊥) - ( bot-in-union))) - ( empty-Prop) - ( λ (y , bot-in-y) → pr2 (pr1 (pr1 y)) bot-in-y) - - is-L-consistent-logic : is-L-consistent-theory logic - is-L-consistent-logic bot-in-logic = - L-is-cons - ( transitive-leq-subtype - ( weak-modal-logic (union-subtype logic logic)) - ( modal-logic (union-subtype logic logic)) - ( logic) - ( transitive-leq-subtype - ( modal-logic (union-subtype logic logic)) - ( modal-logic logic) - ( logic) - ( subset-double-modal-logic axioms) - ( modal-logic-monotic - ( subtype-union-both - ( logic) - ( logic) - ( logic) - ( refl-leq-subtype logic) - ( refl-leq-subtype logic)))) - ( weak-modal-logic-subset-modal-logic (union-subtype logic logic)) - ( ⊥) - ( bot-in-logic)) - - is-inhabited-canonical-kripke-model-world : - is-inhabited canonical-kripke-model-world-type - is-inhabited-canonical-kripke-model-world = - apply-universal-property-trunc-Prop - ( lindenbaum (logic , is-L-consistent-logic)) - ( is-inhabited-Prop canonical-kripke-model-world-type) - ( λ x → unit-trunc-Prop ((pr1 (pr1 x)) , ((pr2 (pr1 x)) , (pr2 (pr2 x))))) - - canonical-kripke-frame : kripke-frame (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) - pr1 (pr1 canonical-kripke-frame) = canonical-kripke-model-world-type - pr2 (pr1 canonical-kripke-frame) = is-inhabited-canonical-kripke-model-world - pr2 canonical-kripke-frame x y = - Π-Prop - ( formula i) - ( λ a → hom-Prop (pr1 x (□ a)) (pr1 y a)) - canonical-kripke-model : kripke-model (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) i (l1 ⊔ l2) - pr1 canonical-kripke-model = canonical-kripke-frame - pr2 canonical-kripke-model n x = pr1 x (var n) + pr1 (pr1 (pr1 canonical-kripke-model)) = + L-complete-theory logic (l1 ⊔ l2) + pr2 (pr1 (pr1 canonical-kripke-model)) = + is-inhabited-L-complete-theory + ( logic) + ( prop-resize) + ( zorn) + ( is-weak-logic) + ( is-cons) + ( contains-ax-k) + ( contains-ax-s) + pr2 (pr1 canonical-kripke-model) x y = + Π-Prop + ( formula i) + ( λ a → + ( hom-Prop + ( modal-theory-L-complete-theory logic x (□ a)) + ( modal-theory-L-complete-theory logic y a))) + pr2 canonical-kripke-model n x = + modal-theory-L-complete-theory logic x (var n) - complete-theory-box : - LEM (l1 ⊔ l2) → - (x : L-consistent-theory) - (is-comp : is-L-complete-theory x) - (a : formula i) → - ( (y : modal-theory (l1 ⊔ l2) i) → - (is-canonical : is-in-subtype canonical-worlds y) → - relation-kripke-model i canonical-kripke-model - ( pr1 x , pr2 x , is-comp) - ( y , is-canonical) → - is-in-subtype y a) → - is-in-subtype (pr1 x) (□ a) - complete-theory-box lem x is-comp a h = - apply-universal-property-trunc-Prop - ( disjuctivity-L-complete-theory lem x is-comp (□ a)) - ( pr1 x (□ a)) - ( λ { (inl box-in-logic) → box-in-logic - ; (inr not-box-in-logic) → - ( ex-falso - ( apply-universal-property-trunc-Prop - ( lindenbaum (w not-box-in-logic)) - ( empty-Prop) - ( λ (y , w-leq-y , y-is-comp) → - ( pr2 y - ( weak-modal-logic-monotic - { ax₁ = pr1 y} - { ax₂ = union-subtype logic (pr1 y)} - ( subtype-union-right logic (pr1 y)) - ( ⊥) - ( weak-modal-logic-mp - ( axioms-subset-weak-modal-logic - ( pr1 y) - ( ~ a) - ( w-leq-y - ( ~ a) - ( subtype-union-left - ( Id-formula-Prop (~ a)) - ( λ b → pr1 x (□ b)) - ( ~ a) - ( refl)))) - ( weak-modal-logic-ax - ( h (pr1 y) (pr2 y , y-is-comp) - ( λ b box-in-logic → - ( w-leq-y b - ( subtype-union-right - ( Id-formula-Prop (~ a)) - ( λ b → pr1 x (□ b)) - ( b) - ( box-in-logic))))))))))))}) + module _ + (x@((theory , is-cons) , is-comp) : L-complete-theory logic (l1 ⊔ l2)) where - list-to-implications : - formula i → - (l : list (formula i)) → - formula i - list-to-implications f nil = f - list-to-implications f (cons g l) = list-to-implications (g →ₘ f) l - - list-to-implications-rev : - formula i → - (l : list (formula i)) → - formula i - list-to-implications-rev f nil = f - list-to-implications-rev f (cons g l) = g →ₘ list-to-implications-rev f l - - list-to-implication-rev-snoc : - (f g : formula i) (l : list (formula i)) → - list-to-implications f (snoc l g) = g →ₘ list-to-implications f l - list-to-implication-rev-snoc f g nil = refl - list-to-implication-rev-snoc f g (cons h l) = - list-to-implication-rev-snoc (h →ₘ f) g l - reverse-list-to-implications : - (f : formula i) (l : list (formula i)) → - list-to-implications f (reverse-list l) = list-to-implications-rev f l - reverse-list-to-implications f nil = refl - reverse-list-to-implications f (cons g l) = - ( list-to-implication-rev-snoc f g (reverse-list l)) ∙ - ( ap (λ x → g →ₘ x) (reverse-list-to-implications f l)) - - move-right : - {l' : Level} (axioms : modal-theory l' i) - (f : formula i) (l : list (formula i)) → - ax-k i ⊆ weak-modal-logic axioms → - ax-s i ⊆ weak-modal-logic axioms → - is-in-subtype (weak-modal-logic (union-subtype axioms (in-list l))) f → - is-in-subtype (weak-modal-logic axioms) (list-to-implications f l) - move-right axioms f nil _ _ = - weak-modal-logic-monotic - ( subtype-union-both - ( axioms) - ( in-list nil) - ( axioms) - ( refl-leq-subtype axioms) - ( λ _ → ex-falso ∘ empty-in-list-nil)) - ( f) - move-right axioms f (cons c l) contains-ax-k contains-ax-s wd = - move-right axioms (c →ₘ f) l contains-ax-k contains-ax-s + private + contains-ax-k' : ax-k i ⊆ theory + contains-ax-k' = + transitive-leq-subtype (ax-k i) logic theory + ( subset-logic-L-complete-theory logic lzero x) + ( contains-ax-k) + + contains-ax-s' : ax-s i ⊆ theory + contains-ax-s' = + transitive-leq-subtype (ax-s i) logic theory + ( subset-logic-L-complete-theory logic lzero x) + ( contains-ax-s) + + contains-ax-dn' : ax-dn i ⊆ theory + contains-ax-dn' = + transitive-leq-subtype (ax-dn i) logic theory + ( subset-logic-L-complete-theory logic lzero x) + ( contains-ax-dn) + + contains-ax-n' : ax-n i ⊆ theory + contains-ax-n' = + transitive-leq-subtype (ax-n i) logic theory + ( subset-logic-L-complete-theory logic lzero x) + ( contains-ax-n) + + contains-ax-k-union : + {l : Level} (t : modal-theory l i) → ax-k i ⊆ logic ∪ t + contains-ax-k-union t = + transitive-leq-subtype (ax-k i) logic (logic ∪ t) + ( subtype-union-left logic t) + ( contains-ax-k) + + contains-ax-s-union : + {l : Level} (t : modal-theory l i) → ax-s i ⊆ logic ∪ t + contains-ax-s-union t = + transitive-leq-subtype (ax-s i) logic (logic ∪ t) + ( subtype-union-left logic t) + ( contains-ax-s) + + is-disjuctive-theory : is-disjuctive-modal-theory theory + is-disjuctive-theory = + is-disjuctive-L-complete-theory logic x + ( contains-ax-k) + ( contains-ax-s) + ( contains-ax-dn) + ( lem) + + L-complete-theory-implication : + {a b : formula i} → + (is-in-subtype theory a → is-in-subtype theory b) → + is-in-subtype theory (a →ₘ b) + L-complete-theory-implication {a} {b} f with is-disjuctive-theory a + ... | inl a-in-logic = + is-weak-modal-logic-L-complete-theory logic lzero x (a →ₘ b) ( forward-implication - ( deduction-lemma - ( union-subtype axioms (in-list l)) - ( transitive-leq-subtype - ( ax-k i) - ( weak-modal-logic axioms) - ( weak-modal-logic (union-subtype axioms (in-list l))) - ( weak-modal-logic-monotic - ( subtype-union-left axioms (in-list l))) - ( contains-ax-k)) - ( transitive-leq-subtype - ( ax-s i) - ( weak-modal-logic axioms) - ( weak-modal-logic (union-subtype axioms (in-list l))) - ( weak-modal-logic-monotic - ( subtype-union-left axioms (in-list l))) - ( contains-ax-s)) - ( c) - ( f)) - ( weak-modal-logic-monotic - ( transitive-leq-subtype - ( union-subtype axioms (in-list (cons c l))) - ( union-subtype axioms (theory-add-formula c (in-list l))) - ( theory-add-formula c (union-subtype axioms (in-list l))) - ( union-swap-1-2 axioms (Id-formula-Prop c) (in-list l)) - ( subset-union-subset-right - ( axioms) - ( in-list (cons c l)) - ( theory-add-formula c (in-list l)) - ( backward-subset-head-add c l))) - ( f) - ( wd))) - - aux'''' : - (l : list (formula i)) → - in-list l ⊆ (λ b → pr1 x (□ b)) → - is-in-subtype - ( weak-modal-logic (union-subtype logic (pr1 x))) - ( □ (list-to-implications-rev a l)) → - is-in-subtype (weak-modal-logic (union-subtype logic (pr1 x))) (□ a) - aux'''' nil sub in-logic = in-logic - aux'''' (cons c l) sub in-logic = - aux'''' l - ( transitive-leq-subtype - ( in-list l) - ( in-list (cons c l)) - ( λ b → pr1 x (□ b)) - ( sub) - ( subset-in-tail c l)) - ( weak-modal-logic-mp - { a = □ c} - { b = □ (list-to-implications-rev a l)} - ( weak-modal-logic-mp - { a = □ (c →ₘ list-to-implications-rev a l)} - { b = □ c →ₘ □ (list-to-implications-rev a l)} - ( weak-modal-logic-monotic - { ax₁ = logic} - { ax₂ = union-subtype logic (pr1 x)} - ( subtype-union-left logic (pr1 x)) - ( □ (c →ₘ list-to-implications-rev a l) →ₘ - □ c →ₘ □ list-to-implications-rev a l) - ( weak-modal-logic-ax - ( contains-ax-n _ (c , list-to-implications-rev a l , refl)))) - ( in-logic)) - ( weak-modal-logic-ax - ( subtype-union-right logic - ( pr1 x) - ( □ c) - ( sub c (unit-trunc-Prop (is-head c l)))))) + ( deduction-lemma theory contains-ax-k' contains-ax-s' a b) + ( weak-modal-logic-closure-monotic + { ax₁ = theory} + { ax₂ = theory-add-formula a theory} + ( subset-add-formula a theory) + ( b) + ( weak-modal-logic-closure-ax (f a-in-logic)))) + ... | inr not-a-in-logic = + is-weak-modal-logic-L-complete-theory logic lzero x (a →ₘ b) + ( forward-implication + ( deduction-lemma theory contains-ax-k' contains-ax-s' a b) + ( logic-ex-falso + ( theory-add-formula a theory) + ( transitive-subset-add-formula a theory (ax-k i) contains-ax-k') + ( transitive-subset-add-formula a theory (ax-s i) contains-ax-s') + ( transitive-subset-add-formula a theory (ax-dn i) contains-ax-dn') + ( a) + ( b) + ( weak-modal-logic-closure-ax (formula-in-add-formula a theory)) + ( weak-modal-logic-closure-monotic + { ax₁ = theory} + { ax₂ = theory-add-formula a theory} + ( subset-add-formula a theory) + ( ~ a) + ( weak-modal-logic-closure-ax not-a-in-logic)))) - aux''' : - (l : list (formula i)) → - in-list l ⊆ (λ b → pr1 x (□ b)) → - is-in-subtype (weak-modal-logic (union-subtype logic (in-list l))) (a) → - is-in-subtype (weak-modal-logic (union-subtype logic (pr1 x))) (□ a) - aux''' l sub in-logic = - aux'''' l sub - ( weak-modal-logic-ax - ( subtype-union-left + L-complete-theory-box : + {a : formula i} → + ( (y : L-complete-theory logic (l1 ⊔ l2)) → + relation-kripke-model i canonical-kripke-model x y → + is-in-subtype (modal-theory-L-complete-theory logic y) a) → + is-in-subtype theory (□ a) + L-complete-theory-box {a} f with is-disjuctive-theory (□ a) + ... | inl box-a-in-logic = box-a-in-logic + ... | inr not-box-a-in-logic = + ex-falso + ( apply-universal-property-trunc-Prop + ( lindenbaum logic contains-ax-k contains-ax-s zorn prop-resize + ( y , is-L-consistent-y)) + ( empty-Prop) + ( λ (w , y-leq-w) → + ( is-consistent-modal-theory-L-complete-theory logic w + ( weak-modal-logic-mp + ( is-weak-modal-logic-L-complete-theory logic lzero w) + ( y-leq-w (~ a) + ( formula-in-add-formula (~ a) (unbox-modal-theory theory))) + ( f w (λ b → y-leq-w b ∘ y-contains-unbox)))))) + where + y : modal-theory (l1 ⊔ l2) i + y = theory-add-formula (~ a) (unbox-modal-theory theory) + + y-contains-unbox : + {b : formula i} → + is-in-subtype theory (□ b) → + is-in-subtype y b + y-contains-unbox {b} = + subset-add-formula (~ a) (unbox-modal-theory theory) b + + list-to-implications : formula i → (l : list (formula i)) → formula i + list-to-implications f nil = f + list-to-implications f (cons g l) = list-to-implications (g →ₘ f) l + + list-to-implications-rev : formula i → (l : list (formula i)) → formula i + list-to-implications-rev f nil = f + list-to-implications-rev f (cons g l) = g →ₘ list-to-implications-rev f l + + list-to-implication-rev-snoc : + (f g : formula i) (l : list (formula i)) → + list-to-implications f (snoc l g) = g →ₘ list-to-implications f l + list-to-implication-rev-snoc f g nil = refl + list-to-implication-rev-snoc f g (cons h l) = + list-to-implication-rev-snoc (h →ₘ f) g l + + eq-reverse-list-to-implications : + (f : formula i) (l : list (formula i)) → + list-to-implications f (reverse-list l) = list-to-implications-rev f l + eq-reverse-list-to-implications f nil = refl + eq-reverse-list-to-implications f (cons g l) = + ( list-to-implication-rev-snoc f g (reverse-list l)) ∙ + ( ap (λ x → g →ₘ x) (eq-reverse-list-to-implications f l)) + + move-assumptions-right : + (f : formula i) (l : list (formula i)) → + is-in-subtype (weak-modal-logic-closure (logic ∪ list-subtype l)) f → + is-in-subtype + ( weak-modal-logic-closure logic) + ( list-to-implications f l) + move-assumptions-right f nil = + weak-modal-logic-closure-monotic + ( subtype-union-both ( logic) - ( pr1 x) - ( □ (list-to-implications-rev a l)) - ( modal-logic-nec + ( list-subtype nil) + ( logic) + ( refl-leq-subtype logic) + ( subset-list-subtype-nil logic)) + ( f) + move-assumptions-right f (cons c l) in-logic = + move-assumptions-right (c →ₘ f) l + ( forward-implication + ( deduction-lemma + ( logic ∪ list-subtype l) + ( contains-ax-k-union (list-subtype l)) + ( contains-ax-s-union (list-subtype l)) + ( c) + ( f)) + ( weak-modal-logic-closure-monotic + ( transitive-leq-subtype + ( logic ∪ list-subtype (cons c l)) + ( logic ∪ theory-add-formula c (list-subtype l)) + ( theory-add-formula c (logic ∪ list-subtype l)) + ( theory-add-formula-union-right c logic (list-subtype l)) + ( subset-union-subset-right logic + ( list-subtype (cons c l)) + ( theory-add-formula c (list-subtype l)) + ( subset-list-subtype-cons + ( theory-add-formula c (list-subtype l)) + ( formula-in-add-formula c (list-subtype l)) + ( subset-add-formula c (list-subtype l))))) + ( f) + ( in-logic))) + + α : + (l : list (formula i)) → + list-subtype l ⊆ unbox-modal-theory theory → + is-in-subtype theory (□ (list-to-implications-rev a l)) → + is-in-subtype theory (□ a) + α nil sub in-logic = in-logic + α (cons c l) sub in-logic = + α l + ( transitive-leq-subtype + ( list-subtype l) + ( list-subtype (cons c l)) + ( unbox-modal-theory theory) + ( sub) + ( subset-tail-list-subtype)) + ( weak-modal-logic-mp + ( is-weak-modal-logic-L-complete-theory logic lzero x) + { a = □ c} + { b = □ list-to-implications-rev a l} + ( weak-modal-logic-mp + ( is-weak-modal-logic-L-complete-theory logic lzero x) + { a = □ (c →ₘ list-to-implications-rev a l)} + { b = □ c →ₘ □ (list-to-implications-rev a l)} + ( contains-ax-n' _ (c , list-to-implications-rev a l , refl)) + ( in-logic)) + ( sub c head-in-list-subtype)) + + β : + (l : list (formula i)) → + list-subtype l ⊆ unbox-modal-theory theory → + is-in-subtype (weak-modal-logic-closure (logic ∪ list-subtype l)) a → + is-in-subtype theory (□ a) + β l sub in-logic = + α l sub + ( subset-logic-L-complete-theory logic lzero x + ( □ list-to-implications-rev a l) + ( modal-logic-nec is-logic ( tr ( is-in-subtype logic) - ( reverse-list-to-implications a l) - ( subset-double-modal-logic - ( axioms) + ( eq-reverse-list-to-implications a l) + ( is-weak-logic ( list-to-implications a (reverse-list l)) - ( weak-modal-logic-subset-modal-logic - ( logic) - ( list-to-implications a (reverse-list l)) - ( move-right logic a (reverse-list l) - ( transitive-leq-subtype - ( ax-k i) - ( logic) - ( weak-modal-logic logic) - ( axioms-subset-weak-modal-logic logic) - ( contains-ax-k)) - ( transitive-leq-subtype - ( ax-s i) - ( logic) - ( weak-modal-logic logic) - ( axioms-subset-weak-modal-logic logic) - ( contains-ax-s)) - ( weak-modal-logic-monotic - ( subset-union-subset-right logic - ( in-list l) - ( in-list (reverse-list l)) - ( subset-reversing-list l)) - ( a) - ( in-logic))))))))) - - aux'' : - (l : list (formula i)) → - in-list l ⊆ (λ b → pr1 x (□ b)) → - is-in-subtype - ( weak-modal-logic - ( theory-add-formula (~ a) (union-subtype logic (in-list l)))) - ( ⊥) → - is-in-subtype (weak-modal-logic (union-subtype logic (pr1 x))) (□ a) - aux'' l sub in-logic = - aux''' l sub - ( weak-modal-logic-mp - { a = ~~ a} - { b = a} - ( weak-modal-logic-ax - ( subtype-union-left logic - ( in-list l) - ( _) - ( contains-ax-dn _ (a , refl)))) - ( forward-implication - ( deduction-lemma - ( union-subtype logic (in-list l)) - ( transitive-leq-subtype - ( ax-k i) - ( logic) - ( weak-modal-logic (union-subtype logic (in-list l))) - ( transitive-leq-subtype - ( logic) - ( union-subtype logic (in-list l)) - ( weak-modal-logic (union-subtype logic (in-list l))) - ( axioms-subset-weak-modal-logic - ( union-subtype logic (in-list l))) - ( subtype-union-left logic (in-list l))) - ( contains-ax-k)) - ( transitive-leq-subtype - ( ax-s i) - ( logic) - ( weak-modal-logic (union-subtype logic (in-list l))) - ( transitive-leq-subtype - ( logic) - ( union-subtype logic (in-list l)) - ( weak-modal-logic (union-subtype logic (in-list l))) - ( axioms-subset-weak-modal-logic - ( union-subtype logic (in-list l))) - ( subtype-union-left logic (in-list l))) - ( contains-ax-s)) - ( ~ a) - ( ⊥)) - ( in-logic))) - - aux' : - (l : list (formula i)) → - in-list l ⊆ (λ b → pr1 x (□ b)) → - is-in-subtype - ( weak-modal-logic - ( union-subtype logic (theory-add-formula (~ a) (in-list l)))) - ( ⊥) → - is-in-subtype (weak-modal-logic (union-subtype logic (pr1 x))) (□ a) - aux' l sub in-logic = - aux'' l sub - ( weak-modal-logic-monotic - ( union-swap-1-2 - ( logic) - ( Id-formula-Prop (~ a)) - ( in-list l)) - ( ⊥) - ( in-logic)) - - aux : - (l : list (formula i)) → - in-list l ⊆ - union-subtype logic (theory-add-formula (~ a) (λ b → pr1 x (□ b))) → - is-in-subtype (weak-modal-logic (in-list l)) ⊥ → - is-in-subtype (weak-modal-logic (union-subtype logic (pr1 x))) (□ a) - aux l sub in-logic = - apply-universal-property-trunc-Prop - ( lists-in-union-lists l logic - ( theory-add-formula (~ a) (λ b → pr1 x (□ b))) - ( sub)) - ( weak-modal-logic (union-subtype logic (pr1 x)) (□ a)) - ( λ (l-ax , l-w , l-sub-union , l-ax-sub-axioms , l-w-sub-w) → - ( apply-universal-property-trunc-Prop - ( lists-in-union-lists l-w - ( Id-formula-Prop (~ a)) - ( λ b → pr1 x (□ b)) - ( l-w-sub-w)) - ( weak-modal-logic (union-subtype logic (pr1 x)) (□ a)) - ( λ (l-not-a , l-boxes , l-sub-union' , l-not-a-sub , l-boxes-sub) → - ( aux' - ( l-boxes) - ( l-boxes-sub) - ( weak-modal-logic-monotic - { ax₁ = in-list l} - { ax₂ = - union-subtype logic - ( theory-add-formula (~ a) (in-list l-boxes))} - ( transitive-leq-subtype - ( in-list l) - ( union-subtype (in-list l-ax) (in-list l-w)) - ( union-subtype - ( logic) - ( theory-add-formula (~ a) (in-list l-boxes))) - ( subset-union-subsets - ( in-list l-ax) - ( in-list l-w) - ( logic) - ( theory-add-formula (~ a) (in-list l-boxes)) - ( l-ax-sub-axioms) + ( move-assumptions-right a (reverse-list l) + ( weak-modal-logic-closure-monotic + ( subset-union-subset-right logic + ( list-subtype l) + ( list-subtype (reverse-list l)) + ( subset-list-subtype-reverse-list l)) + ( a) + ( in-logic))))))) + + γ : + (l : list (formula i)) → + list-subtype l ⊆ unbox-modal-theory theory → + is-contradictory-modal-logic + ( weak-modal-logic-closure + ( theory-add-formula (~ a) (logic ∪ list-subtype l))) → + is-in-subtype theory (□ a) + γ l sub is-cont = + β l sub + ( weak-modal-logic-closure-mp {a = ~~ a} {b = a} + ( weak-modal-logic-closure-ax + ( subtype-union-left logic (list-subtype l) (~~ a →ₘ a) + ( contains-ax-dn (~~ a →ₘ a) (a , refl)))) + ( forward-implication + ( deduction-lemma + ( logic ∪ list-subtype l) + ( contains-ax-k-union (list-subtype l)) + ( contains-ax-s-union (list-subtype l)) + ( ~ a) + ( ⊥)) + ( is-cont))) + + δ : + (l : list (formula i)) → + list-subtype l ⊆ unbox-modal-theory theory → + is-contradictory-modal-logic + ( weak-modal-logic-closure + ( logic ∪ (theory-add-formula (~ a) (list-subtype l)))) → + is-in-subtype theory (□ a) + δ l sub is-cont = + γ l sub + ( is-contradictory-modal-logic-monotic + ( weak-modal-logic-closure + ( logic ∪ theory-add-formula (~ a) (list-subtype l))) + ( weak-modal-logic-closure + ( theory-add-formula (~ a) (logic ∪ list-subtype l))) + ( weak-modal-logic-closure-monotic + ( theory-add-formula-union-right (~ a) logic (list-subtype l))) + ( is-cont)) + + ε : + (l : list (formula i)) → + list-subtype l ⊆ logic ∪ y → + is-contradictory-modal-logic + ( weak-modal-logic-closure (list-subtype l)) → + is-in-subtype theory (□ a) + ε l sub is-cont = + apply-universal-property-trunc-Prop + ( lists-in-union-lists l logic y sub) + ( theory (□ a)) + ( λ ((l-ax , l-y) , l-sub-union , l-ax-sub-logic , l-y-sub-y) → + ( apply-universal-property-trunc-Prop + ( lists-in-union-lists l-y + ( Id-formula-Prop (~ a)) + ( unbox-modal-theory theory) + ( l-y-sub-y)) + ( theory (□ a)) + ( λ ((l-not-a , l-box) , l-sub-union' , l-not-a-sub , l-box-sub) → + ( δ + ( l-box) + ( l-box-sub) + ( is-contradictory-modal-logic-monotic + ( weak-modal-logic-closure (list-subtype l)) + ( weak-modal-logic-closure + ( logic ∪ theory-add-formula (~ a) (list-subtype l-box))) + ( weak-modal-logic-closure-monotic ( transitive-leq-subtype - ( in-list l-w) - ( union-subtype (in-list l-not-a) (in-list l-boxes)) - ( theory-add-formula (~ a) (in-list l-boxes)) - ( subset-union-subset-left - ( in-list l-not-a) - ( Id-formula-Prop (~ a)) - ( in-list l-boxes) - ( l-not-a-sub)) - ( l-sub-union'))) - ( l-sub-union)) - ( ⊥) - ( in-logic)))))) - - w : is-in-subtype (pr1 x) (~ □ a) → L-consistent-theory - pr1 (w b) = theory-add-formula (~ a) (λ b → pr1 x (□ b)) - pr2 (w not-box-in-logic) bot-in-logic = - apply-universal-property-trunc-Prop - ( bot-in-logic) - ( empty-Prop) - ( λ wd-bot → - ( let (l , sub , wd) = list-assumptions-weak-deduction wd-bot - in pr2 x + ( list-subtype l) + ( list-subtype l-ax ∪ list-subtype l-y) + ( logic ∪ theory-add-formula (~ a) (list-subtype l-box)) + ( subset-union-subsets + ( list-subtype l-ax) + ( list-subtype l-y) + ( logic) + ( theory-add-formula (~ a) (list-subtype l-box)) + ( l-ax-sub-logic) + ( transitive-leq-subtype + ( list-subtype l-y) + ( list-subtype l-not-a ∪ list-subtype l-box) + ( theory-add-formula (~ a) (list-subtype l-box)) + ( subtype-union-both + ( list-subtype l-not-a) + ( list-subtype l-box) + ( theory-add-formula (~ a) (list-subtype l-box)) + ( transitive-leq-subtype + ( list-subtype l-not-a) + ( Id-formula-Prop (~ a)) + ( theory-add-formula (~ a) (list-subtype l-box)) + ( subtype-union-left + ( Id-formula-Prop (~ a)) + ( list-subtype l-box)) + ( l-not-a-sub)) + ( subset-add-formula (~ a) (list-subtype l-box))) + ( l-sub-union'))) + ( l-sub-union))) + ( is-cont)))))) + + is-L-consistent-y : is-L-consistent-theory logic y + is-L-consistent-y = + map-universal-property-trunc-Prop + ( empty-Prop) + ( λ d-bot → + ( is-consistent-modal-theory-L-complete-theory logic x ( weak-modal-logic-mp - ( weak-modal-logic-ax - ( subtype-union-right logic (pr1 x) _ not-box-in-logic)) - ( aux l sub (unit-trunc-Prop wd))))) - - module _ - (lem : LEM (l1 ⊔ l2)) + ( is-weak-modal-logic-L-complete-theory logic lzero x) + ( not-box-a-in-logic) + ( ε + ( list-assumptions-weak-deduction d-bot) + ( subset-theory-list-assumptions-weak-deduction d-bot) + ( is-in-weak-modal-logic-closure-weak-deduction + ( is-assumptions-list-assumptions-weak-deduction d-bot)))))) + + canonical-model-theorem-pointwise : + (a : formula i) + (x : L-complete-theory logic (l1 ⊔ l2)) → + type-iff-Prop + ( modal-theory-L-complete-theory logic x a) + ( (canonical-kripke-model , x) ⊨ a) + pr1 (canonical-model-theorem-pointwise (var n) x) = map-raise + pr1 (canonical-model-theorem-pointwise ⊥ x) = + map-raise ∘ is-consistent-modal-theory-L-complete-theory logic x + pr1 (canonical-model-theorem-pointwise (a →ₘ b) x) in-logic f = + forward-implication + ( canonical-model-theorem-pointwise b x) + ( weak-modal-logic-mp + ( is-weak-modal-logic-L-complete-theory logic lzero x) + ( in-logic) + ( backward-implication (canonical-model-theorem-pointwise a x) f)) + pr1 (canonical-model-theorem-pointwise (□ a) x) in-logic y xRy = + forward-implication + ( canonical-model-theorem-pointwise a y) + ( xRy a in-logic) + pr2 (canonical-model-theorem-pointwise (var n) x) = map-inv-raise + pr2 (canonical-model-theorem-pointwise ⊥ x) (map-raise ()) + pr2 (canonical-model-theorem-pointwise (a →ₘ b) x) f = + L-complete-theory-implication x + ( λ in-x → + ( backward-implication + ( canonical-model-theorem-pointwise b x) + ( f + ( forward-implication + ( canonical-model-theorem-pointwise a x) + ( in-x))))) + pr2 (canonical-model-theorem-pointwise (□ a) x) f = + L-complete-theory-box x + ( λ y xRy → + ( backward-implication + ( canonical-model-theorem-pointwise a y) + ( f y xRy))) + + canonical-model-theorem : + (a : formula i) → logic a ⇔ (canonical-kripke-model ⊨M a) + pr1 (canonical-model-theorem a) in-logic x = + forward-implication + ( canonical-model-theorem-pointwise a x) + ( subset-logic-L-complete-theory logic lzero x a in-logic) + pr2 (canonical-model-theorem a) = + contraposition (lower-LEM l1 lem) (logic a) + ( λ na f → + ( apply-universal-property-trunc-Prop + ( lindenbaum logic contains-ax-k contains-ax-s zorn prop-resize + ( x , is-L-consistent-x na)) + ( empty-Prop) + ( λ (w , leq) → + ( is-consistent-modal-theory-L-complete-theory logic w + ( weak-modal-logic-mp + ( is-weak-modal-logic-L-complete-theory logic lzero w) + ( leq (~ a) not-a-in-x) + ( backward-implication + ( canonical-model-theorem-pointwise a w) + ( f w))))))) where - - canonical-model-theorem : - (a : formula i) - (x : canonical-kripke-model-world-type) → - pr1 x a ⇔ ((canonical-kripke-model , x) ⊨ a) - pr1 (canonical-model-theorem (var n) x) in-logic = - map-raise in-logic - pr1 (canonical-model-theorem ⊥ x) in-logic = - map-raise - ( pr1 - ( pr2 x) - ( unit-trunc-Prop - ( w-ax (subtype-union-right logic (pr1 x) ⊥ in-logic)))) - pr1 (canonical-model-theorem (a →ₘ b) x) in-logic fa = - pr1 - ( canonical-model-theorem b x) - ( weak-modal-logic-subset-complete-theory - ( pr1 x , pr1 (pr2 x)) - ( pr2 (pr2 x)) - ( b) - ( unit-trunc-Prop - ( w-mp - ( w-ax in-logic) - ( w-ax (backward-implication (canonical-model-theorem a x) fa))))) - pr1 (canonical-model-theorem (□ a) x) in-logic y xRy = - forward-implication (canonical-model-theorem a y) (xRy a in-logic) - pr2 (canonical-model-theorem (var n) x) f = - map-inv-raise f - pr2 (canonical-model-theorem ⊥ x) (map-raise ()) - pr2 (canonical-model-theorem (a →ₘ b) x) f = - complete-theory-implication - ( lem) - ( pr1 x , pr1 (pr2 x)) - ( pr2 (pr2 x)) - ( λ a-in-logic → - ( backward-implication - ( canonical-model-theorem b x) - ( f - ( forward-implication (canonical-model-theorem a x) a-in-logic)))) - pr2 (canonical-model-theorem (□ a) x) f = - complete-theory-box - ( lem) - ( pr1 x , pr1 (pr2 x)) - ( pr2 (pr2 x)) - ( a) - ( λ y is-canonical access → - ( backward-implication - ( canonical-model-theorem a (y , is-canonical)) - ( f (y , is-canonical) access))) - - canonical-model-theorem' : - (a : formula i) → - neg-Prop (logic a) ⇔ neg-Prop (canonical-kripke-model ⊨M a) - pr1 (canonical-model-theorem' a) nf in-logic = - apply-universal-property-trunc-Prop - ( lindenbaum x) - ( empty-Prop) - ( λ w → - ( pr2 (pr1 w) - ( weak-modal-logic-mp - { a = a} - ( weak-modal-logic-ax - ( transitive-leq-subtype - ( theory-add-formula (~ a) logic) - ( pr1 (pr1 w)) - ( union-subtype logic (pr1 (pr1 w))) - ( subtype-union-right logic (pr1 (pr1 w))) - ( pr1 (pr2 w)) - ( ~ a) - ( formula-in-add-formula (~ a) logic))) - ( weak-modal-logic-ax - ( subtype-union-right - ( logic) - ( pr1 (pr1 w)) - ( a) - ( backward-implication - ( canonical-model-theorem a - ( pr1 (pr1 w) , pr2 (pr1 w) , pr2 (pr2 w))) - ( in-logic - ( pr1 (pr1 w) , pr2 (pr1 w) , pr2 (pr2 w))))))))) - where - x : L-consistent-theory - pr1 x = theory-add-formula (~ a) logic - pr2 x bot-in-logic = - nf - ( modal-logic-mp - { a = ~~ a} - { b = a} - ( contains-ax-dn _ (a , refl)) - ( subset-double-modal-logic - ( axioms) - ( ~~ a) - ( weak-modal-logic-subset-modal-logic - ( logic) - ( ~~ a) - ( forward-implication - ( deduction-lemma logic - ( transitive-leq-subtype - ( ax-k i) - ( logic) - ( weak-modal-logic logic) - ( axioms-subset-weak-modal-logic logic) - ( contains-ax-k)) - ( transitive-leq-subtype - ( ax-s i) - ( logic) - ( weak-modal-logic logic) - ( axioms-subset-weak-modal-logic logic) - ( contains-ax-s)) - ( ~ a) - ( ⊥)) - ( weak-modal-logic-monotic - ( subtype-union-both - ( logic) + x : modal-theory (l1 ⊔ l2) i + x = raise-subtype l2 (Id-formula-Prop (~ a)) + + not-a-in-x : is-in-subtype x ( ~ a) + not-a-in-x = + subset-equiv-subtypes (Id-formula-Prop (~ a)) x + ( compute-raise-subtype l2 (Id-formula-Prop (~ a))) + ( ~ a) + ( refl) + + is-L-consistent-x : + ¬ (is-in-subtype logic a) → is-L-consistent-theory logic x + is-L-consistent-x a-not-in-logic bot-in-logic = + a-not-in-logic + ( modal-logic-mp is-logic + {a = ~~ a} + {b = a} + ( contains-ax-dn (~~ a →ₘ a) (a , refl)) + ( is-logic (~~ a) + ( subset-weak-modal-logic-closure-modal-logic-closure (~~ a) + ( forward-implication + ( deduction-lemma logic contains-ax-k contains-ax-s ( ~ a) ⊥) + ( weak-modal-logic-closure-monotic + ( subtype-union-both logic x (theory-add-formula (~ a) logic) + ( subtype-union-right (Id-formula-Prop (~ a)) logic) + ( transitive-leq-subtype x (Id-formula-Prop (~ a)) ( theory-add-formula (~ a) logic) - ( theory-add-formula (~ a) logic) - ( subtype-union-right (Id-formula-Prop (~ a)) logic) - ( refl-leq-subtype (theory-add-formula (~ a) logic))) - ( ⊥) - ( bot-in-logic)))))) - pr2 (canonical-model-theorem' a) = - map-neg - ( λ in-logic x → - ( forward-implication - ( canonical-model-theorem a x) - ( logic-subset-L-complete-theory - ( pr1 x , pr1 (pr2 x)) - ( pr2 (pr2 x)) - ( a) - ( in-logic)))) - - canonical-model-theorem'' : - (a : formula i) → logic a ⇔ (canonical-kripke-model ⊨M a) - pr1 (canonical-model-theorem'' a) in-logic x = - forward-implication - ( canonical-model-theorem a x) - ( logic-subset-L-complete-theory - ( pr1 x , pr1 (pr2 x)) - ( pr2 (pr2 x)) - ( a) - ( in-logic)) - pr2 (canonical-model-theorem'' a) in-class-logic with lem (logic a) - ... | inl x = x - ... | inr x = - ex-falso - ( forward-implication (canonical-model-theorem' a) x in-class-logic) - - canonical-model-completness : - {l3 : Level} - (C : model-class (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) i (l1 ⊔ l2) l3) → - is-in-subtype C canonical-kripke-model → - completeness logic C - canonical-model-completness C model-in-class a a-in-class-logic = - backward-implication - ( canonical-model-theorem'' a) - ( a-in-class-logic canonical-kripke-model model-in-class) + ( subtype-union-left (Id-formula-Prop (~ a)) logic) + ( inv-subset-equiv-subtypes (Id-formula-Prop (~ a)) x + ( compute-raise-subtype l2 (Id-formula-Prop (~ a)))))) + ( ⊥) + ( bot-in-logic)))))) + + canonical-model-completness : + {l3 : Level} + (C : model-class (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) i (l1 ⊔ l2) l3) → + is-in-subtype C canonical-kripke-model → + completeness logic C + canonical-model-completness C model-in-class a a-in-class-logic = + backward-implication + ( canonical-model-theorem a) + ( a-in-class-logic canonical-kripke-model model-in-class) ``` diff --git a/src/modal-logic/canonical-model.lagda.md b/src/modal-logic/canonical-model.lagda.md deleted file mode 100644 index bad06098c2..0000000000 --- a/src/modal-logic/canonical-model.lagda.md +++ /dev/null @@ -1,72 +0,0 @@ -# Canonical model theorem - -```agda -module modal-logic.canonical-model where -``` - -
Imports - -```agda -open import foundation.dependent-pair-types -open import foundation.propositional-resizing -open import foundation.propositional-truncations -open import foundation.sets -open import foundation.universe-levels - -open import foundation-core.propositions -open import foundation-core.subtypes - -open import modal-logic.axioms -open import modal-logic.formulas -open import modal-logic.kripke-semantics -open import modal-logic.logic-syntax -open import modal-logic.l-complete-theories -open import modal-logic.l-consistent-theories -open import modal-logic.weak-deduction - -open import order-theory.zorn -``` - -
- -## Idea - -TODO - -## Definition - -```agda -module _ - {l1 l2 : Level} {i : Set l1} - (logic : modal-theory l2 i) - (is-logic : is-weak-modal-logic logic) - (is-cons : is-consistent-modal-logic logic) - (contains-ax-k : ax-k i ⊆ logic) - (contains-ax-s : ax-s i ⊆ logic) - (zorn : Zorn-non-empty (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) l2) - (prop-resize : propositional-resizing (l1 ⊔ l2) (lsuc (l1 ⊔ l2))) - where - - canonical-kripke-model : - kripke-model (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) i (l1 ⊔ l2) - pr1 (pr1 (pr1 canonical-kripke-model)) = - L-complete-theory logic (l1 ⊔ l2) - pr2 (pr1 (pr1 canonical-kripke-model)) = - is-inhabited-L-complete-theory - ( logic) - ( prop-resize) - ( zorn) - ( is-logic) - ( is-cons) - ( contains-ax-k) - ( contains-ax-s) - pr2 (pr1 canonical-kripke-model) x y = - Π-Prop - ( formula i) - ( λ a → - ( hom-Prop - ( modal-theory-L-complete-theory logic x (□ a)) - ( modal-theory-L-complete-theory logic y a))) - pr2 canonical-kripke-model n x = - modal-theory-L-complete-theory logic x (var n) -``` diff --git a/src/modal-logic/canonical-theories.lagda.md b/src/modal-logic/canonical-theories.lagda.md index 722b1564f6..faa842a6e9 100644 --- a/src/modal-logic/canonical-theories.lagda.md +++ b/src/modal-logic/canonical-theories.lagda.md @@ -26,13 +26,14 @@ open import modal-logic.canonical-model-theorem open import modal-logic.completeness open import modal-logic.formulas open import modal-logic.kripke-semantics +open import modal-logic.l-complete-theories open import modal-logic.logic-syntax -open import modal-logic.modal-logic-K -open import modal-logic.modal-logic-S5 +open import modal-logic.modal-logic-k +open import modal-logic.modal-logic-s5 open import modal-logic.soundness open import modal-logic.weak-deduction -open import order-theory.maximal-elements-posets +open import order-theory.zorn ```
@@ -47,118 +48,98 @@ TODO module _ {l1 : Level} (i : Set l1) - (lem : LEM l1) - (zorn : Zorn (lsuc l1) l1 l1) + (lem : LEM l1) -- TODO: not needed + (zorn : Zorn-non-empty (lsuc l1) l1 l1) (prop-resize : propositional-resizing l1 (lsuc l1)) - (axioms : modal-theory l1 i) - (is-cons : is-consistent-modal-logic (modal-logic axioms)) - (contains-K : modal-logic-K i ⊆ modal-logic axioms) + (logic : modal-theory l1 i) + (is-logic : is-modal-logic logic) + (is-cons : is-consistent-modal-logic logic) + (is-normal : is-normal-modal-logic logic) where + private + contains-ax-k : ax-k i ⊆ logic + contains-ax-k = + transitive-leq-subtype + ( ax-k i) + ( modal-logic-K i) + ( logic) + ( is-normal) + ( K-contains-ax-k i) + + contains-ax-s : ax-s i ⊆ logic + contains-ax-s = + transitive-leq-subtype + ( ax-s i) + ( modal-logic-K i) + ( logic) + ( is-normal) + ( K-contains-ax-s i) + + contains-ax-dn : ax-dn i ⊆ logic + contains-ax-dn = + transitive-leq-subtype + ( ax-dn i) + ( modal-logic-K i) + ( logic) + ( is-normal) + ( K-contains-ax-dn i) + is-canonical-ax-m : - ax-m i ⊆ modal-logic axioms → + ax-m i ⊆ logic → is-in-subtype ( reflexive-kripke-class (lsuc l1) l1 i l1) - ( canonical-kripke-model axioms zorn prop-resize is-cons contains-K) + ( canonical-kripke-model logic is-logic is-cons is-normal lem zorn + prop-resize) is-canonical-ax-m ax-m-subset x a box-a-in-x = - weak-modal-logic-subset-complete-theory - ( axioms) - ( zorn) - ( prop-resize) - ( is-cons) - ( contains-K) - ( pr1 x , pr1 (pr2 x)) - ( pr2 (pr2 x)) - ( a) - ( weak-modal-logic-mp - ( weak-modal-logic-ax - ( logic-subset-L-complete-theory - ( axioms) - ( zorn) - ( prop-resize) - ( is-cons) - ( contains-K) - ( pr1 x , pr1 (pr2 x)) - ( pr2 (pr2 x)) - ( □ a →ₘ a) - ( ax-m-subset _ (a , refl)))) - ( weak-modal-logic-ax box-a-in-x)) + weak-modal-logic-mp + ( is-weak-modal-logic-L-complete-theory logic lzero x) + ( subset-logic-L-complete-theory logic lzero x + ( □ a →ₘ a) + ( ax-m-subset (□ a →ₘ a) (a , refl))) + ( box-a-in-x) is-canonical-ax-b : - ax-b i ⊆ modal-logic axioms → + ax-b i ⊆ logic → is-in-subtype ( symmetry-kripke-class (lsuc l1) l1 i l1) - ( canonical-kripke-model axioms zorn prop-resize is-cons contains-K) - is-canonical-ax-b ax-b-subset x y r-xy a box-a-in-y = - lemma-box-diamond - ( axioms) - ( zorn) - ( prop-resize) - ( is-cons) - ( contains-K) - ( lem) - ( pr1 y , pr1 (pr2 y)) - ( pr2 (pr2 y)) - ( pr1 x , pr1 (pr2 x)) - ( pr2 (pr2 x)) - ( λ b → - ( map-universal-property-trunc-Prop - ( pr1 y b) - ( λ { (c , refl , c-in-x) → - ( r-xy (◇ c) - ( weak-modal-logic-subset-complete-theory - ( axioms) - ( zorn) - ( prop-resize) - ( is-cons) - ( contains-K) - ( pr1 x , pr1 (pr2 x)) - ( pr2 (pr2 x)) - ( □ ◇ c) - ( weak-modal-logic-mp - ( weak-modal-logic-ax - ( logic-subset-L-complete-theory - ( axioms) - ( zorn) - ( prop-resize) - ( is-cons) - ( contains-K) - ( pr1 x , pr1 (pr2 x)) - ( pr2 (pr2 x)) - ( c →ₘ □ ◇ c) - ( ax-b-subset _ (c , refl)))) - ( weak-modal-logic-ax c-in-x))))}))) - ( a) - ( box-a-in-y) + ( canonical-kripke-model logic is-logic is-cons is-normal lem zorn + prop-resize) + is-canonical-ax-b ax-b-subset x y xRy a box-a-in-y = + lemma-box-diamond-L-complete logic x + ( contains-ax-k) + ( contains-ax-s) + ( contains-ax-dn) + ( lem) + ( is-logic) + ( is-normal) + ( y) + ( λ b → + ( map-universal-property-trunc-Prop + ( modal-theory-L-complete-theory logic y b) + ( λ { (c , c-in-x , refl) → + ( xRy (◇ c) + ( weak-modal-logic-mp + ( is-weak-modal-logic-L-complete-theory logic lzero x) + ( subset-logic-L-complete-theory logic lzero x (c →ₘ □ ◇ c) + ( ax-b-subset (c →ₘ □ ◇ c) (c , refl))) + ( c-in-x)))}))) + ( a) + ( box-a-in-y) is-canonical-ax-4 : - ax-4 i ⊆ modal-logic axioms → + ax-4 i ⊆ logic → is-in-subtype ( transitivity-kripke-class (lsuc l1) l1 i l1) - ( canonical-kripke-model axioms zorn prop-resize is-cons contains-K) - is-canonical-ax-4 ax-4-subset x y z r-yz r-xy a box-a-in-x = - r-yz a - ( r-xy (□ a) - ( weak-modal-logic-subset-complete-theory - ( axioms) - ( zorn) - ( prop-resize) - ( is-cons) - ( contains-K) - ( pr1 x , pr1 (pr2 x)) - ( pr2 (pr2 x)) - ( □ □ a) - ( weak-modal-logic-mp - ( weak-modal-logic-ax - ( logic-subset-L-complete-theory - ( axioms) - ( zorn) - ( prop-resize) - ( is-cons) - ( contains-K) - ( pr1 x , pr1 (pr2 x)) - ( pr2 (pr2 x)) - ( □ a →ₘ □ □ a) - ( ax-4-subset _ (a , refl)))) - ( weak-modal-logic-ax box-a-in-x)))) + ( canonical-kripke-model logic is-logic is-cons is-normal lem zorn + prop-resize) + is-canonical-ax-4 ax-4-subset x y z yRz xRy a box-a-in-x = + yRz a + ( xRy (□ a) + ( weak-modal-logic-mp + ( is-weak-modal-logic-L-complete-theory logic lzero x) + ( subset-logic-L-complete-theory logic lzero x + ( □ a →ₘ □ □ a) + ( ax-4-subset (□ a →ₘ □ □ a) (a , refl))) + ( box-a-in-x))) ``` diff --git a/src/modal-logic/completeness-K.lagda.md b/src/modal-logic/completeness-K.lagda.md index a75a6584a2..23fc1c1437 100644 --- a/src/modal-logic/completeness-K.lagda.md +++ b/src/modal-logic/completeness-K.lagda.md @@ -1,7 +1,7 @@ # Completeness of K ```agda -module modal-logic.completeness-K where +module modal-logic.completeness-k where ```
Imports @@ -38,11 +38,11 @@ open import modal-logic.completeness open import modal-logic.formulas open import modal-logic.kripke-semantics open import modal-logic.logic-syntax -open import modal-logic.modal-logic-K +open import modal-logic.modal-logic-k open import modal-logic.soundness open import modal-logic.weak-deduction -open import order-theory.maximal-elements-posets +open import order-theory.zorn ```
@@ -54,24 +54,33 @@ TODO ## Definition ```agda --- TODO: levels module _ {l1 : Level} (i : Set l1) (lem : LEM l1) - (zorn : Zorn (lsuc l1) l1 l1) + (zorn : Zorn-non-empty (lsuc l1) l1 l1) (prop-resize : propositional-resizing l1 (lsuc l1)) where completeness-K : completeness (modal-logic-K i) (all-models (lsuc l1) l1 i l1) completeness-K = canonical-model-completness - ( modal-logic-K-axioms i) - ( zorn) - ( prop-resize) + ( modal-logic-K i) + ( is-modal-logic-K i) ( is-consistent-K i) - ( refl-leq-subtype (modal-logic-K i)) + ( is-normal-modal-logic-K i) ( lem) + ( zorn) + ( prop-resize) ( all-models (lsuc l1) l1 i l1) ( star) + -- canonical-model-completness + -- ( modal-logic-K-axioms i) + -- ( zorn) + -- ( prop-resize) + -- ( is-consistent-K i) + -- ( refl-leq-subtype (modal-logic-K i)) + -- ( lem) + -- ( all-models (lsuc l1) l1 i l1) + -- ( star) ``` diff --git a/src/modal-logic/completeness-S5.lagda.md b/src/modal-logic/completeness-S5.lagda.md index fac72921a4..9572c5fbd3 100644 --- a/src/modal-logic/completeness-S5.lagda.md +++ b/src/modal-logic/completeness-S5.lagda.md @@ -1,7 +1,7 @@ # Completeness of S5 ```agda -module modal-logic.completeness-S5 where +module modal-logic.completeness-s5 where ```
Imports @@ -39,11 +39,11 @@ open import modal-logic.completeness open import modal-logic.formulas open import modal-logic.kripke-semantics open import modal-logic.logic-syntax -open import modal-logic.modal-logic-S5 +open import modal-logic.modal-logic-s5 open import modal-logic.soundness open import modal-logic.weak-deduction -open import order-theory.maximal-elements-posets +open import order-theory.zorn ```
@@ -55,12 +55,11 @@ TODO ## Definition ```agda --- TODO: levels module _ {l1 : Level} (i : Set l1) (lem : LEM l1) - (zorn : Zorn (lsuc l1) l1 l1) + (zorn : Zorn-non-empty (lsuc l1) l1 l1) (prop-resize : propositional-resizing l1 (lsuc l1)) where @@ -68,39 +67,45 @@ module _ is-in-subtype ( equivalence-kripke-class (lsuc l1) l1 i l1) ( canonical-kripke-model - ( modal-logic-S5-axioms i) - ( zorn) - ( prop-resize) + ( modal-logic-S5 i) + ( is-modal-logic-S5 i) ( is-consistent-S5 i) - ( modal-logic-K-sub-S5 i)) + ( is-normal-modal-logic-S5 i) + ( lem) + ( zorn) + ( prop-resize)) S5-canonical-model-is-equivalence = triple ( is-canonical-ax-m i lem zorn prop-resize - ( modal-logic-S5-axioms i) + ( modal-logic-S5 i) + ( is-modal-logic-S5 i) ( is-consistent-S5 i) - ( modal-logic-K-sub-S5 i) + ( is-normal-modal-logic-S5 i) ( modal-logic-S5-contains-ax-m i)) ( is-canonical-ax-b i lem zorn prop-resize - ( modal-logic-S5-axioms i) + ( modal-logic-S5 i) + ( is-modal-logic-S5 i) ( is-consistent-S5 i) - ( modal-logic-K-sub-S5 i) + ( is-normal-modal-logic-S5 i) ( modal-logic-S5-contains-ax-b i)) ( is-canonical-ax-4 i lem zorn prop-resize - ( modal-logic-S5-axioms i) + ( modal-logic-S5 i) + ( is-modal-logic-S5 i) ( is-consistent-S5 i) - ( modal-logic-K-sub-S5 i) + ( is-normal-modal-logic-S5 i) ( modal-logic-S5-contains-ax-4 i)) completeness-S5 : completeness (modal-logic-S5 i) (equivalence-kripke-class (lsuc l1) l1 i l1) completeness-S5 = canonical-model-completness - ( modal-logic-S5-axioms i) - ( zorn) - ( prop-resize) + ( modal-logic-S5 i) + ( is-modal-logic-S5 i) ( is-consistent-S5 i) - ( modal-logic-K-sub-S5 i) + ( is-normal-modal-logic-S5 i) ( lem) + ( zorn) + ( prop-resize) ( equivalence-kripke-class (lsuc l1) l1 i l1) ( S5-canonical-model-is-equivalence) ``` diff --git a/src/modal-logic/completeness.lagda.md b/src/modal-logic/completeness.lagda.md index d03a0646e5..a288243eb7 100644 --- a/src/modal-logic/completeness.lagda.md +++ b/src/modal-logic/completeness.lagda.md @@ -34,15 +34,6 @@ TODO ## Definition ```agda --- module _ --- {l1 l2 l3 l4 l5 l6 : Level} --- {i : Set l1} (logic : formulas l2 i) --- {w : UU l3} (C : model-class w l4 i l5 l6) --- where - --- completeness : UU (l1 ⊔ l2 ⊔ l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6) --- completeness = class-modal-logic C ⊆ logic - module _ {l1 l2 l3 l4 l5 l6 : Level} {i : Set l1} where @@ -53,24 +44,3 @@ module _ UU (l1 ⊔ l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6) completeness logic C = class-modal-logic C ⊆ logic ``` - -## Properties - -```agda --- module _ --- {l1 l2 l3 l4 l5 l6 l7 : Level} --- {i : Set l1} (logic : formulas l2 i) --- {w : UU l3} --- (C₁ : model-class w l4 i l5 l6) (C₂ : model-class w l4 i l5 l7) --- where - --- completeness-subclass : --- C₁ ⊆ C₂ → completeness logic C₁ → completeness logic C₂ --- completeness-subclass sub complete = --- transitive-leq-subtype --- ( class-modal-logic C₂) --- ( class-modal-logic C₁) --- ( logic) --- ( complete) --- ( class-modal-logic-monotic C₁ C₂ sub) -``` diff --git a/src/modal-logic/finite-approximability.lagda.md b/src/modal-logic/finite-approximability.lagda.md index f51c202428..bdf97c31a7 100644 --- a/src/modal-logic/finite-approximability.lagda.md +++ b/src/modal-logic/finite-approximability.lagda.md @@ -38,13 +38,13 @@ open import foundation-core.equivalence-relations open import foundation-core.invertible-maps open import modal-logic.completeness -open import modal-logic.completeness-K +open import modal-logic.completeness-k open import modal-logic.formulas open import modal-logic.kripke-models-filtrations open import modal-logic.kripke-semantics open import modal-logic.logic-syntax -open import modal-logic.modal-logic-K open import modal-logic.modal-logic-decision +open import modal-logic.modal-logic-k open import modal-logic.soundness open import modal-logic.weak-deduction diff --git a/src/modal-logic/formulas-deduction.lagda.md b/src/modal-logic/formulas-deduction.lagda.md index 09b8f2bea4..208232b757 100644 --- a/src/modal-logic/formulas-deduction.lagda.md +++ b/src/modal-logic/formulas-deduction.lagda.md @@ -19,7 +19,7 @@ open import foundation-core.identity-types open import modal-logic.axioms open import modal-logic.formulas open import modal-logic.logic-syntax -open import modal-logic.modal-logic-K +open import modal-logic.modal-logic-k open import modal-logic.weak-deduction ``` @@ -35,176 +35,214 @@ TODO module _ {l1 l2 : Level} (i : Set l1) (axioms : modal-theory l2 i) - (is-normal : modal-logic-K i ⊆ modal-logic axioms) + (is-normal : is-normal-modal-logic axioms) where private - contains-ax-k : ax-k i ⊆ modal-logic axioms + contains-ax-k : ax-k i ⊆ axioms contains-ax-k = transitive-leq-subtype ( ax-k i) ( modal-logic-K i) - ( modal-logic axioms) - is-normal + ( axioms) + ( is-normal) ( K-contains-ax-k i) - contains-ax-s : ax-s i ⊆ modal-logic axioms + contains-ax-s : ax-s i ⊆ axioms contains-ax-s = transitive-leq-subtype ( ax-s i) ( modal-logic-K i) - ( modal-logic axioms) - is-normal + ( axioms) + ( is-normal) ( K-contains-ax-s i) - contains-ax-n : ax-n i ⊆ modal-logic axioms + contains-ax-n : ax-n i ⊆ axioms contains-ax-n = transitive-leq-subtype ( ax-n i) ( modal-logic-K i) - ( modal-logic axioms) - is-normal + ( axioms) + ( is-normal) ( K-contains-ax-n i) - contains-ax-dn : ax-dn i ⊆ modal-logic axioms + contains-ax-dn : ax-dn i ⊆ axioms contains-ax-dn = transitive-leq-subtype ( ax-dn i) ( modal-logic-K i) - ( modal-logic axioms) - is-normal + ( axioms) + ( is-normal) ( K-contains-ax-dn i) + weak-modal-logic-const : + (a : formula i) {b : formula i} → + is-in-subtype (weak-modal-logic-closure axioms) b → + is-in-subtype (weak-modal-logic-closure axioms) (a →ₘ b) + weak-modal-logic-const a {b} b-in-logic = + weak-modal-logic-closure-mp + ( weak-modal-logic-closure-ax + ( contains-ax-k (b →ₘ a →ₘ b) (b , a , refl))) + ( b-in-logic) + modal-logic-const : (a : formula i) {b : formula i} → - is-in-subtype (modal-logic axioms) b → - is-in-subtype (modal-logic axioms) (a →ₘ b) + is-in-subtype (modal-logic-closure axioms) b → + is-in-subtype (modal-logic-closure axioms) (a →ₘ b) modal-logic-const a {b} b-in-logic = - modal-logic-mp (contains-ax-k _ (b , a , refl)) (b-in-logic) + modal-logic-closure-mp + ( modal-logic-closure-ax + ( contains-ax-k (b →ₘ a →ₘ b) (b , a , refl))) + ( b-in-logic) + + weak-modal-logic-transitivity : + {a b c : formula i} → + is-in-subtype (weak-modal-logic-closure axioms) (b →ₘ c) → + is-in-subtype (weak-modal-logic-closure axioms) (a →ₘ b) → + is-in-subtype (weak-modal-logic-closure axioms) (a →ₘ c) + weak-modal-logic-transitivity {a} {b} {c} bc ab = + weak-modal-logic-closure-mp + ( weak-modal-logic-closure-mp + ( weak-modal-logic-closure-ax + ( contains-ax-s _ (a , b , c , refl))) + ( weak-modal-logic-const a bc)) + ( ab) modal-logic-transitivity : {a b c : formula i} → - is-in-subtype (modal-logic axioms) (b →ₘ c) → - is-in-subtype (modal-logic axioms) (a →ₘ b) → - is-in-subtype (modal-logic axioms) (a →ₘ c) + is-in-subtype (modal-logic-closure axioms) (b →ₘ c) → + is-in-subtype (modal-logic-closure axioms) (a →ₘ b) → + is-in-subtype (modal-logic-closure axioms) (a →ₘ c) modal-logic-transitivity {a} {b} {c} bc ab = - modal-logic-mp - ( modal-logic-mp - (contains-ax-s _ (a , b , c , refl)) - (modal-logic-const a bc)) + modal-logic-closure-mp + ( modal-logic-closure-mp + ( modal-logic-closure-ax + ( contains-ax-s _ (a , b , c , refl))) + ( modal-logic-const a bc)) ( ab) modal-logic-implication-box' : {a b : formula i} → - is-in-subtype (modal-logic axioms) (a →ₘ b) → - is-in-subtype (modal-logic axioms) (□ a →ₘ □ b) + is-in-subtype (modal-logic-closure axioms) (a →ₘ b) → + is-in-subtype (modal-logic-closure axioms) (□ a →ₘ □ b) modal-logic-implication-box' {a} {b} ab = - modal-logic-mp (contains-ax-n _ (a , b , refl)) (modal-logic-nec ab) + modal-logic-closure-mp + ( modal-logic-closure-ax + ( contains-ax-n (□ (a →ₘ b) →ₘ □ a →ₘ □ b) (a , b , refl))) + ( modal-logic-closure-nec ab) modal-logic-implication-box : {a b : formula i} → - is-in-subtype (modal-logic axioms) (a →ₘ b) → - is-in-subtype (modal-logic axioms) (□ a) → - is-in-subtype (modal-logic axioms) (□ b) - modal-logic-implication-box {a} {b} ab boxa = - modal-logic-mp - ( modal-logic-mp (contains-ax-n _ (a , b , refl)) (modal-logic-nec ab)) - ( boxa) + is-in-subtype (modal-logic-closure axioms) (a →ₘ b) → + is-in-subtype (modal-logic-closure axioms) (□ a) → + is-in-subtype (modal-logic-closure axioms) (□ b) + modal-logic-implication-box {a} {b} ab box-a = + modal-logic-closure-mp + ( modal-logic-closure-mp + ( modal-logic-closure-ax + ( contains-ax-n (□ (a →ₘ b) →ₘ □ a →ₘ □ b) (a , b , refl))) + ( modal-logic-closure-nec ab)) + ( box-a) + + weak-modal-logic-implication-dn : + (a : formula i) → + is-in-subtype (weak-modal-logic-closure axioms) (a →ₘ ~~ a) + weak-modal-logic-implication-dn a = + inv-deduction-ex-falso axioms contains-ax-k contains-ax-s contains-ax-dn a ⊥ modal-logic-implication-dn : - (a : formula i) → is-in-subtype (modal-logic axioms) (a →ₘ ~~ a) + (a : formula i) → is-in-subtype (modal-logic-closure axioms) (a →ₘ ~~ a) modal-logic-implication-dn a = - subset-double-modal-logic axioms - ( a →ₘ (a →ₘ ⊥) →ₘ ⊥) - ( weak-modal-logic-subset-modal-logic - ( modal-logic axioms) - ( a →ₘ (a →ₘ ⊥) →ₘ ⊥) - ( forward-implication - ( deduction-lemma - ( modal-logic axioms) - ( transitive-leq-subtype - ( ax-k i) - ( modal-logic axioms) - ( weak-modal-logic (modal-logic axioms)) - ( axioms-subset-weak-modal-logic (modal-logic axioms)) - ( contains-ax-k)) - ( transitive-leq-subtype - ( ax-s i) - ( modal-logic axioms) - ( weak-modal-logic (modal-logic axioms)) - ( axioms-subset-weak-modal-logic (modal-logic axioms)) - ( contains-ax-s)) - ( a) - ( (a →ₘ ⊥) →ₘ ⊥)) - ( forward-implication - ( deduction-lemma - ( theory-add-formula a (modal-logic axioms)) - ( transitive-leq-subtype - ( ax-k i) - ( theory-add-formula a (modal-logic axioms)) - ( weak-modal-logic (theory-add-formula a (modal-logic axioms))) - ( axioms-subset-weak-modal-logic - ( theory-add-formula a (modal-logic axioms))) - ( transitive-leq-subtype - ( ax-k i) - ( modal-logic axioms) - ( theory-add-formula a (modal-logic axioms)) - ( subtype-union-right - ( Id-formula-Prop a) - ( modal-logic axioms)) - ( contains-ax-k))) - ( transitive-leq-subtype - ( ax-s i) - ( theory-add-formula a (modal-logic axioms)) - ( weak-modal-logic (theory-add-formula a (modal-logic axioms))) - ( axioms-subset-weak-modal-logic - ( theory-add-formula a (modal-logic axioms))) - ( transitive-leq-subtype - ( ax-s i) - ( modal-logic axioms) - ( theory-add-formula a (modal-logic axioms)) - ( subtype-union-right - ( Id-formula-Prop a) - ( modal-logic axioms)) - ( contains-ax-s))) - ( a →ₘ ⊥) - ( ⊥)) - ( weak-modal-logic-mp - ( weak-modal-logic-ax - ( subtype-union-left - ( Id-formula-Prop (~ a)) - ( theory-add-formula a (modal-logic axioms)) - ( ~ a) - ( refl))) - ( weak-modal-logic-ax - ( subtype-union-right - ( Id-formula-Prop (~ a)) - ( theory-add-formula a (modal-logic axioms)) - ( a) - ( subtype-union-left - ( Id-formula-Prop a) - ( modal-logic axioms) - ( a) - ( refl)))))))) + subset-weak-modal-logic-closure-modal-logic-closure (a →ₘ ~~ a) + ( weak-modal-logic-implication-dn a) + + weak-modal-logic-implication-negate : + {a b : formula i} → + is-in-subtype axioms (a →ₘ b) → + is-in-subtype (weak-modal-logic-closure axioms) (~ b →ₘ ~ a) + weak-modal-logic-implication-negate {a} {b} ab = + forward-implication + ( deduction-lemma axioms contains-ax-k contains-ax-s (~ b) (~ a)) + ( forward-implication + ( deduction-lemma + ( theory-add-formula (~ b) axioms) + ( contains-ax-k') + ( contains-ax-s') + ( a) + ( ⊥)) + ( logic-ex-falso + ( theory-add-formula a (theory-add-formula (~ b) axioms)) + ( contains-ax-k'') + ( contains-ax-s'') + ( contains-ax-dn'') + ( b) + ( ⊥) + ( weak-modal-logic-closure-mp + ( weak-modal-logic-closure-ax + ( subset-add-formula a (theory-add-formula (~ b) axioms) + ( a →ₘ b) + ( subset-add-formula (~ b) axioms (a →ₘ b) ab))) + ( weak-modal-logic-closure-ax + ( formula-in-add-formula a (theory-add-formula (~ b) axioms)))) + ( weak-modal-logic-closure-ax + ( subset-add-formula a (theory-add-formula (~ b) axioms) + ( ~ b) + ( formula-in-add-formula (~ b) axioms))))) + where + contains-ax-k' : ax-k i ⊆ theory-add-formula (~ b) axioms + contains-ax-k' = + transitive-subset-add-formula (~ b) axioms (ax-k i) contains-ax-k + + contains-ax-s' : ax-s i ⊆ theory-add-formula (~ b) axioms + contains-ax-s' = + transitive-subset-add-formula (~ b) axioms (ax-s i) contains-ax-s + + contains-ax-k'' : + ax-k i ⊆ theory-add-formula a (theory-add-formula (~ b) axioms) + contains-ax-k'' = + transitive-subset-add-formula a (theory-add-formula (~ b) axioms) + ( ax-k i) + ( contains-ax-k') + + contains-ax-s'' : + ax-s i ⊆ theory-add-formula a (theory-add-formula (~ b) axioms) + contains-ax-s'' = + transitive-subset-add-formula a (theory-add-formula (~ b) axioms) + ( ax-s i) + ( contains-ax-s') + + contains-ax-dn'' : + ax-dn i ⊆ theory-add-formula a (theory-add-formula (~ b) axioms) + contains-ax-dn'' = + transitive-subset-add-formula a (theory-add-formula (~ b) axioms) + ( ax-dn i) + ( transitive-subset-add-formula (~ b) axioms (ax-dn i) contains-ax-dn) + + modal-logic-implication-negate : + {a b : formula i} → + is-in-subtype axioms (a →ₘ b) → + is-in-subtype (modal-logic-closure axioms) (~ b →ₘ ~ a) + modal-logic-implication-negate {a} {b} ab = + subset-weak-modal-logic-closure-modal-logic-closure (~ b →ₘ ~ a) + ( weak-modal-logic-implication-negate ab) modal-logic-diamond-negate : {a : formula i} → - is-in-subtype (modal-logic axioms) (◇ ~ a) → - is-in-subtype (modal-logic axioms) (~ □ a) + is-in-subtype (modal-logic-closure axioms) (◇ ~ a) → + is-in-subtype (modal-logic-closure axioms) (~ □ a) modal-logic-diamond-negate {a} dia-a = modal-logic-transitivity ( dia-a) ( modal-logic-implication-box' (modal-logic-implication-dn a)) -module _ - {l1 l2 : Level} (i : Set l1) - (axioms : modal-theory l2 i) - (is-normal : modal-logic-K i ⊆ weak-modal-logic axioms) - where - - postulate - weak-modal-logic-diamond-negate : - {a : formula i} → - is-in-subtype (weak-modal-logic axioms) (◇ ~ a) → - is-in-subtype (weak-modal-logic axioms) (~ □ a) + modal-logic-diamond-negate-implication : + {a : formula i} → + is-modal-logic axioms → + is-in-subtype axioms (◇ ~ a →ₘ ~ □ a) + modal-logic-diamond-negate-implication {a} is-logic = + is-logic (◇ ~ a →ₘ ~ □ a) + ( modal-logic-implication-negate + ( is-logic (□ a →ₘ □ ~~ a) + ( modal-logic-implication-box' + ( modal-logic-implication-dn a)))) ``` diff --git a/src/modal-logic/l-complete-theories.lagda.md b/src/modal-logic/l-complete-theories.lagda.md index 5dfa78384a..f76510e8a4 100644 --- a/src/modal-logic/l-complete-theories.lagda.md +++ b/src/modal-logic/l-complete-theories.lagda.md @@ -16,9 +16,9 @@ open import foundation.existential-quantification open import foundation.inhabited-types open import foundation.law-of-excluded-middle open import foundation.logical-equivalences -open import foundation.subtypes -open import foundation.propositional-truncations open import foundation.propositional-resizing +open import foundation.propositional-truncations +open import foundation.subtypes open import foundation.transport-along-identifications open import foundation.unions-subtypes open import foundation.universe-levels @@ -30,19 +30,20 @@ open import foundation-core.identity-types open import foundation-core.negation open import foundation-core.propositions open import foundation-core.sets -open import foundation-core.subtypes open import lists.lists open import lists.lists-subtypes open import modal-logic.axioms open import modal-logic.formulas -open import modal-logic.logic-syntax +open import modal-logic.formulas-deduction open import modal-logic.l-consistent-theories +open import modal-logic.logic-syntax +open import modal-logic.modal-logic-k open import modal-logic.weak-deduction -open import order-theory.maximal-elements-posets open import order-theory.chains-posets +open import order-theory.maximal-elements-posets open import order-theory.posets open import order-theory.subposets open import order-theory.subtypes-leq-posets @@ -99,6 +100,13 @@ module _ is-L-consistent-theory-modal-theory-L-consistent-theory logic ∘ L-consistent-theory-L-complete-theory + is-consistent-modal-theory-L-complete-theory : + {l3 : Level} (theory : L-complete-theory l3) → + is-consistent-modal-logic (modal-theory-L-complete-theory theory) + is-consistent-modal-theory-L-complete-theory = + is-consistent-modal-theory-L-consistent-theory logic ∘ + L-consistent-theory-L-complete-theory + module _ {l3 : Level} (((theory , is-cons) , is-comp) : L-complete-theory (l1 ⊔ l2 ⊔ l3)) @@ -138,7 +146,7 @@ module _ ( logic) ( logic ∪ theory) ( weak-modal-logic-closure (logic ∪ theory)) - ( subset-axioms-weak-modal-logic) + ( subset-axioms-weak-modal-logic-closure) ( subtype-union-left logic theory)) ( refl-leq-subtype (weak-modal-logic-closure (logic ∪ theory))))) ( ⊥) @@ -160,7 +168,7 @@ module _ ( theory) ( logic ∪ theory) ( weak-modal-logic-closure (logic ∪ theory)) - ( subset-axioms-weak-modal-logic) + ( subset-axioms-weak-modal-logic-closure) ( subtype-union-right logic theory))) subset-union-L-consistent : @@ -180,7 +188,7 @@ module _ ( union-L-consistent (theory , is-cons))) ( theory) ( subset-union-L-consistent) - ( subset-axioms-weak-modal-logic) + ( subset-axioms-weak-modal-logic-closure) subset-logic-L-complete-theory : logic ⊆ theory subset-logic-L-complete-theory = @@ -294,15 +302,43 @@ module _ ( is-L-consistent-add-formula-not-in-logic not-in-logic)) ( formula-in-add-formula (~ a) theory) - is-disjuctive-L-complete-theory : - LEM (l1 ⊔ l2) → - is-disjuctive-modal-theory theory - is-disjuctive-L-complete-theory lem a with lem (theory a) - ... | inl a-in-logic = inl a-in-logic - ... | inr a-not-in-logic = - inr - ( contains-negation-not-contains-formula-L-complete-theory - ( a-not-in-logic)) + module _ + (lem : LEM (l1 ⊔ l2)) + where + + is-disjuctive-L-complete-theory : + is-disjuctive-modal-theory theory + is-disjuctive-L-complete-theory a with lem (theory a) + ... | inl a-in-logic = inl a-in-logic + ... | inr a-not-in-logic = + inr + ( contains-negation-not-contains-formula-L-complete-theory + ( a-not-in-logic)) + + -- TODO: move from module + lemma-box-diamond-L-complete : + is-modal-logic logic → + is-normal-modal-logic logic → + (((theory' , _) , _) : L-complete-theory (l1 ⊔ l2)) → + diamond-modal-theory theory ⊆ theory' → + unbox-modal-theory theory' ⊆ theory + lemma-box-diamond-L-complete is-logic is-normal x leq a box-a-in-x + with is-disjuctive-L-complete-theory a + ... | inl a-in-t = a-in-t + ... | inr not-a-in-t = + ex-falso + ( is-consistent-modal-theory-L-complete-theory x + ( weak-modal-logic-mp + ( is-weak-modal-logic-L-complete-theory lzero x) + { a = □ a} + ( weak-modal-logic-mp + ( is-weak-modal-logic-L-complete-theory lzero x) + { a = ◇ ~ a} + ( subset-logic-L-complete-theory lzero x (◇ ~ a →ₘ ~ □ a) + ( modal-logic-diamond-negate-implication i logic is-normal + ( is-logic))) + ( leq (◇ ~ a) (intro-∃ (~ a) (not-a-in-t , refl)))) + ( box-a-in-x))) is-inhabited-L-complete-exists-complete-L-consistent-theory : {l3 : Level} → @@ -340,7 +376,7 @@ module _ ( modal-theory-chain-element x) ( logic ∪ modal-theory-chain-element x) ( L-union x) - ( subset-axioms-weak-modal-logic) + ( subset-axioms-weak-modal-logic-closure) ( subtype-union-right logic (modal-theory-chain-element x)) leq-L-union : @@ -508,7 +544,7 @@ module _ ( leq-logic)) ( a) ( weak-modal-logic-closure-monotic leq-lists a - ( is-in-weak-deduction-closure-weak-deduction + ( is-in-weak-modal-logic-closure-weak-deduction ( is-assumptions-list-assumptions-weak-deduction ( d))))))))) @@ -674,6 +710,6 @@ module _ ( subtype-union-both theory' theory theory' ( refl-leq-subtype theory') ( leq))) - ( is-consistent-modal-theory-L-consistent-theory logic₂ + ( is-consistent-closure-L-consistent-theory logic₂ ( theory' , is-cons'))) ``` diff --git a/src/modal-logic/l-consistent-theories.lagda.md b/src/modal-logic/l-consistent-theories.lagda.md index 3d07b7beaf..73ed7e888f 100644 --- a/src/modal-logic/l-consistent-theories.lagda.md +++ b/src/modal-logic/l-consistent-theories.lagda.md @@ -75,11 +75,11 @@ module _ is-L-consistent-theory-modal-theory-L-consistent-theory = is-in-subtype-inclusion-subtype is-L-consistent-theory-Prop - is-consistent-modal-theory-L-consistent-theory : + is-consistent-closure-L-consistent-theory : {l3 : Level} (theory : L-consistent-theory l3) → is-consistent-modal-logic ( weak-modal-logic-closure (modal-theory-L-consistent-theory theory)) - is-consistent-modal-theory-L-consistent-theory theory = + is-consistent-closure-L-consistent-theory theory = is-consistent-modal-logic-antimonotic ( weak-modal-logic-closure (modal-theory-L-consistent-theory theory)) ( weak-modal-logic-closure @@ -88,6 +88,16 @@ module _ ( subtype-union-right logic (modal-theory-L-consistent-theory theory))) ( is-L-consistent-theory-modal-theory-L-consistent-theory theory) + is-consistent-modal-theory-L-consistent-theory : + {l3 : Level} (theory : L-consistent-theory l3) → + is-consistent-modal-logic (modal-theory-L-consistent-theory theory) + is-consistent-modal-theory-L-consistent-theory theory = + is-consistent-modal-logic-antimonotic + ( modal-theory-L-consistent-theory theory) + ( weak-modal-logic-closure (modal-theory-L-consistent-theory theory)) + ( subset-axioms-weak-modal-logic-closure) + ( is-consistent-closure-L-consistent-theory theory) + is-L-consistent-antimonotic : {l3 l4 : Level} (theory₁ : modal-theory l3 i) → diff --git a/src/modal-logic/lindenbaum.lagda.md b/src/modal-logic/lindenbaum.lagda.md index 7c7e475cbd..e8021bc781 100644 --- a/src/modal-logic/lindenbaum.lagda.md +++ b/src/modal-logic/lindenbaum.lagda.md @@ -20,9 +20,9 @@ open import foundation-core.sets open import foundation-core.subtypes open import modal-logic.axioms -open import modal-logic.logic-syntax -open import modal-logic.l-consistent-theories open import modal-logic.l-complete-theories +open import modal-logic.l-consistent-theories +open import modal-logic.logic-syntax open import modal-logic.weak-deduction open import order-theory.zorn @@ -44,7 +44,7 @@ module _ (contains-ax-s : ax-s i ⊆ logic) (zorn : Zorn-non-empty (lsuc (l1 ⊔ l2 ⊔ l3)) (l1 ⊔ l2 ⊔ l3) l3) (prop-resize : propositional-resizing (l1 ⊔ l2 ⊔ l3) (lsuc (l1 ⊔ l2 ⊔ l3))) - (x@(theory , is-cons) : L-consistent-theory logic l3) + (x@(theory , is-cons) : L-consistent-theory logic (l1 ⊔ l2 ⊔ l3)) where lindenbaum : @@ -74,7 +74,7 @@ module _ ( logic) ( logic ∪ theory) ( weak-modal-logic-closure (logic ∪ theory)) - ( subset-axioms-weak-modal-logic) + ( subset-axioms-weak-modal-logic-closure) ( subtype-union-left logic theory) contains-ax-k' : ax-k i ⊆ L @@ -130,6 +130,6 @@ module _ ( theory) ( logic ∪ theory) ( L) - ( subset-axioms-weak-modal-logic) + ( subset-axioms-weak-modal-logic-closure) ( subtype-union-right logic theory))) ``` diff --git a/src/modal-logic/logic-syntax.lagda.md b/src/modal-logic/logic-syntax.lagda.md index 104269081e..d552e4d7f9 100644 --- a/src/modal-logic/logic-syntax.lagda.md +++ b/src/modal-logic/logic-syntax.lagda.md @@ -11,6 +11,7 @@ open import foundation.conjunction open import foundation.coproduct-types open import foundation.dependent-pair-types open import foundation.empty-types +open import foundation.existential-quantification open import foundation.function-types open import foundation.identity-types open import foundation.negation @@ -22,6 +23,8 @@ open import foundation.unions-subtypes open import foundation.unit-type open import foundation.universe-levels +open import foundation-core.cartesian-product-types + open import modal-logic.formulas ``` @@ -41,10 +44,6 @@ module _ modal-theory : UU (l1 ⊔ lsuc l2) modal-theory = subtype l2 (formula i) - -- TODO: remove - formulas : UU (l1 ⊔ lsuc l2) - formulas = modal-theory - module _ {l1 l2 : Level} {i : Set l1} where @@ -57,16 +56,21 @@ module _ nec : {a : formula i} → axioms ⊢ a → axioms ⊢ □ a -- TODO: rename to modal-logic-closure - modal-logic : modal-theory l2 i → modal-theory (l1 ⊔ l2) i - modal-logic axioms a = trunc-Prop (axioms ⊢ a) + modal-logic-closure : modal-theory l2 i → modal-theory (l1 ⊔ l2) i + modal-logic-closure axioms a = trunc-Prop (axioms ⊢ a) is-modal-logic-Prop : modal-theory l2 i → Prop (l1 ⊔ l2) is-modal-logic-Prop theory = - leq-prop-subtype (modal-logic theory) theory + leq-prop-subtype (modal-logic-closure theory) theory is-modal-logic : modal-theory l2 i → UU (l1 ⊔ l2) is-modal-logic = type-Prop ∘ is-modal-logic-Prop + is-in-modal-logic-closure-deduction : + {axioms : modal-theory l2 i} {a : formula i} → + axioms ⊢ a → is-in-subtype (modal-logic-closure axioms) a + is-in-modal-logic-closure-deduction = unit-trunc-Prop + is-contradictory-modal-logic-Prop : modal-theory l2 i → Prop l2 is-contradictory-modal-logic-Prop logic = logic ⊥ @@ -99,72 +103,89 @@ module _ is-cons ∘ is-contradictory-modal-logic-monotic ax₁ ax₂ leq module _ - {l1 l2 : Level} {i : Set l1} {axioms : formulas l2 i} + {l1 l2 : Level} {i : Set l1} {axioms : modal-theory l2 i} where - modal-logic-ax : + modal-logic-closure-ax : {a : formula i} → is-in-subtype axioms a → - is-in-subtype (modal-logic axioms) a - modal-logic-ax = unit-trunc-Prop ∘ ax + is-in-subtype (modal-logic-closure axioms) a + modal-logic-closure-ax = unit-trunc-Prop ∘ ax - modal-logic-mp : + modal-logic-closure-mp : {a b : formula i} → - is-in-subtype (modal-logic axioms) (a →ₘ b) → - is-in-subtype (modal-logic axioms) a → - is-in-subtype (modal-logic axioms) b - modal-logic-mp {a} {b} tdab tda = + is-in-subtype (modal-logic-closure axioms) (a →ₘ b) → + is-in-subtype (modal-logic-closure axioms) a → + is-in-subtype (modal-logic-closure axioms) b + modal-logic-closure-mp {a} {b} tdab tda = apply-twice-universal-property-trunc-Prop tdab tda - ( modal-logic axioms b) + ( modal-logic-closure axioms b) ( λ dab da → unit-trunc-Prop (mp dab da)) - modal-logic-nec : + modal-logic-closure-nec : {a : formula i} → - is-in-subtype (modal-logic axioms) a → - is-in-subtype (modal-logic axioms) (□ a) - modal-logic-nec {a} = + is-in-subtype (modal-logic-closure axioms) a → + is-in-subtype (modal-logic-closure axioms) (□ a) + modal-logic-closure-nec {a} = map-universal-property-trunc-Prop - ( modal-logic axioms (□ a)) + ( modal-logic-closure axioms (□ a)) ( λ da → unit-trunc-Prop (nec da)) +module _ + {l1 l2 : Level} {i : Set l1} + {logic : modal-theory l2 i} (is-logic : is-modal-logic logic) + where + + modal-logic-mp : + {a b : formula i} → + is-in-subtype logic (a →ₘ b) → + is-in-subtype logic a → + is-in-subtype logic b + modal-logic-mp {a} {b} dab da = + is-logic b + ( modal-logic-closure-mp + ( modal-logic-closure-ax dab) + ( modal-logic-closure-ax da)) + + modal-logic-nec : + {a : formula i} → + is-in-subtype logic a → + is-in-subtype logic (□ a) + modal-logic-nec {a} d = + is-logic (□ a) (modal-logic-closure-nec (modal-logic-closure-ax d)) + module _ {l1 : Level} {i : Set l1} where axioms-subset-modal-logic : - {l2 : Level} (axioms : formulas l2 i) → axioms ⊆ modal-logic axioms + {l2 : Level} (axioms : modal-theory l2 i) → + axioms ⊆ modal-logic-closure axioms axioms-subset-modal-logic _ a H = unit-trunc-Prop (ax H) modal-logic-closed : - {l2 : Level} {axioms : formulas l2 i} {a : formula i} → - modal-logic axioms ⊢ a → - is-in-subtype (modal-logic axioms) a + {l2 : Level} {axioms : modal-theory l2 i} {a : formula i} → + modal-logic-closure axioms ⊢ a → + is-in-subtype (modal-logic-closure axioms) a modal-logic-closed (ax x) = x modal-logic-closed (mp dab da) = - modal-logic-mp (modal-logic-closed dab) (modal-logic-closed da) + modal-logic-closure-mp (modal-logic-closed dab) (modal-logic-closed da) modal-logic-closed (nec d) = - modal-logic-nec (modal-logic-closed d) + modal-logic-closure-nec (modal-logic-closed d) + -- TODO: refactor subset-double-modal-logic : {l2 : Level} - (axioms : formulas l2 i) → - modal-logic (modal-logic axioms) ⊆ modal-logic axioms + (axioms : modal-theory l2 i) → + is-modal-logic (modal-logic-closure axioms) subset-double-modal-logic axioms a = map-universal-property-trunc-Prop - ( modal-logic axioms a) + ( modal-logic-closure axioms a) ( modal-logic-closed) - modal-logic-idempotent : - {l2 : Level} - (axioms : formulas l2 i) → - modal-logic axioms = modal-logic (modal-logic axioms) - modal-logic-idempotent axioms = - antisymmetric-leq-subtype _ _ - ( axioms-subset-modal-logic (modal-logic axioms)) - ( subset-double-modal-logic axioms) - module _ - {l1 l2 l3 : Level} {i : Set l1} {ax₁ : formulas l2 i} {ax₂ : formulas l3 i} + {l1 l2 l3 : Level} {i : Set l1} + {ax₁ : modal-theory l2 i} {ax₂ : modal-theory l3 i} (leq : ax₁ ⊆ ax₂) where @@ -174,23 +195,25 @@ module _ mp (deduction-monotic dab) (deduction-monotic da) deduction-monotic (nec d) = nec (deduction-monotic d) - modal-logic-monotic : modal-logic ax₁ ⊆ modal-logic ax₂ + modal-logic-monotic : modal-logic-closure ax₁ ⊆ modal-logic-closure ax₂ modal-logic-monotic a = map-universal-property-trunc-Prop - ( modal-logic ax₂ a) + ( modal-logic-closure ax₂ a) ( unit-trunc-Prop ∘ deduction-monotic) module _ - {l1 l2 l3 : Level} {i : Set l1} {ax₁ : formulas l2 i} {ax₂ : formulas l3 i} + {l1 l2 l3 : Level} {i : Set l1} + {ax₁ : modal-theory l2 i} {ax₂ : modal-theory l3 i} where subset-modal-logic-if-subset-axioms : - ax₁ ⊆ modal-logic ax₂ → modal-logic ax₁ ⊆ modal-logic ax₂ + ax₁ ⊆ modal-logic-closure ax₂ → + modal-logic-closure ax₁ ⊆ modal-logic-closure ax₂ subset-modal-logic-if-subset-axioms leq = transitive-leq-subtype - ( modal-logic ax₁) - ( modal-logic (modal-logic ax₂)) - ( modal-logic ax₂) + ( modal-logic-closure ax₁) + ( modal-logic-closure (modal-logic-closure ax₂)) + ( modal-logic-closure ax₂) ( subset-double-modal-logic ax₂) ( modal-logic-monotic leq) @@ -199,7 +222,7 @@ module _ where -- TODO: make Id-formula to be a function for 1 element modal-theory - theory-add-formula : formulas (l1 ⊔ l2) i + theory-add-formula : modal-theory (l1 ⊔ l2) i theory-add-formula = (Id-formula-Prop a) ∪ theory formula-in-add-formula : is-in-subtype theory-add-formula a @@ -208,6 +231,15 @@ module _ subset-add-formula : theory ⊆ theory-add-formula subset-add-formula = subtype-union-right (Id-formula-Prop a) theory + transitive-subset-add-formula : + {l3 : Level} (theory' : modal-theory l3 i) → + theory' ⊆ theory → + theory' ⊆ theory-add-formula + transitive-subset-add-formula theory' leq = + transitive-leq-subtype theory' theory theory-add-formula + ( subset-add-formula) + ( leq) + elim-theory-add-formula : {l3 : Level} (P : formula i → Prop l3) → type-Prop (P a) → @@ -215,11 +247,11 @@ module _ (x : formula i) → is-in-subtype theory-add-formula x → type-Prop (P x) elim-theory-add-formula P H-a H-rest = elim-union-subtype (Id-formula-Prop a) theory P - ( λ { .a refl → H-a }) + ( λ { .a refl → H-a}) ( H-rest) subset-theory-add-formula : - {l3 : Level} (theory' : formulas l3 i) → + {l3 : Level} (theory' : modal-theory l3 i) → is-in-subtype theory' a → theory ⊆ theory' → theory-add-formula ⊆ theory' @@ -228,7 +260,18 @@ module _ ( Id-formula-Prop a) ( theory) ( theory') - ( λ { .a refl → a-in }) + ( λ { .a refl → a-in}) + +module _ + {l1 l2 : Level} {i : Set l1} + where + + unbox-modal-theory : modal-theory l2 i → modal-theory l2 i + unbox-modal-theory theory a = theory (□ a) + + diamond-modal-theory : modal-theory l2 i → modal-theory (l1 ⊔ l2) i + diamond-modal-theory theory a = + ∃-Prop (formula i) (λ b → is-in-subtype theory b × (a = ◇ b)) module _ {l1 : Level} {i : Set l1} @@ -242,8 +285,8 @@ module _ theory-add-formula-union-right : (a : formula i) {l2 l3 : Level} - (theory : formulas l2 i) - (theory' : formulas l3 i) → + (theory : modal-theory l2 i) + (theory' : modal-theory l3 i) → theory ∪ theory-add-formula a theory' ⊆ theory-add-formula a (theory ∪ theory') theory-add-formula-union-right a theory theory' = @@ -252,8 +295,8 @@ module _ inv-theory-add-formula-union-right : (a : formula i) {l2 l3 : Level} - (theory : formulas l2 i) - (theory' : formulas l3 i) → + (theory : modal-theory l2 i) + (theory' : modal-theory l3 i) → theory-add-formula a (theory ∪ theory') ⊆ theory ∪ theory-add-formula a theory' inv-theory-add-formula-union-right a theory theory' = diff --git a/src/modal-logic/modal-logic-K.lagda.md b/src/modal-logic/modal-logic-K.lagda.md index d176c1e385..dfa7284117 100644 --- a/src/modal-logic/modal-logic-K.lagda.md +++ b/src/modal-logic/modal-logic-K.lagda.md @@ -1,7 +1,7 @@ # Modal logic K ```agda -module modal-logic.modal-logic-K where +module modal-logic.modal-logic-k where ```
Imports @@ -45,11 +45,14 @@ module _ {l : Level} (i : Set l) where - modal-logic-K-axioms : formulas l i + modal-logic-K-axioms : modal-theory l i modal-logic-K-axioms = ax-k i ∪ ax-s i ∪ ax-n i ∪ ax-dn i - modal-logic-K : formulas l i - modal-logic-K = modal-logic modal-logic-K-axioms + modal-logic-K : modal-theory l i + modal-logic-K = modal-logic-closure modal-logic-K-axioms + + is-modal-logic-K : is-modal-logic modal-logic-K + is-modal-logic-K = subset-double-modal-logic modal-logic-K-axioms K-axioms-contains-ax-k : ax-k i ⊆ modal-logic-K-axioms K-axioms-contains-ax-k = @@ -128,6 +131,14 @@ module _ ( axioms-subset-modal-logic modal-logic-K-axioms) ( K-axioms-contains-ax-dn) +is-normal-modal-logic : + {l1 l2 : Level} {i : Set l1} → modal-theory l2 i → UU (l1 ⊔ l2) +is-normal-modal-logic {i = i} logic = modal-logic-K i ⊆ logic + +is-normal-modal-logic-K : + {l1 : Level} (i : Set l1) → is-normal-modal-logic (modal-logic-K i) +is-normal-modal-logic-K i = refl-leq-subtype (modal-logic-K i) + module _ {l1 l2 l3 l4 : Level} (i : Set l1) where diff --git a/src/modal-logic/modal-logic-S5.lagda.md b/src/modal-logic/modal-logic-S5.lagda.md index be025827a3..f3e74e3a91 100644 --- a/src/modal-logic/modal-logic-S5.lagda.md +++ b/src/modal-logic/modal-logic-S5.lagda.md @@ -1,7 +1,7 @@ # Modal logic S5 ```agda -module modal-logic.modal-logic-S5 where +module modal-logic.modal-logic-s5 where ```
Imports @@ -25,7 +25,7 @@ open import modal-logic.axioms open import modal-logic.formulas open import modal-logic.kripke-semantics open import modal-logic.logic-syntax -open import modal-logic.modal-logic-K +open import modal-logic.modal-logic-k open import modal-logic.soundness open import univalent-combinatorics.finite-types @@ -44,14 +44,17 @@ module _ {l : Level} (i : Set l) where - modal-logic-S5-axioms : formulas l i + modal-logic-S5-axioms : modal-theory l i modal-logic-S5-axioms = modal-logic-K-axioms i ∪ (ax-m i ∪ (ax-b i ∪ ax-4 i)) - modal-logic-S5 : formulas l i - modal-logic-S5 = modal-logic modal-logic-S5-axioms + modal-logic-S5 : modal-theory l i + modal-logic-S5 = modal-logic-closure modal-logic-S5-axioms - modal-logic-K-sub-S5 : modal-logic-K i ⊆ modal-logic-S5 - modal-logic-K-sub-S5 = + is-modal-logic-S5 : is-modal-logic modal-logic-S5 + is-modal-logic-S5 = subset-double-modal-logic modal-logic-S5-axioms + + is-normal-modal-logic-S5 : is-normal-modal-logic modal-logic-S5 + is-normal-modal-logic-S5 = modal-logic-monotic ( subtype-union-left ( modal-logic-K-axioms i) diff --git a/src/modal-logic/soundness.lagda.md b/src/modal-logic/soundness.lagda.md index 06b7e07b89..17e92d76de 100644 --- a/src/modal-logic/soundness.lagda.md +++ b/src/modal-logic/soundness.lagda.md @@ -47,7 +47,7 @@ soundness logic C = logic ⊆ class-modal-logic C ```agda module _ {l1 l2 l3 l4 l5 l6 : Level} - {i : Set l1} {axioms : formulas l2 i} + {i : Set l1} {axioms : modal-theory l2 i} (C : model-class l3 l4 i l5 l6) where @@ -59,13 +59,14 @@ module _ soundness-axioms H (nec d) M in-class _ y _ = soundness-axioms H d M in-class y - soundness-modal-logic : soundness axioms C → soundness (modal-logic axioms) C + soundness-modal-logic : + soundness axioms C → soundness (modal-logic-closure axioms) C soundness-modal-logic H a = map-universal-property-trunc-Prop (C ⊨C a) (soundness-axioms H) module _ {l1 l2 l3 l4 l5 l6 l7 : Level} - {i : Set l1} (logic : formulas l2 i) + {i : Set l1} (logic : modal-theory l2 i) (C₁ : model-class l3 l4 i l5 l6) (C₂ : model-class l3 l4 i l5 l7) where @@ -79,7 +80,7 @@ module _ module _ {l1 l2 l3 l4 l5 l6 l7 l8 : Level} - {i : Set l1} (theory₁ : formulas l2 i) (theory₂ : formulas l3 i) + {i : Set l1} (theory₁ : modal-theory l2 i) (theory₂ : modal-theory l3 i) (C₁ : model-class l4 l5 i l6 l7) (C₂ : model-class l4 l5 i l6 l8) (sound₁ : soundness theory₁ C₁) (sound₂ : soundness theory₂ C₂) where @@ -104,14 +105,14 @@ module _ soundness-modal-logic-union : soundness - (modal-logic (union-subtype theory₁ theory₂)) + (modal-logic-closure (union-subtype theory₁ theory₂)) (intersection-subtype C₁ C₂) soundness-modal-logic-union = soundness-modal-logic (intersection-subtype C₁ C₂) soundness-union module _ {l1 l2 l3 l4 l5 l6 l7 : Level} - {i : Set l1} (ax₁ : formulas l2 i) (ax₂ : formulas l3 i) + {i : Set l1} (ax₁ : modal-theory l2 i) (ax₂ : modal-theory l3 i) where soundness-union-subclass-left-sublevels : @@ -145,11 +146,11 @@ module _ (C₁ : model-class l4 l5 i l6 (l7 ⊔ l8)) (C₂ : model-class l4 l5 i l6 l7) (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) → C₁ ⊆ C₂ → - soundness (modal-logic (union-subtype ax₁ ax₂)) C₁ + soundness (modal-logic-closure (union-subtype ax₁ ax₂)) C₁ soundness-modal-logic-union-subclass-left-sublevels l8 C₁ C₂ sound₁ sound₂ C₁-sub-C₂ = tr - ( soundness (modal-logic (union-subtype ax₁ ax₂))) + ( soundness (modal-logic-closure (union-subtype ax₁ ax₂))) ( intersection-subtype-left-sublevels l8 C₁ C₂ C₁-sub-C₂) ( soundness-modal-logic-union ax₁ ax₂ C₁ C₂ sound₁ sound₂) @@ -158,11 +159,11 @@ module _ (C₁ : model-class l4 l5 i l6 l7) (C₂ : model-class l4 l5 i l6 (l7 ⊔ l8)) (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) → C₂ ⊆ C₁ → - soundness (modal-logic (union-subtype ax₁ ax₂)) C₂ + soundness (modal-logic-closure (union-subtype ax₁ ax₂)) C₂ soundness-modal-logic-union-subclass-right-sublevels l8 C₁ C₂ sound₁ sound₂ C₂-sub-C₁ = tr - ( soundness (modal-logic (union-subtype ax₁ ax₂))) + ( soundness (modal-logic-closure (union-subtype ax₁ ax₂))) ( intersection-subtype-right-sublevels l8 C₁ C₂ C₂-sub-C₁) ( soundness-modal-logic-union ax₁ ax₂ C₁ C₂ sound₁ sound₂) @@ -170,7 +171,7 @@ module _ (C₁ : model-class l4 l5 i l6 l7) (C₂ : model-class l4 l5 i l6 l7) (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) → C₁ ⊆ C₂ → - soundness (modal-logic (union-subtype ax₁ ax₂)) C₁ + soundness (modal-logic-closure (union-subtype ax₁ ax₂)) C₁ soundness-modal-logic-union-subclass-left = soundness-modal-logic-union-subclass-left-sublevels lzero @@ -178,13 +179,13 @@ module _ (C₁ : model-class l4 l5 i l6 l7) (C₂ : model-class l4 l5 i l6 l7) (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) → C₂ ⊆ C₁ → - soundness (modal-logic (union-subtype ax₁ ax₂)) C₂ + soundness (modal-logic-closure (union-subtype ax₁ ax₂)) C₂ soundness-modal-logic-union-subclass-right = soundness-modal-logic-union-subclass-right-sublevels lzero module _ {l1 l2 l3 l4 l5 l6 l7 : Level} - {i : Set l1} (ax₁ : formulas l2 i) (ax₂ : formulas l3 i) + {i : Set l1} (ax₁ : modal-theory l2 i) (ax₂ : modal-theory l3 i) (C : model-class l4 l5 i l6 l7) (sound₁ : soundness ax₁ C) (sound₂ : soundness ax₂ C) where @@ -197,7 +198,7 @@ module _ ( soundness-union ax₁ ax₂ C C sound₁ sound₂) soundness-modal-logic-union-same-class : - soundness (modal-logic (union-subtype ax₁ ax₂)) C + soundness (modal-logic-closure (union-subtype ax₁ ax₂)) C soundness-modal-logic-union-same-class = soundness-modal-logic C soundness-union-same-class ``` diff --git a/src/modal-logic/weak-deduction.lagda.md b/src/modal-logic/weak-deduction.lagda.md index a8d4553912..56436086db 100644 --- a/src/modal-logic/weak-deduction.lagda.md +++ b/src/modal-logic/weak-deduction.lagda.md @@ -8,8 +8,8 @@ module modal-logic.weak-deduction where ```agda open import foundation.cartesian-product-types -open import foundation.coproduct-types open import foundation.conjunction +open import foundation.coproduct-types open import foundation.dependent-pair-types open import foundation.disjunction open import foundation.empty-types @@ -91,37 +91,6 @@ module _ weak-deduction-mp (dab , is-w-dab) (da , is-w-da) = mp dab da , is-w-dab , is-w-da - -- ind-weak-deduction : - -- {l : Level} - -- {P : {axioms : modal-theory l2 i} {a : formula i} → axioms ⊢w a → UU l} → - -- ( {axioms : modal-theory l2 i} {a : formula i} - -- (in-axioms : is-in-subtype axioms a) → - -- P {axioms} (weak-deduction-ax in-axioms)) → - -- ( {axioms : modal-theory l2 i} {a b : formula i} - -- (dab : axioms ⊢w a →ₘ b) (da : axioms ⊢w a) → - -- P dab → P da → P (weak-deduction-mp dab da)) → - -- {axioms : modal-theory l2 i} {a : formula i} (wd : axioms ⊢w a) → P wd - -- ind-weak-deduction H-ax H-mp (ax in-axioms , _) = - -- H-ax in-axioms - -- ind-weak-deduction {P = P} H-ax H-mp (mp dba db , is-w-dba , is-w-db) = - -- H-mp - -- ( dba , is-w-dba) - -- ( db , is-w-db) - -- ( ind-weak-deduction {P = P} H-ax H-mp (dba , is-w-dba)) - -- ( ind-weak-deduction {P = P} H-ax H-mp (db , is-w-db)) - - -- rec-weak-deduction : - -- {l : Level} - -- {P : UU l} → - -- ( {axioms : modal-theory l2 i} {a : formula i} → - -- is-in-subtype axioms a → P) → - -- ( {axioms : modal-theory l2 i} {a b : formula i} - -- (dab : axioms ⊢w a →ₘ b) (da : axioms ⊢w a) → - -- P → P → P) → - -- {axioms : modal-theory l2 i} {a : formula i} → - -- axioms ⊢w a → P - -- rec-weak-deduction = ind-weak-deduction - ind-weak-deduction : {l : Level} {axioms : modal-theory l2 i} (P : {a : formula i} → axioms ⊢w a → UU l) → @@ -151,15 +120,34 @@ module _ is-weak-modal-logic-Prop : modal-theory l2 i → Prop (l1 ⊔ l2) is-weak-modal-logic-Prop theory = leq-prop-subtype (weak-modal-logic-closure theory) theory - -- implicit-Π-Prop (formula i) (λ a → function-Prop (theory ⊢w a) (theory a)) is-weak-modal-logic : modal-theory l2 i → UU (l1 ⊔ l2) is-weak-modal-logic = type-Prop ∘ is-weak-modal-logic-Prop - is-in-weak-deduction-closure-weak-deduction : + is-in-weak-modal-logic-closure-weak-deduction : {axioms : modal-theory l2 i} {a : formula i} → axioms ⊢w a → is-in-subtype (weak-modal-logic-closure axioms) a - is-in-weak-deduction-closure-weak-deduction = unit-trunc-Prop + is-in-weak-modal-logic-closure-weak-deduction = unit-trunc-Prop + + subset-weak-modal-logic-closure-modal-logic-closure : + {axioms : modal-theory l2 i} → + weak-modal-logic-closure axioms ⊆ modal-logic-closure axioms + subset-weak-modal-logic-closure-modal-logic-closure {axioms} a = + map-universal-property-trunc-Prop + ( modal-logic-closure axioms a) + ( is-in-modal-logic-closure-deduction ∘ deduction-weak-deduction) + + is-weak-modal-logic-is-modal-logic : + {theory : modal-theory l2 i} → + is-modal-logic theory → + is-weak-modal-logic theory + is-weak-modal-logic-is-modal-logic {theory} is-m = + transitive-leq-subtype + ( weak-modal-logic-closure theory) + ( modal-logic-closure theory) + ( theory) + ( is-m) + ( subset-weak-modal-logic-closure-modal-logic-closure) module _ {l1 l2 : Level} {i : Set l1} {axioms : modal-theory l2 i} @@ -170,7 +158,7 @@ module _ is-in-subtype axioms a → is-in-subtype (weak-modal-logic-closure axioms) a weak-modal-logic-closure-ax = - is-in-weak-deduction-closure-weak-deduction ∘ weak-deduction-ax + is-in-weak-modal-logic-closure-weak-deduction ∘ weak-deduction-ax weak-modal-logic-closure-mp : {a b : formula i} → @@ -181,7 +169,7 @@ module _ apply-twice-universal-property-trunc-Prop twdab twda ( weak-modal-logic-closure axioms b) ( λ wdab wda → - ( is-in-weak-deduction-closure-weak-deduction + ( is-in-weak-modal-logic-closure-weak-deduction ( weak-deduction-mp wdab wda))) module _ @@ -221,8 +209,9 @@ module _ ( weak-modal-logic-closure axioms a) ( weak-modal-logic-closed) - subset-axioms-weak-modal-logic : axioms ⊆ weak-modal-logic-closure axioms - subset-axioms-weak-modal-logic a = weak-modal-logic-closure-ax + subset-axioms-weak-modal-logic-closure : + axioms ⊆ weak-modal-logic-closure axioms + subset-axioms-weak-modal-logic-closure a = weak-modal-logic-closure-ax module _ {l1 l2 l3 : Level} {i : Set l1} @@ -242,7 +231,7 @@ module _ weak-modal-logic-closure-monotic a = map-universal-property-trunc-Prop ( weak-modal-logic-closure ax₂ a) - ( is-in-weak-deduction-closure-weak-deduction ∘ weak-deduction-monotic) + ( is-in-weak-modal-logic-closure-weak-deduction ∘ weak-deduction-monotic) module _ {l1 l2 l3 : Level} {i : Set l1} @@ -314,7 +303,7 @@ module _ ( λ {b} b-in-axioms → ( elim-theory-add-formula a axioms ( λ x → weak-modal-logic-closure axioms (a →ₘ x)) - ( is-in-weak-deduction-closure-weak-deduction (deduction-a->a a)) + ( is-in-weak-modal-logic-closure-weak-deduction (deduction-a->a a)) ( λ x x-in-axioms → ( weak-modal-logic-closure-mp ( weak-modal-logic-closure-ax @@ -341,7 +330,7 @@ module _ pr2 (deduction-lemma a b) = map-universal-property-trunc-Prop ( weak-modal-logic-closure (theory-add-formula a axioms) b) - ( is-in-weak-deduction-closure-weak-deduction ∘ + ( is-in-weak-modal-logic-closure-weak-deduction ∘ backward-deduction-lemma) ``` @@ -411,4 +400,139 @@ module _ ( list-subtype (list-assumptions-weak-deduction dab)) ( list-subtype (list-assumptions-weak-deduction da))) ( ra))))) + +module _ + {l1 l2 : Level} {i : Set l1} (axioms : modal-theory l2 i) + (contains-ax-k : ax-k i ⊆ axioms) + (contains-ax-s : ax-s i ⊆ axioms) + (contains-ax-dn : ax-dn i ⊆ axioms) + where + + -- TODO: move to formulas-deduction + + deduction-ex-falso : + (a b : formula i) → + is-in-subtype (weak-modal-logic-closure axioms) (~ a →ₘ a →ₘ b) + deduction-ex-falso a b = + forward-implication + ( deduction-lemma axioms contains-ax-k contains-ax-s (~ a) (a →ₘ b)) + ( forward-implication + ( deduction-lemma + ( theory-add-formula (~ a) axioms) + ( contains-ax-k') + ( contains-ax-s') + ( a) + ( b)) + ( weak-modal-logic-closure-mp {a = ~~ b} + ( weak-modal-logic-closure-ax + ( contains-ax-dn'' (~~ b →ₘ b) (b , refl))) + ( weak-modal-logic-closure-mp {a = ⊥} + ( weak-modal-logic-closure-ax + ( contains-ax-k'' (⊥ →ₘ ~ b →ₘ ⊥) (⊥ , ~ b , refl))) + ( weak-modal-logic-closure-mp {a = a} + ( weak-modal-logic-closure-ax + ( subset-add-formula a + ( theory-add-formula (~ a) axioms) + ( ~ a) + ( formula-in-add-formula (~ a) axioms))) + ( weak-modal-logic-closure-ax + ( formula-in-add-formula a + ( theory-add-formula (~ a) axioms))))))) + where + contains-ax-k' : ax-k i ⊆ theory-add-formula (~ a) axioms + contains-ax-k' = + transitive-subset-add-formula (~ a) axioms (ax-k i) contains-ax-k + + contains-ax-s' : ax-s i ⊆ theory-add-formula (~ a) axioms + contains-ax-s' = + transitive-subset-add-formula (~ a) axioms (ax-s i) contains-ax-s + + contains-ax-k'' : + ax-k i ⊆ theory-add-formula a (theory-add-formula (~ a) axioms) + contains-ax-k'' = + transitive-subset-add-formula a (theory-add-formula (~ a) axioms) + ( ax-k i) + ( contains-ax-k') + + contains-ax-dn'' : + ax-dn i ⊆ theory-add-formula a (theory-add-formula (~ a) axioms) + contains-ax-dn'' = + transitive-subset-add-formula a + ( theory-add-formula (~ a) axioms) + ( ax-dn i) + ( transitive-subset-add-formula (~ a) axioms (ax-dn i) contains-ax-dn) + + logic-ex-falso : + (a b : formula i) → + is-in-subtype (weak-modal-logic-closure axioms) a → + is-in-subtype (weak-modal-logic-closure axioms) (~ a) → + is-in-subtype (weak-modal-logic-closure axioms) b + logic-ex-falso a b a-in-logic not-a-in-logic = + weak-modal-logic-closure-mp + ( weak-modal-logic-closure-mp + ( deduction-ex-falso a b) + ( not-a-in-logic)) + ( a-in-logic) +module _ + {l1 l2 : Level} {i : Set l1} (axioms : modal-theory l2 i) + (contains-ax-k : ax-k i ⊆ axioms) + (contains-ax-s : ax-s i ⊆ axioms) + (contains-ax-dn : ax-dn i ⊆ axioms) + where + + inv-deduction-ex-falso : + (a b : formula i) → + is-in-subtype (weak-modal-logic-closure axioms) (a →ₘ ~ a →ₘ b) + inv-deduction-ex-falso a b = + forward-implication + ( deduction-lemma axioms contains-ax-k contains-ax-s a (~ a →ₘ b)) + ( forward-implication + ( deduction-lemma + ( theory-add-formula a axioms) + ( contains-ax-k') + ( contains-ax-s') + ( ~ a) + ( b)) + ( logic-ex-falso + ( theory-add-formula (a →ₘ ⊥) (theory-add-formula a axioms)) + ( contains-ax-k'') + ( contains-ax-s'') + ( contains-ax-dn'') + ( a) + ( b) + ( weak-modal-logic-closure-ax + ( subset-add-formula (~ a) (theory-add-formula a axioms) a + ( formula-in-add-formula a axioms))) + ( weak-modal-logic-closure-ax + ( formula-in-add-formula (~ a) (theory-add-formula a axioms))))) + where + contains-ax-k' : ax-k i ⊆ theory-add-formula a axioms + contains-ax-k' = + transitive-subset-add-formula a axioms (ax-k i) contains-ax-k + + contains-ax-s' : ax-s i ⊆ theory-add-formula a axioms + contains-ax-s' = + transitive-subset-add-formula a axioms (ax-s i) contains-ax-s + + contains-ax-k'' : + ax-k i ⊆ theory-add-formula (~ a) (theory-add-formula a axioms) + contains-ax-k'' = + transitive-subset-add-formula (~ a) (theory-add-formula a axioms) + ( ax-k i) + ( contains-ax-k') + + contains-ax-s'' : + ax-s i ⊆ theory-add-formula (~ a) (theory-add-formula a axioms) + contains-ax-s'' = + transitive-subset-add-formula (~ a) (theory-add-formula a axioms) + ( ax-s i) + ( contains-ax-s') + + contains-ax-dn'' : + ax-dn i ⊆ theory-add-formula (~ a) (theory-add-formula a axioms) + contains-ax-dn'' = + transitive-subset-add-formula (~ a) + ( theory-add-formula a axioms) + ( ax-dn i) + ( transitive-subset-add-formula a axioms (ax-dn i) contains-ax-dn) ``` diff --git a/src/order-theory.lagda.md b/src/order-theory.lagda.md index 2c49831616..e8a88dd031 100644 --- a/src/order-theory.lagda.md +++ b/src/order-theory.lagda.md @@ -114,4 +114,5 @@ open import order-theory.upper-bounds-posets public open import order-theory.upper-sets-large-posets public open import order-theory.well-founded-orders public open import order-theory.well-founded-relations public +open import order-theory.zorn public ``` diff --git a/src/order-theory/zorn.lagda.md b/src/order-theory/zorn.lagda.md index 0077eedb2c..d0e709d67b 100644 --- a/src/order-theory/zorn.lagda.md +++ b/src/order-theory/zorn.lagda.md @@ -16,12 +16,12 @@ open import foundation.logical-equivalences open import foundation.propositional-truncations open import foundation.universe-levels -open import foundation-core.propositions open import foundation-core.coproduct-types +open import foundation-core.propositions open import order-theory.chains-posets -open import order-theory.posets open import order-theory.maximal-elements-posets +open import order-theory.posets ```
From 7696793a2ffa77a1ab9b56e44edc3a851ad34dbf Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Sun, 5 May 2024 14:47:59 +0300 Subject: [PATCH 44/63] Fix types in filtrations --- src/lists/concatenation-lists.lagda.md | 12 ++ .../finite-approximability.lagda.md | 10 +- src/modal-logic/modal-logic-decision.lagda.md | 196 ++++-------------- 3 files changed, 63 insertions(+), 155 deletions(-) diff --git a/src/lists/concatenation-lists.lagda.md b/src/lists/concatenation-lists.lagda.md index a8e362c224..18a0a0e797 100644 --- a/src/lists/concatenation-lists.lagda.md +++ b/src/lists/concatenation-lists.lagda.md @@ -17,6 +17,8 @@ open import foundation.identity-types open import foundation.sets open import foundation.universe-levels +open import foundation-core.coproduct-types + open import group-theory.monoids open import lists.lists @@ -190,4 +192,14 @@ in-concat-right : in-concat-right nil l2 in-l2 = in-l2 in-concat-right (cons x l1) l2 in-l2 = is-in-tail _ x (concat-list l1 l2) (in-concat-right l1 l2 in-l2) + +in-concat-list : + {l : Level} {A : UU l} (l1 l2 : list A) + {a : A} → a ∈-list (concat-list l1 l2) → (a ∈-list l1) + (a ∈-list l2) +in-concat-list nil l2 {a} in-list = inr in-list +in-concat-list (cons x l1) l2 {.x} (is-head .x _) = inl (is-head x l1) +in-concat-list (cons x l1) l2 {a} (is-in-tail .a .x _ in-list) + with in-concat-list l1 l2 in-list +... | inl in-l1 = inl (is-in-tail a x l1 in-l1) +... | inr in-l2 = inr in-l2 ``` diff --git a/src/modal-logic/finite-approximability.lagda.md b/src/modal-logic/finite-approximability.lagda.md index bdf97c31a7..b7bff428c4 100644 --- a/src/modal-logic/finite-approximability.lagda.md +++ b/src/modal-logic/finite-approximability.lagda.md @@ -48,7 +48,7 @@ open import modal-logic.modal-logic-k open import modal-logic.soundness open import modal-logic.weak-deduction -open import order-theory.maximal-elements-posets +open import order-theory.zorn open import univalent-combinatorics.finite-types ``` @@ -87,7 +87,7 @@ module _ module _ (lem : LEM (lsuc (lsuc l1))) - (zorn : Zorn (lsuc l1) l1 l1) + (zorn : Zorn-non-empty (lsuc l1) l1 l1) (prop-resize : propositional-resizing l1 (lsuc l1)) where @@ -108,7 +108,7 @@ module _ ( λ ((a , M , _) , p) → ( tr (is-in-subtype (filtration-models _ _ i _ _ _ _ _)) (inv p) ( intro-∃ - ( in-list (subformulas-list i a) , M) + ( subformulas i a , M) ( pair (is-finite-subformulas-list ( i) @@ -116,9 +116,9 @@ module _ ( a)) ( is-kripke-model-filtration-minimal-kripke-model-filtration ( i) - ( in-list (subformulas-list i a)) + ( subformulas i a) ( M) - ( is-modal-theory-closed-under-subformulas-subformulas-list + ( is-modal-theory-closed-under-subformulas-subformulas ( i) ( a))))))) diff --git a/src/modal-logic/modal-logic-decision.lagda.md b/src/modal-logic/modal-logic-decision.lagda.md index ab395be441..6bf9173d74 100644 --- a/src/modal-logic/modal-logic-decision.lagda.md +++ b/src/modal-logic/modal-logic-decision.lagda.md @@ -49,6 +49,7 @@ open import foundation-core.propositions open import lists.arrays open import lists.concatenation-lists open import lists.lists +open import lists.lists-subtypes open import lists.reversing-lists open import modal-logic.completeness @@ -116,121 +117,11 @@ module _ with decision-procedure' a ... | inl valid-in-C = complete a (λ M M-in-C → valid-in-C (M , M-in-C)) -module _ - {l : Level} {i : Set l} - where - - collect-vars : formula i → list (type-Set i) - collect-vars (var x) = cons x nil - collect-vars ⊥ = nil - collect-vars (a →ₘ b) = concat-list (collect-vars a) (collect-vars b) - collect-vars (□ a) = collect-vars a - -module _ - {l1 l2 l3 l4 l5 : Level} (i : Set l3) - (F : kripke-frame l1 l2) - where - - valuates-equals-on-vars : - {l6 : Level} → - subtype l6 (type-Set i) → - (type-Set i → type-kripke-frame F → Prop l4) → - (type-Set i → type-kripke-frame F → Prop l5) → - UU (l1 ⊔ l3 ⊔ l4 ⊔ l5 ⊔ l6) - valuates-equals-on-vars vars V V' = - ((a , _) : type-subtype vars) → (x : type-kripke-frame F) → V a x ⇔ V' a x - - valuates-equals-on-vars-subtype : - {l6 l7 : Level} - (vars₁ : subtype l6 (type-Set i)) - (vars₂ : subtype l7 (type-Set i)) → - vars₂ ⊆ vars₁ → - (V : type-Set i → type-kripke-frame F → Prop l4) - (V' : type-Set i → type-kripke-frame F → Prop l5) → - valuates-equals-on-vars (vars₁) V V' → - valuates-equals-on-vars (vars₂) V V' - valuates-equals-on-vars-subtype vars₁ vars₂ sub V V' eq (a , in-vars) = - eq (a , sub a in-vars) - - affect-only-occured-vars : - (a : formula i) → - (V : type-Set i → type-kripke-frame F → Prop l4) → - (V' : type-Set i → type-kripke-frame F → Prop l5) → - valuates-equals-on-vars (in-list (collect-vars a)) V V' → - (x : type-kripke-frame F) → - (((F , V) , x) ⊨ a) ⇔ (((F , V') , x) ⊨ a) - affect-only-occured-vars (var n) V V' eq x = - pair - ( λ v → - ( map-raise - ( forward-implication - ( eq (n , unit-trunc-Prop (is-head n nil)) x) - ( map-inv-raise v)))) - ( λ v → - ( map-raise - ( backward-implication - ( eq (n , unit-trunc-Prop (is-head n nil)) x) - ( map-inv-raise v)))) - affect-only-occured-vars ⊥ V V' eq x = - (map-raise ∘ map-inv-raise) , (map-raise ∘ map-inv-raise) - affect-only-occured-vars (a →ₘ b) V V' eq x = - pair - ( λ fab fa → - ( forward-implication - ( affect-only-occured-vars b V V' - ( valuates-equals-on-vars-subtype - ( in-list (collect-vars (a →ₘ b))) - ( in-list (collect-vars b)) - ( subset-in-concat-right (collect-vars a) (collect-vars b)) - ( V) - ( V') - ( eq)) - ( x)) - ( fab - ( backward-implication - ( affect-only-occured-vars a V V' - ( valuates-equals-on-vars-subtype - ( in-list (collect-vars (a →ₘ b))) - ( in-list (collect-vars a)) - ( subset-in-concat-left (collect-vars a) (collect-vars b)) - ( V) - ( V') - ( eq)) - ( x)) - ( fa))))) - ( λ fab fa → - ( backward-implication - ( affect-only-occured-vars b V V' - ( valuates-equals-on-vars-subtype - ( in-list (collect-vars (a →ₘ b))) - ( in-list (collect-vars b)) - ( subset-in-concat-right (collect-vars a) (collect-vars b)) - ( V) - ( V') - ( eq)) - ( x)) - ( fab - ( forward-implication - ( affect-only-occured-vars a V V' - ( valuates-equals-on-vars-subtype - ( in-list (collect-vars (a →ₘ b))) - ( in-list (collect-vars a)) - ( subset-in-concat-left (collect-vars a) (collect-vars b)) - ( V) - ( V') - ( eq)) - ( x)) - ( fa))))) - affect-only-occured-vars (□ a) V V' eq x = - pair - ( λ f y r → - ( forward-implication (affect-only-occured-vars a V V' eq y) (f y r))) - ( λ f y r → - ( backward-implication (affect-only-occured-vars a V V' eq y) (f y r))) - +-- TODO: move to kuratowsky-finite-sets is-kuratowsky-finite-set-Prop' : {l : Level} → Set l → Prop l is-kuratowsky-finite-set-Prop' X = - ∃-Prop (list (type-Set X)) (λ l → (x : type-Set X) → type-Prop (in-list l x)) + ∃-Prop (list (type-Set X)) + ( λ l → (x : type-Set X) → type-Prop (list-subtype l x)) is-kuratowsky-finite-set' : {l : Level} → Set l → UU l is-kuratowsky-finite-set' X = type-Prop (is-kuratowsky-finite-set-Prop' X) @@ -241,21 +132,22 @@ is-kuratowsky-finite-set-is-kuratowsky-finite-set' : is-kuratowsky-finite-set-is-kuratowsky-finite-set' X = map-universal-property-trunc-Prop ( is-kuratowsky-finite-set-Prop X) - ( λ (l , all-in-list) → + ( λ (l , all-list-subtype) → ( intro-∃ ( length-list l) ( pair ( component-list l) ( λ x → ( apply-universal-property-trunc-Prop - ( all-in-list x) + ( all-list-subtype x) ( trunc-Prop (fiber _ x)) - ( λ x-in-list → + ( λ x-list-subtype → ( unit-trunc-Prop ( pair - ( index-in-list x l x-in-list) + ( index-in-list x l x-list-subtype) ( inv - ( eq-component-list-index-in-list x l x-in-list)))))))))) + ( eq-component-list-index-in-list x l + ( x-list-subtype))))))))))) dependent-map-list : {l1 l2 : Level} {A : UU l1} {B : UU l2} @@ -266,7 +158,7 @@ dependent-map-list {A = A} {B} (cons x l) f = cons (f (x , is-head x l)) (dependent-map-list l f') where f' : Σ A (λ a → a ∈-list l) → B - f' (a , in-list) = f (a , is-in-tail a x l in-list) + f' (a , list-subtype) = f (a , is-in-tail a x l list-subtype) in-dependent-map-list : {l1 l2 : Level} {A : UU l1} {B : UU l2} @@ -291,7 +183,7 @@ module _ rest (□ a) = subformulas-list a subformulas : formula i → modal-theory l i - subformulas a = in-list (subformulas-list a) + subformulas a = list-subtype (subformulas-list a) subformulas-list-has-subimpl : (a : formula i) {x y : formula i} → @@ -310,13 +202,12 @@ module _ ( subformulas-list y) ( is-head y _))) subformulas-list-has-subimpl - (a →ₘ b) {x} {y} (is-in-tail .(x →ₘ y) .(a →ₘ b) _ xy-in-list) + (a →ₘ b) {x} {y} (is-in-tail .(x →ₘ y) .(a →ₘ b) _ xy-list-subtype) with in-concat-list ( subformulas-list a) ( subformulas-list b) - ( x →ₘ y) - ( xy-in-list) + ( xy-list-subtype) ... | inl xy-in-left = let (x-in-tail , y-in-tail) = subformulas-list-has-subimpl a xy-in-left in pair @@ -332,8 +223,9 @@ module _ ( is-in-tail y (a →ₘ b) _ ( in-concat-right (subformulas-list a) (subformulas-list b) y-in-tail)) subformulas-list-has-subimpl - (□ a) {x} {y} (is-in-tail .(x →ₘ y) .(□ a) _ xy-in-list) = - let (x-in-tail , y-in-tail) = subformulas-list-has-subimpl a xy-in-list + (□ a) {x} {y} (is-in-tail .(x →ₘ y) .(□ a) _ xy-list-subtype) = + let (x-in-tail , y-in-tail) = + subformulas-list-has-subimpl a xy-list-subtype in (is-in-tail x (□ a) _ x-in-tail) , (is-in-tail y (□ a) _ y-in-tail) subformulas-list-has-subbox : @@ -343,9 +235,9 @@ module _ subformulas-list-has-subbox .(□ x) {x} (is-head .(□ x) _) = is-in-tail x (□ x) _ (is-head x _) subformulas-list-has-subbox - (a →ₘ b) {x} (is-in-tail .(□ x) .(a →ₘ b) _ x-in-list) + (a →ₘ b) {x} (is-in-tail .(□ x) .(a →ₘ b) _ x-list-subtype) with - in-concat-list (subformulas-list a) (subformulas-list b) (□ x) x-in-list + in-concat-list (subformulas-list a) (subformulas-list b) x-list-subtype ... | inl x-in-left = is-in-tail x (a →ₘ b) _ ( in-concat-left (subformulas-list a) (subformulas-list b) @@ -354,8 +246,9 @@ module _ is-in-tail x (a →ₘ b) _ ( in-concat-right (subformulas-list a) (subformulas-list b) ( subformulas-list-has-subbox b x-in-right)) - subformulas-list-has-subbox (□ a) {x} (is-in-tail .(□ x) .(□ a) _ x-in-list) = - is-in-tail x (□ a) _ (subformulas-list-has-subbox a x-in-list) + subformulas-list-has-subbox + (□ a) {x} (is-in-tail .(□ x) .(□ a) _ x-list-subtype) = + is-in-tail x (□ a) _ (subformulas-list-has-subbox a x-list-subtype) is-modal-theory-closed-under-subformulas-subformulas : (a : formula i) → @@ -363,45 +256,48 @@ module _ is-modal-theory-closed-under-subformulas-subformulas a = is-modal-theory-closed-under-subformulas-condition ( i) - ( in-list (subformulas-list a)) + ( list-subtype (subformulas-list a)) ( map-universal-property-trunc-Prop ( product-Prop - ( in-list (subformulas-list a) _) - ( in-list (subformulas-list a) _)) - ( λ impl-in-list → - ( let (a-in-list , b-in-list) = - subformulas-list-has-subimpl a impl-in-list - in (unit-trunc-Prop a-in-list) , (unit-trunc-Prop b-in-list)))) + ( list-subtype (subformulas-list a) _) + ( list-subtype (subformulas-list a) _)) + ( λ impl-list-subtype → + ( let (a-list-subtype , b-list-subtype) = + subformulas-list-has-subimpl a impl-list-subtype + in pair + ( unit-trunc-Prop a-list-subtype) + ( unit-trunc-Prop b-list-subtype)))) ( map-universal-property-trunc-Prop - ( in-list (subformulas-list a) _) + ( list-subtype (subformulas-list a) _) ( unit-trunc-Prop ∘ subformulas-list-has-subbox a)) subformulas-Set : formula i → Set l - subformulas-Set a = set-subset (formula-Set i) (in-list (subformulas-list a)) + subformulas-Set a = + set-subset (formula-Set i) (list-subtype (subformulas-list a)) subformulas-Set-list : (a : formula i) → list (type-Set (subformulas-Set a)) subformulas-Set-list a = dependent-map-list ( subformulas-list a) - ( λ (x , in-list) → x , unit-trunc-Prop in-list) + ( λ (x , list-subtype) → x , unit-trunc-Prop list-subtype) is-kuratowsky-finite'-subformulas-list : (a : formula i) → is-kuratowsky-finite-set' (subformulas-Set a) is-kuratowsky-finite'-subformulas-list a = intro-∃ ( subformulas-Set-list a) - ( λ (b , trunc-b-in-list) → + ( λ (b , trunc-b-list-subtype) → ( apply-universal-property-trunc-Prop - ( trunc-b-in-list) - ( in-list (subformulas-Set-list a) (b , trunc-b-in-list)) - ( λ b-in-list → + ( trunc-b-list-subtype) + ( list-subtype (subformulas-Set-list a) (b , trunc-b-list-subtype)) + ( λ b-list-subtype → ( unit-trunc-Prop ( tr ( λ i → (b , i) ∈-list subformulas-Set-list a) ( eq-is-prop is-prop-type-trunc-Prop) ( in-dependent-map-list - (λ (x , in-list) → x , unit-trunc-Prop in-list) - ( b-in-list))))))) + (λ (x , list-subtype) → x , unit-trunc-Prop list-subtype) + ( b-list-subtype))))))) is-kuratowsky-finite-subformulas-list : (a : formula i) → is-kuratowsky-finite-set (subformulas-Set a) @@ -413,7 +309,7 @@ module _ is-finite-subformulas-list : LEM l → (a : formula i) → - is-finite (type-subtype (in-list (subformulas-list a))) + is-finite (type-subtype (list-subtype (subformulas-list a))) is-finite-subformulas-list lem a = is-finite-is-kuratowsky-finite-set (subformulas-Set a) lem (is-kuratowsky-finite-subformulas-list a) @@ -423,7 +319,7 @@ module _ LEM l → LEM l2 → (a : formula i) → - is-finite (type-subtype (in-list (subformulas-list a)) → Prop l2) + is-finite (type-subtype (list-subtype (subformulas-list a)) → Prop l2) is-finite-subtypes-subformulas-list lem lem2 a = is-finite-function-type ( is-finite-subformulas-list lem a) @@ -691,13 +587,13 @@ module _ ( unit-trunc-Prop (is-head a _)) ( x)) ( in-logic - ( filtration (in-list (cons a _)) M) + ( filtration (list-subtype (cons a _)) M) ( intro-∃ (a , (M , M-in-class)) refl) ( map-equiv-is-kripke-model-filtration ( i) ( subformulas i a) ( M) - ( filtration (in-list (cons a _)) M) + ( filtration (list-subtype (cons a _)) M) ( is-filt) ( class ( Φ-equivalence i @@ -710,7 +606,7 @@ module _ (C₂ : model-class l6 l7 i l8 l10) → ( ((M , _) : type-subtype C) → (a : formula i) → - is-in-subtype C₂ (filtration (in-list (subformulas-list i a)) M)) → + is-in-subtype C₂ (filtration (list-subtype (subformulas-list i a)) M)) → soundness logic C₂ → soundness logic filtrate-class filtrate-soundness logic C₂ H sound a in-logic M* in-class = From 89951d2143a162d26b36f289ce96d5ad9d40714f Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Sun, 5 May 2024 14:54:20 +0300 Subject: [PATCH 45/63] Remove upper characters from file names --- src/modal-logic.lagda.md | 4 ++++ .../{completeness-K.lagda.md => completeness-k.lagda.md} | 0 .../{completeness-S5.lagda.md => completeness-s5.lagda.md} | 0 .../{modal-logic-K.lagda.md => modal-logic-k.lagda.md} | 0 .../{modal-logic-S5.lagda.md => modal-logic-s5.lagda.md} | 0 5 files changed, 4 insertions(+) rename src/modal-logic/{completeness-K.lagda.md => completeness-k.lagda.md} (100%) rename src/modal-logic/{completeness-S5.lagda.md => completeness-s5.lagda.md} (100%) rename src/modal-logic/{modal-logic-K.lagda.md => modal-logic-k.lagda.md} (100%) rename src/modal-logic/{modal-logic-S5.lagda.md => modal-logic-s5.lagda.md} (100%) diff --git a/src/modal-logic.lagda.md b/src/modal-logic.lagda.md index fe11bc1e13..af5fcf39d7 100644 --- a/src/modal-logic.lagda.md +++ b/src/modal-logic.lagda.md @@ -7,6 +7,8 @@ open import modal-logic.axioms public open import modal-logic.canonical-model-theorem public open import modal-logic.canonical-theories public open import modal-logic.completeness public +open import modal-logic.completeness-k public +open import modal-logic.completeness-s5 public open import modal-logic.finite-approximability public open import modal-logic.formulas public open import modal-logic.formulas-deduction public @@ -17,6 +19,8 @@ open import modal-logic.l-complete-theories public open import modal-logic.l-consistent-theories public open import modal-logic.lindenbaum public open import modal-logic.logic-syntax public +open import modal-logic.modal-logic-k public +open import modal-logic.modal-logic-s5 public open import modal-logic.modal-logic-decision public open import modal-logic.soundness public open import modal-logic.weak-deduction public diff --git a/src/modal-logic/completeness-K.lagda.md b/src/modal-logic/completeness-k.lagda.md similarity index 100% rename from src/modal-logic/completeness-K.lagda.md rename to src/modal-logic/completeness-k.lagda.md diff --git a/src/modal-logic/completeness-S5.lagda.md b/src/modal-logic/completeness-s5.lagda.md similarity index 100% rename from src/modal-logic/completeness-S5.lagda.md rename to src/modal-logic/completeness-s5.lagda.md diff --git a/src/modal-logic/modal-logic-K.lagda.md b/src/modal-logic/modal-logic-k.lagda.md similarity index 100% rename from src/modal-logic/modal-logic-K.lagda.md rename to src/modal-logic/modal-logic-k.lagda.md diff --git a/src/modal-logic/modal-logic-S5.lagda.md b/src/modal-logic/modal-logic-s5.lagda.md similarity index 100% rename from src/modal-logic/modal-logic-S5.lagda.md rename to src/modal-logic/modal-logic-s5.lagda.md From c60b49bfabdd7aced7919ed952b3ddc2cea99d9f Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Sun, 5 May 2024 14:56:17 +0300 Subject: [PATCH 46/63] Make pre-commit one more time --- src/modal-logic.lagda.md | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/modal-logic.lagda.md b/src/modal-logic.lagda.md index af5fcf39d7..83447a76c4 100644 --- a/src/modal-logic.lagda.md +++ b/src/modal-logic.lagda.md @@ -19,9 +19,9 @@ open import modal-logic.l-complete-theories public open import modal-logic.l-consistent-theories public open import modal-logic.lindenbaum public open import modal-logic.logic-syntax public +open import modal-logic.modal-logic-decision public open import modal-logic.modal-logic-k public open import modal-logic.modal-logic-s5 public -open import modal-logic.modal-logic-decision public open import modal-logic.soundness public open import modal-logic.weak-deduction public ``` From 706ff33d61c8ef7aba624769c418458c73f96f28 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Sun, 5 May 2024 16:37:40 +0300 Subject: [PATCH 47/63] Refactor after rebase --- .../propositional-resizing.lagda.md | 5 +- src/foundation/unions-subtypes.lagda.md | 124 ++++++------------ src/lists/lists-subtypes.lagda.md | 89 +++++++------ .../canonical-model-theorem.lagda.md | 2 +- src/modal-logic/completeness-k.lagda.md | 9 -- .../finite-approximability.lagda.md | 6 +- .../kripke-models-filtrations.lagda.md | 14 +- src/modal-logic/kripke-semantics.lagda.md | 7 +- src/modal-logic/l-complete-theories.lagda.md | 57 ++++---- src/modal-logic/lindenbaum.lagda.md | 15 ++- src/modal-logic/logic-syntax.lagda.md | 4 +- src/modal-logic/modal-logic-decision.lagda.md | 12 +- src/modal-logic/modal-logic-k.lagda.md | 19 +-- src/modal-logic/modal-logic-s5.lagda.md | 2 +- src/modal-logic/weak-deduction.lagda.md | 5 +- src/order-theory.lagda.md | 1 - src/order-theory/chains-posets.lagda.md | 2 +- src/order-theory/chains-union.lagda.md | 121 ----------------- src/order-theory/zorn.lagda.md | 11 +- .../kuratowsky-finite-sets.lagda.md | 4 +- 20 files changed, 171 insertions(+), 338 deletions(-) delete mode 100644 src/order-theory/chains-union.lagda.md diff --git a/src/foundation/propositional-resizing.lagda.md b/src/foundation/propositional-resizing.lagda.md index c71522f94a..5b08324f27 100644 --- a/src/foundation/propositional-resizing.lagda.md +++ b/src/foundation/propositional-resizing.lagda.md @@ -11,6 +11,7 @@ open import foundation.decidable-types open import foundation.dependent-pair-types open import foundation.empty-types open import foundation.law-of-excluded-middle +open import foundation.logical-equivalences open import foundation.negation open import foundation.raising-universe-levels open import foundation.unit-type @@ -60,13 +61,13 @@ unit-equiv-true : {l : Level} (P : Prop l) → type-Prop P → type-equiv-Prop unit-Prop P pr1 (unit-equiv-true P p) _ = p pr2 (unit-equiv-true P p) = - is-equiv-is-prop is-prop-unit (is-prop-type-Prop P) (λ _ → star) + is-equiv-has-converse-is-prop is-prop-unit (is-prop-type-Prop P) (λ _ → star) empty-equiv-false : {l : Level} (P : Prop l) → ¬ (type-Prop P) → type-equiv-Prop empty-Prop P pr1 (empty-equiv-false P np) = ex-falso pr2 (empty-equiv-false P np) = - is-equiv-is-prop is-prop-empty (is-prop-type-Prop P) np + is-equiv-has-converse-is-prop is-prop-empty (is-prop-type-Prop P) np propositional-resizing-LEM : (l1 : Level) {l2 : Level} → LEM l2 → propositional-resizing l1 l2 diff --git a/src/foundation/unions-subtypes.lagda.md b/src/foundation/unions-subtypes.lagda.md index 058326ea9a..c36dd462aa 100644 --- a/src/foundation/unions-subtypes.lagda.md +++ b/src/foundation/unions-subtypes.lagda.md @@ -41,9 +41,8 @@ module _ union-subtype : subtype l1 X → subtype l2 X → subtype (l1 ⊔ l2) X union-subtype P Q x = (P x) ∨ (Q x) - union-subtype P Q x = disjunction-Prop (P x) (Q x) - infixl 10 _∪_ + infixr 10 _∪_ _∪_ = union-subtype ``` @@ -103,43 +102,41 @@ module _ {l1 l2 l3 : Level} {X : UU l1} (P : subtype l2 X) (Q : subtype l3 X) where - subtype-union-left : P ⊆ union-subtype P Q - subtype-union-left x = inl-disjunction-Prop (P x) (Q x) + subtype-union-left : P ⊆ P ∪ Q + -- subtype-union-left x = inl-disjunction-Prop (P x) (Q x) + subtype-union-left x = inl-disjunction - subtype-union-right : Q ⊆ union-subtype P Q - subtype-union-right x = inr-disjunction-Prop (P x) (Q x) + subtype-union-right : Q ⊆ P ∪ Q + -- subtype-union-right x = inr-disjunction-Prop (P x) (Q x) + subtype-union-right x = inr-disjunction elim-union-subtype : {l4 : Level} (f : X → Prop l4) → ((x : X) → is-in-subtype P x → type-Prop (f x)) → ((x : X) → is-in-subtype Q x → type-Prop (f x)) → (x : X) → is-in-subtype (P ∪ Q) x → type-Prop (f x) - elim-union-subtype f H-P H-Q x = - elim-disjunction-Prop (P x) (Q x) (f x) (H-P x , H-Q x) + elim-union-subtype f H-P H-Q x = elim-disjunction (f x) (H-P x) (H-Q x) elim-union-subtype' : {l4 : Level} (R : Prop l4) → ((x : X) → is-in-subtype P x → type-Prop R) → ((x : X) → is-in-subtype Q x → type-Prop R) → (x : X) → is-in-subtype (P ∪ Q) x → type-Prop R - elim-union-subtype' R H-P H-Q x = - elim-disjunction-Prop (P x) (Q x) R (H-P x , H-Q x) - -- elim-disjunction-Prop (P x) (Q x) (f x) (H-P x , H-Q x) p + elim-union-subtype' R = elim-union-subtype (λ _ → R) + -- elim-disjunction-Prop (P x) (Q x) R (H-P x , H-Q x) subtype-union-both : - {l4 : Level} (S : subtype l4 X) → P ⊆ S → Q ⊆ S → union-subtype P Q ⊆ S - subtype-union-both S P-sub-S Q-sub-S x = - elim-disjunction-Prop (P x) (Q x) (S x) (P-sub-S x , Q-sub-S x) + {l4 : Level} (S : subtype l4 X) → P ⊆ S → Q ⊆ S → P ∪ Q ⊆ S + subtype-union-both = elim-union-subtype module _ {l1 l2 l3 : Level} {X : UU l1} (P : subtype l2 X) (Q : subtype l3 X) where subset-union-comm : - union-subtype P Q ⊆ union-subtype Q P + P ∪ Q ⊆ Q ∪ P subset-union-comm = - subtype-union-both P Q - ( union-subtype Q P) + subtype-union-both P Q (Q ∪ P) ( subtype-union-right Q P) ( subtype-union-left Q P) @@ -148,50 +145,28 @@ module _ (P : subtype l2 X) (Q : subtype l3 X) (S : subtype l4 X) where - forward-subset-union-assoc : - union-subtype P (union-subtype Q S) ⊆ union-subtype (union-subtype P Q) S + forward-subset-union-assoc : P ∪ (Q ∪ S) ⊆ (P ∪ Q) ∪ S forward-subset-union-assoc = - subtype-union-both - ( P) - ( union-subtype Q S) - ( union-subtype (union-subtype P Q) S) - ( transitive-leq-subtype - ( P) - ( union-subtype P Q) - ( union-subtype (union-subtype P Q) S) - ( subtype-union-left (union-subtype P Q) S) + subtype-union-both P (Q ∪ S) ((P ∪ Q) ∪ S) + ( transitive-leq-subtype P (P ∪ Q) ((P ∪ Q) ∪ S) + ( subtype-union-left (P ∪ Q) S) ( subtype-union-left P Q)) - ( subtype-union-both Q S - ( union-subtype (union-subtype P Q) S) - ( transitive-leq-subtype - ( Q) - ( union-subtype P Q) - ( union-subtype (union-subtype P Q) S) - ( subtype-union-left (union-subtype P Q) S) + ( subtype-union-both Q S ((P ∪ Q) ∪ S) + ( transitive-leq-subtype Q (P ∪ Q) ((P ∪ Q) ∪ S) + ( subtype-union-left (P ∪ Q) S) ( subtype-union-right P Q)) - ( subtype-union-right (union-subtype P Q) S)) + ( subtype-union-right (P ∪ Q) S)) - backward-subset-union-assoc : - union-subtype (union-subtype P Q) S ⊆ union-subtype P (union-subtype Q S) + backward-subset-union-assoc : (P ∪ Q) ∪ S ⊆ P ∪ (Q ∪ S) backward-subset-union-assoc = - subtype-union-both - ( union-subtype P Q) - ( S) - ( union-subtype P (union-subtype Q S)) - ( subtype-union-both P Q - ( union-subtype P (union-subtype Q S)) - ( subtype-union-left P (union-subtype Q S)) - ( transitive-leq-subtype - ( Q) - ( union-subtype Q S) - ( union-subtype P (union-subtype Q S)) - ( subtype-union-right P (union-subtype Q S)) + subtype-union-both (P ∪ Q) S (P ∪ (Q ∪ S)) + ( subtype-union-both P Q (P ∪ (Q ∪ S)) + ( subtype-union-left P (Q ∪ S)) + ( transitive-leq-subtype Q (Q ∪ S) (P ∪ (Q ∪ S)) + ( subtype-union-right P (Q ∪ S)) ( subtype-union-left Q S))) - ( transitive-leq-subtype - ( S) - ( union-subtype Q S) - ( union-subtype P (union-subtype Q S)) - ( subtype-union-right P (union-subtype Q S)) + ( transitive-leq-subtype S (Q ∪ S) (P ∪ (Q ∪ S)) + ( subtype-union-right P (Q ∪ S)) ( subtype-union-right Q S)) module _ @@ -203,13 +178,13 @@ module _ (P1 : subtype l2 X) (Q1 : subtype l3 X) (P2 : subtype l4 X) (Q2 : subtype l5 X) → P1 ⊆ P2 → Q1 ⊆ Q2 → - union-subtype P1 Q1 ⊆ union-subtype P2 Q2 + P1 ∪ Q1 ⊆ P2 ∪ Q2 subset-union-subsets P1 Q1 P2 Q2 P1-sub-P2 Q1-sub-Q2 = - subtype-union-both P1 Q1 (union-subtype P2 Q2) - ( transitive-leq-subtype P1 P2 (union-subtype P2 Q2) + subtype-union-both P1 Q1 (P2 ∪ Q2) + ( transitive-leq-subtype P1 P2 (P2 ∪ Q2) ( subtype-union-left P2 Q2) ( P1-sub-P2)) - ( transitive-leq-subtype Q1 Q2 (union-subtype P2 Q2) + ( transitive-leq-subtype Q1 Q2 (P2 ∪ Q2) ( subtype-union-right P2 Q2) ( Q1-sub-Q2)) @@ -217,13 +192,9 @@ module _ {l2 l3 l4 : Level} (P1 : subtype l2 X) (P2 : subtype l3 X) (Q : subtype l4 X) → P1 ⊆ P2 → - union-subtype P1 Q ⊆ union-subtype P2 Q + P1 ∪ Q ⊆ P2 ∪ Q subset-union-subset-left P1 P2 Q P1-sub-P2 = - subtype-union-both P1 Q (union-subtype P2 Q) - ( transitive-leq-subtype P1 P2 (union-subtype P2 Q) - ( subtype-union-left P2 Q) - ( P1-sub-P2)) - ( subtype-union-right P2 Q) + subset-union-subsets P1 Q P2 Q P1-sub-P2 (refl-leq-subtype Q) subset-union-subset-right : {l2 l3 l4 : Level} @@ -231,11 +202,7 @@ module _ Q1 ⊆ Q2 → union-subtype P Q1 ⊆ union-subtype P Q2 subset-union-subset-right P Q1 Q2 Q1-sub-Q2 = - subtype-union-both P Q1 (union-subtype P Q2) - ( subtype-union-left P Q2) - ( transitive-leq-subtype Q1 Q2 (union-subtype P Q2) - ( subtype-union-right P Q2) - ( Q1-sub-Q2)) + subset-union-subsets P Q1 P Q2 (refl-leq-subtype P) Q1-sub-Q2 module _ {l1 l2 l3 l4 : Level} {X : UU l1} @@ -243,21 +210,12 @@ module _ where union-swap-1-2 : - union-subtype P (union-subtype Q S) ⊆ union-subtype Q (union-subtype P S) + P ∪ (Q ∪ S) ⊆ Q ∪ (P ∪ S) union-swap-1-2 = - transitive-leq-subtype - ( union-subtype P (union-subtype Q S)) - ( union-subtype (union-subtype Q P) S) - ( union-subtype Q (union-subtype P S)) + transitive-leq-subtype (P ∪ (Q ∪ S)) ((Q ∪ P) ∪ S) (Q ∪ (P ∪ S)) ( backward-subset-union-assoc Q P S) - ( transitive-leq-subtype - ( union-subtype P (union-subtype Q S)) - ( union-subtype (union-subtype P Q) S) - ( union-subtype (union-subtype Q P) S) - ( subset-union-subset-left - ( union-subtype P Q) - ( union-subtype Q P) - ( S) + ( transitive-leq-subtype (P ∪ (Q ∪ S)) ((P ∪ Q) ∪ S) ((Q ∪ P) ∪ S) + ( subset-union-subset-left (P ∪ Q) (Q ∪ P) S ( subset-union-comm P Q)) ( forward-subset-union-assoc P Q S)) diff --git a/src/lists/lists-subtypes.lagda.md b/src/lists/lists-subtypes.lagda.md index e2ef1b4afc..8c56701be9 100644 --- a/src/lists/lists-subtypes.lagda.md +++ b/src/lists/lists-subtypes.lagda.md @@ -149,13 +149,13 @@ module _ {l2 l3 : Level} (xs : list A) (a : subtype l2 A) (b : subtype l3 A) → list-subtype xs ⊆ union-subtype a b → - ∃ ( list A × list A) + exists-structure ( list A × list A) ( λ (xsl , xsr) → ( list-subtype xs ⊆ list-subtype xsl ∪ list-subtype xsr) × ( list-subtype xsl ⊆ a) × ( list-subtype xsr ⊆ b)) lists-in-union-lists nil a b sub = - intro-∃ (nil , nil) + intro-exists (nil , nil) ( triple ( subset-list-subtype-nil (list-subtype nil ∪ list-subtype nil)) ( subset-list-subtype-nil a) @@ -169,55 +169,54 @@ module _ ( union-subtype a b) ( leq) ( subset-tail-list-subtype))) - ( ∃-Prop _ _) + ( exists-structure-Prop _ _) ( λ ((xsl , xsr) , leq-lists , leq-xsl , leq-xsr) → - ( elim-disjunction-Prop (a x) (b x) (∃-Prop _ _) - ( pair - ( λ x-in-a → - ( intro-∃ (cons x xsl , xsr) - ( triple - ( subset-list-subtype-cons + ( elim-disjunction (exists-structure-Prop _ _) + ( λ x-in-a → + ( intro-exists (cons x xsl , xsr) + ( triple + ( subset-list-subtype-cons + ( list-subtype (cons x xsl) ∪ list-subtype xsr) + ( subtype-union-left + ( list-subtype (cons x xsl)) + ( list-subtype xsr) + ( x) + ( head-in-list-subtype)) + ( transitive-leq-subtype + ( list-subtype xs) + ( list-subtype xsl ∪ list-subtype xsr) ( list-subtype (cons x xsl) ∪ list-subtype xsr) - ( subtype-union-left + ( subset-union-subset-left + ( list-subtype xsl) ( list-subtype (cons x xsl)) ( list-subtype xsr) - ( x) - ( head-in-list-subtype)) - ( transitive-leq-subtype - ( list-subtype xs) - ( list-subtype xsl ∪ list-subtype xsr) - ( list-subtype (cons x xsl) ∪ list-subtype xsr) - ( subset-union-subset-left - ( list-subtype xsl) - ( list-subtype (cons x xsl)) - ( list-subtype xsr) - ( subset-tail-list-subtype)) - ( leq-lists))) - ( backward-implication (iff-subset-head-tail x xsl a) - ( x-in-a , leq-xsl)) - ( leq-xsr)))) - ( λ x-in-b → - ( intro-∃ (xsl , cons x xsr) - ( triple - ( subset-list-subtype-cons + ( subset-tail-list-subtype)) + ( leq-lists))) + ( backward-implication (iff-subset-head-tail x xsl a) + ( x-in-a , leq-xsl)) + ( leq-xsr)))) + ( λ x-in-b → + ( intro-exists (xsl , cons x xsr) + ( triple + ( subset-list-subtype-cons + ( list-subtype xsl ∪ list-subtype (cons x xsr)) + ( subtype-union-right + ( list-subtype xsl) + ( list-subtype (cons x xsr)) + ( x) + ( head-in-list-subtype)) + ( transitive-leq-subtype + ( list-subtype xs) + ( list-subtype xsl ∪ list-subtype xsr) ( list-subtype xsl ∪ list-subtype (cons x xsr)) - ( subtype-union-right + ( subset-union-subset-right ( list-subtype xsl) + ( list-subtype xsr) ( list-subtype (cons x xsr)) - ( x) - ( head-in-list-subtype)) - ( transitive-leq-subtype - ( list-subtype xs) - ( list-subtype xsl ∪ list-subtype xsr) - ( list-subtype xsl ∪ list-subtype (cons x xsr)) - ( subset-union-subset-right - ( list-subtype xsl) - ( list-subtype xsr) - ( list-subtype (cons x xsr)) - ( subset-tail-list-subtype)) - ( leq-lists))) - ( leq-xsl) - ( backward-implication (iff-subset-head-tail x xsr b) - ( x-in-b , leq-xsr)))))) + ( subset-tail-list-subtype)) + ( leq-lists))) + ( leq-xsl) + ( backward-implication (iff-subset-head-tail x xsr b) + ( x-in-b , leq-xsr))))) ( leq x head-in-list-subtype))) ``` diff --git a/src/modal-logic/canonical-model-theorem.lagda.md b/src/modal-logic/canonical-model-theorem.lagda.md index b27845e7f8..4ee33fe8f2 100644 --- a/src/modal-logic/canonical-model-theorem.lagda.md +++ b/src/modal-logic/canonical-model-theorem.lagda.md @@ -497,7 +497,7 @@ module _ ( f y xRy))) canonical-model-theorem : - (a : formula i) → logic a ⇔ (canonical-kripke-model ⊨M a) + (a : formula i) → type-iff-Prop (logic a) (canonical-kripke-model ⊨M a) pr1 (canonical-model-theorem a) in-logic x = forward-implication ( canonical-model-theorem-pointwise a x) diff --git a/src/modal-logic/completeness-k.lagda.md b/src/modal-logic/completeness-k.lagda.md index 23fc1c1437..a622cc4e7b 100644 --- a/src/modal-logic/completeness-k.lagda.md +++ b/src/modal-logic/completeness-k.lagda.md @@ -74,13 +74,4 @@ module _ ( prop-resize) ( all-models (lsuc l1) l1 i l1) ( star) - -- canonical-model-completness - -- ( modal-logic-K-axioms i) - -- ( zorn) - -- ( prop-resize) - -- ( is-consistent-K i) - -- ( refl-leq-subtype (modal-logic-K i)) - -- ( lem) - -- ( all-models (lsuc l1) l1 i l1) - -- ( star) ``` diff --git a/src/modal-logic/finite-approximability.lagda.md b/src/modal-logic/finite-approximability.lagda.md index b7bff428c4..dc8bbfbcca 100644 --- a/src/modal-logic/finite-approximability.lagda.md +++ b/src/modal-logic/finite-approximability.lagda.md @@ -71,7 +71,7 @@ module _ modal-theory l2 i → Prop (l1 ⊔ l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ lsuc l6) is-finitely-approximable-Prop l3 l4 l5 l6 logic = - ∃-Prop + exists-structure-Prop ( model-class l3 l4 i l5 l6) ( λ C → ( product @@ -107,7 +107,7 @@ module _ ( filtration-models _ _ i _ _ _ _ _ M*) ( λ ((a , M , _) , p) → ( tr (is-in-subtype (filtration-models _ _ i _ _ _ _ _)) (inv p) - ( intro-∃ + ( intro-exists ( subformulas i a , M) ( pair (is-finite-subformulas-list @@ -156,7 +156,7 @@ module _ ( lsuc (lsuc (lsuc l1))) ( modal-logic-K i) is-finitely-approximable-K = - intro-∃ + intro-exists ( K-finite-class) ( pair ( soundness-subclass diff --git a/src/modal-logic/kripke-models-filtrations.lagda.md b/src/modal-logic/kripke-models-filtrations.lagda.md index 93d47aa1e0..61d871eb51 100644 --- a/src/modal-logic/kripke-models-filtrations.lagda.md +++ b/src/modal-logic/kripke-models-filtrations.lagda.md @@ -202,7 +202,7 @@ module _ ( (a , in-theory) : type-subtype theory) ( (x , _) : type-subtype (subtype-equivalence-class Φ-equivalence class)) → - f (a , in-theory) ⇔ ((M , x) ⊨ a)) + type-iff-Prop (f (a , in-theory)) ((M , x) ⊨ a)) is-prop-valuate-function-equivalence-class : (class : equivalence-class Φ-equivalence) → @@ -508,7 +508,7 @@ module _ minimal-kripke-model-filtration-relation : Relation-Prop (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) (equivalence-class Φ-equivalence) minimal-kripke-model-filtration-relation x* y* = - ∃-Prop + exists-structure-Prop ( type-kripke-model i M × type-kripke-model i M) ( λ (x , y) → ( product @@ -596,12 +596,12 @@ module _ ( class Φ-equivalence x) ( class Φ-equivalence y)) proof-lower-bound x y r = - intro-∃ (x , y) (r , (λ _ _ → id , id) , (λ _ _ → id , id)) + intro-exists (x , y) (r , (λ _ _ → id , id) , (λ _ _ → id , id)) is-kripke-model-filtration-minimal-kripke-model-filtration : is-kripke-model-filtration minimal-kripke-model-filtration is-kripke-model-filtration-minimal-kripke-model-filtration = - intro-∃ + intro-exists ( id-equiv) ( triple ( λ n x → @@ -674,7 +674,7 @@ module _ is-kripke-model-filtration minimal-transitive-kripke-model-filtration is-kripke-model-filtration-minimal-transitive-kripke-model-filtration M-is-trans = - intro-∃ + intro-exists ( id-equiv) ( triple ( λ n x → @@ -712,7 +712,7 @@ module _ map-universal-property-trunc-Prop ( relation-Prop-kripke-model i minimal-kripke-model-filtration y* x*) ( λ ((x , y) , r-xy , x-in-x* , y-in-y*) → - ( intro-∃ (y , x) (is-sym x y r-xy , y-in-y* , x-in-x*))) + ( intro-exists (y , x) (is-sym x y r-xy , y-in-y* , x-in-x*))) minimal-filtration-preserves-symmetry : is-in-subtype (symmetry-kripke-class l1 l2 i l4) M → @@ -793,7 +793,7 @@ module _ model-class l1 l2 i l4 (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4 ⊔ lsuc l5 ⊔ lsuc l6 ⊔ lsuc l7 ⊔ lsuc l8) filtration-models M* = - ∃-Prop + exists-structure-Prop ( modal-theory l5 i × kripke-model l6 l7 i l8) ( λ (theory , M) → ( product diff --git a/src/modal-logic/kripke-semantics.lagda.md b/src/modal-logic/kripke-semantics.lagda.md index d8754eed50..34724ec443 100644 --- a/src/modal-logic/kripke-semantics.lagda.md +++ b/src/modal-logic/kripke-semantics.lagda.md @@ -146,7 +146,8 @@ module _ where is-serial : UU (l1 ⊔ l2) - is-serial = (x : A) → ∃ A (λ y → R x y) + -- is-serial = (x : A) → ∃ A (λ y → R x y) + is-serial = (x : A) → exists-structure A (λ y → R x y) is-euclidean : UU (l1 ⊔ l2) is-euclidean = (x y z : A) → R x y → R x z → R y z @@ -160,7 +161,7 @@ module _ is-prop-is-serial-Relation-Prop : is-prop is-serial-Relation-Prop is-prop-is-serial-Relation-Prop = - is-prop-Π (λ x → is-prop-∃ A _) + is-prop-Π (λ x → is-prop-exists-structure A _) is-euclidean-Relation-Prop : UU (l1 ⊔ l2) is-euclidean-Relation-Prop = is-euclidean (type-Relation-Prop R) @@ -244,7 +245,7 @@ module _ Prop (l1 ⊔ l2 ⊔ l4) (M , x) ⊨ var n = raise-Prop (l1 ⊔ l2) (valuate-kripke-model i M n x) (M , x) ⊨ ⊥ = raise-empty-Prop (l1 ⊔ l2 ⊔ l4) - (M , x) ⊨ a →ₘ b = implication-Prop ((M , x) ⊨ a) ((M , x) ⊨ b) + (M , x) ⊨ a →ₘ b = hom-Prop ((M , x) ⊨ a) ((M , x) ⊨ b) (M , x) ⊨ □ a = Π-Prop ( type-kripke-model i M) diff --git a/src/modal-logic/l-complete-theories.lagda.md b/src/modal-logic/l-complete-theories.lagda.md index f76510e8a4..e7f71de3c8 100644 --- a/src/modal-logic/l-complete-theories.lagda.md +++ b/src/modal-logic/l-complete-theories.lagda.md @@ -337,16 +337,15 @@ module _ ( subset-logic-L-complete-theory lzero x (◇ ~ a →ₘ ~ □ a) ( modal-logic-diamond-negate-implication i logic is-normal ( is-logic))) - ( leq (◇ ~ a) (intro-∃ (~ a) (not-a-in-t , refl)))) + ( leq (◇ ~ a) (intro-exists (~ a) (not-a-in-t , refl)))) ( box-a-in-x))) is-inhabited-L-complete-exists-complete-L-consistent-theory : {l3 : Level} → - ∃ (L-consistent-theory logic l3) is-L-complete-theory → + exists (L-consistent-theory logic l3) is-L-complete-theory-Prop → is-inhabited (L-complete-theory l3) is-inhabited-L-complete-exists-complete-L-consistent-theory {l3} = - elim-exists-Prop - ( is-L-complete-theory-Prop) + elim-exists ( is-inhabited-Prop (L-complete-theory l3)) ( λ theory is-comp → unit-trunc-Prop (theory , is-comp)) @@ -393,17 +392,14 @@ module _ chain-union-modal-theory : modal-theory (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4) i chain-union-modal-theory a = - ∃-Prop (type-chain-Poset P C) - ( λ x → is-in-subtype (modal-theory-chain-element x) a) + ∃ (type-chain-Poset P C) (λ x → modal-theory-chain-element x a) exists-chain-element-with-formula-Prop : (a : formula i) → Prop (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4) exists-chain-element-with-formula-Prop a = - ∃-Prop (type-chain-Poset P C) + ∃ (type-chain-Poset P C) ( λ x → - ( is-in-subtype - ( weak-modal-logic-closure (logic ∪ modal-theory-chain-element x)) - ( a))) + ( weak-modal-logic-closure (logic ∪ modal-theory-chain-element x) a)) exists-chain-element-with-formula : (a : formula i) → UU (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4) @@ -471,7 +467,7 @@ module _ ( is-inh) ( exists-chain-element-with-formula-Prop a) ( λ x → - ( intro-∃ x + ( intro-exists x ( weak-modal-logic-closure-monotic ( subtype-union-both logic (list-subtype nil) ( logic ∪ modal-theory-chain-element x) @@ -493,26 +489,25 @@ module _ ( L-union-deduction-lemma l h a in-logic)) ( exists-chain-element-with-formula-Prop a) ( λ (x , h-in-x) (y , ha-in-y) → - ( elim-disjunction-Prop - ( leq-Poset-Prop P - ( type-Poset-type-chain-Poset P C x) - ( type-Poset-type-chain-Poset P C y)) - ( leq-Poset-Prop P - ( type-Poset-type-chain-Poset P C y) - ( type-Poset-type-chain-Poset P C x)) + ( elim-disjunction + -- ( leq-Poset-Prop P + -- ( type-Poset-type-chain-Poset P C x) + -- ( type-Poset-type-chain-Poset P C y)) + -- ( leq-Poset-Prop P + -- ( type-Poset-type-chain-Poset P C y) + -- ( type-Poset-type-chain-Poset P C x)) ( exists-chain-element-with-formula-Prop a) - ( pair - ( λ x-leq-y → - ( intro-∃ y - ( weak-modal-logic-closure-mp - ( ha-in-y) - ( leq-L-union x y x-leq-y h - ( theory-subset-L-union x h h-in-x))))) - ( λ y-leq-x → - ( intro-∃ x - ( weak-modal-logic-closure-mp - ( leq-L-union y x y-leq-x (h →ₘ a) ha-in-y) + ( λ x-leq-y → + ( intro-exists y + ( weak-modal-logic-closure-mp + ( ha-in-y) + ( leq-L-union x y x-leq-y h ( theory-subset-L-union x h h-in-x))))) + ( λ y-leq-x → + ( intro-exists x + ( weak-modal-logic-closure-mp + ( leq-L-union y x y-leq-x (h →ₘ a) ha-in-y) + ( theory-subset-L-union x h h-in-x)))) ( is-chain-Subposet-chain-Poset P C x y))) in-chain-in-chain-union : @@ -633,7 +628,7 @@ module _ ( resized-chain-union-modal-theory C) ( chain-union-modal-theory C) ( equiv-resized-chain-union-modal-theory C)) - ( λ a in-theory → intro-∃ x in-theory) + ( λ a in-theory → intro-exists x in-theory) extend-L-consistent-theory : (zorn : Zorn-non-empty (l1 ⊔ l2 ⊔ lsuc l3) (l1 ⊔ l3) l4) → @@ -645,7 +640,7 @@ module _ ( L-consistent-theories-Poset logic l3) ( is-inh) ( λ C C-is-inh → - ( intro-∃ + ( intro-exists ( resized-chain-union-L-consistent-theory C C-is-inh) ( union-is-chain-upper-bound C C-is-inh)))) diff --git a/src/modal-logic/lindenbaum.lagda.md b/src/modal-logic/lindenbaum.lagda.md index e8021bc781..1361ed879a 100644 --- a/src/modal-logic/lindenbaum.lagda.md +++ b/src/modal-logic/lindenbaum.lagda.md @@ -48,10 +48,11 @@ module _ where lindenbaum : - ∃ ( L-complete-theory logic (l1 ⊔ l2 ⊔ l3)) + exists (L-complete-theory logic (l1 ⊔ l2 ⊔ l3)) ( λ y → - ( modal-theory-L-consistent-theory logic x ⊆ - modal-theory-L-complete-theory logic y)) + ( leq-prop-subtype + ( modal-theory-L-consistent-theory logic x) + ( modal-theory-L-complete-theory logic y))) lindenbaum = apply-universal-property-trunc-Prop ( extend-L-consistent-theory L prop-resize contains-ax-k' contains-ax-s' @@ -59,9 +60,13 @@ module _ ( is-inhabited-L-consistent-theory L ( is-weak-modal-logic-weak-modal-logic-closure) ( is-cons))) - ( ∃-Prop _ _) + ( ∃ (L-complete-theory logic (l1 ⊔ l2 ⊔ l3)) + ( λ y → + ( leq-prop-subtype + ( modal-theory-L-consistent-theory logic x) + ( modal-theory-L-complete-theory logic y)))) ( λ y → - ( intro-∃ + ( intro-exists ( result-L-complete y) ( subset-theory-transofrm-L-complete y))) where diff --git a/src/modal-logic/logic-syntax.lagda.md b/src/modal-logic/logic-syntax.lagda.md index d552e4d7f9..35ca1441b8 100644 --- a/src/modal-logic/logic-syntax.lagda.md +++ b/src/modal-logic/logic-syntax.lagda.md @@ -271,7 +271,9 @@ module _ diamond-modal-theory : modal-theory l2 i → modal-theory (l1 ⊔ l2) i diamond-modal-theory theory a = - ∃-Prop (formula i) (λ b → is-in-subtype theory b × (a = ◇ b)) + exists-structure-Prop + ( formula i) + ( λ b → is-in-subtype theory b × (a = ◇ b)) module _ {l1 : Level} {i : Set l1} diff --git a/src/modal-logic/modal-logic-decision.lagda.md b/src/modal-logic/modal-logic-decision.lagda.md index 6bf9173d74..99234f755b 100644 --- a/src/modal-logic/modal-logic-decision.lagda.md +++ b/src/modal-logic/modal-logic-decision.lagda.md @@ -107,7 +107,7 @@ module _ soundness theory C → completeness theory C → (a : formula i) → - theory a ⇔ prop-bool (decision-procedure a) + is-in-subtype theory a ↔ type-prop-bool (decision-procedure a) pr1 (decision-procedure-correctness sound complete a) in-theory with decision-procedure' a ... | inl _ = star @@ -120,7 +120,7 @@ module _ -- TODO: move to kuratowsky-finite-sets is-kuratowsky-finite-set-Prop' : {l : Level} → Set l → Prop l is-kuratowsky-finite-set-Prop' X = - ∃-Prop (list (type-Set X)) + exists-structure-Prop (list (type-Set X)) ( λ l → (x : type-Set X) → type-Prop (list-subtype l x)) is-kuratowsky-finite-set' : {l : Level} → Set l → UU l @@ -133,7 +133,7 @@ is-kuratowsky-finite-set-is-kuratowsky-finite-set' X = map-universal-property-trunc-Prop ( is-kuratowsky-finite-set-Prop X) ( λ (l , all-list-subtype) → - ( intro-∃ + ( intro-exists ( length-list l) ( pair ( component-list l) @@ -284,7 +284,7 @@ module _ is-kuratowsky-finite'-subformulas-list : (a : formula i) → is-kuratowsky-finite-set' (subformulas-Set a) is-kuratowsky-finite'-subformulas-list a = - intro-∃ + intro-exists ( subformulas-Set-list a) ( λ (b , trunc-b-list-subtype) → ( apply-universal-property-trunc-Prop @@ -549,7 +549,7 @@ module _ filtrate-class : model-class l6 l7 i l8 ( l3 ⊔ l5 ⊔ lsuc (l1 ⊔ l2 ⊔ l4 ⊔ l6 ⊔ l7 ⊔ l8)) filtrate-class M* = - ∃-Prop (formula i × type-subtype C) + exists-structure-Prop (formula i × type-subtype C) ( λ (a , (M , _)) → M* = filtration (subformulas i a) M) module _ @@ -588,7 +588,7 @@ module _ ( x)) ( in-logic ( filtration (list-subtype (cons a _)) M) - ( intro-∃ (a , (M , M-in-class)) refl) + ( intro-exists (a , (M , M-in-class)) refl) ( map-equiv-is-kripke-model-filtration ( i) ( subformulas i a) diff --git a/src/modal-logic/modal-logic-k.lagda.md b/src/modal-logic/modal-logic-k.lagda.md index dfa7284117..05aa4f94d7 100644 --- a/src/modal-logic/modal-logic-k.lagda.md +++ b/src/modal-logic/modal-logic-k.lagda.md @@ -45,8 +45,9 @@ module _ {l : Level} (i : Set l) where + -- TODO: refactor modal-logic-K-axioms : modal-theory l i - modal-logic-K-axioms = ax-k i ∪ ax-s i ∪ ax-n i ∪ ax-dn i + modal-logic-K-axioms = ((ax-k i ∪ ax-s i) ∪ ax-n i) ∪ ax-dn i modal-logic-K : modal-theory l i modal-logic-K = modal-logic-closure modal-logic-K-axioms @@ -62,9 +63,9 @@ module _ ( modal-logic-K-axioms) ( transitive-leq-subtype ( ax-k i ∪ ax-s i) - ( ax-k i ∪ ax-s i ∪ ax-n i) + ( (ax-k i ∪ ax-s i) ∪ ax-n i) ( modal-logic-K-axioms) - ( subtype-union-left (ax-k i ∪ ax-s i ∪ ax-n i) (ax-dn i)) + ( subtype-union-left ((ax-k i ∪ ax-s i) ∪ ax-n i) (ax-dn i)) ( subtype-union-left (ax-k i ∪ ax-s i) (ax-n i))) ( subtype-union-left (ax-k i) (ax-s i)) @@ -76,9 +77,9 @@ module _ ( modal-logic-K-axioms) ( transitive-leq-subtype ( ax-k i ∪ ax-s i) - ( ax-k i ∪ ax-s i ∪ ax-n i) + ( (ax-k i ∪ ax-s i) ∪ ax-n i) ( modal-logic-K-axioms) - ( subtype-union-left (ax-k i ∪ ax-s i ∪ ax-n i) (ax-dn i)) + ( subtype-union-left ((ax-k i ∪ ax-s i) ∪ ax-n i) (ax-dn i)) ( subtype-union-left (ax-k i ∪ ax-s i) (ax-n i))) ( subtype-union-right (ax-k i) (ax-s i)) @@ -86,14 +87,14 @@ module _ K-axioms-contains-ax-n = transitive-leq-subtype ( ax-n i) - ( ax-k i ∪ ax-s i ∪ ax-n i) + ( (ax-k i ∪ ax-s i) ∪ ax-n i) ( modal-logic-K-axioms) - ( subtype-union-left (ax-k i ∪ ax-s i ∪ ax-n i) (ax-dn i)) + ( subtype-union-left ((ax-k i ∪ ax-s i) ∪ ax-n i) (ax-dn i)) ( subtype-union-right (ax-k i ∪ ax-s i) (ax-n i)) K-axioms-contains-ax-dn : ax-dn i ⊆ modal-logic-K-axioms K-axioms-contains-ax-dn = - subtype-union-right (ax-k i ∪ ax-s i ∪ ax-n i) (ax-dn i) + subtype-union-right ((ax-k i ∪ ax-s i) ∪ ax-n i) (ax-dn i) K-contains-ax-k : ax-k i ⊆ modal-logic-K K-contains-ax-k = @@ -147,7 +148,7 @@ module _ soundness (modal-logic-K-axioms i) (decidable-models l2 l3 i l4) soundness-K-axioms = soundness-union-subclass-right-sublevels - ( ax-k i ∪ ax-s i ∪ ax-n i) + ( (ax-k i ∪ ax-s i) ∪ ax-n i) ( ax-dn i) ( l1 ⊔ l2 ⊔ l3 ⊔ l4) ( all-models l2 l3 i l4) diff --git a/src/modal-logic/modal-logic-s5.lagda.md b/src/modal-logic/modal-logic-s5.lagda.md index f3e74e3a91..c1b2afced3 100644 --- a/src/modal-logic/modal-logic-s5.lagda.md +++ b/src/modal-logic/modal-logic-s5.lagda.md @@ -45,7 +45,7 @@ module _ where modal-logic-S5-axioms : modal-theory l i - modal-logic-S5-axioms = modal-logic-K-axioms i ∪ (ax-m i ∪ (ax-b i ∪ ax-4 i)) + modal-logic-S5-axioms = modal-logic-K-axioms i ∪ ax-m i ∪ ax-b i ∪ ax-4 i modal-logic-S5 : modal-theory l i modal-logic-S5 = modal-logic-closure modal-logic-S5-axioms diff --git a/src/modal-logic/weak-deduction.lagda.md b/src/modal-logic/weak-deduction.lagda.md index 56436086db..cc5c68bf9f 100644 --- a/src/modal-logic/weak-deduction.lagda.md +++ b/src/modal-logic/weak-deduction.lagda.md @@ -321,8 +321,9 @@ module _ deduction-lemma : (a b : formula i) → - weak-modal-logic-closure (theory-add-formula a axioms) b ⇔ - weak-modal-logic-closure axioms (a →ₘ b) + type-iff-Prop + ( weak-modal-logic-closure (theory-add-formula a axioms) b) + ( weak-modal-logic-closure axioms (a →ₘ b)) pr1 (deduction-lemma a b) = map-universal-property-trunc-Prop ( weak-modal-logic-closure axioms (a →ₘ b)) diff --git a/src/order-theory.lagda.md b/src/order-theory.lagda.md index e8a88dd031..2ce1ea63d7 100644 --- a/src/order-theory.lagda.md +++ b/src/order-theory.lagda.md @@ -10,7 +10,6 @@ open import order-theory.bottom-elements-posets public open import order-theory.bottom-elements-preorders public open import order-theory.chains-posets public open import order-theory.chains-preorders public -open import order-theory.chains-union public open import order-theory.closure-operators-large-locales public open import order-theory.closure-operators-large-posets public open import order-theory.commuting-squares-of-galois-connections-large-posets public diff --git a/src/order-theory/chains-posets.lagda.md b/src/order-theory/chains-posets.lagda.md index 72ce9a41dc..2361ce49d2 100644 --- a/src/order-theory/chains-posets.lagda.md +++ b/src/order-theory/chains-posets.lagda.md @@ -106,7 +106,7 @@ module _ is-chain-upper-bound = type-Prop ∘ is-chain-upper-bound-Prop has-chain-upper-bound-Prop : Prop (l1 ⊔ l2 ⊔ l3) - has-chain-upper-bound-Prop = ∃-Prop (type-Poset X) is-chain-upper-bound + has-chain-upper-bound-Prop = ∃ (type-Poset X) is-chain-upper-bound-Prop has-chain-upper-bound : UU (l1 ⊔ l2 ⊔ l3) has-chain-upper-bound = type-Prop has-chain-upper-bound-Prop diff --git a/src/order-theory/chains-union.lagda.md b/src/order-theory/chains-union.lagda.md deleted file mode 100644 index 3acee548c3..0000000000 --- a/src/order-theory/chains-union.lagda.md +++ /dev/null @@ -1,121 +0,0 @@ -# Chain union - -```agda -module order-theory.chains-union where -``` - -
Imports - -```agda -open import foundation.dependent-pair-types -open import foundation.equivalences -open import foundation.existential-quantification -open import foundation.function-types -open import foundation.propositional-resizing -open import foundation.propositional-truncations -open import foundation.propositions -open import foundation.subtypes -open import foundation.unit-type -open import foundation.universe-levels - -open import order-theory.chains-posets -open import order-theory.maximal-elements-posets -open import order-theory.posets -open import order-theory.subposets -open import order-theory.subtypes-leq-posets -``` - -
- -## Idea - -TODO - -## Definition - -```agda -module _ - {l1 l2 l3 : Level} (A : UU l1) (C : chain-Poset l3 (subtypes-leq-Poset l2 A)) - where - - in-chain-union : subtype (l1 ⊔ lsuc l2 ⊔ l3) A - in-chain-union a = - exists-Prop - ( type-chain-Poset (subtypes-leq-Poset l2 A) C) - ( λ x → - ( inclusion-subtype - ( sub-preorder-chain-Poset (subtypes-leq-Poset l2 A) C) - ( x) - ( a))) - - is-chain-union : - {l4 : Level} (u : subtype l4 A) → UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ l4) - is-chain-union = equiv-subtypes in-chain-union - - chain-union-is-upper-bound : - (u : subtype l2 A) → - is-chain-union u → - is-chain-upper-bound (subtypes-leq-Poset l2 A) C u - chain-union-is-upper-bound u u-is-union s a a-in-s = - map-equiv (u-is-union a) (intro-∃ s a-in-s) - - module _ - (prop-resize : propositional-resizing l2 (l1 ⊔ lsuc l2 ⊔ l3)) - where - - chain-union : Σ (subtype l2 A) is-chain-union - pr1 chain-union = resize prop-resize ∘ in-chain-union - pr2 chain-union = inv-equiv ∘ is-equiv-resize prop-resize ∘ in-chain-union - -module _ - {l1 l2 l3 l4 : Level} (A : UU l1) - (S : subtype l3 (subtype l2 A)) - (C : chain-Poset l4 (poset-Subposet (subtypes-leq-Poset l2 A) S)) - where - - private - poset : Poset (l1 ⊔ lsuc l2 ⊔ l3) (l1 ⊔ l2) - poset = poset-Subposet (subtypes-leq-Poset l2 A) S - - in-sub-chain-union : subtype (l1 ⊔ lsuc l2 ⊔ l3 ⊔ l4) A - in-sub-chain-union a = - exists-Prop - ( type-chain-Poset poset C) - ( λ x → - ( inclusion-subtype S - ( inclusion-subtype (sub-preorder-chain-Poset poset C) x) a)) - - is-sub-chain-union : (u : type-Poset poset) → UU (l1 ⊔ lsuc l2 ⊔ l3 ⊔ l4) - is-sub-chain-union u = - equiv-subtypes in-sub-chain-union (inclusion-subtype S u) - - sub-chain-union-is-upper-bound : - (u : type-Poset poset) → - is-sub-chain-union u → - is-chain-upper-bound poset C u - sub-chain-union-is-upper-bound u u-is-union s a a-in-s = - map-equiv (u-is-union a) (intro-∃ s a-in-s) - - module _ - (x : type-chain-Poset poset C) - (subtype-monotic : - (y : subtype l2 A) → - inclusion-subtype S - ( inclusion-subtype (sub-preorder-chain-Poset poset C) x) ⊆ y → - is-in-subtype S y) - (prop-resize : propositional-resizing l2 (l1 ⊔ lsuc l2 ⊔ l3 ⊔ l4)) - where - - sub-chain-union : Σ (type-Poset poset) is-sub-chain-union - pr1 (pr1 sub-chain-union) = - resize prop-resize ∘ in-sub-chain-union - pr2 (pr1 sub-chain-union) = - subtype-monotic - ( resize prop-resize ∘ in-sub-chain-union) - ( λ a a-in-x → - ( map-equiv - ( inv-equiv (is-equiv-resize prop-resize (in-sub-chain-union a))) - ( intro-∃ x a-in-x))) - pr2 sub-chain-union = - inv-equiv ∘ is-equiv-resize prop-resize ∘ in-sub-chain-union -``` diff --git a/src/order-theory/zorn.lagda.md b/src/order-theory/zorn.lagda.md index d0e709d67b..01977b413f 100644 --- a/src/order-theory/zorn.lagda.md +++ b/src/order-theory/zorn.lagda.md @@ -44,7 +44,7 @@ module _ ( λ X → ( function-Prop ( (C : chain-Poset l3 X) → has-chain-upper-bound X C) - ( ∃-Prop (type-Poset X) (is-maximal-element-Poset X)))) + ( ∃ (type-Poset X) (is-maximal-element-Poset-Prop X)))) Zorn : UU (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3) Zorn = type-Prop Zorn-Prop @@ -60,7 +60,7 @@ module _ ( (C : chain-Poset l3 X) → is-inhabited (type-chain-Poset X C) → has-chain-upper-bound X C) - ( ∃-Prop (type-Poset X) (is-maximal-element-Poset X))))) + ( ∃ (type-Poset X) (is-maximal-element-Poset-Prop X))))) Zorn-non-empty : UU (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3) Zorn-non-empty = type-Prop Zorn-non-empty-Prop @@ -92,10 +92,11 @@ module _ ( is-inhabited-X) ( has-chain-upper-bound-Prop X C) ( λ x → - ( intro-∃ x - ( λ (y , y-in-C) → ex-falso (is-empty-C (intro-∃ y y-in-C))))) + ( intro-exists x + ( λ (y , y-in-C) → + ( ex-falso (is-empty-C (intro-exists y y-in-C)))))) - iff-Zorn-non-empty-Zorn : Zorn-non-empty-Prop ⇔ Zorn-Prop + iff-Zorn-non-empty-Zorn : type-iff-Prop Zorn-non-empty-Prop Zorn-Prop pr1 (iff-Zorn-non-empty-Zorn) = Zorn-Zorn-non-empty pr2 (iff-Zorn-non-empty-Zorn) = Zorn-non-empty-Zorn ``` diff --git a/src/univalent-combinatorics/kuratowsky-finite-sets.lagda.md b/src/univalent-combinatorics/kuratowsky-finite-sets.lagda.md index 78a8b2a4c6..4ca2dce32d 100644 --- a/src/univalent-combinatorics/kuratowsky-finite-sets.lagda.md +++ b/src/univalent-combinatorics/kuratowsky-finite-sets.lagda.md @@ -103,7 +103,7 @@ is-kuratowsky-finite-set-surjection : is-kuratowsky-finite-set-surjection X Y f = map-universal-property-trunc-Prop ( is-kuratowsky-finite-set-Prop Y) - ( λ (n , g) → (intro-∃ n (surjection-comp f g))) + ( λ (n , g) → (intro-exists n (surjection-comp f g))) is-kuratowsky-finite-set-is-finite : {l : Level} (X : Set l) → @@ -112,7 +112,7 @@ is-kuratowsky-finite-set-is-finite : is-kuratowsky-finite-set-is-finite X = map-universal-property-trunc-Prop ( is-kuratowsky-finite-set-Prop X) - ( λ (n , e) → intro-∃ n (map-equiv e , is-surjective-map-equiv e)) + ( λ (n , e) → intro-exists n (map-equiv e , is-surjective-map-equiv e)) ``` ### Classical facts From 424667a4f6ee9de0a47f6dc420dd4c9d6175ac26 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Sun, 5 May 2024 17:07:07 +0300 Subject: [PATCH 48/63] Small refactor filtrations --- ...kripke-models-filtrations-theorem.lagda.md | 2 +- .../kripke-models-filtrations.lagda.md | 136 +++++++-------- src/modal-logic/modal-logic-decision.lagda.md | 162 +++++------------- 3 files changed, 100 insertions(+), 200 deletions(-) diff --git a/src/modal-logic/kripke-models-filtrations-theorem.lagda.md b/src/modal-logic/kripke-models-filtrations-theorem.lagda.md index 01026a8dbc..b1ec052e76 100644 --- a/src/modal-logic/kripke-models-filtrations-theorem.lagda.md +++ b/src/modal-logic/kripke-models-filtrations-theorem.lagda.md @@ -58,7 +58,7 @@ module _ where kripke-models-filtrations-theorem' : - (is-filtration : is-kripke-model-filtration-Σ i theory M M*) + (is-filtration : is-kripke-model-filtration i theory M M*) (a : formula i) → is-in-subtype theory a → (x : type-kripke-model i M) → diff --git a/src/modal-logic/kripke-models-filtrations.lagda.md b/src/modal-logic/kripke-models-filtrations.lagda.md index 61d871eb51..45f81455b0 100644 --- a/src/modal-logic/kripke-models-filtrations.lagda.md +++ b/src/modal-logic/kripke-models-filtrations.lagda.md @@ -275,27 +275,31 @@ module _ ( is-set-equivalence-class Φ-equivalence x-class y-class)) ( λ (x , x-in-class) (y , y-in-class) → ( equational-reasoning - x-class = class Φ-equivalence x by - ( inv - ( eq-class-equivalence-class - ( Φ-equivalence) - ( x-class) - ( x-in-class))) - = class Φ-equivalence y by - ( apply-effectiveness-class' - ( Φ-equivalence) - ( λ a a-in-theory → - ( g-val (a , a-in-theory) (y , y-in-class) - ∘iff iff-eq (htpy-eq p (a , a-in-theory)) - ∘iff inv-iff - ( f-val - ( a , a-in-theory) - ( x , x-in-class))))) - = y-class by - ( eq-class-equivalence-class - ( Φ-equivalence) - ( y-class) - ( y-in-class)))) + x-class + = class Φ-equivalence x + by + inv + ( eq-class-equivalence-class + ( Φ-equivalence) + ( x-class) + ( x-in-class)) + = class Φ-equivalence y + by + apply-effectiveness-class' + ( Φ-equivalence) + ( λ a a-in-theory → + ( g-val (a , a-in-theory) (y , y-in-class) + ∘iff iff-eq (htpy-eq p (a , a-in-theory)) + ∘iff inv-iff + ( f-val + ( a , a-in-theory) + ( x , x-in-class)))) + = y-class + by + eq-class-equivalence-class + ( Φ-equivalence) + ( y-class) + ( y-in-class))) injection-map-function-equivalence-class : injection @@ -370,20 +374,9 @@ module _ filtration-relation-upper-bound = type-Prop ∘ filtration-relation-upper-bound-Prop - -- is-alpha-Σ : - -- UU (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6 ⊔ l8) - -- is-alpha-Σ = - -- product - -- ( equivalence-class Φ-equivalence ≃ type-kripke-model i M*) - -- ( is-bounded-valuate (valuate-kripke-model i M*)) - - -- is-alpha-Prop : - -- Prop (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6 ⊔ l8) - -- is-alpha-Prop = trunc-Prop is-alpha-Σ - - is-kripke-model-filtration-Σ : - UU (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6 ⊔ l7 ⊔ l8) - is-kripke-model-filtration-Σ = + is-kripke-model-filtration : + UU (lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) ⊔ l6 ⊔ l7 ⊔ l8) + is-kripke-model-filtration = Σ ( equivalence-class Φ-equivalence ≃ type-kripke-model i M*) ( λ e → ( product @@ -392,44 +385,36 @@ module _ ( filtration-relation-lower-bound e) ( filtration-relation-upper-bound e)))) - is-kripke-model-filtration-Prop : - Prop (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6 ⊔ l7 ⊔ l8) - is-kripke-model-filtration-Prop = trunc-Prop is-kripke-model-filtration-Σ - - is-kripke-model-filtration : - UU (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ l6 ⊔ l7 ⊔ l8) - is-kripke-model-filtration = type-Prop is-kripke-model-filtration-Prop - equiv-is-kripke-model-filtration : - is-kripke-model-filtration-Σ → + is-kripke-model-filtration → equivalence-class Φ-equivalence ≃ type-kripke-model i M* equiv-is-kripke-model-filtration = pr1 map-equiv-is-kripke-model-filtration : - is-kripke-model-filtration-Σ → + is-kripke-model-filtration → equivalence-class Φ-equivalence → type-kripke-model i M* map-equiv-is-kripke-model-filtration = map-equiv ∘ equiv-is-kripke-model-filtration map-inv-equiv-is-kripke-model-filtration : - is-kripke-model-filtration-Σ → + is-kripke-model-filtration → type-kripke-model i M* → equivalence-class Φ-equivalence map-inv-equiv-is-kripke-model-filtration = map-inv-equiv ∘ equiv-is-kripke-model-filtration is-filtration-valuate-is-kripke-model-filtration : - (e : is-kripke-model-filtration-Σ) → + (e : is-kripke-model-filtration) → is-filtration-valuate (equiv-is-kripke-model-filtration e) is-filtration-valuate-is-kripke-model-filtration = pr1 ∘ pr2 filtration-relation-lower-bound-is-kripke-model-filtration : - (e : is-kripke-model-filtration-Σ) → + (e : is-kripke-model-filtration) → filtration-relation-lower-bound (equiv-is-kripke-model-filtration e) filtration-relation-lower-bound-is-kripke-model-filtration = pr1 ∘ pr2 ∘ pr2 filtration-relation-upper-bound-is-kripke-model-filtration : - (e : is-kripke-model-filtration-Σ) → + (e : is-kripke-model-filtration) → filtration-relation-upper-bound (equiv-is-kripke-model-filtration e) filtration-relation-upper-bound-is-kripke-model-filtration = pr2 ∘ pr2 ∘ pr2 @@ -451,7 +436,7 @@ module _ ( is-section-map-section-map-equiv e x*) class-x-eq-x* : - (is-filt : is-kripke-model-filtration-Σ) + (is-filt : is-kripke-model-filtration) (x : type-kripke-model i M) (x* : type-kripke-model i M*) → is-in-equivalence-class Φ-equivalence @@ -479,24 +464,20 @@ module _ is-kripke-model-filtration → is-in-subtype (reflexive-kripke-class l1 l2 i l4) M → is-in-subtype (reflexive-kripke-class l6 l7 i l8) M* - filtration-preserves-reflexivity t-is-filt is-refl class = + filtration-preserves-reflexivity is-filt is-refl class = apply-universal-property-trunc-Prop - ( t-is-filt) + ( is-inhabited-subtype-equivalence-class Φ-equivalence + ( map-inv-equiv-is-kripke-model-filtration is-filt class)) ( relation-Prop-kripke-model i M* class class) - ( λ is-filt → - ( apply-universal-property-trunc-Prop - ( is-inhabited-subtype-equivalence-class Φ-equivalence - ( map-inv-equiv-is-kripke-model-filtration is-filt class)) - ( relation-Prop-kripke-model i M* class class) - ( λ (x , in-class) → - ( tr - ( λ y → relation-kripke-model i M* y y) - ( class-x-eq-x* is-filt x class in-class) - ( filtration-relation-lower-bound-is-kripke-model-filtration - ( is-filt) - ( x) - ( x) - ( is-refl x)))))) + ( λ (x , in-class) → + ( tr + ( λ y → relation-kripke-model i M* y y) + ( class-x-eq-x* is-filt x class in-class) + ( filtration-relation-lower-bound-is-kripke-model-filtration + ( is-filt) + ( x) + ( x) + ( is-refl x)))) is-inhabited-equivalence-classes : is-inhabited (equivalence-class Φ-equivalence) @@ -548,8 +529,8 @@ module _ minimal-transitive-kripke-model-filtration : kripke-model - ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) - ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) ( i) ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) pr1 (pr1 (pr1 minimal-transitive-kripke-model-filtration)) = @@ -601,8 +582,7 @@ module _ is-kripke-model-filtration-minimal-kripke-model-filtration : is-kripke-model-filtration minimal-kripke-model-filtration is-kripke-model-filtration-minimal-kripke-model-filtration = - intro-exists - ( id-equiv) + pair id-equiv ( triple ( λ n x → ( pair @@ -674,8 +654,7 @@ module _ is-kripke-model-filtration minimal-transitive-kripke-model-filtration is-kripke-model-filtration-minimal-transitive-kripke-model-filtration M-is-trans = - intro-exists - ( id-equiv) + pair id-equiv ( triple ( λ n x → ( pair @@ -695,7 +674,7 @@ module _ is-in-subtype (reflexive-kripke-class l1 l2 i l4) M → is-in-subtype ( reflexive-kripke-class - ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) ( i) ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) @@ -718,7 +697,7 @@ module _ is-in-subtype (symmetry-kripke-class l1 l2 i l4) M → is-in-subtype ( symmetry-kripke-class - ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) ( i) ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) @@ -730,8 +709,8 @@ module _ is-in-subtype (reflexive-kripke-class l1 l2 i l4) M → is-in-subtype ( reflexive-kripke-class - ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) - ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) ( i) ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) ( minimal-transitive-kripke-model-filtration) @@ -744,8 +723,8 @@ module _ is-in-subtype (symmetry-kripke-class l1 l2 i l4) M → is-in-subtype ( symmetry-kripke-class - ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) - ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) ( i) ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) ( minimal-transitive-kripke-model-filtration) @@ -789,6 +768,7 @@ module _ (l4 : Level) (l5 l6 l7 l8 : Level) where + filtration-models : model-class l1 l2 i l4 (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4 ⊔ lsuc l5 ⊔ lsuc l6 ⊔ lsuc l7 ⊔ lsuc l8) diff --git a/src/modal-logic/modal-logic-decision.lagda.md b/src/modal-logic/modal-logic-decision.lagda.md index 99234f755b..94661f19f5 100644 --- a/src/modal-logic/modal-logic-decision.lagda.md +++ b/src/modal-logic/modal-logic-decision.lagda.md @@ -438,8 +438,6 @@ module _ module _ {l1 l2 l3 l4 : Level} (i : Set l3) - -- (theory : modal-theory l5 i) - -- (is-fin : is-finite (type-subtype theory)) (l5 l6 l7 l8 : Level) (lem : LEM (l2 ⊔ lsuc l3 ⊔ l4 ⊔ lsuc l5 ⊔ lsuc l6 ⊔ lsuc l7 ⊔ lsuc l8)) where @@ -449,94 +447,24 @@ module _ filtration-models-subset-finite-models M* = map-universal-property-trunc-Prop ( finite-models l1 l2 i l4 M*) - ( λ ((theory , M) , is-fin , t-is-filt) → - ( apply-universal-property-trunc-Prop - ( t-is-filt) - ( finite-models l1 l2 i l4 M*) - ( λ is-filt → - ( triple - ( is-finite-equiv - ( equiv-is-kripke-model-filtration i theory M M* is-filt) - ( is-finite-equivalence-classes-filtration i M - ( lower-LEM (l2 ⊔ l4) lem) - ( theory) - ( is-fin))) - ( λ x y → - ( lower-LEM - ( lsuc l3 ⊔ l4 ⊔ lsuc l5 ⊔ lsuc l6 ⊔ lsuc l7 ⊔ lsuc l8) - ( lem) - ( relation-Prop-kripke-model i M* x y))) - ( λ x n → - ( lower-LEM - ( l2 ⊔ lsuc l3 ⊔ lsuc l5 ⊔ lsuc l6 ⊔ lsuc l7 ⊔ lsuc l8) - ( lem) - ( valuate-kripke-model i M* n x))))))) - - -- module _ - -- (l6 l7 l8 : Level) - -- where - - -- private - -- w' = type-is-small (is-small-equivalence-classes-filtration l6) - - -- extend : - -- (type-subtype (λ n → theory (var n)) → w' → Prop l8) → - -- type-Set i → w' → Prop l8 - -- extend f n x - -- with - -- lower-LEM - -- ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) - -- ( lem) - -- ( theory (var n)) - -- ... | inl in-theory = f (n , in-theory) x - -- ... | inr _ = raise-empty-Prop l8 - - -- is-bounded-valuate-extend : - -- (f : type-subtype (λ n → theory (var n)) → w' → Prop l8) → - -- is-bounded-valuate i theory M (extend f) - -- is-bounded-valuate-extend f n not-in-theory x val - -- with - -- lower-LEM - -- ( lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) - -- ( lem) - -- ( theory (var n)) - -- ... | inl in-theory = not-in-theory in-theory - -- ... | inr not-in-theory' = map-inv-raise val - - -- last : - -- (w' → w' → Prop l7) - -- × (type-subtype (λ n → theory (var n)) → w' → Prop l8) → - -- type-subtype (is-alpha-Prop i theory M {l6} {l7} {l8}) - -- last (r , v) = - -- pair - -- ( pair - -- ( pair - -- ( pair - -- ( w') - -- ( map-is-inhabited - -- ( λ x → - -- ( map-equiv-is-small - -- ( is-small-equivalence-classes-filtration l6) - -- ( class (Φ-equivalence i theory M) x))) - -- ( is-inhabited-type-kripke-model i M))) - -- ( r)) - -- ( extend v)) - -- ( unit-trunc-Prop - -- ( pair - -- ( equiv-is-small (is-small-equivalence-classes-filtration l6)) - -- ( is-bounded-valuate-extend v))) - - -- is-surjective-last : is-surjective last - -- is-surjective-last (r , v) = - -- unit-trunc-Prop - -- (pair - -- ( pair - -- ( λ x y → - -- ( relation-Prop-kripke-model i r - -- ( {! !}) - -- ( {! !}))) - -- ( {! !})) - -- ( {! !})) + ( λ ((theory , M) , is-fin , is-filt) → + ( triple + ( is-finite-equiv + ( equiv-is-kripke-model-filtration i theory M M* is-filt) + ( is-finite-equivalence-classes-filtration i M + ( lower-LEM (l2 ⊔ l4) lem) + ( theory) + ( is-fin))) + ( λ x y → + ( lower-LEM + ( lsuc l3 ⊔ l4 ⊔ lsuc l5 ⊔ lsuc l6 ⊔ lsuc l7 ⊔ lsuc l8) + ( lem) + ( relation-Prop-kripke-model i M* x y))) + ( λ x n → + ( lower-LEM + ( l2 ⊔ lsuc l3 ⊔ lsuc l5 ⊔ lsuc l6 ⊔ lsuc l7 ⊔ lsuc l8) + ( lem) + ( valuate-kripke-model i M* n x))))) module _ {l1 l2 l3 l4 l5 l6 l7 l8 : Level} (i : Set l3) @@ -567,39 +495,31 @@ module _ filtrate-completeness logic complete a in-logic = complete a ( λ M M-in-class x → - apply-universal-property-trunc-Prop - ( filtration-is-filtration - ( M , M-in-class) + ( backward-implication + ( kripke-models-filtrations-theorem' i ( subformulas i a) - ( is-modal-theory-closed-under-subformulas-subformulas i a)) - ( (M , x) ⊨ a) - ( λ is-filt → - ( backward-implication - ( kripke-models-filtrations-theorem' i + ( is-modal-theory-closed-under-subformulas-subformulas i a) + ( M) + ( filtration (subformulas i a) M) + ( filtration-is-filtration + ( M , M-in-class) + ( subformulas i a) + ( is-modal-theory-closed-under-subformulas-subformulas i a)) + ( a) + ( head-in-list-subtype) + ( x)) + ( in-logic + ( filtration (subformulas i a) M) + ( intro-exists (a , M , M-in-class) refl) + ( map-equiv-is-kripke-model-filtration i + ( subformulas i a) + ( M) + ( filtration (subformulas i a) M) + ( filtration-is-filtration + ( M , M-in-class) ( subformulas i a) - ( is-modal-theory-closed-under-subformulas-subformulas - ( i) - ( a)) - ( M) - ( filtration (subformulas i a) M) - ( is-filt) - ( a) - ( unit-trunc-Prop (is-head a _)) - ( x)) - ( in-logic - ( filtration (list-subtype (cons a _)) M) - ( intro-exists (a , (M , M-in-class)) refl) - ( map-equiv-is-kripke-model-filtration - ( i) - ( subformulas i a) - ( M) - ( filtration (list-subtype (cons a _)) M) - ( is-filt) - ( class - ( Φ-equivalence i - ( subformulas i a) - ( M)) - ( x))))))) + ( is-modal-theory-closed-under-subformulas-subformulas i a)) + ( class (Φ-equivalence i (subformulas i a) M) x))))) filtrate-soundness : {l9 l10 : Level} (logic : modal-theory l9 i) → From 873449a385f24f59e316d536b2a1cd7586c626ca Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Sun, 5 May 2024 17:13:59 +0300 Subject: [PATCH 49/63] Refactor subtypes --- .../intersections-subtypes.lagda.md | 38 +++++++++---------- src/modal-logic/kripke-semantics.lagda.md | 11 ------ src/modal-logic/modal-logic-s5.lagda.md | 4 +- src/modal-logic/soundness.lagda.md | 32 ++++++++-------- 4 files changed, 35 insertions(+), 50 deletions(-) diff --git a/src/foundation/intersections-subtypes.lagda.md b/src/foundation/intersections-subtypes.lagda.md index 31fdec6e97..2838e427d0 100644 --- a/src/foundation/intersections-subtypes.lagda.md +++ b/src/foundation/intersections-subtypes.lagda.md @@ -62,7 +62,7 @@ module _ pr2 (pr2 (is-greatest-binary-lower-bound-intersection-subtype R) p) x r = pr2 (p x r) - infixl 15 _∩_ + infixr 15 _∩_ _∩_ = intersection-subtype ``` @@ -100,47 +100,47 @@ module _ -- where -- is-commutative-subtype-intersection : --- intersection-subtype P Q ⊆ intersection-subtype Q P +-- P ∩ Q ⊆ Q ∩ P -- is-commutative-subtype-intersection x (in-P , in-Q) = in-Q , in-P module _ {l1 l2 l3 : Level} {X : UU l1} (P : subtype l2 X) (Q : subtype l3 X) where - subtype-intersection-left : intersection-subtype P Q ⊆ P + subtype-intersection-left : P ∩ Q ⊆ P subtype-intersection-left _ = pr1 - subtype-intersection-right : intersection-subtype P Q ⊆ Q + subtype-intersection-right : P ∩ Q ⊆ Q subtype-intersection-right _ = pr2 subtype-both-intersection : {l4 : Level} (S : subtype l4 X) → - S ⊆ P → S ⊆ Q → S ⊆ intersection-subtype P Q + S ⊆ P → S ⊆ Q → S ⊆ P ∩ Q pr1 (subtype-both-intersection S S-sub-P S-sub-Q x in-S) = S-sub-P x in-S pr2 (subtype-both-intersection S S-sub-P S-sub-Q x in-S) = S-sub-Q x in-S - intersection-subtype-left-subtype : P ⊆ Q → P ⊆ intersection-subtype P Q + intersection-subtype-left-subtype : P ⊆ Q → P ⊆ P ∩ Q intersection-subtype-left-subtype P-sub-Q = subtype-both-intersection P (refl-leq-subtype P) P-sub-Q - intersection-subtype-right-subtype : Q ⊆ P → Q ⊆ intersection-subtype P Q + intersection-subtype-right-subtype : Q ⊆ P → Q ⊆ P ∩ Q intersection-subtype-right-subtype Q-sub-P = subtype-both-intersection Q Q-sub-P (refl-leq-subtype Q) equiv-intersection-subtype-left : - P ⊆ Q → equiv-subtypes (intersection-subtype P Q) P + P ⊆ Q → equiv-subtypes (P ∩ Q) P equiv-intersection-subtype-left P-sub-Q = equiv-antisymmetric-leq-subtype - ( intersection-subtype P Q) + ( P ∩ Q) ( P) ( subtype-intersection-left) ( intersection-subtype-left-subtype P-sub-Q) equiv-intersection-subtype-right : - Q ⊆ P → equiv-subtypes (intersection-subtype P Q) Q + Q ⊆ P → equiv-subtypes (P ∩ Q) Q equiv-intersection-subtype-right Q-sub-P = equiv-antisymmetric-leq-subtype - ( intersection-subtype P Q) + ( P ∩ Q) ( Q) ( subtype-intersection-right) ( intersection-subtype-right-subtype Q-sub-P) @@ -150,7 +150,7 @@ module _ where is-reflexivity-intersection : - {l2 : Level} (P : subtype l2 X) → intersection-subtype P P = P + {l2 : Level} (P : subtype l2 X) → P ∩ P = P is-reflexivity-intersection P = antisymmetric-leq-subtype _ _ ( subtype-intersection-left P P) @@ -163,16 +163,16 @@ module _ is-commutative-subtype-intersection : {l2 l3 : Level} (P : subtype l2 X) (Q : subtype l3 X) → - intersection-subtype P Q ⊆ intersection-subtype Q P + P ∩ Q ⊆ Q ∩ P is-commutative-subtype-intersection P Q = subtype-both-intersection Q P - ( intersection-subtype P Q) + ( P ∩ Q) ( subtype-intersection-right P Q) ( subtype-intersection-left P Q) is-commutative-intersection : {l2 l3 : Level} (P : subtype l2 X) (Q : subtype l3 X) → - intersection-subtype P Q = intersection-subtype Q P + P ∩ Q = Q ∩ P is-commutative-intersection P Q = antisymmetric-leq-subtype _ _ ( is-commutative-subtype-intersection P Q) @@ -180,7 +180,7 @@ module _ intersection-subtype-left-sublevels : {l2 : Level} (l3 : Level) (P : subtype (l2 ⊔ l3) X) (Q : subtype l2 X) → - P ⊆ Q → intersection-subtype P Q = P + P ⊆ Q → P ∩ Q = P intersection-subtype-left-sublevels _ P Q P-sub-Q = antisymmetric-leq-subtype _ _ ( subtype-intersection-left P Q) @@ -188,7 +188,7 @@ module _ intersection-subtype-right-sublevels : {l2 : Level} (l3 : Level) (P : subtype l2 X) (Q : subtype (l2 ⊔ l3) X) → - Q ⊆ P → intersection-subtype P Q = Q + Q ⊆ P → P ∩ Q = Q intersection-subtype-right-sublevels l3 P Q Q-sub-P = tr ( _= Q) @@ -197,11 +197,11 @@ module _ intersection-subtype-left : {l2 : Level} (P Q : subtype l2 X) → - P ⊆ Q → intersection-subtype P Q = P + P ⊆ Q → P ∩ Q = P intersection-subtype-left = intersection-subtype-left-sublevels lzero intersection-subtype-right : {l2 : Level} (P Q : subtype l2 X) → - Q ⊆ P → intersection-subtype P Q = Q + Q ⊆ P → P ∩ Q = Q intersection-subtype-right = intersection-subtype-right-sublevels lzero ``` diff --git a/src/modal-logic/kripke-semantics.lagda.md b/src/modal-logic/kripke-semantics.lagda.md index 34724ec443..9f0240a0e8 100644 --- a/src/modal-logic/kripke-semantics.lagda.md +++ b/src/modal-logic/kripke-semantics.lagda.md @@ -184,17 +184,6 @@ module _ relation-property-class property M = property (relation-Prop-kripke-model i M) - -- intersection-relation-property-class : - -- {l5 l6 : Level} → - -- (R₁ : {A : UU l1} → subtype l5 (Relation-Prop l2 A)) → - -- (R₂ : {A : UU l1} → subtype l5 (Relation-Prop l2 A)) → - -- Id - -- ( intersection-subtype - -- ( relation-property-class R₁) - -- ( relation-property-class R₂)) - -- ( relation-property-class (intersection-subtype R₁ R₂)) - -- intersection-relation-property-class R₁ R₂ = refl - reflexive-kripke-class : model-class l1 l2 i l4 (l1 ⊔ l2) reflexive-kripke-class = relation-property-class diff --git a/src/modal-logic/modal-logic-s5.lagda.md b/src/modal-logic/modal-logic-s5.lagda.md index c1b2afced3..5e3305d1d8 100644 --- a/src/modal-logic/modal-logic-s5.lagda.md +++ b/src/modal-logic/modal-logic-s5.lagda.md @@ -143,9 +143,7 @@ module _ ( ax-m i) ( ax-b i ∪ ax-4 i) ( reflexive-kripke-class l2 l3 i l4) - ( intersection-subtype - ( symmetry-kripke-class l2 l3 i l4) - ( transitivity-kripke-class l2 l3 i l4)) + ( symmetry-kripke-class l2 l3 i l4 ∩ transitivity-kripke-class l2 l3 i l4) ( ax-m-soundness i l3 l4) ( soundness-union ( ax-b i) diff --git a/src/modal-logic/soundness.lagda.md b/src/modal-logic/soundness.lagda.md index 17e92d76de..8296c2f350 100644 --- a/src/modal-logic/soundness.lagda.md +++ b/src/modal-logic/soundness.lagda.md @@ -104,11 +104,9 @@ module _ ( forces-in-intersection M in-class a) soundness-modal-logic-union : - soundness - (modal-logic-closure (union-subtype theory₁ theory₂)) - (intersection-subtype C₁ C₂) + soundness (modal-logic-closure (theory₁ ∪ theory₂)) (C₁ ∩ C₂) soundness-modal-logic-union = - soundness-modal-logic (intersection-subtype C₁ C₂) soundness-union + soundness-modal-logic (C₁ ∩ C₂) soundness-union module _ {l1 l2 l3 l4 l5 l6 l7 : Level} @@ -120,11 +118,11 @@ module _ (C₁ : model-class l4 l5 i l6 (l7 ⊔ l8)) (C₂ : model-class l4 l5 i l6 l7) (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) → C₁ ⊆ C₂ → - soundness (union-subtype ax₁ ax₂) C₁ + soundness (ax₁ ∪ ax₂) C₁ soundness-union-subclass-left-sublevels l8 C₁ C₂ sound₁ sound₂ C₁-sub-C₂ = tr - ( soundness (union-subtype ax₁ ax₂)) + ( soundness (ax₁ ∪ ax₂)) ( intersection-subtype-left-sublevels l8 C₁ C₂ C₁-sub-C₂) ( soundness-union ax₁ ax₂ C₁ C₂ sound₁ sound₂) @@ -133,11 +131,11 @@ module _ (C₁ : model-class l4 l5 i l6 l7) (C₂ : model-class l4 l5 i l6 (l7 ⊔ l8)) (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) → C₂ ⊆ C₁ → - soundness (union-subtype ax₁ ax₂) C₂ + soundness (ax₁ ∪ ax₂) C₂ soundness-union-subclass-right-sublevels l8 C₁ C₂ sound₁ sound₂ C₂-sub-C₁ = tr - ( soundness (union-subtype ax₁ ax₂)) + ( soundness (ax₁ ∪ ax₂)) ( intersection-subtype-right-sublevels l8 C₁ C₂ C₂-sub-C₁) ( soundness-union ax₁ ax₂ C₁ C₂ sound₁ sound₂) @@ -146,11 +144,11 @@ module _ (C₁ : model-class l4 l5 i l6 (l7 ⊔ l8)) (C₂ : model-class l4 l5 i l6 l7) (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) → C₁ ⊆ C₂ → - soundness (modal-logic-closure (union-subtype ax₁ ax₂)) C₁ + soundness (modal-logic-closure (ax₁ ∪ ax₂)) C₁ soundness-modal-logic-union-subclass-left-sublevels l8 C₁ C₂ sound₁ sound₂ C₁-sub-C₂ = tr - ( soundness (modal-logic-closure (union-subtype ax₁ ax₂))) + ( soundness (modal-logic-closure (ax₁ ∪ ax₂))) ( intersection-subtype-left-sublevels l8 C₁ C₂ C₁-sub-C₂) ( soundness-modal-logic-union ax₁ ax₂ C₁ C₂ sound₁ sound₂) @@ -159,11 +157,11 @@ module _ (C₁ : model-class l4 l5 i l6 l7) (C₂ : model-class l4 l5 i l6 (l7 ⊔ l8)) (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) → C₂ ⊆ C₁ → - soundness (modal-logic-closure (union-subtype ax₁ ax₂)) C₂ + soundness (modal-logic-closure (ax₁ ∪ ax₂)) C₂ soundness-modal-logic-union-subclass-right-sublevels l8 C₁ C₂ sound₁ sound₂ C₂-sub-C₁ = tr - ( soundness (modal-logic-closure (union-subtype ax₁ ax₂))) + ( soundness (modal-logic-closure (ax₁ ∪ ax₂))) ( intersection-subtype-right-sublevels l8 C₁ C₂ C₂-sub-C₁) ( soundness-modal-logic-union ax₁ ax₂ C₁ C₂ sound₁ sound₂) @@ -171,7 +169,7 @@ module _ (C₁ : model-class l4 l5 i l6 l7) (C₂ : model-class l4 l5 i l6 l7) (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) → C₁ ⊆ C₂ → - soundness (modal-logic-closure (union-subtype ax₁ ax₂)) C₁ + soundness (modal-logic-closure (ax₁ ∪ ax₂)) C₁ soundness-modal-logic-union-subclass-left = soundness-modal-logic-union-subclass-left-sublevels lzero @@ -179,7 +177,7 @@ module _ (C₁ : model-class l4 l5 i l6 l7) (C₂ : model-class l4 l5 i l6 l7) (sound₁ : soundness ax₁ C₁) (sound₂ : soundness ax₂ C₂) → C₂ ⊆ C₁ → - soundness (modal-logic-closure (union-subtype ax₁ ax₂)) C₂ + soundness (modal-logic-closure (ax₁ ∪ ax₂)) C₂ soundness-modal-logic-union-subclass-right = soundness-modal-logic-union-subclass-right-sublevels lzero @@ -190,15 +188,15 @@ module _ (sound₁ : soundness ax₁ C) (sound₂ : soundness ax₂ C) where - soundness-union-same-class : soundness (union-subtype ax₁ ax₂) C + soundness-union-same-class : soundness (ax₁ ∪ ax₂) C soundness-union-same-class = tr - ( soundness (union-subtype ax₁ ax₂)) + ( soundness (ax₁ ∪ ax₂)) ( is-reflexivity-intersection C) ( soundness-union ax₁ ax₂ C C sound₁ sound₂) soundness-modal-logic-union-same-class : - soundness (modal-logic-closure (union-subtype ax₁ ax₂)) C + soundness (modal-logic-closure (ax₁ ∪ ax₂)) C soundness-modal-logic-union-same-class = soundness-modal-logic C soundness-union-same-class ``` From 4111374d5a144ebcca9d0b6b04850515d46fd80f Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Sun, 5 May 2024 17:31:47 +0300 Subject: [PATCH 50/63] Refactor formulas and module names --- src/modal-logic.lagda.md | 4 +- src/modal-logic/axioms.lagda.md | 57 +++--- .../canonical-model-theorem.lagda.md | 169 +++++++++--------- src/modal-logic/canonical-theories.lagda.md | 18 +- src/modal-logic/completeness-k.lagda.md | 2 +- src/modal-logic/completeness-s5.lagda.md | 2 +- src/modal-logic/completeness.lagda.md | 2 +- ...n.lagda.md => decision-procedure.lagda.md} | 59 +++--- ...gic-syntax.lagda.md => deduction.lagda.md} | 62 +++---- .../finite-approximability.lagda.md | 4 +- src/modal-logic/formulas-deduction.lagda.md | 105 +++++------ src/modal-logic/formulas.lagda.md | 143 +++++++-------- ...kripke-models-filtrations-theorem.lagda.md | 12 +- .../kripke-models-filtrations.lagda.md | 50 +++--- src/modal-logic/kripke-semantics.lagda.md | 22 +-- src/modal-logic/l-complete-theories.lagda.md | 60 +++---- .../l-consistent-theories.lagda.md | 4 +- src/modal-logic/lindenbaum.lagda.md | 2 +- src/modal-logic/modal-logic-k.lagda.md | 4 +- src/modal-logic/modal-logic-s5.lagda.md | 4 +- src/modal-logic/soundness.lagda.md | 6 +- src/modal-logic/weak-deduction.lagda.md | 134 +++++++------- 22 files changed, 472 insertions(+), 453 deletions(-) rename src/modal-logic/{modal-logic-decision.lagda.md => decision-procedure.lagda.md} (92%) rename src/modal-logic/{logic-syntax.lagda.md => deduction.lagda.md} (83%) diff --git a/src/modal-logic.lagda.md b/src/modal-logic.lagda.md index 83447a76c4..2b0ad79bbc 100644 --- a/src/modal-logic.lagda.md +++ b/src/modal-logic.lagda.md @@ -9,6 +9,8 @@ open import modal-logic.canonical-theories public open import modal-logic.completeness public open import modal-logic.completeness-k public open import modal-logic.completeness-s5 public +open import modal-logic.decision-procedure public +open import modal-logic.deduction public open import modal-logic.finite-approximability public open import modal-logic.formulas public open import modal-logic.formulas-deduction public @@ -18,8 +20,6 @@ open import modal-logic.kripke-semantics public open import modal-logic.l-complete-theories public open import modal-logic.l-consistent-theories public open import modal-logic.lindenbaum public -open import modal-logic.logic-syntax public -open import modal-logic.modal-logic-decision public open import modal-logic.modal-logic-k public open import modal-logic.modal-logic-s5 public open import modal-logic.soundness public diff --git a/src/modal-logic/axioms.lagda.md b/src/modal-logic/axioms.lagda.md index 58cac7d685..149ad41c3b 100644 --- a/src/modal-logic/axioms.lagda.md +++ b/src/modal-logic/axioms.lagda.md @@ -32,9 +32,9 @@ open import foundation-core.sets open import foundation-core.subtypes open import foundation-core.transport-along-identifications +open import modal-logic.deduction open import modal-logic.formulas open import modal-logic.kripke-semantics -open import modal-logic.logic-syntax open import modal-logic.soundness ``` @@ -52,8 +52,8 @@ module _ where ax-1-parameter : - (h : formula i → formula i) → is-injective h → modal-theory l i - pr1 (ax-1-parameter h inj f) = Σ (formula i) (λ a → f = h a) + (h : modal-formula i → modal-formula i) → is-injective h → modal-theory l i + pr1 (ax-1-parameter h inj f) = Σ (modal-formula i) (λ a → f = h a) pr2 (ax-1-parameter h inj f) (a , refl) = is-prop-is-contr ( is-contr-Σ-is-prop a refl @@ -62,12 +62,12 @@ module _ ( a , refl) ax-2-parameters : - (h : formula i → formula i → formula i) → - ({x x' y y' : formula i} → h x y = h x' y' → x = x') → - ({x x' y y' : formula i} → h x y = h x' y' → y = y') → + (h : modal-formula i → modal-formula i → modal-formula i) → + ({x x' y y' : modal-formula i} → h x y = h x' y' → x = x') → + ({x x' y y' : modal-formula i} → h x y = h x' y' → y = y') → modal-theory l i pr1 (ax-2-parameters h inj-1 inj-2 f) = - Σ (formula i) (λ a → Σ (formula i) (λ b → f = h a b)) + Σ (modal-formula i) (λ a → Σ (modal-formula i) (λ b → f = h a b)) pr2 (ax-2-parameters h inj-1 inj-2 f) (a , b , refl) = is-prop-is-contr ( is-contr-Σ-is-prop a (b , refl) @@ -76,14 +76,17 @@ module _ ( a , b , refl) ax-3-parameters : - (h : formula i → formula i → formula i → formula i) → - ({x x' y y' z z' : formula i} → h x y z = h x' y' z' → x = x') → - ({x x' y y' z z' : formula i} → h x y z = h x' y' z' → y = y') → - ({x x' y y' z z' : formula i} → h x y z = h x' y' z' → z = z') → + (h : + modal-formula i → modal-formula i → modal-formula i → modal-formula i) → + ({x x' y y' z z' : modal-formula i} → h x y z = h x' y' z' → x = x') → + ({x x' y y' z z' : modal-formula i} → h x y z = h x' y' z' → y = y') → + ({x x' y y' z z' : modal-formula i} → h x y z = h x' y' z' → z = z') → modal-theory l i pr1 (ax-3-parameters h inj-1 inj-2 inj-3 f) = - Σ ( formula i) - ( λ a → Σ (formula i) (λ b → Σ (formula i) ( λ c → f = h a b c))) + Σ ( modal-formula i) + ( λ a → + ( Σ (modal-formula i) + ( λ b → Σ (modal-formula i) ( λ c → f = h a b c)))) pr2 (ax-3-parameters h inj-1 inj-2 inj-3 f) (a , b , c , refl) = is-prop-is-contr ( is-contr-Σ-is-prop a (b , c , refl) @@ -109,28 +112,28 @@ module _ ax-n : modal-theory l i ax-n = ax-2-parameters - ( λ a b → □ (a →ₘ b) →ₘ □ a →ₘ □ b) + ( λ a b → □ₘ (a →ₘ b) →ₘ □ₘ a →ₘ □ₘ b) ( eq-implication-left ∘ eq-box ∘ eq-implication-left) ( eq-implication-right ∘ eq-box ∘ eq-implication-left) ax-dn : modal-theory l i - ax-dn = ax-1-parameter (λ a → ~~ a →ₘ a) eq-implication-right + ax-dn = ax-1-parameter (λ a → ¬¬ₘ a →ₘ a) eq-implication-right ax-m : modal-theory l i - ax-m = ax-1-parameter (λ a → □ a →ₘ a) eq-implication-right + ax-m = ax-1-parameter (λ a → □ₘ a →ₘ a) eq-implication-right ax-b : modal-theory l i - ax-b = ax-1-parameter (λ a → a →ₘ □ ◇ a) eq-implication-left + ax-b = ax-1-parameter (λ a → a →ₘ □ₘ ◇ₘ a) eq-implication-left ax-d : modal-theory l i - ax-d = ax-1-parameter (λ a → □ a →ₘ ◇ a) (eq-box ∘ eq-implication-left) + ax-d = ax-1-parameter (λ a → □ₘ a →ₘ ◇ₘ a) (eq-box ∘ eq-implication-left) ax-4 : modal-theory l i - ax-4 = ax-1-parameter (λ a → □ a →ₘ □ □ a) (eq-box ∘ eq-implication-left) + ax-4 = ax-1-parameter (λ a → □ₘ a →ₘ □ₘ □ₘ a) (eq-box ∘ eq-implication-left) ax-5 : modal-theory l i ax-5 = - ax-1-parameter ( λ a → ◇ a →ₘ □ ◇ a) ( eq-diamond ∘ eq-implication-left) + ax-1-parameter ( λ a → ◇ₘ a →ₘ □ₘ ◇ₘ a) ( eq-diamond ∘ eq-implication-left) module _ {l1 l2 : Level} @@ -150,26 +153,26 @@ module _ ax-n-soundness : soundness (ax-n i) (all-models l2 l3 i l4) ax-n-soundness - .(□ (a →ₘ b) →ₘ □ a →ₘ □ b) + .(□ₘ (a →ₘ b) →ₘ □ₘ a →ₘ □ₘ b) (a , b , refl) M in-class x fab fa y r = fab y r (fa y r) ax-dn-soundness : soundness (ax-dn i) (decidable-models l2 l3 i l4) - ax-dn-soundness .(~~ a →ₘ a) (a , refl) M is-dec x f + ax-dn-soundness .(¬¬ₘ a →ₘ a) (a , refl) M is-dec x f with (is-dec a x) ... | inl fa = fa ... | inr fna = raise-ex-falso _ (f (λ fa -> map-raise (fna fa))) ax-m-soundness : soundness (ax-m i) (reflexive-kripke-class l2 l3 i l4) - ax-m-soundness .(□ a →ₘ a) (a , refl) M is-refl x fa = fa x (is-refl x) + ax-m-soundness .(□ₘ a →ₘ a) (a , refl) M is-refl x fa = fa x (is-refl x) ax-b-soundness : soundness (ax-b i) (symmetry-kripke-class l2 l3 i l4) - ax-b-soundness .(a →ₘ □ ◇ a) (a , refl) M is-sym x fa y r contra = + ax-b-soundness .(a →ₘ □ₘ ◇ₘ a) (a , refl) M is-sym x fa y r contra = contra x (is-sym x y r) fa ax-d-soundness : soundness (ax-d i) (serial-kripke-class l2 l3 i l4) - ax-d-soundness .(□ a →ₘ ◇ a) (a , refl) M is-serial x fa contra = + ax-d-soundness .(□ₘ a →ₘ ◇ₘ a) (a , refl) M is-serial x fa contra = map-raise ( apply-universal-property-trunc-Prop ( is-serial x) @@ -177,10 +180,10 @@ module _ ( λ (y , r) → map-inv-raise (contra y r (fa y r)))) ax-4-soundness : soundness (ax-4 i) (transitivity-kripke-class l2 l3 i l4) - ax-4-soundness .(□ a →ₘ □ □ a) (a , refl) M is-trans x fa y r-xy z r-yz = + ax-4-soundness .(□ₘ a →ₘ □ₘ □ₘ a) (a , refl) M is-trans x fa y r-xy z r-yz = fa z (is-trans x y z r-yz r-xy) ax-5-sooundness : soundness (ax-5 i) (euclidean-kripke-class l2 l3 i l4) - ax-5-sooundness .(◇ a →ₘ □ ◇ a) (a , refl) M is-eucl x fa y r-xy contra = + ax-5-sooundness .(◇ₘ a →ₘ □ₘ ◇ₘ a) (a , refl) M is-eucl x fa y r-xy contra = fa (λ z r-xz → contra z (is-eucl x y z r-xy r-xz)) ``` diff --git a/src/modal-logic/canonical-model-theorem.lagda.md b/src/modal-logic/canonical-model-theorem.lagda.md index 4ee33fe8f2..46f7c70de5 100644 --- a/src/modal-logic/canonical-model-theorem.lagda.md +++ b/src/modal-logic/canonical-model-theorem.lagda.md @@ -33,12 +33,12 @@ open import lists.reversing-lists open import modal-logic.axioms open import modal-logic.completeness +open import modal-logic.deduction open import modal-logic.formulas open import modal-logic.kripke-semantics open import modal-logic.l-complete-theories open import modal-logic.l-consistent-theories open import modal-logic.lindenbaum -open import modal-logic.logic-syntax open import modal-logic.modal-logic-k open import modal-logic.weak-deduction @@ -108,10 +108,10 @@ module _ ( contains-ax-s) pr2 (pr1 canonical-kripke-model) x y = Π-Prop - ( formula i) + ( modal-formula i) ( λ a → ( hom-Prop - ( modal-theory-L-complete-theory logic x (□ a)) + ( modal-theory-L-complete-theory logic x (□ₘ a)) ( modal-theory-L-complete-theory logic y a))) pr2 canonical-kripke-model n x = modal-theory-L-complete-theory logic x (var n) @@ -168,7 +168,7 @@ module _ ( lem) L-complete-theory-implication : - {a b : formula i} → + {a b : modal-formula i} → (is-in-subtype theory a → is-in-subtype theory b) → is-in-subtype theory (a →ₘ b) L-complete-theory-implication {a} {b} f with is-disjuctive-theory a @@ -198,16 +198,16 @@ module _ { ax₁ = theory} { ax₂ = theory-add-formula a theory} ( subset-add-formula a theory) - ( ~ a) + ( ¬ₘ a) ( weak-modal-logic-closure-ax not-a-in-logic)))) L-complete-theory-box : - {a : formula i} → + {a : modal-formula i} → ( (y : L-complete-theory logic (l1 ⊔ l2)) → relation-kripke-model i canonical-kripke-model x y → is-in-subtype (modal-theory-L-complete-theory logic y) a) → - is-in-subtype theory (□ a) - L-complete-theory-box {a} f with is-disjuctive-theory (□ a) + is-in-subtype theory (□ₘ a) + L-complete-theory-box {a} f with is-disjuctive-theory (□ₘ a) ... | inl box-a-in-logic = box-a-in-logic ... | inr not-box-a-in-logic = ex-falso @@ -219,37 +219,39 @@ module _ ( is-consistent-modal-theory-L-complete-theory logic w ( weak-modal-logic-mp ( is-weak-modal-logic-L-complete-theory logic lzero w) - ( y-leq-w (~ a) - ( formula-in-add-formula (~ a) (unbox-modal-theory theory))) + ( y-leq-w (¬ₘ a) + ( formula-in-add-formula (¬ₘ a) (unbox-modal-theory theory))) ( f w (λ b → y-leq-w b ∘ y-contains-unbox)))))) where y : modal-theory (l1 ⊔ l2) i - y = theory-add-formula (~ a) (unbox-modal-theory theory) + y = theory-add-formula (¬ₘ a) (unbox-modal-theory theory) y-contains-unbox : - {b : formula i} → - is-in-subtype theory (□ b) → + {b : modal-formula i} → + is-in-subtype theory (□ₘ b) → is-in-subtype y b y-contains-unbox {b} = - subset-add-formula (~ a) (unbox-modal-theory theory) b + subset-add-formula (¬ₘ a) (unbox-modal-theory theory) b - list-to-implications : formula i → (l : list (formula i)) → formula i + list-to-implications : + modal-formula i → (l : list (modal-formula i)) → modal-formula i list-to-implications f nil = f list-to-implications f (cons g l) = list-to-implications (g →ₘ f) l - list-to-implications-rev : formula i → (l : list (formula i)) → formula i + list-to-implications-rev : + modal-formula i → (l : list (modal-formula i)) → modal-formula i list-to-implications-rev f nil = f list-to-implications-rev f (cons g l) = g →ₘ list-to-implications-rev f l list-to-implication-rev-snoc : - (f g : formula i) (l : list (formula i)) → + (f g : modal-formula i) (l : list (modal-formula i)) → list-to-implications f (snoc l g) = g →ₘ list-to-implications f l list-to-implication-rev-snoc f g nil = refl list-to-implication-rev-snoc f g (cons h l) = list-to-implication-rev-snoc (h →ₘ f) g l eq-reverse-list-to-implications : - (f : formula i) (l : list (formula i)) → + (f : modal-formula i) (l : list (modal-formula i)) → list-to-implications f (reverse-list l) = list-to-implications-rev f l eq-reverse-list-to-implications f nil = refl eq-reverse-list-to-implications f (cons g l) = @@ -257,7 +259,7 @@ module _ ( ap (λ x → g →ₘ x) (eq-reverse-list-to-implications f l)) move-assumptions-right : - (f : formula i) (l : list (formula i)) → + (f : modal-formula i) (l : list (modal-formula i)) → is-in-subtype (weak-modal-logic-closure (logic ∪ list-subtype l)) f → is-in-subtype ( weak-modal-logic-closure logic) @@ -297,10 +299,10 @@ module _ ( in-logic))) α : - (l : list (formula i)) → + (l : list (modal-formula i)) → list-subtype l ⊆ unbox-modal-theory theory → - is-in-subtype theory (□ (list-to-implications-rev a l)) → - is-in-subtype theory (□ a) + is-in-subtype theory (□ₘ (list-to-implications-rev a l)) → + is-in-subtype theory (□ₘ a) α nil sub in-logic = in-logic α (cons c l) sub in-logic = α l @@ -312,25 +314,25 @@ module _ ( subset-tail-list-subtype)) ( weak-modal-logic-mp ( is-weak-modal-logic-L-complete-theory logic lzero x) - { a = □ c} - { b = □ list-to-implications-rev a l} + { a = □ₘ c} + { b = □ₘ list-to-implications-rev a l} ( weak-modal-logic-mp ( is-weak-modal-logic-L-complete-theory logic lzero x) - { a = □ (c →ₘ list-to-implications-rev a l)} - { b = □ c →ₘ □ (list-to-implications-rev a l)} + { a = □ₘ (c →ₘ list-to-implications-rev a l)} + { b = □ₘ c →ₘ □ₘ (list-to-implications-rev a l)} ( contains-ax-n' _ (c , list-to-implications-rev a l , refl)) ( in-logic)) ( sub c head-in-list-subtype)) β : - (l : list (formula i)) → + (l : list (modal-formula i)) → list-subtype l ⊆ unbox-modal-theory theory → is-in-subtype (weak-modal-logic-closure (logic ∪ list-subtype l)) a → - is-in-subtype theory (□ a) + is-in-subtype theory (□ₘ a) β l sub in-logic = α l sub ( subset-logic-L-complete-theory logic lzero x - ( □ list-to-implications-rev a l) + ( □ₘ list-to-implications-rev a l) ( modal-logic-nec is-logic ( tr ( is-in-subtype logic) @@ -347,62 +349,62 @@ module _ ( in-logic))))))) γ : - (l : list (formula i)) → + (l : list (modal-formula i)) → list-subtype l ⊆ unbox-modal-theory theory → is-contradictory-modal-logic ( weak-modal-logic-closure - ( theory-add-formula (~ a) (logic ∪ list-subtype l))) → - is-in-subtype theory (□ a) + ( theory-add-formula (¬ₘ a) (logic ∪ list-subtype l))) → + is-in-subtype theory (□ₘ a) γ l sub is-cont = β l sub - ( weak-modal-logic-closure-mp {a = ~~ a} {b = a} + ( weak-modal-logic-closure-mp {a = ¬¬ₘ a} {b = a} ( weak-modal-logic-closure-ax - ( subtype-union-left logic (list-subtype l) (~~ a →ₘ a) - ( contains-ax-dn (~~ a →ₘ a) (a , refl)))) + ( subtype-union-left logic (list-subtype l) (¬¬ₘ a →ₘ a) + ( contains-ax-dn (¬¬ₘ a →ₘ a) (a , refl)))) ( forward-implication ( deduction-lemma ( logic ∪ list-subtype l) ( contains-ax-k-union (list-subtype l)) ( contains-ax-s-union (list-subtype l)) - ( ~ a) - ( ⊥)) + ( ¬ₘ a) + ( ⊥ₘ)) ( is-cont))) δ : - (l : list (formula i)) → + (l : list (modal-formula i)) → list-subtype l ⊆ unbox-modal-theory theory → is-contradictory-modal-logic ( weak-modal-logic-closure - ( logic ∪ (theory-add-formula (~ a) (list-subtype l)))) → - is-in-subtype theory (□ a) + ( logic ∪ (theory-add-formula (¬ₘ a) (list-subtype l)))) → + is-in-subtype theory (□ₘ a) δ l sub is-cont = γ l sub ( is-contradictory-modal-logic-monotic ( weak-modal-logic-closure - ( logic ∪ theory-add-formula (~ a) (list-subtype l))) + ( logic ∪ theory-add-formula (¬ₘ a) (list-subtype l))) ( weak-modal-logic-closure - ( theory-add-formula (~ a) (logic ∪ list-subtype l))) + ( theory-add-formula (¬ₘ a) (logic ∪ list-subtype l))) ( weak-modal-logic-closure-monotic - ( theory-add-formula-union-right (~ a) logic (list-subtype l))) + ( theory-add-formula-union-right (¬ₘ a) logic (list-subtype l))) ( is-cont)) ε : - (l : list (formula i)) → + (l : list (modal-formula i)) → list-subtype l ⊆ logic ∪ y → is-contradictory-modal-logic ( weak-modal-logic-closure (list-subtype l)) → - is-in-subtype theory (□ a) + is-in-subtype theory (□ₘ a) ε l sub is-cont = apply-universal-property-trunc-Prop ( lists-in-union-lists l logic y sub) - ( theory (□ a)) + ( theory (□ₘ a)) ( λ ((l-ax , l-y) , l-sub-union , l-ax-sub-logic , l-y-sub-y) → ( apply-universal-property-trunc-Prop ( lists-in-union-lists l-y - ( Id-formula-Prop (~ a)) + ( Id-formula-Prop (¬ₘ a)) ( unbox-modal-theory theory) ( l-y-sub-y)) - ( theory (□ a)) + ( theory (□ₘ a)) ( λ ((l-not-a , l-box) , l-sub-union' , l-not-a-sub , l-box-sub) → ( δ ( l-box) @@ -410,35 +412,37 @@ module _ ( is-contradictory-modal-logic-monotic ( weak-modal-logic-closure (list-subtype l)) ( weak-modal-logic-closure - ( logic ∪ theory-add-formula (~ a) (list-subtype l-box))) + ( logic ∪ theory-add-formula (¬ₘ a) (list-subtype l-box))) ( weak-modal-logic-closure-monotic ( transitive-leq-subtype ( list-subtype l) ( list-subtype l-ax ∪ list-subtype l-y) - ( logic ∪ theory-add-formula (~ a) (list-subtype l-box)) + ( logic ∪ + theory-add-formula (¬ₘ a) (list-subtype l-box)) ( subset-union-subsets ( list-subtype l-ax) ( list-subtype l-y) ( logic) - ( theory-add-formula (~ a) (list-subtype l-box)) + ( theory-add-formula (¬ₘ a) (list-subtype l-box)) ( l-ax-sub-logic) ( transitive-leq-subtype ( list-subtype l-y) ( list-subtype l-not-a ∪ list-subtype l-box) - ( theory-add-formula (~ a) (list-subtype l-box)) + ( theory-add-formula (¬ₘ a) (list-subtype l-box)) ( subtype-union-both ( list-subtype l-not-a) ( list-subtype l-box) - ( theory-add-formula (~ a) (list-subtype l-box)) + ( theory-add-formula (¬ₘ a) (list-subtype l-box)) ( transitive-leq-subtype ( list-subtype l-not-a) - ( Id-formula-Prop (~ a)) - ( theory-add-formula (~ a) (list-subtype l-box)) + ( Id-formula-Prop (¬ₘ a)) + ( theory-add-formula (¬ₘ a) + ( list-subtype l-box)) ( subtype-union-left - ( Id-formula-Prop (~ a)) + ( Id-formula-Prop (¬ₘ a)) ( list-subtype l-box)) ( l-not-a-sub)) - ( subset-add-formula (~ a) (list-subtype l-box))) + ( subset-add-formula (¬ₘ a) (list-subtype l-box))) ( l-sub-union'))) ( l-sub-union))) ( is-cont)))))) @@ -459,13 +463,13 @@ module _ ( is-assumptions-list-assumptions-weak-deduction d-bot)))))) canonical-model-theorem-pointwise : - (a : formula i) + (a : modal-formula i) (x : L-complete-theory logic (l1 ⊔ l2)) → type-iff-Prop ( modal-theory-L-complete-theory logic x a) ( (canonical-kripke-model , x) ⊨ a) pr1 (canonical-model-theorem-pointwise (var n) x) = map-raise - pr1 (canonical-model-theorem-pointwise ⊥ x) = + pr1 (canonical-model-theorem-pointwise ⊥ₘ x) = map-raise ∘ is-consistent-modal-theory-L-complete-theory logic x pr1 (canonical-model-theorem-pointwise (a →ₘ b) x) in-logic f = forward-implication @@ -474,12 +478,12 @@ module _ ( is-weak-modal-logic-L-complete-theory logic lzero x) ( in-logic) ( backward-implication (canonical-model-theorem-pointwise a x) f)) - pr1 (canonical-model-theorem-pointwise (□ a) x) in-logic y xRy = + pr1 (canonical-model-theorem-pointwise (□ₘ a) x) in-logic y xRy = forward-implication ( canonical-model-theorem-pointwise a y) ( xRy a in-logic) pr2 (canonical-model-theorem-pointwise (var n) x) = map-inv-raise - pr2 (canonical-model-theorem-pointwise ⊥ x) (map-raise ()) + pr2 (canonical-model-theorem-pointwise ⊥ₘ x) (map-raise ()) pr2 (canonical-model-theorem-pointwise (a →ₘ b) x) f = L-complete-theory-implication x ( λ in-x → @@ -489,7 +493,7 @@ module _ ( forward-implication ( canonical-model-theorem-pointwise a x) ( in-x))))) - pr2 (canonical-model-theorem-pointwise (□ a) x) f = + pr2 (canonical-model-theorem-pointwise (□ₘ a) x) f = L-complete-theory-box x ( λ y xRy → ( backward-implication @@ -497,7 +501,8 @@ module _ ( f y xRy))) canonical-model-theorem : - (a : formula i) → type-iff-Prop (logic a) (canonical-kripke-model ⊨M a) + (a : modal-formula i) → + type-iff-Prop (logic a) (canonical-kripke-model ⊨M a) pr1 (canonical-model-theorem a) in-logic x = forward-implication ( canonical-model-theorem-pointwise a x) @@ -513,19 +518,19 @@ module _ ( is-consistent-modal-theory-L-complete-theory logic w ( weak-modal-logic-mp ( is-weak-modal-logic-L-complete-theory logic lzero w) - ( leq (~ a) not-a-in-x) + ( leq (¬ₘ a) not-a-in-x) ( backward-implication ( canonical-model-theorem-pointwise a w) ( f w))))))) where x : modal-theory (l1 ⊔ l2) i - x = raise-subtype l2 (Id-formula-Prop (~ a)) + x = raise-subtype l2 (Id-formula-Prop (¬ₘ a)) - not-a-in-x : is-in-subtype x ( ~ a) + not-a-in-x : is-in-subtype x ( ¬ₘ a) not-a-in-x = - subset-equiv-subtypes (Id-formula-Prop (~ a)) x - ( compute-raise-subtype l2 (Id-formula-Prop (~ a))) - ( ~ a) + subset-equiv-subtypes (Id-formula-Prop (¬ₘ a)) x + ( compute-raise-subtype l2 (Id-formula-Prop (¬ₘ a))) + ( ¬ₘ a) ( refl) is-L-consistent-x : @@ -533,22 +538,22 @@ module _ is-L-consistent-x a-not-in-logic bot-in-logic = a-not-in-logic ( modal-logic-mp is-logic - {a = ~~ a} + {a = ¬¬ₘ a} {b = a} - ( contains-ax-dn (~~ a →ₘ a) (a , refl)) - ( is-logic (~~ a) - ( subset-weak-modal-logic-closure-modal-logic-closure (~~ a) + ( contains-ax-dn (¬¬ₘ a →ₘ a) (a , refl)) + ( is-logic (¬¬ₘ a) + ( subset-weak-modal-logic-closure-modal-logic-closure (¬¬ₘ a) ( forward-implication - ( deduction-lemma logic contains-ax-k contains-ax-s ( ~ a) ⊥) + ( deduction-lemma logic contains-ax-k contains-ax-s ( ¬ₘ a) ⊥ₘ) ( weak-modal-logic-closure-monotic - ( subtype-union-both logic x (theory-add-formula (~ a) logic) - ( subtype-union-right (Id-formula-Prop (~ a)) logic) - ( transitive-leq-subtype x (Id-formula-Prop (~ a)) - ( theory-add-formula (~ a) logic) - ( subtype-union-left (Id-formula-Prop (~ a)) logic) - ( inv-subset-equiv-subtypes (Id-formula-Prop (~ a)) x - ( compute-raise-subtype l2 (Id-formula-Prop (~ a)))))) - ( ⊥) + ( subtype-union-both logic x (theory-add-formula (¬ₘ a) logic) + ( subtype-union-right (Id-formula-Prop (¬ₘ a)) logic) + ( transitive-leq-subtype x (Id-formula-Prop (¬ₘ a)) + ( theory-add-formula (¬ₘ a) logic) + ( subtype-union-left (Id-formula-Prop (¬ₘ a)) logic) + ( inv-subset-equiv-subtypes (Id-formula-Prop (¬ₘ a)) x + ( compute-raise-subtype l2 (Id-formula-Prop (¬ₘ a)))))) + ( ⊥ₘ) ( bot-in-logic)))))) canonical-model-completness : diff --git a/src/modal-logic/canonical-theories.lagda.md b/src/modal-logic/canonical-theories.lagda.md index faa842a6e9..23b428b971 100644 --- a/src/modal-logic/canonical-theories.lagda.md +++ b/src/modal-logic/canonical-theories.lagda.md @@ -24,10 +24,10 @@ open import foundation-core.coproduct-types open import modal-logic.axioms open import modal-logic.canonical-model-theorem open import modal-logic.completeness +open import modal-logic.deduction open import modal-logic.formulas open import modal-logic.kripke-semantics open import modal-logic.l-complete-theories -open import modal-logic.logic-syntax open import modal-logic.modal-logic-k open import modal-logic.modal-logic-s5 open import modal-logic.soundness @@ -95,8 +95,8 @@ module _ weak-modal-logic-mp ( is-weak-modal-logic-L-complete-theory logic lzero x) ( subset-logic-L-complete-theory logic lzero x - ( □ a →ₘ a) - ( ax-m-subset (□ a →ₘ a) (a , refl))) + ( □ₘ a →ₘ a) + ( ax-m-subset (□ₘ a →ₘ a) (a , refl))) ( box-a-in-x) is-canonical-ax-b : @@ -118,11 +118,11 @@ module _ ( map-universal-property-trunc-Prop ( modal-theory-L-complete-theory logic y b) ( λ { (c , c-in-x , refl) → - ( xRy (◇ c) + ( xRy (◇ₘ c) ( weak-modal-logic-mp ( is-weak-modal-logic-L-complete-theory logic lzero x) - ( subset-logic-L-complete-theory logic lzero x (c →ₘ □ ◇ c) - ( ax-b-subset (c →ₘ □ ◇ c) (c , refl))) + ( subset-logic-L-complete-theory logic lzero x (c →ₘ □ₘ ◇ₘ c) + ( ax-b-subset (c →ₘ □ₘ ◇ₘ c) (c , refl))) ( c-in-x)))}))) ( a) ( box-a-in-y) @@ -135,11 +135,11 @@ module _ prop-resize) is-canonical-ax-4 ax-4-subset x y z yRz xRy a box-a-in-x = yRz a - ( xRy (□ a) + ( xRy (□ₘ a) ( weak-modal-logic-mp ( is-weak-modal-logic-L-complete-theory logic lzero x) ( subset-logic-L-complete-theory logic lzero x - ( □ a →ₘ □ □ a) - ( ax-4-subset (□ a →ₘ □ □ a) (a , refl))) + ( □ₘ a →ₘ □ₘ □ₘ a) + ( ax-4-subset (□ₘ a →ₘ □ₘ □ₘ a) (a , refl))) ( box-a-in-x))) ``` diff --git a/src/modal-logic/completeness-k.lagda.md b/src/modal-logic/completeness-k.lagda.md index a622cc4e7b..2d2401f843 100644 --- a/src/modal-logic/completeness-k.lagda.md +++ b/src/modal-logic/completeness-k.lagda.md @@ -35,9 +35,9 @@ open import foundation.universe-levels open import modal-logic.axioms open import modal-logic.canonical-model-theorem open import modal-logic.completeness +open import modal-logic.deduction open import modal-logic.formulas open import modal-logic.kripke-semantics -open import modal-logic.logic-syntax open import modal-logic.modal-logic-k open import modal-logic.soundness open import modal-logic.weak-deduction diff --git a/src/modal-logic/completeness-s5.lagda.md b/src/modal-logic/completeness-s5.lagda.md index 9572c5fbd3..f3a88e16d1 100644 --- a/src/modal-logic/completeness-s5.lagda.md +++ b/src/modal-logic/completeness-s5.lagda.md @@ -36,9 +36,9 @@ open import modal-logic.axioms open import modal-logic.canonical-model-theorem open import modal-logic.canonical-theories open import modal-logic.completeness +open import modal-logic.deduction open import modal-logic.formulas open import modal-logic.kripke-semantics -open import modal-logic.logic-syntax open import modal-logic.modal-logic-s5 open import modal-logic.soundness open import modal-logic.weak-deduction diff --git a/src/modal-logic/completeness.lagda.md b/src/modal-logic/completeness.lagda.md index a288243eb7..4ef7f28e94 100644 --- a/src/modal-logic/completeness.lagda.md +++ b/src/modal-logic/completeness.lagda.md @@ -20,9 +20,9 @@ open import foundation.universe-levels open import foundation-core.coproduct-types +open import modal-logic.deduction open import modal-logic.formulas open import modal-logic.kripke-semantics -open import modal-logic.logic-syntax ```
diff --git a/src/modal-logic/modal-logic-decision.lagda.md b/src/modal-logic/decision-procedure.lagda.md similarity index 92% rename from src/modal-logic/modal-logic-decision.lagda.md rename to src/modal-logic/decision-procedure.lagda.md index 94661f19f5..1578b6ef72 100644 --- a/src/modal-logic/modal-logic-decision.lagda.md +++ b/src/modal-logic/decision-procedure.lagda.md @@ -1,7 +1,7 @@ # Modal logic decision ```agda -module modal-logic.modal-logic-decision where +module modal-logic.decision-procedure where ```
Imports @@ -53,11 +53,11 @@ open import lists.lists-subtypes open import lists.reversing-lists open import modal-logic.completeness +open import modal-logic.deduction open import modal-logic.formulas open import modal-logic.kripke-models-filtrations open import modal-logic.kripke-models-filtrations-theorem open import modal-logic.kripke-semantics -open import modal-logic.logic-syntax open import modal-logic.soundness open import modal-logic.weak-deduction @@ -89,7 +89,7 @@ module _ where decision-procedure' : - (a : formula i) → + (a : modal-formula i) → is-decidable ( (M : type-subtype C) → type-Prop (inclusion-subtype C M ⊨M a)) decision-procedure' a = @@ -98,7 +98,7 @@ module _ ( λ (M , M-in-C) → ( is-finite-model-valuate-decidable-models i M (C-sub-fin M M-in-C) a)) - decision-procedure : (a : formula i) → bool + decision-procedure : (a : modal-formula i) → bool decision-procedure a with decision-procedure' a ... | inl _ = true ... | inr _ = false @@ -106,7 +106,7 @@ module _ decision-procedure-correctness : soundness theory C → completeness theory C → - (a : formula i) → + (a : modal-formula i) → is-in-subtype theory a ↔ type-prop-bool (decision-procedure a) pr1 (decision-procedure-correctness sound complete a) in-theory with decision-procedure' a @@ -173,20 +173,20 @@ module _ {l : Level} (i : Set l) where - subformulas-list : formula i → list (formula i) + subformulas-list : modal-formula i → list (modal-formula i) subformulas-list a = cons a (rest a) where - rest : formula i → list (formula i) + rest : modal-formula i → list (modal-formula i) rest (var x) = nil - rest ⊥ = nil + rest ⊥ₘ = nil rest (a →ₘ b) = concat-list (subformulas-list a) (subformulas-list b) - rest (□ a) = subformulas-list a + rest (□ₘ a) = subformulas-list a - subformulas : formula i → modal-theory l i + subformulas : modal-formula i → modal-theory l i subformulas a = list-subtype (subformulas-list a) subformulas-list-has-subimpl : - (a : formula i) {x y : formula i} → + (a : modal-formula i) {x y : modal-formula i} → (x →ₘ y) ∈-list subformulas-list a → (x ∈-list subformulas-list a) × (y ∈-list subformulas-list a) subformulas-list-has-subimpl .(x →ₘ y) {x} {y} (is-head .(x →ₘ y) _) = @@ -223,19 +223,19 @@ module _ ( is-in-tail y (a →ₘ b) _ ( in-concat-right (subformulas-list a) (subformulas-list b) y-in-tail)) subformulas-list-has-subimpl - (□ a) {x} {y} (is-in-tail .(x →ₘ y) .(□ a) _ xy-list-subtype) = + (□ₘ a) {x} {y} (is-in-tail .(x →ₘ y) .(□ₘ a) _ xy-list-subtype) = let (x-in-tail , y-in-tail) = subformulas-list-has-subimpl a xy-list-subtype - in (is-in-tail x (□ a) _ x-in-tail) , (is-in-tail y (□ a) _ y-in-tail) + in (is-in-tail x (□ₘ a) _ x-in-tail) , (is-in-tail y (□ₘ a) _ y-in-tail) subformulas-list-has-subbox : - (a : formula i) {x : formula i} → - □ x ∈-list subformulas-list a → + (a : modal-formula i) {x : modal-formula i} → + □ₘ x ∈-list subformulas-list a → x ∈-list subformulas-list a - subformulas-list-has-subbox .(□ x) {x} (is-head .(□ x) _) = - is-in-tail x (□ x) _ (is-head x _) + subformulas-list-has-subbox .(□ₘ x) {x} (is-head .(□ₘ x) _) = + is-in-tail x (□ₘ x) _ (is-head x _) subformulas-list-has-subbox - (a →ₘ b) {x} (is-in-tail .(□ x) .(a →ₘ b) _ x-list-subtype) + (a →ₘ b) {x} (is-in-tail .(□ₘ x) .(a →ₘ b) _ x-list-subtype) with in-concat-list (subformulas-list a) (subformulas-list b) x-list-subtype ... | inl x-in-left = @@ -247,11 +247,11 @@ module _ ( in-concat-right (subformulas-list a) (subformulas-list b) ( subformulas-list-has-subbox b x-in-right)) subformulas-list-has-subbox - (□ a) {x} (is-in-tail .(□ x) .(□ a) _ x-list-subtype) = - is-in-tail x (□ a) _ (subformulas-list-has-subbox a x-list-subtype) + (□ₘ a) {x} (is-in-tail .(□ₘ x) .(□ₘ a) _ x-list-subtype) = + is-in-tail x (□ₘ a) _ (subformulas-list-has-subbox a x-list-subtype) is-modal-theory-closed-under-subformulas-subformulas : - (a : formula i) → + (a : modal-formula i) → is-modal-theory-closed-under-subformulas i (subformulas a) is-modal-theory-closed-under-subformulas-subformulas a = is-modal-theory-closed-under-subformulas-condition @@ -271,18 +271,19 @@ module _ ( list-subtype (subformulas-list a) _) ( unit-trunc-Prop ∘ subformulas-list-has-subbox a)) - subformulas-Set : formula i → Set l + subformulas-Set : modal-formula i → Set l subformulas-Set a = set-subset (formula-Set i) (list-subtype (subformulas-list a)) - subformulas-Set-list : (a : formula i) → list (type-Set (subformulas-Set a)) + subformulas-Set-list : + (a : modal-formula i) → list (type-Set (subformulas-Set a)) subformulas-Set-list a = dependent-map-list ( subformulas-list a) ( λ (x , list-subtype) → x , unit-trunc-Prop list-subtype) is-kuratowsky-finite'-subformulas-list : - (a : formula i) → is-kuratowsky-finite-set' (subformulas-Set a) + (a : modal-formula i) → is-kuratowsky-finite-set' (subformulas-Set a) is-kuratowsky-finite'-subformulas-list a = intro-exists ( subformulas-Set-list a) @@ -300,7 +301,7 @@ module _ ( b-list-subtype))))))) is-kuratowsky-finite-subformulas-list : - (a : formula i) → is-kuratowsky-finite-set (subformulas-Set a) + (a : modal-formula i) → is-kuratowsky-finite-set (subformulas-Set a) is-kuratowsky-finite-subformulas-list a = is-kuratowsky-finite-set-is-kuratowsky-finite-set' ( subformulas-Set a) @@ -308,7 +309,7 @@ module _ is-finite-subformulas-list : LEM l → - (a : formula i) → + (a : modal-formula i) → is-finite (type-subtype (list-subtype (subformulas-list a))) is-finite-subformulas-list lem a = is-finite-is-kuratowsky-finite-set @@ -318,7 +319,7 @@ module _ {l2 : Level} → LEM l → LEM l2 → - (a : formula i) → + (a : modal-formula i) → is-finite (type-subtype (list-subtype (subformulas-list a)) → Prop l2) is-finite-subtypes-subformulas-list lem lem2 a = is-finite-function-type @@ -477,7 +478,7 @@ module _ filtrate-class : model-class l6 l7 i l8 ( l3 ⊔ l5 ⊔ lsuc (l1 ⊔ l2 ⊔ l4 ⊔ l6 ⊔ l7 ⊔ l8)) filtrate-class M* = - exists-structure-Prop (formula i × type-subtype C) + exists-structure-Prop (modal-formula i × type-subtype C) ( λ (a , (M , _)) → M* = filtration (subformulas i a) M) module _ @@ -525,7 +526,7 @@ module _ {l9 l10 : Level} (logic : modal-theory l9 i) → (C₂ : model-class l6 l7 i l8 l10) → ( ((M , _) : type-subtype C) → - (a : formula i) → + (a : modal-formula i) → is-in-subtype C₂ (filtration (list-subtype (subformulas-list i a)) M)) → soundness logic C₂ → soundness logic filtrate-class diff --git a/src/modal-logic/logic-syntax.lagda.md b/src/modal-logic/deduction.lagda.md similarity index 83% rename from src/modal-logic/logic-syntax.lagda.md rename to src/modal-logic/deduction.lagda.md index 35ca1441b8..4ad1c4f583 100644 --- a/src/modal-logic/logic-syntax.lagda.md +++ b/src/modal-logic/deduction.lagda.md @@ -1,7 +1,7 @@ # Modal logic syntax ```agda -module modal-logic.logic-syntax where +module modal-logic.deduction where ```
Imports @@ -42,7 +42,7 @@ module _ where modal-theory : UU (l1 ⊔ lsuc l2) - modal-theory = subtype l2 (formula i) + modal-theory = subtype l2 (modal-formula i) module _ {l1 l2 : Level} {i : Set l1} @@ -50,10 +50,10 @@ module _ infix 5 _⊢_ - data _⊢_ (axioms : modal-theory l2 i) : formula i → UU (l1 ⊔ l2) where - ax : {a : formula i} → is-in-subtype axioms a → axioms ⊢ a - mp : {a b : formula i} → axioms ⊢ a →ₘ b → axioms ⊢ a → axioms ⊢ b - nec : {a : formula i} → axioms ⊢ a → axioms ⊢ □ a + data _⊢_ (axioms : modal-theory l2 i) : modal-formula i → UU (l1 ⊔ l2) where + ax : {a : modal-formula i} → is-in-subtype axioms a → axioms ⊢ a + mp : {a b : modal-formula i} → axioms ⊢ a →ₘ b → axioms ⊢ a → axioms ⊢ b + nec : {a : modal-formula i} → axioms ⊢ a → axioms ⊢ □ₘ a -- TODO: rename to modal-logic-closure modal-logic-closure : modal-theory l2 i → modal-theory (l1 ⊔ l2) i @@ -67,12 +67,12 @@ module _ is-modal-logic = type-Prop ∘ is-modal-logic-Prop is-in-modal-logic-closure-deduction : - {axioms : modal-theory l2 i} {a : formula i} → + {axioms : modal-theory l2 i} {a : modal-formula i} → axioms ⊢ a → is-in-subtype (modal-logic-closure axioms) a is-in-modal-logic-closure-deduction = unit-trunc-Prop is-contradictory-modal-logic-Prop : modal-theory l2 i → Prop l2 - is-contradictory-modal-logic-Prop logic = logic ⊥ + is-contradictory-modal-logic-Prop logic = logic ⊥ₘ is-contradictory-modal-logic : modal-theory l2 i → UU l2 is-contradictory-modal-logic = type-Prop ∘ is-contradictory-modal-logic-Prop @@ -92,7 +92,7 @@ module _ ax₁ ⊆ ax₂ → is-contradictory-modal-logic ax₁ → is-contradictory-modal-logic ax₂ - is-contradictory-modal-logic-monotic ax₁ ax₂ leq = leq ⊥ + is-contradictory-modal-logic-monotic ax₁ ax₂ leq = leq ⊥ₘ is-consistent-modal-logic-antimonotic : {l2 l3 : Level} (ax₁ : modal-theory l2 i) (ax₂ : modal-theory l3 i) → @@ -107,13 +107,13 @@ module _ where modal-logic-closure-ax : - {a : formula i} → + {a : modal-formula i} → is-in-subtype axioms a → is-in-subtype (modal-logic-closure axioms) a modal-logic-closure-ax = unit-trunc-Prop ∘ ax modal-logic-closure-mp : - {a b : formula i} → + {a b : modal-formula i} → is-in-subtype (modal-logic-closure axioms) (a →ₘ b) → is-in-subtype (modal-logic-closure axioms) a → is-in-subtype (modal-logic-closure axioms) b @@ -123,12 +123,12 @@ module _ ( λ dab da → unit-trunc-Prop (mp dab da)) modal-logic-closure-nec : - {a : formula i} → + {a : modal-formula i} → is-in-subtype (modal-logic-closure axioms) a → - is-in-subtype (modal-logic-closure axioms) (□ a) + is-in-subtype (modal-logic-closure axioms) (□ₘ a) modal-logic-closure-nec {a} = map-universal-property-trunc-Prop - ( modal-logic-closure axioms (□ a)) + ( modal-logic-closure axioms (□ₘ a)) ( λ da → unit-trunc-Prop (nec da)) module _ @@ -137,7 +137,7 @@ module _ where modal-logic-mp : - {a b : formula i} → + {a b : modal-formula i} → is-in-subtype logic (a →ₘ b) → is-in-subtype logic a → is-in-subtype logic b @@ -148,11 +148,11 @@ module _ ( modal-logic-closure-ax da)) modal-logic-nec : - {a : formula i} → + {a : modal-formula i} → is-in-subtype logic a → - is-in-subtype logic (□ a) + is-in-subtype logic (□ₘ a) modal-logic-nec {a} d = - is-logic (□ a) (modal-logic-closure-nec (modal-logic-closure-ax d)) + is-logic (□ₘ a) (modal-logic-closure-nec (modal-logic-closure-ax d)) module _ {l1 : Level} {i : Set l1} @@ -164,7 +164,7 @@ module _ axioms-subset-modal-logic _ a H = unit-trunc-Prop (ax H) modal-logic-closed : - {l2 : Level} {axioms : modal-theory l2 i} {a : formula i} → + {l2 : Level} {axioms : modal-theory l2 i} {a : modal-formula i} → modal-logic-closure axioms ⊢ a → is-in-subtype (modal-logic-closure axioms) a modal-logic-closed (ax x) = x @@ -189,7 +189,7 @@ module _ (leq : ax₁ ⊆ ax₂) where - deduction-monotic : {a : formula i} → ax₁ ⊢ a → ax₂ ⊢ a + deduction-monotic : {a : modal-formula i} → ax₁ ⊢ a → ax₂ ⊢ a deduction-monotic (ax x) = ax (leq _ x) deduction-monotic (mp dab da) = mp (deduction-monotic dab) (deduction-monotic da) @@ -218,7 +218,9 @@ module _ ( modal-logic-monotic leq) module _ - {l1 l2 : Level} {i : Set l1} (a : formula i) (theory : modal-theory l2 i) + {l1 l2 : Level} {i : Set l1} + (a : modal-formula i) + (theory : modal-theory l2 i) where -- TODO: make Id-formula to be a function for 1 element modal-theory @@ -241,10 +243,10 @@ module _ ( leq) elim-theory-add-formula : - {l3 : Level} (P : formula i → Prop l3) → + {l3 : Level} (P : modal-formula i → Prop l3) → type-Prop (P a) → - ((x : formula i) → is-in-subtype theory x → type-Prop (P x)) → - (x : formula i) → is-in-subtype theory-add-formula x → type-Prop (P x) + ((x : modal-formula i) → is-in-subtype theory x → type-Prop (P x)) → + (x : modal-formula i) → is-in-subtype theory-add-formula x → type-Prop (P x) elim-theory-add-formula P H-a H-rest = elim-union-subtype (Id-formula-Prop a) theory P ( λ { .a refl → H-a}) @@ -267,13 +269,13 @@ module _ where unbox-modal-theory : modal-theory l2 i → modal-theory l2 i - unbox-modal-theory theory a = theory (□ a) + unbox-modal-theory theory a = theory (□ₘ a) diamond-modal-theory : modal-theory l2 i → modal-theory (l1 ⊔ l2) i diamond-modal-theory theory a = exists-structure-Prop - ( formula i) - ( λ b → is-in-subtype theory b × (a = ◇ b)) + ( modal-formula i) + ( λ b → is-in-subtype theory b × (a = ◇ₘ b)) module _ {l1 : Level} {i : Set l1} @@ -282,10 +284,10 @@ module _ is-disjuctive-modal-theory : {l2 : Level} → modal-theory l2 i → UU (l1 ⊔ l2) is-disjuctive-modal-theory theory = - (a : formula i) → is-in-subtype theory a + is-in-subtype theory (~ a) + (a : modal-formula i) → is-in-subtype theory a + is-in-subtype theory (¬ₘ a) theory-add-formula-union-right : - (a : formula i) + (a : modal-formula i) {l2 l3 : Level} (theory : modal-theory l2 i) (theory' : modal-theory l3 i) → @@ -295,7 +297,7 @@ module _ union-swap-1-2 theory (Id-formula-Prop a) theory' inv-theory-add-formula-union-right : - (a : formula i) + (a : modal-formula i) {l2 l3 : Level} (theory : modal-theory l2 i) (theory' : modal-theory l3 i) → diff --git a/src/modal-logic/finite-approximability.lagda.md b/src/modal-logic/finite-approximability.lagda.md index dc8bbfbcca..a7cdfc4c50 100644 --- a/src/modal-logic/finite-approximability.lagda.md +++ b/src/modal-logic/finite-approximability.lagda.md @@ -39,11 +39,11 @@ open import foundation-core.invertible-maps open import modal-logic.completeness open import modal-logic.completeness-k +open import modal-logic.decision-procedure +open import modal-logic.deduction open import modal-logic.formulas open import modal-logic.kripke-models-filtrations open import modal-logic.kripke-semantics -open import modal-logic.logic-syntax -open import modal-logic.modal-logic-decision open import modal-logic.modal-logic-k open import modal-logic.soundness open import modal-logic.weak-deduction diff --git a/src/modal-logic/formulas-deduction.lagda.md b/src/modal-logic/formulas-deduction.lagda.md index 208232b757..3209aff63e 100644 --- a/src/modal-logic/formulas-deduction.lagda.md +++ b/src/modal-logic/formulas-deduction.lagda.md @@ -17,8 +17,8 @@ open import foundation.universe-levels open import foundation-core.identity-types open import modal-logic.axioms +open import modal-logic.deduction open import modal-logic.formulas -open import modal-logic.logic-syntax open import modal-logic.modal-logic-k open import modal-logic.weak-deduction ``` @@ -76,7 +76,7 @@ module _ ( K-contains-ax-dn i) weak-modal-logic-const : - (a : formula i) {b : formula i} → + (a : modal-formula i) {b : modal-formula i} → is-in-subtype (weak-modal-logic-closure axioms) b → is-in-subtype (weak-modal-logic-closure axioms) (a →ₘ b) weak-modal-logic-const a {b} b-in-logic = @@ -86,7 +86,7 @@ module _ ( b-in-logic) modal-logic-const : - (a : formula i) {b : formula i} → + (a : modal-formula i) {b : modal-formula i} → is-in-subtype (modal-logic-closure axioms) b → is-in-subtype (modal-logic-closure axioms) (a →ₘ b) modal-logic-const a {b} b-in-logic = @@ -96,7 +96,7 @@ module _ ( b-in-logic) weak-modal-logic-transitivity : - {a b c : formula i} → + {a b c : modal-formula i} → is-in-subtype (weak-modal-logic-closure axioms) (b →ₘ c) → is-in-subtype (weak-modal-logic-closure axioms) (a →ₘ b) → is-in-subtype (weak-modal-logic-closure axioms) (a →ₘ c) @@ -109,7 +109,7 @@ module _ ( ab) modal-logic-transitivity : - {a b c : formula i} → + {a b c : modal-formula i} → is-in-subtype (modal-logic-closure axioms) (b →ₘ c) → is-in-subtype (modal-logic-closure axioms) (a →ₘ b) → is-in-subtype (modal-logic-closure axioms) (a →ₘ c) @@ -122,127 +122,130 @@ module _ ( ab) modal-logic-implication-box' : - {a b : formula i} → + {a b : modal-formula i} → is-in-subtype (modal-logic-closure axioms) (a →ₘ b) → - is-in-subtype (modal-logic-closure axioms) (□ a →ₘ □ b) + is-in-subtype (modal-logic-closure axioms) (□ₘ a →ₘ □ₘ b) modal-logic-implication-box' {a} {b} ab = modal-logic-closure-mp ( modal-logic-closure-ax - ( contains-ax-n (□ (a →ₘ b) →ₘ □ a →ₘ □ b) (a , b , refl))) + ( contains-ax-n (□ₘ (a →ₘ b) →ₘ □ₘ a →ₘ □ₘ b) (a , b , refl))) ( modal-logic-closure-nec ab) modal-logic-implication-box : - {a b : formula i} → + {a b : modal-formula i} → is-in-subtype (modal-logic-closure axioms) (a →ₘ b) → - is-in-subtype (modal-logic-closure axioms) (□ a) → - is-in-subtype (modal-logic-closure axioms) (□ b) + is-in-subtype (modal-logic-closure axioms) (□ₘ a) → + is-in-subtype (modal-logic-closure axioms) (□ₘ b) modal-logic-implication-box {a} {b} ab box-a = modal-logic-closure-mp ( modal-logic-closure-mp ( modal-logic-closure-ax - ( contains-ax-n (□ (a →ₘ b) →ₘ □ a →ₘ □ b) (a , b , refl))) + ( contains-ax-n (□ₘ (a →ₘ b) →ₘ □ₘ a →ₘ □ₘ b) (a , b , refl))) ( modal-logic-closure-nec ab)) ( box-a) weak-modal-logic-implication-dn : - (a : formula i) → - is-in-subtype (weak-modal-logic-closure axioms) (a →ₘ ~~ a) + (a : modal-formula i) → + is-in-subtype (weak-modal-logic-closure axioms) (a →ₘ ¬¬ₘ a) weak-modal-logic-implication-dn a = - inv-deduction-ex-falso axioms contains-ax-k contains-ax-s contains-ax-dn a ⊥ + inv-deduction-ex-falso axioms contains-ax-k contains-ax-s contains-ax-dn + ( a) + ( ⊥ₘ) modal-logic-implication-dn : - (a : formula i) → is-in-subtype (modal-logic-closure axioms) (a →ₘ ~~ a) + (a : modal-formula i) → + is-in-subtype (modal-logic-closure axioms) (a →ₘ ¬¬ₘ a) modal-logic-implication-dn a = - subset-weak-modal-logic-closure-modal-logic-closure (a →ₘ ~~ a) + subset-weak-modal-logic-closure-modal-logic-closure (a →ₘ ¬¬ₘ a) ( weak-modal-logic-implication-dn a) weak-modal-logic-implication-negate : - {a b : formula i} → + {a b : modal-formula i} → is-in-subtype axioms (a →ₘ b) → - is-in-subtype (weak-modal-logic-closure axioms) (~ b →ₘ ~ a) + is-in-subtype (weak-modal-logic-closure axioms) (¬ₘ b →ₘ ¬ₘ a) weak-modal-logic-implication-negate {a} {b} ab = forward-implication - ( deduction-lemma axioms contains-ax-k contains-ax-s (~ b) (~ a)) + ( deduction-lemma axioms contains-ax-k contains-ax-s (¬ₘ b) (¬ₘ a)) ( forward-implication ( deduction-lemma - ( theory-add-formula (~ b) axioms) + ( theory-add-formula (¬ₘ b) axioms) ( contains-ax-k') ( contains-ax-s') ( a) - ( ⊥)) + ( ⊥ₘ)) ( logic-ex-falso - ( theory-add-formula a (theory-add-formula (~ b) axioms)) + ( theory-add-formula a (theory-add-formula (¬ₘ b) axioms)) ( contains-ax-k'') ( contains-ax-s'') ( contains-ax-dn'') ( b) - ( ⊥) + ( ⊥ₘ) ( weak-modal-logic-closure-mp ( weak-modal-logic-closure-ax - ( subset-add-formula a (theory-add-formula (~ b) axioms) + ( subset-add-formula a (theory-add-formula (¬ₘ b) axioms) ( a →ₘ b) - ( subset-add-formula (~ b) axioms (a →ₘ b) ab))) + ( subset-add-formula (¬ₘ b) axioms (a →ₘ b) ab))) ( weak-modal-logic-closure-ax - ( formula-in-add-formula a (theory-add-formula (~ b) axioms)))) + ( formula-in-add-formula a (theory-add-formula (¬ₘ b) axioms)))) ( weak-modal-logic-closure-ax - ( subset-add-formula a (theory-add-formula (~ b) axioms) - ( ~ b) - ( formula-in-add-formula (~ b) axioms))))) + ( subset-add-formula a (theory-add-formula (¬ₘ b) axioms) + ( ¬ₘ b) + ( formula-in-add-formula (¬ₘ b) axioms))))) where - contains-ax-k' : ax-k i ⊆ theory-add-formula (~ b) axioms + contains-ax-k' : ax-k i ⊆ theory-add-formula (¬ₘ b) axioms contains-ax-k' = - transitive-subset-add-formula (~ b) axioms (ax-k i) contains-ax-k + transitive-subset-add-formula (¬ₘ b) axioms (ax-k i) contains-ax-k - contains-ax-s' : ax-s i ⊆ theory-add-formula (~ b) axioms + contains-ax-s' : ax-s i ⊆ theory-add-formula (¬ₘ b) axioms contains-ax-s' = - transitive-subset-add-formula (~ b) axioms (ax-s i) contains-ax-s + transitive-subset-add-formula (¬ₘ b) axioms (ax-s i) contains-ax-s contains-ax-k'' : - ax-k i ⊆ theory-add-formula a (theory-add-formula (~ b) axioms) + ax-k i ⊆ theory-add-formula a (theory-add-formula (¬ₘ b) axioms) contains-ax-k'' = - transitive-subset-add-formula a (theory-add-formula (~ b) axioms) + transitive-subset-add-formula a (theory-add-formula (¬ₘ b) axioms) ( ax-k i) ( contains-ax-k') contains-ax-s'' : - ax-s i ⊆ theory-add-formula a (theory-add-formula (~ b) axioms) + ax-s i ⊆ theory-add-formula a (theory-add-formula (¬ₘ b) axioms) contains-ax-s'' = - transitive-subset-add-formula a (theory-add-formula (~ b) axioms) + transitive-subset-add-formula a (theory-add-formula (¬ₘ b) axioms) ( ax-s i) ( contains-ax-s') contains-ax-dn'' : - ax-dn i ⊆ theory-add-formula a (theory-add-formula (~ b) axioms) + ax-dn i ⊆ theory-add-formula a (theory-add-formula (¬ₘ b) axioms) contains-ax-dn'' = - transitive-subset-add-formula a (theory-add-formula (~ b) axioms) + transitive-subset-add-formula a (theory-add-formula (¬ₘ b) axioms) ( ax-dn i) - ( transitive-subset-add-formula (~ b) axioms (ax-dn i) contains-ax-dn) + ( transitive-subset-add-formula (¬ₘ b) axioms (ax-dn i) contains-ax-dn) modal-logic-implication-negate : - {a b : formula i} → + {a b : modal-formula i} → is-in-subtype axioms (a →ₘ b) → - is-in-subtype (modal-logic-closure axioms) (~ b →ₘ ~ a) + is-in-subtype (modal-logic-closure axioms) (¬ₘ b →ₘ ¬ₘ a) modal-logic-implication-negate {a} {b} ab = - subset-weak-modal-logic-closure-modal-logic-closure (~ b →ₘ ~ a) + subset-weak-modal-logic-closure-modal-logic-closure (¬ₘ b →ₘ ¬ₘ a) ( weak-modal-logic-implication-negate ab) modal-logic-diamond-negate : - {a : formula i} → - is-in-subtype (modal-logic-closure axioms) (◇ ~ a) → - is-in-subtype (modal-logic-closure axioms) (~ □ a) + {a : modal-formula i} → + is-in-subtype (modal-logic-closure axioms) (◇ₘ ¬ₘ a) → + is-in-subtype (modal-logic-closure axioms) (¬ₘ □ₘ a) modal-logic-diamond-negate {a} dia-a = modal-logic-transitivity ( dia-a) ( modal-logic-implication-box' (modal-logic-implication-dn a)) modal-logic-diamond-negate-implication : - {a : formula i} → + {a : modal-formula i} → is-modal-logic axioms → - is-in-subtype axioms (◇ ~ a →ₘ ~ □ a) + is-in-subtype axioms (◇ₘ ¬ₘ a →ₘ ¬ₘ □ₘ a) modal-logic-diamond-negate-implication {a} is-logic = - is-logic (◇ ~ a →ₘ ~ □ a) + is-logic (◇ₘ ¬ₘ a →ₘ ¬ₘ □ₘ a) ( modal-logic-implication-negate - ( is-logic (□ a →ₘ □ ~~ a) + ( is-logic (□ₘ a →ₘ □ₘ ¬¬ₘ a) ( modal-logic-implication-box' ( modal-logic-implication-dn a)))) ``` diff --git a/src/modal-logic/formulas.lagda.md b/src/modal-logic/formulas.lagda.md index 0c72d6bbea..925062c064 100644 --- a/src/modal-logic/formulas.lagda.md +++ b/src/modal-logic/formulas.lagda.md @@ -37,52 +37,52 @@ module _ where infixr 7 _→ₘ_ - infixr 25 □_ + infixr 25 □ₘ_ - data formula (i : Set l) : UU l where - var : type-Set i → formula i - ⊥ : formula i - _→ₘ_ : formula i → formula i → formula i - □_ : formula i → formula i + data modal-formula (i : Set l) : UU l where + var : type-Set i → modal-formula i + ⊥ₘ : modal-formula i + _→ₘ_ : modal-formula i → modal-formula i → modal-formula i + □ₘ_ : modal-formula i → modal-formula i module _ {l : Level} {i : Set l} where - infixr 25 ~_ - infixr 25 ~~_ + infixr 25 ¬ₘ_ + infixr 25 ¬¬ₘ_ infixl 10 _∨ₘ_ infixl 15 _∧ₘ_ - infixr 25 ◇_ + infixr 25 ◇ₘ_ - ~_ : formula i → formula i - ~ a = a →ₘ ⊥ + ¬ₘ_ : modal-formula i → modal-formula i + ¬ₘ a = a →ₘ ⊥ₘ - ~~_ : formula i → formula i - ~~ a = ~ ~ a + ¬¬ₘ_ : modal-formula i → modal-formula i + ¬¬ₘ a = ¬ₘ ¬ₘ a - _∨ₘ_ : formula i → formula i → formula i - a ∨ₘ b = ~ a →ₘ b + _∨ₘ_ : modal-formula i → modal-formula i → modal-formula i + a ∨ₘ b = ¬ₘ a →ₘ b - _∧ₘ_ : formula i → formula i → formula i - a ∧ₘ b = ~ (a →ₘ ~ b) + _∧ₘ_ : modal-formula i → modal-formula i → modal-formula i + a ∧ₘ b = ¬ₘ (a →ₘ ¬ₘ b) - ⊤ : formula i - ⊤ = ~ ⊥ + ⊤ₘ : modal-formula i + ⊤ₘ = ¬ₘ ⊥ₘ - ◇_ : formula i → formula i - ◇ a = ~ □ ~ a + ◇ₘ_ : modal-formula i → modal-formula i + ◇ₘ a = ¬ₘ □ₘ ¬ₘ a - eq-implication-left : {a b c d : formula i} → a →ₘ b = c →ₘ d → a = c + eq-implication-left : {a b c d : modal-formula i} → a →ₘ b = c →ₘ d → a = c eq-implication-left refl = refl - eq-implication-right : {a b c d : formula i} → a →ₘ b = c →ₘ d → b = d + eq-implication-right : {a b c d : modal-formula i} → a →ₘ b = c →ₘ d → b = d eq-implication-right refl = refl - eq-box : {a b : formula i} → □ a = □ b → a = b + eq-box : {a b : modal-formula i} → □ₘ a = □ₘ b → a = b eq-box refl = refl - eq-diamond : {a b : formula i} → ◇ a = ◇ b → a = b + eq-diamond : {a b : modal-formula i} → ◇ₘ a = ◇ₘ b → a = b eq-diamond refl = refl ``` @@ -95,72 +95,72 @@ module _ {l : Level} (i : Set l) where - Eq-formula : formula i → formula i → UU l + Eq-formula : modal-formula i → modal-formula i → UU l Eq-formula (var n) (var m) = n = m - Eq-formula (var _) ⊥ = raise-empty l + Eq-formula (var _) ⊥ₘ = raise-empty l Eq-formula (var _) (_ →ₘ _) = raise-empty l - Eq-formula (var -) (□ _) = raise-empty l - Eq-formula ⊥ (var _) = raise-empty l - Eq-formula ⊥ ⊥ = raise-unit l - Eq-formula ⊥ (_ →ₘ _) = raise-empty l - Eq-formula ⊥ (□ _) = raise-empty l + Eq-formula (var -) (□ₘ _) = raise-empty l + Eq-formula ⊥ₘ (var _) = raise-empty l + Eq-formula ⊥ₘ ⊥ₘ = raise-unit l + Eq-formula ⊥ₘ (_ →ₘ _) = raise-empty l + Eq-formula ⊥ₘ (□ₘ _) = raise-empty l Eq-formula (_ →ₘ _) (var _) = raise-empty l - Eq-formula (_ →ₘ _) ⊥ = raise-empty l + Eq-formula (_ →ₘ _) ⊥ₘ = raise-empty l Eq-formula (a →ₘ b) (c →ₘ d) = (Eq-formula a c) × (Eq-formula b d) - Eq-formula (_ →ₘ _) (□ _) = raise-empty l - Eq-formula (□ _) (var _) = raise-empty l - Eq-formula (□ _) ⊥ = raise-empty l - Eq-formula (□ _) (_ →ₘ _) = raise-empty l - Eq-formula (□ a) (□ c) = Eq-formula a c + Eq-formula (_ →ₘ _) (□ₘ _) = raise-empty l + Eq-formula (□ₘ _) (var _) = raise-empty l + Eq-formula (□ₘ _) ⊥ₘ = raise-empty l + Eq-formula (□ₘ _) (_ →ₘ _) = raise-empty l + Eq-formula (□ₘ a) (□ₘ c) = Eq-formula a c - refl-Eq-formula : (a : formula i) → Eq-formula a a + refl-Eq-formula : (a : modal-formula i) → Eq-formula a a refl-Eq-formula (var n) = refl - refl-Eq-formula ⊥ = raise-star + refl-Eq-formula ⊥ₘ = raise-star refl-Eq-formula (a →ₘ b) = (refl-Eq-formula a) , (refl-Eq-formula b) - refl-Eq-formula (□ a) = refl-Eq-formula a + refl-Eq-formula (□ₘ a) = refl-Eq-formula a - Eq-eq-formula : {a b : formula i} → a = b → Eq-formula a b + Eq-eq-formula : {a b : modal-formula i} → a = b → Eq-formula a b Eq-eq-formula {a} refl = refl-Eq-formula a - eq-Eq-formula : {a b : formula i} → Eq-formula a b → a = b + eq-Eq-formula : {a b : modal-formula i} → Eq-formula a b → a = b eq-Eq-formula {var _} {var _} refl = refl - eq-Eq-formula {var _} {⊥} (map-raise ()) + eq-Eq-formula {var _} {⊥ₘ} (map-raise ()) eq-Eq-formula {var _} {_ →ₘ _} (map-raise ()) - eq-Eq-formula {var _} {□ _} (map-raise ()) - eq-Eq-formula {⊥} {var _} (map-raise ()) - eq-Eq-formula {⊥} {⊥} _ = refl - eq-Eq-formula {⊥} {_ →ₘ _} (map-raise ()) - eq-Eq-formula {⊥} {□ _} (map-raise ()) + eq-Eq-formula {var _} {□ₘ _} (map-raise ()) + eq-Eq-formula {⊥ₘ} {var _} (map-raise ()) + eq-Eq-formula {⊥ₘ} {⊥ₘ} _ = refl + eq-Eq-formula {⊥ₘ} {_ →ₘ _} (map-raise ()) + eq-Eq-formula {⊥ₘ} {□ₘ _} (map-raise ()) eq-Eq-formula {_ →ₘ _} {var _} (map-raise ()) - eq-Eq-formula {_ →ₘ _} {⊥} (map-raise ()) + eq-Eq-formula {_ →ₘ _} {⊥ₘ} (map-raise ()) eq-Eq-formula {a →ₘ b} {c →ₘ d} (eq1 , eq2) = ap (λ (x , y) → x →ₘ y) (eq-pair (eq-Eq-formula eq1) (eq-Eq-formula eq2)) - eq-Eq-formula {_ →ₘ _} {□ _} (map-raise ()) - eq-Eq-formula {□ _} {var _} (map-raise ()) - eq-Eq-formula {□ _} {⊥} (map-raise ()) - eq-Eq-formula {□ _} {_ →ₘ _} (map-raise ()) - eq-Eq-formula {□ _} {□ _} eq = ap □_ (eq-Eq-formula eq) + eq-Eq-formula {_ →ₘ _} {□ₘ _} (map-raise ()) + eq-Eq-formula {□ₘ _} {var _} (map-raise ()) + eq-Eq-formula {□ₘ _} {⊥ₘ} (map-raise ()) + eq-Eq-formula {□ₘ _} {_ →ₘ _} (map-raise ()) + eq-Eq-formula {□ₘ _} {□ₘ _} eq = ap □ₘ_ (eq-Eq-formula eq) - is-prop-Eq-formula : (a b : formula i) → is-prop (Eq-formula a b) + is-prop-Eq-formula : (a b : modal-formula i) → is-prop (Eq-formula a b) is-prop-Eq-formula (var n) (var m) = is-prop-type-Prop (Id-Prop i n m) - is-prop-Eq-formula (var _) ⊥ = is-prop-raise-empty + is-prop-Eq-formula (var _) ⊥ₘ = is-prop-raise-empty is-prop-Eq-formula (var _) (_ →ₘ _) = is-prop-raise-empty - is-prop-Eq-formula (var -) (□ _) = is-prop-raise-empty - is-prop-Eq-formula ⊥ (var _) = is-prop-raise-empty - is-prop-Eq-formula ⊥ ⊥ = is-prop-raise-unit - is-prop-Eq-formula ⊥ (_ →ₘ _) = is-prop-raise-empty - is-prop-Eq-formula ⊥ (□ _) = is-prop-raise-empty + is-prop-Eq-formula (var -) (□ₘ _) = is-prop-raise-empty + is-prop-Eq-formula ⊥ₘ (var _) = is-prop-raise-empty + is-prop-Eq-formula ⊥ₘ ⊥ₘ = is-prop-raise-unit + is-prop-Eq-formula ⊥ₘ (_ →ₘ _) = is-prop-raise-empty + is-prop-Eq-formula ⊥ₘ (□ₘ _) = is-prop-raise-empty is-prop-Eq-formula (_ →ₘ _) (var _) = is-prop-raise-empty - is-prop-Eq-formula (_ →ₘ _) ⊥ = is-prop-raise-empty + is-prop-Eq-formula (_ →ₘ _) ⊥ₘ = is-prop-raise-empty is-prop-Eq-formula (a →ₘ b) (c →ₘ d) = is-prop-product (is-prop-Eq-formula a c) (is-prop-Eq-formula b d) - is-prop-Eq-formula (_ →ₘ _) (□ _) = is-prop-raise-empty - is-prop-Eq-formula (□ _) (var _) = is-prop-raise-empty - is-prop-Eq-formula (□ _) ⊥ = is-prop-raise-empty - is-prop-Eq-formula (□ _) (_ →ₘ _) = is-prop-raise-empty - is-prop-Eq-formula (□ a) (□ c) = is-prop-Eq-formula a c + is-prop-Eq-formula (_ →ₘ _) (□ₘ _) = is-prop-raise-empty + is-prop-Eq-formula (□ₘ _) (var _) = is-prop-raise-empty + is-prop-Eq-formula (□ₘ _) ⊥ₘ = is-prop-raise-empty + is-prop-Eq-formula (□ₘ _) (_ →ₘ _) = is-prop-raise-empty + is-prop-Eq-formula (□ₘ a) (□ₘ c) = is-prop-Eq-formula a c - is-set-formula : is-set (formula i) + is-set-formula : is-set (modal-formula i) is-set-formula = is-set-prop-in-id ( Eq-formula) @@ -169,9 +169,10 @@ module _ ( λ _ _ → eq-Eq-formula) formula-Set : Set l - pr1 formula-Set = formula i + pr1 formula-Set = modal-formula i pr2 formula-Set = is-set-formula -Id-formula-Prop : {l : Level} {i : Set l} → formula i → formula i → Prop l +Id-formula-Prop : + {l : Level} {i : Set l} → modal-formula i → modal-formula i → Prop l Id-formula-Prop = Id-Prop (formula-Set _) ``` diff --git a/src/modal-logic/kripke-models-filtrations-theorem.lagda.md b/src/modal-logic/kripke-models-filtrations-theorem.lagda.md index b1ec052e76..4fac92f4bd 100644 --- a/src/modal-logic/kripke-models-filtrations-theorem.lagda.md +++ b/src/modal-logic/kripke-models-filtrations-theorem.lagda.md @@ -35,10 +35,10 @@ open import foundation-core.equivalence-relations open import foundation-core.invertible-maps open import modal-logic.completeness +open import modal-logic.deduction open import modal-logic.formulas open import modal-logic.kripke-models-filtrations open import modal-logic.kripke-semantics -open import modal-logic.logic-syntax ```
@@ -59,7 +59,7 @@ module _ kripke-models-filtrations-theorem' : (is-filtration : is-kripke-model-filtration i theory M M*) - (a : formula i) → + (a : modal-formula i) → is-in-subtype theory a → (x : type-kripke-model i M) → type-iff-Prop @@ -81,7 +81,7 @@ module _ ( n) ( x)) ( in-theory , map-inv-raise f)) - pr1 (kripke-models-filtrations-theorem' is-filtration ⊥ in-theory x) = + pr1 (kripke-models-filtrations-theorem' is-filtration ⊥ₘ in-theory x) = map-raise ∘ map-inv-raise pr1 (kripke-models-filtrations-theorem' is-filtration (a →ₘ b) in-theory x) fab fa = @@ -95,7 +95,7 @@ module _ ( pr1 (theory-is-closed in-theory)) ( x)) ( fa))) - pr1 (kripke-models-filtrations-theorem' is-filtration (□ a) in-theory x) + pr1 (kripke-models-filtrations-theorem' is-filtration (□ₘ a) in-theory x) f y* r-xy = apply-universal-property-trunc-Prop ( is-inhabited-subtype-equivalence-class @@ -157,7 +157,7 @@ module _ ( n) ( x)) ( map-inv-raise f))) - pr2 (kripke-models-filtrations-theorem' is-filtration ⊥ in-theory x) = + pr2 (kripke-models-filtrations-theorem' is-filtration ⊥ₘ in-theory x) = map-raise ∘ map-inv-raise pr2 (kripke-models-filtrations-theorem' is-filtration (a →ₘ b) in-theory x) fab fa = @@ -171,7 +171,7 @@ module _ ( pr1 (theory-is-closed in-theory)) ( x)) ( fa))) - pr2 (kripke-models-filtrations-theorem' is-filtration (□ a) in-theory x) + pr2 (kripke-models-filtrations-theorem' is-filtration (□ₘ a) in-theory x) f y r-xy = backward-implication ( kripke-models-filtrations-theorem' is-filtration a diff --git a/src/modal-logic/kripke-models-filtrations.lagda.md b/src/modal-logic/kripke-models-filtrations.lagda.md index 45f81455b0..5693b0f3fe 100644 --- a/src/modal-logic/kripke-models-filtrations.lagda.md +++ b/src/modal-logic/kripke-models-filtrations.lagda.md @@ -38,9 +38,9 @@ open import foundation-core.injective-maps open import foundation-core.transport-along-identifications open import modal-logic.axioms +open import modal-logic.deduction open import modal-logic.formulas open import modal-logic.kripke-semantics -open import modal-logic.logic-syntax open import univalent-combinatorics.finite-types ``` @@ -133,21 +133,21 @@ module _ {l2 : Level} (theory : modal-theory l2 i) where - is-modal-theory-has-subformulas-Prop : formula i → Prop l2 + is-modal-theory-has-subformulas-Prop : modal-formula i → Prop l2 is-modal-theory-has-subformulas-Prop (var _) = raise-unit-Prop l2 - is-modal-theory-has-subformulas-Prop ⊥ = raise-unit-Prop l2 + is-modal-theory-has-subformulas-Prop ⊥ₘ = raise-unit-Prop l2 is-modal-theory-has-subformulas-Prop (a →ₘ b) = product-Prop (theory a) (theory b) - is-modal-theory-has-subformulas-Prop (□ a) = theory a + is-modal-theory-has-subformulas-Prop (□ₘ a) = theory a - is-modal-theory-has-subformulas : formula i → UU l2 + is-modal-theory-has-subformulas : modal-formula i → UU l2 is-modal-theory-has-subformulas = type-Prop ∘ is-modal-theory-has-subformulas-Prop is-modal-theory-closed-under-subformulas-Prop : Prop (l1 ⊔ l2) is-modal-theory-closed-under-subformulas-Prop = implicit-Π-Prop - ( formula i) + ( modal-formula i) ( λ a → ( function-Prop ( is-in-subtype theory a) @@ -158,19 +158,21 @@ module _ type-Prop (is-modal-theory-closed-under-subformulas-Prop) is-modal-theory-closed-under-subformulas-condition : - ( {a b : formula i} → + ( {a b : modal-formula i} → is-in-subtype theory (a →ₘ b) → is-in-subtype theory a × is-in-subtype theory b) → - ( {a : formula i} → is-in-subtype theory (□ a) → is-in-subtype theory a) → + ( {a : modal-formula i} → + is-in-subtype theory (□ₘ a) → + is-in-subtype theory a) → is-modal-theory-closed-under-subformulas is-modal-theory-closed-under-subformulas-condition h-impl h-box {var n} _ = raise-star is-modal-theory-closed-under-subformulas-condition - h-impl h-box {⊥} _ = raise-star + h-impl h-box {⊥ₘ} _ = raise-star is-modal-theory-closed-under-subformulas-condition h-impl h-box {a →ₘ b} = h-impl is-modal-theory-closed-under-subformulas-condition - h-impl h-box {□ a} = h-box + h-impl h-box {□ₘ a} = h-box module _ {l1 l2 l3 l4 l5 : Level} (i : Set l3) @@ -182,7 +184,7 @@ module _ equivalence-relation (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) (type-kripke-model i M) pr1 Φ-equivalence x y = Π-Prop - ( formula i) + ( modal-formula i) ( λ a → ( function-Prop ( is-in-subtype theory a) @@ -354,9 +356,9 @@ module _ equivalence-class Φ-equivalence ≃ type-kripke-model i M* → Prop (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5 ⊔ l7) filtration-relation-upper-bound-Prop e = - Π-Prop (formula i) + Π-Prop (modal-formula i) ( λ a → - ( function-Prop (is-in-subtype theory (□ a)) + ( function-Prop (is-in-subtype theory (□ₘ a)) ( Π-Prop (type-kripke-model i M) ( λ x → ( Π-Prop (type-kripke-model i M) @@ -365,7 +367,7 @@ module _ ( relation-kripke-model i M* ( map-equiv e (class Φ-equivalence x)) ( map-equiv e (class Φ-equivalence y))) - ( hom-Prop ((M , x) ⊨ □ a) + ( hom-Prop ((M , x) ⊨ □ₘ a) ( (M , y) ⊨ a))))))))) filtration-relation-upper-bound : @@ -548,13 +550,13 @@ module _ where proof-upper-bound : - (a : formula i) → - is-in-subtype theory (□ a) → + (a : modal-formula i) → + is-in-subtype theory (□ₘ a) → (x y : type-kripke-model i M) → relation-kripke-model i minimal-kripke-model-filtration ( class Φ-equivalence x) ( class Φ-equivalence y) → - type-Prop ((M , x) ⊨ □ a) → + type-Prop ((M , x) ⊨ □ₘ a) → type-Prop ((M , y) ⊨ a) proof-upper-bound a box-in-theory x y r-xy x-forces-box = apply-universal-property-trunc-Prop @@ -564,7 +566,7 @@ module _ ( backward-implication ( iff-y a (theory-is-closed box-in-theory)) ( forward-implication - ( iff-x (□ a) box-in-theory) + ( iff-x (□ₘ a) box-in-theory) ( x-forces-box) ( y') ( r-xy')))) @@ -599,14 +601,14 @@ module _ helper : is-in-subtype (transitivity-kripke-class l1 l2 i l4) M → - (a : formula i) → - is-in-subtype theory (□ a) → + (a : modal-formula i) → + is-in-subtype theory (□ₘ a) → (x y : type-kripke-model i M) → transitive-closure ( relation-kripke-model i minimal-kripke-model-filtration) ( class Φ-equivalence x) ( class Φ-equivalence y) → - type-Prop ((M , x) ⊨ □ a) → + type-Prop ((M , x) ⊨ □ₘ a) → type-Prop ((M , y) ⊨ a) helper M-is-trans a box-in-theory x y (base* r-xy) x-forces-box = proof-upper-bound a box-in-theory x y r-xy x-forces-box @@ -629,11 +631,11 @@ module _ ( λ ((x' , z') , r-xz' , iff-x , iff-z) → ( helper M-is-trans a box-in-theory z y c-zy ( backward-implication - ( iff-z (□ a) box-in-theory) - ( ax-4-soundness i l2 l4 (□ a →ₘ □ □ a) (a , refl) M M-is-trans + ( iff-z (□ₘ a) box-in-theory) + ( ax-4-soundness i l2 l4 _ (a , refl) M M-is-trans ( x') ( forward-implication - ( iff-x (□ a) box-in-theory) + ( iff-x (□ₘ a) box-in-theory) ( x-forces-box)) ( z') ( r-xz'))))) diff --git a/src/modal-logic/kripke-semantics.lagda.md b/src/modal-logic/kripke-semantics.lagda.md index 9f0240a0e8..6778632e19 100644 --- a/src/modal-logic/kripke-semantics.lagda.md +++ b/src/modal-logic/kripke-semantics.lagda.md @@ -33,8 +33,8 @@ open import foundation.universe-levels open import foundation-core.equivalence-relations open import foundation-core.identity-types +open import modal-logic.deduction open import modal-logic.formulas -open import modal-logic.logic-syntax open import univalent-combinatorics.decidable-dependent-function-types open import univalent-combinatorics.finite-types @@ -230,23 +230,23 @@ module _ _⊨_ : Σ (kripke-model l1 l2 i l4) (type-kripke-model i) → - formula i → + modal-formula i → Prop (l1 ⊔ l2 ⊔ l4) (M , x) ⊨ var n = raise-Prop (l1 ⊔ l2) (valuate-kripke-model i M n x) - (M , x) ⊨ ⊥ = raise-empty-Prop (l1 ⊔ l2 ⊔ l4) + (M , x) ⊨ ⊥ₘ = raise-empty-Prop (l1 ⊔ l2 ⊔ l4) (M , x) ⊨ a →ₘ b = hom-Prop ((M , x) ⊨ a) ((M , x) ⊨ b) - (M , x) ⊨ □ a = + (M , x) ⊨ □ₘ a = Π-Prop ( type-kripke-model i M) ( λ y → function-Prop (relation-kripke-model i M x y) ((M , y) ⊨ a)) - _⊨M_ : kripke-model l1 l2 i l4 → formula i → Prop (l1 ⊔ l2 ⊔ l4) + _⊨M_ : kripke-model l1 l2 i l4 → modal-formula i → Prop (l1 ⊔ l2 ⊔ l4) M ⊨M a = Π-Prop (type-kripke-model i M) (λ x → (M , x) ⊨ a) _⊨C_ : {l5 : Level} → model-class l1 l2 i l4 l5 → - formula i → + modal-formula i → Prop (lsuc l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ l5) C ⊨C a = Π-Prop @@ -278,7 +278,7 @@ module _ decidable-models : model-class l1 l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4) decidable-models M = Π-Prop - ( formula i) + ( modal-formula i) ( λ a → ( Π-Prop ( type-kripke-model i M) @@ -312,15 +312,15 @@ module _ finite-models-subclass-decidable-models M (w-is-fin , dec-r , dec-v) = lemma where lemma : - (a : formula i) (x : type-kripke-model i M) → + (a : modal-formula i) (x : type-kripke-model i M) → is-decidable (type-Prop ((M , x) ⊨ a)) lemma (var n) x = is-decidable-raise (l1 ⊔ l2) _ (dec-v x n) - lemma ⊥ x = + lemma ⊥ₘ x = inr map-inv-raise lemma (a →ₘ b) x = is-decidable-function-type (lemma a x) (lemma b x) - lemma (□ a) x = + lemma (□ₘ a) x = is-decidable-Π-is-finite ( w-is-fin) ( λ y → is-decidable-function-type (dec-r x y) (lemma a y)) @@ -328,7 +328,7 @@ module _ is-finite-model-valuate-decidable-models : (M : kripke-model l1 l2 i l4) → is-in-subtype (finite-models l1 l2 i l4) M → - (a : formula i) → + (a : modal-formula i) → is-decidable (type-Prop (M ⊨M a)) is-finite-model-valuate-decidable-models M sub-fin a = is-decidable-Π-is-finite diff --git a/src/modal-logic/l-complete-theories.lagda.md b/src/modal-logic/l-complete-theories.lagda.md index e7f71de3c8..d58be908bf 100644 --- a/src/modal-logic/l-complete-theories.lagda.md +++ b/src/modal-logic/l-complete-theories.lagda.md @@ -35,10 +35,10 @@ open import lists.lists open import lists.lists-subtypes open import modal-logic.axioms +open import modal-logic.deduction open import modal-logic.formulas open import modal-logic.formulas-deduction open import modal-logic.l-consistent-theories -open import modal-logic.logic-syntax open import modal-logic.modal-logic-k open import modal-logic.weak-deduction @@ -149,7 +149,7 @@ module _ ( subset-axioms-weak-modal-logic-closure) ( subtype-union-left logic theory)) ( refl-leq-subtype (weak-modal-logic-closure (logic ∪ theory))))) - ( ⊥) + ( ⊥ₘ) ( bot-in-logic)) module _ @@ -261,46 +261,46 @@ module _ ( contains-ax-dn) is-L-consistent-add-formula-not-in-logic : - {a : formula i} → + {a : modal-formula i} → ¬ (is-in-subtype theory a) → - is-L-consistent-theory logic (theory-add-formula (~ a) theory) + is-L-consistent-theory logic (theory-add-formula (¬ₘ a) theory) is-L-consistent-add-formula-not-in-logic {a} not-in-logic bot-in-logic = not-in-logic ( weak-modal-logic-mp ( is-weak-modal-logic-L-complete-theory lzero t) - ( contains-ax-dn' (~~ a →ₘ a) (a , refl)) - ( is-weak-modal-logic-L-complete-theory lzero t (~~ a) + ( contains-ax-dn' (¬¬ₘ a →ₘ a) (a , refl)) + ( is-weak-modal-logic-L-complete-theory lzero t (¬¬ₘ a) ( forward-implication ( deduction-lemma ( theory) ( contains-ax-k') ( contains-ax-s') - ( ~ a) - ( ⊥)) + ( ¬ₘ a) + ( ⊥ₘ)) ( weak-modal-logic-closure-monotic - ( subtype-union-both logic (theory-add-formula (~ a) theory) - ( theory-add-formula (~ a) theory) + ( subtype-union-both logic (theory-add-formula (¬ₘ a) theory) + ( theory-add-formula (¬ₘ a) theory) ( transitive-leq-subtype ( logic) ( theory) - ( theory-add-formula (~ a) theory) - ( subset-add-formula (~ a) theory) + ( theory-add-formula (¬ₘ a) theory) + ( subset-add-formula (¬ₘ a) theory) ( subset-logic-L-complete-theory lzero t)) - ( refl-leq-subtype (theory-add-formula (~ a) theory))) - ( ⊥) + ( refl-leq-subtype (theory-add-formula (¬ₘ a) theory))) + ( ⊥ₘ) ( bot-in-logic))))) contains-negation-not-contains-formula-L-complete-theory : - {a : formula i} → + {a : modal-formula i} → ¬ (is-in-subtype theory a) → - is-in-subtype theory (~ a) + is-in-subtype theory (¬ₘ a) contains-negation-not-contains-formula-L-complete-theory {a} not-in-logic = tr - ( λ t → is-in-subtype t (~ a)) + ( λ t → is-in-subtype t (¬ₘ a)) ( eq-is-L-consistent-union-L-complete t - ( Id-formula-Prop (~ a)) + ( Id-formula-Prop (¬ₘ a)) ( is-L-consistent-add-formula-not-in-logic not-in-logic)) - ( formula-in-add-formula (~ a) theory) + ( formula-in-add-formula (¬ₘ a) theory) module _ (lem : LEM (l1 ⊔ l2)) @@ -330,14 +330,14 @@ module _ ( is-consistent-modal-theory-L-complete-theory x ( weak-modal-logic-mp ( is-weak-modal-logic-L-complete-theory lzero x) - { a = □ a} + { a = □ₘ a} ( weak-modal-logic-mp ( is-weak-modal-logic-L-complete-theory lzero x) - { a = ◇ ~ a} - ( subset-logic-L-complete-theory lzero x (◇ ~ a →ₘ ~ □ a) + { a = ◇ₘ ¬ₘ a} + ( subset-logic-L-complete-theory lzero x (◇ₘ ¬ₘ a →ₘ ¬ₘ □ₘ a) ( modal-logic-diamond-negate-implication i logic is-normal ( is-logic))) - ( leq (◇ ~ a) (intro-exists (~ a) (not-a-in-t , refl)))) + ( leq (◇ₘ ¬ₘ a) (intro-exists (¬ₘ a) (not-a-in-t , refl)))) ( box-a-in-x))) is-inhabited-L-complete-exists-complete-L-consistent-theory : @@ -395,14 +395,14 @@ module _ ∃ (type-chain-Poset P C) (λ x → modal-theory-chain-element x a) exists-chain-element-with-formula-Prop : - (a : formula i) → Prop (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4) + (a : modal-formula i) → Prop (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4) exists-chain-element-with-formula-Prop a = ∃ (type-chain-Poset P C) ( λ x → ( weak-modal-logic-closure (logic ∪ modal-theory-chain-element x) a)) exists-chain-element-with-formula : - (a : formula i) → UU (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4) + (a : modal-formula i) → UU (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4) exists-chain-element-with-formula = type-Prop ∘ exists-chain-element-with-formula-Prop @@ -427,8 +427,8 @@ module _ ( contains-ax-s) L-union-deduction-lemma : - (l : list (formula i)) → - (h a : formula i) → + (l : list (modal-formula i)) → + (h a : modal-formula i) → is-in-subtype ( weak-modal-logic-closure (logic ∪ list-subtype (cons h l))) a → is-in-subtype @@ -457,9 +457,9 @@ module _ in-chain-in-chain-union-assumptions : is-inhabited (type-chain-Poset P C) → - (l : list (formula i)) → + (l : list (modal-formula i)) → list-subtype l ⊆ chain-union-modal-theory → - {a : formula i} → + {a : modal-formula i} → is-in-subtype (weak-modal-logic-closure (logic ∪ list-subtype l)) a → exists-chain-element-with-formula a in-chain-in-chain-union-assumptions is-inh nil leq {a} in-logic = @@ -512,7 +512,7 @@ module _ in-chain-in-chain-union : is-inhabited (type-chain-Poset P C) → - {a : formula i} → + {a : modal-formula i} → is-in-subtype ( weak-modal-logic-closure (logic ∪ chain-union-modal-theory)) ( a) → diff --git a/src/modal-logic/l-consistent-theories.lagda.md b/src/modal-logic/l-consistent-theories.lagda.md index 73ed7e888f..45522f1883 100644 --- a/src/modal-logic/l-consistent-theories.lagda.md +++ b/src/modal-logic/l-consistent-theories.lagda.md @@ -28,8 +28,8 @@ open import foundation-core.sets open import foundation-core.subtypes open import modal-logic.axioms +open import modal-logic.deduction open import modal-logic.formulas -open import modal-logic.logic-syntax open import modal-logic.weak-deduction open import order-theory.maximal-elements-posets @@ -124,7 +124,7 @@ module _ L-consistent-theory-leq = type-Relation-Prop L-consistent-theory-leq-Prop theories-Poset : (l3 : Level) → Poset (l1 ⊔ lsuc l3) (l1 ⊔ l3) - theories-Poset l3 = subtypes-leq-Poset l3 (formula i) + theories-Poset l3 = subtypes-leq-Poset l3 (modal-formula i) L-consistent-theories-Poset : (l3 : Level) → Poset (l1 ⊔ l2 ⊔ lsuc l3) (l1 ⊔ l3) diff --git a/src/modal-logic/lindenbaum.lagda.md b/src/modal-logic/lindenbaum.lagda.md index 1361ed879a..4200ca90d7 100644 --- a/src/modal-logic/lindenbaum.lagda.md +++ b/src/modal-logic/lindenbaum.lagda.md @@ -20,9 +20,9 @@ open import foundation-core.sets open import foundation-core.subtypes open import modal-logic.axioms +open import modal-logic.deduction open import modal-logic.l-complete-theories open import modal-logic.l-consistent-theories -open import modal-logic.logic-syntax open import modal-logic.weak-deduction open import order-theory.zorn diff --git a/src/modal-logic/modal-logic-k.lagda.md b/src/modal-logic/modal-logic-k.lagda.md index 05aa4f94d7..8f31830b90 100644 --- a/src/modal-logic/modal-logic-k.lagda.md +++ b/src/modal-logic/modal-logic-k.lagda.md @@ -24,9 +24,9 @@ open import foundation.universe-levels open import foundation-core.coproduct-types open import modal-logic.axioms +open import modal-logic.deduction open import modal-logic.formulas open import modal-logic.kripke-semantics -open import modal-logic.logic-syntax open import modal-logic.soundness open import univalent-combinatorics.finite-types @@ -189,7 +189,7 @@ module _ map-inv-raise ( soundness-K-finite ( i) - ( ⊥) + ( ⊥ₘ) ( bot-in-logic) ( pair ( pair (unit , unit-trunc-Prop star) (λ _ _ → empty-Prop)) diff --git a/src/modal-logic/modal-logic-s5.lagda.md b/src/modal-logic/modal-logic-s5.lagda.md index 5e3305d1d8..4bb29a447b 100644 --- a/src/modal-logic/modal-logic-s5.lagda.md +++ b/src/modal-logic/modal-logic-s5.lagda.md @@ -22,9 +22,9 @@ open import foundation.universe-levels open import foundation-core.coproduct-types open import modal-logic.axioms +open import modal-logic.deduction open import modal-logic.formulas open import modal-logic.kripke-semantics -open import modal-logic.logic-syntax open import modal-logic.modal-logic-k open import modal-logic.soundness @@ -210,7 +210,7 @@ module _ map-inv-raise ( soundness-S5-finite ( i) - ( ⊥) + ( ⊥ₘ) ( bot-in-logic) ( pair ( pair (unit , unit-trunc-Prop star) (λ _ _ → unit-Prop)) diff --git a/src/modal-logic/soundness.lagda.md b/src/modal-logic/soundness.lagda.md index 8296c2f350..5e7a04cee6 100644 --- a/src/modal-logic/soundness.lagda.md +++ b/src/modal-logic/soundness.lagda.md @@ -20,9 +20,9 @@ open import foundation.universe-levels open import foundation-core.coproduct-types +open import modal-logic.deduction open import modal-logic.formulas open import modal-logic.kripke-semantics -open import modal-logic.logic-syntax ```
@@ -52,7 +52,7 @@ module _ where soundness-axioms : - soundness axioms C → {a : formula i} → axioms ⊢ a → type-Prop (C ⊨C a) + soundness axioms C → {a : modal-formula i} → axioms ⊢ a → type-Prop (C ⊨C a) soundness-axioms H (ax x) = H _ x soundness-axioms H (mp dab da) M in-class x = soundness-axioms H dab M in-class x (soundness-axioms H da M in-class x) @@ -88,7 +88,7 @@ module _ forces-in-intersection : (M : kripke-model l4 l5 i l6) → is-in-subtype (C₁ ∩ C₂) M → - (a : formula i) → + (a : modal-formula i) → is-in-subtype theory₁ a + is-in-subtype theory₂ a → type-Prop (M ⊨M a) forces-in-intersection M in-class a (inl d) = diff --git a/src/modal-logic/weak-deduction.lagda.md b/src/modal-logic/weak-deduction.lagda.md index cc5c68bf9f..2db8b615a9 100644 --- a/src/modal-logic/weak-deduction.lagda.md +++ b/src/modal-logic/weak-deduction.lagda.md @@ -34,8 +34,8 @@ open import lists.lists-subtypes open import lists.reversing-lists open import modal-logic.axioms +open import modal-logic.deduction open import modal-logic.formulas -open import modal-logic.logic-syntax ``` @@ -52,39 +52,41 @@ module _ where is-weak-deduction-Prop : - {axioms : modal-theory l2 i} {a : formula i} → axioms ⊢ a → Prop lzero + {axioms : modal-theory l2 i} {a : modal-formula i} → axioms ⊢ a → Prop lzero is-weak-deduction-Prop (ax x) = unit-Prop is-weak-deduction-Prop (mp dab da) = conjunction-Prop (is-weak-deduction-Prop dab) (is-weak-deduction-Prop da) is-weak-deduction-Prop (nec d) = empty-Prop is-weak-deduction : - {axioms : modal-theory l2 i} {a : formula i} → axioms ⊢ a → UU lzero + {axioms : modal-theory l2 i} {a : modal-formula i} → axioms ⊢ a → UU lzero is-weak-deduction = type-Prop ∘ is-weak-deduction-Prop infix 5 _⊢w_ - _⊢w_ : modal-theory l2 i → formula i → UU (l1 ⊔ l2) + _⊢w_ : modal-theory l2 i → modal-formula i → UU (l1 ⊔ l2) axioms ⊢w a = type-subtype (is-weak-deduction-Prop {axioms} {a}) deduction-weak-deduction : - {axioms : modal-theory l2 i} {a : formula i} → axioms ⊢w a → axioms ⊢ a + {axioms : modal-theory l2 i} {a : modal-formula i} → + axioms ⊢w a → + axioms ⊢ a deduction-weak-deduction = inclusion-subtype is-weak-deduction-Prop is-weak-deduction-deduction-weak-deduction : - {axioms : modal-theory l2 i} {a : formula i} (d : axioms ⊢w a) → + {axioms : modal-theory l2 i} {a : modal-formula i} (d : axioms ⊢w a) → is-weak-deduction (deduction-weak-deduction d) is-weak-deduction-deduction-weak-deduction = is-in-subtype-inclusion-subtype is-weak-deduction-Prop weak-deduction-ax : - {axioms : modal-theory l2 i} {a : formula i} → + {axioms : modal-theory l2 i} {a : modal-formula i} → is-in-subtype axioms a → axioms ⊢w a weak-deduction-ax in-axioms = ax in-axioms , star weak-deduction-mp : - {axioms : modal-theory l2 i} {a b : formula i} → + {axioms : modal-theory l2 i} {a b : modal-formula i} → axioms ⊢w (a →ₘ b) → axioms ⊢w a → axioms ⊢w b @@ -93,12 +95,12 @@ module _ ind-weak-deduction : {l : Level} {axioms : modal-theory l2 i} - (P : {a : formula i} → axioms ⊢w a → UU l) → - ( {a : formula i} (in-axioms : is-in-subtype axioms a) → + (P : {a : modal-formula i} → axioms ⊢w a → UU l) → + ( {a : modal-formula i} (in-axioms : is-in-subtype axioms a) → P (weak-deduction-ax in-axioms)) → - ( {a b : formula i} (dab : axioms ⊢w a →ₘ b) (da : axioms ⊢w a) → + ( {a b : modal-formula i} (dab : axioms ⊢w a →ₘ b) (da : axioms ⊢w a) → P dab → P da → P (weak-deduction-mp dab da)) → - {a : formula i} (wd : axioms ⊢w a) → P wd + {a : modal-formula i} (wd : axioms ⊢w a) → P wd ind-weak-deduction P H-ax H-mp (ax in-axioms , _) = H-ax in-axioms ind-weak-deduction P H-ax H-mp (mp dba db , is-w-dba , is-w-db) = @@ -108,10 +110,10 @@ module _ rec-weak-deduction : {l : Level} {axioms : modal-theory l2 i} {P : UU l} → - ( {a : formula i} (in-axioms : is-in-subtype axioms a) → P) → - ( {a b : formula i} (dab : axioms ⊢w a →ₘ b) (da : axioms ⊢w a) → + ( {a : modal-formula i} (in-axioms : is-in-subtype axioms a) → P) → + ( {a b : modal-formula i} (dab : axioms ⊢w a →ₘ b) (da : axioms ⊢w a) → P → P → P) → - {a : formula i} (wd : axioms ⊢w a) → P + {a : modal-formula i} (wd : axioms ⊢w a) → P rec-weak-deduction {P = P} = ind-weak-deduction (λ _ → P) weak-modal-logic-closure : modal-theory l2 i → modal-theory (l1 ⊔ l2) i @@ -125,7 +127,7 @@ module _ is-weak-modal-logic = type-Prop ∘ is-weak-modal-logic-Prop is-in-weak-modal-logic-closure-weak-deduction : - {axioms : modal-theory l2 i} {a : formula i} → + {axioms : modal-theory l2 i} {a : modal-formula i} → axioms ⊢w a → is-in-subtype (weak-modal-logic-closure axioms) a is-in-weak-modal-logic-closure-weak-deduction = unit-trunc-Prop @@ -154,14 +156,14 @@ module _ where weak-modal-logic-closure-ax : - {a : formula i} → + {a : modal-formula i} → is-in-subtype axioms a → is-in-subtype (weak-modal-logic-closure axioms) a weak-modal-logic-closure-ax = is-in-weak-modal-logic-closure-weak-deduction ∘ weak-deduction-ax weak-modal-logic-closure-mp : - {a b : formula i} → + {a b : modal-formula i} → is-in-subtype (weak-modal-logic-closure axioms) (a →ₘ b) → is-in-subtype (weak-modal-logic-closure axioms) a → is-in-subtype (weak-modal-logic-closure axioms) b @@ -179,7 +181,7 @@ module _ where weak-modal-logic-mp : - {a b : formula i} → + {a b : modal-formula i} → is-in-subtype theory (a →ₘ b) → is-in-subtype theory a → is-in-subtype theory b @@ -194,7 +196,7 @@ module _ where weak-modal-logic-closed : - {a : formula i} → weak-modal-logic-closure axioms ⊢w a → + {a : modal-formula i} → weak-modal-logic-closure axioms ⊢w a → is-in-subtype (weak-modal-logic-closure axioms) a weak-modal-logic-closed = ind-weak-deduction _ @@ -220,7 +222,7 @@ module _ (leq : ax₁ ⊆ ax₂) where - weak-deduction-monotic : {a : formula i} → ax₁ ⊢w a → ax₂ ⊢w a + weak-deduction-monotic : {a : modal-formula i} → ax₁ ⊢w a → ax₂ ⊢w a weak-deduction-monotic = ind-weak-deduction _ ( λ {a} in-axioms → weak-deduction-ax (leq a in-axioms)) @@ -255,7 +257,7 @@ module _ where backward-subset-head-add : - (a : formula i) (l : list (formula i)) → + (a : modal-formula i) (l : list (modal-formula i)) → list-subtype (cons a l) ⊆ theory-add-formula a (list-subtype l) backward-subset-head-add a l = subset-list-subtype-cons @@ -269,7 +271,7 @@ module _ where backward-deduction-lemma : - {a b : formula i} → + {a b : modal-formula i} → axioms ⊢w a →ₘ b → theory-add-formula a axioms ⊢w b backward-deduction-lemma {a} wab = @@ -286,7 +288,7 @@ module _ -- TODO: move to file with deduction deduction-a->a : - (a : formula i) → axioms ⊢w a →ₘ a + (a : modal-formula i) → axioms ⊢w a →ₘ a deduction-a->a a = weak-deduction-mp ( weak-deduction-mp @@ -295,7 +297,7 @@ module _ ( weak-deduction-ax (contains-ax-k _ (a , a , refl))) forward-deduction-lemma : - (a : formula i) {b : formula i} → + (a : modal-formula i) {b : modal-formula i} → theory-add-formula a axioms ⊢w b → is-in-subtype (weak-modal-logic-closure axioms) (a →ₘ b) forward-deduction-lemma a = @@ -320,7 +322,7 @@ module _ ( dab))) deduction-lemma : - (a b : formula i) → + (a b : modal-formula i) → type-iff-Prop ( weak-modal-logic-closure (theory-add-formula a axioms) b) ( weak-modal-logic-closure axioms (a →ₘ b)) @@ -343,15 +345,15 @@ module _ where list-assumptions-weak-deduction : - {l2 : Level} {theory : modal-theory l2 i} {a : formula i} → - theory ⊢w a → list (formula i) + {l2 : Level} {theory : modal-theory l2 i} {a : modal-formula i} → + theory ⊢w a → list (modal-formula i) list-assumptions-weak-deduction = rec-weak-deduction ( λ {a} _ → cons a nil) ( λ _ _ l1 l2 → concat-list l1 l2) subset-theory-list-assumptions-weak-deduction : - {l2 : Level} {theory : modal-theory l2 i} {a : formula i} → + {l2 : Level} {theory : modal-theory l2 i} {a : modal-formula i} → (d : theory ⊢w a) → list-subtype (list-assumptions-weak-deduction d) ⊆ theory subset-theory-list-assumptions-weak-deduction {theory = theory} = @@ -378,7 +380,7 @@ module _ ( subset-list-subtype-concat-union))) is-assumptions-list-assumptions-weak-deduction : - {l2 : Level} {theory : modal-theory l2 i} {a : formula i} → + {l2 : Level} {theory : modal-theory l2 i} {a : modal-formula i} → (d : theory ⊢w a) → list-subtype (list-assumptions-weak-deduction d) ⊢w a is-assumptions-list-assumptions-weak-deduction {theory = theory} = @@ -412,61 +414,61 @@ module _ -- TODO: move to formulas-deduction deduction-ex-falso : - (a b : formula i) → - is-in-subtype (weak-modal-logic-closure axioms) (~ a →ₘ a →ₘ b) + (a b : modal-formula i) → + is-in-subtype (weak-modal-logic-closure axioms) (¬ₘ a →ₘ a →ₘ b) deduction-ex-falso a b = forward-implication - ( deduction-lemma axioms contains-ax-k contains-ax-s (~ a) (a →ₘ b)) + ( deduction-lemma axioms contains-ax-k contains-ax-s (¬ₘ a) (a →ₘ b)) ( forward-implication ( deduction-lemma - ( theory-add-formula (~ a) axioms) + ( theory-add-formula (¬ₘ a) axioms) ( contains-ax-k') ( contains-ax-s') ( a) ( b)) - ( weak-modal-logic-closure-mp {a = ~~ b} + ( weak-modal-logic-closure-mp {a = ¬¬ₘ b} ( weak-modal-logic-closure-ax - ( contains-ax-dn'' (~~ b →ₘ b) (b , refl))) - ( weak-modal-logic-closure-mp {a = ⊥} + ( contains-ax-dn'' (¬¬ₘ b →ₘ b) (b , refl))) + ( weak-modal-logic-closure-mp {a = ⊥ₘ} ( weak-modal-logic-closure-ax - ( contains-ax-k'' (⊥ →ₘ ~ b →ₘ ⊥) (⊥ , ~ b , refl))) + ( contains-ax-k'' (⊥ₘ →ₘ ¬ₘ b →ₘ ⊥ₘ) (⊥ₘ , ¬ₘ b , refl))) ( weak-modal-logic-closure-mp {a = a} ( weak-modal-logic-closure-ax ( subset-add-formula a - ( theory-add-formula (~ a) axioms) - ( ~ a) - ( formula-in-add-formula (~ a) axioms))) + ( theory-add-formula (¬ₘ a) axioms) + ( ¬ₘ a) + ( formula-in-add-formula (¬ₘ a) axioms))) ( weak-modal-logic-closure-ax ( formula-in-add-formula a - ( theory-add-formula (~ a) axioms))))))) + ( theory-add-formula (¬ₘ a) axioms))))))) where - contains-ax-k' : ax-k i ⊆ theory-add-formula (~ a) axioms + contains-ax-k' : ax-k i ⊆ theory-add-formula (¬ₘ a) axioms contains-ax-k' = - transitive-subset-add-formula (~ a) axioms (ax-k i) contains-ax-k + transitive-subset-add-formula (¬ₘ a) axioms (ax-k i) contains-ax-k - contains-ax-s' : ax-s i ⊆ theory-add-formula (~ a) axioms + contains-ax-s' : ax-s i ⊆ theory-add-formula (¬ₘ a) axioms contains-ax-s' = - transitive-subset-add-formula (~ a) axioms (ax-s i) contains-ax-s + transitive-subset-add-formula (¬ₘ a) axioms (ax-s i) contains-ax-s contains-ax-k'' : - ax-k i ⊆ theory-add-formula a (theory-add-formula (~ a) axioms) + ax-k i ⊆ theory-add-formula a (theory-add-formula (¬ₘ a) axioms) contains-ax-k'' = - transitive-subset-add-formula a (theory-add-formula (~ a) axioms) + transitive-subset-add-formula a (theory-add-formula (¬ₘ a) axioms) ( ax-k i) ( contains-ax-k') contains-ax-dn'' : - ax-dn i ⊆ theory-add-formula a (theory-add-formula (~ a) axioms) + ax-dn i ⊆ theory-add-formula a (theory-add-formula (¬ₘ a) axioms) contains-ax-dn'' = transitive-subset-add-formula a - ( theory-add-formula (~ a) axioms) + ( theory-add-formula (¬ₘ a) axioms) ( ax-dn i) - ( transitive-subset-add-formula (~ a) axioms (ax-dn i) contains-ax-dn) + ( transitive-subset-add-formula (¬ₘ a) axioms (ax-dn i) contains-ax-dn) logic-ex-falso : - (a b : formula i) → + (a b : modal-formula i) → is-in-subtype (weak-modal-logic-closure axioms) a → - is-in-subtype (weak-modal-logic-closure axioms) (~ a) → + is-in-subtype (weak-modal-logic-closure axioms) (¬ₘ a) → is-in-subtype (weak-modal-logic-closure axioms) b logic-ex-falso a b a-in-logic not-a-in-logic = weak-modal-logic-closure-mp @@ -482,30 +484,30 @@ module _ where inv-deduction-ex-falso : - (a b : formula i) → - is-in-subtype (weak-modal-logic-closure axioms) (a →ₘ ~ a →ₘ b) + (a b : modal-formula i) → + is-in-subtype (weak-modal-logic-closure axioms) (a →ₘ ¬ₘ a →ₘ b) inv-deduction-ex-falso a b = forward-implication - ( deduction-lemma axioms contains-ax-k contains-ax-s a (~ a →ₘ b)) + ( deduction-lemma axioms contains-ax-k contains-ax-s a (¬ₘ a →ₘ b)) ( forward-implication ( deduction-lemma ( theory-add-formula a axioms) ( contains-ax-k') ( contains-ax-s') - ( ~ a) + ( ¬ₘ a) ( b)) ( logic-ex-falso - ( theory-add-formula (a →ₘ ⊥) (theory-add-formula a axioms)) + ( theory-add-formula (a →ₘ ⊥ₘ) (theory-add-formula a axioms)) ( contains-ax-k'') ( contains-ax-s'') ( contains-ax-dn'') ( a) ( b) ( weak-modal-logic-closure-ax - ( subset-add-formula (~ a) (theory-add-formula a axioms) a + ( subset-add-formula (¬ₘ a) (theory-add-formula a axioms) a ( formula-in-add-formula a axioms))) ( weak-modal-logic-closure-ax - ( formula-in-add-formula (~ a) (theory-add-formula a axioms))))) + ( formula-in-add-formula (¬ₘ a) (theory-add-formula a axioms))))) where contains-ax-k' : ax-k i ⊆ theory-add-formula a axioms contains-ax-k' = @@ -516,23 +518,23 @@ module _ transitive-subset-add-formula a axioms (ax-s i) contains-ax-s contains-ax-k'' : - ax-k i ⊆ theory-add-formula (~ a) (theory-add-formula a axioms) + ax-k i ⊆ theory-add-formula (¬ₘ a) (theory-add-formula a axioms) contains-ax-k'' = - transitive-subset-add-formula (~ a) (theory-add-formula a axioms) + transitive-subset-add-formula (¬ₘ a) (theory-add-formula a axioms) ( ax-k i) ( contains-ax-k') contains-ax-s'' : - ax-s i ⊆ theory-add-formula (~ a) (theory-add-formula a axioms) + ax-s i ⊆ theory-add-formula (¬ₘ a) (theory-add-formula a axioms) contains-ax-s'' = - transitive-subset-add-formula (~ a) (theory-add-formula a axioms) + transitive-subset-add-formula (¬ₘ a) (theory-add-formula a axioms) ( ax-s i) ( contains-ax-s') contains-ax-dn'' : - ax-dn i ⊆ theory-add-formula (~ a) (theory-add-formula a axioms) + ax-dn i ⊆ theory-add-formula (¬ₘ a) (theory-add-formula a axioms) contains-ax-dn'' = - transitive-subset-add-formula (~ a) + transitive-subset-add-formula (¬ₘ a) ( theory-add-formula a axioms) ( ax-dn i) ( transitive-subset-add-formula a axioms (ax-dn i) contains-ax-dn) From 662bc0f30c061dd5b84bd2764c6722f08a7c1113 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Sun, 5 May 2024 17:39:00 +0300 Subject: [PATCH 51/63] Refactor var --- .../canonical-model-theorem.lagda.md | 6 +-- src/modal-logic/decision-procedure.lagda.md | 2 +- src/modal-logic/formulas.lagda.md | 46 +++++++++---------- ...kripke-models-filtrations-theorem.lagda.md | 44 +++++++++--------- .../kripke-models-filtrations.lagda.md | 12 ++--- src/modal-logic/kripke-semantics.lagda.md | 4 +- 6 files changed, 58 insertions(+), 56 deletions(-) diff --git a/src/modal-logic/canonical-model-theorem.lagda.md b/src/modal-logic/canonical-model-theorem.lagda.md index 46f7c70de5..077875b17d 100644 --- a/src/modal-logic/canonical-model-theorem.lagda.md +++ b/src/modal-logic/canonical-model-theorem.lagda.md @@ -114,7 +114,7 @@ module _ ( modal-theory-L-complete-theory logic x (□ₘ a)) ( modal-theory-L-complete-theory logic y a))) pr2 canonical-kripke-model n x = - modal-theory-L-complete-theory logic x (var n) + modal-theory-L-complete-theory logic x (varₘ n) module _ (x@((theory , is-cons) , is-comp) : L-complete-theory logic (l1 ⊔ l2)) @@ -468,7 +468,7 @@ module _ type-iff-Prop ( modal-theory-L-complete-theory logic x a) ( (canonical-kripke-model , x) ⊨ a) - pr1 (canonical-model-theorem-pointwise (var n) x) = map-raise + pr1 (canonical-model-theorem-pointwise (varₘ n) x) = map-raise pr1 (canonical-model-theorem-pointwise ⊥ₘ x) = map-raise ∘ is-consistent-modal-theory-L-complete-theory logic x pr1 (canonical-model-theorem-pointwise (a →ₘ b) x) in-logic f = @@ -482,7 +482,7 @@ module _ forward-implication ( canonical-model-theorem-pointwise a y) ( xRy a in-logic) - pr2 (canonical-model-theorem-pointwise (var n) x) = map-inv-raise + pr2 (canonical-model-theorem-pointwise (varₘ n) x) = map-inv-raise pr2 (canonical-model-theorem-pointwise ⊥ₘ x) (map-raise ()) pr2 (canonical-model-theorem-pointwise (a →ₘ b) x) f = L-complete-theory-implication x diff --git a/src/modal-logic/decision-procedure.lagda.md b/src/modal-logic/decision-procedure.lagda.md index 1578b6ef72..3347d59b6a 100644 --- a/src/modal-logic/decision-procedure.lagda.md +++ b/src/modal-logic/decision-procedure.lagda.md @@ -177,7 +177,7 @@ module _ subformulas-list a = cons a (rest a) where rest : modal-formula i → list (modal-formula i) - rest (var x) = nil + rest (varₘ x) = nil rest ⊥ₘ = nil rest (a →ₘ b) = concat-list (subformulas-list a) (subformulas-list b) rest (□ₘ a) = subformulas-list a diff --git a/src/modal-logic/formulas.lagda.md b/src/modal-logic/formulas.lagda.md index 925062c064..1ae8949ec6 100644 --- a/src/modal-logic/formulas.lagda.md +++ b/src/modal-logic/formulas.lagda.md @@ -40,7 +40,7 @@ module _ infixr 25 □ₘ_ data modal-formula (i : Set l) : UU l where - var : type-Set i → modal-formula i + varₘ : type-Set i → modal-formula i ⊥ₘ : modal-formula i _→ₘ_ : modal-formula i → modal-formula i → modal-formula i □ₘ_ : modal-formula i → modal-formula i @@ -96,25 +96,25 @@ module _ where Eq-formula : modal-formula i → modal-formula i → UU l - Eq-formula (var n) (var m) = n = m - Eq-formula (var _) ⊥ₘ = raise-empty l - Eq-formula (var _) (_ →ₘ _) = raise-empty l - Eq-formula (var -) (□ₘ _) = raise-empty l - Eq-formula ⊥ₘ (var _) = raise-empty l + Eq-formula (varₘ n) (varₘ m) = n = m + Eq-formula (varₘ _) ⊥ₘ = raise-empty l + Eq-formula (varₘ _) (_ →ₘ _) = raise-empty l + Eq-formula (varₘ -) (□ₘ _) = raise-empty l + Eq-formula ⊥ₘ (varₘ _) = raise-empty l Eq-formula ⊥ₘ ⊥ₘ = raise-unit l Eq-formula ⊥ₘ (_ →ₘ _) = raise-empty l Eq-formula ⊥ₘ (□ₘ _) = raise-empty l - Eq-formula (_ →ₘ _) (var _) = raise-empty l + Eq-formula (_ →ₘ _) (varₘ _) = raise-empty l Eq-formula (_ →ₘ _) ⊥ₘ = raise-empty l Eq-formula (a →ₘ b) (c →ₘ d) = (Eq-formula a c) × (Eq-formula b d) Eq-formula (_ →ₘ _) (□ₘ _) = raise-empty l - Eq-formula (□ₘ _) (var _) = raise-empty l + Eq-formula (□ₘ _) (varₘ _) = raise-empty l Eq-formula (□ₘ _) ⊥ₘ = raise-empty l Eq-formula (□ₘ _) (_ →ₘ _) = raise-empty l Eq-formula (□ₘ a) (□ₘ c) = Eq-formula a c refl-Eq-formula : (a : modal-formula i) → Eq-formula a a - refl-Eq-formula (var n) = refl + refl-Eq-formula (varₘ n) = refl refl-Eq-formula ⊥ₘ = raise-star refl-Eq-formula (a →ₘ b) = (refl-Eq-formula a) , (refl-Eq-formula b) refl-Eq-formula (□ₘ a) = refl-Eq-formula a @@ -123,39 +123,39 @@ module _ Eq-eq-formula {a} refl = refl-Eq-formula a eq-Eq-formula : {a b : modal-formula i} → Eq-formula a b → a = b - eq-Eq-formula {var _} {var _} refl = refl - eq-Eq-formula {var _} {⊥ₘ} (map-raise ()) - eq-Eq-formula {var _} {_ →ₘ _} (map-raise ()) - eq-Eq-formula {var _} {□ₘ _} (map-raise ()) - eq-Eq-formula {⊥ₘ} {var _} (map-raise ()) + eq-Eq-formula {varₘ _} {varₘ _} refl = refl + eq-Eq-formula {varₘ _} {⊥ₘ} (map-raise ()) + eq-Eq-formula {varₘ _} {_ →ₘ _} (map-raise ()) + eq-Eq-formula {varₘ _} {□ₘ _} (map-raise ()) + eq-Eq-formula {⊥ₘ} {varₘ _} (map-raise ()) eq-Eq-formula {⊥ₘ} {⊥ₘ} _ = refl eq-Eq-formula {⊥ₘ} {_ →ₘ _} (map-raise ()) eq-Eq-formula {⊥ₘ} {□ₘ _} (map-raise ()) - eq-Eq-formula {_ →ₘ _} {var _} (map-raise ()) + eq-Eq-formula {_ →ₘ _} {varₘ _} (map-raise ()) eq-Eq-formula {_ →ₘ _} {⊥ₘ} (map-raise ()) eq-Eq-formula {a →ₘ b} {c →ₘ d} (eq1 , eq2) = ap (λ (x , y) → x →ₘ y) (eq-pair (eq-Eq-formula eq1) (eq-Eq-formula eq2)) eq-Eq-formula {_ →ₘ _} {□ₘ _} (map-raise ()) - eq-Eq-formula {□ₘ _} {var _} (map-raise ()) + eq-Eq-formula {□ₘ _} {varₘ _} (map-raise ()) eq-Eq-formula {□ₘ _} {⊥ₘ} (map-raise ()) eq-Eq-formula {□ₘ _} {_ →ₘ _} (map-raise ()) eq-Eq-formula {□ₘ _} {□ₘ _} eq = ap □ₘ_ (eq-Eq-formula eq) is-prop-Eq-formula : (a b : modal-formula i) → is-prop (Eq-formula a b) - is-prop-Eq-formula (var n) (var m) = is-prop-type-Prop (Id-Prop i n m) - is-prop-Eq-formula (var _) ⊥ₘ = is-prop-raise-empty - is-prop-Eq-formula (var _) (_ →ₘ _) = is-prop-raise-empty - is-prop-Eq-formula (var -) (□ₘ _) = is-prop-raise-empty - is-prop-Eq-formula ⊥ₘ (var _) = is-prop-raise-empty + is-prop-Eq-formula (varₘ n) (varₘ m) = is-prop-type-Prop (Id-Prop i n m) + is-prop-Eq-formula (varₘ _) ⊥ₘ = is-prop-raise-empty + is-prop-Eq-formula (varₘ _) (_ →ₘ _) = is-prop-raise-empty + is-prop-Eq-formula (varₘ -) (□ₘ _) = is-prop-raise-empty + is-prop-Eq-formula ⊥ₘ (varₘ _) = is-prop-raise-empty is-prop-Eq-formula ⊥ₘ ⊥ₘ = is-prop-raise-unit is-prop-Eq-formula ⊥ₘ (_ →ₘ _) = is-prop-raise-empty is-prop-Eq-formula ⊥ₘ (□ₘ _) = is-prop-raise-empty - is-prop-Eq-formula (_ →ₘ _) (var _) = is-prop-raise-empty + is-prop-Eq-formula (_ →ₘ _) (varₘ _) = is-prop-raise-empty is-prop-Eq-formula (_ →ₘ _) ⊥ₘ = is-prop-raise-empty is-prop-Eq-formula (a →ₘ b) (c →ₘ d) = is-prop-product (is-prop-Eq-formula a c) (is-prop-Eq-formula b d) is-prop-Eq-formula (_ →ₘ _) (□ₘ _) = is-prop-raise-empty - is-prop-Eq-formula (□ₘ _) (var _) = is-prop-raise-empty + is-prop-Eq-formula (□ₘ _) (varₘ _) = is-prop-raise-empty is-prop-Eq-formula (□ₘ _) ⊥ₘ = is-prop-raise-empty is-prop-Eq-formula (□ₘ _) (_ →ₘ _) = is-prop-raise-empty is-prop-Eq-formula (□ₘ a) (□ₘ c) = is-prop-Eq-formula a c diff --git a/src/modal-logic/kripke-models-filtrations-theorem.lagda.md b/src/modal-logic/kripke-models-filtrations-theorem.lagda.md index 4fac92f4bd..4cd5178a0b 100644 --- a/src/modal-logic/kripke-models-filtrations-theorem.lagda.md +++ b/src/modal-logic/kripke-models-filtrations-theorem.lagda.md @@ -69,18 +69,19 @@ module _ ( map-equiv-is-kripke-model-filtration i theory M M* is-filtration ( class (Φ-equivalence i theory M) x)) ⊨ a) - pr1 (kripke-models-filtrations-theorem' is-filtration (var n) in-theory x) f = - map-raise - ( forward-implication - ( is-filtration-valuate-is-kripke-model-filtration - ( i) - ( theory) - ( M) - ( M*) - ( is-filtration) - ( n) - ( x)) - ( in-theory , map-inv-raise f)) + pr1 (kripke-models-filtrations-theorem' is-filtration (varₘ n) in-theory x) + f = + map-raise + ( forward-implication + ( is-filtration-valuate-is-kripke-model-filtration + ( i) + ( theory) + ( M) + ( M*) + ( is-filtration) + ( n) + ( x)) + ( in-theory , map-inv-raise f)) pr1 (kripke-models-filtrations-theorem' is-filtration ⊥ₘ in-theory x) = map-raise ∘ map-inv-raise pr1 (kripke-models-filtrations-theorem' is-filtration (a →ₘ b) in-theory x) @@ -148,15 +149,16 @@ module _ ( inv y*-id-class) ( r-xy)) (f))))) - pr2 (kripke-models-filtrations-theorem' is-filtration (var n) in-theory x) f = - map-raise - ( pr2 - ( backward-implication - ( is-filtration-valuate-is-kripke-model-filtration i theory M M* - ( is-filtration) - ( n) - ( x)) - ( map-inv-raise f))) + pr2 (kripke-models-filtrations-theorem' is-filtration (varₘ n) in-theory x) + f = + map-raise + ( pr2 + ( backward-implication + ( is-filtration-valuate-is-kripke-model-filtration i theory M M* + ( is-filtration) + ( n) + ( x)) + ( map-inv-raise f))) pr2 (kripke-models-filtrations-theorem' is-filtration ⊥ₘ in-theory x) = map-raise ∘ map-inv-raise pr2 (kripke-models-filtrations-theorem' is-filtration (a →ₘ b) in-theory x) diff --git a/src/modal-logic/kripke-models-filtrations.lagda.md b/src/modal-logic/kripke-models-filtrations.lagda.md index 5693b0f3fe..cf456f0287 100644 --- a/src/modal-logic/kripke-models-filtrations.lagda.md +++ b/src/modal-logic/kripke-models-filtrations.lagda.md @@ -134,7 +134,7 @@ module _ where is-modal-theory-has-subformulas-Prop : modal-formula i → Prop l2 - is-modal-theory-has-subformulas-Prop (var _) = raise-unit-Prop l2 + is-modal-theory-has-subformulas-Prop (varₘ _) = raise-unit-Prop l2 is-modal-theory-has-subformulas-Prop ⊥ₘ = raise-unit-Prop l2 is-modal-theory-has-subformulas-Prop (a →ₘ b) = product-Prop (theory a) (theory b) @@ -166,7 +166,7 @@ module _ is-in-subtype theory a) → is-modal-theory-closed-under-subformulas is-modal-theory-closed-under-subformulas-condition - h-impl h-box {var n} _ = raise-star + h-impl h-box {varₘ n} _ = raise-star is-modal-theory-closed-under-subformulas-condition h-impl h-box {⊥ₘ} _ = raise-star is-modal-theory-closed-under-subformulas-condition @@ -324,7 +324,7 @@ module _ ( λ n → ( Π-Prop (type-kripke-model i M) ( λ x → iff-Prop - ( product-Prop (theory (var n)) (valuate-kripke-model i M n x)) + ( product-Prop (theory (varₘ n)) (valuate-kripke-model i M n x)) ( valuate-kripke-model i M* n ( map-equiv e (class Φ-equivalence x)))))) @@ -506,7 +506,7 @@ module _ Prop (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) minimal-kripke-model-filtration-valuate n x* = product-Prop - ( theory (var n)) + ( theory (varₘ n)) ( Π-Prop ( type-kripke-model i M) ( λ x → @@ -593,7 +593,7 @@ module _ ( λ y eq-xy → ( map-inv-raise ( forward-implication - ( eq-xy (var n) in-theory) + ( eq-xy (varₘ n) in-theory) ( map-raise val-n))))))) ( λ (in-theory , val-n) → in-theory , val-n x (λ _ _ → id , id))) ( proof-lower-bound) @@ -665,7 +665,7 @@ module _ ( λ y eq-xy → ( map-inv-raise ( forward-implication - ( eq-xy (var n) in-theory) + ( eq-xy (varₘ n) in-theory) ( map-raise val-n))))))) ( λ (in-theory , val-n) → in-theory , val-n x (λ _ _ → id , id))) ( λ x y r → unit-trunc-Prop (base* (proof-lower-bound x y r))) diff --git a/src/modal-logic/kripke-semantics.lagda.md b/src/modal-logic/kripke-semantics.lagda.md index 6778632e19..34acf2eda1 100644 --- a/src/modal-logic/kripke-semantics.lagda.md +++ b/src/modal-logic/kripke-semantics.lagda.md @@ -232,7 +232,7 @@ module _ Σ (kripke-model l1 l2 i l4) (type-kripke-model i) → modal-formula i → Prop (l1 ⊔ l2 ⊔ l4) - (M , x) ⊨ var n = raise-Prop (l1 ⊔ l2) (valuate-kripke-model i M n x) + (M , x) ⊨ varₘ n = raise-Prop (l1 ⊔ l2) (valuate-kripke-model i M n x) (M , x) ⊨ ⊥ₘ = raise-empty-Prop (l1 ⊔ l2 ⊔ l4) (M , x) ⊨ a →ₘ b = hom-Prop ((M , x) ⊨ a) ((M , x) ⊨ b) (M , x) ⊨ □ₘ a = @@ -314,7 +314,7 @@ module _ lemma : (a : modal-formula i) (x : type-kripke-model i M) → is-decidable (type-Prop ((M , x) ⊨ a)) - lemma (var n) x = + lemma (varₘ n) x = is-decidable-raise (l1 ⊔ l2) _ (dec-v x n) lemma ⊥ₘ x = inr map-inv-raise From 29eed3314899471532ac4c8edb0c96ce5765ca12 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Sun, 5 May 2024 17:49:49 +0300 Subject: [PATCH 52/63] Fix filtration lemma and deduction theorem --- .../canonical-model-theorem.lagda.md | 12 ++++--- src/modal-logic/decision-procedure.lagda.md | 4 +-- src/modal-logic/formulas-deduction.lagda.md | 4 +-- ...kripke-models-filtrations-theorem.lagda.md | 32 +++++++++---------- src/modal-logic/l-complete-theories.lagda.md | 10 +++--- src/modal-logic/weak-deduction.lagda.md | 26 +++++++-------- 6 files changed, 45 insertions(+), 43 deletions(-) diff --git a/src/modal-logic/canonical-model-theorem.lagda.md b/src/modal-logic/canonical-model-theorem.lagda.md index 077875b17d..0ff03bd00e 100644 --- a/src/modal-logic/canonical-model-theorem.lagda.md +++ b/src/modal-logic/canonical-model-theorem.lagda.md @@ -175,7 +175,7 @@ module _ ... | inl a-in-logic = is-weak-modal-logic-L-complete-theory logic lzero x (a →ₘ b) ( forward-implication - ( deduction-lemma theory contains-ax-k' contains-ax-s' a b) + ( deduction-theorem theory contains-ax-k' contains-ax-s' a b) ( weak-modal-logic-closure-monotic { ax₁ = theory} { ax₂ = theory-add-formula a theory} @@ -185,7 +185,7 @@ module _ ... | inr not-a-in-logic = is-weak-modal-logic-L-complete-theory logic lzero x (a →ₘ b) ( forward-implication - ( deduction-lemma theory contains-ax-k' contains-ax-s' a b) + ( deduction-theorem theory contains-ax-k' contains-ax-s' a b) ( logic-ex-falso ( theory-add-formula a theory) ( transitive-subset-add-formula a theory (ax-k i) contains-ax-k') @@ -276,7 +276,7 @@ module _ move-assumptions-right f (cons c l) in-logic = move-assumptions-right (c →ₘ f) l ( forward-implication - ( deduction-lemma + ( deduction-theorem ( logic ∪ list-subtype l) ( contains-ax-k-union (list-subtype l)) ( contains-ax-s-union (list-subtype l)) @@ -362,7 +362,7 @@ module _ ( subtype-union-left logic (list-subtype l) (¬¬ₘ a →ₘ a) ( contains-ax-dn (¬¬ₘ a →ₘ a) (a , refl)))) ( forward-implication - ( deduction-lemma + ( deduction-theorem ( logic ∪ list-subtype l) ( contains-ax-k-union (list-subtype l)) ( contains-ax-s-union (list-subtype l)) @@ -544,7 +544,9 @@ module _ ( is-logic (¬¬ₘ a) ( subset-weak-modal-logic-closure-modal-logic-closure (¬¬ₘ a) ( forward-implication - ( deduction-lemma logic contains-ax-k contains-ax-s ( ¬ₘ a) ⊥ₘ) + ( deduction-theorem logic contains-ax-k contains-ax-s + ( ¬ₘ a) + ( ⊥ₘ)) ( weak-modal-logic-closure-monotic ( subtype-union-both logic x (theory-add-formula (¬ₘ a) logic) ( subtype-union-right (Id-formula-Prop (¬ₘ a)) logic) diff --git a/src/modal-logic/decision-procedure.lagda.md b/src/modal-logic/decision-procedure.lagda.md index 3347d59b6a..2ea9d21d0a 100644 --- a/src/modal-logic/decision-procedure.lagda.md +++ b/src/modal-logic/decision-procedure.lagda.md @@ -54,9 +54,9 @@ open import lists.reversing-lists open import modal-logic.completeness open import modal-logic.deduction +open import modal-logic.filtration-lemma open import modal-logic.formulas open import modal-logic.kripke-models-filtrations -open import modal-logic.kripke-models-filtrations-theorem open import modal-logic.kripke-semantics open import modal-logic.soundness open import modal-logic.weak-deduction @@ -497,7 +497,7 @@ module _ complete a ( λ M M-in-class x → ( backward-implication - ( kripke-models-filtrations-theorem' i + ( filtration-lemma i ( subformulas i a) ( is-modal-theory-closed-under-subformulas-subformulas i a) ( M) diff --git a/src/modal-logic/formulas-deduction.lagda.md b/src/modal-logic/formulas-deduction.lagda.md index 3209aff63e..a49a213e4c 100644 --- a/src/modal-logic/formulas-deduction.lagda.md +++ b/src/modal-logic/formulas-deduction.lagda.md @@ -165,9 +165,9 @@ module _ is-in-subtype (weak-modal-logic-closure axioms) (¬ₘ b →ₘ ¬ₘ a) weak-modal-logic-implication-negate {a} {b} ab = forward-implication - ( deduction-lemma axioms contains-ax-k contains-ax-s (¬ₘ b) (¬ₘ a)) + ( deduction-theorem axioms contains-ax-k contains-ax-s (¬ₘ b) (¬ₘ a)) ( forward-implication - ( deduction-lemma + ( deduction-theorem ( theory-add-formula (¬ₘ b) axioms) ( contains-ax-k') ( contains-ax-s') diff --git a/src/modal-logic/kripke-models-filtrations-theorem.lagda.md b/src/modal-logic/kripke-models-filtrations-theorem.lagda.md index 4cd5178a0b..a33f68d2f4 100644 --- a/src/modal-logic/kripke-models-filtrations-theorem.lagda.md +++ b/src/modal-logic/kripke-models-filtrations-theorem.lagda.md @@ -1,7 +1,7 @@ # Kripke models filtrations theorem ```agda -module modal-logic.kripke-models-filtrations-theorem where +module modal-logic.filtration-lemma where ```
Imports @@ -57,7 +57,7 @@ module _ (M : kripke-model l1 l2 i l4) (M* : kripke-model l6 l7 i l8) where - kripke-models-filtrations-theorem' : + filtration-lemma : (is-filtration : is-kripke-model-filtration i theory M M*) (a : modal-formula i) → is-in-subtype theory a → @@ -69,7 +69,7 @@ module _ ( map-equiv-is-kripke-model-filtration i theory M M* is-filtration ( class (Φ-equivalence i theory M) x)) ⊨ a) - pr1 (kripke-models-filtrations-theorem' is-filtration (varₘ n) in-theory x) + pr1 (filtration-lemma is-filtration (varₘ n) in-theory x) f = map-raise ( forward-implication @@ -82,21 +82,21 @@ module _ ( n) ( x)) ( in-theory , map-inv-raise f)) - pr1 (kripke-models-filtrations-theorem' is-filtration ⊥ₘ in-theory x) = + pr1 (filtration-lemma is-filtration ⊥ₘ in-theory x) = map-raise ∘ map-inv-raise - pr1 (kripke-models-filtrations-theorem' is-filtration (a →ₘ b) in-theory x) + pr1 (filtration-lemma is-filtration (a →ₘ b) in-theory x) fab fa = forward-implication - ( kripke-models-filtrations-theorem' is-filtration b + ( filtration-lemma is-filtration b ( pr2 (theory-is-closed in-theory)) ( x)) ( fab ( backward-implication - ( kripke-models-filtrations-theorem' is-filtration a + ( filtration-lemma is-filtration a ( pr1 (theory-is-closed in-theory)) ( x)) ( fa))) - pr1 (kripke-models-filtrations-theorem' is-filtration (□ₘ a) in-theory x) + pr1 (filtration-lemma is-filtration (□ₘ a) in-theory x) f y* r-xy = apply-universal-property-trunc-Prop ( is-inhabited-subtype-equivalence-class @@ -128,7 +128,7 @@ module _ ( λ z* → type-Prop ((M* , z*) ⊨ a)) ( y*-id-class) ( forward-implication - ( kripke-models-filtrations-theorem' is-filtration a + ( filtration-lemma is-filtration a ( theory-is-closed in-theory) ( y)) ( filtration-relation-upper-bound-is-kripke-model-filtration @@ -149,7 +149,7 @@ module _ ( inv y*-id-class) ( r-xy)) (f))))) - pr2 (kripke-models-filtrations-theorem' is-filtration (varₘ n) in-theory x) + pr2 (filtration-lemma is-filtration (varₘ n) in-theory x) f = map-raise ( pr2 @@ -159,24 +159,24 @@ module _ ( n) ( x)) ( map-inv-raise f))) - pr2 (kripke-models-filtrations-theorem' is-filtration ⊥ₘ in-theory x) = + pr2 (filtration-lemma is-filtration ⊥ₘ in-theory x) = map-raise ∘ map-inv-raise - pr2 (kripke-models-filtrations-theorem' is-filtration (a →ₘ b) in-theory x) + pr2 (filtration-lemma is-filtration (a →ₘ b) in-theory x) fab fa = backward-implication - ( kripke-models-filtrations-theorem' is-filtration b + ( filtration-lemma is-filtration b ( pr2 (theory-is-closed in-theory)) ( x)) ( fab ( forward-implication - ( kripke-models-filtrations-theorem' is-filtration a + ( filtration-lemma is-filtration a ( pr1 (theory-is-closed in-theory)) ( x)) ( fa))) - pr2 (kripke-models-filtrations-theorem' is-filtration (□ₘ a) in-theory x) + pr2 (filtration-lemma is-filtration (□ₘ a) in-theory x) f y r-xy = backward-implication - ( kripke-models-filtrations-theorem' is-filtration a + ( filtration-lemma is-filtration a ( theory-is-closed in-theory) ( y)) ( f diff --git a/src/modal-logic/l-complete-theories.lagda.md b/src/modal-logic/l-complete-theories.lagda.md index d58be908bf..674c47be3b 100644 --- a/src/modal-logic/l-complete-theories.lagda.md +++ b/src/modal-logic/l-complete-theories.lagda.md @@ -271,7 +271,7 @@ module _ ( contains-ax-dn' (¬¬ₘ a →ₘ a) (a , refl)) ( is-weak-modal-logic-L-complete-theory lzero t (¬¬ₘ a) ( forward-implication - ( deduction-lemma + ( deduction-theorem ( theory) ( contains-ax-k') ( contains-ax-s') @@ -426,16 +426,16 @@ module _ ( subtype-union-left logic theory) ( contains-ax-s) - L-union-deduction-lemma : + L-union-deduction-theorem : (l : list (modal-formula i)) → (h a : modal-formula i) → is-in-subtype ( weak-modal-logic-closure (logic ∪ list-subtype (cons h l))) a → is-in-subtype ( weak-modal-logic-closure (logic ∪ list-subtype l)) (h →ₘ a) - L-union-deduction-lemma l h a in-logic = + L-union-deduction-theorem l h a in-logic = forward-implication - ( deduction-lemma + ( deduction-theorem ( logic ∪ list-subtype l) ( contains-ax-k' (list-subtype l)) ( contains-ax-s' (list-subtype l)) @@ -486,7 +486,7 @@ module _ ( chain-union-modal-theory) ( leq) ( subset-tail-list-subtype)) - ( L-union-deduction-lemma l h a in-logic)) + ( L-union-deduction-theorem l h a in-logic)) ( exists-chain-element-with-formula-Prop a) ( λ (x , h-in-x) (y , ha-in-y) → ( elim-disjunction diff --git a/src/modal-logic/weak-deduction.lagda.md b/src/modal-logic/weak-deduction.lagda.md index 2db8b615a9..c4b9db49d5 100644 --- a/src/modal-logic/weak-deduction.lagda.md +++ b/src/modal-logic/weak-deduction.lagda.md @@ -270,11 +270,11 @@ module _ (axioms : modal-theory l2 i) where - backward-deduction-lemma : + backward-deduction-theorem : {a b : modal-formula i} → axioms ⊢w a →ₘ b → theory-add-formula a axioms ⊢w b - backward-deduction-lemma {a} wab = + backward-deduction-theorem {a} wab = weak-deduction-mp ( weak-deduction-monotic ( subset-add-formula a axioms) @@ -296,11 +296,11 @@ module _ ( weak-deduction-ax (contains-ax-k _ (a , a →ₘ a , refl)))) ( weak-deduction-ax (contains-ax-k _ (a , a , refl))) - forward-deduction-lemma : + forward-deduction-theorem : (a : modal-formula i) {b : modal-formula i} → theory-add-formula a axioms ⊢w b → is-in-subtype (weak-modal-logic-closure axioms) (a →ₘ b) - forward-deduction-lemma a = + forward-deduction-theorem a = ind-weak-deduction _ ( λ {b} b-in-axioms → ( elim-theory-add-formula a axioms @@ -321,20 +321,20 @@ module _ ( dabc)) ( dab))) - deduction-lemma : + deduction-theorem : (a b : modal-formula i) → type-iff-Prop ( weak-modal-logic-closure (theory-add-formula a axioms) b) ( weak-modal-logic-closure axioms (a →ₘ b)) - pr1 (deduction-lemma a b) = + pr1 (deduction-theorem a b) = map-universal-property-trunc-Prop ( weak-modal-logic-closure axioms (a →ₘ b)) - ( forward-deduction-lemma a) - pr2 (deduction-lemma a b) = + ( forward-deduction-theorem a) + pr2 (deduction-theorem a b) = map-universal-property-trunc-Prop ( weak-modal-logic-closure (theory-add-formula a axioms) b) ( is-in-weak-modal-logic-closure-weak-deduction ∘ - backward-deduction-lemma) + backward-deduction-theorem) ``` ### TODO: List of assumptions @@ -418,9 +418,9 @@ module _ is-in-subtype (weak-modal-logic-closure axioms) (¬ₘ a →ₘ a →ₘ b) deduction-ex-falso a b = forward-implication - ( deduction-lemma axioms contains-ax-k contains-ax-s (¬ₘ a) (a →ₘ b)) + ( deduction-theorem axioms contains-ax-k contains-ax-s (¬ₘ a) (a →ₘ b)) ( forward-implication - ( deduction-lemma + ( deduction-theorem ( theory-add-formula (¬ₘ a) axioms) ( contains-ax-k') ( contains-ax-s') @@ -488,9 +488,9 @@ module _ is-in-subtype (weak-modal-logic-closure axioms) (a →ₘ ¬ₘ a →ₘ b) inv-deduction-ex-falso a b = forward-implication - ( deduction-lemma axioms contains-ax-k contains-ax-s a (¬ₘ a →ₘ b)) + ( deduction-theorem axioms contains-ax-k contains-ax-s a (¬ₘ a →ₘ b)) ( forward-implication - ( deduction-lemma + ( deduction-theorem ( theory-add-formula a axioms) ( contains-ax-k') ( contains-ax-s') From b2d85c4f32e725933502dcc68bcbb2eaa6e7882f Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Sun, 5 May 2024 18:01:31 +0300 Subject: [PATCH 53/63] Fix filtratiom lemma file name --- src/modal-logic.lagda.md | 2 +- ...s-filtrations-theorem.lagda.md => filtration-lemma.lagda.md} | 0 2 files changed, 1 insertion(+), 1 deletion(-) rename src/modal-logic/{kripke-models-filtrations-theorem.lagda.md => filtration-lemma.lagda.md} (100%) diff --git a/src/modal-logic.lagda.md b/src/modal-logic.lagda.md index 2b0ad79bbc..d4f78a1f4e 100644 --- a/src/modal-logic.lagda.md +++ b/src/modal-logic.lagda.md @@ -11,11 +11,11 @@ open import modal-logic.completeness-k public open import modal-logic.completeness-s5 public open import modal-logic.decision-procedure public open import modal-logic.deduction public +open import modal-logic.filtration-lemma public open import modal-logic.finite-approximability public open import modal-logic.formulas public open import modal-logic.formulas-deduction public open import modal-logic.kripke-models-filtrations public -open import modal-logic.kripke-models-filtrations-theorem public open import modal-logic.kripke-semantics public open import modal-logic.l-complete-theories public open import modal-logic.l-consistent-theories public diff --git a/src/modal-logic/kripke-models-filtrations-theorem.lagda.md b/src/modal-logic/filtration-lemma.lagda.md similarity index 100% rename from src/modal-logic/kripke-models-filtrations-theorem.lagda.md rename to src/modal-logic/filtration-lemma.lagda.md From 294ae8272bbbee886042dc5bc73c93f9ada68c29 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Mon, 6 May 2024 01:28:46 +0300 Subject: [PATCH 54/63] Add recurser for equivalence classes and refactor --- src/foundation/equivalence-classes.lagda.md | 88 +- .../canonical-model-theorem.lagda.md | 870 +++++++++--------- src/modal-logic/canonical-theories.lagda.md | 31 +- src/modal-logic/completeness-k.lagda.md | 2 +- src/modal-logic/completeness-s5.lagda.md | 10 +- src/modal-logic/deduction.lagda.md | 50 +- .../kripke-models-filtrations.lagda.md | 9 +- src/modal-logic/modal-logic-k.lagda.md | 2 +- src/modal-logic/modal-logic-s5.lagda.md | 2 +- src/modal-logic/soundness.lagda.md | 11 +- src/modal-logic/weak-deduction.lagda.md | 76 +- 11 files changed, 635 insertions(+), 516 deletions(-) diff --git a/src/foundation/equivalence-classes.lagda.md b/src/foundation/equivalence-classes.lagda.md index 71954bae57..329ad16806 100644 --- a/src/foundation/equivalence-classes.lagda.md +++ b/src/foundation/equivalence-classes.lagda.md @@ -10,6 +10,7 @@ module foundation.equivalence-classes where open import foundation.conjunction open import foundation.dependent-pair-types open import foundation.effective-maps-equivalence-relations +open import foundation.equality-dependent-pair-types open import foundation.existential-quantification open import foundation.functoriality-propositional-truncation open import foundation.fundamental-theorem-of-identity-types @@ -18,6 +19,7 @@ open import foundation.locally-small-types open import foundation.logical-equivalences open import foundation.propositional-truncations open import foundation.reflecting-maps-equivalence-relations +open import foundation.sets open import foundation.slice open import foundation.small-types open import foundation.subtype-identity-principle @@ -27,14 +29,16 @@ open import foundation.universal-property-image open import foundation.universe-levels open import foundation-core.cartesian-product-types +open import foundation-core.dependent-identifications open import foundation-core.embeddings open import foundation-core.equivalence-relations open import foundation-core.equivalences +open import foundation-core.function-types open import foundation-core.functoriality-dependent-pair-types open import foundation-core.identity-types open import foundation-core.propositions -open import foundation-core.sets open import foundation-core.torsorial-type-families +open import foundation-core.transport-along-identifications ```
@@ -383,6 +387,88 @@ module _ map-inv-equiv (is-effective-class x y) ``` +### TODO: title + +```agda + eq-class-in-common-class : + (c : equivalence-class R) {a a' : A} → + is-in-equivalence-class R c a → + is-in-equivalence-class R c a' → + class R a = class R a' + eq-class-in-common-class c {a} {a'} a-in-c a'-in-c = + equational-reasoning + class R a + = c by eq-effective-quotient' a c a-in-c + = class R a' by inv (eq-effective-quotient' a' c a'-in-c) + + sim-equivalence-relation-in-same-class : + (c : equivalence-class R) {a a' : A} → + is-in-equivalence-class R c a → + is-in-equivalence-class R c a' → + sim-equivalence-relation R a a' + sim-equivalence-relation-in-same-class c {a} {a'} a-in-c a'-in-c = + apply-effectiveness-class (eq-class-in-common-class c a-in-c a'-in-c) +``` + +### TODO: Eliminator + +```agda + rec-equivalence-class : + {l3 : Level} (B : Set l3) → + (f : A → type-Set B) → + ((a a' : A) → sim-equivalence-relation R a a' → f a = f a') → + equivalence-class R → type-Set B + rec-equivalence-class {l3} B f H c = + pr1 + ( apply-universal-property-trunc-Prop + ( is-inhabited-subtype-equivalence-class R c) + ( b , is-prop-b) + ( λ (a , a-in-c) → + ( pair (f a) + ( λ a' a'-in-c → + ( H a a' + ( sim-equivalence-relation-in-same-class c a-in-c a'-in-c)))))) + where + b : UU (l1 ⊔ l2 ⊔ l3) + b = + Σ (type-Set B) (λ b → (a : A) → is-in-equivalence-class R c a → b = f a) + + is-prop-b : is-prop b + is-prop-b = + is-prop-all-elements-equal + ( λ (b , h) (b' , h') → + ( eq-pair-Σ + ( apply-universal-property-trunc-Prop + ( is-inhabited-subtype-equivalence-class R c) + ( b = b' , is-set-type-Set B b b') + ( λ (a , a-in-c) → h a a-in-c ∙ inv (h' a a-in-c))) + ( eq-is-prop + ( is-prop-Π + ( λ x → is-prop-function-type (is-set-type-Set B b' (f x))))))) + + rec-equivalence-class-Prop : + {l3 : Level} (B : Prop l3) → + (f : A → type-Prop B) → + equivalence-class R → type-Prop B + rec-equivalence-class-Prop B f = + rec-equivalence-class (set-Prop B) f + ( λ a a' _ → eq-is-prop (is-prop-type-Prop B)) + + ind-equivalence-class-Prop : + {l3 : Level} (B : equivalence-class R → Prop l3) → + (f : (a : A) → type-Prop (B (class R a))) → + (c : equivalence-class R) → type-Prop (B c) + ind-equivalence-class-Prop B f c = + apply-universal-property-trunc-Prop + ( is-inhabited-subtype-equivalence-class R c) + ( B c) + ( λ (a , a-in-c) → + ( tr + ( type-Prop ∘ B) + ( eq-effective-quotient' a c a-in-c) + ( rec-equivalence-class-Prop (Π-Prop A (B ∘ class R)) (λ _ → f) c a))) +``` + ### The map `class` into the type of equivalence classes is surjective and effective ```agda diff --git a/src/modal-logic/canonical-model-theorem.lagda.md b/src/modal-logic/canonical-model-theorem.lagda.md index 0ff03bd00e..b76752b205 100644 --- a/src/modal-logic/canonical-model-theorem.lagda.md +++ b/src/modal-logic/canonical-model-theorem.lagda.md @@ -60,7 +60,6 @@ module _ (is-logic : is-modal-logic logic) (is-cons : is-consistent-modal-logic logic) (is-normal : is-normal-modal-logic logic) - (lem : LEM (l1 ⊔ l2)) (zorn : Zorn-non-empty (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) l2) (prop-resize : propositional-resizing (l1 ⊔ l2) (lsuc (l1 ⊔ l2))) where @@ -117,454 +116,471 @@ module _ modal-theory-L-complete-theory logic x (varₘ n) module _ - (x@((theory , is-cons) , is-comp) : L-complete-theory logic (l1 ⊔ l2)) + (lem : LEM (l1 ⊔ l2)) where - private - contains-ax-k' : ax-k i ⊆ theory - contains-ax-k' = - transitive-leq-subtype (ax-k i) logic theory - ( subset-logic-L-complete-theory logic lzero x) - ( contains-ax-k) - - contains-ax-s' : ax-s i ⊆ theory - contains-ax-s' = - transitive-leq-subtype (ax-s i) logic theory - ( subset-logic-L-complete-theory logic lzero x) - ( contains-ax-s) - - contains-ax-dn' : ax-dn i ⊆ theory - contains-ax-dn' = - transitive-leq-subtype (ax-dn i) logic theory - ( subset-logic-L-complete-theory logic lzero x) - ( contains-ax-dn) - - contains-ax-n' : ax-n i ⊆ theory - contains-ax-n' = - transitive-leq-subtype (ax-n i) logic theory - ( subset-logic-L-complete-theory logic lzero x) - ( contains-ax-n) - - contains-ax-k-union : - {l : Level} (t : modal-theory l i) → ax-k i ⊆ logic ∪ t - contains-ax-k-union t = - transitive-leq-subtype (ax-k i) logic (logic ∪ t) - ( subtype-union-left logic t) - ( contains-ax-k) - - contains-ax-s-union : - {l : Level} (t : modal-theory l i) → ax-s i ⊆ logic ∪ t - contains-ax-s-union t = - transitive-leq-subtype (ax-s i) logic (logic ∪ t) - ( subtype-union-left logic t) - ( contains-ax-s) - - is-disjuctive-theory : is-disjuctive-modal-theory theory - is-disjuctive-theory = - is-disjuctive-L-complete-theory logic x - ( contains-ax-k) - ( contains-ax-s) - ( contains-ax-dn) - ( lem) - - L-complete-theory-implication : - {a b : modal-formula i} → - (is-in-subtype theory a → is-in-subtype theory b) → - is-in-subtype theory (a →ₘ b) - L-complete-theory-implication {a} {b} f with is-disjuctive-theory a - ... | inl a-in-logic = - is-weak-modal-logic-L-complete-theory logic lzero x (a →ₘ b) - ( forward-implication - ( deduction-theorem theory contains-ax-k' contains-ax-s' a b) - ( weak-modal-logic-closure-monotic - { ax₁ = theory} - { ax₂ = theory-add-formula a theory} - ( subset-add-formula a theory) - ( b) - ( weak-modal-logic-closure-ax (f a-in-logic)))) - ... | inr not-a-in-logic = - is-weak-modal-logic-L-complete-theory logic lzero x (a →ₘ b) - ( forward-implication - ( deduction-theorem theory contains-ax-k' contains-ax-s' a b) - ( logic-ex-falso - ( theory-add-formula a theory) - ( transitive-subset-add-formula a theory (ax-k i) contains-ax-k') - ( transitive-subset-add-formula a theory (ax-s i) contains-ax-s') - ( transitive-subset-add-formula a theory (ax-dn i) contains-ax-dn') - ( a) - ( b) - ( weak-modal-logic-closure-ax (formula-in-add-formula a theory)) + module _ + (x@((theory , is-cons) , is-comp) : L-complete-theory logic (l1 ⊔ l2)) + where + + private + contains-ax-k' : ax-k i ⊆ theory + contains-ax-k' = + transitive-leq-subtype (ax-k i) logic theory + ( subset-logic-L-complete-theory logic lzero x) + ( contains-ax-k) + + contains-ax-s' : ax-s i ⊆ theory + contains-ax-s' = + transitive-leq-subtype (ax-s i) logic theory + ( subset-logic-L-complete-theory logic lzero x) + ( contains-ax-s) + + contains-ax-dn' : ax-dn i ⊆ theory + contains-ax-dn' = + transitive-leq-subtype (ax-dn i) logic theory + ( subset-logic-L-complete-theory logic lzero x) + ( contains-ax-dn) + + contains-ax-n' : ax-n i ⊆ theory + contains-ax-n' = + transitive-leq-subtype (ax-n i) logic theory + ( subset-logic-L-complete-theory logic lzero x) + ( contains-ax-n) + + contains-ax-k-union : + {l : Level} (t : modal-theory l i) → ax-k i ⊆ logic ∪ t + contains-ax-k-union t = + transitive-leq-subtype (ax-k i) logic (logic ∪ t) + ( subtype-union-left logic t) + ( contains-ax-k) + + contains-ax-s-union : + {l : Level} (t : modal-theory l i) → ax-s i ⊆ logic ∪ t + contains-ax-s-union t = + transitive-leq-subtype (ax-s i) logic (logic ∪ t) + ( subtype-union-left logic t) + ( contains-ax-s) + + is-disjuctive-theory : is-disjuctive-modal-theory theory + is-disjuctive-theory = + is-disjuctive-L-complete-theory logic x + ( contains-ax-k) + ( contains-ax-s) + ( contains-ax-dn) + ( lem) + + L-complete-theory-implication : + {a b : modal-formula i} → + (is-in-subtype theory a → is-in-subtype theory b) → + is-in-subtype theory (a →ₘ b) + L-complete-theory-implication {a} {b} f with is-disjuctive-theory a + ... | inl a-in-logic = + is-weak-modal-logic-L-complete-theory logic lzero x (a →ₘ b) + ( forward-implication + ( deduction-theorem theory contains-ax-k' contains-ax-s' a b) ( weak-modal-logic-closure-monotic { ax₁ = theory} { ax₂ = theory-add-formula a theory} ( subset-add-formula a theory) - ( ¬ₘ a) - ( weak-modal-logic-closure-ax not-a-in-logic)))) - - L-complete-theory-box : - {a : modal-formula i} → - ( (y : L-complete-theory logic (l1 ⊔ l2)) → - relation-kripke-model i canonical-kripke-model x y → - is-in-subtype (modal-theory-L-complete-theory logic y) a) → - is-in-subtype theory (□ₘ a) - L-complete-theory-box {a} f with is-disjuctive-theory (□ₘ a) - ... | inl box-a-in-logic = box-a-in-logic - ... | inr not-box-a-in-logic = - ex-falso - ( apply-universal-property-trunc-Prop - ( lindenbaum logic contains-ax-k contains-ax-s zorn prop-resize - ( y , is-L-consistent-y)) - ( empty-Prop) - ( λ (w , y-leq-w) → - ( is-consistent-modal-theory-L-complete-theory logic w - ( weak-modal-logic-mp - ( is-weak-modal-logic-L-complete-theory logic lzero w) - ( y-leq-w (¬ₘ a) - ( formula-in-add-formula (¬ₘ a) (unbox-modal-theory theory))) - ( f w (λ b → y-leq-w b ∘ y-contains-unbox)))))) - where - y : modal-theory (l1 ⊔ l2) i - y = theory-add-formula (¬ₘ a) (unbox-modal-theory theory) - - y-contains-unbox : - {b : modal-formula i} → - is-in-subtype theory (□ₘ b) → - is-in-subtype y b - y-contains-unbox {b} = - subset-add-formula (¬ₘ a) (unbox-modal-theory theory) b - - list-to-implications : - modal-formula i → (l : list (modal-formula i)) → modal-formula i - list-to-implications f nil = f - list-to-implications f (cons g l) = list-to-implications (g →ₘ f) l - - list-to-implications-rev : - modal-formula i → (l : list (modal-formula i)) → modal-formula i - list-to-implications-rev f nil = f - list-to-implications-rev f (cons g l) = g →ₘ list-to-implications-rev f l - - list-to-implication-rev-snoc : - (f g : modal-formula i) (l : list (modal-formula i)) → - list-to-implications f (snoc l g) = g →ₘ list-to-implications f l - list-to-implication-rev-snoc f g nil = refl - list-to-implication-rev-snoc f g (cons h l) = - list-to-implication-rev-snoc (h →ₘ f) g l - - eq-reverse-list-to-implications : - (f : modal-formula i) (l : list (modal-formula i)) → - list-to-implications f (reverse-list l) = list-to-implications-rev f l - eq-reverse-list-to-implications f nil = refl - eq-reverse-list-to-implications f (cons g l) = - ( list-to-implication-rev-snoc f g (reverse-list l)) ∙ - ( ap (λ x → g →ₘ x) (eq-reverse-list-to-implications f l)) - - move-assumptions-right : - (f : modal-formula i) (l : list (modal-formula i)) → - is-in-subtype (weak-modal-logic-closure (logic ∪ list-subtype l)) f → - is-in-subtype - ( weak-modal-logic-closure logic) - ( list-to-implications f l) - move-assumptions-right f nil = - weak-modal-logic-closure-monotic - ( subtype-union-both - ( logic) - ( list-subtype nil) - ( logic) - ( refl-leq-subtype logic) - ( subset-list-subtype-nil logic)) - ( f) - move-assumptions-right f (cons c l) in-logic = - move-assumptions-right (c →ₘ f) l + ( b) + ( weak-modal-logic-closure-ax (f a-in-logic)))) + ... | inr not-a-in-logic = + is-weak-modal-logic-L-complete-theory logic lzero x (a →ₘ b) ( forward-implication - ( deduction-theorem - ( logic ∪ list-subtype l) - ( contains-ax-k-union (list-subtype l)) - ( contains-ax-s-union (list-subtype l)) - ( c) - ( f)) - ( weak-modal-logic-closure-monotic - ( transitive-leq-subtype - ( logic ∪ list-subtype (cons c l)) - ( logic ∪ theory-add-formula c (list-subtype l)) - ( theory-add-formula c (logic ∪ list-subtype l)) - ( theory-add-formula-union-right c logic (list-subtype l)) - ( subset-union-subset-right logic - ( list-subtype (cons c l)) - ( theory-add-formula c (list-subtype l)) - ( subset-list-subtype-cons - ( theory-add-formula c (list-subtype l)) - ( formula-in-add-formula c (list-subtype l)) - ( subset-add-formula c (list-subtype l))))) - ( f) - ( in-logic))) - - α : - (l : list (modal-formula i)) → - list-subtype l ⊆ unbox-modal-theory theory → - is-in-subtype theory (□ₘ (list-to-implications-rev a l)) → - is-in-subtype theory (□ₘ a) - α nil sub in-logic = in-logic - α (cons c l) sub in-logic = - α l - ( transitive-leq-subtype - ( list-subtype l) - ( list-subtype (cons c l)) - ( unbox-modal-theory theory) - ( sub) - ( subset-tail-list-subtype)) - ( weak-modal-logic-mp - ( is-weak-modal-logic-L-complete-theory logic lzero x) - { a = □ₘ c} - { b = □ₘ list-to-implications-rev a l} - ( weak-modal-logic-mp - ( is-weak-modal-logic-L-complete-theory logic lzero x) - { a = □ₘ (c →ₘ list-to-implications-rev a l)} - { b = □ₘ c →ₘ □ₘ (list-to-implications-rev a l)} - ( contains-ax-n' _ (c , list-to-implications-rev a l , refl)) - ( in-logic)) - ( sub c head-in-list-subtype)) - - β : - (l : list (modal-formula i)) → - list-subtype l ⊆ unbox-modal-theory theory → - is-in-subtype (weak-modal-logic-closure (logic ∪ list-subtype l)) a → - is-in-subtype theory (□ₘ a) - β l sub in-logic = - α l sub - ( subset-logic-L-complete-theory logic lzero x - ( □ₘ list-to-implications-rev a l) - ( modal-logic-nec is-logic - ( tr - ( is-in-subtype logic) - ( eq-reverse-list-to-implications a l) - ( is-weak-logic - ( list-to-implications a (reverse-list l)) - ( move-assumptions-right a (reverse-list l) - ( weak-modal-logic-closure-monotic - ( subset-union-subset-right logic - ( list-subtype l) - ( list-subtype (reverse-list l)) - ( subset-list-subtype-reverse-list l)) - ( a) - ( in-logic))))))) - - γ : - (l : list (modal-formula i)) → - list-subtype l ⊆ unbox-modal-theory theory → - is-contradictory-modal-logic - ( weak-modal-logic-closure - ( theory-add-formula (¬ₘ a) (logic ∪ list-subtype l))) → + ( deduction-theorem theory contains-ax-k' contains-ax-s' a b) + ( logic-ex-falso + ( theory-add-formula a theory) + ( transitive-subset-add-formula a theory (ax-k i) contains-ax-k') + ( transitive-subset-add-formula a theory (ax-s i) contains-ax-s') + ( transitive-subset-add-formula a theory + ( ax-dn i) + ( contains-ax-dn')) + ( a) + ( b) + ( weak-modal-logic-closure-ax (formula-in-add-formula a theory)) + ( weak-modal-logic-closure-monotic + { ax₁ = theory} + { ax₂ = theory-add-formula a theory} + ( subset-add-formula a theory) + ( ¬ₘ a) + ( weak-modal-logic-closure-ax not-a-in-logic)))) + + L-complete-theory-box : + {a : modal-formula i} → + ( (y : L-complete-theory logic (l1 ⊔ l2)) → + relation-kripke-model i canonical-kripke-model x y → + is-in-subtype (modal-theory-L-complete-theory logic y) a) → is-in-subtype theory (□ₘ a) - γ l sub is-cont = - β l sub - ( weak-modal-logic-closure-mp {a = ¬¬ₘ a} {b = a} - ( weak-modal-logic-closure-ax - ( subtype-union-left logic (list-subtype l) (¬¬ₘ a →ₘ a) - ( contains-ax-dn (¬¬ₘ a →ₘ a) (a , refl)))) + L-complete-theory-box {a} f with is-disjuctive-theory (□ₘ a) + ... | inl box-a-in-logic = box-a-in-logic + ... | inr not-box-a-in-logic = + ex-falso + ( apply-universal-property-trunc-Prop + ( lindenbaum logic contains-ax-k contains-ax-s zorn prop-resize + ( y , is-L-consistent-y)) + ( empty-Prop) + ( λ (w , y-leq-w) → + ( is-consistent-modal-theory-L-complete-theory logic w + ( weak-modal-logic-mp + ( is-weak-modal-logic-L-complete-theory logic lzero w) + ( y-leq-w (¬ₘ a) + ( formula-in-add-formula + ( ¬ₘ a) + ( unbox-modal-theory theory))) + ( f w (λ b → y-leq-w b ∘ y-contains-unbox)))))) + where + y : modal-theory (l1 ⊔ l2) i + y = theory-add-formula (¬ₘ a) (unbox-modal-theory theory) + + y-contains-unbox : + {b : modal-formula i} → + is-in-subtype theory (□ₘ b) → + is-in-subtype y b + y-contains-unbox {b} = + subset-add-formula (¬ₘ a) (unbox-modal-theory theory) b + + list-to-implications : + modal-formula i → (l : list (modal-formula i)) → modal-formula i + list-to-implications f nil = f + list-to-implications f (cons g l) = list-to-implications (g →ₘ f) l + + list-to-implications-rev : + modal-formula i → (l : list (modal-formula i)) → modal-formula i + list-to-implications-rev f nil = f + list-to-implications-rev f (cons g l) = + g →ₘ list-to-implications-rev f l + + list-to-implication-rev-snoc : + (f g : modal-formula i) (l : list (modal-formula i)) → + list-to-implications f (snoc l g) = g →ₘ list-to-implications f l + list-to-implication-rev-snoc f g nil = refl + list-to-implication-rev-snoc f g (cons h l) = + list-to-implication-rev-snoc (h →ₘ f) g l + + eq-reverse-list-to-implications : + (f : modal-formula i) (l : list (modal-formula i)) → + list-to-implications f (reverse-list l) = list-to-implications-rev f l + eq-reverse-list-to-implications f nil = refl + eq-reverse-list-to-implications f (cons g l) = + ( list-to-implication-rev-snoc f g (reverse-list l)) ∙ + ( ap (λ x → g →ₘ x) (eq-reverse-list-to-implications f l)) + + move-assumptions-right : + (f : modal-formula i) (l : list (modal-formula i)) → + is-in-subtype (weak-modal-logic-closure (logic ∪ list-subtype l)) f → + is-in-subtype + ( weak-modal-logic-closure logic) + ( list-to-implications f l) + move-assumptions-right f nil = + weak-modal-logic-closure-monotic + ( subtype-union-both + ( logic) + ( list-subtype nil) + ( logic) + ( refl-leq-subtype logic) + ( subset-list-subtype-nil logic)) + ( f) + move-assumptions-right f (cons c l) in-logic = + move-assumptions-right (c →ₘ f) l ( forward-implication ( deduction-theorem ( logic ∪ list-subtype l) ( contains-ax-k-union (list-subtype l)) ( contains-ax-s-union (list-subtype l)) - ( ¬ₘ a) - ( ⊥ₘ)) - ( is-cont))) - - δ : - (l : list (modal-formula i)) → - list-subtype l ⊆ unbox-modal-theory theory → - is-contradictory-modal-logic - ( weak-modal-logic-closure - ( logic ∪ (theory-add-formula (¬ₘ a) (list-subtype l)))) → - is-in-subtype theory (□ₘ a) - δ l sub is-cont = - γ l sub - ( is-contradictory-modal-logic-monotic + ( c) + ( f)) + ( weak-modal-logic-closure-monotic + ( transitive-leq-subtype + ( logic ∪ list-subtype (cons c l)) + ( logic ∪ theory-add-formula c (list-subtype l)) + ( theory-add-formula c (logic ∪ list-subtype l)) + ( theory-add-formula-union-right c logic (list-subtype l)) + ( subset-union-subset-right logic + ( list-subtype (cons c l)) + ( theory-add-formula c (list-subtype l)) + ( subset-list-subtype-cons + ( theory-add-formula c (list-subtype l)) + ( formula-in-add-formula c (list-subtype l)) + ( subset-add-formula c (list-subtype l))))) + ( f) + ( in-logic))) + + α : + (l : list (modal-formula i)) → + list-subtype l ⊆ unbox-modal-theory theory → + is-in-subtype theory (□ₘ (list-to-implications-rev a l)) → + is-in-subtype theory (□ₘ a) + α nil sub in-logic = in-logic + α (cons c l) sub in-logic = + α l + ( transitive-leq-subtype + ( list-subtype l) + ( list-subtype (cons c l)) + ( unbox-modal-theory theory) + ( sub) + ( subset-tail-list-subtype)) + ( weak-modal-logic-mp + ( is-weak-modal-logic-L-complete-theory logic lzero x) + { a = □ₘ c} + { b = □ₘ list-to-implications-rev a l} + ( weak-modal-logic-mp + ( is-weak-modal-logic-L-complete-theory logic lzero x) + { a = □ₘ (c →ₘ list-to-implications-rev a l)} + { b = □ₘ c →ₘ □ₘ (list-to-implications-rev a l)} + ( contains-ax-n' _ (c , list-to-implications-rev a l , refl)) + ( in-logic)) + ( sub c head-in-list-subtype)) + + β : + (l : list (modal-formula i)) → + list-subtype l ⊆ unbox-modal-theory theory → + is-in-subtype (weak-modal-logic-closure (logic ∪ list-subtype l)) a → + is-in-subtype theory (□ₘ a) + β l sub in-logic = + α l sub + ( subset-logic-L-complete-theory logic lzero x + ( □ₘ list-to-implications-rev a l) + ( modal-logic-nec is-logic + ( tr + ( is-in-subtype logic) + ( eq-reverse-list-to-implications a l) + ( is-weak-logic + ( list-to-implications a (reverse-list l)) + ( move-assumptions-right a (reverse-list l) + ( weak-modal-logic-closure-monotic + ( subset-union-subset-right logic + ( list-subtype l) + ( list-subtype (reverse-list l)) + ( subset-list-subtype-reverse-list l)) + ( a) + ( in-logic))))))) + + γ : + (l : list (modal-formula i)) → + list-subtype l ⊆ unbox-modal-theory theory → + is-contradictory-modal-logic ( weak-modal-logic-closure - ( logic ∪ theory-add-formula (¬ₘ a) (list-subtype l))) + ( theory-add-formula (¬ₘ a) (logic ∪ list-subtype l))) → + is-in-subtype theory (□ₘ a) + γ l sub is-cont = + β l sub + ( weak-modal-logic-closure-mp {a = ¬¬ₘ a} {b = a} + ( weak-modal-logic-closure-ax + ( subtype-union-left logic (list-subtype l) (¬¬ₘ a →ₘ a) + ( contains-ax-dn (¬¬ₘ a →ₘ a) (a , refl)))) + ( forward-implication + ( deduction-theorem + ( logic ∪ list-subtype l) + ( contains-ax-k-union (list-subtype l)) + ( contains-ax-s-union (list-subtype l)) + ( ¬ₘ a) + ( ⊥ₘ)) + ( is-cont))) + + δ : + (l : list (modal-formula i)) → + list-subtype l ⊆ unbox-modal-theory theory → + is-contradictory-modal-logic ( weak-modal-logic-closure - ( theory-add-formula (¬ₘ a) (logic ∪ list-subtype l))) - ( weak-modal-logic-closure-monotic - ( theory-add-formula-union-right (¬ₘ a) logic (list-subtype l))) - ( is-cont)) - - ε : - (l : list (modal-formula i)) → - list-subtype l ⊆ logic ∪ y → - is-contradictory-modal-logic - ( weak-modal-logic-closure (list-subtype l)) → - is-in-subtype theory (□ₘ a) - ε l sub is-cont = - apply-universal-property-trunc-Prop - ( lists-in-union-lists l logic y sub) - ( theory (□ₘ a)) - ( λ ((l-ax , l-y) , l-sub-union , l-ax-sub-logic , l-y-sub-y) → - ( apply-universal-property-trunc-Prop - ( lists-in-union-lists l-y - ( Id-formula-Prop (¬ₘ a)) - ( unbox-modal-theory theory) - ( l-y-sub-y)) - ( theory (□ₘ a)) - ( λ ((l-not-a , l-box) , l-sub-union' , l-not-a-sub , l-box-sub) → - ( δ - ( l-box) - ( l-box-sub) - ( is-contradictory-modal-logic-monotic - ( weak-modal-logic-closure (list-subtype l)) - ( weak-modal-logic-closure - ( logic ∪ theory-add-formula (¬ₘ a) (list-subtype l-box))) - ( weak-modal-logic-closure-monotic - ( transitive-leq-subtype - ( list-subtype l) - ( list-subtype l-ax ∪ list-subtype l-y) + ( logic ∪ (theory-add-formula (¬ₘ a) (list-subtype l)))) → + is-in-subtype theory (□ₘ a) + δ l sub is-cont = + γ l sub + ( is-contradictory-modal-logic-monotic + ( weak-modal-logic-closure + ( logic ∪ theory-add-formula (¬ₘ a) (list-subtype l))) + ( weak-modal-logic-closure + ( theory-add-formula (¬ₘ a) (logic ∪ list-subtype l))) + ( weak-modal-logic-closure-monotic + ( theory-add-formula-union-right (¬ₘ a) logic (list-subtype l))) + ( is-cont)) + + ε : + (l : list (modal-formula i)) → + list-subtype l ⊆ logic ∪ y → + is-contradictory-modal-logic + ( weak-modal-logic-closure (list-subtype l)) → + is-in-subtype theory (□ₘ a) + ε l sub is-cont = + apply-universal-property-trunc-Prop + ( lists-in-union-lists l logic y sub) + ( theory (□ₘ a)) + ( λ ((l-ax , l-y) , sub-union , l-ax-sub-logic , l-y-sub-y) → + ( apply-universal-property-trunc-Prop + ( lists-in-union-lists l-y + ( Id-formula-Prop (¬ₘ a)) + ( unbox-modal-theory theory) + ( l-y-sub-y)) + ( theory (□ₘ a)) + ( λ ((l-not-a , l-box) , sub-union' , l-not-a-sub , l-box-sub) → + ( δ + ( l-box) + ( l-box-sub) + ( is-contradictory-modal-logic-monotic + ( weak-modal-logic-closure (list-subtype l)) + ( weak-modal-logic-closure ( logic ∪ - theory-add-formula (¬ₘ a) (list-subtype l-box)) - ( subset-union-subsets - ( list-subtype l-ax) - ( list-subtype l-y) - ( logic) - ( theory-add-formula (¬ₘ a) (list-subtype l-box)) - ( l-ax-sub-logic) - ( transitive-leq-subtype + theory-add-formula (¬ₘ a) (list-subtype l-box))) + ( weak-modal-logic-closure-monotic + ( transitive-leq-subtype + ( list-subtype l) + ( list-subtype l-ax ∪ list-subtype l-y) + ( logic ∪ + theory-add-formula (¬ₘ a) (list-subtype l-box)) + ( subset-union-subsets + ( list-subtype l-ax) ( list-subtype l-y) - ( list-subtype l-not-a ∪ list-subtype l-box) + ( logic) ( theory-add-formula (¬ₘ a) (list-subtype l-box)) - ( subtype-union-both - ( list-subtype l-not-a) - ( list-subtype l-box) + ( l-ax-sub-logic) + ( transitive-leq-subtype + ( list-subtype l-y) + ( list-subtype l-not-a ∪ list-subtype l-box) ( theory-add-formula (¬ₘ a) (list-subtype l-box)) - ( transitive-leq-subtype + ( subtype-union-both ( list-subtype l-not-a) - ( Id-formula-Prop (¬ₘ a)) - ( theory-add-formula (¬ₘ a) + ( list-subtype l-box) + ( theory-add-formula + ( ¬ₘ a) ( list-subtype l-box)) - ( subtype-union-left + ( transitive-leq-subtype + ( list-subtype l-not-a) ( Id-formula-Prop (¬ₘ a)) - ( list-subtype l-box)) - ( l-not-a-sub)) - ( subset-add-formula (¬ₘ a) (list-subtype l-box))) - ( l-sub-union'))) - ( l-sub-union))) - ( is-cont)))))) - - is-L-consistent-y : is-L-consistent-theory logic y - is-L-consistent-y = - map-universal-property-trunc-Prop - ( empty-Prop) - ( λ d-bot → - ( is-consistent-modal-theory-L-complete-theory logic x - ( weak-modal-logic-mp - ( is-weak-modal-logic-L-complete-theory logic lzero x) - ( not-box-a-in-logic) - ( ε - ( list-assumptions-weak-deduction d-bot) - ( subset-theory-list-assumptions-weak-deduction d-bot) - ( is-in-weak-modal-logic-closure-weak-deduction - ( is-assumptions-list-assumptions-weak-deduction d-bot)))))) - - canonical-model-theorem-pointwise : - (a : modal-formula i) - (x : L-complete-theory logic (l1 ⊔ l2)) → - type-iff-Prop - ( modal-theory-L-complete-theory logic x a) - ( (canonical-kripke-model , x) ⊨ a) - pr1 (canonical-model-theorem-pointwise (varₘ n) x) = map-raise - pr1 (canonical-model-theorem-pointwise ⊥ₘ x) = - map-raise ∘ is-consistent-modal-theory-L-complete-theory logic x - pr1 (canonical-model-theorem-pointwise (a →ₘ b) x) in-logic f = - forward-implication - ( canonical-model-theorem-pointwise b x) - ( weak-modal-logic-mp - ( is-weak-modal-logic-L-complete-theory logic lzero x) - ( in-logic) - ( backward-implication (canonical-model-theorem-pointwise a x) f)) - pr1 (canonical-model-theorem-pointwise (□ₘ a) x) in-logic y xRy = - forward-implication - ( canonical-model-theorem-pointwise a y) - ( xRy a in-logic) - pr2 (canonical-model-theorem-pointwise (varₘ n) x) = map-inv-raise - pr2 (canonical-model-theorem-pointwise ⊥ₘ x) (map-raise ()) - pr2 (canonical-model-theorem-pointwise (a →ₘ b) x) f = - L-complete-theory-implication x - ( λ in-x → - ( backward-implication - ( canonical-model-theorem-pointwise b x) - ( f - ( forward-implication - ( canonical-model-theorem-pointwise a x) - ( in-x))))) - pr2 (canonical-model-theorem-pointwise (□ₘ a) x) f = - L-complete-theory-box x - ( λ y xRy → - ( backward-implication - ( canonical-model-theorem-pointwise a y) - ( f y xRy))) - - canonical-model-theorem : - (a : modal-formula i) → - type-iff-Prop (logic a) (canonical-kripke-model ⊨M a) - pr1 (canonical-model-theorem a) in-logic x = - forward-implication - ( canonical-model-theorem-pointwise a x) - ( subset-logic-L-complete-theory logic lzero x a in-logic) - pr2 (canonical-model-theorem a) = - contraposition (lower-LEM l1 lem) (logic a) - ( λ na f → - ( apply-universal-property-trunc-Prop - ( lindenbaum logic contains-ax-k contains-ax-s zorn prop-resize - ( x , is-L-consistent-x na)) - ( empty-Prop) - ( λ (w , leq) → - ( is-consistent-modal-theory-L-complete-theory logic w - ( weak-modal-logic-mp - ( is-weak-modal-logic-L-complete-theory logic lzero w) - ( leq (¬ₘ a) not-a-in-x) - ( backward-implication - ( canonical-model-theorem-pointwise a w) - ( f w))))))) - where - x : modal-theory (l1 ⊔ l2) i - x = raise-subtype l2 (Id-formula-Prop (¬ₘ a)) - - not-a-in-x : is-in-subtype x ( ¬ₘ a) - not-a-in-x = - subset-equiv-subtypes (Id-formula-Prop (¬ₘ a)) x - ( compute-raise-subtype l2 (Id-formula-Prop (¬ₘ a))) - ( ¬ₘ a) - ( refl) - - is-L-consistent-x : - ¬ (is-in-subtype logic a) → is-L-consistent-theory logic x - is-L-consistent-x a-not-in-logic bot-in-logic = - a-not-in-logic - ( modal-logic-mp is-logic - {a = ¬¬ₘ a} - {b = a} - ( contains-ax-dn (¬¬ₘ a →ₘ a) (a , refl)) - ( is-logic (¬¬ₘ a) - ( subset-weak-modal-logic-closure-modal-logic-closure (¬¬ₘ a) + ( theory-add-formula (¬ₘ a) + ( list-subtype l-box)) + ( subtype-union-left + ( Id-formula-Prop (¬ₘ a)) + ( list-subtype l-box)) + ( l-not-a-sub)) + ( subset-add-formula + ( ¬ₘ a) + ( list-subtype l-box))) + ( sub-union'))) + ( sub-union))) + ( is-cont)))))) + + is-L-consistent-y : is-L-consistent-theory logic y + is-L-consistent-y = + map-universal-property-trunc-Prop + ( empty-Prop) + ( λ d-bot → + ( is-consistent-modal-theory-L-complete-theory logic x + ( weak-modal-logic-mp + ( is-weak-modal-logic-L-complete-theory logic lzero x) + ( not-box-a-in-logic) + ( ε + ( list-assumptions-weak-deduction d-bot) + ( subset-theory-list-assumptions-weak-deduction d-bot) + ( is-in-weak-modal-logic-closure-weak-deduction + ( is-assumptions-list-assumptions-weak-deduction + ( d-bot))))))) + + canonical-model-theorem-pointwise : + (a : modal-formula i) + (x : L-complete-theory logic (l1 ⊔ l2)) → + type-iff-Prop + ( modal-theory-L-complete-theory logic x a) + ( (canonical-kripke-model , x) ⊨ a) + pr1 (canonical-model-theorem-pointwise (varₘ n) x) = map-raise + pr1 (canonical-model-theorem-pointwise ⊥ₘ x) = + map-raise ∘ is-consistent-modal-theory-L-complete-theory logic x + pr1 (canonical-model-theorem-pointwise (a →ₘ b) x) in-logic f = + forward-implication + ( canonical-model-theorem-pointwise b x) + ( weak-modal-logic-mp + ( is-weak-modal-logic-L-complete-theory logic lzero x) + ( in-logic) + ( backward-implication (canonical-model-theorem-pointwise a x) f)) + pr1 (canonical-model-theorem-pointwise (□ₘ a) x) in-logic y xRy = + forward-implication + ( canonical-model-theorem-pointwise a y) + ( xRy a in-logic) + pr2 (canonical-model-theorem-pointwise (varₘ n) x) = map-inv-raise + pr2 (canonical-model-theorem-pointwise ⊥ₘ x) (map-raise ()) + pr2 (canonical-model-theorem-pointwise (a →ₘ b) x) f = + L-complete-theory-implication x + ( λ in-x → + ( backward-implication + ( canonical-model-theorem-pointwise b x) + ( f ( forward-implication - ( deduction-theorem logic contains-ax-k contains-ax-s - ( ¬ₘ a) - ( ⊥ₘ)) - ( weak-modal-logic-closure-monotic - ( subtype-union-both logic x (theory-add-formula (¬ₘ a) logic) - ( subtype-union-right (Id-formula-Prop (¬ₘ a)) logic) - ( transitive-leq-subtype x (Id-formula-Prop (¬ₘ a)) + ( canonical-model-theorem-pointwise a x) + ( in-x))))) + pr2 (canonical-model-theorem-pointwise (□ₘ a) x) f = + L-complete-theory-box x + ( λ y xRy → + ( backward-implication + ( canonical-model-theorem-pointwise a y) + ( f y xRy))) + + canonical-model-theorem : + (a : modal-formula i) → + type-iff-Prop (logic a) (canonical-kripke-model ⊨M a) + pr1 (canonical-model-theorem a) in-logic x = + forward-implication + ( canonical-model-theorem-pointwise a x) + ( subset-logic-L-complete-theory logic lzero x a in-logic) + pr2 (canonical-model-theorem a) = + contraposition (lower-LEM l1 lem) (logic a) + ( λ na f → + ( apply-universal-property-trunc-Prop + ( lindenbaum logic contains-ax-k contains-ax-s zorn prop-resize + ( x , is-L-consistent-x na)) + ( empty-Prop) + ( λ (w , leq) → + ( is-consistent-modal-theory-L-complete-theory logic w + ( weak-modal-logic-mp + ( is-weak-modal-logic-L-complete-theory logic lzero w) + ( leq (¬ₘ a) not-a-in-x) + ( backward-implication + ( canonical-model-theorem-pointwise a w) + ( f w))))))) + where + x : modal-theory (l1 ⊔ l2) i + x = raise-subtype l2 (Id-formula-Prop (¬ₘ a)) + + not-a-in-x : is-in-subtype x ( ¬ₘ a) + not-a-in-x = + subset-equiv-subtypes (Id-formula-Prop (¬ₘ a)) x + ( compute-raise-subtype l2 (Id-formula-Prop (¬ₘ a))) + ( ¬ₘ a) + ( refl) + + is-L-consistent-x : + ¬ (is-in-subtype logic a) → is-L-consistent-theory logic x + is-L-consistent-x a-not-in-logic bot-in-logic = + a-not-in-logic + ( modal-logic-mp is-logic + {a = ¬¬ₘ a} + {b = a} + ( contains-ax-dn (¬¬ₘ a →ₘ a) (a , refl)) + ( is-logic (¬¬ₘ a) + ( subset-weak-modal-logic-closure-modal-logic-closure (¬¬ₘ a) + ( forward-implication + ( deduction-theorem logic contains-ax-k contains-ax-s + ( ¬ₘ a) + ( ⊥ₘ)) + ( weak-modal-logic-closure-monotic + ( subtype-union-both logic x ( theory-add-formula (¬ₘ a) logic) - ( subtype-union-left (Id-formula-Prop (¬ₘ a)) logic) - ( inv-subset-equiv-subtypes (Id-formula-Prop (¬ₘ a)) x - ( compute-raise-subtype l2 (Id-formula-Prop (¬ₘ a)))))) - ( ⊥ₘ) - ( bot-in-logic)))))) - - canonical-model-completness : - {l3 : Level} - (C : model-class (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) i (l1 ⊔ l2) l3) → - is-in-subtype C canonical-kripke-model → - completeness logic C - canonical-model-completness C model-in-class a a-in-class-logic = - backward-implication - ( canonical-model-theorem a) - ( a-in-class-logic canonical-kripke-model model-in-class) + ( subtype-union-right (Id-formula-Prop (¬ₘ a)) logic) + ( transitive-leq-subtype x (Id-formula-Prop (¬ₘ a)) + ( theory-add-formula (¬ₘ a) logic) + ( subtype-union-left (Id-formula-Prop (¬ₘ a)) logic) + ( inv-subset-equiv-subtypes (Id-formula-Prop (¬ₘ a)) x + ( compute-raise-subtype l2 + ( Id-formula-Prop (¬ₘ a)))))) + ( ⊥ₘ) + ( bot-in-logic)))))) + + canonical-model-completness : + {l3 : Level} + (C : model-class (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) i (l1 ⊔ l2) l3) → + is-in-subtype C canonical-kripke-model → + completeness logic C + canonical-model-completness C model-in-class a a-in-class-logic = + backward-implication + ( canonical-model-theorem a) + ( a-in-class-logic canonical-kripke-model model-in-class) ``` diff --git a/src/modal-logic/canonical-theories.lagda.md b/src/modal-logic/canonical-theories.lagda.md index 23b428b971..ccc8dcd2fe 100644 --- a/src/modal-logic/canonical-theories.lagda.md +++ b/src/modal-logic/canonical-theories.lagda.md @@ -48,7 +48,6 @@ TODO module _ {l1 : Level} (i : Set l1) - (lem : LEM l1) -- TODO: not needed (zorn : Zorn-non-empty (lsuc l1) l1 l1) (prop-resize : propositional-resizing l1 (lsuc l1)) (logic : modal-theory l1 i) @@ -89,8 +88,13 @@ module _ ax-m i ⊆ logic → is-in-subtype ( reflexive-kripke-class (lsuc l1) l1 i l1) - ( canonical-kripke-model logic is-logic is-cons is-normal lem zorn - prop-resize) + ( canonical-kripke-model + ( logic) + ( is-logic) + ( is-cons) + ( is-normal) + ( zorn) + ( prop-resize)) is-canonical-ax-m ax-m-subset x a box-a-in-x = weak-modal-logic-mp ( is-weak-modal-logic-L-complete-theory logic lzero x) @@ -100,12 +104,18 @@ module _ ( box-a-in-x) is-canonical-ax-b : + LEM l1 → ax-b i ⊆ logic → is-in-subtype ( symmetry-kripke-class (lsuc l1) l1 i l1) - ( canonical-kripke-model logic is-logic is-cons is-normal lem zorn - prop-resize) - is-canonical-ax-b ax-b-subset x y xRy a box-a-in-y = + ( canonical-kripke-model + ( logic) + ( is-logic) + ( is-cons) + ( is-normal) + ( zorn) + ( prop-resize)) + is-canonical-ax-b lem ax-b-subset x y xRy a box-a-in-y = lemma-box-diamond-L-complete logic x ( contains-ax-k) ( contains-ax-s) @@ -131,8 +141,13 @@ module _ ax-4 i ⊆ logic → is-in-subtype ( transitivity-kripke-class (lsuc l1) l1 i l1) - ( canonical-kripke-model logic is-logic is-cons is-normal lem zorn - prop-resize) + ( canonical-kripke-model + ( logic) + ( is-logic) + ( is-cons) + ( is-normal) + ( zorn) + ( prop-resize)) is-canonical-ax-4 ax-4-subset x y z yRz xRy a box-a-in-x = yRz a ( xRy (□ₘ a) diff --git a/src/modal-logic/completeness-k.lagda.md b/src/modal-logic/completeness-k.lagda.md index 2d2401f843..4f614d5756 100644 --- a/src/modal-logic/completeness-k.lagda.md +++ b/src/modal-logic/completeness-k.lagda.md @@ -69,9 +69,9 @@ module _ ( is-modal-logic-K i) ( is-consistent-K i) ( is-normal-modal-logic-K i) - ( lem) ( zorn) ( prop-resize) + ( lem) ( all-models (lsuc l1) l1 i l1) ( star) ``` diff --git a/src/modal-logic/completeness-s5.lagda.md b/src/modal-logic/completeness-s5.lagda.md index f3a88e16d1..758b39623a 100644 --- a/src/modal-logic/completeness-s5.lagda.md +++ b/src/modal-logic/completeness-s5.lagda.md @@ -71,24 +71,24 @@ module _ ( is-modal-logic-S5 i) ( is-consistent-S5 i) ( is-normal-modal-logic-S5 i) - ( lem) ( zorn) ( prop-resize)) S5-canonical-model-is-equivalence = triple - ( is-canonical-ax-m i lem zorn prop-resize + ( is-canonical-ax-m i zorn prop-resize ( modal-logic-S5 i) ( is-modal-logic-S5 i) ( is-consistent-S5 i) ( is-normal-modal-logic-S5 i) ( modal-logic-S5-contains-ax-m i)) - ( is-canonical-ax-b i lem zorn prop-resize + ( is-canonical-ax-b i zorn prop-resize ( modal-logic-S5 i) ( is-modal-logic-S5 i) ( is-consistent-S5 i) ( is-normal-modal-logic-S5 i) + ( lem) ( modal-logic-S5-contains-ax-b i)) - ( is-canonical-ax-4 i lem zorn prop-resize + ( is-canonical-ax-4 i zorn prop-resize ( modal-logic-S5 i) ( is-modal-logic-S5 i) ( is-consistent-S5 i) @@ -103,9 +103,9 @@ module _ ( is-modal-logic-S5 i) ( is-consistent-S5 i) ( is-normal-modal-logic-S5 i) - ( lem) ( zorn) ( prop-resize) + ( lem) ( equivalence-kripke-class (lsuc l1) l1 i l1) ( S5-canonical-model-is-equivalence) ``` diff --git a/src/modal-logic/deduction.lagda.md b/src/modal-logic/deduction.lagda.md index 4ad1c4f583..f9f91e8d57 100644 --- a/src/modal-logic/deduction.lagda.md +++ b/src/modal-logic/deduction.lagda.md @@ -1,4 +1,4 @@ -# Modal logic syntax +# Modal logic deduction ```agda module modal-logic.deduction where @@ -48,16 +48,16 @@ module _ {l1 l2 : Level} {i : Set l1} where - infix 5 _⊢_ + infix 5 _⊢ₘ_ - data _⊢_ (axioms : modal-theory l2 i) : modal-formula i → UU (l1 ⊔ l2) where - ax : {a : modal-formula i} → is-in-subtype axioms a → axioms ⊢ a - mp : {a b : modal-formula i} → axioms ⊢ a →ₘ b → axioms ⊢ a → axioms ⊢ b - nec : {a : modal-formula i} → axioms ⊢ a → axioms ⊢ □ₘ a + data _⊢ₘ_ (axioms : modal-theory l2 i) : modal-formula i → UU (l1 ⊔ l2) where + modal-ax : {a : modal-formula i} → is-in-subtype axioms a → axioms ⊢ₘ a + modal-mp : + {a b : modal-formula i} → axioms ⊢ₘ a →ₘ b → axioms ⊢ₘ a → axioms ⊢ₘ b + modal-nec : {a : modal-formula i} → axioms ⊢ₘ a → axioms ⊢ₘ □ₘ a - -- TODO: rename to modal-logic-closure modal-logic-closure : modal-theory l2 i → modal-theory (l1 ⊔ l2) i - modal-logic-closure axioms a = trunc-Prop (axioms ⊢ a) + modal-logic-closure axioms a = trunc-Prop (axioms ⊢ₘ a) is-modal-logic-Prop : modal-theory l2 i → Prop (l1 ⊔ l2) is-modal-logic-Prop theory = @@ -68,7 +68,7 @@ module _ is-in-modal-logic-closure-deduction : {axioms : modal-theory l2 i} {a : modal-formula i} → - axioms ⊢ a → is-in-subtype (modal-logic-closure axioms) a + axioms ⊢ₘ a → is-in-subtype (modal-logic-closure axioms) a is-in-modal-logic-closure-deduction = unit-trunc-Prop is-contradictory-modal-logic-Prop : modal-theory l2 i → Prop l2 @@ -110,7 +110,7 @@ module _ {a : modal-formula i} → is-in-subtype axioms a → is-in-subtype (modal-logic-closure axioms) a - modal-logic-closure-ax = unit-trunc-Prop ∘ ax + modal-logic-closure-ax = unit-trunc-Prop ∘ modal-ax modal-logic-closure-mp : {a b : modal-formula i} → @@ -120,7 +120,7 @@ module _ modal-logic-closure-mp {a} {b} tdab tda = apply-twice-universal-property-trunc-Prop tdab tda ( modal-logic-closure axioms b) - ( λ dab da → unit-trunc-Prop (mp dab da)) + ( λ dab da → unit-trunc-Prop (modal-mp dab da)) modal-logic-closure-nec : {a : modal-formula i} → @@ -129,7 +129,7 @@ module _ modal-logic-closure-nec {a} = map-universal-property-trunc-Prop ( modal-logic-closure axioms (□ₘ a)) - ( λ da → unit-trunc-Prop (nec da)) + ( λ da → unit-trunc-Prop (modal-nec da)) module _ {l1 l2 : Level} {i : Set l1} @@ -161,24 +161,24 @@ module _ axioms-subset-modal-logic : {l2 : Level} (axioms : modal-theory l2 i) → axioms ⊆ modal-logic-closure axioms - axioms-subset-modal-logic _ a H = unit-trunc-Prop (ax H) + axioms-subset-modal-logic _ a H = unit-trunc-Prop (modal-ax H) modal-logic-closed : {l2 : Level} {axioms : modal-theory l2 i} {a : modal-formula i} → - modal-logic-closure axioms ⊢ a → + modal-logic-closure axioms ⊢ₘ a → is-in-subtype (modal-logic-closure axioms) a - modal-logic-closed (ax x) = x - modal-logic-closed (mp dab da) = + modal-logic-closed (modal-ax x) = x + modal-logic-closed (modal-mp dab da) = modal-logic-closure-mp (modal-logic-closed dab) (modal-logic-closed da) - modal-logic-closed (nec d) = + modal-logic-closed (modal-nec d) = modal-logic-closure-nec (modal-logic-closed d) -- TODO: refactor - subset-double-modal-logic : + is-modal-logic-modal-logic-closure : {l2 : Level} (axioms : modal-theory l2 i) → is-modal-logic (modal-logic-closure axioms) - subset-double-modal-logic axioms a = + is-modal-logic-modal-logic-closure axioms a = map-universal-property-trunc-Prop ( modal-logic-closure axioms a) ( modal-logic-closed) @@ -189,11 +189,11 @@ module _ (leq : ax₁ ⊆ ax₂) where - deduction-monotic : {a : modal-formula i} → ax₁ ⊢ a → ax₂ ⊢ a - deduction-monotic (ax x) = ax (leq _ x) - deduction-monotic (mp dab da) = - mp (deduction-monotic dab) (deduction-monotic da) - deduction-monotic (nec d) = nec (deduction-monotic d) + deduction-monotic : {a : modal-formula i} → ax₁ ⊢ₘ a → ax₂ ⊢ₘ a + deduction-monotic (modal-ax x) = modal-ax (leq _ x) + deduction-monotic (modal-mp dab da) = + modal-mp (deduction-monotic dab) (deduction-monotic da) + deduction-monotic (modal-nec d) = modal-nec (deduction-monotic d) modal-logic-monotic : modal-logic-closure ax₁ ⊆ modal-logic-closure ax₂ modal-logic-monotic a = @@ -214,7 +214,7 @@ module _ ( modal-logic-closure ax₁) ( modal-logic-closure (modal-logic-closure ax₂)) ( modal-logic-closure ax₂) - ( subset-double-modal-logic ax₂) + ( is-modal-logic-modal-logic-closure ax₂) ( modal-logic-monotic leq) module _ diff --git a/src/modal-logic/kripke-models-filtrations.lagda.md b/src/modal-logic/kripke-models-filtrations.lagda.md index cf456f0287..018e1293e3 100644 --- a/src/modal-logic/kripke-models-filtrations.lagda.md +++ b/src/modal-logic/kripke-models-filtrations.lagda.md @@ -290,12 +290,9 @@ module _ apply-effectiveness-class' ( Φ-equivalence) ( λ a a-in-theory → - ( g-val (a , a-in-theory) (y , y-in-class) - ∘iff iff-eq (htpy-eq p (a , a-in-theory)) - ∘iff inv-iff - ( f-val - ( a , a-in-theory) - ( x , x-in-class)))) + ( g-val (a , a-in-theory) (y , y-in-class) ∘iff + iff-eq (htpy-eq p (a , a-in-theory)) ∘iff + inv-iff (f-val (a , a-in-theory) (x , x-in-class)))) = y-class by eq-class-equivalence-class diff --git a/src/modal-logic/modal-logic-k.lagda.md b/src/modal-logic/modal-logic-k.lagda.md index 8f31830b90..87bdd294df 100644 --- a/src/modal-logic/modal-logic-k.lagda.md +++ b/src/modal-logic/modal-logic-k.lagda.md @@ -53,7 +53,7 @@ module _ modal-logic-K = modal-logic-closure modal-logic-K-axioms is-modal-logic-K : is-modal-logic modal-logic-K - is-modal-logic-K = subset-double-modal-logic modal-logic-K-axioms + is-modal-logic-K = is-modal-logic-modal-logic-closure modal-logic-K-axioms K-axioms-contains-ax-k : ax-k i ⊆ modal-logic-K-axioms K-axioms-contains-ax-k = diff --git a/src/modal-logic/modal-logic-s5.lagda.md b/src/modal-logic/modal-logic-s5.lagda.md index 4bb29a447b..fc9fecbeed 100644 --- a/src/modal-logic/modal-logic-s5.lagda.md +++ b/src/modal-logic/modal-logic-s5.lagda.md @@ -51,7 +51,7 @@ module _ modal-logic-S5 = modal-logic-closure modal-logic-S5-axioms is-modal-logic-S5 : is-modal-logic modal-logic-S5 - is-modal-logic-S5 = subset-double-modal-logic modal-logic-S5-axioms + is-modal-logic-S5 = is-modal-logic-modal-logic-closure modal-logic-S5-axioms is-normal-modal-logic-S5 : is-normal-modal-logic modal-logic-S5 is-normal-modal-logic-S5 = diff --git a/src/modal-logic/soundness.lagda.md b/src/modal-logic/soundness.lagda.md index 5e7a04cee6..f583a78c93 100644 --- a/src/modal-logic/soundness.lagda.md +++ b/src/modal-logic/soundness.lagda.md @@ -52,11 +52,14 @@ module _ where soundness-axioms : - soundness axioms C → {a : modal-formula i} → axioms ⊢ a → type-Prop (C ⊨C a) - soundness-axioms H (ax x) = H _ x - soundness-axioms H (mp dab da) M in-class x = + soundness axioms C → + {a : modal-formula i} → + axioms ⊢ₘ a → + type-Prop (C ⊨C a) + soundness-axioms H (modal-ax x) = H _ x + soundness-axioms H (modal-mp dab da) M in-class x = soundness-axioms H dab M in-class x (soundness-axioms H da M in-class x) - soundness-axioms H (nec d) M in-class _ y _ = + soundness-axioms H (modal-nec d) M in-class _ y _ = soundness-axioms H d M in-class y soundness-modal-logic : diff --git a/src/modal-logic/weak-deduction.lagda.md b/src/modal-logic/weak-deduction.lagda.md index c4b9db49d5..6428a55323 100644 --- a/src/modal-logic/weak-deduction.lagda.md +++ b/src/modal-logic/weak-deduction.lagda.md @@ -52,29 +52,31 @@ module _ where is-weak-deduction-Prop : - {axioms : modal-theory l2 i} {a : modal-formula i} → axioms ⊢ a → Prop lzero - is-weak-deduction-Prop (ax x) = unit-Prop - is-weak-deduction-Prop (mp dab da) = + {axioms : modal-theory l2 i} {a : modal-formula i} → + axioms ⊢ₘ a → + Prop lzero + is-weak-deduction-Prop (modal-ax x) = unit-Prop + is-weak-deduction-Prop (modal-mp dab da) = conjunction-Prop (is-weak-deduction-Prop dab) (is-weak-deduction-Prop da) - is-weak-deduction-Prop (nec d) = empty-Prop + is-weak-deduction-Prop (modal-nec d) = empty-Prop is-weak-deduction : - {axioms : modal-theory l2 i} {a : modal-formula i} → axioms ⊢ a → UU lzero + {axioms : modal-theory l2 i} {a : modal-formula i} → axioms ⊢ₘ a → UU lzero is-weak-deduction = type-Prop ∘ is-weak-deduction-Prop - infix 5 _⊢w_ + infix 5 _⊢ₘw_ - _⊢w_ : modal-theory l2 i → modal-formula i → UU (l1 ⊔ l2) - axioms ⊢w a = type-subtype (is-weak-deduction-Prop {axioms} {a}) + _⊢ₘw_ : modal-theory l2 i → modal-formula i → UU (l1 ⊔ l2) + axioms ⊢ₘw a = type-subtype (is-weak-deduction-Prop {axioms} {a}) deduction-weak-deduction : {axioms : modal-theory l2 i} {a : modal-formula i} → - axioms ⊢w a → - axioms ⊢ a + axioms ⊢ₘw a → + axioms ⊢ₘ a deduction-weak-deduction = inclusion-subtype is-weak-deduction-Prop is-weak-deduction-deduction-weak-deduction : - {axioms : modal-theory l2 i} {a : modal-formula i} (d : axioms ⊢w a) → + {axioms : modal-theory l2 i} {a : modal-formula i} (d : axioms ⊢ₘw a) → is-weak-deduction (deduction-weak-deduction d) is-weak-deduction-deduction-weak-deduction = is-in-subtype-inclusion-subtype is-weak-deduction-Prop @@ -82,28 +84,28 @@ module _ weak-deduction-ax : {axioms : modal-theory l2 i} {a : modal-formula i} → is-in-subtype axioms a → - axioms ⊢w a - weak-deduction-ax in-axioms = ax in-axioms , star + axioms ⊢ₘw a + weak-deduction-ax in-axioms = modal-ax in-axioms , star weak-deduction-mp : {axioms : modal-theory l2 i} {a b : modal-formula i} → - axioms ⊢w (a →ₘ b) → - axioms ⊢w a → - axioms ⊢w b + axioms ⊢ₘw (a →ₘ b) → + axioms ⊢ₘw a → + axioms ⊢ₘw b weak-deduction-mp (dab , is-w-dab) (da , is-w-da) = - mp dab da , is-w-dab , is-w-da + modal-mp dab da , is-w-dab , is-w-da ind-weak-deduction : {l : Level} {axioms : modal-theory l2 i} - (P : {a : modal-formula i} → axioms ⊢w a → UU l) → + (P : {a : modal-formula i} → axioms ⊢ₘw a → UU l) → ( {a : modal-formula i} (in-axioms : is-in-subtype axioms a) → P (weak-deduction-ax in-axioms)) → - ( {a b : modal-formula i} (dab : axioms ⊢w a →ₘ b) (da : axioms ⊢w a) → + ( {a b : modal-formula i} (dab : axioms ⊢ₘw a →ₘ b) (da : axioms ⊢ₘw a) → P dab → P da → P (weak-deduction-mp dab da)) → - {a : modal-formula i} (wd : axioms ⊢w a) → P wd - ind-weak-deduction P H-ax H-mp (ax in-axioms , _) = + {a : modal-formula i} (wd : axioms ⊢ₘw a) → P wd + ind-weak-deduction P H-ax H-mp (modal-ax in-axioms , _) = H-ax in-axioms - ind-weak-deduction P H-ax H-mp (mp dba db , is-w-dba , is-w-db) = + ind-weak-deduction P H-ax H-mp (modal-mp dba db , is-w-dba , is-w-db) = H-mp _ _ ( ind-weak-deduction P H-ax H-mp (dba , is-w-dba)) ( ind-weak-deduction P H-ax H-mp (db , is-w-db)) @@ -111,13 +113,13 @@ module _ rec-weak-deduction : {l : Level} {axioms : modal-theory l2 i} {P : UU l} → ( {a : modal-formula i} (in-axioms : is-in-subtype axioms a) → P) → - ( {a b : modal-formula i} (dab : axioms ⊢w a →ₘ b) (da : axioms ⊢w a) → + ( {a b : modal-formula i} (dab : axioms ⊢ₘw a →ₘ b) (da : axioms ⊢ₘw a) → P → P → P) → - {a : modal-formula i} (wd : axioms ⊢w a) → P + {a : modal-formula i} (wd : axioms ⊢ₘw a) → P rec-weak-deduction {P = P} = ind-weak-deduction (λ _ → P) weak-modal-logic-closure : modal-theory l2 i → modal-theory (l1 ⊔ l2) i - weak-modal-logic-closure axioms a = trunc-Prop (axioms ⊢w a) + weak-modal-logic-closure axioms a = trunc-Prop (axioms ⊢ₘw a) is-weak-modal-logic-Prop : modal-theory l2 i → Prop (l1 ⊔ l2) is-weak-modal-logic-Prop theory = @@ -128,7 +130,7 @@ module _ is-in-weak-modal-logic-closure-weak-deduction : {axioms : modal-theory l2 i} {a : modal-formula i} → - axioms ⊢w a → is-in-subtype (weak-modal-logic-closure axioms) a + axioms ⊢ₘw a → is-in-subtype (weak-modal-logic-closure axioms) a is-in-weak-modal-logic-closure-weak-deduction = unit-trunc-Prop subset-weak-modal-logic-closure-modal-logic-closure : @@ -196,7 +198,7 @@ module _ where weak-modal-logic-closed : - {a : modal-formula i} → weak-modal-logic-closure axioms ⊢w a → + {a : modal-formula i} → weak-modal-logic-closure axioms ⊢ₘw a → is-in-subtype (weak-modal-logic-closure axioms) a weak-modal-logic-closed = ind-weak-deduction _ @@ -222,7 +224,7 @@ module _ (leq : ax₁ ⊆ ax₂) where - weak-deduction-monotic : {a : modal-formula i} → ax₁ ⊢w a → ax₂ ⊢w a + weak-deduction-monotic : {a : modal-formula i} → ax₁ ⊢ₘw a → ax₂ ⊢ₘw a weak-deduction-monotic = ind-weak-deduction _ ( λ {a} in-axioms → weak-deduction-ax (leq a in-axioms)) @@ -272,8 +274,8 @@ module _ backward-deduction-theorem : {a b : modal-formula i} → - axioms ⊢w a →ₘ b → - theory-add-formula a axioms ⊢w b + axioms ⊢ₘw a →ₘ b → + theory-add-formula a axioms ⊢ₘw b backward-deduction-theorem {a} wab = weak-deduction-mp ( weak-deduction-monotic @@ -288,7 +290,7 @@ module _ -- TODO: move to file with deduction deduction-a->a : - (a : modal-formula i) → axioms ⊢w a →ₘ a + (a : modal-formula i) → axioms ⊢ₘw a →ₘ a deduction-a->a a = weak-deduction-mp ( weak-deduction-mp @@ -298,7 +300,7 @@ module _ forward-deduction-theorem : (a : modal-formula i) {b : modal-formula i} → - theory-add-formula a axioms ⊢w b → + theory-add-formula a axioms ⊢ₘw b → is-in-subtype (weak-modal-logic-closure axioms) (a →ₘ b) forward-deduction-theorem a = ind-weak-deduction _ @@ -346,7 +348,7 @@ module _ list-assumptions-weak-deduction : {l2 : Level} {theory : modal-theory l2 i} {a : modal-formula i} → - theory ⊢w a → list (modal-formula i) + theory ⊢ₘw a → list (modal-formula i) list-assumptions-weak-deduction = rec-weak-deduction ( λ {a} _ → cons a nil) @@ -354,7 +356,7 @@ module _ subset-theory-list-assumptions-weak-deduction : {l2 : Level} {theory : modal-theory l2 i} {a : modal-formula i} → - (d : theory ⊢w a) → + (d : theory ⊢ₘw a) → list-subtype (list-assumptions-weak-deduction d) ⊆ theory subset-theory-list-assumptions-weak-deduction {theory = theory} = ind-weak-deduction @@ -381,11 +383,11 @@ module _ is-assumptions-list-assumptions-weak-deduction : {l2 : Level} {theory : modal-theory l2 i} {a : modal-formula i} → - (d : theory ⊢w a) → - list-subtype (list-assumptions-weak-deduction d) ⊢w a + (d : theory ⊢ₘw a) → + list-subtype (list-assumptions-weak-deduction d) ⊢ₘw a is-assumptions-list-assumptions-weak-deduction {theory = theory} = ind-weak-deduction - ( λ {a} d → list-subtype (list-assumptions-weak-deduction d) ⊢w a) + ( λ {a} d → list-subtype (list-assumptions-weak-deduction d) ⊢ₘw a) ( λ _ → weak-deduction-ax head-in-list-subtype) ( λ dab da rab ra → ( weak-deduction-monotic From b978ed19dcb9ac6a3e5fe56cf29204b2a7337ed2 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Mon, 6 May 2024 18:35:32 +0300 Subject: [PATCH 55/63] Prove inductor for equivalence classes --- src/foundation/equivalence-classes.lagda.md | 220 +++++++++++++++--- .../canonical-model-theorem.lagda.md | 5 +- src/modal-logic/decision-procedure.lagda.md | 6 +- .../kripke-models-filtrations.lagda.md | 156 +++++-------- src/modal-logic/kripke-semantics.lagda.md | 8 +- 5 files changed, 264 insertions(+), 131 deletions(-) diff --git a/src/foundation/equivalence-classes.lagda.md b/src/foundation/equivalence-classes.lagda.md index 329ad16806..a077a35775 100644 --- a/src/foundation/equivalence-classes.lagda.md +++ b/src/foundation/equivalence-classes.lagda.md @@ -7,13 +7,16 @@ module foundation.equivalence-classes where
Imports ```agda +open import foundation.action-on-identifications-functions open import foundation.conjunction +open import foundation.constant-type-families open import foundation.dependent-pair-types open import foundation.effective-maps-equivalence-relations open import foundation.equality-dependent-pair-types open import foundation.existential-quantification open import foundation.functoriality-propositional-truncation open import foundation.fundamental-theorem-of-identity-types +open import foundation.homotopies open import foundation.inhabited-subtypes open import foundation.locally-small-types open import foundation.logical-equivalences @@ -25,6 +28,7 @@ open import foundation.small-types open import foundation.subtype-identity-principle open import foundation.subtypes open import foundation.surjective-maps +open import foundation.transport-along-identifications open import foundation.universal-property-image open import foundation.universe-levels @@ -38,7 +42,6 @@ open import foundation-core.functoriality-dependent-pair-types open import foundation-core.identity-types open import foundation-core.propositions open import foundation-core.torsorial-type-families -open import foundation-core.transport-along-identifications ```
@@ -390,6 +393,12 @@ module _ ### TODO: title ```agda + is-retraction-apply-effectiveness-class : + (x y : A) → + apply-effectiveness-class' {x} {y} ∘ apply-effectiveness-class {x} {y} ~ id + is-retraction-apply-effectiveness-class x y p = + eq-is-prop (is-set-equivalence-class R (class R x) (class R y)) + eq-class-in-common-class : (c : equivalence-class R) {a a' : A} → is-in-equivalence-class R c a → @@ -413,25 +422,104 @@ module _ ### TODO: Eliminator ```agda - rec-equivalence-class : - {l3 : Level} (B : Set l3) → - (f : A → type-Set B) → - ((a a' : A) → sim-equivalence-relation R a a' → f a = f a') → - equivalence-class R → type-Set B - rec-equivalence-class {l3} B f H c = + ind-equivalence-class : + {l3 : Level} (B : equivalence-class R → Set l3) → + (f : (a : A) → type-Set (B (class R a))) → + ( (a a' : A) → + (s : sim-equivalence-relation R a a') → + tr (type-Set ∘ B) (apply-effectiveness-class' s) (f a) = f a') → + (c : equivalence-class R) → type-Set (B c) + ind-equivalence-class {l3} B f H c = pr1 ( apply-universal-property-trunc-Prop ( is-inhabited-subtype-equivalence-class R c) ( b , is-prop-b) ( λ (a , a-in-c) → - ( pair (f a) + ( pair + ( tr (type-Set ∘ B) (eq-class-equivalence-class R c a-in-c) (f a)) ( λ a' a'-in-c → - ( H a a' - ( sim-equivalence-relation-in-same-class c a-in-c a'-in-c)))))) + ( helper a a' a-in-c a'-in-c + ( H a a' + ( sim-equivalence-relation-in-same-class c + ( a-in-c) + ( a'-in-c)))))))) where + helper : + (a a' : A) → + (a-in-c : is-in-equivalence-class R c a) → + (a'-in-c : is-in-equivalence-class R c a') → + tr (type-Set ∘ B) + ( apply-effectiveness-class' + (sim-equivalence-relation-in-same-class c a-in-c a'-in-c)) + (f a) + = f a' → + tr (type-Set ∘ B) (eq-class-equivalence-class R c a'-in-c) (f a') + = tr (type-Set ∘ B) (eq-class-equivalence-class R c a-in-c) (f a) + helper a a' a-in-c a'-in-c p = + let q = is-retraction-apply-effectiveness-class a a' + ( eq-class-in-common-class c a-in-c a'-in-c) + s = tr (λ t → tr (type-Set ∘ B) t (f a) = f a') q p + u = ap (tr (type-Set ∘ B) (eq-effective-quotient' a' c a'-in-c)) s + in + equational-reasoning + tr (type-Set ∘ B) (eq-class-equivalence-class R c a'-in-c) (f a') + = tr (type-Set ∘ B) (eq-effective-quotient' a' c a'-in-c) (f a') + by + ap + ( λ t → tr (type-Set ∘ B) t (f a')) + ( eq-is-prop (is-set-equivalence-class R (class R a') c)) + = tr (type-Set ∘ B) + ( eq-effective-quotient' a' c a'-in-c) + ( tr (type-Set ∘ B) + ( eq-class-in-common-class c a-in-c a'-in-c) + ( f a)) + by inv u + = tr (type-Set ∘ B) + ( eq-effective-quotient' a c a-in-c ∙ + inv (eq-effective-quotient' a' c a'-in-c) ∙ + eq-effective-quotient' a' c a'-in-c) + ( f a) + by + inv + ( tr-concat + ( eq-class-in-common-class c a-in-c a'-in-c) + ( eq-effective-quotient' a' c a'-in-c) + ( f a)) + = tr (type-Set ∘ B) + ( eq-effective-quotient' a c a-in-c ∙ + ( inv (eq-effective-quotient' a' c a'-in-c) ∙ + eq-effective-quotient' a' c a'-in-c)) + ( f a) + by + ap + ( λ t → tr (type-Set ∘ B) t (f a)) + ( assoc + ( eq-effective-quotient' a c a-in-c) + ( inv (eq-effective-quotient' a' c a'-in-c)) + ( eq-effective-quotient' a' c a'-in-c)) + = tr (type-Set ∘ B) (eq-effective-quotient' a c a-in-c ∙ refl) (f a) + by + ap + ( λ t → + ( tr (type-Set ∘ B) + ( eq-effective-quotient' a c a-in-c ∙ t) + ( f a))) + ( left-inv (eq-effective-quotient' a' c a'-in-c)) + = tr (type-Set ∘ B) (eq-effective-quotient' a c a-in-c) (f a) + by ap (λ t → tr (type-Set ∘ B) t (f a)) right-unit + = tr (type-Set ∘ B) (eq-class-equivalence-class R c a-in-c) (f a) + by + ap + ( λ t → tr (type-Set ∘ B) t (f a)) + ( eq-is-prop (is-set-equivalence-class R (class R a) c)) + b : UU (l1 ⊔ l2 ⊔ l3) b = - Σ (type-Set B) (λ b → (a : A) → is-in-equivalence-class R c a → b = f a) + Σ ( type-Set (B c)) + ( λ b → + (a : A) → + (a-in-c : is-in-equivalence-class R c a) → + tr (type-Set ∘ B) (eq-class-equivalence-class R c a-in-c) (f a) = b) is-prop-b : is-prop b is-prop-b = @@ -440,33 +528,107 @@ module _ ( eq-pair-Σ ( apply-universal-property-trunc-Prop ( is-inhabited-subtype-equivalence-class R c) - ( b = b' , is-set-type-Set B b b') - ( λ (a , a-in-c) → h a a-in-c ∙ inv (h' a a-in-c))) + ( b = b' , is-set-type-Set (B c) b b') + ( λ (a , a-in-c) → inv (h a a-in-c) ∙ h' a a-in-c)) ( eq-is-prop ( is-prop-Π - ( λ x → is-prop-function-type (is-set-type-Set B b' (f x))))))) + ( λ a → + ( is-prop-Π + ( λ a-in-c → + ( is-set-type-Set (B c) + ( tr (type-Set ∘ B) + ( eq-class-equivalence-class R c a-in-c) + ( f a)) + ( b'))))))))) - rec-equivalence-class-Prop : + rec-equivalence-class : + {l3 : Level} (B : Set l3) → + (f : A → type-Set B) → + ((a a' : A) → sim-equivalence-relation R a a' → f a = f a') → + equivalence-class R → type-Set B + rec-equivalence-class {l3} B f H = + ind-equivalence-class (λ _ → B) f + ( λ a a' s → + equational-reasoning + tr (type-Set ∘ (λ _ → B)) (apply-effectiveness-class' s) (f a) + = f a + by tr-constant-type-family (apply-effectiveness-class' s) (f a) + = f a' + by H a a' s) + + ind-equivalence-class-prop : + {l3 : Level} (B : equivalence-class R → Prop l3) → + (f : (a : A) → type-Prop (B (class R a))) → + (c : equivalence-class R) → type-Prop (B c) + ind-equivalence-class-prop B f = + ind-equivalence-class (λ a → set-Prop (B a)) f + ( λ a a' _ → eq-is-prop (is-prop-type-Prop (B (class R a')))) + + rec-equivalence-class-prop : {l3 : Level} (B : Prop l3) → (f : A → type-Prop B) → equivalence-class R → type-Prop B - rec-equivalence-class-Prop B f = + rec-equivalence-class-prop B f = rec-equivalence-class (set-Prop B) f ( λ a a' _ → eq-is-prop (is-prop-type-Prop B)) - ind-equivalence-class-Prop : - {l3 : Level} (B : equivalence-class R → Prop l3) → - (f : (a : A) → type-Prop (B (class R a))) → - (c : equivalence-class R) → type-Prop (B c) - ind-equivalence-class-Prop B f c = - apply-universal-property-trunc-Prop - ( is-inhabited-subtype-equivalence-class R c) - ( B c) - ( λ (a , a-in-c) → - ( tr - ( type-Prop ∘ B) - ( eq-effective-quotient' a c a-in-c) - ( rec-equivalence-class-Prop (Π-Prop A (B ∘ class R)) (λ _ → f) c a))) + -- TODO: proof + postulate + compute-rec-equivalence-class' : + {l3 : Level} (B : Set l3) → + (f : A → type-Set B) → + (H : (a a' : A) → sim-equivalence-relation R a a' → f a = f a') → + (a : A) → + rec-equivalence-class B f H (class R a) = f a + + -- TODO: complete + -- compute-rec-equivalence-class'' : + -- {l3 : Level} (B : Set l3) → + -- (f : A → type-Set B) → + -- (H : (a a' : A) → sim-equivalence-relation R a a' → f a = f a') → + -- (c : equivalence-class R) → + -- (a : A) → + -- (a-in-c : is-in-equivalence-class R c a) → + -- rec-equivalence-class B f H c = f a + -- compute-rec-equivalence-class'' B f H c a a-in-c = + -- let q = + -- ind-equivalence-class + -- ( λ c' → + -- ( Σ-Set B + -- ( λ b → set-Prop (exists-structure-Prop A (λ a' → f a = b))))) + -- ( λ a' → (f a) , intro-exists a refl) + -- ( λ x y s → + -- ( eq-pair-Σ + -- ( equational-reasoning + -- {! !} + -- -- tr {! !} (apply-effectiveness-class' s) (f a) {! !} + -- -- = {! !} by {! !} + -- = f a by {! !}) + -- ( eq-is-prop (is-prop-exists-structure _ _)))) + -- ( c) + -- in apply-universal-property-trunc-Prop + -- ( pr2 q) + -- ( Id-Prop B (rec-equivalence-class B f H c) (f a)) + -- ( λ (a , h) → inv (h ∙ {! !})) + + compute-rec-equivalence-class : + {l3 : Level} (B : Set l3) → + (f : A → type-Set B) → + (H : (a a' : A) → sim-equivalence-relation R a a' → f a = f a') → + (c : equivalence-class R) → + (a : A) → + (a-in-c : is-in-equivalence-class R c a) → + rec-equivalence-class B f H c = f a + compute-rec-equivalence-class B f H c a a-in-c = + equational-reasoning + rec-equivalence-class B f H c + = rec-equivalence-class B f H (class R a) + by + ap + ( rec-equivalence-class B f H) + ( inv (eq-effective-quotient' a c a-in-c)) + = f a + by compute-rec-equivalence-class' B f H a ``` ### The map `class` into the type of equivalence classes is surjective and effective diff --git a/src/modal-logic/canonical-model-theorem.lagda.md b/src/modal-logic/canonical-model-theorem.lagda.md index b76752b205..fd9f12613b 100644 --- a/src/modal-logic/canonical-model-theorem.lagda.md +++ b/src/modal-logic/canonical-model-theorem.lagda.md @@ -109,9 +109,8 @@ module _ Π-Prop ( modal-formula i) ( λ a → - ( hom-Prop - ( modal-theory-L-complete-theory logic x (□ₘ a)) - ( modal-theory-L-complete-theory logic y a))) + ( modal-theory-L-complete-theory logic x (□ₘ a) ⇒ + modal-theory-L-complete-theory logic y a)) pr2 canonical-kripke-model n x = modal-theory-L-complete-theory logic x (varₘ n) diff --git a/src/modal-logic/decision-procedure.lagda.md b/src/modal-logic/decision-procedure.lagda.md index 2ea9d21d0a..df05d75ed1 100644 --- a/src/modal-logic/decision-procedure.lagda.md +++ b/src/modal-logic/decision-procedure.lagda.md @@ -349,12 +349,12 @@ module _ ... | inl x = pr1 x ... | inr x = a - is-invertable-surjection-from-injection : + is-invertible-surjection-from-injection : (a0 : type-Set A) → (inj@(f , _) : injection (type-Set A) (type-Set B)) → (a : type-Set A) → surjection-from-injection a0 inj (f a) = a - is-invertable-surjection-from-injection a0 (f , is-inj) a + is-invertible-surjection-from-injection a0 (f , is-inj) a with lem2 ( pair @@ -373,7 +373,7 @@ module _ is-surjective (surjection-from-injection a inj) is-surjective-surjection-from-injection a0 (f , is-inj) a = unit-trunc-Prop - ( f a , is-invertable-surjection-from-injection a0 (f , is-inj) a) + ( f a , is-invertible-surjection-from-injection a0 (f , is-inj) a) kuratowsky-finite-set-injection : injection (type-Set A) (type-Set B) → diff --git a/src/modal-logic/kripke-models-filtrations.lagda.md b/src/modal-logic/kripke-models-filtrations.lagda.md index 018e1293e3..3fda21311d 100644 --- a/src/modal-logic/kripke-models-filtrations.lagda.md +++ b/src/modal-logic/kripke-models-filtrations.lagda.md @@ -195,110 +195,83 @@ module _ pr2 (pr2 (pr2 Φ-equivalence)) x y z r-xy r-yz a in-theory = r-xy a in-theory ∘iff r-yz a in-theory - valuate-function-equivalence-class : - equivalence-class Φ-equivalence → - UU (lsuc l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ l5) - valuate-function-equivalence-class class = - Σ ( type-subtype theory → Prop (l1 ⊔ l2 ⊔ l4)) - ( λ f → - ( (a , in-theory) : type-subtype theory) - ( (x , _) : - type-subtype (subtype-equivalence-class Φ-equivalence class)) → - type-iff-Prop (f (a , in-theory)) ((M , x) ⊨ a)) - - is-prop-valuate-function-equivalence-class : - (class : equivalence-class Φ-equivalence) → - is-prop (valuate-function-equivalence-class class) - is-prop-valuate-function-equivalence-class class = - apply-universal-property-trunc-Prop - ( is-inhabited-subtype-equivalence-class Φ-equivalence class) - ( is-prop _ , is-prop-is-prop _) - ( λ x → - ( is-prop-all-elements-equal - ( λ (f , f-val) (g , g-val) → - ( eq-pair-Σ - ( eq-htpy - ( λ a → - ( eq-iff - ( λ fa → - ( backward-implication - ( g-val a x) - ( forward-implication (f-val a x) fa))) - ( λ ga → - ( backward-implication - ( f-val a x) - ( forward-implication (g-val a x) ga)))))) - ( eq-is-prop - ( is-prop-Π - ( λ (a , in-theory) → - ( is-prop-Π - ( λ (x , _) → - ( is-prop-iff-Prop - ( g (a , in-theory)) - ( (M , x) ⊨ a))))))))))) - - function-equivalence-class : - (class : equivalence-class Φ-equivalence) → - valuate-function-equivalence-class class - function-equivalence-class class = - apply-universal-property-trunc-Prop - ( is-inhabited-subtype-equivalence-class Φ-equivalence class) - ( pair - ( valuate-function-equivalence-class class) - ( is-prop-valuate-function-equivalence-class class)) - ( λ (x , x-in-class) → - ( pair - ( λ (a , _) → (M , x) ⊨ a) - ( λ (a , in-theory) (y , y-in-class) → - ( apply-effectiveness-class Φ-equivalence {x} {y} - ( concat - ( eq-effective-quotient' Φ-equivalence x class x-in-class) - ( _) - ( inv - ( eq-effective-quotient' Φ-equivalence y class y-in-class))) - ( a) - ( in-theory))))) + map-function-equivalence-class-Set : + Set (lsuc l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ l5) + map-function-equivalence-class-Set = + function-Set (type-subtype theory) (Prop-Set (l1 ⊔ l2 ⊔ l4)) + + map-function-worlds : + type-kripke-model i M → type-Set map-function-equivalence-class-Set + map-function-worlds x (a , _) = (M , x) ⊨ a + + map-function-worlds-correct : + (x y : type-kripke-model i M) → + sim-equivalence-relation Φ-equivalence x y → + map-function-worlds x = map-function-worlds y + map-function-worlds-correct x y s = + eq-htpy + ( λ (a , a-in-theory) → + ( eq-iff' ((M , x) ⊨ a) ((M , y) ⊨ a) (s a a-in-theory))) map-function-equivalence-class : equivalence-class Φ-equivalence → type-subtype theory → Prop (l1 ⊔ l2 ⊔ l4) - map-function-equivalence-class = pr1 ∘ function-equivalence-class + map-function-equivalence-class = + rec-equivalence-class Φ-equivalence + ( map-function-equivalence-class-Set) + ( map-function-worlds) + ( map-function-worlds-correct) is-injective-map-function-equivalence-class : is-injective map-function-equivalence-class is-injective-map-function-equivalence-class {x-class} {y-class} p = - let (f , f-val) = function-equivalence-class x-class - (g , g-val) = function-equivalence-class y-class - in apply-twice-universal-property-trunc-Prop + apply-twice-universal-property-trunc-Prop ( is-inhabited-subtype-equivalence-class Φ-equivalence x-class) ( is-inhabited-subtype-equivalence-class Φ-equivalence y-class) ( pair ( x-class = y-class) ( is-set-equivalence-class Φ-equivalence x-class y-class)) ( λ (x , x-in-class) (y , y-in-class) → - ( equational-reasoning - x-class - = class Φ-equivalence x - by - inv - ( eq-class-equivalence-class - ( Φ-equivalence) - ( x-class) - ( x-in-class)) - = class Φ-equivalence y - by - apply-effectiveness-class' - ( Φ-equivalence) - ( λ a a-in-theory → - ( g-val (a , a-in-theory) (y , y-in-class) ∘iff - iff-eq (htpy-eq p (a , a-in-theory)) ∘iff - inv-iff (f-val (a , a-in-theory) (x , x-in-class)))) - = y-class - by - eq-class-equivalence-class - ( Φ-equivalence) - ( y-class) - ( y-in-class))) + ( eq-share-common-element-equivalence-class Φ-equivalence + ( x-class) + ( y-class) + ( intro-exists x + ( pair + ( x-in-class) + ( transitive-is-in-equivalence-class Φ-equivalence + ( y-class) + ( y) + ( x) + ( y-in-class) + ( λ a a-in-theory → + ( iff-eq + ( inv + ( ap (λ f → f (a , a-in-theory)) + ( equational-reasoning + map-function-worlds x + = map-function-equivalence-class x-class + by + inv + ( compute-rec-equivalence-class + ( Φ-equivalence) + ( map-function-equivalence-class-Set) + ( map-function-worlds) + ( map-function-worlds-correct) + ( x-class) + ( x) + ( x-in-class)) + = map-function-equivalence-class y-class + by p + = map-function-worlds y + by + compute-rec-equivalence-class + ( Φ-equivalence) + ( map-function-equivalence-class-Set) + ( map-function-worlds) + ( map-function-worlds-correct) + ( y-class) + ( y) + ( y-in-class))))))))))) injection-map-function-equivalence-class : injection @@ -364,8 +337,7 @@ module _ ( relation-kripke-model i M* ( map-equiv e (class Φ-equivalence x)) ( map-equiv e (class Φ-equivalence y))) - ( hom-Prop ((M , x) ⊨ □ₘ a) - ( (M , y) ⊨ a))))))))) + ( (M , x) ⊨ □ₘ a ⇒ (M , y) ⊨ a)))))))) filtration-relation-upper-bound : equivalence-class Φ-equivalence ≃ type-kripke-model i M* → diff --git a/src/modal-logic/kripke-semantics.lagda.md b/src/modal-logic/kripke-semantics.lagda.md index 34acf2eda1..6ef9aac0ee 100644 --- a/src/modal-logic/kripke-semantics.lagda.md +++ b/src/modal-logic/kripke-semantics.lagda.md @@ -224,9 +224,9 @@ module _ {l1 l2 l3 l4 : Level} {i : Set l3} where - infix 5 _⊨_ - infix 5 _⊨M_ - infix 5 _⊨C_ + infix 6 _⊨_ + infix 6 _⊨M_ + infix 6 _⊨C_ _⊨_ : Σ (kripke-model l1 l2 i l4) (type-kripke-model i) → @@ -234,7 +234,7 @@ module _ Prop (l1 ⊔ l2 ⊔ l4) (M , x) ⊨ varₘ n = raise-Prop (l1 ⊔ l2) (valuate-kripke-model i M n x) (M , x) ⊨ ⊥ₘ = raise-empty-Prop (l1 ⊔ l2 ⊔ l4) - (M , x) ⊨ a →ₘ b = hom-Prop ((M , x) ⊨ a) ((M , x) ⊨ b) + (M , x) ⊨ a →ₘ b = (M , x) ⊨ a ⇒ (M , x) ⊨ b (M , x) ⊨ □ₘ a = Π-Prop ( type-kripke-model i M) From e44f24f60d831943948d61edc6d75c9d45f825cf Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Tue, 7 May 2024 13:29:09 +0300 Subject: [PATCH 56/63] Replace zorn lemma in canonical model theorem --- src/lists/lists-subtypes.lagda.md | 13 ++ .../canonical-model-theorem.lagda.md | 2 +- src/modal-logic/canonical-theories.lagda.md | 2 +- src/modal-logic/completeness-k.lagda.md | 2 +- src/modal-logic/completeness-s5.lagda.md | 2 +- .../finite-approximability.lagda.md | 2 +- src/modal-logic/l-complete-theories.lagda.md | 112 ++++++++++++------ .../l-consistent-theories.lagda.md | 12 ++ src/modal-logic/lindenbaum.lagda.md | 2 +- src/modal-logic/weak-deduction.lagda.md | 4 + 10 files changed, 113 insertions(+), 40 deletions(-) diff --git a/src/lists/lists-subtypes.lagda.md b/src/lists/lists-subtypes.lagda.md index 8c56701be9..ab5e45d5e0 100644 --- a/src/lists/lists-subtypes.lagda.md +++ b/src/lists/lists-subtypes.lagda.md @@ -7,9 +7,11 @@ module lists.lists-subtypes where
Imports ```agda +open import foundation.decidable-types open import foundation.dependent-pair-types open import foundation.disjunction open import foundation.existential-quantification +open import foundation.inhabited-subtypes open import foundation.intersections-subtypes open import foundation.logical-equivalences open import foundation.propositional-truncations @@ -17,6 +19,7 @@ open import foundation.unions-subtypes open import foundation.universe-levels open import foundation-core.cartesian-product-types +open import foundation-core.coproduct-types open import foundation-core.empty-types open import foundation-core.function-types open import foundation-core.negation @@ -66,6 +69,16 @@ module _ {a : A} {l : list A} → is-in-subtype (list-subtype (cons a l)) a head-in-list-subtype {a} {l} = in-list-subtype-in-list (is-head a l) + is-decidable-is-inhabited-list-subtype : + (l : list A) → is-decidable (is-inhabited-subtype (list-subtype l)) + is-decidable-is-inhabited-list-subtype nil = + inr + ( map-universal-property-trunc-Prop + ( empty-Prop) + ( λ (x , in-list) → not-in-list-nil in-list)) + is-decidable-is-inhabited-list-subtype (cons x xs) = + inl (unit-trunc-Prop (x , head-in-list-subtype)) + subset-list-subtype-cons : {a : A} {l : list A} → {l2 : Level} (S : subtype l2 A) → diff --git a/src/modal-logic/canonical-model-theorem.lagda.md b/src/modal-logic/canonical-model-theorem.lagda.md index fd9f12613b..3c243f110a 100644 --- a/src/modal-logic/canonical-model-theorem.lagda.md +++ b/src/modal-logic/canonical-model-theorem.lagda.md @@ -60,7 +60,7 @@ module _ (is-logic : is-modal-logic logic) (is-cons : is-consistent-modal-logic logic) (is-normal : is-normal-modal-logic logic) - (zorn : Zorn-non-empty (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) l2) + (zorn : Zorn (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) l2) (prop-resize : propositional-resizing (l1 ⊔ l2) (lsuc (l1 ⊔ l2))) where diff --git a/src/modal-logic/canonical-theories.lagda.md b/src/modal-logic/canonical-theories.lagda.md index ccc8dcd2fe..2d40e12775 100644 --- a/src/modal-logic/canonical-theories.lagda.md +++ b/src/modal-logic/canonical-theories.lagda.md @@ -48,7 +48,7 @@ TODO module _ {l1 : Level} (i : Set l1) - (zorn : Zorn-non-empty (lsuc l1) l1 l1) + (zorn : Zorn (lsuc l1) l1 l1) (prop-resize : propositional-resizing l1 (lsuc l1)) (logic : modal-theory l1 i) (is-logic : is-modal-logic logic) diff --git a/src/modal-logic/completeness-k.lagda.md b/src/modal-logic/completeness-k.lagda.md index 4f614d5756..c0ae16bff5 100644 --- a/src/modal-logic/completeness-k.lagda.md +++ b/src/modal-logic/completeness-k.lagda.md @@ -58,7 +58,7 @@ module _ {l1 : Level} (i : Set l1) (lem : LEM l1) - (zorn : Zorn-non-empty (lsuc l1) l1 l1) + (zorn : Zorn (lsuc l1) l1 l1) (prop-resize : propositional-resizing l1 (lsuc l1)) where diff --git a/src/modal-logic/completeness-s5.lagda.md b/src/modal-logic/completeness-s5.lagda.md index 758b39623a..de6d6ed443 100644 --- a/src/modal-logic/completeness-s5.lagda.md +++ b/src/modal-logic/completeness-s5.lagda.md @@ -59,7 +59,7 @@ module _ {l1 : Level} (i : Set l1) (lem : LEM l1) - (zorn : Zorn-non-empty (lsuc l1) l1 l1) + (zorn : Zorn (lsuc l1) l1 l1) (prop-resize : propositional-resizing l1 (lsuc l1)) where diff --git a/src/modal-logic/finite-approximability.lagda.md b/src/modal-logic/finite-approximability.lagda.md index a7cdfc4c50..e5aa3d2c23 100644 --- a/src/modal-logic/finite-approximability.lagda.md +++ b/src/modal-logic/finite-approximability.lagda.md @@ -87,7 +87,7 @@ module _ module _ (lem : LEM (lsuc (lsuc l1))) - (zorn : Zorn-non-empty (lsuc l1) l1 l1) + (zorn : Zorn (lsuc l1) l1 l1) (prop-resize : propositional-resizing l1 (lsuc l1)) where diff --git a/src/modal-logic/l-complete-theories.lagda.md b/src/modal-logic/l-complete-theories.lagda.md index 674c47be3b..32a58cd784 100644 --- a/src/modal-logic/l-complete-theories.lagda.md +++ b/src/modal-logic/l-complete-theories.lagda.md @@ -13,6 +13,7 @@ open import foundation.dependent-pair-types open import foundation.disjunction open import foundation.empty-types open import foundation.existential-quantification +open import foundation.inhabited-subtypes open import foundation.inhabited-types open import foundation.law-of-excluded-middle open import foundation.logical-equivalences @@ -394,6 +395,18 @@ module _ chain-union-modal-theory a = ∃ (type-chain-Poset P C) (λ x → modal-theory-chain-element x a) + is-inhabited-chain-is-inhabited-chain-union : + is-inhabited-subtype (chain-union-modal-theory) → + is-inhabited (type-chain-Poset P C) + is-inhabited-chain-is-inhabited-chain-union = + map-universal-property-trunc-Prop + ( is-inhabited-Prop (type-chain-Poset P C)) + ( λ (x , x-in-union) → + ( apply-universal-property-trunc-Prop + ( x-in-union) + ( is-inhabited-Prop (type-chain-Poset P C)) + ( λ (c , _) → unit-trunc-Prop c))) + exists-chain-element-with-formula-Prop : (a : modal-formula i) → Prop (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4) exists-chain-element-with-formula-Prop a = @@ -490,12 +503,6 @@ module _ ( exists-chain-element-with-formula-Prop a) ( λ (x , h-in-x) (y , ha-in-y) → ( elim-disjunction - -- ( leq-Poset-Prop P - -- ( type-Poset-type-chain-Poset P C x) - -- ( type-Poset-type-chain-Poset P C y)) - -- ( leq-Poset-Prop P - -- ( type-Poset-type-chain-Poset P C y) - -- ( type-Poset-type-chain-Poset P C x)) ( exists-chain-element-with-formula-Prop a) ( λ x-leq-y → ( intro-exists y @@ -510,14 +517,45 @@ module _ ( theory-subset-L-union x h h-in-x)))) ( is-chain-Subposet-chain-Poset P C x y))) + in-chain-in-list : + (l : list (modal-formula i)) → + list-subtype l ⊆ chain-union-modal-theory → + {a : modal-formula i} → + ¬ is-in-subtype (weak-modal-logic-closure logic) a → + is-in-subtype (weak-modal-logic-closure (logic ∪ list-subtype l)) a → + exists-chain-element-with-formula a + in-chain-in-list l leq {a} not-in-logic in-union = + in-chain-in-chain-union-assumptions + ( is-inhabited-chain-is-inhabited-chain-union + ( rec-coproduct + ( map-is-inhabited (λ (x , in-list) → x , leq x in-list)) + ( λ not-inh → + ( ex-falso + ( not-in-logic + ( weak-modal-logic-closure-monotic + ( subtype-union-both + ( logic) + ( list-subtype l) + ( logic) + ( refl-leq-subtype logic) + ( λ x in-list → + ( ex-falso + ( not-inh (unit-trunc-Prop (x , in-list)))))) + ( a) + ( in-union))))) + ( is-decidable-is-inhabited-list-subtype l))) + ( l) + ( leq) + ( in-union) + in-chain-in-chain-union : - is-inhabited (type-chain-Poset P C) → {a : modal-formula i} → + ¬ is-in-subtype (weak-modal-logic-closure logic) a → is-in-subtype ( weak-modal-logic-closure (logic ∪ chain-union-modal-theory)) ( a) → exists-chain-element-with-formula a - in-chain-in-chain-union is-inh {a} = + in-chain-in-chain-union {a} not-in-logic = map-universal-property-trunc-Prop ( exists-chain-element-with-formula-Prop a) ( λ d → @@ -529,9 +567,9 @@ module _ ( subset-theory-list-assumptions-weak-deduction d)) ( exists-chain-element-with-formula-Prop a) ( λ ((logic-l , theory-l) , leq-lists , leq-logic , leq-theory) → - ( in-chain-in-chain-union-assumptions is-inh theory-l leq-theory + ( in-chain-in-list theory-l leq-theory not-in-logic ( weak-modal-logic-closure-monotic - {ax₁ = list-subtype logic-l ∪ list-subtype theory-l} + { ax₁ = list-subtype logic-l ∪ list-subtype theory-l} ( subset-union-subset-left ( list-subtype logic-l) ( logic) @@ -540,15 +578,15 @@ module _ ( a) ( weak-modal-logic-closure-monotic leq-lists a ( is-in-weak-modal-logic-closure-weak-deduction - ( is-assumptions-list-assumptions-weak-deduction + (is-assumptions-list-assumptions-weak-deduction ( d))))))))) is-L-consistent-theory-chain-union-modal-theory : - is-inhabited (type-chain-Poset P C) → + is-consistent-modal-logic (weak-modal-logic-closure logic) → is-L-consistent-theory logic chain-union-modal-theory - is-L-consistent-theory-chain-union-modal-theory is-inh in-logic = + is-L-consistent-theory-chain-union-modal-theory is-cons in-logic = apply-universal-property-trunc-Prop - ( in-chain-in-chain-union is-inh in-logic) + ( in-chain-in-chain-union is-cons in-logic) ( empty-Prop) ( λ (x , in-logic') → ( is-L-consistent-theory-modal-theory-L-consistent-theory @@ -586,9 +624,9 @@ module _ is-L-consistent-resized-chain-union-modal-theory : (C : chain-Poset l4 P) → - is-inhabited (type-chain-Poset P C) → + is-consistent-modal-logic (weak-modal-logic-closure logic) → is-L-consistent-theory logic (resized-chain-union-modal-theory C) - is-L-consistent-resized-chain-union-modal-theory C is-inh = + is-L-consistent-resized-chain-union-modal-theory C is-cons = is-L-consistent-antimonotic logic ( resized-chain-union-modal-theory C) ( chain-union-modal-theory C) @@ -599,23 +637,23 @@ module _ ( is-L-consistent-theory-chain-union-modal-theory C ( contains-ax-k) ( contains-ax-s) - ( is-inh)) + ( is-cons)) resized-chain-union-L-consistent-theory : (C : chain-Poset l4 P) → - is-inhabited (type-chain-Poset P C) → + is-consistent-modal-logic (weak-modal-logic-closure logic) → L-consistent-theory logic l3 - pr1 (resized-chain-union-L-consistent-theory C is-inh) = + pr1 (resized-chain-union-L-consistent-theory C is-cons) = resized-chain-union-modal-theory C - pr2 (resized-chain-union-L-consistent-theory C is-inh) = - is-L-consistent-resized-chain-union-modal-theory C is-inh + pr2 (resized-chain-union-L-consistent-theory C is-cons) = + is-L-consistent-resized-chain-union-modal-theory C is-cons union-is-chain-upper-bound : (C : chain-Poset l4 P) → - (is-inh : is-inhabited (type-chain-Poset P C)) → + (is-cons : is-consistent-modal-logic (weak-modal-logic-closure logic)) → is-chain-upper-bound P C - ( resized-chain-union-L-consistent-theory C is-inh) - union-is-chain-upper-bound C is-inh x = + ( resized-chain-union-L-consistent-theory C is-cons) + union-is-chain-upper-bound C _ x = transitive-leq-subtype ( modal-theory-L-consistent-theory logic ( type-Poset-type-chain-Poset @@ -631,23 +669,29 @@ module _ ( λ a in-theory → intro-exists x in-theory) extend-L-consistent-theory : - (zorn : Zorn-non-empty (l1 ⊔ l2 ⊔ lsuc l3) (l1 ⊔ l3) l4) → + Zorn (l1 ⊔ l2 ⊔ lsuc l3) (l1 ⊔ l3) l4 → is-inhabited (L-consistent-theory logic l3) → is-inhabited (L-complete-theory l3) extend-L-consistent-theory zorn is-inh = - is-inhabited-L-complete-exists-complete-L-consistent-theory - ( zorn - ( L-consistent-theories-Poset logic l3) - ( is-inh) - ( λ C C-is-inh → - ( intro-exists - ( resized-chain-union-L-consistent-theory C C-is-inh) - ( union-is-chain-upper-bound C C-is-inh)))) + map-universal-property-trunc-Prop + ( is-inhabited-Prop (L-complete-theory l3)) + ( λ theory → + ( zorn + ( L-consistent-theories-Poset logic l3) + ( λ C → + ( intro-exists + ( resized-chain-union-L-consistent-theory C + ( is-consistent-closure-logic-L-consistent-theory logic + ( theory))) + ( union-is-chain-upper-bound C + ( is-consistent-closure-logic-L-consistent-theory logic + ( theory))))))) + ( is-inh) module _ {l3 : Level} (prop-resize : propositional-resizing (l1 ⊔ l2) (lsuc l1 ⊔ lsuc l2 ⊔ l3)) - (zorn : Zorn-non-empty (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) l3) + (zorn : Zorn (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) l3) (is-logic : is-weak-modal-logic logic) (is-cons : is-consistent-modal-logic logic) (contains-ax-k : ax-k i ⊆ logic) diff --git a/src/modal-logic/l-consistent-theories.lagda.md b/src/modal-logic/l-consistent-theories.lagda.md index 45522f1883..da8b6be749 100644 --- a/src/modal-logic/l-consistent-theories.lagda.md +++ b/src/modal-logic/l-consistent-theories.lagda.md @@ -98,6 +98,18 @@ module _ ( subset-axioms-weak-modal-logic-closure) ( is-consistent-closure-L-consistent-theory theory) + is-consistent-closure-logic-L-consistent-theory : + {l3 : Level} (theory : L-consistent-theory l3) → + is-consistent-modal-logic (weak-modal-logic-closure logic) + is-consistent-closure-logic-L-consistent-theory theory = + is-consistent-modal-logic-antimonotic + ( weak-modal-logic-closure logic) + ( weak-modal-logic-closure + ( logic ∪ modal-theory-L-consistent-theory theory)) + ( weak-modal-logic-closure-monotic + ( subtype-union-left logic (modal-theory-L-consistent-theory theory))) + ( is-L-consistent-theory-modal-theory-L-consistent-theory theory) + is-L-consistent-antimonotic : {l3 l4 : Level} (theory₁ : modal-theory l3 i) → diff --git a/src/modal-logic/lindenbaum.lagda.md b/src/modal-logic/lindenbaum.lagda.md index 4200ca90d7..5bfa7b4a9a 100644 --- a/src/modal-logic/lindenbaum.lagda.md +++ b/src/modal-logic/lindenbaum.lagda.md @@ -42,7 +42,7 @@ module _ (logic : modal-theory l2 i) (contains-ax-k : ax-k i ⊆ logic) (contains-ax-s : ax-s i ⊆ logic) - (zorn : Zorn-non-empty (lsuc (l1 ⊔ l2 ⊔ l3)) (l1 ⊔ l2 ⊔ l3) l3) + (zorn : Zorn (lsuc (l1 ⊔ l2 ⊔ l3)) (l1 ⊔ l2 ⊔ l3) l3) (prop-resize : propositional-resizing (l1 ⊔ l2 ⊔ l3) (lsuc (l1 ⊔ l2 ⊔ l3))) (x@(theory , is-cons) : L-consistent-theory logic (l1 ⊔ l2 ⊔ l3)) where diff --git a/src/modal-logic/weak-deduction.lagda.md b/src/modal-logic/weak-deduction.lagda.md index 6428a55323..2ede344697 100644 --- a/src/modal-logic/weak-deduction.lagda.md +++ b/src/modal-logic/weak-deduction.lagda.md @@ -65,10 +65,14 @@ module _ is-weak-deduction = type-Prop ∘ is-weak-deduction-Prop infix 5 _⊢ₘw_ + infix 5 _⊬ₘw_ _⊢ₘw_ : modal-theory l2 i → modal-formula i → UU (l1 ⊔ l2) axioms ⊢ₘw a = type-subtype (is-weak-deduction-Prop {axioms} {a}) + _⊬ₘw_ : modal-theory l2 i → modal-formula i → UU (l1 ⊔ l2) + axioms ⊬ₘw a = ¬ (axioms ⊢ₘw a) + deduction-weak-deduction : {axioms : modal-theory l2 i} {a : modal-formula i} → axioms ⊢ₘw a → From 081aa5371e519c240999cc0dc56812baf8acc200 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Wed, 8 May 2024 22:57:44 +0300 Subject: [PATCH 57/63] Refactor kripke semantics --- src/modal-logic/axioms.lagda.md | 12 +- .../canonical-model-theorem.lagda.md | 12 +- src/modal-logic/canonical-theories.lagda.md | 6 +- src/modal-logic/completeness-s5.lagda.md | 8 +- src/modal-logic/decision-procedure.lagda.md | 19 ++- src/modal-logic/filtration-lemma.lagda.md | 8 +- .../finite-approximability.lagda.md | 37 ++++- src/modal-logic/formulas.lagda.md | 2 +- .../kripke-models-filtrations.lagda.md | 80 +++++----- src/modal-logic/kripke-semantics.lagda.md | 146 +++++++----------- src/modal-logic/modal-logic-k.lagda.md | 26 ++-- src/modal-logic/modal-logic-s5.lagda.md | 55 +++---- src/modal-logic/soundness.lagda.md | 8 +- src/modal-logic/weak-deduction.lagda.md | 6 +- 14 files changed, 216 insertions(+), 209 deletions(-) diff --git a/src/modal-logic/axioms.lagda.md b/src/modal-logic/axioms.lagda.md index 149ad41c3b..0ba6f82786 100644 --- a/src/modal-logic/axioms.lagda.md +++ b/src/modal-logic/axioms.lagda.md @@ -158,20 +158,20 @@ module _ M in-class x fab fa y r = fab y r (fa y r) - ax-dn-soundness : soundness (ax-dn i) (decidable-models l2 l3 i l4) + ax-dn-soundness : soundness (ax-dn i) (decidable-kripke-models l2 l3 i l4) ax-dn-soundness .(¬¬ₘ a →ₘ a) (a , refl) M is-dec x f with (is-dec a x) ... | inl fa = fa ... | inr fna = raise-ex-falso _ (f (λ fa -> map-raise (fna fa))) - ax-m-soundness : soundness (ax-m i) (reflexive-kripke-class l2 l3 i l4) + ax-m-soundness : soundness (ax-m i) (reflexive-kripke-models l2 l3 i l4) ax-m-soundness .(□ₘ a →ₘ a) (a , refl) M is-refl x fa = fa x (is-refl x) - ax-b-soundness : soundness (ax-b i) (symmetry-kripke-class l2 l3 i l4) + ax-b-soundness : soundness (ax-b i) (symmetry-kripke-models l2 l3 i l4) ax-b-soundness .(a →ₘ □ₘ ◇ₘ a) (a , refl) M is-sym x fa y r contra = contra x (is-sym x y r) fa - ax-d-soundness : soundness (ax-d i) (serial-kripke-class l2 l3 i l4) + ax-d-soundness : soundness (ax-d i) (serial-kripke-models l2 l3 i l4) ax-d-soundness .(□ₘ a →ₘ ◇ₘ a) (a , refl) M is-serial x fa contra = map-raise ( apply-universal-property-trunc-Prop @@ -179,11 +179,11 @@ module _ ( empty-Prop) ( λ (y , r) → map-inv-raise (contra y r (fa y r)))) - ax-4-soundness : soundness (ax-4 i) (transitivity-kripke-class l2 l3 i l4) + ax-4-soundness : soundness (ax-4 i) (transitive-kripke-models l2 l3 i l4) ax-4-soundness .(□ₘ a →ₘ □ₘ □ₘ a) (a , refl) M is-trans x fa y r-xy z r-yz = fa z (is-trans x y z r-yz r-xy) - ax-5-sooundness : soundness (ax-5 i) (euclidean-kripke-class l2 l3 i l4) + ax-5-sooundness : soundness (ax-5 i) (euclidean-kripke-models l2 l3 i l4) ax-5-sooundness .(◇ₘ a →ₘ □ₘ ◇ₘ a) (a , refl) M is-eucl x fa y r-xy contra = fa (λ z r-xz → contra z (is-eucl x y z r-xy r-xz)) ``` diff --git a/src/modal-logic/canonical-model-theorem.lagda.md b/src/modal-logic/canonical-model-theorem.lagda.md index 3c243f110a..f368ab32d8 100644 --- a/src/modal-logic/canonical-model-theorem.lagda.md +++ b/src/modal-logic/canonical-model-theorem.lagda.md @@ -94,9 +94,9 @@ module _ canonical-kripke-model : kripke-model (lsuc l1 ⊔ lsuc l2) (l1 ⊔ l2) i (l1 ⊔ l2) - pr1 (pr1 (pr1 canonical-kripke-model)) = + pr1 (pr1 canonical-kripke-model) = L-complete-theory logic (l1 ⊔ l2) - pr2 (pr1 (pr1 canonical-kripke-model)) = + pr2 (pr1 canonical-kripke-model) = is-inhabited-L-complete-theory ( logic) ( prop-resize) @@ -105,13 +105,13 @@ module _ ( is-cons) ( contains-ax-k) ( contains-ax-s) - pr2 (pr1 canonical-kripke-model) x y = + pr1 (pr2 canonical-kripke-model) x y = Π-Prop ( modal-formula i) ( λ a → ( modal-theory-L-complete-theory logic x (□ₘ a) ⇒ modal-theory-L-complete-theory logic y a)) - pr2 canonical-kripke-model n x = + pr2 (pr2 canonical-kripke-model) n x = modal-theory-L-complete-theory logic x (varₘ n) module _ @@ -480,7 +480,7 @@ module _ (x : L-complete-theory logic (l1 ⊔ l2)) → type-iff-Prop ( modal-theory-L-complete-theory logic x a) - ( (canonical-kripke-model , x) ⊨ a) + ( (canonical-kripke-model , x) ⊨ₘ a) pr1 (canonical-model-theorem-pointwise (varₘ n) x) = map-raise pr1 (canonical-model-theorem-pointwise ⊥ₘ x) = map-raise ∘ is-consistent-modal-theory-L-complete-theory logic x @@ -515,7 +515,7 @@ module _ canonical-model-theorem : (a : modal-formula i) → - type-iff-Prop (logic a) (canonical-kripke-model ⊨M a) + type-iff-Prop (logic a) (canonical-kripke-model ⊨Mₘ a) pr1 (canonical-model-theorem a) in-logic x = forward-implication ( canonical-model-theorem-pointwise a x) diff --git a/src/modal-logic/canonical-theories.lagda.md b/src/modal-logic/canonical-theories.lagda.md index 2d40e12775..1f1a1ce17e 100644 --- a/src/modal-logic/canonical-theories.lagda.md +++ b/src/modal-logic/canonical-theories.lagda.md @@ -87,7 +87,7 @@ module _ is-canonical-ax-m : ax-m i ⊆ logic → is-in-subtype - ( reflexive-kripke-class (lsuc l1) l1 i l1) + ( reflexive-kripke-models (lsuc l1) l1 i l1) ( canonical-kripke-model ( logic) ( is-logic) @@ -107,7 +107,7 @@ module _ LEM l1 → ax-b i ⊆ logic → is-in-subtype - ( symmetry-kripke-class (lsuc l1) l1 i l1) + ( symmetry-kripke-models (lsuc l1) l1 i l1) ( canonical-kripke-model ( logic) ( is-logic) @@ -140,7 +140,7 @@ module _ is-canonical-ax-4 : ax-4 i ⊆ logic → is-in-subtype - ( transitivity-kripke-class (lsuc l1) l1 i l1) + ( transitive-kripke-models (lsuc l1) l1 i l1) ( canonical-kripke-model ( logic) ( is-logic) diff --git a/src/modal-logic/completeness-s5.lagda.md b/src/modal-logic/completeness-s5.lagda.md index de6d6ed443..6be7337d8e 100644 --- a/src/modal-logic/completeness-s5.lagda.md +++ b/src/modal-logic/completeness-s5.lagda.md @@ -65,7 +65,7 @@ module _ S5-canonical-model-is-equivalence : is-in-subtype - ( equivalence-kripke-class (lsuc l1) l1 i l1) + ( equivalence-kripke-models (lsuc l1) l1 i l1) ( canonical-kripke-model ( modal-logic-S5 i) ( is-modal-logic-S5 i) @@ -96,7 +96,9 @@ module _ ( modal-logic-S5-contains-ax-4 i)) completeness-S5 : - completeness (modal-logic-S5 i) (equivalence-kripke-class (lsuc l1) l1 i l1) + completeness + ( modal-logic-S5 i) + ( equivalence-kripke-models (lsuc l1) l1 i l1) completeness-S5 = canonical-model-completness ( modal-logic-S5 i) @@ -106,6 +108,6 @@ module _ ( zorn) ( prop-resize) ( lem) - ( equivalence-kripke-class (lsuc l1) l1 i l1) + ( equivalence-kripke-models (lsuc l1) l1 i l1) ( S5-canonical-model-is-equivalence) ``` diff --git a/src/modal-logic/decision-procedure.lagda.md b/src/modal-logic/decision-procedure.lagda.md index df05d75ed1..aaba16d86e 100644 --- a/src/modal-logic/decision-procedure.lagda.md +++ b/src/modal-logic/decision-procedure.lagda.md @@ -84,19 +84,21 @@ module _ (i : Set l3) (theory : modal-theory l2 i) (C : model-class l1 l2 i l4 l5) - (C-sub-fin : C ⊆ finite-models l1 l2 i l4) + (C-sub-fin : C ⊆ finite-decidable-kripke-models l1 l2 i l4) (C-is-fin : is-finite (type-subtype C)) where decision-procedure' : (a : modal-formula i) → is-decidable - ( (M : type-subtype C) → type-Prop (inclusion-subtype C M ⊨M a)) + ( (M : type-subtype C) → type-Prop (inclusion-subtype C M ⊨Mₘ a)) decision-procedure' a = is-decidable-Π-is-finite ( C-is-fin) ( λ (M , M-in-C) → - ( is-finite-model-valuate-decidable-models i M (C-sub-fin M M-in-C) a)) + ( is-finite-model-valuate-decidable-kripke-models i M + ( C-sub-fin M M-in-C) + ( a))) decision-procedure : (a : modal-formula i) → bool decision-procedure a with decision-procedure' a @@ -443,11 +445,12 @@ module _ (lem : LEM (l2 ⊔ lsuc l3 ⊔ l4 ⊔ lsuc l5 ⊔ lsuc l6 ⊔ lsuc l7 ⊔ lsuc l8)) where - filtration-models-subset-finite-models : - filtration-models l1 l2 i l4 l5 l6 l7 l8 ⊆ finite-models l1 l2 i l4 - filtration-models-subset-finite-models M* = + filtration-models-subset-finite-decidable-kripke-models : + filtration-models l1 l2 i l4 l5 l6 l7 l8 ⊆ + finite-decidable-kripke-models l1 l2 i l4 + filtration-models-subset-finite-decidable-kripke-models M* = map-universal-property-trunc-Prop - ( finite-models l1 l2 i l4 M*) + ( finite-decidable-kripke-models l1 l2 i l4 M*) ( λ ((theory , M) , is-fin , is-filt) → ( triple ( is-finite-equiv @@ -533,7 +536,7 @@ module _ filtrate-soundness logic C₂ H sound a in-logic M* in-class = apply-universal-property-trunc-Prop ( in-class) - ( M* ⊨M a) + ( M* ⊨Mₘ a) ( λ ((b , (M , in-C)) , p) → ( sound a in-logic M* ( tr (is-in-subtype C₂) (inv p) (H (M , in-C) b)))) diff --git a/src/modal-logic/filtration-lemma.lagda.md b/src/modal-logic/filtration-lemma.lagda.md index a33f68d2f4..1ed91e361a 100644 --- a/src/modal-logic/filtration-lemma.lagda.md +++ b/src/modal-logic/filtration-lemma.lagda.md @@ -63,12 +63,12 @@ module _ is-in-subtype theory a → (x : type-kripke-model i M) → type-iff-Prop - ( (M , x) ⊨ a) + ( (M , x) ⊨ₘ a) ( pair ( M*) ( map-equiv-is-kripke-model-filtration i theory M M* is-filtration ( class (Φ-equivalence i theory M) x)) - ⊨ a) + ⊨ₘ a) pr1 (filtration-lemma is-filtration (varₘ n) in-theory x) f = map-raise @@ -104,7 +104,7 @@ module _ ( map-inv-equiv-is-kripke-model-filtration i theory M M* ( is-filtration) ( y*))) - ( (M* , y*) ⊨ a) + ( (M* , y*) ⊨ₘ a) ( λ (y , y-in-class) → ( let y*-id-class = concat @@ -125,7 +125,7 @@ module _ ( is-filtration))) ( y*)) in tr - ( λ z* → type-Prop ((M* , z*) ⊨ a)) + ( λ z* → type-Prop ((M* , z*) ⊨ₘ a)) ( y*-id-class) ( forward-implication ( filtration-lemma is-filtration a diff --git a/src/modal-logic/finite-approximability.lagda.md b/src/modal-logic/finite-approximability.lagda.md index e5aa3d2c23..df29fe1bfc 100644 --- a/src/modal-logic/finite-approximability.lagda.md +++ b/src/modal-logic/finite-approximability.lagda.md @@ -78,6 +78,20 @@ module _ ( soundness logic (finite-subclass i C)) ( completeness logic (finite-subclass i C)))) + is-finitely-approximable-Prop' : + {l2 : Level} (l3 l4 l5 l6 : Level) → + modal-theory l2 i → + Prop (l1 ⊔ l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ lsuc l6) + is-finitely-approximable-Prop' l3 l4 l5 l6 logic = + exists-structure-Prop + ( model-class l3 l4 i l5 l6) + ( λ C → + ( product + ( C ⊆ finite-kripke-models l3 l4 i l5) + ( product + ( soundness logic (finite-subclass i C)) + ( completeness logic (finite-subclass i C))))) + is-finitely-approximable : {l2 : Level} (l3 l4 l5 l6 : Level) → modal-theory l2 i → @@ -146,7 +160,7 @@ module _ ( all-models (lsuc (lsuc l1)) (lsuc l1) i (lsuc l1)) ( λ _ _ → star) ( λ a in-logic M _ x → - ( soundness-K i a in-logic M (λ b y → lem ((M , y) ⊨ b)) x)) + ( soundness-K i a in-logic M (λ b y → lem ((M , y) ⊨ₘ b)) x)) is-finitely-approximable-K : is-finitely-approximable @@ -164,7 +178,11 @@ module _ ( K-finite-class) ( finite-subclass i K-finite-class) ( subtype-intersection-right - ( finite-models (lsuc (lsuc l1)) (lsuc l1) i (lsuc l1)) + ( finite-decidable-kripke-models + ( lsuc (lsuc l1)) + ( lsuc l1) + ( i) + ( lsuc l1)) ( K-finite-class)) ( soundness-K-filtration)) ( transitive-leq-subtype @@ -186,9 +204,18 @@ module _ ( lsuc l1) ( l1) ( l1)) - ( finite-models (lsuc (lsuc l1)) (lsuc l1) i (lsuc l1)) - ( filtration-models-subset-finite-models i l1 - (lsuc l1) l1 l1 lem) + ( finite-decidable-kripke-models + ( lsuc (lsuc l1)) + ( lsuc l1) + ( i) + ( lsuc l1)) + ( filtration-models-subset-finite-decidable-kripke-models + ( i) + ( l1) + ( lsuc l1) + ( l1) + ( l1) + ( lem)) ( K-finite-class-sub-filtration-models) ( M) ( M-in-class)) diff --git a/src/modal-logic/formulas.lagda.md b/src/modal-logic/formulas.lagda.md index 1ae8949ec6..0272cadcc4 100644 --- a/src/modal-logic/formulas.lagda.md +++ b/src/modal-logic/formulas.lagda.md @@ -36,7 +36,7 @@ module _ {l : Level} where - infixr 7 _→ₘ_ + infixr 8 _→ₘ_ infixr 25 □ₘ_ data modal-formula (i : Set l) : UU l where diff --git a/src/modal-logic/kripke-models-filtrations.lagda.md b/src/modal-logic/kripke-models-filtrations.lagda.md index 3fda21311d..49794e093d 100644 --- a/src/modal-logic/kripke-models-filtrations.lagda.md +++ b/src/modal-logic/kripke-models-filtrations.lagda.md @@ -188,7 +188,7 @@ module _ ( λ a → ( function-Prop ( is-in-subtype theory a) - ( iff-Prop ((M , x) ⊨ a) ((M , y) ⊨ a)))) + ( (M , x) ⊨ₘ a ⇔ (M , y) ⊨ₘ a))) pr1 (pr2 Φ-equivalence) x a in-theory = id , id pr1 (pr2 (pr2 Φ-equivalence)) x y r a in-theory = inv-iff (r a in-theory) @@ -202,7 +202,7 @@ module _ map-function-worlds : type-kripke-model i M → type-Set map-function-equivalence-class-Set - map-function-worlds x (a , _) = (M , x) ⊨ a + map-function-worlds x (a , _) = (M , x) ⊨ₘ a map-function-worlds-correct : (x y : type-kripke-model i M) → @@ -211,7 +211,7 @@ module _ map-function-worlds-correct x y s = eq-htpy ( λ (a , a-in-theory) → - ( eq-iff' ((M , x) ⊨ a) ((M , y) ⊨ a) (s a a-in-theory))) + ( eq-iff' ((M , x) ⊨ₘ a) ((M , y) ⊨ₘ a) (s a a-in-theory))) map-function-equivalence-class : equivalence-class Φ-equivalence → @@ -337,7 +337,7 @@ module _ ( relation-kripke-model i M* ( map-equiv e (class Φ-equivalence x)) ( map-equiv e (class Φ-equivalence y))) - ( (M , x) ⊨ □ₘ a ⇒ (M , y) ⊨ a)))))))) + ( (M , x) ⊨ₘ □ₘ a ⇒ (M , y) ⊨ₘ a)))))))) filtration-relation-upper-bound : equivalence-class Φ-equivalence ≃ type-kripke-model i M* → @@ -418,8 +418,8 @@ module _ filtration-relation-preserves-reflexivity : (e : equivalence-class Φ-equivalence ≃ type-kripke-model i M*) → type-Prop (filtration-relation-lower-bound-Prop e) → - is-in-subtype (reflexive-kripke-class l1 l2 i l4) M → - is-in-subtype (reflexive-kripke-class l6 l7 i l8) M* + is-in-subtype (reflexive-kripke-models l1 l2 i l4) M → + is-in-subtype (reflexive-kripke-models l6 l7 i l8) M* filtration-relation-preserves-reflexivity e bound is-refl x* = apply-universal-property-trunc-Prop ( is-inhabited-subtype-equivalence-class Φ-equivalence @@ -433,8 +433,8 @@ module _ filtration-preserves-reflexivity : is-kripke-model-filtration → - is-in-subtype (reflexive-kripke-class l1 l2 i l4) M → - is-in-subtype (reflexive-kripke-class l6 l7 i l8) M* + is-in-subtype (reflexive-kripke-models l1 l2 i l4) M → + is-in-subtype (reflexive-kripke-models l6 l7 i l8) M* filtration-preserves-reflexivity is-filt is-refl class = apply-universal-property-trunc-Prop ( is-inhabited-subtype-equivalence-class Φ-equivalence @@ -489,13 +489,13 @@ module _ ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) ( i) ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) - pr1 (pr1 (pr1 minimal-kripke-model-filtration)) = + pr1 (pr1 minimal-kripke-model-filtration) = equivalence-class Φ-equivalence - pr2 (pr1 (pr1 minimal-kripke-model-filtration)) = - is-inhabited-equivalence-classes pr2 (pr1 minimal-kripke-model-filtration) = + is-inhabited-equivalence-classes + pr1 (pr2 minimal-kripke-model-filtration) = minimal-kripke-model-filtration-relation - pr2 minimal-kripke-model-filtration = + pr2 (pr2 minimal-kripke-model-filtration) = minimal-kripke-model-filtration-valuate minimal-transitive-kripke-model-filtration : @@ -504,14 +504,14 @@ module _ ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) ( i) ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) - pr1 (pr1 (pr1 minimal-transitive-kripke-model-filtration)) = + pr1 (pr1 minimal-transitive-kripke-model-filtration) = equivalence-class Φ-equivalence - pr2 (pr1 (pr1 minimal-transitive-kripke-model-filtration)) = - is-inhabited-equivalence-classes pr2 (pr1 minimal-transitive-kripke-model-filtration) = + is-inhabited-equivalence-classes + pr1 (pr2 minimal-transitive-kripke-model-filtration) = transitive-closure-Prop ( type-Relation-Prop minimal-kripke-model-filtration-relation) - pr2 minimal-transitive-kripke-model-filtration = + pr2 (pr2 minimal-transitive-kripke-model-filtration) = minimal-kripke-model-filtration-valuate module _ @@ -525,12 +525,12 @@ module _ relation-kripke-model i minimal-kripke-model-filtration ( class Φ-equivalence x) ( class Φ-equivalence y) → - type-Prop ((M , x) ⊨ □ₘ a) → - type-Prop ((M , y) ⊨ a) + type-Prop ((M , x) ⊨ₘ □ₘ a) → + type-Prop ((M , y) ⊨ₘ a) proof-upper-bound a box-in-theory x y r-xy x-forces-box = apply-universal-property-trunc-Prop ( r-xy) - ( (M , y) ⊨ a) + ( (M , y) ⊨ₘ a) ( λ ((x' , y') , r-xy' , iff-x , iff-y) → ( backward-implication ( iff-y a (theory-is-closed box-in-theory)) @@ -569,7 +569,7 @@ module _ ( proof-upper-bound)) helper : - is-in-subtype (transitivity-kripke-class l1 l2 i l4) M → + is-in-subtype (transitive-kripke-models l1 l2 i l4) M → (a : modal-formula i) → is-in-subtype theory (□ₘ a) → (x y : type-kripke-model i M) → @@ -577,26 +577,26 @@ module _ ( relation-kripke-model i minimal-kripke-model-filtration) ( class Φ-equivalence x) ( class Φ-equivalence y) → - type-Prop ((M , x) ⊨ □ₘ a) → - type-Prop ((M , y) ⊨ a) + type-Prop ((M , x) ⊨ₘ □ₘ a) → + type-Prop ((M , y) ⊨ₘ a) helper M-is-trans a box-in-theory x y (base* r-xy) x-forces-box = proof-upper-bound a box-in-theory x y r-xy x-forces-box helper M-is-trans a box-in-theory x y (step* {y = z*} r-xz c-zy) x-forces-box = apply-universal-property-trunc-Prop ( is-inhabited-subtype-equivalence-class Φ-equivalence z*) - ( (M , y) ⊨ a) + ( (M , y) ⊨ₘ a) ( λ (z , z-in-z*) → aux z (eq-class-equivalence-class Φ-equivalence z* z-in-z*)) where aux : (z : type-kripke-model i M) → class Φ-equivalence z = z* → - type-Prop ((M , y) ⊨ a) + type-Prop ((M , y) ⊨ₘ a) aux z refl = apply-universal-property-trunc-Prop ( r-xz) - ( (M , y) ⊨ a) + ( (M , y) ⊨ₘ a) ( λ ((x' , z') , r-xz' , iff-x , iff-z) → ( helper M-is-trans a box-in-theory z y c-zy ( backward-implication @@ -610,18 +610,18 @@ module _ ( r-xz'))))) filtration-relation-upper-bound-Prop-minimal-transitive-kripke-model-filtration : - is-in-subtype (transitivity-kripke-class l1 l2 i l4) M → + is-in-subtype (transitive-kripke-models l1 l2 i l4) M → filtration-relation-upper-bound minimal-transitive-kripke-model-filtration id-equiv filtration-relation-upper-bound-Prop-minimal-transitive-kripke-model-filtration M-is-trans a box-in-theory x y r-xy x-forces-box = apply-universal-property-trunc-Prop ( r-xy) - ( (M , y) ⊨ a) + ( (M , y) ⊨ₘ a) ( λ r → helper M-is-trans a box-in-theory x y r x-forces-box) is-kripke-model-filtration-minimal-transitive-kripke-model-filtration : - is-in-subtype (transitivity-kripke-class l1 l2 i l4) M → + is-in-subtype (transitive-kripke-models l1 l2 i l4) M → is-kripke-model-filtration minimal-transitive-kripke-model-filtration is-kripke-model-filtration-minimal-transitive-kripke-model-filtration M-is-trans = @@ -642,9 +642,9 @@ module _ ( M-is-trans))) minimal-filtration-preserves-reflexivity : - is-in-subtype (reflexive-kripke-class l1 l2 i l4) M → + is-in-subtype (reflexive-kripke-models l1 l2 i l4) M → is-in-subtype - ( reflexive-kripke-class + ( reflexive-kripke-models ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) ( i) @@ -656,7 +656,7 @@ module _ ( is-kripke-model-filtration-minimal-kripke-model-filtration) minimal-kripke-model-filtration-relation-preserves-symmetry : - is-in-subtype (symmetry-kripke-class l1 l2 i l4) M → + is-in-subtype (symmetry-kripke-models l1 l2 i l4) M → is-symmetric-Relation-Prop minimal-kripke-model-filtration-relation minimal-kripke-model-filtration-relation-preserves-symmetry is-sym x* y* = map-universal-property-trunc-Prop @@ -665,9 +665,9 @@ module _ ( intro-exists (y , x) (is-sym x y r-xy , y-in-y* , x-in-x*))) minimal-filtration-preserves-symmetry : - is-in-subtype (symmetry-kripke-class l1 l2 i l4) M → + is-in-subtype (symmetry-kripke-models l1 l2 i l4) M → is-in-subtype - ( symmetry-kripke-class + ( symmetry-kripke-models ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) ( i) @@ -677,9 +677,9 @@ module _ minimal-kripke-model-filtration-relation-preserves-symmetry minimal-transitive-filtration-preserves-reflexivity : - is-in-subtype (reflexive-kripke-class l1 l2 i l4) M → + is-in-subtype (reflexive-kripke-models l1 l2 i l4) M → is-in-subtype - ( reflexive-kripke-class + ( reflexive-kripke-models ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) ( i) @@ -691,9 +691,9 @@ module _ ( minimal-filtration-preserves-reflexivity is-refl) minimal-transitive-kripke-model-filtration-preserves-symmetry : - is-in-subtype (symmetry-kripke-class l1 l2 i l4) M → + is-in-subtype (symmetry-kripke-models l1 l2 i l4) M → is-in-subtype - ( symmetry-kripke-class + ( symmetry-kripke-models ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) ( i) @@ -706,7 +706,7 @@ module _ minimal-transitive-kripke-model-filtration-is-transitive : is-in-subtype - ( transitivity-kripke-class + ( transitive-kripke-models ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) ( i) @@ -717,9 +717,9 @@ module _ ( type-Relation-Prop minimal-kripke-model-filtration-relation) minimal-transitive-filtration-preserves-equivalence : - is-in-subtype (equivalence-kripke-class l1 l2 i l4) M → + is-in-subtype (equivalence-kripke-models l1 l2 i l4) M → is-in-subtype - ( equivalence-kripke-class + ( equivalence-kripke-models ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) ( i) diff --git a/src/modal-logic/kripke-semantics.lagda.md b/src/modal-logic/kripke-semantics.lagda.md index 6ef9aac0ee..24fb2eea9b 100644 --- a/src/modal-logic/kripke-semantics.lagda.md +++ b/src/modal-logic/kripke-semantics.lagda.md @@ -51,76 +51,44 @@ TODO ### Semantics ```agda -module _ - (l1 l2 : Level) - where - - kripke-frame : UU (lsuc l1 ⊔ lsuc l2) - kripke-frame = Σ (Inhabited-Type l1) (Relation-Prop l2 ∘ type-Inhabited-Type) - -module _ - {l1 l2 : Level} - where - - Inhabited-Type-kripke-frame : kripke-frame l1 l2 → Inhabited-Type l1 - Inhabited-Type-kripke-frame = pr1 - - type-kripke-frame : kripke-frame l1 l2 → UU l1 - type-kripke-frame = type-Inhabited-Type ∘ Inhabited-Type-kripke-frame - - is-inhabited-type-kripke-frame : - (F : kripke-frame l1 l2) → is-inhabited (type-kripke-frame F) - is-inhabited-type-kripke-frame = - is-inhabited-type-Inhabited-Type ∘ Inhabited-Type-kripke-frame - - relation-Prop-kripke-frame : - (F : kripke-frame l1 l2) → Relation-Prop l2 (type-kripke-frame F) - relation-Prop-kripke-frame = pr2 - - relation-kripke-frame : - (F : kripke-frame l1 l2) → Relation l2 (type-kripke-frame F) - relation-kripke-frame = type-Relation-Prop ∘ relation-Prop-kripke-frame - module _ (l1 l2 : Level) {l3 : Level} (i : Set l3) (l4 : Level) where kripke-model : UU (lsuc l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4) kripke-model = - Σ (kripke-frame l1 l2) (λ F → type-Set i → type-kripke-frame F → Prop l4) + Σ ( Inhabited-Type l1) + ( λ w → + ( product + ( Relation-Prop l2 (type-Inhabited-Type w)) + ( type-Set i → type-Inhabited-Type w → Prop l4))) module _ {l1 l2 l3 l4 : Level} (i : Set l3) where - kripke-frame-kripke-model : kripke-model l1 l2 i l4 → kripke-frame l1 l2 - kripke-frame-kripke-model = pr1 - Inhabited-Type-kripke-model : kripke-model l1 l2 i l4 → Inhabited-Type l1 - Inhabited-Type-kripke-model = - Inhabited-Type-kripke-frame ∘ kripke-frame-kripke-model + Inhabited-Type-kripke-model = pr1 type-kripke-model : kripke-model l1 l2 i l4 → UU l1 - type-kripke-model = type-kripke-frame ∘ kripke-frame-kripke-model + type-kripke-model = type-Inhabited-Type ∘ Inhabited-Type-kripke-model is-inhabited-type-kripke-model : (M : kripke-model l1 l2 i l4) → is-inhabited (type-kripke-model M) is-inhabited-type-kripke-model = - is-inhabited-type-kripke-frame ∘ kripke-frame-kripke-model + is-inhabited-type-Inhabited-Type ∘ Inhabited-Type-kripke-model relation-Prop-kripke-model : (M : kripke-model l1 l2 i l4) → Relation-Prop l2 (type-kripke-model M) - relation-Prop-kripke-model = - relation-Prop-kripke-frame ∘ kripke-frame-kripke-model + relation-Prop-kripke-model = pr1 ∘ pr2 relation-kripke-model : (M : kripke-model l1 l2 i l4) → Relation l2 (type-kripke-model M) - relation-kripke-model = - relation-kripke-frame ∘ kripke-frame-kripke-model + relation-kripke-model = type-Relation-Prop ∘ relation-Prop-kripke-model valuate-kripke-model : (M : kripke-model l1 l2 i l4) → type-Set i → type-kripke-model M → Prop l4 - valuate-kripke-model = pr2 + valuate-kripke-model = pr2 ∘ pr2 module _ (l1 l2 : Level) {l3 : Level} (i : Set l3) (l4 : Level) @@ -184,80 +152,80 @@ module _ relation-property-class property M = property (relation-Prop-kripke-model i M) - reflexive-kripke-class : model-class l1 l2 i l4 (l1 ⊔ l2) - reflexive-kripke-class = + reflexive-kripke-models : model-class l1 l2 i l4 (l1 ⊔ l2) + reflexive-kripke-models = relation-property-class ( λ x → ( is-reflexive-Relation-Prop x , is-prop-is-reflexive-Relation-Prop x)) - symmetry-kripke-class : model-class l1 l2 i l4 (l1 ⊔ l2) - symmetry-kripke-class = + symmetry-kripke-models : model-class l1 l2 i l4 (l1 ⊔ l2) + symmetry-kripke-models = relation-property-class ( λ x → ( is-symmetric-Relation-Prop x , is-prop-is-symmetric-Relation-Prop x)) - transitivity-kripke-class : model-class l1 l2 i l4 (l1 ⊔ l2) - transitivity-kripke-class = + transitive-kripke-models : model-class l1 l2 i l4 (l1 ⊔ l2) + transitive-kripke-models = relation-property-class ( λ x → ( pair ( is-transitive-Relation-Prop x) ( is-prop-is-transitive-Relation-Prop x))) - serial-kripke-class : model-class l1 l2 i l4 (l1 ⊔ l2) - serial-kripke-class = + serial-kripke-models : model-class l1 l2 i l4 (l1 ⊔ l2) + serial-kripke-models = relation-property-class ( λ x → ( is-serial-Relation-Prop x , is-prop-is-serial-Relation-Prop x)) - euclidean-kripke-class : model-class l1 l2 i l4 (l1 ⊔ l2) - euclidean-kripke-class = + euclidean-kripke-models : model-class l1 l2 i l4 (l1 ⊔ l2) + euclidean-kripke-models = relation-property-class ( λ x → ( is-euclidean-Relation-Prop x , is-prop-is-euclidean-Relation-Prop x)) - equivalence-kripke-class : model-class l1 l2 i l4 (l1 ⊔ l2) - equivalence-kripke-class = + equivalence-kripke-models : model-class l1 l2 i l4 (l1 ⊔ l2) + equivalence-kripke-models = relation-property-class is-equivalence-relation-Prop module _ {l1 l2 l3 l4 : Level} {i : Set l3} where - infix 6 _⊨_ - infix 6 _⊨M_ - infix 6 _⊨C_ + infix 7 _⊨ₘ_ + infix 7 _⊨Mₘ_ + infix 7 _⊨Cₘ_ - _⊨_ : + _⊨ₘ_ : Σ (kripke-model l1 l2 i l4) (type-kripke-model i) → modal-formula i → Prop (l1 ⊔ l2 ⊔ l4) - (M , x) ⊨ varₘ n = raise-Prop (l1 ⊔ l2) (valuate-kripke-model i M n x) - (M , x) ⊨ ⊥ₘ = raise-empty-Prop (l1 ⊔ l2 ⊔ l4) - (M , x) ⊨ a →ₘ b = (M , x) ⊨ a ⇒ (M , x) ⊨ b - (M , x) ⊨ □ₘ a = + (M , x) ⊨ₘ varₘ n = raise-Prop (l1 ⊔ l2) (valuate-kripke-model i M n x) + (M , x) ⊨ₘ ⊥ₘ = raise-empty-Prop (l1 ⊔ l2 ⊔ l4) + (M , x) ⊨ₘ a →ₘ b = (M , x) ⊨ₘ a ⇒ (M , x) ⊨ₘ b + (M , x) ⊨ₘ □ₘ a = Π-Prop ( type-kripke-model i M) - ( λ y → function-Prop (relation-kripke-model i M x y) ((M , y) ⊨ a)) + ( λ y → function-Prop (relation-kripke-model i M x y) ((M , y) ⊨ₘ a)) - _⊨M_ : kripke-model l1 l2 i l4 → modal-formula i → Prop (l1 ⊔ l2 ⊔ l4) - M ⊨M a = Π-Prop (type-kripke-model i M) (λ x → (M , x) ⊨ a) + _⊨Mₘ_ : kripke-model l1 l2 i l4 → modal-formula i → Prop (l1 ⊔ l2 ⊔ l4) + M ⊨Mₘ a = Π-Prop (type-kripke-model i M) (λ x → (M , x) ⊨ₘ a) - _⊨C_ : + _⊨Cₘ_ : {l5 : Level} → model-class l1 l2 i l4 l5 → modal-formula i → Prop (lsuc l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ l5) - C ⊨C a = + C ⊨Cₘ a = Π-Prop ( kripke-model l1 l2 i l4) - ( λ M → function-Prop (is-in-subtype C M) (M ⊨M a)) + ( λ M → function-Prop (is-in-subtype C M) (M ⊨Mₘ a)) class-modal-logic : {l5 : Level} → model-class l1 l2 i l4 l5 → modal-theory (lsuc l1 ⊔ lsuc l2 ⊔ l3 ⊔ lsuc l4 ⊔ l5) i - class-modal-logic = _⊨C_ + class-modal-logic = _⊨Cₘ_ -- TODO: rename class-modal-logic-monotic : @@ -275,18 +243,20 @@ module _ (l4 : Level) where - decidable-models : model-class l1 l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4) - decidable-models M = + decidable-kripke-models : model-class l1 l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4) + decidable-kripke-models M = Π-Prop ( modal-formula i) ( λ a → ( Π-Prop ( type-kripke-model i M) - ( λ x → is-decidable-Prop ((M , x) ⊨ a)))) + ( λ x → is-decidable-Prop ((M , x) ⊨ₘ a)))) + + finite-kripke-models : model-class l1 l2 i l4 l1 + finite-kripke-models = is-finite-Prop ∘ type-kripke-model i - -- TODO: maybe dicidable-finite? - finite-models : model-class l1 l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4) - finite-models M = + finite-decidable-kripke-models : model-class l1 l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4) + finite-decidable-kripke-models M = product-Prop ( is-finite-Prop (type-kripke-model i M)) ( product-Prop @@ -307,13 +277,15 @@ module _ {l1 l2 l3 l4 : Level} (i : Set l3) where - finite-models-subclass-decidable-models : - finite-models l1 l2 i l4 ⊆ decidable-models l1 l2 i l4 - finite-models-subclass-decidable-models M (w-is-fin , dec-r , dec-v) = lemma + finite-decidable-subclass-decidable-models : + finite-decidable-kripke-models l1 l2 i l4 ⊆ + decidable-kripke-models l1 l2 i l4 + finite-decidable-subclass-decidable-models M (w-is-fin , dec-r , dec-v) = + lemma where lemma : (a : modal-formula i) (x : type-kripke-model i M) → - is-decidable (type-Prop ((M , x) ⊨ a)) + is-decidable (type-Prop ((M , x) ⊨ₘ a)) lemma (varₘ n) x = is-decidable-raise (l1 ⊔ l2) _ (dec-v x n) lemma ⊥ₘ x = @@ -325,25 +297,25 @@ module _ ( w-is-fin) ( λ y → is-decidable-function-type (dec-r x y) (lemma a y)) - is-finite-model-valuate-decidable-models : + is-finite-model-valuate-decidable-kripke-models : (M : kripke-model l1 l2 i l4) → - is-in-subtype (finite-models l1 l2 i l4) M → + is-in-subtype (finite-decidable-kripke-models l1 l2 i l4) M → (a : modal-formula i) → - is-decidable (type-Prop (M ⊨M a)) - is-finite-model-valuate-decidable-models M sub-fin a = + is-decidable (type-Prop (M ⊨Mₘ a)) + is-finite-model-valuate-decidable-kripke-models M sub-fin a = is-decidable-Π-is-finite ( pr1 (sub-fin)) - ( finite-models-subclass-decidable-models M sub-fin a) + ( finite-decidable-subclass-decidable-models M sub-fin a) decidable-subclass : {l5 : Level} → model-class l1 l2 i l4 l5 → model-class l1 l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) - decidable-subclass C = (decidable-models l1 l2 i l4) ∩ C + decidable-subclass C = (decidable-kripke-models l1 l2 i l4) ∩ C finite-subclass : {l5 : Level} → model-class l1 l2 i l4 l5 → model-class l1 l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) - finite-subclass C = (finite-models l1 l2 i l4) ∩ C + finite-subclass C = (finite-decidable-kripke-models l1 l2 i l4) ∩ C ``` diff --git a/src/modal-logic/modal-logic-k.lagda.md b/src/modal-logic/modal-logic-k.lagda.md index 87bdd294df..fcaf873cdb 100644 --- a/src/modal-logic/modal-logic-k.lagda.md +++ b/src/modal-logic/modal-logic-k.lagda.md @@ -145,14 +145,14 @@ module _ where soundness-K-axioms : - soundness (modal-logic-K-axioms i) (decidable-models l2 l3 i l4) + soundness (modal-logic-K-axioms i) (decidable-kripke-models l2 l3 i l4) soundness-K-axioms = soundness-union-subclass-right-sublevels ( (ax-k i ∪ ax-s i) ∪ ax-n i) ( ax-dn i) ( l1 ⊔ l2 ⊔ l3 ⊔ l4) ( all-models l2 l3 i l4) - ( decidable-models l2 l3 i l4) + ( decidable-kripke-models l2 l3 i l4) ( soundness-union-same-class ( ax-k i ∪ ax-s i) ( ax-n i) @@ -165,19 +165,22 @@ module _ ( ax-s-soundness i l3 l4)) ( ax-n-soundness i l3 l4)) ( ax-dn-soundness i l3 l4) - ( all-models-is-largest-class i (decidable-models l2 l3 i l4)) + ( all-models-is-largest-class i (decidable-kripke-models l2 l3 i l4)) - soundness-K : soundness (modal-logic-K i) (decidable-models l2 l3 i l4) + soundness-K : soundness (modal-logic-K i) (decidable-kripke-models l2 l3 i l4) soundness-K = - soundness-modal-logic (decidable-models l2 l3 i l4) soundness-K-axioms + soundness-modal-logic + ( decidable-kripke-models l2 l3 i l4) + ( soundness-K-axioms) - soundness-K-finite : soundness (modal-logic-K i) (finite-models l2 l3 i l4) + soundness-K-finite : + soundness (modal-logic-K i) (finite-decidable-kripke-models l2 l3 i l4) soundness-K-finite = soundness-subclass ( modal-logic-K i) - ( decidable-models l2 l3 i l4) - ( finite-models l2 l3 i l4) - ( finite-models-subclass-decidable-models i) + ( decidable-kripke-models l2 l3 i l4) + ( finite-decidable-kripke-models l2 l3 i l4) + ( finite-decidable-subclass-decidable-models i) ( soundness-K) module _ @@ -191,8 +194,9 @@ module _ ( i) ( ⊥ₘ) ( bot-in-logic) - ( pair - ( pair (unit , unit-trunc-Prop star) (λ _ _ → empty-Prop)) + ( triple + ( unit , unit-trunc-Prop star) + ( λ _ _ → empty-Prop) ( λ _ _ → empty-Prop)) ( triple ( is-finite-unit) diff --git a/src/modal-logic/modal-logic-s5.lagda.md b/src/modal-logic/modal-logic-s5.lagda.md index fc9fecbeed..0e83448786 100644 --- a/src/modal-logic/modal-logic-s5.lagda.md +++ b/src/modal-logic/modal-logic-s5.lagda.md @@ -137,68 +137,70 @@ module _ where soundness-S5-additional-axioms : - soundness (ax-m i ∪ (ax-b i ∪ ax-4 i)) (equivalence-kripke-class l2 l3 i l4) + soundness + ( ax-m i ∪ ax-b i ∪ ax-4 i) + (equivalence-kripke-models l2 l3 i l4) soundness-S5-additional-axioms = soundness-union ( ax-m i) ( ax-b i ∪ ax-4 i) - ( reflexive-kripke-class l2 l3 i l4) - ( symmetry-kripke-class l2 l3 i l4 ∩ transitivity-kripke-class l2 l3 i l4) + ( reflexive-kripke-models l2 l3 i l4) + ( symmetry-kripke-models l2 l3 i l4 ∩ transitive-kripke-models l2 l3 i l4) ( ax-m-soundness i l3 l4) ( soundness-union ( ax-b i) ( ax-4 i) - ( symmetry-kripke-class l2 l3 i l4) - ( transitivity-kripke-class l2 l3 i l4) + ( symmetry-kripke-models l2 l3 i l4) + ( transitive-kripke-models l2 l3 i l4) ( ax-b-soundness i l3 l4) ( ax-4-soundness i l3 l4)) soundness-S5-axioms : soundness ( modal-logic-S5-axioms i) - ( decidable-subclass i (equivalence-kripke-class l2 l3 i l4)) + ( decidable-subclass i (equivalence-kripke-models l2 l3 i l4)) soundness-S5-axioms = soundness-union ( modal-logic-K-axioms i) ( ax-m i ∪ (ax-b i ∪ ax-4 i)) - ( decidable-models l2 l3 i l4) - ( equivalence-kripke-class l2 l3 i l4) + ( decidable-kripke-models l2 l3 i l4) + ( equivalence-kripke-models l2 l3 i l4) ( soundness-K-axioms i) ( soundness-S5-additional-axioms) soundness-S5 : soundness ( modal-logic-S5 i) - ( decidable-subclass i (equivalence-kripke-class l2 l3 i l4)) + ( decidable-subclass i (equivalence-kripke-models l2 l3 i l4)) soundness-S5 = soundness-modal-logic - ( decidable-subclass i (equivalence-kripke-class l2 l3 i l4)) + ( decidable-subclass i (equivalence-kripke-models l2 l3 i l4)) ( soundness-S5-axioms) soundness-S5-finite : soundness ( modal-logic-S5 i) - ( finite-subclass i (equivalence-kripke-class l2 l3 i l4)) + ( finite-subclass i (equivalence-kripke-models l2 l3 i l4)) soundness-S5-finite = soundness-subclass ( modal-logic-S5 i) - ( decidable-subclass i (equivalence-kripke-class l2 l3 i l4)) - ( finite-subclass i (equivalence-kripke-class l2 l3 i l4)) + ( decidable-subclass i (equivalence-kripke-models l2 l3 i l4)) + ( finite-subclass i (equivalence-kripke-models l2 l3 i l4)) ( subtype-both-intersection - ( decidable-models l2 l3 i l4) - ( equivalence-kripke-class l2 l3 i l4) - ( finite-subclass i (equivalence-kripke-class l2 l3 i l4)) + ( decidable-kripke-models l2 l3 i l4) + ( equivalence-kripke-models l2 l3 i l4) + ( finite-subclass i (equivalence-kripke-models l2 l3 i l4)) ( transitive-leq-subtype - ( finite-subclass i (equivalence-kripke-class l2 l3 i l4)) - ( finite-models l2 l3 i l4) - ( decidable-models l2 l3 i l4) - ( finite-models-subclass-decidable-models i) + ( finite-subclass i (equivalence-kripke-models l2 l3 i l4)) + ( finite-decidable-kripke-models l2 l3 i l4) + ( decidable-kripke-models l2 l3 i l4) + ( finite-decidable-subclass-decidable-models i) ( subtype-intersection-left - ( finite-models l2 l3 i l4) - ( equivalence-kripke-class l2 l3 i l4))) + ( finite-decidable-kripke-models l2 l3 i l4) + ( equivalence-kripke-models l2 l3 i l4))) ( subtype-intersection-right - ( finite-models l2 l3 i l4) - ( equivalence-kripke-class l2 l3 i l4))) + ( finite-decidable-kripke-models l2 l3 i l4) + ( equivalence-kripke-models l2 l3 i l4))) ( soundness-S5) module _ @@ -212,8 +214,9 @@ module _ ( i) ( ⊥ₘ) ( bot-in-logic) - ( pair - ( pair (unit , unit-trunc-Prop star) (λ _ _ → unit-Prop)) + ( triple + ( unit , unit-trunc-Prop star) + ( λ _ _ → unit-Prop) ( λ _ _ → unit-Prop)) ( pair ( triple diff --git a/src/modal-logic/soundness.lagda.md b/src/modal-logic/soundness.lagda.md index f583a78c93..105a142e77 100644 --- a/src/modal-logic/soundness.lagda.md +++ b/src/modal-logic/soundness.lagda.md @@ -55,7 +55,7 @@ module _ soundness axioms C → {a : modal-formula i} → axioms ⊢ₘ a → - type-Prop (C ⊨C a) + type-Prop (C ⊨Cₘ a) soundness-axioms H (modal-ax x) = H _ x soundness-axioms H (modal-mp dab da) M in-class x = soundness-axioms H dab M in-class x (soundness-axioms H da M in-class x) @@ -65,7 +65,7 @@ module _ soundness-modal-logic : soundness axioms C → soundness (modal-logic-closure axioms) C soundness-modal-logic H a = - map-universal-property-trunc-Prop (C ⊨C a) (soundness-axioms H) + map-universal-property-trunc-Prop (C ⊨Cₘ a) (soundness-axioms H) module _ {l1 l2 l3 l4 l5 l6 l7 : Level} @@ -93,7 +93,7 @@ module _ is-in-subtype (C₁ ∩ C₂) M → (a : modal-formula i) → is-in-subtype theory₁ a + is-in-subtype theory₂ a → - type-Prop (M ⊨M a) + type-Prop (M ⊨Mₘ a) forces-in-intersection M in-class a (inl d) = sound₁ a d M (subtype-intersection-left C₁ C₂ M in-class) forces-in-intersection M in-class a (inr d) = @@ -103,7 +103,7 @@ module _ soundness-union a is-theory M in-class = apply-universal-property-trunc-Prop ( is-theory) - ( M ⊨M a) + ( M ⊨Mₘ a) ( forces-in-intersection M in-class a) soundness-modal-logic-union : diff --git a/src/modal-logic/weak-deduction.lagda.md b/src/modal-logic/weak-deduction.lagda.md index 2ede344697..72df82f3a4 100644 --- a/src/modal-logic/weak-deduction.lagda.md +++ b/src/modal-logic/weak-deduction.lagda.md @@ -57,7 +57,7 @@ module _ Prop lzero is-weak-deduction-Prop (modal-ax x) = unit-Prop is-weak-deduction-Prop (modal-mp dab da) = - conjunction-Prop (is-weak-deduction-Prop dab) (is-weak-deduction-Prop da) + is-weak-deduction-Prop dab ∧ is-weak-deduction-Prop da is-weak-deduction-Prop (modal-nec d) = empty-Prop is-weak-deduction : @@ -65,14 +65,10 @@ module _ is-weak-deduction = type-Prop ∘ is-weak-deduction-Prop infix 5 _⊢ₘw_ - infix 5 _⊬ₘw_ _⊢ₘw_ : modal-theory l2 i → modal-formula i → UU (l1 ⊔ l2) axioms ⊢ₘw a = type-subtype (is-weak-deduction-Prop {axioms} {a}) - _⊬ₘw_ : modal-theory l2 i → modal-formula i → UU (l1 ⊔ l2) - axioms ⊬ₘw a = ¬ (axioms ⊢ₘw a) - deduction-weak-deduction : {axioms : modal-theory l2 i} {a : modal-formula i} → axioms ⊢ₘw a → From 0df9fd48bf29b28cecc7495becac292d752ee643 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Thu, 9 May 2024 04:24:18 +0300 Subject: [PATCH 58/63] Fix filtration definition --- src/modal-logic/filtration-lemma.lagda.md | 17 +++++----- .../kripke-models-filtrations.lagda.md | 33 ++++++++++--------- src/modal-logic/l-complete-theories.lagda.md | 4 +-- 3 files changed, 28 insertions(+), 26 deletions(-) diff --git a/src/modal-logic/filtration-lemma.lagda.md b/src/modal-logic/filtration-lemma.lagda.md index 1ed91e361a..4034315e63 100644 --- a/src/modal-logic/filtration-lemma.lagda.md +++ b/src/modal-logic/filtration-lemma.lagda.md @@ -80,8 +80,9 @@ module _ ( M*) ( is-filtration) ( n) + ( in-theory) ( x)) - ( in-theory , map-inv-raise f)) + ( map-inv-raise f)) pr1 (filtration-lemma is-filtration ⊥ₘ in-theory x) = map-raise ∘ map-inv-raise pr1 (filtration-lemma is-filtration (a →ₘ b) in-theory x) @@ -152,13 +153,13 @@ module _ pr2 (filtration-lemma is-filtration (varₘ n) in-theory x) f = map-raise - ( pr2 - ( backward-implication - ( is-filtration-valuate-is-kripke-model-filtration i theory M M* - ( is-filtration) - ( n) - ( x)) - ( map-inv-raise f))) + ( backward-implication + ( is-filtration-valuate-is-kripke-model-filtration i theory M M* + ( is-filtration) + ( n) + ( in-theory) + ( x)) + ( map-inv-raise f)) pr2 (filtration-lemma is-filtration ⊥ₘ in-theory x) = map-raise ∘ map-inv-raise pr2 (filtration-lemma is-filtration (a →ₘ b) in-theory x) diff --git a/src/modal-logic/kripke-models-filtrations.lagda.md b/src/modal-logic/kripke-models-filtrations.lagda.md index 49794e093d..399c7c9c26 100644 --- a/src/modal-logic/kripke-models-filtrations.lagda.md +++ b/src/modal-logic/kripke-models-filtrations.lagda.md @@ -292,11 +292,12 @@ module _ is-filtration-valuate-Prop e = Π-Prop (type-Set i) ( λ n → - ( Π-Prop (type-kripke-model i M) - ( λ x → iff-Prop - ( product-Prop (theory (varₘ n)) (valuate-kripke-model i M n x)) - ( valuate-kripke-model i M* n - ( map-equiv e (class Φ-equivalence x)))))) + ( theory (varₘ n) ⇒ + ( Π-Prop (type-kripke-model i M) + ( λ x → iff-Prop + ( valuate-kripke-model i M n x) + ( valuate-kripke-model i M* n + ( map-equiv e (class Φ-equivalence x))))))) is-filtration-valuate : equivalence-class Φ-equivalence ≃ type-kripke-model i M* → @@ -555,16 +556,17 @@ module _ is-kripke-model-filtration-minimal-kripke-model-filtration = pair id-equiv ( triple - ( λ n x → + ( λ n in-theory x → ( pair - ( λ (in-theory , val-n) → - ( pair in-theory + ( λ val-n → + ( pair + ( in-theory) ( λ y eq-xy → ( map-inv-raise ( forward-implication ( eq-xy (varₘ n) in-theory) - ( map-raise val-n))))))) - ( λ (in-theory , val-n) → in-theory , val-n x (λ _ _ → id , id))) + ( map-raise val-n)))))) + ( λ (in-theory , val-n) → val-n x (λ _ _ → id , id)))) ( proof-lower-bound) ( proof-upper-bound)) @@ -627,16 +629,17 @@ module _ M-is-trans = pair id-equiv ( triple - ( λ n x → + ( λ n in-theory x → ( pair - ( λ (in-theory , val-n) → - ( pair in-theory + ( λ val-n → + ( pair + ( in-theory) ( λ y eq-xy → ( map-inv-raise ( forward-implication ( eq-xy (varₘ n) in-theory) - ( map-raise val-n))))))) - ( λ (in-theory , val-n) → in-theory , val-n x (λ _ _ → id , id))) + ( map-raise val-n)))))) + ( λ (in-theory , val-n) → val-n x (λ _ _ → id , id)))) ( λ x y r → unit-trunc-Prop (base* (proof-lower-bound x y r))) ( filtration-relation-upper-bound-Prop-minimal-transitive-kripke-model-filtration ( M-is-trans))) diff --git a/src/modal-logic/l-complete-theories.lagda.md b/src/modal-logic/l-complete-theories.lagda.md index 32a58cd784..4ceb90b95c 100644 --- a/src/modal-logic/l-complete-theories.lagda.md +++ b/src/modal-logic/l-complete-theories.lagda.md @@ -603,9 +603,7 @@ module _ P : Poset (l1 ⊔ l2 ⊔ lsuc l3) (l1 ⊔ l3) P = L-consistent-theories-Poset logic l3 - resized-chain-union-modal-theory : - chain-Poset l4 P → - modal-theory l3 i + resized-chain-union-modal-theory : chain-Poset l4 P → modal-theory l3 i resized-chain-union-modal-theory C = resize prop-resize ∘ chain-union-modal-theory C From 355578452f8b271f6d75e25f08f5a25a1b021e65 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Sat, 11 May 2024 20:52:38 +0300 Subject: [PATCH 59/63] Complete eliminators for equivalence class --- src/foundation/equivalence-classes.lagda.md | 266 ++++++++---------- src/lists/lists-subtypes.lagda.md | 5 +- src/modal-logic/axioms.lagda.md | 2 +- src/modal-logic/canonical-theories.lagda.md | 15 +- src/modal-logic/decision-procedure.lagda.md | 8 +- src/modal-logic/deduction.lagda.md | 35 ++- src/modal-logic/formulas.lagda.md | 33 ++- .../kripke-models-filtrations.lagda.md | 4 +- 8 files changed, 187 insertions(+), 181 deletions(-) diff --git a/src/foundation/equivalence-classes.lagda.md b/src/foundation/equivalence-classes.lagda.md index a077a35775..24a76700ba 100644 --- a/src/foundation/equivalence-classes.lagda.md +++ b/src/foundation/equivalence-classes.lagda.md @@ -10,6 +10,7 @@ module foundation.equivalence-classes where open import foundation.action-on-identifications-functions open import foundation.conjunction open import foundation.constant-type-families +open import foundation.dependent-identifications open import foundation.dependent-pair-types open import foundation.effective-maps-equivalence-relations open import foundation.equality-dependent-pair-types @@ -33,7 +34,6 @@ open import foundation.universal-property-image open import foundation.universe-levels open import foundation-core.cartesian-product-types -open import foundation-core.dependent-identifications open import foundation-core.embeddings open import foundation-core.equivalence-relations open import foundation-core.equivalences @@ -144,6 +144,9 @@ module _ is-set-equivalence-class = is-set-type-subtype is-equivalence-class-Prop is-set-subtype + is-in-self-equivalence-class : (a : A) → is-in-equivalence-class (class a) a + is-in-self-equivalence-class a = refl-equivalence-relation R a + equivalence-class-Set : Set (l1 ⊔ lsuc l2) pr1 equivalence-class-Set = equivalence-class pr2 equivalence-class-Set = is-set-equivalence-class @@ -410,6 +413,15 @@ module _ = c by eq-effective-quotient' a c a-in-c = class R a' by inv (eq-effective-quotient' a' c a'-in-c) + -- sim-equivalence-relation-in-class : + -- {a a' : A} → + -- is-in-equivalence-class R (class R a) a' → + -- sim-equivalence-relation R a a' + -- sim-equivalence-relation-in-class {a} {a'} a'-in-c = + -- apply-effectiveness-class + -- ( eq-class-in-common-class (class R a') {! !} {! !}) + -- -- apply-effectiveness-class (eq-class-in-common-class c a-in-c a'-in-c) + sim-equivalence-relation-in-same-class : (c : equivalence-class R) {a a' : A} → is-in-equivalence-class R c a → @@ -417,129 +429,94 @@ module _ sim-equivalence-relation R a a' sim-equivalence-relation-in-same-class c {a} {a'} a-in-c a'-in-c = apply-effectiveness-class (eq-class-in-common-class c a-in-c a'-in-c) + + sim-equivalence-relation-in-class : + {a a' : A} → + is-in-equivalence-class R (class R a) a' → + sim-equivalence-relation R a a' + sim-equivalence-relation-in-class {a} {a'} a'-in-c = + sim-equivalence-relation-in-same-class + ( class R a) + ( is-in-self-equivalence-class R a) + ( a'-in-c) ``` ### TODO: Eliminator ```agda - ind-equivalence-class : - {l3 : Level} (B : equivalence-class R → Set l3) → - (f : (a : A) → type-Set (B (class R a))) → - ( (a a' : A) → + module _ + {l3 : Level} (B : equivalence-class R → Set l3) + (f : (a : A) → type-Set (B (class R a))) + (H : + (a a' : A) → (s : sim-equivalence-relation R a a') → - tr (type-Set ∘ B) (apply-effectiveness-class' s) (f a) = f a') → - (c : equivalence-class R) → type-Set (B c) - ind-equivalence-class {l3} B f H c = - pr1 - ( apply-universal-property-trunc-Prop - ( is-inhabited-subtype-equivalence-class R c) - ( b , is-prop-b) - ( λ (a , a-in-c) → - ( pair - ( tr (type-Set ∘ B) (eq-class-equivalence-class R c a-in-c) (f a)) - ( λ a' a'-in-c → - ( helper a a' a-in-c a'-in-c - ( H a a' - ( sim-equivalence-relation-in-same-class c - ( a-in-c) - ( a'-in-c)))))))) + dependent-identification + ( type-Set ∘ B) + ( apply-effectiveness-class' s) + ( f a) + ( f a')) where - helper : - (a a' : A) → - (a-in-c : is-in-equivalence-class R c a) → - (a'-in-c : is-in-equivalence-class R c a') → - tr (type-Set ∘ B) - ( apply-effectiveness-class' - (sim-equivalence-relation-in-same-class c a-in-c a'-in-c)) - (f a) - = f a' → - tr (type-Set ∘ B) (eq-class-equivalence-class R c a'-in-c) (f a') - = tr (type-Set ∘ B) (eq-class-equivalence-class R c a-in-c) (f a) - helper a a' a-in-c a'-in-c p = - let q = is-retraction-apply-effectiveness-class a a' - ( eq-class-in-common-class c a-in-c a'-in-c) - s = tr (λ t → tr (type-Set ∘ B) t (f a) = f a') q p - u = ap (tr (type-Set ∘ B) (eq-effective-quotient' a' c a'-in-c)) s - in + + private + + b : equivalence-class R → UU (l1 ⊔ l2 ⊔ l3) + b c = + Σ ( type-Set (B c)) + ( λ b → + (a : A) → + (a-in-c : is-in-equivalence-class R c a) → + tr (type-Set ∘ B) (eq-class-equivalence-class R c a-in-c) (f a) = b) + + is-prop-b : (c : equivalence-class R) → is-prop (b c) + is-prop-b c = + is-prop-all-elements-equal + ( λ (b , h) (b' , h') → + ( eq-pair-Σ + ( apply-universal-property-trunc-Prop + ( is-inhabited-subtype-equivalence-class R c) + ( b = b' , is-set-type-Set (B c) b b') + ( λ (a , a-in-c) → inv (h a a-in-c) ∙ h' a a-in-c)) + ( eq-is-prop + ( is-prop-Π + ( λ a → + ( is-prop-Π + ( λ a-in-c → + ( is-set-type-Set (B c) + ( tr (type-Set ∘ B) + ( eq-class-equivalence-class R c a-in-c) + ( f a)) + ( b'))))))))) + + b-instance-class : (a : A) → b (class R a) + pr1 (b-instance-class a) = f a + pr2 (b-instance-class a) a' a'-in-c = equational-reasoning - tr (type-Set ∘ B) (eq-class-equivalence-class R c a'-in-c) (f a') - = tr (type-Set ∘ B) (eq-effective-quotient' a' c a'-in-c) (f a') - by - ap - ( λ t → tr (type-Set ∘ B) t (f a')) - ( eq-is-prop (is-set-equivalence-class R (class R a') c)) - = tr (type-Set ∘ B) - ( eq-effective-quotient' a' c a'-in-c) - ( tr (type-Set ∘ B) - ( eq-class-in-common-class c a-in-c a'-in-c) - ( f a)) - by inv u - = tr (type-Set ∘ B) - ( eq-effective-quotient' a c a-in-c ∙ - inv (eq-effective-quotient' a' c a'-in-c) ∙ - eq-effective-quotient' a' c a'-in-c) - ( f a) - by - inv - ( tr-concat - ( eq-class-in-common-class c a-in-c a'-in-c) - ( eq-effective-quotient' a' c a'-in-c) - ( f a)) - = tr (type-Set ∘ B) - ( eq-effective-quotient' a c a-in-c ∙ - ( inv (eq-effective-quotient' a' c a'-in-c) ∙ - eq-effective-quotient' a' c a'-in-c)) - ( f a) - by - ap - ( λ t → tr (type-Set ∘ B) t (f a)) - ( assoc - ( eq-effective-quotient' a c a-in-c) - ( inv (eq-effective-quotient' a' c a'-in-c)) - ( eq-effective-quotient' a' c a'-in-c)) - = tr (type-Set ∘ B) (eq-effective-quotient' a c a-in-c ∙ refl) (f a) - by - ap - ( λ t → - ( tr (type-Set ∘ B) - ( eq-effective-quotient' a c a-in-c ∙ t) - ( f a))) - ( left-inv (eq-effective-quotient' a' c a'-in-c)) - = tr (type-Set ∘ B) (eq-effective-quotient' a c a-in-c) (f a) - by ap (λ t → tr (type-Set ∘ B) t (f a)) right-unit - = tr (type-Set ∘ B) (eq-class-equivalence-class R c a-in-c) (f a) + tr (type-Set ∘ B) _ (f a') + = tr (type-Set ∘ B) _ (f a') by ap - ( λ t → tr (type-Set ∘ B) t (f a)) - ( eq-is-prop (is-set-equivalence-class R (class R a) c)) - - b : UU (l1 ⊔ l2 ⊔ l3) - b = - Σ ( type-Set (B c)) - ( λ b → - (a : A) → - (a-in-c : is-in-equivalence-class R c a) → - tr (type-Set ∘ B) (eq-class-equivalence-class R c a-in-c) (f a) = b) - - is-prop-b : is-prop b - is-prop-b = - is-prop-all-elements-equal - ( λ (b , h) (b' , h') → - ( eq-pair-Σ - ( apply-universal-property-trunc-Prop - ( is-inhabited-subtype-equivalence-class R c) - ( b = b' , is-set-type-Set (B c) b b') - ( λ (a , a-in-c) → inv (h a a-in-c) ∙ h' a a-in-c)) - ( eq-is-prop - ( is-prop-Π - ( λ a → - ( is-prop-Π - ( λ a-in-c → - ( is-set-type-Set (B c) - ( tr (type-Set ∘ B) - ( eq-class-equivalence-class R c a-in-c) - ( f a)) - ( b'))))))))) + ( λ p → tr (type-Set ∘ B) p (f a')) + ( eq-is-prop (is-set-equivalence-class R _ _)) + = f a by inv-dependent-identification _ _ (H a a' a'-in-c) + + b-instance : (c : equivalence-class R) → b c + b-instance c = + apply-universal-property-trunc-Prop + ( is-inhabited-subtype-equivalence-class R c) + ( b c , is-prop-b c) + ( λ (a , a-in-c) → + ( tr b (eq-effective-quotient' a c a-in-c) (b-instance-class a))) + + ind-equivalence-class : (c : equivalence-class R) → type-Set (B c) + ind-equivalence-class = pr1 ∘ b-instance + + compute-ind-equivalence-class : + (a : A) → ind-equivalence-class (class R a) = f a + compute-ind-equivalence-class a = + ap pr1 + { x = b-instance (class R a)} + { y = b-instance-class a} + ( eq-is-prop (is-prop-b (class R a))) rec-equivalence-class : {l3 : Level} (B : Set l3) → @@ -548,6 +525,7 @@ module _ equivalence-class R → type-Set B rec-equivalence-class {l3} B f H = ind-equivalence-class (λ _ → B) f + -- TODO: duplicate code ( λ a a' s → equational-reasoning tr (type-Set ∘ (λ _ → B)) (apply-effectiveness-class' s) (f a) @@ -572,46 +550,24 @@ module _ rec-equivalence-class (set-Prop B) f ( λ a a' _ → eq-is-prop (is-prop-type-Prop B)) - -- TODO: proof - postulate - compute-rec-equivalence-class' : - {l3 : Level} (B : Set l3) → - (f : A → type-Set B) → - (H : (a a' : A) → sim-equivalence-relation R a a' → f a = f a') → - (a : A) → - rec-equivalence-class B f H (class R a) = f a - - -- TODO: complete - -- compute-rec-equivalence-class'' : - -- {l3 : Level} (B : Set l3) → - -- (f : A → type-Set B) → - -- (H : (a a' : A) → sim-equivalence-relation R a a' → f a = f a') → - -- (c : equivalence-class R) → - -- (a : A) → - -- (a-in-c : is-in-equivalence-class R c a) → - -- rec-equivalence-class B f H c = f a - -- compute-rec-equivalence-class'' B f H c a a-in-c = - -- let q = - -- ind-equivalence-class - -- ( λ c' → - -- ( Σ-Set B - -- ( λ b → set-Prop (exists-structure-Prop A (λ a' → f a = b))))) - -- ( λ a' → (f a) , intro-exists a refl) - -- ( λ x y s → - -- ( eq-pair-Σ - -- ( equational-reasoning - -- {! !} - -- -- tr {! !} (apply-effectiveness-class' s) (f a) {! !} - -- -- = {! !} by {! !} - -- = f a by {! !}) - -- ( eq-is-prop (is-prop-exists-structure _ _)))) - -- ( c) - -- in apply-universal-property-trunc-Prop - -- ( pr2 q) - -- ( Id-Prop B (rec-equivalence-class B f H c) (f a)) - -- ( λ (a , h) → inv (h ∙ {! !})) - compute-rec-equivalence-class : + {l3 : Level} (B : Set l3) → + (f : A → type-Set B) → + (H : (a a' : A) → sim-equivalence-relation R a a' → f a = f a') → + (a : A) → + rec-equivalence-class B f H (class R a) = f a + compute-rec-equivalence-class B f H = + compute-ind-equivalence-class (λ _ → B) f + -- TODO: duplicate code + ( λ a a' s → + equational-reasoning + tr (type-Set ∘ (λ _ → B)) (apply-effectiveness-class' s) (f a) + = f a + by tr-constant-type-family (apply-effectiveness-class' s) (f a) + = f a' + by H a a' s) + + compute-rec-equivalence-class' : {l3 : Level} (B : Set l3) → (f : A → type-Set B) → (H : (a a' : A) → sim-equivalence-relation R a a' → f a = f a') → @@ -619,7 +575,7 @@ module _ (a : A) → (a-in-c : is-in-equivalence-class R c a) → rec-equivalence-class B f H c = f a - compute-rec-equivalence-class B f H c a a-in-c = + compute-rec-equivalence-class' B f H c a a-in-c = equational-reasoning rec-equivalence-class B f H c = rec-equivalence-class B f H (class R a) @@ -628,7 +584,7 @@ module _ ( rec-equivalence-class B f H) ( inv (eq-effective-quotient' a c a-in-c)) = f a - by compute-rec-equivalence-class' B f H a + by compute-rec-equivalence-class B f H a ``` ### The map `class` into the type of equivalence classes is surjective and effective diff --git a/src/lists/lists-subtypes.lagda.md b/src/lists/lists-subtypes.lagda.md index ab5e45d5e0..350adad9f9 100644 --- a/src/lists/lists-subtypes.lagda.md +++ b/src/lists/lists-subtypes.lagda.md @@ -88,8 +88,9 @@ module _ subset-list-subtype-cons {a} {l} S a-in-S l-sub-S x = map-universal-property-trunc-Prop ( S x) - ( λ { (is-head .x .l) → a-in-S - ; (is-in-tail .x .a .l t) → l-sub-S x (in-list-subtype-in-list t)}) + ( λ where + (is-head .x .l) → a-in-S + (is-in-tail .x .a .l t) → l-sub-S x (in-list-subtype-in-list t)) subset-list-subtype-reverse-list : (l : list A) → list-subtype l ⊆ list-subtype (reverse-list l) diff --git a/src/modal-logic/axioms.lagda.md b/src/modal-logic/axioms.lagda.md index 0ba6f82786..1d3b62f009 100644 --- a/src/modal-logic/axioms.lagda.md +++ b/src/modal-logic/axioms.lagda.md @@ -57,7 +57,7 @@ module _ pr2 (ax-1-parameter h inj f) (a , refl) = is-prop-is-contr ( is-contr-Σ-is-prop a refl - ( λ _ → is-set-formula i _ _) + ( λ b → is-set-formula (h a) (h b)) ( λ x → inj)) ( a , refl) diff --git a/src/modal-logic/canonical-theories.lagda.md b/src/modal-logic/canonical-theories.lagda.md index 1f1a1ce17e..126b5be099 100644 --- a/src/modal-logic/canonical-theories.lagda.md +++ b/src/modal-logic/canonical-theories.lagda.md @@ -127,13 +127,14 @@ module _ ( λ b → ( map-universal-property-trunc-Prop ( modal-theory-L-complete-theory logic y b) - ( λ { (c , c-in-x , refl) → - ( xRy (◇ₘ c) - ( weak-modal-logic-mp - ( is-weak-modal-logic-L-complete-theory logic lzero x) - ( subset-logic-L-complete-theory logic lzero x (c →ₘ □ₘ ◇ₘ c) - ( ax-b-subset (c →ₘ □ₘ ◇ₘ c) (c , refl))) - ( c-in-x)))}))) + ( λ where + (c , c-in-x , refl) → + ( xRy (◇ₘ c) + ( weak-modal-logic-mp + ( is-weak-modal-logic-L-complete-theory logic lzero x) + ( subset-logic-L-complete-theory logic lzero x (c →ₘ □ₘ ◇ₘ c) + ( ax-b-subset (c →ₘ □ₘ ◇ₘ c) (c , refl))) + ( c-in-x)))))) ( a) ( box-a-in-y) diff --git a/src/modal-logic/decision-procedure.lagda.md b/src/modal-logic/decision-procedure.lagda.md index aaba16d86e..5d15d2ccf2 100644 --- a/src/modal-logic/decision-procedure.lagda.md +++ b/src/modal-logic/decision-procedure.lagda.md @@ -80,11 +80,11 @@ TODO ```agda module _ - {l1 l2 l3 l4 l5 : Level} - (i : Set l3) + {l1 l2 l3 l4 l5 l6 : Level} + (i : Set l1) (theory : modal-theory l2 i) - (C : model-class l1 l2 i l4 l5) - (C-sub-fin : C ⊆ finite-decidable-kripke-models l1 l2 i l4) + (C : model-class l3 l4 i l5 l6) + (C-sub-fin : C ⊆ finite-decidable-kripke-models l3 l4 i l5) (C-is-fin : is-finite (type-subtype C)) where diff --git a/src/modal-logic/deduction.lagda.md b/src/modal-logic/deduction.lagda.md index f9f91e8d57..04655ee4ae 100644 --- a/src/modal-logic/deduction.lagda.md +++ b/src/modal-logic/deduction.lagda.md @@ -32,10 +32,23 @@ open import modal-logic.formulas ## Idea -TODO +The deduction system of modal logic is defined inductively as follows: + +a ∈ T ------ (AX) T ⊢ a + +T ⊢ a → b T ⊢ a ------------------ (MP) T ⊢ b + +T ⊢ a ------ (NEC) T ⊢ □a + +where `T` is a set of formulas and `a`, `b` are formulas. + +Modal logic generated by a set of axioms is the smallest set of formulas that +contains the axioms and is closed under the deduction system. ## Definition +### Theory of modal formulas + ```agda module _ {l1 : Level} (l2 : Level) (i : Set l1) @@ -43,7 +56,11 @@ module _ modal-theory : UU (l1 ⊔ lsuc l2) modal-theory = subtype l2 (modal-formula i) +``` +### Deduction system + +```agda module _ {l1 l2 : Level} {i : Set l1} where @@ -55,17 +72,29 @@ module _ modal-mp : {a b : modal-formula i} → axioms ⊢ₘ a →ₘ b → axioms ⊢ₘ a → axioms ⊢ₘ b modal-nec : {a : modal-formula i} → axioms ⊢ₘ a → axioms ⊢ₘ □ₘ a +``` + +### Closure of a theory under the deduction system +```agda modal-logic-closure : modal-theory l2 i → modal-theory (l1 ⊔ l2) i modal-logic-closure axioms a = trunc-Prop (axioms ⊢ₘ a) +``` + +### Modal logic predicate +```agda is-modal-logic-Prop : modal-theory l2 i → Prop (l1 ⊔ l2) is-modal-logic-Prop theory = leq-prop-subtype (modal-logic-closure theory) theory is-modal-logic : modal-theory l2 i → UU (l1 ⊔ l2) is-modal-logic = type-Prop ∘ is-modal-logic-Prop +``` +### TODO: Title + +```agda is-in-modal-logic-closure-deduction : {axioms : modal-theory l2 i} {a : modal-formula i} → axioms ⊢ₘ a → is-in-subtype (modal-logic-closure axioms) a @@ -249,7 +278,7 @@ module _ (x : modal-formula i) → is-in-subtype theory-add-formula x → type-Prop (P x) elim-theory-add-formula P H-a H-rest = elim-union-subtype (Id-formula-Prop a) theory P - ( λ { .a refl → H-a}) + ( λ where .a refl → H-a) ( H-rest) subset-theory-add-formula : @@ -262,7 +291,7 @@ module _ ( Id-formula-Prop a) ( theory) ( theory') - ( λ { .a refl → a-in}) + ( λ where .a refl → a-in) module _ {l1 l2 : Level} {i : Set l1} diff --git a/src/modal-logic/formulas.lagda.md b/src/modal-logic/formulas.lagda.md index 0272cadcc4..bae7cecd24 100644 --- a/src/modal-logic/formulas.lagda.md +++ b/src/modal-logic/formulas.lagda.md @@ -27,10 +27,17 @@ open import foundation.universe-levels ## Idea -TODO +The formula of modal logic is defined inductively as follows: + +- A variable is a formula. +- ⊥ is a formula. +- If `a` and `b` are formulas, then `a → b` is a formula. +- If `a` is a formula, then □`a` is a formula. ## Definition +### Inductive definition of formulas + ```agda module _ {l : Level} @@ -44,7 +51,11 @@ module _ ⊥ₘ : modal-formula i _→ₘ_ : modal-formula i → modal-formula i → modal-formula i □ₘ_ : modal-formula i → modal-formula i +``` +### Additional notations + +```agda module _ {l : Level} {i : Set l} where @@ -72,7 +83,11 @@ module _ ◇ₘ_ : modal-formula i → modal-formula i ◇ₘ a = ¬ₘ □ₘ ¬ₘ a +``` +### If formulas are equal, then their subformulas are equal + +```agda eq-implication-left : {a b c d : modal-formula i} → a →ₘ b = c →ₘ d → a = c eq-implication-left refl = refl @@ -88,11 +103,11 @@ module _ ## Properties -### The formulas are a set +### Characterizing the identity type of lists ```agda module _ - {l : Level} (i : Set l) + {l : Level} {i : Set l} where Eq-formula : modal-formula i → modal-formula i → UU l @@ -167,12 +182,16 @@ module _ ( is-prop-Eq-formula) ( refl-Eq-formula) ( λ _ _ → eq-Eq-formula) +``` + +### A formula is a set - formula-Set : Set l - pr1 formula-Set = modal-formula i - pr2 formula-Set = is-set-formula +```agda +formula-Set : {l : Level} (i : Set l) → Set l +pr1 (formula-Set i) = modal-formula i +pr2 (formula-Set i) = is-set-formula Id-formula-Prop : {l : Level} {i : Set l} → modal-formula i → modal-formula i → Prop l -Id-formula-Prop = Id-Prop (formula-Set _) +Id-formula-Prop {i = i} = Id-Prop (formula-Set i) ``` diff --git a/src/modal-logic/kripke-models-filtrations.lagda.md b/src/modal-logic/kripke-models-filtrations.lagda.md index 399c7c9c26..578d339019 100644 --- a/src/modal-logic/kripke-models-filtrations.lagda.md +++ b/src/modal-logic/kripke-models-filtrations.lagda.md @@ -252,7 +252,7 @@ module _ = map-function-equivalence-class x-class by inv - ( compute-rec-equivalence-class + ( compute-rec-equivalence-class' ( Φ-equivalence) ( map-function-equivalence-class-Set) ( map-function-worlds) @@ -264,7 +264,7 @@ module _ by p = map-function-worlds y by - compute-rec-equivalence-class + compute-rec-equivalence-class' ( Φ-equivalence) ( map-function-equivalence-class-Set) ( map-function-worlds) From e1bc75f5218d148a03170af3cb91630e211c5117 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Thu, 16 May 2024 21:34:50 +0300 Subject: [PATCH 60/63] Refactor finite approximability --- .../decidable-propositions.lagda.md | 1 + src/foundation/equivalence-classes.lagda.md | 1 - src/modal-logic/decision-procedure.lagda.md | 105 ++++---- .../finite-approximability.lagda.md | 229 +++++++++--------- src/modal-logic/formulas.lagda.md | 10 +- src/modal-logic/kripke-semantics.lagda.md | 23 ++ 6 files changed, 194 insertions(+), 175 deletions(-) diff --git a/src/foundation-core/decidable-propositions.lagda.md b/src/foundation-core/decidable-propositions.lagda.md index fbb712b15a..ca6173c9fc 100644 --- a/src/foundation-core/decidable-propositions.lagda.md +++ b/src/foundation-core/decidable-propositions.lagda.md @@ -214,6 +214,7 @@ equiv-empty-or-unit-is-decidable-prop {l} {A} H@(_ , is-d) with is-d ... | inl contra = inl (equiv-unit-is-decidable-prop H contra) ... | inr a = inr (equiv-empty-is-decidable-prop H a) +-- TODO: move to foundation is-small-prop-is-decidable-prop : {l1 : Level} (l2 : Level) (A : UU l1) → is-decidable-prop A → is-small l2 A is-small-prop-is-decidable-prop l2 A H diff --git a/src/foundation/equivalence-classes.lagda.md b/src/foundation/equivalence-classes.lagda.md index 24a76700ba..0d2aee702b 100644 --- a/src/foundation/equivalence-classes.lagda.md +++ b/src/foundation/equivalence-classes.lagda.md @@ -458,7 +458,6 @@ module _ where private - b : equivalence-class R → UU (l1 ⊔ l2 ⊔ l3) b c = Σ ( type-Set (B c)) diff --git a/src/modal-logic/decision-procedure.lagda.md b/src/modal-logic/decision-procedure.lagda.md index 5d15d2ccf2..5c3e1a7824 100644 --- a/src/modal-logic/decision-procedure.lagda.md +++ b/src/modal-logic/decision-procedure.lagda.md @@ -80,9 +80,8 @@ TODO ```agda module _ - {l1 l2 l3 l4 l5 l6 : Level} + {l1 l3 l4 l5 l6 : Level} (i : Set l1) - (theory : modal-theory l2 i) (C : model-class l3 l4 i l5 l6) (C-sub-fin : C ⊆ finite-decidable-kripke-models l3 l4 i l5) (C-is-fin : is-finite (type-subtype C)) @@ -106,16 +105,17 @@ module _ ... | inr _ = false decision-procedure-correctness : + {l2 : Level} (theory : modal-theory l2 i) → soundness theory C → completeness theory C → (a : modal-formula i) → is-in-subtype theory a ↔ type-prop-bool (decision-procedure a) - pr1 (decision-procedure-correctness sound complete a) in-theory + pr1 (decision-procedure-correctness theory sound complete a) in-theory with decision-procedure' a ... | inl _ = star ... | inr not-valid-in-C = not-valid-in-C (λ (M , M-in-C) x → sound a in-theory M M-in-C x) - pr2 (decision-procedure-correctness sound complete a) _ + pr2 (decision-procedure-correctness theory sound complete a) _ with decision-procedure' a ... | inl valid-in-C = complete a (λ M M-in-C → valid-in-C (M , M-in-C)) @@ -172,7 +172,7 @@ in-dependent-map-list {A = A} {B} f (is-in-tail _ x l a-in-l) = is-in-tail _ _ _ (in-dependent-map-list _ a-in-l) module _ - {l : Level} (i : Set l) + {l : Level} {i : Set l} where subformulas-list : modal-formula i → list (modal-formula i) @@ -445,30 +445,19 @@ module _ (lem : LEM (l2 ⊔ lsuc l3 ⊔ l4 ⊔ lsuc l5 ⊔ lsuc l6 ⊔ lsuc l7 ⊔ lsuc l8)) where - filtration-models-subset-finite-decidable-kripke-models : + filtration-models-subset-finite-kripke-models : filtration-models l1 l2 i l4 l5 l6 l7 l8 ⊆ - finite-decidable-kripke-models l1 l2 i l4 - filtration-models-subset-finite-decidable-kripke-models M* = + finite-kripke-models l1 l2 i l4 + filtration-models-subset-finite-kripke-models M* = map-universal-property-trunc-Prop - ( finite-decidable-kripke-models l1 l2 i l4 M*) + ( finite-kripke-models l1 l2 i l4 M*) ( λ ((theory , M) , is-fin , is-filt) → - ( triple - ( is-finite-equiv - ( equiv-is-kripke-model-filtration i theory M M* is-filt) - ( is-finite-equivalence-classes-filtration i M - ( lower-LEM (l2 ⊔ l4) lem) - ( theory) - ( is-fin))) - ( λ x y → - ( lower-LEM - ( lsuc l3 ⊔ l4 ⊔ lsuc l5 ⊔ lsuc l6 ⊔ lsuc l7 ⊔ lsuc l8) - ( lem) - ( relation-Prop-kripke-model i M* x y))) - ( λ x n → - ( lower-LEM - ( l2 ⊔ lsuc l3 ⊔ lsuc l5 ⊔ lsuc l6 ⊔ lsuc l7 ⊔ lsuc l8) - ( lem) - ( valuate-kripke-model i M* n x))))) + ( is-finite-equiv + ( equiv-is-kripke-model-filtration i theory M M* is-filt) + ( is-finite-equivalence-classes-filtration i M + ( lower-LEM (l2 ⊔ l4) lem) + ( theory) + ( is-fin)))) module _ {l1 l2 l3 l4 l5 l6 l7 l8 : Level} (i : Set l3) @@ -482,7 +471,7 @@ module _ model-class l6 l7 i l8 ( l3 ⊔ l5 ⊔ lsuc (l1 ⊔ l2 ⊔ l4 ⊔ l6 ⊔ l7 ⊔ l8)) filtrate-class M* = exists-structure-Prop (modal-formula i × type-subtype C) - ( λ (a , (M , _)) → M* = filtration (subformulas i a) M) + ( λ (a , M , _) → M* = filtration (subformulas a) M) module _ (filtration-is-filtration : @@ -492,6 +481,29 @@ module _ is-kripke-model-filtration i theory M (filtration theory M)) where + is-finite-filtrate-class : + LEM (lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4)) → + filtrate-class ⊆ finite-kripke-models l6 l7 i l8 + is-finite-filtrate-class lem M* = + elim-exists + ( finite-kripke-models l6 l7 i l8 M*) + ( λ where + (a , M , M-in-C) refl → + ( is-finite-equiv + ( equiv-is-kripke-model-filtration i + ( subformulas a) + ( M) + ( M*) + ( filtration-is-filtration + ( M , M-in-C) + ( subformulas a) + ( is-modal-theory-closed-under-subformulas-subformulas a))) + ( is-finite-equivalence-classes-filtration i M lem + ( subformulas a) + ( is-finite-subformulas-list + ( lower-LEM (lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4)) lem) + ( a))))) + filtrate-completeness : {l9 : Level} (logic : modal-theory l9 i) → completeness logic C → @@ -501,43 +513,40 @@ module _ ( λ M M-in-class x → ( backward-implication ( filtration-lemma i - ( subformulas i a) - ( is-modal-theory-closed-under-subformulas-subformulas i a) + ( subformulas a) + ( is-modal-theory-closed-under-subformulas-subformulas a) ( M) - ( filtration (subformulas i a) M) + ( filtration (subformulas a) M) ( filtration-is-filtration ( M , M-in-class) - ( subformulas i a) - ( is-modal-theory-closed-under-subformulas-subformulas i a)) + ( subformulas a) + ( is-modal-theory-closed-under-subformulas-subformulas a)) ( a) ( head-in-list-subtype) ( x)) ( in-logic - ( filtration (subformulas i a) M) + ( filtration (subformulas a) M) ( intro-exists (a , M , M-in-class) refl) ( map-equiv-is-kripke-model-filtration i - ( subformulas i a) + ( subformulas a) ( M) - ( filtration (subformulas i a) M) + ( filtration (subformulas a) M) ( filtration-is-filtration ( M , M-in-class) - ( subformulas i a) - ( is-modal-theory-closed-under-subformulas-subformulas i a)) - ( class (Φ-equivalence i (subformulas i a) M) x))))) + ( subformulas a) + ( is-modal-theory-closed-under-subformulas-subformulas a)) + ( class (Φ-equivalence i (subformulas a) M) x))))) filtrate-soundness : {l9 l10 : Level} (logic : modal-theory l9 i) → (C₂ : model-class l6 l7 i l8 l10) → - ( ((M , _) : type-subtype C) → - (a : modal-formula i) → - is-in-subtype C₂ (filtration (list-subtype (subformulas-list i a)) M)) → + filtrate-class ⊆ C₂ → soundness logic C₂ → soundness logic filtrate-class - filtrate-soundness logic C₂ H sound a in-logic M* in-class = - apply-universal-property-trunc-Prop - ( in-class) - ( M* ⊨Mₘ a) - ( λ ((b , (M , in-C)) , p) → - ( sound a in-logic M* - ( tr (is-in-subtype C₂) (inv p) (H (M , in-C) b)))) + filtrate-soundness logic C₂ H = + transitive-leq-subtype + ( logic) + ( class-modal-logic C₂) + ( class-modal-logic filtrate-class) + ( class-modal-logic-monotic filtrate-class C₂ H) ``` diff --git a/src/modal-logic/finite-approximability.lagda.md b/src/modal-logic/finite-approximability.lagda.md index df29fe1bfc..82772920c4 100644 --- a/src/modal-logic/finite-approximability.lagda.md +++ b/src/modal-logic/finite-approximability.lagda.md @@ -39,12 +39,14 @@ open import foundation-core.invertible-maps open import modal-logic.completeness open import modal-logic.completeness-k +open import modal-logic.completeness-s5 open import modal-logic.decision-procedure open import modal-logic.deduction open import modal-logic.formulas open import modal-logic.kripke-models-filtrations open import modal-logic.kripke-semantics open import modal-logic.modal-logic-k +open import modal-logic.modal-logic-s5 open import modal-logic.soundness open import modal-logic.weak-deduction @@ -65,7 +67,6 @@ TODO module _ {l1 : Level} (i : Set l1) where - is-finitely-approximable-Prop : {l2 : Level} (l3 l4 l5 l6 : Level) → modal-theory l2 i → @@ -74,23 +75,9 @@ module _ exists-structure-Prop ( model-class l3 l4 i l5 l6) ( λ C → - ( product - ( soundness logic (finite-subclass i C)) - ( completeness logic (finite-subclass i C)))) - - is-finitely-approximable-Prop' : - {l2 : Level} (l3 l4 l5 l6 : Level) → - modal-theory l2 i → - Prop (l1 ⊔ l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5 ⊔ lsuc l6) - is-finitely-approximable-Prop' l3 l4 l5 l6 logic = - exists-structure-Prop - ( model-class l3 l4 i l5 l6) - ( λ C → - ( product - ( C ⊆ finite-kripke-models l3 l4 i l5) - ( product - ( soundness logic (finite-subclass i C)) - ( completeness logic (finite-subclass i C))))) + ( C ⊆ finite-kripke-models l3 l4 i l5) × + ( soundness logic C) × + ( completeness logic C)) is-finitely-approximable : {l2 : Level} (l3 l4 l5 l6 : Level) → @@ -99,125 +86,125 @@ module _ is-finitely-approximable l3 l4 l5 l6 logic = type-Prop (is-finitely-approximable-Prop l3 l4 l5 l6 logic) + module _ + {l2 l3 l4 l5 l6 l7 l8 l9 l10 : Level} + (C : model-class l2 l3 i l4 l5) + (filtration : modal-theory l1 i → + kripke-model l2 l3 i l4 → + kripke-model l6 l7 i l8) + (is-filtration : + ((M , _) : type-subtype C) (theory : modal-theory l1 i) → + is-modal-theory-closed-under-subformulas i theory → + is-kripke-model-filtration i theory M (filtration theory M)) + (logic : modal-theory l9 i) + (complete : completeness logic C) + (C₂ : model-class l6 l7 i l8 l10) + (leq : filtrate-class i C filtration ⊆ C₂) + (sound : soundness logic C₂) + where + + is-finitely-approximable-filtration : + LEM (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4) → + is-finitely-approximable l6 l7 l8 + (l1 ⊔ l5 ⊔ lsuc (l2 ⊔ l3 ⊔ l4 ⊔ l6 ⊔ l7 ⊔ l8)) logic + is-finitely-approximable-filtration lem = + intro-exists (filtrate-class i C filtration) + ( triple + ( is-finite-filtrate-class i C filtration is-filtration lem) + ( filtrate-soundness i C filtration logic C₂ leq sound) + ( filtrate-completeness i C filtration is-filtration logic complete)) + module _ (lem : LEM (lsuc (lsuc l1))) (zorn : Zorn (lsuc l1) l1 l1) (prop-resize : propositional-resizing l1 (lsuc l1)) where - K-finite-class : - model-class (lsuc (lsuc l1)) (lsuc l1) i (lsuc l1) (lsuc (lsuc (lsuc l1))) - K-finite-class = - filtrate-class i - ( all-models (lsuc l1) l1 i l1) - ( minimal-kripke-model-filtration i) - - K-finite-class-sub-filtration-models : - K-finite-class ⊆ - filtration-models - (lsuc (lsuc l1)) (lsuc l1) i (lsuc l1) l1 (lsuc l1) l1 l1 - K-finite-class-sub-filtration-models M* = - map-universal-property-trunc-Prop - ( filtration-models _ _ i _ _ _ _ _ M*) - ( λ ((a , M , _) , p) → - ( tr (is-in-subtype (filtration-models _ _ i _ _ _ _ _)) (inv p) - ( intro-exists - ( subformulas i a , M) - ( pair - (is-finite-subformulas-list - ( i) - ( lower-LEM (lsuc (lsuc l1)) lem) - ( a)) - ( is-kripke-model-filtration-minimal-kripke-model-filtration - ( i) - ( subformulas i a) - ( M) - ( is-modal-theory-closed-under-subformulas-subformulas - ( i) - ( a))))))) - - completeness-K-filtration : - completeness (modal-logic-K i) (K-finite-class) - completeness-K-filtration = - filtrate-completeness i + is-finitely-approximable-K : + is-finitely-approximable + ( lsuc (lsuc l1)) + ( lsuc l1) + ( lsuc l1) + ( lsuc (lsuc (lsuc l1))) + ( modal-logic-K i) + is-finitely-approximable-K = + is-finitely-approximable-filtration ( all-models (lsuc l1) l1 i l1) - ( minimal-kripke-model-filtration i) - ( λ (M , _) theory theory-is-closed → + ( minimal-kripke-model-filtration _) -- TODO: make i implicit + ( λ (M , _) theory is-closed → ( is-kripke-model-filtration-minimal-kripke-model-filtration i ( theory) ( M) - ( theory-is-closed))) + ( is-closed))) ( modal-logic-K i) ( completeness-K i (lower-LEM (lsuc (lsuc l1)) lem) zorn prop-resize) - - soundness-K-filtration : - soundness (modal-logic-K i) (K-finite-class) - soundness-K-filtration = - filtrate-soundness i - ( all-models (lsuc l1) l1 i l1) - ( minimal-kripke-model-filtration i) - ( modal-logic-K i) ( all-models (lsuc (lsuc l1)) (lsuc l1) i (lsuc l1)) ( λ _ _ → star) - ( λ a in-logic M _ x → - ( soundness-K i a in-logic M (λ b y → lem ((M , y) ⊨ₘ b)) x)) - - is-finitely-approximable-K : + ( transitive-leq-subtype + ( modal-logic-K i) + ( class-modal-logic + ( decidable-kripke-models (lsuc (lsuc l1)) (lsuc l1) i (lsuc l1))) + ( class-modal-logic + ( all-models (lsuc (lsuc l1)) (lsuc l1) i (lsuc l1))) + ( class-modal-logic-monotic + ( all-models (lsuc (lsuc l1)) (lsuc l1) i (lsuc l1)) + ( decidable-kripke-models (lsuc (lsuc l1)) (lsuc l1) i (lsuc l1)) + ( all-models-is-decidable i lem)) + ( soundness-K i)) + ( lem) + + is-finitely-approximable-S5 : is-finitely-approximable ( lsuc (lsuc l1)) - ( lsuc l1) + ( lsuc (lsuc l1)) ( lsuc l1) ( lsuc (lsuc (lsuc l1))) - ( modal-logic-K i) - is-finitely-approximable-K = - intro-exists - ( K-finite-class) - ( pair - ( soundness-subclass - ( modal-logic-K i) - ( K-finite-class) - ( finite-subclass i K-finite-class) - ( subtype-intersection-right - ( finite-decidable-kripke-models - ( lsuc (lsuc l1)) - ( lsuc l1) - ( i) - ( lsuc l1)) - ( K-finite-class)) - ( soundness-K-filtration)) - ( transitive-leq-subtype - ( class-modal-logic (finite-subclass i K-finite-class)) - ( class-modal-logic K-finite-class) - ( modal-logic-K i) - ( completeness-K-filtration) - ( λ a a-in-logic M M-in-class → - ( a-in-logic M - ( pair - ( transitive-leq-subtype - ( K-finite-class) - ( filtration-models - ( lsuc (lsuc l1)) - ( lsuc l1) - ( i) - ( lsuc l1) - ( l1) - ( lsuc l1) - ( l1) - ( l1)) - ( finite-decidable-kripke-models - ( lsuc (lsuc l1)) - ( lsuc l1) - ( i) - ( lsuc l1)) - ( filtration-models-subset-finite-decidable-kripke-models - ( i) - ( l1) - ( lsuc l1) - ( l1) - ( l1) - ( lem)) - ( K-finite-class-sub-filtration-models) - ( M) - ( M-in-class)) - ( M-in-class)))))) + ( modal-logic-S5 i) + is-finitely-approximable-S5 = + is-finitely-approximable-filtration + ( equivalence-kripke-models (lsuc l1) l1 i l1) + ( minimal-transitive-kripke-model-filtration _) -- TODO: make i implicit + ( λ (M , in-equiv) theory is-closed → + ( is-kripke-model-filtration-minimal-transitive-kripke-model-filtration + ( i) + ( theory) + ( M) + ( is-closed) + ( pr2 (pr2 in-equiv)))) -- TODO: refactor + ( modal-logic-S5 i) + ( completeness-S5 i (lower-LEM (lsuc (lsuc l1)) lem) zorn prop-resize) + ( equivalence-kripke-models + (lsuc (lsuc l1)) (lsuc (lsuc l1)) i (lsuc l1)) + ( λ M* → + ( elim-exists + ( equivalence-kripke-models + (lsuc (lsuc l1)) (lsuc (lsuc l1)) i (lsuc l1) M*) + ( λ where + ( a , M , in-equiv) refl → + ( minimal-transitive-filtration-preserves-equivalence + _ -- TODO: make implicit + ( subformulas a) + ( M) + ( is-modal-theory-closed-under-subformulas-subformulas a) + ( in-equiv))))) + ( transitive-leq-subtype + ( modal-logic-S5 i) + ( class-modal-logic + ( decidable-subclass i + ( equivalence-kripke-models + (lsuc (lsuc l1)) (lsuc (lsuc l1)) i (lsuc l1)))) + ( class-modal-logic + ( equivalence-kripke-models + (lsuc (lsuc l1)) (lsuc (lsuc l1)) i (lsuc l1))) + ( class-modal-logic-monotic + ( equivalence-kripke-models + (lsuc (lsuc l1)) (lsuc (lsuc l1)) i (lsuc l1)) + ( decidable-subclass i + ( equivalence-kripke-models + (lsuc (lsuc l1)) (lsuc (lsuc l1)) i (lsuc l1))) + ( subset-decidable-subclass-lem i lem + ( equivalence-kripke-models + (lsuc (lsuc l1)) (lsuc (lsuc l1)) i (lsuc l1)))) + ( soundness-S5 i)) + ( lem) ``` diff --git a/src/modal-logic/formulas.lagda.md b/src/modal-logic/formulas.lagda.md index bae7cecd24..a7d5e4a7e0 100644 --- a/src/modal-logic/formulas.lagda.md +++ b/src/modal-logic/formulas.lagda.md @@ -103,7 +103,7 @@ module _ ## Properties -### Characterizing the identity type of lists +### Characterizing the identity type of formulas ```agda module _ @@ -174,7 +174,11 @@ module _ is-prop-Eq-formula (□ₘ _) ⊥ₘ = is-prop-raise-empty is-prop-Eq-formula (□ₘ _) (_ →ₘ _) = is-prop-raise-empty is-prop-Eq-formula (□ₘ a) (□ₘ c) = is-prop-Eq-formula a c +``` + +### A formula is a set +```agda is-set-formula : is-set (modal-formula i) is-set-formula = is-set-prop-in-id @@ -182,11 +186,7 @@ module _ ( is-prop-Eq-formula) ( refl-Eq-formula) ( λ _ _ → eq-Eq-formula) -``` - -### A formula is a set -```agda formula-Set : {l : Level} (i : Set l) → Set l pr1 (formula-Set i) = modal-formula i pr2 (formula-Set i) = is-set-formula diff --git a/src/modal-logic/kripke-semantics.lagda.md b/src/modal-logic/kripke-semantics.lagda.md index 24fb2eea9b..c5e15b3d91 100644 --- a/src/modal-logic/kripke-semantics.lagda.md +++ b/src/modal-logic/kripke-semantics.lagda.md @@ -21,6 +21,7 @@ open import foundation.function-types open import foundation.inhabited-types open import foundation.intersections-subtypes open import foundation.iterated-dependent-product-types +open import foundation.law-of-excluded-middle open import foundation.negation open import foundation.propositional-extensionality open import foundation.propositions @@ -318,4 +319,26 @@ module _ model-class l1 l2 i l4 l5 → model-class l1 l2 i l4 (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) finite-subclass C = (finite-decidable-kripke-models l1 l2 i l4) ∩ C + +module _ + {l1 l2 l3 l4 : Level} (i : Set l3) + (lem : LEM (l1 ⊔ l2 ⊔ l4)) + where + + all-models-is-decidable : + all-models l1 l2 i l4 ⊆ decidable-kripke-models l1 l2 i l4 + all-models-is-decidable M _ a x = lem ((M , x) ⊨ₘ a) + + subset-decidable-subclass-lem : + {l5 : Level} (C : model-class l1 l2 i l4 l5) → + C ⊆ decidable-subclass i C + subset-decidable-subclass-lem C = + subtype-both-intersection (decidable-kripke-models l1 l2 i l4) C C + ( transitive-leq-subtype + ( C) + ( all-models l1 l2 i l4) + ( decidable-kripke-models l1 l2 i l4) + ( all-models-is-decidable) + ( all-models-is-largest-class i C)) + ( refl-leq-subtype C) ``` From a9d4e80b30b623537669d4f04928f62bfeb07446 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Sat, 25 May 2024 03:22:13 +0300 Subject: [PATCH 61/63] Refactor --- src/foundation.lagda.md | 1 + ...ary-relations-transitive-closures.lagda.md | 120 +++++++++ src/foundation/subtypes.lagda.md | 4 + src/lists/functoriality-lists.lagda.md | 24 ++ src/lists/lists-subtypes.lagda.md | 1 + src/modal-logic.lagda.md | 2 + src/modal-logic/axioms.lagda.md | 2 +- .../canonical-model-theorem.lagda.md | 27 ++- ...closed-under-subformulas-theories.lagda.md | 95 ++++++++ src/modal-logic/decision-procedure.lagda.md | 227 ++---------------- src/modal-logic/deduction.lagda.md | 15 +- src/modal-logic/filtration-lemma.lagda.md | 97 ++++---- .../finite-approximability.lagda.md | 41 +--- src/modal-logic/formulas.lagda.md | 101 ++++---- .../kripke-models-filtrations.lagda.md | 169 +++---------- src/modal-logic/l-complete-theories.lagda.md | 3 +- src/modal-logic/subformulas.lagda.md | 146 +++++++++++ src/order-theory.lagda.md | 1 + src/order-theory/chains-posets.lagda.md | 23 -- .../upper-bounds-chains-posets.lagda.md | 48 ++++ src/order-theory/zorn.lagda.md | 1 + .../kuratowsky-finite-sets.lagda.md | 93 ++++++- 22 files changed, 728 insertions(+), 513 deletions(-) create mode 100644 src/foundation/binary-relations-transitive-closures.lagda.md create mode 100644 src/modal-logic/closed-under-subformulas-theories.lagda.md create mode 100644 src/modal-logic/subformulas.lagda.md create mode 100644 src/order-theory/upper-bounds-chains-posets.lagda.md diff --git a/src/foundation.lagda.md b/src/foundation.lagda.md index 563cc1c995..28dcaf1f4a 100644 --- a/src/foundation.lagda.md +++ b/src/foundation.lagda.md @@ -38,6 +38,7 @@ open import foundation.binary-homotopies public open import foundation.binary-operations-unordered-pairs-of-types public open import foundation.binary-reflecting-maps-equivalence-relations public open import foundation.binary-relations public +open import foundation.binary-relations-transitive-closures public open import foundation.binary-relations-with-extensions public open import foundation.binary-relations-with-lifts public open import foundation.binary-transport public diff --git a/src/foundation/binary-relations-transitive-closures.lagda.md b/src/foundation/binary-relations-transitive-closures.lagda.md new file mode 100644 index 0000000000..13c5a9b454 --- /dev/null +++ b/src/foundation/binary-relations-transitive-closures.lagda.md @@ -0,0 +1,120 @@ +# Transitive Closures + +```agda +module foundation.binary-relations-transitive-closures where +``` + +
Imports + +```agda +open import foundation.binary-relations +open import foundation.propositional-truncations +open import foundation.universe-levels + +open import foundation-core.function-types +open import foundation-core.propositions +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l1 l2 : Level} {A : UU l1} + where + + data transitive-closure (R : Relation l2 A) : Relation (l1 ⊔ l2) A + where + transitive-closure-base : {x y : A} → R x y → transitive-closure R x y + transitive-closure-step : + {x y z : A} → R x y → transitive-closure R y z → transitive-closure R x z + + is-transitive-transitive-closure : + (R : Relation l2 A) → is-transitive (transitive-closure R) + is-transitive-transitive-closure + R x y z c-yz (transitive-closure-base r-xy) = + transitive-closure-step r-xy c-yz + is-transitive-transitive-closure + R x y z c-yz (transitive-closure-step {y = u} r-xu c-uy) = + transitive-closure-step r-xu + ( is-transitive-transitive-closure R u y z c-yz c-uy) + + transitive-closure-preserves-reflexivity : + (R : Relation l2 A) → + is-reflexive R → + is-reflexive (transitive-closure R) + transitive-closure-preserves-reflexivity R is-refl x = + transitive-closure-base (is-refl x) + + transitive-closure-preserves-symmetry : + (R : Relation l2 A) → + is-symmetric R → + is-symmetric (transitive-closure R) + transitive-closure-preserves-symmetry R is-sym x y + (transitive-closure-base r) = + transitive-closure-base (is-sym x y r) + transitive-closure-preserves-symmetry + R is-sym x y (transitive-closure-step {y = u} r-xu c-uy) = + is-transitive-transitive-closure R y u x + ( transitive-closure-base (is-sym x u r-xu)) + ( transitive-closure-preserves-symmetry R is-sym u y c-uy) +``` + +### Transitive closure of a relation valued in propositions + +```agda + transitive-closure-Prop : + (R : Relation-Prop l2 A) → Relation-Prop (l1 ⊔ l2) A + transitive-closure-Prop R x y = + trunc-Prop (transitive-closure (type-Relation-Prop R) x y) + + is-transitive-transitive-closure-Prop : + (R : Relation-Prop l2 A) → + is-transitive-Relation-Prop (transitive-closure-Prop R) + is-transitive-transitive-closure-Prop R x y z c-yz c-xy = + apply-twice-universal-property-trunc-Prop + ( c-yz) + ( c-xy) + ( transitive-closure-Prop R x z) + ( λ r-yz r-xy → + unit-trunc-Prop + ( is-transitive-transitive-closure + ( type-Relation-Prop R) + ( x) + ( y) + ( z) + ( r-yz) + ( r-xy))) + + transitive-closure-Prop-preserves-reflexivity : + (R : Relation-Prop l2 A) → + is-reflexive-Relation-Prop R → + is-reflexive-Relation-Prop (transitive-closure-Prop R) + transitive-closure-Prop-preserves-reflexivity R is-refl x = + unit-trunc-Prop + ( transitive-closure-preserves-reflexivity + ( type-Relation-Prop R) + ( is-refl) + ( x)) + + transitive-closure-Prop-preserves-symmetry : + (R : Relation-Prop l2 A) → + is-symmetric-Relation-Prop R → + is-symmetric-Relation-Prop (transitive-closure-Prop R) + transitive-closure-Prop-preserves-symmetry R is-sym x y = + map-universal-property-trunc-Prop + ( transitive-closure-Prop R y x) + ( λ r-xy → + unit-trunc-Prop + ( transitive-closure-preserves-symmetry + ( type-Relation-Prop R) + ( is-sym) + ( x) + ( y) + ( r-xy))) +``` diff --git a/src/foundation/subtypes.lagda.md b/src/foundation/subtypes.lagda.md index 5405df24ea..b202c8798e 100644 --- a/src/foundation/subtypes.lagda.md +++ b/src/foundation/subtypes.lagda.md @@ -166,6 +166,10 @@ module _ equiv-subtypes P Q → equiv-subtypes Q P inv-equiv-subtypes P Q e x = inv-equiv (e x) + id-equiv-subtypes : + {l2 : Level} (P : subtype l2 A) → equiv-subtypes P P + id-equiv-subtypes P x = id-equiv + equiv-antisymmetric-leq-subtype : {l2 l3 : Level} (P : subtype l2 A) (Q : subtype l3 A) → P ⊆ Q → Q ⊆ P → equiv-subtypes P Q diff --git a/src/lists/functoriality-lists.lagda.md b/src/lists/functoriality-lists.lagda.md index ce61ae6805..9a90d76af6 100644 --- a/src/lists/functoriality-lists.lagda.md +++ b/src/lists/functoriality-lists.lagda.md @@ -257,3 +257,27 @@ map-snoc-list : map-snoc-list f nil a = refl map-snoc-list f (cons b x) a = ap (cons (f b)) (map-snoc-list f x a) ``` + +### TODO: maybe another file + +```agda +dependent-map-list : + {l1 l2 : Level} {A : UU l1} {B : UU l2} + (l : list A) (f : (a : A) → a ∈-list l → B) → + list B +dependent-map-list nil f = nil +dependent-map-list {A = A} {B} (cons x l) f = + cons (f x (is-head x l)) (dependent-map-list l f') + where + f' : (a : A) → a ∈-list l → B + f' a list-subtype = f a (is-in-tail a x l list-subtype) + +in-dependent-map-list : + {l1 l2 : Level} {A : UU l1} {B : UU l2} + {l : list A} (f : (a : A) → a ∈-list l → B) + {a : A} (a-in-l : a ∈-list l) → + f a a-in-l ∈-list dependent-map-list l f +in-dependent-map-list f (is-head _ l) = is-head _ _ +in-dependent-map-list {A = A} {B} f (is-in-tail _ x l a-in-l) = + is-in-tail _ _ _ (in-dependent-map-list _ a-in-l) +``` diff --git a/src/lists/lists-subtypes.lagda.md b/src/lists/lists-subtypes.lagda.md index 350adad9f9..4b0d2ab0a4 100644 --- a/src/lists/lists-subtypes.lagda.md +++ b/src/lists/lists-subtypes.lagda.md @@ -24,6 +24,7 @@ open import foundation-core.empty-types open import foundation-core.function-types open import foundation-core.negation open import foundation-core.propositions +open import foundation-core.sets open import foundation-core.subtypes open import lists.concatenation-lists diff --git a/src/modal-logic.lagda.md b/src/modal-logic.lagda.md index d4f78a1f4e..28ed4dfca1 100644 --- a/src/modal-logic.lagda.md +++ b/src/modal-logic.lagda.md @@ -6,6 +6,7 @@ module modal-logic where open import modal-logic.axioms public open import modal-logic.canonical-model-theorem public open import modal-logic.canonical-theories public +open import modal-logic.closed-under-subformulas-theories public open import modal-logic.completeness public open import modal-logic.completeness-k public open import modal-logic.completeness-s5 public @@ -23,5 +24,6 @@ open import modal-logic.lindenbaum public open import modal-logic.modal-logic-k public open import modal-logic.modal-logic-s5 public open import modal-logic.soundness public +open import modal-logic.subformulas public open import modal-logic.weak-deduction public ``` diff --git a/src/modal-logic/axioms.lagda.md b/src/modal-logic/axioms.lagda.md index 1d3b62f009..2bb56ee567 100644 --- a/src/modal-logic/axioms.lagda.md +++ b/src/modal-logic/axioms.lagda.md @@ -57,7 +57,7 @@ module _ pr2 (ax-1-parameter h inj f) (a , refl) = is-prop-is-contr ( is-contr-Σ-is-prop a refl - ( λ b → is-set-formula (h a) (h b)) + ( λ b → is-set-modal-formula (h a) (h b)) ( λ x → inj)) ( a , refl) diff --git a/src/modal-logic/canonical-model-theorem.lagda.md b/src/modal-logic/canonical-model-theorem.lagda.md index f368ab32d8..dd310c9297 100644 --- a/src/modal-logic/canonical-model-theorem.lagda.md +++ b/src/modal-logic/canonical-model-theorem.lagda.md @@ -408,7 +408,7 @@ module _ ( λ ((l-ax , l-y) , sub-union , l-ax-sub-logic , l-y-sub-y) → ( apply-universal-property-trunc-Prop ( lists-in-union-lists l-y - ( Id-formula-Prop (¬ₘ a)) + ( Id-modal-formula-Prop (¬ₘ a)) ( unbox-modal-theory theory) ( l-y-sub-y)) ( theory (□ₘ a)) @@ -445,11 +445,11 @@ module _ ( list-subtype l-box)) ( transitive-leq-subtype ( list-subtype l-not-a) - ( Id-formula-Prop (¬ₘ a)) + ( Id-modal-formula-Prop (¬ₘ a)) ( theory-add-formula (¬ₘ a) ( list-subtype l-box)) ( subtype-union-left - ( Id-formula-Prop (¬ₘ a)) + ( Id-modal-formula-Prop (¬ₘ a)) ( list-subtype l-box)) ( l-not-a-sub)) ( subset-add-formula @@ -537,12 +537,12 @@ module _ ( f w))))))) where x : modal-theory (l1 ⊔ l2) i - x = raise-subtype l2 (Id-formula-Prop (¬ₘ a)) + x = raise-subtype l2 (Id-modal-formula-Prop (¬ₘ a)) not-a-in-x : is-in-subtype x ( ¬ₘ a) not-a-in-x = - subset-equiv-subtypes (Id-formula-Prop (¬ₘ a)) x - ( compute-raise-subtype l2 (Id-formula-Prop (¬ₘ a))) + subset-equiv-subtypes (Id-modal-formula-Prop (¬ₘ a)) x + ( compute-raise-subtype l2 (Id-modal-formula-Prop (¬ₘ a))) ( ¬ₘ a) ( refl) @@ -563,13 +563,18 @@ module _ ( weak-modal-logic-closure-monotic ( subtype-union-both logic x ( theory-add-formula (¬ₘ a) logic) - ( subtype-union-right (Id-formula-Prop (¬ₘ a)) logic) - ( transitive-leq-subtype x (Id-formula-Prop (¬ₘ a)) + ( subtype-union-right + ( Id-modal-formula-Prop (¬ₘ a)) logic) + ( transitive-leq-subtype x ( Id-modal-formula-Prop (¬ₘ a)) ( theory-add-formula (¬ₘ a) logic) - ( subtype-union-left (Id-formula-Prop (¬ₘ a)) logic) - ( inv-subset-equiv-subtypes (Id-formula-Prop (¬ₘ a)) x + ( subtype-union-left + ( Id-modal-formula-Prop (¬ₘ a)) + ( logic)) + ( inv-subset-equiv-subtypes + (Id-modal-formula-Prop (¬ₘ a)) + ( x) ( compute-raise-subtype l2 - ( Id-formula-Prop (¬ₘ a)))))) + ( Id-modal-formula-Prop (¬ₘ a)))))) ( ⊥ₘ) ( bot-in-logic)))))) diff --git a/src/modal-logic/closed-under-subformulas-theories.lagda.md b/src/modal-logic/closed-under-subformulas-theories.lagda.md new file mode 100644 index 0000000000..eeb4f8dc8b --- /dev/null +++ b/src/modal-logic/closed-under-subformulas-theories.lagda.md @@ -0,0 +1,95 @@ +# Closed under subformulas theories + +```agda +module modal-logic.closed-under-subformulas-theories where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.binary-relations +open import foundation.binary-relations-transitive-closures +open import foundation.coproduct-types +open import foundation.decidable-types +open import foundation.dependent-pair-types +open import foundation.equivalence-classes +open import foundation.equivalences +open import foundation.existential-quantification +open import foundation.function-extensionality +open import foundation.function-types +open import foundation.identity-types +open import foundation.inhabited-types +open import foundation.logical-equivalences +open import foundation.negation +open import foundation.propositional-extensionality +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.raising-universe-levels +open import foundation.sets +open import foundation.subtypes +open import foundation.unit-type +open import foundation.universe-levels + +open import foundation-core.cartesian-product-types +open import foundation-core.embeddings +open import foundation-core.equality-dependent-pair-types +open import foundation-core.equivalence-relations +open import foundation-core.injective-maps +open import foundation-core.transport-along-identifications + +open import modal-logic.axioms +open import modal-logic.deduction +open import modal-logic.formulas +open import modal-logic.kripke-semantics + +open import univalent-combinatorics.finite-types +``` + +
+ +## Idea + +Theory is closed under subformulas if every subformula of a formula in the +theory is also in the theory. + +## Definition + +```agda +module _ + {l1 l2 : Level} {i : Set l1} (theory : modal-theory l2 i) + where + + is-modal-theory-has-subboxes-Prop : Prop (l1 ⊔ l2) + is-modal-theory-has-subboxes-Prop = + implicit-Π-Prop (modal-formula i) (λ a → theory (□ₘ a) ⇒ theory a) + + is-modal-theory-has-subboxes : UU (l1 ⊔ l2) + is-modal-theory-has-subboxes = type-Prop is-modal-theory-has-subboxes-Prop + + is-modal-theory-has-subimps-Prop : Prop (l1 ⊔ l2) + is-modal-theory-has-subimps-Prop = + implicit-Π-Prop (modal-formula i × modal-formula i) + ( λ (a , b) → theory (a →ₘ b) ⇒ product-Prop (theory a) (theory b)) + + is-modal-theory-has-subimps : UU (l1 ⊔ l2) + is-modal-theory-has-subimps = type-Prop is-modal-theory-has-subimps-Prop + + is-modal-theory-closed-under-subformulas-Prop : Prop (l1 ⊔ l2) + is-modal-theory-closed-under-subformulas-Prop = + product-Prop + ( is-modal-theory-has-subboxes-Prop) + ( is-modal-theory-has-subimps-Prop) + + is-modal-theory-closed-under-subformulas : UU (l1 ⊔ l2) + is-modal-theory-closed-under-subformulas = + type-Prop is-modal-theory-closed-under-subformulas-Prop + + is-has-subboxes-is-closed-under-subformulas : + is-modal-theory-closed-under-subformulas → is-modal-theory-has-subboxes + is-has-subboxes-is-closed-under-subformulas = pr1 + + is-has-subimps-is-closed-under-subformulas : + is-modal-theory-closed-under-subformulas → is-modal-theory-has-subimps + is-has-subimps-is-closed-under-subformulas = pr2 +``` diff --git a/src/modal-logic/decision-procedure.lagda.md b/src/modal-logic/decision-procedure.lagda.md index 5c3e1a7824..e9f32eb2da 100644 --- a/src/modal-logic/decision-procedure.lagda.md +++ b/src/modal-logic/decision-procedure.lagda.md @@ -52,6 +52,7 @@ open import lists.lists open import lists.lists-subtypes open import lists.reversing-lists +open import modal-logic.closed-under-subformulas-theories open import modal-logic.completeness open import modal-logic.deduction open import modal-logic.filtration-lemma @@ -59,6 +60,7 @@ open import modal-logic.formulas open import modal-logic.kripke-models-filtrations open import modal-logic.kripke-semantics open import modal-logic.soundness +open import modal-logic.subformulas open import modal-logic.weak-deduction open import univalent-combinatorics.counting @@ -119,214 +121,29 @@ module _ with decision-procedure' a ... | inl valid-in-C = complete a (λ M M-in-C → valid-in-C (M , M-in-C)) --- TODO: move to kuratowsky-finite-sets -is-kuratowsky-finite-set-Prop' : {l : Level} → Set l → Prop l -is-kuratowsky-finite-set-Prop' X = - exists-structure-Prop (list (type-Set X)) - ( λ l → (x : type-Set X) → type-Prop (list-subtype l x)) - -is-kuratowsky-finite-set' : {l : Level} → Set l → UU l -is-kuratowsky-finite-set' X = type-Prop (is-kuratowsky-finite-set-Prop' X) - -is-kuratowsky-finite-set-is-kuratowsky-finite-set' : - {l : Level} (X : Set l) → - is-kuratowsky-finite-set' X → is-kuratowsky-finite-set X -is-kuratowsky-finite-set-is-kuratowsky-finite-set' X = - map-universal-property-trunc-Prop - ( is-kuratowsky-finite-set-Prop X) - ( λ (l , all-list-subtype) → - ( intro-exists - ( length-list l) - ( pair - ( component-list l) - ( λ x → - ( apply-universal-property-trunc-Prop - ( all-list-subtype x) - ( trunc-Prop (fiber _ x)) - ( λ x-list-subtype → - ( unit-trunc-Prop - ( pair - ( index-in-list x l x-list-subtype) - ( inv - ( eq-component-list-index-in-list x l - ( x-list-subtype))))))))))) - -dependent-map-list : - {l1 l2 : Level} {A : UU l1} {B : UU l2} - (l : list A) (f : Σ A (λ a → a ∈-list l) → B) → - list B -dependent-map-list nil f = nil -dependent-map-list {A = A} {B} (cons x l) f = - cons (f (x , is-head x l)) (dependent-map-list l f') - where - f' : Σ A (λ a → a ∈-list l) → B - f' (a , list-subtype) = f (a , is-in-tail a x l list-subtype) - -in-dependent-map-list : - {l1 l2 : Level} {A : UU l1} {B : UU l2} - {l : list A} (f : Σ A (λ a → a ∈-list l) → B) - {a : A} (a-in-l : a ∈-list l) → - f (a , a-in-l) ∈-list dependent-map-list l f -in-dependent-map-list f (is-head _ l) = is-head _ _ -in-dependent-map-list {A = A} {B} f (is-in-tail _ x l a-in-l) = - is-in-tail _ _ _ (in-dependent-map-list _ a-in-l) - module _ {l : Level} {i : Set l} where - subformulas-list : modal-formula i → list (modal-formula i) - subformulas-list a = cons a (rest a) - where - rest : modal-formula i → list (modal-formula i) - rest (varₘ x) = nil - rest ⊥ₘ = nil - rest (a →ₘ b) = concat-list (subformulas-list a) (subformulas-list b) - rest (□ₘ a) = subformulas-list a - - subformulas : modal-formula i → modal-theory l i - subformulas a = list-subtype (subformulas-list a) - - subformulas-list-has-subimpl : - (a : modal-formula i) {x y : modal-formula i} → - (x →ₘ y) ∈-list subformulas-list a → - (x ∈-list subformulas-list a) × (y ∈-list subformulas-list a) - subformulas-list-has-subimpl .(x →ₘ y) {x} {y} (is-head .(x →ₘ y) _) = - pair - ( is-in-tail x (x →ₘ y) _ - ( in-concat-left - ( subformulas-list x) - ( subformulas-list y) - ( is-head x _))) - ( is-in-tail y (x →ₘ y) _ - ( in-concat-right - ( subformulas-list x) - ( subformulas-list y) - ( is-head y _))) - subformulas-list-has-subimpl - (a →ₘ b) {x} {y} (is-in-tail .(x →ₘ y) .(a →ₘ b) _ xy-list-subtype) - with - in-concat-list - ( subformulas-list a) - ( subformulas-list b) - ( xy-list-subtype) - ... | inl xy-in-left = - let (x-in-tail , y-in-tail) = subformulas-list-has-subimpl a xy-in-left - in pair - ( is-in-tail x (a →ₘ b) _ - ( in-concat-left (subformulas-list a) (subformulas-list b) x-in-tail)) - ( is-in-tail y (a →ₘ b) _ - ( in-concat-left (subformulas-list a) (subformulas-list b) y-in-tail)) - ... | inr xy-in-right = - let (x-in-tail , y-in-tail) = subformulas-list-has-subimpl b xy-in-right - in pair - ( is-in-tail x (a →ₘ b) _ - ( in-concat-right (subformulas-list a) (subformulas-list b) x-in-tail)) - ( is-in-tail y (a →ₘ b) _ - ( in-concat-right (subformulas-list a) (subformulas-list b) y-in-tail)) - subformulas-list-has-subimpl - (□ₘ a) {x} {y} (is-in-tail .(x →ₘ y) .(□ₘ a) _ xy-list-subtype) = - let (x-in-tail , y-in-tail) = - subformulas-list-has-subimpl a xy-list-subtype - in (is-in-tail x (□ₘ a) _ x-in-tail) , (is-in-tail y (□ₘ a) _ y-in-tail) - - subformulas-list-has-subbox : - (a : modal-formula i) {x : modal-formula i} → - □ₘ x ∈-list subformulas-list a → - x ∈-list subformulas-list a - subformulas-list-has-subbox .(□ₘ x) {x} (is-head .(□ₘ x) _) = - is-in-tail x (□ₘ x) _ (is-head x _) - subformulas-list-has-subbox - (a →ₘ b) {x} (is-in-tail .(□ₘ x) .(a →ₘ b) _ x-list-subtype) - with - in-concat-list (subformulas-list a) (subformulas-list b) x-list-subtype - ... | inl x-in-left = - is-in-tail x (a →ₘ b) _ - ( in-concat-left (subformulas-list a) (subformulas-list b) - ( subformulas-list-has-subbox a x-in-left)) - ... | inr x-in-right = - is-in-tail x (a →ₘ b) _ - ( in-concat-right (subformulas-list a) (subformulas-list b) - ( subformulas-list-has-subbox b x-in-right)) - subformulas-list-has-subbox - (□ₘ a) {x} (is-in-tail .(□ₘ x) .(□ₘ a) _ x-list-subtype) = - is-in-tail x (□ₘ a) _ (subformulas-list-has-subbox a x-list-subtype) - - is-modal-theory-closed-under-subformulas-subformulas : - (a : modal-formula i) → - is-modal-theory-closed-under-subformulas i (subformulas a) - is-modal-theory-closed-under-subformulas-subformulas a = - is-modal-theory-closed-under-subformulas-condition - ( i) - ( list-subtype (subformulas-list a)) - ( map-universal-property-trunc-Prop - ( product-Prop - ( list-subtype (subformulas-list a) _) - ( list-subtype (subformulas-list a) _)) - ( λ impl-list-subtype → - ( let (a-list-subtype , b-list-subtype) = - subformulas-list-has-subimpl a impl-list-subtype - in pair - ( unit-trunc-Prop a-list-subtype) - ( unit-trunc-Prop b-list-subtype)))) - ( map-universal-property-trunc-Prop - ( list-subtype (subformulas-list a) _) - ( unit-trunc-Prop ∘ subformulas-list-has-subbox a)) - - subformulas-Set : modal-formula i → Set l - subformulas-Set a = - set-subset (formula-Set i) (list-subtype (subformulas-list a)) - - subformulas-Set-list : - (a : modal-formula i) → list (type-Set (subformulas-Set a)) - subformulas-Set-list a = - dependent-map-list - ( subformulas-list a) - ( λ (x , list-subtype) → x , unit-trunc-Prop list-subtype) - - is-kuratowsky-finite'-subformulas-list : - (a : modal-formula i) → is-kuratowsky-finite-set' (subformulas-Set a) - is-kuratowsky-finite'-subformulas-list a = - intro-exists - ( subformulas-Set-list a) - ( λ (b , trunc-b-list-subtype) → - ( apply-universal-property-trunc-Prop - ( trunc-b-list-subtype) - ( list-subtype (subformulas-Set-list a) (b , trunc-b-list-subtype)) - ( λ b-list-subtype → - ( unit-trunc-Prop - ( tr - ( λ i → (b , i) ∈-list subformulas-Set-list a) - ( eq-is-prop is-prop-type-trunc-Prop) - ( in-dependent-map-list - (λ (x , list-subtype) → x , unit-trunc-Prop list-subtype) - ( b-list-subtype))))))) - - is-kuratowsky-finite-subformulas-list : - (a : modal-formula i) → is-kuratowsky-finite-set (subformulas-Set a) - is-kuratowsky-finite-subformulas-list a = - is-kuratowsky-finite-set-is-kuratowsky-finite-set' - ( subformulas-Set a) - ( is-kuratowsky-finite'-subformulas-list a) - is-finite-subformulas-list : LEM l → (a : modal-formula i) → - is-finite (type-subtype (list-subtype (subformulas-list a))) + is-finite (type-subtype (subformulas a)) is-finite-subformulas-list lem a = is-finite-is-kuratowsky-finite-set - (subformulas-Set a) lem (is-kuratowsky-finite-subformulas-list a) + ( set-subset (modal-formula-Set i) (subformulas a)) + ( lem) + ( is-kuratowsky-finite-subformulas-list a) is-finite-subtypes-subformulas-list : {l2 : Level} → - LEM l → - LEM l2 → + LEM (l ⊔ l2) → (a : modal-formula i) → is-finite (type-subtype (list-subtype (subformulas-list a)) → Prop l2) - is-finite-subtypes-subformulas-list lem lem2 a = + is-finite-subtypes-subformulas-list {l2} lem a = is-finite-function-type - ( is-finite-subformulas-list lem a) - ( is-finite-Prop-LEM lem2) + ( is-finite-subformulas-list (lower-LEM l2 lem) a) + ( is-finite-Prop-LEM (lower-LEM l lem)) module _ {l1 l2 : Level} (A : Set l1) (B : Set l2) @@ -418,25 +235,25 @@ module _ where is-finite-equivalence-classes-filtration : - is-finite (equivalence-class (Φ-equivalence i theory M)) + is-finite (equivalence-class (Φ-equivalence theory M)) is-finite-equivalence-classes-filtration = is-finite-injection - ( equivalence-class-Set (Φ-equivalence i theory M)) + ( equivalence-class-Set (Φ-equivalence theory M)) ( function-Set (type-subtype theory) (Prop-Set (l1 ⊔ l2 ⊔ l4))) ( lem) ( lem) - ( injection-map-function-equivalence-class i theory M) + ( injection-map-function-equivalence-class theory M) ( is-finite-function-type ( is-fin) ( is-finite-Prop-LEM ( lower-LEM (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) lem))) is-small-equivalence-classes-filtration : - (l : Level) → is-small l (equivalence-class (Φ-equivalence i theory M)) + (l : Level) → is-small l (equivalence-class (Φ-equivalence theory M)) is-small-equivalence-classes-filtration l = is-small-is-finite l ( pair - ( equivalence-class (Φ-equivalence i theory M)) + ( equivalence-class (Φ-equivalence theory M)) ( is-finite-equivalence-classes-filtration)) module _ @@ -453,7 +270,7 @@ module _ ( finite-kripke-models l1 l2 i l4 M*) ( λ ((theory , M) , is-fin , is-filt) → ( is-finite-equiv - ( equiv-is-kripke-model-filtration i theory M M* is-filt) + ( equiv-is-kripke-model-filtration theory M M* is-filt) ( is-finite-equivalence-classes-filtration i M ( lower-LEM (l2 ⊔ l4) lem) ( theory) @@ -477,8 +294,8 @@ module _ (filtration-is-filtration : ((M , _) : type-subtype C) (theory : modal-theory l3 i) → - is-modal-theory-closed-under-subformulas i theory → - is-kripke-model-filtration i theory M (filtration theory M)) + is-modal-theory-closed-under-subformulas theory → + is-kripke-model-filtration theory M (filtration theory M)) where is-finite-filtrate-class : @@ -490,7 +307,7 @@ module _ ( λ where (a , M , M-in-C) refl → ( is-finite-equiv - ( equiv-is-kripke-model-filtration i + ( equiv-is-kripke-model-filtration ( subformulas a) ( M) ( M*) @@ -512,7 +329,7 @@ module _ complete a ( λ M M-in-class x → ( backward-implication - ( filtration-lemma i + ( filtration-lemma ( subformulas a) ( is-modal-theory-closed-under-subformulas-subformulas a) ( M) @@ -527,7 +344,7 @@ module _ ( in-logic ( filtration (subformulas a) M) ( intro-exists (a , M , M-in-class) refl) - ( map-equiv-is-kripke-model-filtration i + ( map-equiv-is-kripke-model-filtration ( subformulas a) ( M) ( filtration (subformulas a) M) @@ -535,7 +352,7 @@ module _ ( M , M-in-class) ( subformulas a) ( is-modal-theory-closed-under-subformulas-subformulas a)) - ( class (Φ-equivalence i (subformulas a) M) x))))) + ( class (Φ-equivalence (subformulas a) M) x))))) filtrate-soundness : {l9 l10 : Level} (logic : modal-theory l9 i) → diff --git a/src/modal-logic/deduction.lagda.md b/src/modal-logic/deduction.lagda.md index 04655ee4ae..2b0c57d3eb 100644 --- a/src/modal-logic/deduction.lagda.md +++ b/src/modal-logic/deduction.lagda.md @@ -254,13 +254,14 @@ module _ -- TODO: make Id-formula to be a function for 1 element modal-theory theory-add-formula : modal-theory (l1 ⊔ l2) i - theory-add-formula = (Id-formula-Prop a) ∪ theory + theory-add-formula = (Id-modal-formula-Prop a) ∪ theory formula-in-add-formula : is-in-subtype theory-add-formula a - formula-in-add-formula = subtype-union-left (Id-formula-Prop a) theory a refl + formula-in-add-formula = + subtype-union-left (Id-modal-formula-Prop a) theory a refl subset-add-formula : theory ⊆ theory-add-formula - subset-add-formula = subtype-union-right (Id-formula-Prop a) theory + subset-add-formula = subtype-union-right (Id-modal-formula-Prop a) theory transitive-subset-add-formula : {l3 : Level} (theory' : modal-theory l3 i) → @@ -277,7 +278,7 @@ module _ ((x : modal-formula i) → is-in-subtype theory x → type-Prop (P x)) → (x : modal-formula i) → is-in-subtype theory-add-formula x → type-Prop (P x) elim-theory-add-formula P H-a H-rest = - elim-union-subtype (Id-formula-Prop a) theory P + elim-union-subtype (Id-modal-formula-Prop a) theory P ( λ where .a refl → H-a) ( H-rest) @@ -288,7 +289,7 @@ module _ theory-add-formula ⊆ theory' subset-theory-add-formula theory' a-in = subtype-union-both - ( Id-formula-Prop a) + ( Id-modal-formula-Prop a) ( theory) ( theory') ( λ where .a refl → a-in) @@ -323,7 +324,7 @@ module _ theory ∪ theory-add-formula a theory' ⊆ theory-add-formula a (theory ∪ theory') theory-add-formula-union-right a theory theory' = - union-swap-1-2 theory (Id-formula-Prop a) theory' + union-swap-1-2 theory (Id-modal-formula-Prop a) theory' inv-theory-add-formula-union-right : (a : modal-formula i) @@ -333,5 +334,5 @@ module _ theory-add-formula a (theory ∪ theory') ⊆ theory ∪ theory-add-formula a theory' inv-theory-add-formula-union-right a theory theory' = - union-swap-1-2 (Id-formula-Prop a) theory theory' + union-swap-1-2 (Id-modal-formula-Prop a) theory theory' ``` diff --git a/src/modal-logic/filtration-lemma.lagda.md b/src/modal-logic/filtration-lemma.lagda.md index 4034315e63..54c6434001 100644 --- a/src/modal-logic/filtration-lemma.lagda.md +++ b/src/modal-logic/filtration-lemma.lagda.md @@ -34,6 +34,7 @@ open import foundation-core.embeddings open import foundation-core.equivalence-relations open import foundation-core.invertible-maps +open import modal-logic.closed-under-subformulas-theories open import modal-logic.completeness open import modal-logic.deduction open import modal-logic.formulas @@ -51,14 +52,14 @@ TODO ```agda module _ - {l1 l2 l3 l4 l5 l6 l7 l8 : Level} (i : Set l3) + {l1 l2 l3 l4 l5 l6 l7 l8 : Level} {i : Set l3} (theory : modal-theory l5 i) - (theory-is-closed : is-modal-theory-closed-under-subformulas i theory) + (theory-is-closed : is-modal-theory-closed-under-subformulas theory) (M : kripke-model l1 l2 i l4) (M* : kripke-model l6 l7 i l8) where filtration-lemma : - (is-filtration : is-kripke-model-filtration i theory M M*) + (is-filtration : is-kripke-model-filtration theory M M*) (a : modal-formula i) → is-in-subtype theory a → (x : type-kripke-model i M) → @@ -66,15 +67,14 @@ module _ ( (M , x) ⊨ₘ a) ( pair ( M*) - ( map-equiv-is-kripke-model-filtration i theory M M* is-filtration - ( class (Φ-equivalence i theory M) x)) + ( map-equiv-is-kripke-model-filtration theory M M* is-filtration + ( class (Φ-equivalence theory M) x)) ⊨ₘ a) pr1 (filtration-lemma is-filtration (varₘ n) in-theory x) f = map-raise ( forward-implication ( is-filtration-valuate-is-kripke-model-filtration - ( i) ( theory) ( M) ( M*) @@ -87,53 +87,60 @@ module _ map-raise ∘ map-inv-raise pr1 (filtration-lemma is-filtration (a →ₘ b) in-theory x) fab fa = - forward-implication + let (a-in-theory , b-in-theory) = + is-has-subimps-is-closed-under-subformulas + ( theory) + ( theory-is-closed) + ( in-theory) + in forward-implication ( filtration-lemma is-filtration b - ( pr2 (theory-is-closed in-theory)) + ( b-in-theory) ( x)) ( fab ( backward-implication ( filtration-lemma is-filtration a - ( pr1 (theory-is-closed in-theory)) + ( a-in-theory) ( x)) ( fa))) pr1 (filtration-lemma is-filtration (□ₘ a) in-theory x) f y* r-xy = apply-universal-property-trunc-Prop ( is-inhabited-subtype-equivalence-class - ( Φ-equivalence i theory M) - ( map-inv-equiv-is-kripke-model-filtration i theory M M* + ( Φ-equivalence theory M) + ( map-inv-equiv-is-kripke-model-filtration theory M M* ( is-filtration) ( y*))) ( (M* , y*) ⊨ₘ a) ( λ (y , y-in-class) → - ( let y*-id-class - = concat - ( ap - ( map-equiv-is-kripke-model-filtration i theory M M* - ( is-filtration)) - ( eq-class-equivalence-class - ( Φ-equivalence i theory M) - ( map-inv-equiv-is-kripke-model-filtration i theory M - ( M*) - ( is-filtration) - ( y*)) - ( y-in-class))) - ( y*) - ( is-retraction-map-retraction-map-equiv - ( inv-equiv - ( equiv-is-kripke-model-filtration i theory M M* - ( is-filtration))) - ( y*)) + ( let y*-id-class = + concat + ( ap + ( map-equiv-is-kripke-model-filtration theory M M* + ( is-filtration)) + ( eq-class-equivalence-class + ( Φ-equivalence theory M) + ( map-inv-equiv-is-kripke-model-filtration theory M + ( M*) + ( is-filtration) + ( y*)) + ( y-in-class))) + ( y*) + ( is-retraction-map-retraction-map-equiv + ( inv-equiv + ( equiv-is-kripke-model-filtration theory M M* + ( is-filtration))) + ( y*)) in tr ( λ z* → type-Prop ((M* , z*) ⊨ₘ a)) ( y*-id-class) ( forward-implication ( filtration-lemma is-filtration a - ( theory-is-closed in-theory) + ( is-has-subboxes-is-closed-under-subformulas + ( theory) + ( theory-is-closed) + ( in-theory)) ( y)) ( filtration-relation-upper-bound-is-kripke-model-filtration - ( i) ( theory) ( M) ( M*) @@ -144,9 +151,9 @@ module _ ( y) ( tr ( relation-kripke-model i M* - ( map-equiv-is-kripke-model-filtration i theory M M* + ( map-equiv-is-kripke-model-filtration theory M M* ( is-filtration) - ( class (Φ-equivalence i theory M) x))) + ( class (Φ-equivalence theory M) x))) ( inv y*-id-class) ( r-xy)) (f))))) @@ -154,7 +161,7 @@ module _ f = map-raise ( backward-implication - ( is-filtration-valuate-is-kripke-model-filtration i theory M M* + ( is-filtration-valuate-is-kripke-model-filtration theory M M* ( is-filtration) ( n) ( in-theory) @@ -164,26 +171,34 @@ module _ map-raise ∘ map-inv-raise pr2 (filtration-lemma is-filtration (a →ₘ b) in-theory x) fab fa = - backward-implication + let (a-in-theory , b-in-theory) = + is-has-subimps-is-closed-under-subformulas + ( theory) + ( theory-is-closed) + ( in-theory) + in backward-implication ( filtration-lemma is-filtration b - ( pr2 (theory-is-closed in-theory)) + ( b-in-theory) ( x)) ( fab ( forward-implication ( filtration-lemma is-filtration a - ( pr1 (theory-is-closed in-theory)) + ( a-in-theory) ( x)) ( fa))) pr2 (filtration-lemma is-filtration (□ₘ a) in-theory x) f y r-xy = backward-implication ( filtration-lemma is-filtration a - ( theory-is-closed in-theory) + ( is-has-subboxes-is-closed-under-subformulas + ( theory) + ( theory-is-closed) + ( in-theory)) ( y)) ( f - ( map-equiv-is-kripke-model-filtration i theory M M* is-filtration - ( class (Φ-equivalence i theory M) y)) - ( filtration-relation-lower-bound-is-kripke-model-filtration i theory + ( map-equiv-is-kripke-model-filtration theory M M* is-filtration + ( class (Φ-equivalence theory M) y)) + ( filtration-relation-lower-bound-is-kripke-model-filtration theory ( M) ( M*) ( is-filtration) diff --git a/src/modal-logic/finite-approximability.lagda.md b/src/modal-logic/finite-approximability.lagda.md index 82772920c4..6bab0be54d 100644 --- a/src/modal-logic/finite-approximability.lagda.md +++ b/src/modal-logic/finite-approximability.lagda.md @@ -7,48 +7,31 @@ module modal-logic.finite-approximability where
Imports ```agda -open import foundation.action-on-identifications-functions -open import foundation.binary-relations -open import foundation.coproduct-types -open import foundation.decidable-types open import foundation.dependent-pair-types -open import foundation.equivalence-classes -open import foundation.equivalences open import foundation.existential-quantification -open import foundation.function-types -open import foundation.identity-types -open import foundation.inhabited-types -open import foundation.intersections-subtypes open import foundation.law-of-excluded-middle -open import foundation.logical-equivalences open import foundation.propositional-resizing -open import foundation.propositional-truncations -open import foundation.propositions -open import foundation.raising-universe-levels -open import foundation.sets -open import foundation.subtypes -open import foundation.transport-along-identifications -open import foundation.unions-subtypes open import foundation.unit-type open import foundation.universe-levels open import foundation-core.cartesian-product-types -open import foundation-core.embeddings -open import foundation-core.equivalence-relations -open import foundation-core.invertible-maps +open import foundation-core.identity-types +open import foundation-core.propositions +open import foundation-core.sets +open import foundation-core.subtypes +open import modal-logic.closed-under-subformulas-theories open import modal-logic.completeness open import modal-logic.completeness-k open import modal-logic.completeness-s5 open import modal-logic.decision-procedure open import modal-logic.deduction -open import modal-logic.formulas open import modal-logic.kripke-models-filtrations open import modal-logic.kripke-semantics open import modal-logic.modal-logic-k open import modal-logic.modal-logic-s5 open import modal-logic.soundness -open import modal-logic.weak-deduction +open import modal-logic.subformulas open import order-theory.zorn @@ -94,8 +77,8 @@ module _ kripke-model l6 l7 i l8) (is-filtration : ((M , _) : type-subtype C) (theory : modal-theory l1 i) → - is-modal-theory-closed-under-subformulas i theory → - is-kripke-model-filtration i theory M (filtration theory M)) + is-modal-theory-closed-under-subformulas theory → + is-kripke-model-filtration theory M (filtration theory M)) (logic : modal-theory l9 i) (complete : completeness logic C) (C₂ : model-class l6 l7 i l8 l10) @@ -130,9 +113,9 @@ module _ is-finitely-approximable-K = is-finitely-approximable-filtration ( all-models (lsuc l1) l1 i l1) - ( minimal-kripke-model-filtration _) -- TODO: make i implicit + ( minimal-kripke-model-filtration) ( λ (M , _) theory is-closed → - ( is-kripke-model-filtration-minimal-kripke-model-filtration i + ( is-kripke-model-filtration-minimal-kripke-model-filtration ( theory) ( M) ( is-closed))) @@ -163,10 +146,9 @@ module _ is-finitely-approximable-S5 = is-finitely-approximable-filtration ( equivalence-kripke-models (lsuc l1) l1 i l1) - ( minimal-transitive-kripke-model-filtration _) -- TODO: make i implicit + ( minimal-transitive-kripke-model-filtration) ( λ (M , in-equiv) theory is-closed → ( is-kripke-model-filtration-minimal-transitive-kripke-model-filtration - ( i) ( theory) ( M) ( is-closed) @@ -182,7 +164,6 @@ module _ ( λ where ( a , M , in-equiv) refl → ( minimal-transitive-filtration-preserves-equivalence - _ -- TODO: make implicit ( subformulas a) ( M) ( is-modal-theory-closed-under-subformulas-subformulas a) diff --git a/src/modal-logic/formulas.lagda.md b/src/modal-logic/formulas.lagda.md index a7d5e4a7e0..73012bcf6b 100644 --- a/src/modal-logic/formulas.lagda.md +++ b/src/modal-logic/formulas.lagda.md @@ -134,64 +134,67 @@ module _ refl-Eq-formula (a →ₘ b) = (refl-Eq-formula a) , (refl-Eq-formula b) refl-Eq-formula (□ₘ a) = refl-Eq-formula a - Eq-eq-formula : {a b : modal-formula i} → a = b → Eq-formula a b - Eq-eq-formula {a} refl = refl-Eq-formula a - - eq-Eq-formula : {a b : modal-formula i} → Eq-formula a b → a = b - eq-Eq-formula {varₘ _} {varₘ _} refl = refl - eq-Eq-formula {varₘ _} {⊥ₘ} (map-raise ()) - eq-Eq-formula {varₘ _} {_ →ₘ _} (map-raise ()) - eq-Eq-formula {varₘ _} {□ₘ _} (map-raise ()) - eq-Eq-formula {⊥ₘ} {varₘ _} (map-raise ()) - eq-Eq-formula {⊥ₘ} {⊥ₘ} _ = refl - eq-Eq-formula {⊥ₘ} {_ →ₘ _} (map-raise ()) - eq-Eq-formula {⊥ₘ} {□ₘ _} (map-raise ()) - eq-Eq-formula {_ →ₘ _} {varₘ _} (map-raise ()) - eq-Eq-formula {_ →ₘ _} {⊥ₘ} (map-raise ()) - eq-Eq-formula {a →ₘ b} {c →ₘ d} (eq1 , eq2) = - ap (λ (x , y) → x →ₘ y) (eq-pair (eq-Eq-formula eq1) (eq-Eq-formula eq2)) - eq-Eq-formula {_ →ₘ _} {□ₘ _} (map-raise ()) - eq-Eq-formula {□ₘ _} {varₘ _} (map-raise ()) - eq-Eq-formula {□ₘ _} {⊥ₘ} (map-raise ()) - eq-Eq-formula {□ₘ _} {_ →ₘ _} (map-raise ()) - eq-Eq-formula {□ₘ _} {□ₘ _} eq = ap □ₘ_ (eq-Eq-formula eq) - - is-prop-Eq-formula : (a b : modal-formula i) → is-prop (Eq-formula a b) - is-prop-Eq-formula (varₘ n) (varₘ m) = is-prop-type-Prop (Id-Prop i n m) - is-prop-Eq-formula (varₘ _) ⊥ₘ = is-prop-raise-empty - is-prop-Eq-formula (varₘ _) (_ →ₘ _) = is-prop-raise-empty - is-prop-Eq-formula (varₘ -) (□ₘ _) = is-prop-raise-empty - is-prop-Eq-formula ⊥ₘ (varₘ _) = is-prop-raise-empty - is-prop-Eq-formula ⊥ₘ ⊥ₘ = is-prop-raise-unit - is-prop-Eq-formula ⊥ₘ (_ →ₘ _) = is-prop-raise-empty - is-prop-Eq-formula ⊥ₘ (□ₘ _) = is-prop-raise-empty - is-prop-Eq-formula (_ →ₘ _) (varₘ _) = is-prop-raise-empty - is-prop-Eq-formula (_ →ₘ _) ⊥ₘ = is-prop-raise-empty - is-prop-Eq-formula (a →ₘ b) (c →ₘ d) = - is-prop-product (is-prop-Eq-formula a c) (is-prop-Eq-formula b d) - is-prop-Eq-formula (_ →ₘ _) (□ₘ _) = is-prop-raise-empty - is-prop-Eq-formula (□ₘ _) (varₘ _) = is-prop-raise-empty - is-prop-Eq-formula (□ₘ _) ⊥ₘ = is-prop-raise-empty - is-prop-Eq-formula (□ₘ _) (_ →ₘ _) = is-prop-raise-empty - is-prop-Eq-formula (□ₘ a) (□ₘ c) = is-prop-Eq-formula a c + Eq-eq-modal-formula : {a b : modal-formula i} → a = b → Eq-formula a b + Eq-eq-modal-formula {a} refl = refl-Eq-formula a + + eq-Eq-modal-formula : {a b : modal-formula i} → Eq-formula a b → a = b + eq-Eq-modal-formula {varₘ _} {varₘ _} refl = refl + eq-Eq-modal-formula {varₘ _} {⊥ₘ} (map-raise ()) + eq-Eq-modal-formula {varₘ _} {_ →ₘ _} (map-raise ()) + eq-Eq-modal-formula {varₘ _} {□ₘ _} (map-raise ()) + eq-Eq-modal-formula {⊥ₘ} {varₘ _} (map-raise ()) + eq-Eq-modal-formula {⊥ₘ} {⊥ₘ} _ = refl + eq-Eq-modal-formula {⊥ₘ} {_ →ₘ _} (map-raise ()) + eq-Eq-modal-formula {⊥ₘ} {□ₘ _} (map-raise ()) + eq-Eq-modal-formula {_ →ₘ _} {varₘ _} (map-raise ()) + eq-Eq-modal-formula {_ →ₘ _} {⊥ₘ} (map-raise ()) + eq-Eq-modal-formula {a →ₘ b} {c →ₘ d} (eq1 , eq2) = + ap (λ (x , y) → x →ₘ y) + ( eq-pair (eq-Eq-modal-formula eq1) (eq-Eq-modal-formula eq2)) + eq-Eq-modal-formula {_ →ₘ _} {□ₘ _} (map-raise ()) + eq-Eq-modal-formula {□ₘ _} {varₘ _} (map-raise ()) + eq-Eq-modal-formula {□ₘ _} {⊥ₘ} (map-raise ()) + eq-Eq-modal-formula {□ₘ _} {_ →ₘ _} (map-raise ()) + eq-Eq-modal-formula {□ₘ _} {□ₘ _} eq = ap □ₘ_ (eq-Eq-modal-formula eq) + + is-prop-Eq-modal-formula : (a b : modal-formula i) → is-prop (Eq-formula a b) + is-prop-Eq-modal-formula (varₘ n) (varₘ m) = is-prop-type-Prop (Id-Prop i n m) + is-prop-Eq-modal-formula (varₘ _) ⊥ₘ = is-prop-raise-empty + is-prop-Eq-modal-formula (varₘ _) (_ →ₘ _) = is-prop-raise-empty + is-prop-Eq-modal-formula (varₘ -) (□ₘ _) = is-prop-raise-empty + is-prop-Eq-modal-formula ⊥ₘ (varₘ _) = is-prop-raise-empty + is-prop-Eq-modal-formula ⊥ₘ ⊥ₘ = is-prop-raise-unit + is-prop-Eq-modal-formula ⊥ₘ (_ →ₘ _) = is-prop-raise-empty + is-prop-Eq-modal-formula ⊥ₘ (□ₘ _) = is-prop-raise-empty + is-prop-Eq-modal-formula (_ →ₘ _) (varₘ _) = is-prop-raise-empty + is-prop-Eq-modal-formula (_ →ₘ _) ⊥ₘ = is-prop-raise-empty + is-prop-Eq-modal-formula (a →ₘ b) (c →ₘ d) = + is-prop-product + ( is-prop-Eq-modal-formula a c) + ( is-prop-Eq-modal-formula b d) + is-prop-Eq-modal-formula (_ →ₘ _) (□ₘ _) = is-prop-raise-empty + is-prop-Eq-modal-formula (□ₘ _) (varₘ _) = is-prop-raise-empty + is-prop-Eq-modal-formula (□ₘ _) ⊥ₘ = is-prop-raise-empty + is-prop-Eq-modal-formula (□ₘ _) (_ →ₘ _) = is-prop-raise-empty + is-prop-Eq-modal-formula (□ₘ a) (□ₘ c) = is-prop-Eq-modal-formula a c ``` ### A formula is a set ```agda - is-set-formula : is-set (modal-formula i) - is-set-formula = + is-set-modal-formula : is-set (modal-formula i) + is-set-modal-formula = is-set-prop-in-id ( Eq-formula) - ( is-prop-Eq-formula) + ( is-prop-Eq-modal-formula) ( refl-Eq-formula) - ( λ _ _ → eq-Eq-formula) + ( λ _ _ → eq-Eq-modal-formula) -formula-Set : {l : Level} (i : Set l) → Set l -pr1 (formula-Set i) = modal-formula i -pr2 (formula-Set i) = is-set-formula +modal-formula-Set : {l : Level} (i : Set l) → Set l +pr1 (modal-formula-Set i) = modal-formula i +pr2 (modal-formula-Set i) = is-set-modal-formula -Id-formula-Prop : +Id-modal-formula-Prop : {l : Level} {i : Set l} → modal-formula i → modal-formula i → Prop l -Id-formula-Prop {i = i} = Id-Prop (formula-Set i) +Id-modal-formula-Prop {i = i} = Id-Prop (modal-formula-Set i) ``` diff --git a/src/modal-logic/kripke-models-filtrations.lagda.md b/src/modal-logic/kripke-models-filtrations.lagda.md index 578d339019..b22652967a 100644 --- a/src/modal-logic/kripke-models-filtrations.lagda.md +++ b/src/modal-logic/kripke-models-filtrations.lagda.md @@ -9,6 +9,7 @@ module modal-logic.kripke-models-filtrations where ```agda open import foundation.action-on-identifications-functions open import foundation.binary-relations +open import foundation.binary-relations-transitive-closures open import foundation.coproduct-types open import foundation.decidable-types open import foundation.dependent-pair-types @@ -31,13 +32,13 @@ open import foundation.unit-type open import foundation.universe-levels open import foundation-core.cartesian-product-types -open import foundation-core.embeddings open import foundation-core.equality-dependent-pair-types open import foundation-core.equivalence-relations open import foundation-core.injective-maps open import foundation-core.transport-along-identifications open import modal-logic.axioms +open import modal-logic.closed-under-subformulas-theories open import modal-logic.deduction open import modal-logic.formulas open import modal-logic.kripke-semantics @@ -54,128 +55,8 @@ TODO ## Definition ```agda --- TODO: move to new file module _ - {l1 l2 : Level} {A : UU l1} - where - - data transitive-closure (R : Relation l2 A) : A → A → UU (l1 ⊔ l2) - where - base* : {x y : A} → R x y → transitive-closure R x y - step* : - {x y z : A} → - R x y → - transitive-closure R y z → - transitive-closure R x z - - transitive-closure-Prop : - (R : Relation l2 A) → Relation-Prop (l1 ⊔ l2) A - transitive-closure-Prop R x y = - trunc-Prop (transitive-closure R x y) - - is-transitive-transitive-closure : - (R : Relation l2 A) → is-transitive (transitive-closure R) - is-transitive-transitive-closure R x y z c-yz (base* r-xy) = - step* r-xy c-yz - is-transitive-transitive-closure - R x y z c-yz (step* {y = u} r-xu c-uy) = - step* r-xu (is-transitive-transitive-closure R u y z c-yz c-uy) - - is-transitive-transitive-closure-Prop : - (R : Relation l2 A) → - is-transitive-Relation-Prop (transitive-closure-Prop R) - is-transitive-transitive-closure-Prop R x y z c-yz c-xy = - apply-twice-universal-property-trunc-Prop - ( c-yz) - ( c-xy) - ( transitive-closure-Prop R x z) - ( λ r-yz r-xy → - ( unit-trunc-Prop (is-transitive-transitive-closure R x y z r-yz r-xy))) - - transitive-closure-preserves-reflexivity : - (R : Relation l2 A) → - is-reflexive R → - is-reflexive (transitive-closure R) - transitive-closure-preserves-reflexivity R is-refl x = base* (is-refl x) - - transitive-closure-preserves-symmetry : - (R : Relation l2 A) → - is-symmetric R → - is-symmetric (transitive-closure R) - transitive-closure-preserves-symmetry R is-sym x y (base* r) = - base* (is-sym x y r) - transitive-closure-preserves-symmetry R is-sym x y (step* {y = u} r-xu c-uy) = - is-transitive-transitive-closure R y u x - ( base* (is-sym x u r-xu)) - ( transitive-closure-preserves-symmetry R is-sym u y c-uy) - - transitive-closure-Prop-preserves-reflexivity : - (R : Relation l2 A) → - is-reflexive R → - is-reflexive-Relation-Prop (transitive-closure-Prop R) - transitive-closure-Prop-preserves-reflexivity R is-refl x = - unit-trunc-Prop (transitive-closure-preserves-reflexivity R is-refl x) - - transitive-closure-Prop-preserves-symmetry : - (R : Relation l2 A) → - is-symmetric R → - is-symmetric-Relation-Prop (transitive-closure-Prop R) - transitive-closure-Prop-preserves-symmetry R is-sym x y = - map-universal-property-trunc-Prop - ( transitive-closure-Prop R y x) - ( unit-trunc-Prop ∘ transitive-closure-preserves-symmetry R is-sym x y) - -module _ - {l1 : Level} (i : Set l1) - where - - module _ - {l2 : Level} (theory : modal-theory l2 i) - where - - is-modal-theory-has-subformulas-Prop : modal-formula i → Prop l2 - is-modal-theory-has-subformulas-Prop (varₘ _) = raise-unit-Prop l2 - is-modal-theory-has-subformulas-Prop ⊥ₘ = raise-unit-Prop l2 - is-modal-theory-has-subformulas-Prop (a →ₘ b) = - product-Prop (theory a) (theory b) - is-modal-theory-has-subformulas-Prop (□ₘ a) = theory a - - is-modal-theory-has-subformulas : modal-formula i → UU l2 - is-modal-theory-has-subformulas = - type-Prop ∘ is-modal-theory-has-subformulas-Prop - - is-modal-theory-closed-under-subformulas-Prop : Prop (l1 ⊔ l2) - is-modal-theory-closed-under-subformulas-Prop = - implicit-Π-Prop - ( modal-formula i) - ( λ a → - ( function-Prop - ( is-in-subtype theory a) - ( is-modal-theory-has-subformulas-Prop a))) - - is-modal-theory-closed-under-subformulas : UU (l1 ⊔ l2) - is-modal-theory-closed-under-subformulas = - type-Prop (is-modal-theory-closed-under-subformulas-Prop) - - is-modal-theory-closed-under-subformulas-condition : - ( {a b : modal-formula i} → - is-in-subtype theory (a →ₘ b) → - is-in-subtype theory a × is-in-subtype theory b) → - ( {a : modal-formula i} → - is-in-subtype theory (□ₘ a) → - is-in-subtype theory a) → - is-modal-theory-closed-under-subformulas - is-modal-theory-closed-under-subformulas-condition - h-impl h-box {varₘ n} _ = raise-star - is-modal-theory-closed-under-subformulas-condition - h-impl h-box {⊥ₘ} _ = raise-star - is-modal-theory-closed-under-subformulas-condition - h-impl h-box {a →ₘ b} = h-impl - is-modal-theory-closed-under-subformulas-condition - h-impl h-box {□ₘ a} = h-box - -module _ - {l1 l2 l3 l4 l5 : Level} (i : Set l3) + {l1 l2 l3 l4 l5 : Level} {i : Set l3} (theory : modal-theory l5 i) (M : kripke-model l1 l2 i l4) where @@ -510,13 +391,12 @@ module _ pr2 (pr1 minimal-transitive-kripke-model-filtration) = is-inhabited-equivalence-classes pr1 (pr2 minimal-transitive-kripke-model-filtration) = - transitive-closure-Prop - ( type-Relation-Prop minimal-kripke-model-filtration-relation) + transitive-closure-Prop minimal-kripke-model-filtration-relation pr2 (pr2 minimal-transitive-kripke-model-filtration) = minimal-kripke-model-filtration-valuate module _ - (theory-is-closed : is-modal-theory-closed-under-subformulas i theory) + (theory-is-closed : is-modal-theory-closed-under-subformulas theory) where proof-upper-bound : @@ -534,7 +414,12 @@ module _ ( (M , y) ⊨ₘ a) ( λ ((x' , y') , r-xy' , iff-x , iff-y) → ( backward-implication - ( iff-y a (theory-is-closed box-in-theory)) + -- ( iff-y a (theory-is-closed box-in-theory)) + ( iff-y a + ( is-has-subboxes-is-closed-under-subformulas + ( theory) + ( theory-is-closed) + ( box-in-theory))) ( forward-implication ( iff-x (□ₘ a) box-in-theory) ( x-forces-box) @@ -581,15 +466,17 @@ module _ ( class Φ-equivalence y) → type-Prop ((M , x) ⊨ₘ □ₘ a) → type-Prop ((M , y) ⊨ₘ a) - helper M-is-trans a box-in-theory x y (base* r-xy) x-forces-box = - proof-upper-bound a box-in-theory x y r-xy x-forces-box - helper M-is-trans a box-in-theory x y (step* {y = z*} r-xz c-zy) - x-forces-box = - apply-universal-property-trunc-Prop - ( is-inhabited-subtype-equivalence-class Φ-equivalence z*) - ( (M , y) ⊨ₘ a) - ( λ (z , z-in-z*) → - aux z (eq-class-equivalence-class Φ-equivalence z* z-in-z*)) + helper M-is-trans a box-in-theory x y + (transitive-closure-base r-xy) x-forces-box = + proof-upper-bound a box-in-theory x y r-xy x-forces-box + helper M-is-trans a box-in-theory x y + (transitive-closure-step {y = z*} r-xz c-zy) + x-forces-box = + apply-universal-property-trunc-Prop + ( is-inhabited-subtype-equivalence-class Φ-equivalence z*) + ( (M , y) ⊨ₘ a) + ( λ (z , z-in-z*) → + aux z (eq-class-equivalence-class Φ-equivalence z* z-in-z*)) where aux : (z : type-kripke-model i M) → @@ -640,7 +527,9 @@ module _ ( eq-xy (varₘ n) in-theory) ( map-raise val-n)))))) ( λ (in-theory , val-n) → val-n x (λ _ _ → id , id)))) - ( λ x y r → unit-trunc-Prop (base* (proof-lower-bound x y r))) + ( λ x y r → + unit-trunc-Prop + ( transitive-closure-base (proof-lower-bound x y r))) ( filtration-relation-upper-bound-Prop-minimal-transitive-kripke-model-filtration ( M-is-trans))) @@ -690,7 +579,7 @@ module _ ( minimal-transitive-kripke-model-filtration) minimal-transitive-filtration-preserves-reflexivity is-refl = transitive-closure-Prop-preserves-reflexivity - ( type-Relation-Prop minimal-kripke-model-filtration-relation) + ( minimal-kripke-model-filtration-relation) ( minimal-filtration-preserves-reflexivity is-refl) minimal-transitive-kripke-model-filtration-preserves-symmetry : @@ -704,7 +593,7 @@ module _ ( minimal-transitive-kripke-model-filtration) minimal-transitive-kripke-model-filtration-preserves-symmetry is-sym = transitive-closure-Prop-preserves-symmetry - ( type-Relation-Prop minimal-kripke-model-filtration-relation) + ( minimal-kripke-model-filtration-relation) ( minimal-filtration-preserves-symmetry is-sym) minimal-transitive-kripke-model-filtration-is-transitive : @@ -717,7 +606,7 @@ module _ ( minimal-transitive-kripke-model-filtration) minimal-transitive-kripke-model-filtration-is-transitive = is-transitive-transitive-closure-Prop - ( type-Relation-Prop minimal-kripke-model-filtration-relation) + ( minimal-kripke-model-filtration-relation) minimal-transitive-filtration-preserves-equivalence : is-in-subtype (equivalence-kripke-models l1 l2 i l4) M → @@ -752,5 +641,5 @@ module _ ( λ (theory , M) → ( product ( is-finite (type-subtype theory)) - ( is-kripke-model-filtration i theory M M*))) + ( is-kripke-model-filtration theory M M*))) ``` diff --git a/src/modal-logic/l-complete-theories.lagda.md b/src/modal-logic/l-complete-theories.lagda.md index 4ceb90b95c..349a083394 100644 --- a/src/modal-logic/l-complete-theories.lagda.md +++ b/src/modal-logic/l-complete-theories.lagda.md @@ -48,6 +48,7 @@ open import order-theory.maximal-elements-posets open import order-theory.posets open import order-theory.subposets open import order-theory.subtypes-leq-posets +open import order-theory.upper-bounds-chains-posets open import order-theory.zorn ``` @@ -299,7 +300,7 @@ module _ tr ( λ t → is-in-subtype t (¬ₘ a)) ( eq-is-L-consistent-union-L-complete t - ( Id-formula-Prop (¬ₘ a)) + ( Id-modal-formula-Prop (¬ₘ a)) ( is-L-consistent-add-formula-not-in-logic not-in-logic)) ( formula-in-add-formula (¬ₘ a) theory) diff --git a/src/modal-logic/subformulas.lagda.md b/src/modal-logic/subformulas.lagda.md new file mode 100644 index 0000000000..737a786848 --- /dev/null +++ b/src/modal-logic/subformulas.lagda.md @@ -0,0 +1,146 @@ +# Modal logic subformulas + +```agda +module modal-logic.subformulas where +``` + +
Imports + +```agda +open import foundation.dependent-pair-types +open import foundation.functoriality-propositional-truncation +open import foundation.propositional-truncations +open import foundation.universe-levels + +open import foundation-core.cartesian-product-types +open import foundation-core.coproduct-types +open import foundation-core.function-types +open import foundation-core.sets +open import foundation-core.subtypes + +open import lists.concatenation-lists +open import lists.lists +open import lists.lists-subtypes + +open import modal-logic.closed-under-subformulas-theories +open import modal-logic.deduction +open import modal-logic.formulas + +open import univalent-combinatorics.kuratowsky-finite-sets +``` + +
+ +## Idea + +The subformulas of a modal formula are the formulas that occur in the formula. + +## Definition + +```agda +module _ + {l : Level} {i : Set l} + where + + subformulas-list : modal-formula i → list (modal-formula i) + subformulas-list a = cons a (rest a) + where + rest : modal-formula i → list (modal-formula i) + rest (varₘ x) = nil + rest ⊥ₘ = nil + rest (a →ₘ b) = concat-list (subformulas-list a) (subformulas-list b) + rest (□ₘ a) = subformulas-list a + + subformulas : modal-formula i → modal-theory l i + subformulas a = list-subtype (subformulas-list a) + + subformulas-list-has-subimpl : + (a : modal-formula i) {x y : modal-formula i} → + (x →ₘ y) ∈-list subformulas-list a → + (x ∈-list subformulas-list a) × (y ∈-list subformulas-list a) + subformulas-list-has-subimpl .(x →ₘ y) {x} {y} (is-head .(x →ₘ y) _) = + pair + ( is-in-tail x (x →ₘ y) _ + ( in-concat-left + ( subformulas-list x) + ( subformulas-list y) + ( is-head x _))) + ( is-in-tail y (x →ₘ y) _ + ( in-concat-right + ( subformulas-list x) + ( subformulas-list y) + ( is-head y _))) + subformulas-list-has-subimpl + (a →ₘ b) {x} {y} (is-in-tail .(x →ₘ y) .(a →ₘ b) _ xy-list-subtype) + with + in-concat-list + ( subformulas-list a) + ( subformulas-list b) + ( xy-list-subtype) + ... | inl xy-in-left = + let (x-in-tail , y-in-tail) = subformulas-list-has-subimpl a xy-in-left + in pair + ( is-in-tail x (a →ₘ b) _ + ( in-concat-left (subformulas-list a) (subformulas-list b) x-in-tail)) + ( is-in-tail y (a →ₘ b) _ + ( in-concat-left (subformulas-list a) (subformulas-list b) y-in-tail)) + ... | inr xy-in-right = + let (x-in-tail , y-in-tail) = subformulas-list-has-subimpl b xy-in-right + in pair + ( is-in-tail x (a →ₘ b) _ + ( in-concat-right (subformulas-list a) (subformulas-list b) x-in-tail)) + ( is-in-tail y (a →ₘ b) _ + ( in-concat-right (subformulas-list a) (subformulas-list b) y-in-tail)) + subformulas-list-has-subimpl + (□ₘ a) {x} {y} (is-in-tail .(x →ₘ y) .(□ₘ a) _ xy-list-subtype) = + let (x-in-tail , y-in-tail) = + subformulas-list-has-subimpl a xy-list-subtype + in (is-in-tail x (□ₘ a) _ x-in-tail) , (is-in-tail y (□ₘ a) _ y-in-tail) + + subformulas-list-has-subbox : + (a : modal-formula i) {x : modal-formula i} → + □ₘ x ∈-list subformulas-list a → + x ∈-list subformulas-list a + subformulas-list-has-subbox .(□ₘ x) {x} (is-head .(□ₘ x) _) = + is-in-tail x (□ₘ x) _ (is-head x _) + subformulas-list-has-subbox + (a →ₘ b) {x} (is-in-tail .(□ₘ x) .(a →ₘ b) _ x-list-subtype) + with + in-concat-list (subformulas-list a) (subformulas-list b) x-list-subtype + ... | inl x-in-left = + is-in-tail x (a →ₘ b) _ + ( in-concat-left (subformulas-list a) (subformulas-list b) + ( subformulas-list-has-subbox a x-in-left)) + ... | inr x-in-right = + is-in-tail x (a →ₘ b) _ + ( in-concat-right (subformulas-list a) (subformulas-list b) + ( subformulas-list-has-subbox b x-in-right)) + subformulas-list-has-subbox + (□ₘ a) {x} (is-in-tail .(□ₘ x) .(□ₘ a) _ x-list-subtype) = + is-in-tail x (□ₘ a) _ (subformulas-list-has-subbox a x-list-subtype) + + is-modal-theory-has-subboxes-subformulas : + (a : modal-formula i) → is-modal-theory-has-subboxes (subformulas a) + is-modal-theory-has-subboxes-subformulas a = + map-trunc-Prop (subformulas-list-has-subbox a) + + is-modal-theory-has-subimps-subformulas : + (a : modal-formula i) → is-modal-theory-has-subimps (subformulas a) + is-modal-theory-has-subimps-subformulas a = + map-distributive-trunc-product-Prop ∘ + map-trunc-Prop (subformulas-list-has-subimpl a) + + is-modal-theory-closed-under-subformulas-subformulas : + (a : modal-formula i) → + is-modal-theory-closed-under-subformulas (subformulas a) + is-modal-theory-closed-under-subformulas-subformulas a = + pair + ( is-modal-theory-has-subboxes-subformulas a) + ( is-modal-theory-has-subimps-subformulas a) + + is-kuratowsky-finite-subformulas-list : + (a : modal-formula i) → + is-kuratowsky-finite-set (set-subset (modal-formula-Set i) (subformulas a)) + is-kuratowsky-finite-subformulas-list a = + is-kuratowski-finite-list-subtype (modal-formula-Set i) (subformulas-list a) +``` diff --git a/src/order-theory.lagda.md b/src/order-theory.lagda.md index 2ce1ea63d7..128c153630 100644 --- a/src/order-theory.lagda.md +++ b/src/order-theory.lagda.md @@ -108,6 +108,7 @@ open import order-theory.top-elements-posets public open import order-theory.top-elements-preorders public open import order-theory.total-orders public open import order-theory.total-preorders public +open import order-theory.upper-bounds-chains-posets public open import order-theory.upper-bounds-large-posets public open import order-theory.upper-bounds-posets public open import order-theory.upper-sets-large-posets public diff --git a/src/order-theory/chains-posets.lagda.md b/src/order-theory/chains-posets.lagda.md index 2361ce49d2..c3eaf21e1b 100644 --- a/src/order-theory/chains-posets.lagda.md +++ b/src/order-theory/chains-posets.lagda.md @@ -87,27 +87,4 @@ module _ is-prop (inclusion-chain-Poset C D) is-prop-inclusion-chain-Poset = is-prop-inclusion-chain-Preorder (preorder-Poset X) - --- TODO: move to separate file -module _ - {l1 l2 l3 : Level} (X : Poset l1 l2) (C : chain-Poset l3 X) - where - - is-chain-upper-bound-Prop : (x : type-Poset X) → Prop (l1 ⊔ l2 ⊔ l3) - is-chain-upper-bound-Prop x = - Π-Prop - ( type-chain-Poset X C) - ( λ y → - ( leq-Poset-Prop X - ( inclusion-subtype (sub-preorder-chain-Poset X C) y) - ( x))) - - is-chain-upper-bound : (x : type-Poset X) → UU (l1 ⊔ l2 ⊔ l3) - is-chain-upper-bound = type-Prop ∘ is-chain-upper-bound-Prop - - has-chain-upper-bound-Prop : Prop (l1 ⊔ l2 ⊔ l3) - has-chain-upper-bound-Prop = ∃ (type-Poset X) is-chain-upper-bound-Prop - - has-chain-upper-bound : UU (l1 ⊔ l2 ⊔ l3) - has-chain-upper-bound = type-Prop has-chain-upper-bound-Prop ``` diff --git a/src/order-theory/upper-bounds-chains-posets.lagda.md b/src/order-theory/upper-bounds-chains-posets.lagda.md new file mode 100644 index 0000000000..aba3dbbe68 --- /dev/null +++ b/src/order-theory/upper-bounds-chains-posets.lagda.md @@ -0,0 +1,48 @@ +# Upper bounds of chains in posets + +```agda +module order-theory.upper-bounds-chains-posets where +``` + +
Imports + +```agda +open import foundation.existential-quantification +open import foundation.universe-levels + +open import foundation-core.function-types +open import foundation-core.propositions + +open import order-theory.chains-posets +open import order-theory.posets +open import order-theory.upper-bounds-posets +``` + +
+ +## Idea + +An **upper bound** of a chain `C` in a poset `P` is an element `x` such that for +every element `y` in `C`, `y ≤ x` holds. + +## Definition + +```agda +module _ + {l1 l2 l3 : Level} (X : Poset l1 l2) (C : chain-Poset l3 X) + where + + is-chain-upper-bound-Prop : type-Poset X → Prop (l1 ⊔ l2 ⊔ l3) + is-chain-upper-bound-Prop = + is-upper-bound-family-of-elements-Poset-Prop X + ( type-Poset-type-chain-Poset X C) + + is-chain-upper-bound : type-Poset X → UU (l1 ⊔ l2 ⊔ l3) + is-chain-upper-bound = type-Prop ∘ is-chain-upper-bound-Prop + + has-chain-upper-bound-Prop : Prop (l1 ⊔ l2 ⊔ l3) + has-chain-upper-bound-Prop = ∃ (type-Poset X) is-chain-upper-bound-Prop + + has-chain-upper-bound : UU (l1 ⊔ l2 ⊔ l3) + has-chain-upper-bound = type-Prop has-chain-upper-bound-Prop +``` diff --git a/src/order-theory/zorn.lagda.md b/src/order-theory/zorn.lagda.md index 01977b413f..ecc4021460 100644 --- a/src/order-theory/zorn.lagda.md +++ b/src/order-theory/zorn.lagda.md @@ -22,6 +22,7 @@ open import foundation-core.propositions open import order-theory.chains-posets open import order-theory.maximal-elements-posets open import order-theory.posets +open import order-theory.upper-bounds-chains-posets ```
diff --git a/src/univalent-combinatorics/kuratowsky-finite-sets.lagda.md b/src/univalent-combinatorics/kuratowsky-finite-sets.lagda.md index 4ca2dce32d..793e834905 100644 --- a/src/univalent-combinatorics/kuratowsky-finite-sets.lagda.md +++ b/src/univalent-combinatorics/kuratowsky-finite-sets.lagda.md @@ -12,17 +12,24 @@ open import elementary-number-theory.natural-numbers open import foundation.decidable-equality open import foundation.dependent-pair-types open import foundation.existential-quantification +open import foundation.functoriality-propositional-truncation open import foundation.law-of-excluded-middle open import foundation.logical-equivalences open import foundation.propositional-truncations open import foundation.propositions open import foundation.sets +open import foundation.subtypes open import foundation.surjective-maps open import foundation.universe-levels open import foundation-core.equivalences open import foundation-core.function-types open import foundation-core.identity-types +open import foundation-core.transport-along-identifications + +open import lists.functoriality-lists +open import lists.lists +open import lists.lists-subtypes open import univalent-combinatorics.equality-finite-types open import univalent-combinatorics.finite-types @@ -69,6 +76,78 @@ module _ is-kuratowsky-finite-set-𝔽-Kuratowsky = pr2 X ``` +## Second definition + +```agda +is-kuratowsky-finite-set-list-Prop : {l : Level} → Set l → Prop l +is-kuratowsky-finite-set-list-Prop X = + exists-structure-Prop (list (type-Set X)) + ( λ l → (x : type-Set X) → is-in-subtype (list-subtype l) x) + +is-kuratowsky-finite-set-list : {l : Level} → Set l → UU l +is-kuratowsky-finite-set-list X = + type-Prop (is-kuratowsky-finite-set-list-Prop X) + +is-kuratowsky-finite-set-is-kuratowsky-finite-set-list : + {l : Level} (X : Set l) → + is-kuratowsky-finite-set-list X → is-kuratowsky-finite-set X +is-kuratowsky-finite-set-is-kuratowsky-finite-set-list X = + elim-exists + ( is-kuratowsky-finite-set-Prop X) + ( λ l all-list-subtype → + intro-exists (length-list l) + ( pair + ( component-list l) + ( λ x → + map-trunc-Prop + ( λ in-list → + pair + ( index-in-list x l in-list) + ( inv (eq-component-list-index-in-list x l in-list))) + ( all-list-subtype x)))) + +-- TODO: prove another implication +``` + +### Kuratowky finite subsets + +```agda +is-kuratowsky-finite-subset : + {l1 l2 : Level} (A : Set l1) (B : subtype l2 (type-Set A)) + (l : list (type-Set A)) → + equiv-subtypes B (list-subtype l) → + is-kuratowsky-finite-set (set-subset A B) +is-kuratowsky-finite-subset A B l e = + is-kuratowsky-finite-set-is-kuratowsky-finite-set-list + ( set-subset A B) + ( intro-exists l' + ( λ (a , in-B) → + map-trunc-Prop + ( λ a-in-l → + tr + ( λ i → (a , i) ∈-list l') + ( eq-is-prop (is-prop-type-Prop (B a))) + ( in-dependent-map-list _ a-in-l)) + ( map-equiv (e a) in-B))) + where + l' : list (type-subtype B) + l' = + dependent-map-list l + ( λ a in-list → + a , map-section-map-equiv (e a) (in-list-subtype-in-list in-list)) +``` + +### List subtype is Kuratowky finite + +```agda +is-kuratowski-finite-list-subtype : + {l1 : Level} (A : Set l1) (l : list (type-Set A)) → + is-kuratowsky-finite-set (set-subset A (list-subtype l)) +is-kuratowski-finite-list-subtype A l = + is-kuratowsky-finite-subset A (list-subtype l) l + ( id-equiv-subtypes (list-subtype l)) +``` + ## Properties ### A Kuratowsky finite set is finite if and only if it has decidable equality @@ -92,7 +171,7 @@ has-decidable-equality-is-finite-type-𝔽-Kuratowsky X H = has-decidable-equality-is-finite H ``` -### TODO: change title +### Kuratowsky finite sets are closed under surjections ```agda is-kuratowsky-finite-set-surjection : @@ -101,18 +180,22 @@ is-kuratowsky-finite-set-surjection : is-kuratowsky-finite-set X → is-kuratowsky-finite-set Y is-kuratowsky-finite-set-surjection X Y f = - map-universal-property-trunc-Prop + elim-exists ( is-kuratowsky-finite-set-Prop Y) - ( λ (n , g) → (intro-exists n (surjection-comp f g))) + ( λ n g → intro-exists n (surjection-comp f g)) +``` +### Any finite set is Kuratowsky finite + +```agda is-kuratowsky-finite-set-is-finite : {l : Level} (X : Set l) → is-finite (type-Set X) → is-kuratowsky-finite-set X is-kuratowsky-finite-set-is-finite X = - map-universal-property-trunc-Prop + elim-exists ( is-kuratowsky-finite-set-Prop X) - ( λ (n , e) → intro-exists n (map-equiv e , is-surjective-map-equiv e)) + ( λ n e → intro-exists n (map-equiv e , is-surjective-map-equiv e)) ``` ### Classical facts From 2cc2b803a925f4fab3a13d4a9468830e4baa1d5b Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Fri, 31 May 2024 20:34:09 +0300 Subject: [PATCH 62/63] Refactor filtrations --- src/foundation-core/injective-maps.lagda.md | 5 + src/foundation/injective-maps.lagda.md | 38 ++ .../law-of-excluded-middle.lagda.md | 3 + src/modal-logic.lagda.md | 3 + src/modal-logic/decision-procedure.lagda.md | 292 +-------------- .../filtrated-kripke-classes.lagda.md | 137 +++++++ .../finite-approximability.lagda.md | 2 +- .../kripke-models-filtrations.lagda.md | 340 +----------------- .../minimal-kripke-filtration.lagda.md | 163 +++++++++ ...imal-transitive-kripke-filtration.lagda.md | 233 ++++++++++++ src/modal-logic/subformulas.lagda.md | 34 ++ src/univalent-combinatorics.lagda.md | 1 + .../kuratowsky-finite-sets.lagda.md | 11 + .../subfinite-types.lagda.md | 120 +++++++ 14 files changed, 770 insertions(+), 612 deletions(-) create mode 100644 src/modal-logic/filtrated-kripke-classes.lagda.md create mode 100644 src/modal-logic/minimal-kripke-filtration.lagda.md create mode 100644 src/modal-logic/minimal-transitive-kripke-filtration.lagda.md create mode 100644 src/univalent-combinatorics/subfinite-types.lagda.md diff --git a/src/foundation-core/injective-maps.lagda.md b/src/foundation-core/injective-maps.lagda.md index ea67f3fc16..5abe1559c3 100644 --- a/src/foundation-core/injective-maps.lagda.md +++ b/src/foundation-core/injective-maps.lagda.md @@ -90,6 +90,11 @@ module _ is-injective h → is-injective g → is-injective (g ∘ h) is-injective-comp is-inj-h is-inj-g = is-inj-h ∘ is-inj-g + injection-comp : + injection A B → injection B C → injection A C + injection-comp (f , is-inj-f) (g , is-inj-g) = + g ∘ f , is-injective-comp is-inj-f is-inj-g + is-injective-left-map-triangle : (f : A → C) (g : B → C) (h : A → B) → f ~ (g ∘ h) → is-injective h → is-injective g → is-injective f diff --git a/src/foundation/injective-maps.lagda.md b/src/foundation/injective-maps.lagda.md index 38991f22b2..d8e2ce7c56 100644 --- a/src/foundation/injective-maps.lagda.md +++ b/src/foundation/injective-maps.lagda.md @@ -9,16 +9,23 @@ open import foundation-core.injective-maps public
Imports ```agda +open import foundation.decidable-types open import foundation.dependent-pair-types +open import foundation.functoriality-propositional-truncation +open import foundation.inhabited-types open import foundation.logical-equivalences +open import foundation.surjective-maps open import foundation.universe-levels +open import foundation-core.coproduct-types open import foundation-core.embeddings open import foundation-core.empty-types +open import foundation-core.fibers-of-maps open import foundation-core.identity-types open import foundation-core.negation open import foundation-core.propositional-maps open import foundation-core.propositions +open import foundation-core.retractions open import foundation-core.sets ``` @@ -104,3 +111,34 @@ module _ pr1 (is-injective-Prop H f) = is-injective f pr2 (is-injective-Prop H f) = is-prop-is-injective H f ``` + +### TODO: Title + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + (f : A → B) + (is-injective-f : is-injective f) + (dec : (b : B) → is-decidable (fiber f b)) + where + + retraction-is-injective : A → retraction f + pr1 (retraction-is-injective a0) b with dec b + ... | inl (a , _) = a + ... | inr _ = a0 + pr2 (retraction-is-injective _) a with dec (f a) + ... | inl (a , eq) = is-injective-f eq + ... | inr contra = ex-falso (contra (a , refl)) + + surjective-is-injective : A → B ↠ A + pr1 (surjective-is-injective a) = + map-retraction f (retraction-is-injective a) + pr2 (surjective-is-injective a) = + is-surjective-has-section + ( pair f + ( is-retraction-map-retraction f + ( retraction-is-injective a))) + + is-inhabited-inv-surjections : is-inhabited A → is-inhabited (B ↠ A) + is-inhabited-inv-surjections = map-trunc-Prop surjective-is-injective +``` diff --git a/src/foundation/law-of-excluded-middle.lagda.md b/src/foundation/law-of-excluded-middle.lagda.md index b0b59d0a22..bad4ef5d35 100644 --- a/src/foundation/law-of-excluded-middle.lagda.md +++ b/src/foundation/law-of-excluded-middle.lagda.md @@ -47,6 +47,9 @@ module _ LEM : UU (lsuc l) LEM = type-Prop LEM-Prop + +apply-LEM : {l : Level} → LEM l → {P : UU l} → is-prop P → is-decidable P +apply-LEM lem {P} is-prop-P = lem (P , is-prop-P) ``` ## Properties diff --git a/src/modal-logic.lagda.md b/src/modal-logic.lagda.md index 28ed4dfca1..eda8e000aa 100644 --- a/src/modal-logic.lagda.md +++ b/src/modal-logic.lagda.md @@ -12,6 +12,7 @@ open import modal-logic.completeness-k public open import modal-logic.completeness-s5 public open import modal-logic.decision-procedure public open import modal-logic.deduction public +open import modal-logic.filtrated-kripke-classes public open import modal-logic.filtration-lemma public open import modal-logic.finite-approximability public open import modal-logic.formulas public @@ -21,6 +22,8 @@ open import modal-logic.kripke-semantics public open import modal-logic.l-complete-theories public open import modal-logic.l-consistent-theories public open import modal-logic.lindenbaum public +open import modal-logic.minimal-kripke-filtration public +open import modal-logic.minimal-transitive-kripke-filtration public open import modal-logic.modal-logic-k public open import modal-logic.modal-logic-s5 public open import modal-logic.soundness public diff --git a/src/modal-logic/decision-procedure.lagda.md b/src/modal-logic/decision-procedure.lagda.md index e9f32eb2da..2cee89ab6f 100644 --- a/src/modal-logic/decision-procedure.lagda.md +++ b/src/modal-logic/decision-procedure.lagda.md @@ -7,69 +7,27 @@ module modal-logic.decision-procedure where
Imports ```agda -open import elementary-number-theory.natural-numbers - -open import foundation.action-on-identifications-binary-functions -open import foundation.action-on-identifications-dependent-functions -open import foundation.action-on-identifications-functions open import foundation.booleans -open import foundation.contractible-types open import foundation.decidable-types open import foundation.dependent-pair-types -open import foundation.embeddings -open import foundation.empty-types -open import foundation.equivalence-classes -open import foundation.equivalences -open import foundation.existential-quantification -open import foundation.function-extensionality -open import foundation.inhabited-types open import foundation.law-of-excluded-middle open import foundation.logical-equivalences -open import foundation.propositional-extensionality -open import foundation.propositional-truncations -open import foundation.raising-universe-levels -open import foundation.sets -open import foundation.subtypes -open import foundation.surjective-maps -open import foundation.transport-along-identifications -open import foundation.unions-subtypes open import foundation.unit-type open import foundation.universe-levels -open import foundation-core.cartesian-product-types open import foundation-core.coproduct-types -open import foundation-core.dependent-identifications -open import foundation-core.equality-dependent-pair-types -open import foundation-core.fibers-of-maps -open import foundation-core.function-types -open import foundation-core.identity-types -open import foundation-core.injective-maps open import foundation-core.propositions +open import foundation-core.sets +open import foundation-core.subtypes -open import lists.arrays -open import lists.concatenation-lists -open import lists.lists -open import lists.lists-subtypes -open import lists.reversing-lists - -open import modal-logic.closed-under-subformulas-theories open import modal-logic.completeness open import modal-logic.deduction -open import modal-logic.filtration-lemma open import modal-logic.formulas -open import modal-logic.kripke-models-filtrations open import modal-logic.kripke-semantics open import modal-logic.soundness -open import modal-logic.subformulas -open import modal-logic.weak-deduction -open import univalent-combinatorics.counting open import univalent-combinatorics.decidable-dependent-function-types open import univalent-combinatorics.finite-types -open import univalent-combinatorics.function-types -open import univalent-combinatorics.kuratowsky-finite-sets -open import univalent-combinatorics.small-types -open import univalent-combinatorics.standard-finite-types ```
@@ -120,250 +78,4 @@ module _ pr2 (decision-procedure-correctness theory sound complete a) _ with decision-procedure' a ... | inl valid-in-C = complete a (λ M M-in-C → valid-in-C (M , M-in-C)) - -module _ - {l : Level} {i : Set l} - where - - is-finite-subformulas-list : - LEM l → - (a : modal-formula i) → - is-finite (type-subtype (subformulas a)) - is-finite-subformulas-list lem a = - is-finite-is-kuratowsky-finite-set - ( set-subset (modal-formula-Set i) (subformulas a)) - ( lem) - ( is-kuratowsky-finite-subformulas-list a) - - is-finite-subtypes-subformulas-list : - {l2 : Level} → - LEM (l ⊔ l2) → - (a : modal-formula i) → - is-finite (type-subtype (list-subtype (subformulas-list a)) → Prop l2) - is-finite-subtypes-subformulas-list {l2} lem a = - is-finite-function-type - ( is-finite-subformulas-list (lower-LEM l2 lem) a) - ( is-finite-Prop-LEM (lower-LEM l lem)) - -module _ - {l1 l2 : Level} (A : Set l1) (B : Set l2) - (lem : LEM l1) - (lem2 : LEM (l1 ⊔ l2)) - where - - surjection-from-injection : - type-Set A → - injection (type-Set A) (type-Set B) → - type-Set B → type-Set A - surjection-from-injection a (f , is-inj) b - with - lem2 - ( pair - ( Σ (type-Set A) (λ a → f a = b)) - ( is-prop-all-elements-equal - ( λ x y → - ( eq-pair-Σ - ( is-inj (pr2 x ∙ inv (pr2 y))) - ( eq-is-prop (is-set-type-Set B _ _)))))) - ... | inl x = pr1 x - ... | inr x = a - - is-invertible-surjection-from-injection : - (a0 : type-Set A) → - (inj@(f , _) : injection (type-Set A) (type-Set B)) → - (a : type-Set A) → - surjection-from-injection a0 inj (f a) = a - is-invertible-surjection-from-injection a0 (f , is-inj) a - with - lem2 - ( pair - ( Σ (type-Set A) (λ a' → f a' = f a)) - ( is-prop-all-elements-equal - ( λ x y → - ( eq-pair-Σ - ( is-inj (pr2 x ∙ inv (pr2 y))) - ( eq-is-prop (is-set-type-Set B _ _)))))) - ... | inl x = is-inj (pr2 x) - ... | inr x = ex-falso (x (a , refl)) - - is-surjective-surjection-from-injection : - (a : type-Set A) → - (inj : injection (type-Set A) (type-Set B)) → - is-surjective (surjection-from-injection a inj) - is-surjective-surjection-from-injection a0 (f , is-inj) a = - unit-trunc-Prop - ( f a , is-invertible-surjection-from-injection a0 (f , is-inj) a) - - kuratowsky-finite-set-injection : - injection (type-Set A) (type-Set B) → - is-kuratowsky-finite-set B → - is-kuratowsky-finite-set A - kuratowsky-finite-set-injection inj b-is-fin - with lem (is-inhabited-Prop (type-Set A)) - ... | inl x = - apply-twice-universal-property-trunc-Prop - ( x) - ( b-is-fin) - ( is-kuratowsky-finite-set-Prop A) - ( λ a (n , e) → - ( unit-trunc-Prop - ( triple - ( n) - ( surjection-from-injection a inj ∘ pr1 e) - ( is-surjective-comp - ( is-surjective-surjection-from-injection a inj) - ( pr2 e))))) - ... | inr x = - is-kuratowsky-finite-set-is-finite A - ( is-finite-is-empty (x ∘ unit-trunc-Prop)) - - is-finite-injection : - injection (type-Set A) (type-Set B) → - is-finite (type-Set B) → - is-finite (type-Set A) - is-finite-injection inj b-is-fin = - is-finite-is-kuratowsky-finite-set A - lem (kuratowsky-finite-set-injection inj - ( is-kuratowsky-finite-set-is-finite B b-is-fin)) - -module _ - {l1 l2 l3 l4 l5 : Level} (i : Set l3) - (M : kripke-model l1 l2 i l4) - (lem : LEM (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5)) - (theory : modal-theory l5 i) - (is-fin : is-finite (type-subtype theory)) - where - - is-finite-equivalence-classes-filtration : - is-finite (equivalence-class (Φ-equivalence theory M)) - is-finite-equivalence-classes-filtration = - is-finite-injection - ( equivalence-class-Set (Φ-equivalence theory M)) - ( function-Set (type-subtype theory) (Prop-Set (l1 ⊔ l2 ⊔ l4))) - ( lem) - ( lem) - ( injection-map-function-equivalence-class theory M) - ( is-finite-function-type - ( is-fin) - ( is-finite-Prop-LEM - ( lower-LEM (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3 ⊔ lsuc l4 ⊔ lsuc l5) lem))) - - is-small-equivalence-classes-filtration : - (l : Level) → is-small l (equivalence-class (Φ-equivalence theory M)) - is-small-equivalence-classes-filtration l = - is-small-is-finite l - ( pair - ( equivalence-class (Φ-equivalence theory M)) - ( is-finite-equivalence-classes-filtration)) - -module _ - {l1 l2 l3 l4 : Level} (i : Set l3) - (l5 l6 l7 l8 : Level) - (lem : LEM (l2 ⊔ lsuc l3 ⊔ l4 ⊔ lsuc l5 ⊔ lsuc l6 ⊔ lsuc l7 ⊔ lsuc l8)) - where - - filtration-models-subset-finite-kripke-models : - filtration-models l1 l2 i l4 l5 l6 l7 l8 ⊆ - finite-kripke-models l1 l2 i l4 - filtration-models-subset-finite-kripke-models M* = - map-universal-property-trunc-Prop - ( finite-kripke-models l1 l2 i l4 M*) - ( λ ((theory , M) , is-fin , is-filt) → - ( is-finite-equiv - ( equiv-is-kripke-model-filtration theory M M* is-filt) - ( is-finite-equivalence-classes-filtration i M - ( lower-LEM (l2 ⊔ l4) lem) - ( theory) - ( is-fin)))) - -module _ - {l1 l2 l3 l4 l5 l6 l7 l8 : Level} (i : Set l3) - (C : model-class l1 l2 i l4 l5) - (filtration : modal-theory l3 i → - kripke-model l1 l2 i l4 → - kripke-model l6 l7 i l8) - where - - filtrate-class : - model-class l6 l7 i l8 ( l3 ⊔ l5 ⊔ lsuc (l1 ⊔ l2 ⊔ l4 ⊔ l6 ⊔ l7 ⊔ l8)) - filtrate-class M* = - exists-structure-Prop (modal-formula i × type-subtype C) - ( λ (a , M , _) → M* = filtration (subformulas a) M) - - module _ - (filtration-is-filtration : - ((M , _) : type-subtype C) - (theory : modal-theory l3 i) → - is-modal-theory-closed-under-subformulas theory → - is-kripke-model-filtration theory M (filtration theory M)) - where - - is-finite-filtrate-class : - LEM (lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4)) → - filtrate-class ⊆ finite-kripke-models l6 l7 i l8 - is-finite-filtrate-class lem M* = - elim-exists - ( finite-kripke-models l6 l7 i l8 M*) - ( λ where - (a , M , M-in-C) refl → - ( is-finite-equiv - ( equiv-is-kripke-model-filtration - ( subformulas a) - ( M) - ( M*) - ( filtration-is-filtration - ( M , M-in-C) - ( subformulas a) - ( is-modal-theory-closed-under-subformulas-subformulas a))) - ( is-finite-equivalence-classes-filtration i M lem - ( subformulas a) - ( is-finite-subformulas-list - ( lower-LEM (lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4)) lem) - ( a))))) - - filtrate-completeness : - {l9 : Level} (logic : modal-theory l9 i) → - completeness logic C → - completeness logic filtrate-class - filtrate-completeness logic complete a in-logic = - complete a - ( λ M M-in-class x → - ( backward-implication - ( filtration-lemma - ( subformulas a) - ( is-modal-theory-closed-under-subformulas-subformulas a) - ( M) - ( filtration (subformulas a) M) - ( filtration-is-filtration - ( M , M-in-class) - ( subformulas a) - ( is-modal-theory-closed-under-subformulas-subformulas a)) - ( a) - ( head-in-list-subtype) - ( x)) - ( in-logic - ( filtration (subformulas a) M) - ( intro-exists (a , M , M-in-class) refl) - ( map-equiv-is-kripke-model-filtration - ( subformulas a) - ( M) - ( filtration (subformulas a) M) - ( filtration-is-filtration - ( M , M-in-class) - ( subformulas a) - ( is-modal-theory-closed-under-subformulas-subformulas a)) - ( class (Φ-equivalence (subformulas a) M) x))))) - - filtrate-soundness : - {l9 l10 : Level} (logic : modal-theory l9 i) → - (C₂ : model-class l6 l7 i l8 l10) → - filtrate-class ⊆ C₂ → - soundness logic C₂ → - soundness logic filtrate-class - filtrate-soundness logic C₂ H = - transitive-leq-subtype - ( logic) - ( class-modal-logic C₂) - ( class-modal-logic filtrate-class) - ( class-modal-logic-monotic filtrate-class C₂ H) ``` diff --git a/src/modal-logic/filtrated-kripke-classes.lagda.md b/src/modal-logic/filtrated-kripke-classes.lagda.md new file mode 100644 index 0000000000..74498cab63 --- /dev/null +++ b/src/modal-logic/filtrated-kripke-classes.lagda.md @@ -0,0 +1,137 @@ +# Filtrated Kripke Classes + +```agda +module modal-logic.filtrated-kripke-classes where +``` + +
Imports + +```agda +open import foundation.dependent-pair-types +open import foundation.equivalence-classes +open import foundation.existential-quantification +open import foundation.law-of-excluded-middle +open import foundation.logical-equivalences +open import foundation.universe-levels + +open import foundation-core.cartesian-product-types +open import foundation-core.identity-types +open import foundation-core.propositions +open import foundation-core.sets +open import foundation-core.subtypes + +open import lists.lists-subtypes + +open import modal-logic.closed-under-subformulas-theories +open import modal-logic.completeness +open import modal-logic.deduction +open import modal-logic.filtration-lemma +open import modal-logic.formulas +open import modal-logic.kripke-models-filtrations +open import modal-logic.kripke-semantics +open import modal-logic.soundness +open import modal-logic.subformulas + +open import univalent-combinatorics.finite-types +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l1 l2 l3 l4 l5 l6 l7 l8 : Level} (i : Set l3) + (C : model-class l1 l2 i l4 l5) + (filtration : modal-theory l3 i → + kripke-model l1 l2 i l4 → + kripke-model l6 l7 i l8) + where + + filtrate-class : + model-class l6 l7 i l8 ( l3 ⊔ l5 ⊔ lsuc (l1 ⊔ l2 ⊔ l4 ⊔ l6 ⊔ l7 ⊔ l8)) + filtrate-class M* = + exists-structure-Prop (modal-formula i × type-subtype C) + ( λ (a , M , _) → M* = filtration (subformulas a) M) + + module _ + (filtration-is-filtration : + ((M , _) : type-subtype C) + (theory : modal-theory l3 i) → + is-modal-theory-closed-under-subformulas theory → + is-kripke-model-filtration theory M (filtration theory M)) + where + + is-finite-filtrate-class : + LEM (lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4)) → + filtrate-class ⊆ finite-kripke-models l6 l7 i l8 + is-finite-filtrate-class lem M* = + elim-exists + ( finite-kripke-models l6 l7 i l8 M*) + ( λ where + (a , M , M-in-C) refl → + ( is-finite-equiv + ( equiv-is-kripke-model-filtration + ( subformulas a) + ( M) + ( M*) + ( filtration-is-filtration + ( M , M-in-C) + ( subformulas a) + ( is-modal-theory-closed-under-subformulas-subformulas a))) + ( is-finite-equivalence-classes-filtration i M lem + ( subformulas a) + ( is-finite-subformulas-list + ( lower-LEM (lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4)) lem) + ( a))))) + + filtrate-completeness : + {l9 : Level} (logic : modal-theory l9 i) → + completeness logic C → + completeness logic filtrate-class + filtrate-completeness logic complete a in-logic = + complete a + ( λ M M-in-class x → + ( backward-implication + ( filtration-lemma + ( subformulas a) + ( is-modal-theory-closed-under-subformulas-subformulas a) + ( M) + ( filtration (subformulas a) M) + ( filtration-is-filtration + ( M , M-in-class) + ( subformulas a) + ( is-modal-theory-closed-under-subformulas-subformulas a)) + ( a) + ( head-in-list-subtype) + ( x)) + ( in-logic + ( filtration (subformulas a) M) + ( intro-exists (a , M , M-in-class) refl) + ( map-equiv-is-kripke-model-filtration + ( subformulas a) + ( M) + ( filtration (subformulas a) M) + ( filtration-is-filtration + ( M , M-in-class) + ( subformulas a) + ( is-modal-theory-closed-under-subformulas-subformulas a)) + ( class (Φ-equivalence (subformulas a) M) x))))) + + filtrate-soundness : + {l9 l10 : Level} (logic : modal-theory l9 i) → + (C₂ : model-class l6 l7 i l8 l10) → + filtrate-class ⊆ C₂ → + soundness logic C₂ → + soundness logic filtrate-class + filtrate-soundness logic C₂ H = + transitive-leq-subtype + ( logic) + ( class-modal-logic C₂) + ( class-modal-logic filtrate-class) + ( class-modal-logic-monotic filtrate-class C₂ H) +``` diff --git a/src/modal-logic/finite-approximability.lagda.md b/src/modal-logic/finite-approximability.lagda.md index 6bab0be54d..7e20523fb6 100644 --- a/src/modal-logic/finite-approximability.lagda.md +++ b/src/modal-logic/finite-approximability.lagda.md @@ -148,7 +148,7 @@ module _ ( equivalence-kripke-models (lsuc l1) l1 i l1) ( minimal-transitive-kripke-model-filtration) ( λ (M , in-equiv) theory is-closed → - ( is-kripke-model-filtration-minimal-transitive-kripke-model-filtration + ( is-filtration-minimal-transitive-kripke-model-filtration ( theory) ( M) ( is-closed) diff --git a/src/modal-logic/kripke-models-filtrations.lagda.md b/src/modal-logic/kripke-models-filtrations.lagda.md index b22652967a..8dc5679228 100644 --- a/src/modal-logic/kripke-models-filtrations.lagda.md +++ b/src/modal-logic/kripke-models-filtrations.lagda.md @@ -8,33 +8,25 @@ module modal-logic.kripke-models-filtrations where ```agda open import foundation.action-on-identifications-functions -open import foundation.binary-relations -open import foundation.binary-relations-transitive-closures -open import foundation.coproduct-types -open import foundation.decidable-types open import foundation.dependent-pair-types open import foundation.equivalence-classes -open import foundation.equivalences open import foundation.existential-quantification open import foundation.function-extensionality -open import foundation.function-types -open import foundation.identity-types -open import foundation.inhabited-types +open import foundation.law-of-excluded-middle open import foundation.logical-equivalences -open import foundation.negation open import foundation.propositional-extensionality open import foundation.propositional-truncations -open import foundation.propositions -open import foundation.raising-universe-levels open import foundation.sets -open import foundation.subtypes -open import foundation.unit-type open import foundation.universe-levels open import foundation-core.cartesian-product-types -open import foundation-core.equality-dependent-pair-types open import foundation-core.equivalence-relations +open import foundation-core.equivalences +open import foundation-core.function-types +open import foundation-core.identity-types open import foundation-core.injective-maps +open import foundation-core.propositions +open import foundation-core.subtypes open import foundation-core.transport-along-identifications open import modal-logic.axioms @@ -44,6 +36,8 @@ open import modal-logic.formulas open import modal-logic.kripke-semantics open import univalent-combinatorics.finite-types +open import univalent-combinatorics.function-types +open import univalent-combinatorics.subfinite-types ```
@@ -332,314 +326,18 @@ module _ ( x) ( is-refl x)))) - is-inhabited-equivalence-classes : - is-inhabited (equivalence-class Φ-equivalence) - is-inhabited-equivalence-classes = - map-is-inhabited - ( class Φ-equivalence) - ( is-inhabited-type-kripke-model i M) - - minimal-kripke-model-filtration-relation : - Relation-Prop (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) (equivalence-class Φ-equivalence) - minimal-kripke-model-filtration-relation x* y* = - exists-structure-Prop - ( type-kripke-model i M × type-kripke-model i M) - ( λ (x , y) → - ( product - ( relation-kripke-model i M x y) - ( product - ( is-in-equivalence-class Φ-equivalence x* x) - ( is-in-equivalence-class Φ-equivalence y* y)))) - - minimal-kripke-model-filtration-valuate : - type-Set i → - equivalence-class Φ-equivalence → - Prop (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) - minimal-kripke-model-filtration-valuate n x* = - product-Prop - ( theory (varₘ n)) - ( Π-Prop - ( type-kripke-model i M) - ( λ x → - ( function-Prop - ( is-in-equivalence-class Φ-equivalence x* x) - ( valuate-kripke-model i M n x)))) - - minimal-kripke-model-filtration : - kripke-model - ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) - ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) - ( i) - ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) - pr1 (pr1 minimal-kripke-model-filtration) = - equivalence-class Φ-equivalence - pr2 (pr1 minimal-kripke-model-filtration) = - is-inhabited-equivalence-classes - pr1 (pr2 minimal-kripke-model-filtration) = - minimal-kripke-model-filtration-relation - pr2 (pr2 minimal-kripke-model-filtration) = - minimal-kripke-model-filtration-valuate - - minimal-transitive-kripke-model-filtration : - kripke-model - ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) - ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) - ( i) - ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) - pr1 (pr1 minimal-transitive-kripke-model-filtration) = - equivalence-class Φ-equivalence - pr2 (pr1 minimal-transitive-kripke-model-filtration) = - is-inhabited-equivalence-classes - pr1 (pr2 minimal-transitive-kripke-model-filtration) = - transitive-closure-Prop minimal-kripke-model-filtration-relation - pr2 (pr2 minimal-transitive-kripke-model-filtration) = - minimal-kripke-model-filtration-valuate - - module _ - (theory-is-closed : is-modal-theory-closed-under-subformulas theory) - where - - proof-upper-bound : - (a : modal-formula i) → - is-in-subtype theory (□ₘ a) → - (x y : type-kripke-model i M) → - relation-kripke-model i minimal-kripke-model-filtration - ( class Φ-equivalence x) - ( class Φ-equivalence y) → - type-Prop ((M , x) ⊨ₘ □ₘ a) → - type-Prop ((M , y) ⊨ₘ a) - proof-upper-bound a box-in-theory x y r-xy x-forces-box = - apply-universal-property-trunc-Prop - ( r-xy) - ( (M , y) ⊨ₘ a) - ( λ ((x' , y') , r-xy' , iff-x , iff-y) → - ( backward-implication - -- ( iff-y a (theory-is-closed box-in-theory)) - ( iff-y a - ( is-has-subboxes-is-closed-under-subformulas - ( theory) - ( theory-is-closed) - ( box-in-theory))) - ( forward-implication - ( iff-x (□ₘ a) box-in-theory) - ( x-forces-box) - ( y') - ( r-xy')))) - - proof-lower-bound : - (x y : type-kripke-model i M) → - relation-kripke-model i M x y → - type-Prop - ( minimal-kripke-model-filtration-relation - ( class Φ-equivalence x) - ( class Φ-equivalence y)) - proof-lower-bound x y r = - intro-exists (x , y) (r , (λ _ _ → id , id) , (λ _ _ → id , id)) - - is-kripke-model-filtration-minimal-kripke-model-filtration : - is-kripke-model-filtration minimal-kripke-model-filtration - is-kripke-model-filtration-minimal-kripke-model-filtration = - pair id-equiv - ( triple - ( λ n in-theory x → - ( pair - ( λ val-n → - ( pair - ( in-theory) - ( λ y eq-xy → - ( map-inv-raise - ( forward-implication - ( eq-xy (varₘ n) in-theory) - ( map-raise val-n)))))) - ( λ (in-theory , val-n) → val-n x (λ _ _ → id , id)))) - ( proof-lower-bound) - ( proof-upper-bound)) - - helper : - is-in-subtype (transitive-kripke-models l1 l2 i l4) M → - (a : modal-formula i) → - is-in-subtype theory (□ₘ a) → - (x y : type-kripke-model i M) → - transitive-closure - ( relation-kripke-model i minimal-kripke-model-filtration) - ( class Φ-equivalence x) - ( class Φ-equivalence y) → - type-Prop ((M , x) ⊨ₘ □ₘ a) → - type-Prop ((M , y) ⊨ₘ a) - helper M-is-trans a box-in-theory x y - (transitive-closure-base r-xy) x-forces-box = - proof-upper-bound a box-in-theory x y r-xy x-forces-box - helper M-is-trans a box-in-theory x y - (transitive-closure-step {y = z*} r-xz c-zy) - x-forces-box = - apply-universal-property-trunc-Prop - ( is-inhabited-subtype-equivalence-class Φ-equivalence z*) - ( (M , y) ⊨ₘ a) - ( λ (z , z-in-z*) → - aux z (eq-class-equivalence-class Φ-equivalence z* z-in-z*)) - where - aux : - (z : type-kripke-model i M) → - class Φ-equivalence z = z* → - type-Prop ((M , y) ⊨ₘ a) - aux z refl = - apply-universal-property-trunc-Prop - ( r-xz) - ( (M , y) ⊨ₘ a) - ( λ ((x' , z') , r-xz' , iff-x , iff-z) → - ( helper M-is-trans a box-in-theory z y c-zy - ( backward-implication - ( iff-z (□ₘ a) box-in-theory) - ( ax-4-soundness i l2 l4 _ (a , refl) M M-is-trans - ( x') - ( forward-implication - ( iff-x (□ₘ a) box-in-theory) - ( x-forces-box)) - ( z') - ( r-xz'))))) - - filtration-relation-upper-bound-Prop-minimal-transitive-kripke-model-filtration : - is-in-subtype (transitive-kripke-models l1 l2 i l4) M → - filtration-relation-upper-bound - minimal-transitive-kripke-model-filtration id-equiv - filtration-relation-upper-bound-Prop-minimal-transitive-kripke-model-filtration - M-is-trans a box-in-theory x y r-xy x-forces-box = - apply-universal-property-trunc-Prop - ( r-xy) - ( (M , y) ⊨ₘ a) - ( λ r → helper M-is-trans a box-in-theory x y r x-forces-box) - - is-kripke-model-filtration-minimal-transitive-kripke-model-filtration : - is-in-subtype (transitive-kripke-models l1 l2 i l4) M → - is-kripke-model-filtration minimal-transitive-kripke-model-filtration - is-kripke-model-filtration-minimal-transitive-kripke-model-filtration - M-is-trans = - pair id-equiv - ( triple - ( λ n in-theory x → - ( pair - ( λ val-n → - ( pair - ( in-theory) - ( λ y eq-xy → - ( map-inv-raise - ( forward-implication - ( eq-xy (varₘ n) in-theory) - ( map-raise val-n)))))) - ( λ (in-theory , val-n) → val-n x (λ _ _ → id , id)))) - ( λ x y r → - unit-trunc-Prop - ( transitive-closure-base (proof-lower-bound x y r))) - ( filtration-relation-upper-bound-Prop-minimal-transitive-kripke-model-filtration - ( M-is-trans))) - - minimal-filtration-preserves-reflexivity : - is-in-subtype (reflexive-kripke-models l1 l2 i l4) M → - is-in-subtype - ( reflexive-kripke-models - ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) - ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) - ( i) - ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) - ( minimal-kripke-model-filtration) - minimal-filtration-preserves-reflexivity = - filtration-preserves-reflexivity - ( minimal-kripke-model-filtration) - ( is-kripke-model-filtration-minimal-kripke-model-filtration) - - minimal-kripke-model-filtration-relation-preserves-symmetry : - is-in-subtype (symmetry-kripke-models l1 l2 i l4) M → - is-symmetric-Relation-Prop minimal-kripke-model-filtration-relation - minimal-kripke-model-filtration-relation-preserves-symmetry is-sym x* y* = - map-universal-property-trunc-Prop - ( relation-Prop-kripke-model i minimal-kripke-model-filtration y* x*) - ( λ ((x , y) , r-xy , x-in-x* , y-in-y*) → - ( intro-exists (y , x) (is-sym x y r-xy , y-in-y* , x-in-x*))) - - minimal-filtration-preserves-symmetry : - is-in-subtype (symmetry-kripke-models l1 l2 i l4) M → - is-in-subtype - ( symmetry-kripke-models - ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) - ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) - ( i) - ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) - ( minimal-kripke-model-filtration) - minimal-filtration-preserves-symmetry = - minimal-kripke-model-filtration-relation-preserves-symmetry - - minimal-transitive-filtration-preserves-reflexivity : - is-in-subtype (reflexive-kripke-models l1 l2 i l4) M → - is-in-subtype - ( reflexive-kripke-models - ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) - ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) - ( i) - ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) - ( minimal-transitive-kripke-model-filtration) - minimal-transitive-filtration-preserves-reflexivity is-refl = - transitive-closure-Prop-preserves-reflexivity - ( minimal-kripke-model-filtration-relation) - ( minimal-filtration-preserves-reflexivity is-refl) - - minimal-transitive-kripke-model-filtration-preserves-symmetry : - is-in-subtype (symmetry-kripke-models l1 l2 i l4) M → - is-in-subtype - ( symmetry-kripke-models - ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) - ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) - ( i) - ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) - ( minimal-transitive-kripke-model-filtration) - minimal-transitive-kripke-model-filtration-preserves-symmetry is-sym = - transitive-closure-Prop-preserves-symmetry - ( minimal-kripke-model-filtration-relation) - ( minimal-filtration-preserves-symmetry is-sym) - - minimal-transitive-kripke-model-filtration-is-transitive : - is-in-subtype - ( transitive-kripke-models - ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) - ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) - ( i) - ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) - ( minimal-transitive-kripke-model-filtration) - minimal-transitive-kripke-model-filtration-is-transitive = - is-transitive-transitive-closure-Prop - ( minimal-kripke-model-filtration-relation) - - minimal-transitive-filtration-preserves-equivalence : - is-in-subtype (equivalence-kripke-models l1 l2 i l4) M → - is-in-subtype - ( equivalence-kripke-models - ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) - ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) - ( i) - ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) - ( minimal-transitive-kripke-model-filtration) - minimal-transitive-filtration-preserves-equivalence - (is-refl , is-sym , _) = - triple - ( minimal-transitive-filtration-preserves-reflexivity is-refl) - ( minimal-transitive-kripke-model-filtration-preserves-symmetry - ( is-sym)) - ( minimal-transitive-kripke-model-filtration-is-transitive) - module _ - (l1 l2 : Level) - {l3 : Level} (i : Set l3) - (l4 : Level) - (l5 l6 l7 l8 : Level) + {l1 l2 l3 l4 l5 : Level} (i : Set l3) + (M : kripke-model l1 l2 i l4) + (lem : LEM (lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5))) + (theory : modal-theory l5 i) + (is-fin : is-finite (type-subtype theory)) where - filtration-models : - model-class l1 l2 i l4 - (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4 ⊔ lsuc l5 ⊔ lsuc l6 ⊔ lsuc l7 ⊔ lsuc l8) - filtration-models M* = - exists-structure-Prop - ( modal-theory l5 i × kripke-model l6 l7 i l8) - ( λ (theory , M) → - ( product - ( is-finite (type-subtype theory)) - ( is-kripke-model-filtration theory M M*))) + is-finite-equivalence-classes-filtration : + is-finite (equivalence-class (Φ-equivalence theory M)) + is-finite-equivalence-classes-filtration = + is-finite-injection lem + ( injection-map-function-equivalence-class theory M) + ( is-finite-function-type is-fin (is-finite-Prop-LEM (lower-LEM _ lem))) ``` diff --git a/src/modal-logic/minimal-kripke-filtration.lagda.md b/src/modal-logic/minimal-kripke-filtration.lagda.md new file mode 100644 index 0000000000..d8ee284138 --- /dev/null +++ b/src/modal-logic/minimal-kripke-filtration.lagda.md @@ -0,0 +1,163 @@ +# Minimal Kripke filtration + +```agda +module modal-logic.minimal-kripke-filtration where +``` + +
Imports + +```agda +open import foundation.dependent-pair-types +open import foundation.equivalence-classes +open import foundation.existential-quantification +open import foundation.inhabited-types +open import foundation.logical-equivalences +open import foundation.raising-universe-levels +open import foundation.universe-levels + +open import foundation-core.cartesian-product-types +open import foundation-core.equivalences +open import foundation-core.propositions +open import foundation-core.sets +open import foundation-core.subtypes + +open import modal-logic.closed-under-subformulas-theories +open import modal-logic.deduction +open import modal-logic.formulas +open import modal-logic.kripke-models-filtrations +open import modal-logic.kripke-semantics +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l1 l2 l3 l4 l5 : Level} {i : Set l3} + (theory : modal-theory l5 i) + (M : kripke-model l1 l2 i l4) + where + + minimal-kripke-model-filtration : + kripke-model + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) + ( i) + ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) + pr1 (pr1 minimal-kripke-model-filtration) = + equivalence-class (Φ-equivalence theory M) + pr2 (pr1 minimal-kripke-model-filtration) = + map-is-inhabited + ( class (Φ-equivalence theory M)) + ( is-inhabited-type-kripke-model i M) + pr1 (pr2 minimal-kripke-model-filtration) x* y* = + exists-structure-Prop ( type-kripke-model i M × type-kripke-model i M) + ( λ (x , y) → + relation-kripke-model i M x y × + is-in-equivalence-class (Φ-equivalence theory M) x* x × + is-in-equivalence-class (Φ-equivalence theory M) y* y) + pr2 (pr2 minimal-kripke-model-filtration) n x* = + function-Prop + ( is-in-subtype theory (varₘ n)) + ( Π-Prop (type-kripke-model i M) + ( λ x → + function-Prop + ( is-in-equivalence-class (Φ-equivalence theory M) x* x) + ( valuate-kripke-model i M n x))) +``` + +## Properties + +### Minimal kripke filtration is a filtration function + +```agda + module _ + (theory-is-closed : is-modal-theory-closed-under-subformulas theory) + where + + is-kripke-model-filtration-minimal-kripke-model-filtration : + is-kripke-model-filtration theory M minimal-kripke-model-filtration + pr1 is-kripke-model-filtration-minimal-kripke-model-filtration = + id-equiv + pr1 (pr2 is-kripke-model-filtration-minimal-kripke-model-filtration) + n in-theory x = + pair + ( λ val-n in-theory y eq-xy → + map-inv-raise + ( forward-implication + ( eq-xy (varₘ n) in-theory) + ( map-raise val-n))) + ( λ val-n → val-n in-theory x (λ _ _ → id-iff)) + pr1 (pr2 (pr2 is-kripke-model-filtration-minimal-kripke-model-filtration)) + x y r = + intro-exists (x , y) (r , (λ _ _ → id-iff) , (λ _ _ → id-iff)) + pr2 (pr2 (pr2 is-kripke-model-filtration-minimal-kripke-model-filtration)) + a box-in-theory x y r-xy x-forces-box = + elim-exists + ( (M , y) ⊨ₘ a) + ( λ (x' , y') (r-xy' , iff-x , iff-y) → + backward-implication + ( iff-y a + ( is-has-subboxes-is-closed-under-subformulas + ( theory) + ( theory-is-closed) + ( box-in-theory))) + ( forward-implication + ( iff-x (□ₘ a) box-in-theory) + ( x-forces-box) + ( y') + ( r-xy'))) + ( r-xy) +``` + +### Minimal kripke filtration preserves reflexivity + +```agda + minimal-filtration-preserves-reflexivity : + is-in-subtype (reflexive-kripke-models l1 l2 i l4) M → + is-in-subtype + ( reflexive-kripke-models + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) + ( i) + ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( minimal-kripke-model-filtration) + minimal-filtration-preserves-reflexivity = + filtration-preserves-reflexivity theory M + ( minimal-kripke-model-filtration) + ( is-kripke-model-filtration-minimal-kripke-model-filtration) +``` + +### Minimal kripke filtration preserves symmetry + +```agda + minimal-filtration-preserves-symmetry : + is-in-subtype (symmetry-kripke-models l1 l2 i l4) M → + is-in-subtype + ( symmetry-kripke-models + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) + ( i) + ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( minimal-kripke-model-filtration) + minimal-filtration-preserves-symmetry is-sym x* y* = + elim-exists + ( relation-Prop-kripke-model i minimal-kripke-model-filtration y* x*) + ( λ (x , y) (r-xy , x-in-x* , y-in-y*) → + intro-exists (y , x) (is-sym x y r-xy , y-in-y* , x-in-x*)) + +-- module _ +-- {l1 l2 l3 l4 l5 : Level} {i : Set l3} +-- where + +-- test : +-- is-class-filtration-function i l2 l3 l4 l5 (all-models _ _ i _) +-- ( minimal-kripke-model-filtration) +-- test theory is-closed (M , _) = +-- is-kripke-model-filtration-minimal-kripke-model-filtration theory M is-closed +``` diff --git a/src/modal-logic/minimal-transitive-kripke-filtration.lagda.md b/src/modal-logic/minimal-transitive-kripke-filtration.lagda.md new file mode 100644 index 0000000000..3ef2311931 --- /dev/null +++ b/src/modal-logic/minimal-transitive-kripke-filtration.lagda.md @@ -0,0 +1,233 @@ +# Minimal transitive Kripke filtration + +```agda +module modal-logic.minimal-transitive-kripke-filtration where +``` + +
Imports + +```agda +open import foundation.binary-relations-transitive-closures +open import foundation.dependent-pair-types +open import foundation.equivalence-classes +open import foundation.logical-equivalences +open import foundation.propositional-truncations +open import foundation.unit-type +open import foundation.universe-levels + +open import foundation-core.identity-types +open import foundation-core.propositions +open import foundation-core.sets +open import foundation-core.subtypes + +open import modal-logic.axioms +open import modal-logic.closed-under-subformulas-theories +open import modal-logic.deduction +open import modal-logic.formulas +open import modal-logic.kripke-models-filtrations +open import modal-logic.kripke-semantics +open import modal-logic.minimal-kripke-filtration +``` + +
+ +## Idea + +TODO + +## Definition + +```agda +module _ + {l1 l2 l3 l4 l5 : Level} {i : Set l3} + (theory : modal-theory l5 i) + (M : kripke-model l1 l2 i l4) + where + + minimal-transitive-kripke-model-filtration : + kripke-model + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( i) + ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5) + pr1 minimal-transitive-kripke-model-filtration = + Inhabited-Type-kripke-model i (minimal-kripke-model-filtration theory M) + pr1 (pr2 minimal-transitive-kripke-model-filtration) = + transitive-closure-Prop + ( relation-Prop-kripke-model i (minimal-kripke-model-filtration theory M)) + pr2 (pr2 minimal-transitive-kripke-model-filtration) = + valuate-kripke-model i (minimal-kripke-model-filtration theory M) +``` + +## Properties + +### Minimal transitive kripke filtration is transitive + +```agda + minimal-transitive-filtration-is-transitive : + is-in-subtype + ( transitive-kripke-models + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( i) + ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( minimal-transitive-kripke-model-filtration) + minimal-transitive-filtration-is-transitive = + is-transitive-transitive-closure-Prop + ( relation-Prop-kripke-model i + ( minimal-kripke-model-filtration theory M)) +``` + +### Minimal transitive filtration is a filtration function + +```agda + module _ + (theory-is-closed : is-modal-theory-closed-under-subformulas theory) + where + + is-filtration-minimal-transitive-kripke-model-filtration : + is-in-subtype (transitive-kripke-models l1 l2 i l4) M → + is-kripke-model-filtration theory M + ( minimal-transitive-kripke-model-filtration) + pr1 (is-filtration-minimal-transitive-kripke-model-filtration M-is-trans) = + equiv-is-kripke-model-filtration theory M + ( minimal-kripke-model-filtration theory M) + ( is-kripke-model-filtration-minimal-kripke-model-filtration theory M + ( theory-is-closed)) + pr1 (pr2 + (is-filtration-minimal-transitive-kripke-model-filtration M-is-trans)) = + is-filtration-valuate-is-kripke-model-filtration theory M + ( minimal-kripke-model-filtration theory M) + ( is-kripke-model-filtration-minimal-kripke-model-filtration theory M + ( theory-is-closed)) + pr1 (pr2 (pr2 + (is-filtration-minimal-transitive-kripke-model-filtration M-is-trans))) + x y r = + -- TODO: Refactor + unit-trunc-Prop + ( transitive-closure-base + ( filtration-relation-lower-bound-is-kripke-model-filtration + ( theory) + ( M) + ( minimal-kripke-model-filtration theory M) + ( is-kripke-model-filtration-minimal-kripke-model-filtration + ( theory) + ( M) + ( theory-is-closed)) + ( x) + ( y) + ( r))) + pr2 (pr2 (pr2 + (is-filtration-minimal-transitive-kripke-model-filtration M-is-trans))) + a box-in-theory x y r-xy x-forces-box = + apply-universal-property-trunc-Prop + ( r-xy) + ( (M , y) ⊨ₘ a) + ( λ r → α x r x-forces-box) + where + α : + (x : type-kripke-model i M) → + transitive-closure + ( relation-kripke-model i (minimal-kripke-model-filtration theory M)) + ( class (Φ-equivalence theory M) x) + ( class (Φ-equivalence theory M) y) → + type-Prop ((M , x) ⊨ₘ □ₘ a) → + type-Prop ((M , y) ⊨ₘ a) + α x (transitive-closure-base r) x-forces-box = + filtration-relation-upper-bound-is-kripke-model-filtration theory M + ( minimal-kripke-model-filtration theory M) + ( is-kripke-model-filtration-minimal-kripke-model-filtration theory M + ( theory-is-closed)) + ( a) + ( box-in-theory) + ( x) + ( y) + ( r) + ( x-forces-box) + α x (transitive-closure-step {y = z*} r-xz c-zy) x-forces-box = + apply-universal-property-trunc-Prop + ( is-inhabited-subtype-equivalence-class (Φ-equivalence theory M) z*) + ( (M , y) ⊨ₘ a) + ( λ (z , z-in-z*) → + β z + (eq-class-equivalence-class (Φ-equivalence theory M) z* z-in-z*)) + where + β : + (z : type-kripke-model i M) → + class (Φ-equivalence theory M) z = z* → + type-Prop ((M , y) ⊨ₘ a) + β z refl = + apply-universal-property-trunc-Prop + ( r-xz) + ( (M , y) ⊨ₘ a) + ( λ ((x' , z') , r-xz' , iff-x , iff-z) → + α z c-zy + ( backward-implication + ( iff-z (□ₘ a) box-in-theory) + ( ax-4-soundness i _ _ _ (a , refl) M M-is-trans + ( x') + ( forward-implication + ( iff-x (□ₘ a) box-in-theory) + ( x-forces-box)) + ( z') + ( r-xz')))) +``` + +### Minimal transitive kripke filtration preserves symmetry + +```agda + minimal-transitive-filtration-preserves-reflexivity : + is-in-subtype (reflexive-kripke-models l1 l2 i l4) M → + is-in-subtype + ( reflexive-kripke-models + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( i) + ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( minimal-transitive-kripke-model-filtration) + minimal-transitive-filtration-preserves-reflexivity is-refl = + transitive-closure-Prop-preserves-reflexivity + ( relation-Prop-kripke-model i + ( minimal-kripke-model-filtration theory M)) + ( minimal-filtration-preserves-reflexivity theory M theory-is-closed + ( is-refl)) +``` + +### Minimal transitive kripke filtration preserves symmetry + +```agda + minimal-transitive-filtration-preserves-symmetry : + is-in-subtype (symmetry-kripke-models l1 l2 i l4) M → + is-in-subtype + ( symmetry-kripke-models + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( i) + ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( minimal-transitive-kripke-model-filtration) + minimal-transitive-filtration-preserves-symmetry is-sym = + transitive-closure-Prop-preserves-symmetry + ( relation-Prop-kripke-model i + ( minimal-kripke-model-filtration theory M)) + ( minimal-filtration-preserves-symmetry theory M theory-is-closed + ( is-sym)) +``` + +### Minimal transitive kripke filtration preserves equivalence + +```agda + minimal-transitive-filtration-preserves-equivalence : + is-in-subtype (equivalence-kripke-models l1 l2 i l4) M → + is-in-subtype + ( equivalence-kripke-models + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( lsuc (l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( i) + ( l1 ⊔ l2 ⊔ l3 ⊔ l4 ⊔ l5)) + ( minimal-transitive-kripke-model-filtration) + minimal-transitive-filtration-preserves-equivalence (is-refl , is-sym , _) = + triple + ( minimal-transitive-filtration-preserves-reflexivity is-refl) + ( minimal-transitive-filtration-preserves-symmetry is-sym) + ( minimal-transitive-filtration-is-transitive) +``` diff --git a/src/modal-logic/subformulas.lagda.md b/src/modal-logic/subformulas.lagda.md index 737a786848..283bd22de1 100644 --- a/src/modal-logic/subformulas.lagda.md +++ b/src/modal-logic/subformulas.lagda.md @@ -9,12 +9,14 @@ module modal-logic.subformulas where ```agda open import foundation.dependent-pair-types open import foundation.functoriality-propositional-truncation +open import foundation.law-of-excluded-middle open import foundation.propositional-truncations open import foundation.universe-levels open import foundation-core.cartesian-product-types open import foundation-core.coproduct-types open import foundation-core.function-types +open import foundation-core.propositions open import foundation-core.sets open import foundation-core.subtypes @@ -26,6 +28,8 @@ open import modal-logic.closed-under-subformulas-theories open import modal-logic.deduction open import modal-logic.formulas +open import univalent-combinatorics.finite-types +open import univalent-combinatorics.function-types open import univalent-combinatorics.kuratowsky-finite-sets ``` @@ -53,7 +57,11 @@ module _ subformulas : modal-formula i → modal-theory l i subformulas a = list-subtype (subformulas-list a) +``` + +### Subformulas is closed under subformulas +```agda subformulas-list-has-subimpl : (a : modal-formula i) {x y : modal-formula i} → (x →ₘ y) ∈-list subformulas-list a → @@ -144,3 +152,29 @@ module _ is-kuratowsky-finite-subformulas-list a = is-kuratowski-finite-list-subtype (modal-formula-Set i) (subformulas-list a) ``` + +TODO: Remove + +### TODO: Classical facts + +```agda + is-finite-subformulas-list : + LEM l → + (a : modal-formula i) → + is-finite (type-subtype (subformulas a)) + is-finite-subformulas-list lem a = + is-finite-is-kuratowsky-finite-set + ( set-subset (modal-formula-Set i) (subformulas a)) + ( lem) + ( is-kuratowsky-finite-subformulas-list a) + + is-finite-subtypes-subformulas-list : + {l2 : Level} → + LEM (l ⊔ l2) → + (a : modal-formula i) → + is-finite (type-subtype (subformulas a) → Prop l2) + is-finite-subtypes-subformulas-list {l2} lem a = + is-finite-function-type + ( is-finite-subformulas-list (lower-LEM l2 lem) a) + ( is-finite-Prop-LEM (lower-LEM l lem)) +``` diff --git a/src/univalent-combinatorics.lagda.md b/src/univalent-combinatorics.lagda.md index 0cbecf9212..08dc38f9f3 100644 --- a/src/univalent-combinatorics.lagda.md +++ b/src/univalent-combinatorics.lagda.md @@ -98,6 +98,7 @@ open import univalent-combinatorics.standard-finite-trees public open import univalent-combinatorics.standard-finite-types public open import univalent-combinatorics.steiner-systems public open import univalent-combinatorics.steiner-triple-systems public +open import univalent-combinatorics.subfinite-types public open import univalent-combinatorics.sums-of-natural-numbers public open import univalent-combinatorics.surjective-maps public open import univalent-combinatorics.symmetric-difference public diff --git a/src/univalent-combinatorics/kuratowsky-finite-sets.lagda.md b/src/univalent-combinatorics/kuratowsky-finite-sets.lagda.md index 793e834905..c89b708ef4 100644 --- a/src/univalent-combinatorics/kuratowsky-finite-sets.lagda.md +++ b/src/univalent-combinatorics/kuratowsky-finite-sets.lagda.md @@ -209,4 +209,15 @@ is-finite-is-kuratowsky-finite-set X lem is-fin = is-finite-has-decidable-equality-type-𝔽-Kuratowsky ( X , is-fin) ( λ x y → lem (Id-Prop X x y)) + +is-finite-surjection : + {l1 l2 : Level} (X : Set l1) (Y : Set l2) → + LEM l2 → + type-Set X ↠ type-Set Y → + is-finite (type-Set X) → + is-finite (type-Set Y) +is-finite-surjection X Y lem f is-fin = + is-finite-is-kuratowsky-finite-set Y lem + ( is-kuratowsky-finite-set-surjection X Y f + ( is-kuratowsky-finite-set-is-finite X is-fin)) ``` diff --git a/src/univalent-combinatorics/subfinite-types.lagda.md b/src/univalent-combinatorics/subfinite-types.lagda.md new file mode 100644 index 0000000000..bc071e76fd --- /dev/null +++ b/src/univalent-combinatorics/subfinite-types.lagda.md @@ -0,0 +1,120 @@ +# Subfinite types + +```agda +module univalent-combinatorics.subfinite-types where +``` + +
Imports + +```agda +open import elementary-number-theory.natural-numbers + +open import foundation.dependent-pair-types +open import foundation.existential-quantification +open import foundation.inhabited-types +open import foundation.injective-maps +open import foundation.law-of-excluded-middle +open import foundation.propositional-truncations +open import foundation.sets +open import foundation.universe-levels + +open import foundation-core.coproduct-types +open import foundation-core.equivalences +open import foundation-core.function-types +open import foundation-core.propositions + +open import univalent-combinatorics.finite-types +open import univalent-combinatorics.kuratowsky-finite-sets +open import univalent-combinatorics.standard-finite-types +``` + +
+ +## Idea + +A subfinite type is a set `X` for which there exists an injection into `X` from +a standard finite type. In other words, the subfinite types are the subset of a +standard finite type. + +## Definition + +```agda +is-subfinite-Prop : {l : Level} → UU l → Prop l +is-subfinite-Prop X = + exists-structure-Prop ℕ (λ n → injection X (Fin n)) + +is-subfinite : {l : Level} → UU l → UU l +is-subfinite X = type-Prop (is-subfinite-Prop X) +``` + +### Subfinite types are sets + +```agda +is-set-is-subfinite : {l : Level} {X : UU l} → is-subfinite X → is-set X +is-set-is-subfinite {X = X} = + elim-exists + ( is-set-Prop X) + ( λ n (f , is-inj) → is-set-is-injective (is-set-Fin n) is-inj) + +Set-is-subfinite : {l : Level} {X : UU l} → is-subfinite X → Set l +Set-is-subfinite {X = X} is-sub = X , is-set-is-subfinite is-sub +``` + +### Subfinite sets are closed under injections + +```agda +is-subfinite-injection : + {l1 l2 : Level} {X : UU l1} {Y : UU l2} → + injection X Y → + is-subfinite Y → + is-subfinite X +is-subfinite-injection {X = X} f = + elim-exists + ( is-subfinite-Prop X) + ( λ n g → intro-exists n (injection-comp f g)) +``` + +### Any finite set is subfinite + +```agda +is-subfinite-is-finite : {l : Level} {X : UU l} → is-finite X → is-subfinite X +is-subfinite-is-finite {X = X} = + elim-exists + ( is-subfinite-Prop X) + ( λ n e → + intro-exists n + ( map-inv-equiv e , λ {_} {_} → is-injective-map-inv-equiv e)) +``` + +### Classical facts + +```agda +is-finite-is-subfinite : + {l : Level} {X : UU l} → LEM l → is-subfinite X → is-finite X +is-finite-is-subfinite {X = X} lem is-sub with lem (is-inhabited-Prop X) +... | inl is-inh = + elim-exists + ( is-finite-Prop X) + ( λ n (f , is-inj) → + ( apply-universal-property-trunc-Prop + ( is-inhabited-inv-surjections f is-inj + ( apply-LEM lem ∘ is-prop-map-is-injective (is-set-Fin n) is-inj) + ( is-inh)) + ( is-finite-Prop X) + ( λ g → + is-finite-is-kuratowsky-finite-set + ( Set-is-subfinite is-sub) + ( lem) + ( intro-exists n g)))) + ( is-sub) +... | inr not-inh = is-finite-is-empty (not-inh ∘ unit-trunc-Prop) + +is-finite-injection : + {l1 l2 : Level} {X : UU l1} {Y : UU l2} → + LEM l1 → + injection X Y → + is-finite Y → + is-finite X +is-finite-injection lem f = + is-finite-is-subfinite lem ∘ is-subfinite-injection f ∘ is-subfinite-is-finite +``` From ba358737e6a8d3d88925f3c0a3e1beb7395e6ad7 Mon Sep 17 00:00:00 2001 From: Viktor Yudov Date: Fri, 31 May 2024 21:05:07 +0300 Subject: [PATCH 63/63] Fix import in finite-approximability --- src/modal-logic/finite-approximability.lagda.md | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/src/modal-logic/finite-approximability.lagda.md b/src/modal-logic/finite-approximability.lagda.md index 7e20523fb6..b3179123e5 100644 --- a/src/modal-logic/finite-approximability.lagda.md +++ b/src/modal-logic/finite-approximability.lagda.md @@ -24,10 +24,12 @@ open import modal-logic.closed-under-subformulas-theories open import modal-logic.completeness open import modal-logic.completeness-k open import modal-logic.completeness-s5 -open import modal-logic.decision-procedure open import modal-logic.deduction +open import modal-logic.filtrated-kripke-classes open import modal-logic.kripke-models-filtrations open import modal-logic.kripke-semantics +open import modal-logic.minimal-kripke-filtration +open import modal-logic.minimal-transitive-kripke-filtration open import modal-logic.modal-logic-k open import modal-logic.modal-logic-s5 open import modal-logic.soundness